summaryrefslogtreecommitdiff
path: root/snapshots
diff options
context:
space:
mode:
Diffstat (limited to 'snapshots')
-rw-r--r--snapshots/coq-riscv/sail/lib/coq/Sail2_instr_kinds.v6
-rw-r--r--snapshots/coq-riscv/sail/lib/coq/Sail2_prompt.v2
-rw-r--r--snapshots/coq-riscv/sail/lib/coq/Sail2_string.v131
-rw-r--r--snapshots/coq-riscv/sail/lib/coq/Sail2_values.v77
-rw-r--r--snapshots/coq-riscv/sail/riscv/riscv.v30284
-rw-r--r--snapshots/coq-riscv/sail/riscv/riscv_extras.v74
-rw-r--r--snapshots/coq-riscv/sail/riscv/riscv_types.v53
7 files changed, 14033 insertions, 16594 deletions
diff --git a/snapshots/coq-riscv/sail/lib/coq/Sail2_instr_kinds.v b/snapshots/coq-riscv/sail/lib/coq/Sail2_instr_kinds.v
index c93e9e93..eadc567a 100644
--- a/snapshots/coq-riscv/sail/lib/coq/Sail2_instr_kinds.v
+++ b/snapshots/coq-riscv/sail/lib/coq/Sail2_instr_kinds.v
@@ -144,7 +144,7 @@ Inductive barrier_kind :=
(* AArch64 barriers *)
| Barrier_DMB | Barrier_DMB_ST | Barrier_DMB_LD | Barrier_DSB
| Barrier_DSB_ST | Barrier_DSB_LD | Barrier_ISB
- | Barrier_TM_COMMIT
+ (* | Barrier_TM_COMMIT*)
(* MIPS barriers *)
| Barrier_MIPS_SYNC
(* RISC-V barriers *)
@@ -204,11 +204,11 @@ Inductive instruction_kind :=
| IK_mem_read : read_kind -> instruction_kind
| IK_mem_write : write_kind -> instruction_kind
| IK_mem_rmw : (read_kind * write_kind) -> instruction_kind
- | IK_branch (* this includes conditional-branch (multiple nias, none of which is NIA_indirect_address),
+ | IK_branch : unit -> instruction_kind (* this includes conditional-branch (multiple nias, none of which is NIA_indirect_address),
indirect/computed-branch (single nia of kind NIA_indirect_address)
and branch/jump (single nia of kind NIA_concrete_address) *)
| IK_trans : trans_kind -> instruction_kind
- | IK_simple : instruction_kind.
+ | IK_simple : unit -> instruction_kind.
(*
instance (Show instruction_kind)
diff --git a/snapshots/coq-riscv/sail/lib/coq/Sail2_prompt.v b/snapshots/coq-riscv/sail/lib/coq/Sail2_prompt.v
index 0b3a2cd8..a0ef467e 100644
--- a/snapshots/coq-riscv/sail/lib/coq/Sail2_prompt.v
+++ b/snapshots/coq-riscv/sail/lib/coq/Sail2_prompt.v
@@ -125,3 +125,5 @@ Definition build_ex_m {rv e} {T:Type} (x:monad rv T e) {P:T -> Prop} `{H:forall
Definition projT1_m {rv e} {P:Z -> Prop} (x: monad rv {x : Z & P x} e) : monad rv Z e :=
x >>= fun y => returnm (projT1 y).
+Definition derive_m {rv e} {P Q:Z -> Prop} (x : monad rv {x : Z & P x} e) `{forall x, ArithFact (P x) -> ArithFact (Q x)} : monad rv {x : Z & (ArithFact (Q x))} e :=
+ x >>= fun y => returnm (build_ex (projT1 y)).
diff --git a/snapshots/coq-riscv/sail/lib/coq/Sail2_string.v b/snapshots/coq-riscv/sail/lib/coq/Sail2_string.v
index 9ca9cb67..c8bf5f9f 100644
--- a/snapshots/coq-riscv/sail/lib/coq/Sail2_string.v
+++ b/snapshots/coq-riscv/sail/lib/coq/Sail2_string.v
@@ -16,7 +16,82 @@ Definition string_length s : {n : Z & ArithFact (n >= 0)} :=
Definition string_append := String.append.
-(* TODO: maybe_int_of_prefix, maybe_int_of_string *)
+Local Open Scope char_scope.
+Local Definition hex_char (c : Ascii.ascii) : option Z :=
+match c with
+| "0" => Some 0
+| "1" => Some 1
+| "2" => Some 2
+| "3" => Some 3
+| "4" => Some 4
+| "5" => Some 5
+| "6" => Some 6
+| "7" => Some 7
+| "8" => Some 8
+| "9" => Some 9
+| "a" => Some 10
+| "b" => Some 11
+| "c" => Some 12
+| "d" => Some 13
+| "e" => Some 14
+| "f" => Some 15
+| _ => None
+end.
+Local Close Scope char_scope.
+Local Fixpoint more_digits (s : string) (base : Z) (acc : Z) (len : nat) : Z * nat :=
+match s with
+| EmptyString => (acc, len)
+| String "_" t => more_digits t base acc (S len)
+| String h t =>
+ match hex_char h with
+ | None => (acc, len)
+ | Some i =>
+ if i <? base
+ then more_digits t base (base * acc + i) (S len)
+ else (acc, len)
+ end
+end.
+Local Definition int_of (s : string) (base : Z) (len : nat) : option (Z * {n : Z & ArithFact (n >= 0)}) :=
+match s with
+| EmptyString => None
+| String h t =>
+ match hex_char h with
+ | None => None
+ | Some i =>
+ if i <? base
+ then
+ let (i, len') := more_digits t base i (S len) in
+ Some (i, build_ex (Z.of_nat len'))
+ else None
+ end
+end.
+
+(* I've stuck closely to OCaml's int_of_string, because that's what's currently
+ used elsewhere. *)
+
+Definition maybe_int_of_prefix (s : string) : option (Z * {n : Z & ArithFact (n >= 0)}) :=
+match s with
+| EmptyString => None
+| String "0" (String ("x"|"X") t) => int_of t 16 2
+| String "0" (String ("o"|"O") t) => int_of t 8 2
+| String "0" (String ("b"|"B") t) => int_of t 2 2
+| String "0" (String "u" t) => int_of t 10 2
+| String "-" t =>
+ match int_of t 10 1 with
+ | None => None
+ | Some (i,len) => Some (-i,len)
+ end
+| _ => int_of s 10 0
+end.
+
+Definition maybe_int_of_string (s : string) : option Z :=
+match maybe_int_of_prefix s with
+| None => None
+| Some (i,len) =>
+ if projT1 len =? projT1 (string_length s)
+ then Some i
+ else None
+end.
Fixpoint n_leading_spaces (s:string) : nat :=
match s with
@@ -32,4 +107,56 @@ Definition spc_matches_prefix s : option (unit * {n : Z & ArithFact (n >= 0)}) :
match n_leading_spaces s with
| O => None
| S n => Some (tt, build_ex (Z.of_nat (S n)))
- end. \ No newline at end of file
+ end.
+
+Definition hex_bits_n_matches_prefix sz `{ArithFact (sz >= 0)} s : option (mword sz * {n : Z & ArithFact (n >= 0)}) :=
+ match maybe_int_of_prefix s with
+ | None => None
+ | Some (n, len) =>
+ if andb (0 <=? n) (n <? pow 2 sz)
+ then Some (of_int sz n, len)
+ else None
+ end.
+
+Definition hex_bits_3_matches_prefix s := hex_bits_n_matches_prefix 3 s.
+Definition hex_bits_4_matches_prefix s := hex_bits_n_matches_prefix 4 s.
+Definition hex_bits_5_matches_prefix s := hex_bits_n_matches_prefix 5 s.
+Definition hex_bits_6_matches_prefix s := hex_bits_n_matches_prefix 6 s.
+Definition hex_bits_7_matches_prefix s := hex_bits_n_matches_prefix 7 s.
+Definition hex_bits_8_matches_prefix s := hex_bits_n_matches_prefix 8 s.
+Definition hex_bits_9_matches_prefix s := hex_bits_n_matches_prefix 9 s.
+Definition hex_bits_10_matches_prefix s := hex_bits_n_matches_prefix 10 s.
+Definition hex_bits_11_matches_prefix s := hex_bits_n_matches_prefix 11 s.
+Definition hex_bits_12_matches_prefix s := hex_bits_n_matches_prefix 12 s.
+Definition hex_bits_13_matches_prefix s := hex_bits_n_matches_prefix 13 s.
+Definition hex_bits_14_matches_prefix s := hex_bits_n_matches_prefix 14 s.
+Definition hex_bits_15_matches_prefix s := hex_bits_n_matches_prefix 15 s.
+Definition hex_bits_16_matches_prefix s := hex_bits_n_matches_prefix 16 s.
+Definition hex_bits_20_matches_prefix s := hex_bits_n_matches_prefix 20 s.
+Definition hex_bits_21_matches_prefix s := hex_bits_n_matches_prefix 21 s.
+Definition hex_bits_28_matches_prefix s := hex_bits_n_matches_prefix 28 s.
+Definition hex_bits_32_matches_prefix s := hex_bits_n_matches_prefix 32 s.
+Definition hex_bits_33_matches_prefix s := hex_bits_n_matches_prefix 33 s.
+
+Local Definition zero : N := Ascii.N_of_ascii "0".
+Local Fixpoint string_of_N (limit : nat) (n : N) (acc : string) : string :=
+match limit with
+| O => acc
+| S limit' =>
+ let (d,m) := N.div_eucl n 10 in
+ let acc := String (Ascii.ascii_of_N (m + zero)) acc in
+ if N.ltb 0 d then string_of_N limit' d acc else acc
+end.
+Local Fixpoint pos_limit p :=
+match p with
+| xH => S O
+| xI p | xO p => S (pos_limit p)
+end.
+Definition string_of_int (z : Z) : string :=
+match z with
+| Z0 => "0"
+| Zpos p => string_of_N (pos_limit p) (Npos p) ""
+| Zneg p => String "-" (string_of_N (pos_limit p) (Npos p) "")
+end.
+
+
diff --git a/snapshots/coq-riscv/sail/lib/coq/Sail2_values.v b/snapshots/coq-riscv/sail/lib/coq/Sail2_values.v
index def6e248..eaf75840 100644
--- a/snapshots/coq-riscv/sail/lib/coq/Sail2_values.v
+++ b/snapshots/coq-riscv/sail/lib/coq/Sail2_values.v
@@ -61,14 +61,14 @@ Instance Decidable_eq_string : forall (x y : string), Decidable (x = y) :=
(* Project away range constraints in comparisons *)
-Definition ltb_range_l {P} (l : sigT P) r := Z.ltb (projT1 l) r.
-Definition leb_range_l {P} (l : sigT P) r := Z.leb (projT1 l) r.
-Definition gtb_range_l {P} (l : sigT P) r := Z.gtb (projT1 l) r.
-Definition geb_range_l {P} (l : sigT P) r := Z.geb (projT1 l) r.
-Definition ltb_range_r {P} l (r : sigT P) := Z.ltb l (projT1 r).
-Definition leb_range_r {P} l (r : sigT P) := Z.leb l (projT1 r).
-Definition gtb_range_r {P} l (r : sigT P) := Z.gtb l (projT1 r).
-Definition geb_range_r {P} l (r : sigT P) := Z.geb l (projT1 r).
+Definition ltb_range_l {lo hi} (l : {x & ArithFact (lo <= x /\ x <= hi)}) r := Z.ltb (projT1 l) r.
+Definition leb_range_l {lo hi} (l : {x & ArithFact (lo <= x /\ x <= hi)}) r := Z.leb (projT1 l) r.
+Definition gtb_range_l {lo hi} (l : {x & ArithFact (lo <= x /\ x <= hi)}) r := Z.gtb (projT1 l) r.
+Definition geb_range_l {lo hi} (l : {x & ArithFact (lo <= x /\ x <= hi)}) r := Z.geb (projT1 l) r.
+Definition ltb_range_r {lo hi} l (r : {x & ArithFact (lo <= x /\ x <= hi)}) := Z.ltb l (projT1 r).
+Definition leb_range_r {lo hi} l (r : {x & ArithFact (lo <= x /\ x <= hi)}) := Z.leb l (projT1 r).
+Definition gtb_range_r {lo hi} l (r : {x & ArithFact (lo <= x /\ x <= hi)}) := Z.gtb l (projT1 r).
+Definition geb_range_r {lo hi} l (r : {x & ArithFact (lo <= x /\ x <= hi)}) := Z.geb l (projT1 r).
Definition ii := Z.
Definition nn := nat.
@@ -952,7 +952,7 @@ auto using Z.le_ge, Zle_0_pos.
destruct w.
Qed.
Ltac unwrap_ArithFacts :=
- repeat match goal with H:(ArithFact _) |- _ => let H' := fresh H in case H as [H'] end.
+ repeat match goal with H:(ArithFact _) |- _ => let H' := fresh H in case H as [H']; clear H end.
Ltac unbool_comparisons :=
repeat match goal with
| H:context [Z.geb _ _] |- _ => rewrite Z.geb_leb in H
@@ -977,15 +977,37 @@ Ltac unbool_comparisons :=
(* Split up dependent pairs to get at proofs of properties *)
Ltac extract_properties :=
+ (* Properties of local definitions *)
repeat match goal with H := (projT1 ?X) |- _ =>
let x := fresh "x" in
let Hx := fresh "Hx" in
destruct X as [x Hx] in *;
change (projT1 (existT _ x Hx)) with x in *; unfold H in * end;
+ (* Properties in the goal *)
repeat match goal with |- context [projT1 ?X] =>
let x := fresh "x" in
let Hx := fresh "Hx" in
destruct X as [x Hx] in *;
+ change (projT1 (existT _ x Hx)) with x in * end;
+ (* Properties with proofs embedded by build_ex; uses revert/generalize
+ rather than destruct because it seemed to be more efficient, but
+ some experimentation would be needed to be sure.
+ repeat (
+ match goal with H:context [@build_ex ?T ?n ?P ?prf] |- _ =>
+ let x := fresh "x" in
+ let zz := constr:(@build_ex T n P prf) in
+ revert dependent H(*; generalize zz; intros*)
+ end;
+ match goal with |- context [@build_ex ?T ?n ?P ?prf] =>
+ let x := fresh "x" in
+ let zz := constr:(@build_ex T n P prf) in
+ generalize zz as x
+ end;
+ intros).*)
+ repeat match goal with _:context [projT1 ?X] |- _ =>
+ let x := fresh "x" in
+ let Hx := fresh "Hx" in
+ destruct X as [x Hx] in *;
change (projT1 (existT _ x Hx)) with x in * end.
(* TODO: hyps, too? *)
Ltac reduce_list_lengths :=
@@ -1008,32 +1030,55 @@ Ltac dump_context :=
| H:=?X |- _ => idtac H ":=" X; fail
| H:?X |- _ => idtac H ":" X; fail end;
match goal with |- ?X => idtac "Goal:" X end.
+Ltac split_cases :=
+ repeat match goal with
+ |- context [match ?X with _ => _ end] => destruct X
+ end.
Ltac prepare_for_solver :=
(*dump_context;*)
clear_irrelevant_defns;
clear_non_Z_defns;
+ autounfold with sail in * |- *; (* You can add Hint Unfold ... : sail to let omega see through fns *)
+ split_cases;
extract_properties;
repeat match goal with w:mword ?n |- _ => apply ArithFact_mword in w end;
unwrap_ArithFacts;
unfold_In;
- autounfold with sail in * |- *; (* You can add Hint Unfold ... : sail to let omega see through fns *)
unbool_comparisons;
reduce_list_lengths;
reduce_pow.
+
+Lemma trivial_range {x : Z} : ArithFact (x <= x /\ x <= x).
+constructor.
+auto with zarith.
+Qed.
+
+Lemma ArithFact_self_proof {P} : forall x : {y : Z & ArithFact (P y)}, ArithFact (P (projT1 x)).
+intros [x H].
+exact H.
+Qed.
+
+Ltac fill_in_evar_eq :=
+ match goal with |- ArithFact (?x = ?y) =>
+ (is_evar x || is_evar y);
+ (* compute to allow projections to remove proofs that might not be allowed in the evar *)
+(* Disabled because cbn may reduce definitions, even after clearbody
+ let x := eval cbn in x in
+ let y := eval cbn in y in*)
+ idtac "Warning: unknown equality constraint"; constructor; exact (eq_refl _ : x = y) end.
+
Ltac solve_arithfact :=
(* Attempt a simple proof first to avoid lengthy preparation steps (especially
as the large proof terms can upset subsequent proofs). *)
+intros; (* To solve implications for derive_m *)
+try (exact trivial_range);
+try fill_in_evar_eq;
+try match goal with |- context [projT1 ?X] => apply (ArithFact_self_proof X) end;
try (constructor; omega);
prepare_for_solver;
(*dump_context;*)
solve
[ match goal with |- ArithFact (?x _) => is_evar x; idtac "Warning: unknown constraint"; constructor; exact (I : (fun _ => True) _) end
- | match goal with |- ArithFact (?x = ?y) =>
- (is_evar x || is_evar y);
- (* compute to allow projections to remove proofs that might not be allowed in the evar *)
- let x := eval cbn in x in
- let y := eval cbn in y in
- idtac "Warning: unknown equality constraint"; constructor; exact (eq_refl _ : x = y) end
| apply ArithFact_mword; assumption
| constructor; omega with Z
(* The datatypes hints give us some list handling, esp In *)
diff --git a/snapshots/coq-riscv/sail/riscv/riscv.v b/snapshots/coq-riscv/sail/riscv/riscv.v
index 86296adb..21350812 100644
--- a/snapshots/coq-riscv/sail/riscv/riscv.v
+++ b/snapshots/coq-riscv/sail/riscv/riscv.v
@@ -13,37 +13,47 @@ Open Scope string.
Open Scope bool.
Section Content.
-Definition spc_forwards '(tt : unit) : string := " ".
+Definition eq_unit (g__22 : unit) (g__23 : unit) : bool := true.
-Definition spc_backwards (s : string) : unit := tt.
-Definition opt_spc_forwards '(tt : unit) : string := "".
-Definition opt_spc_backwards (s : string) : unit := tt.
+Definition neq_range {n : Z} {m : Z} {o : Z} {p : Z} '(existT _ x _ : {rangevar : Z & ArithFact (n <=
+ rangevar /\
+ rangevar <= m)}) '(existT _ y _ : {rangevar : Z & ArithFact (o <= rangevar /\ rangevar <= p)})
+: bool :=
+ negb (eq_range (build_ex x) (build_ex y)).
-Definition def_spc_forwards '(tt : unit) : string := " ".
+Definition neq_int (x : Z) (y : Z) : bool := negb (Z.eqb x y).
-Definition def_spc_backwards (s : string) : unit := tt.
+Definition neq_bool (x : bool) (y : bool) : bool := negb (Bool.eqb x y).
Axiom eq_real : forall (_ : real) (_ : real) , bool.
-Axiom vcons : forall {n : Z} {a : Type} (_ : a) (_ : vec a n) , vec a (n + 1).
+Axiom lteq_real : forall (_ : real) (_ : real) , bool.
-Axiom vector_concat : forall {n : Z} {m : Z} {a : Type} (_ : vec a n) (_ : vec a m) , vec a (n + m).
+Axiom lt_real : forall (_ : real) (_ : real) , bool.
-Definition neq_atom (x : Z) (y : Z) : bool := negb (Z.eqb x y).
+Axiom gteq_real : forall (_ : real) (_ : real) , bool.
-Definition neq_int (x : Z) (y : Z) : bool := negb (Z.eqb x y).
+Axiom gt_real : forall (_ : real) (_ : real) , bool.
-Definition neq_vec {n : Z} (x : mword n) (y : mword n) : bool := negb (eq_vec x y).
+Definition spc_forwards '(tt : unit) : string := " ".
+Definition spc_backwards (s : string) : unit := tt.
+Definition opt_spc_forwards '(tt : unit) : string := "".
-Axiom builtin_and_vec : forall {n : Z} (_ : bits n) (_ : bits n) , bits n.
+Definition opt_spc_backwards (s : string) : unit := tt.
+Definition def_spc_forwards '(tt : unit) : string := " ".
+Definition def_spc_backwards (s : string) : unit := tt.
-Axiom builtin_or_vec : forall {n : Z} (_ : bits n) (_ : bits n) , bits n.
+Axiom vcons : forall {n : Z} {a : Type} (_ : a) (_ : vec a n) , vec a (n + 1).
+
+Axiom vector_concat : forall {n : Z} {m : Z} {a : Type} (_ : vec a n) (_ : vec a m) , vec a (n + m).
+
+Definition neq_vec {n : Z} (x : mword n) (y : mword n) : bool := negb (eq_vec x y).
@@ -74,8 +84,6 @@ Definition cast_unit_vec (b : bitU)
end)
: M (mword 1).
-Axiom string_of_int : forall (_ : Z) , string.
-
Axiom DecStr : forall (_ : Z) , string.
Axiom HexStr : forall (_ : Z) , string.
@@ -92,14 +100,6 @@ Axiom mult_real : forall (_ : real) (_ : real) , real.
Axiom Sqrt : forall (_ : real) , real.
-Axiom gteq_real : forall (_ : real) (_ : real) , bool.
-
-Axiom lteq_real : forall (_ : real) (_ : real) , bool.
-
-Axiom gt_real : forall (_ : real) (_ : real) , bool.
-
-Axiom lt_real : forall (_ : real) (_ : real) , bool.
-
Axiom RoundDown : forall (_ : real) , Z.
Axiom RoundUp : forall (_ : real) , Z.
@@ -159,7 +159,8 @@ Definition __RISCV_read (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (r
B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
B0]
: mword 64) addr >>= fun w__0 : mword (8 * width) =>
- returnm ((Some w__0)
+ returnm ((Some
+ (w__0))
: option (mword (8 * width)))
| (true, false, false) =>
read_ram 64 width
@@ -168,7 +169,8 @@ Definition __RISCV_read (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (r
B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
B0]
: mword 64) addr >>= fun w__1 : mword (8 * width) =>
- returnm ((Some w__1)
+ returnm ((Some
+ (w__1))
: option (mword (8 * width)))
| (true, true, false) =>
read_ram 64 width
@@ -177,7 +179,8 @@ Definition __RISCV_read (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (r
B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
B0]
: mword 64) addr >>= fun w__2 : mword (8 * width) =>
- returnm ((Some w__2)
+ returnm ((Some
+ (w__2))
: option (mword (8 * width)))
| (false, false, true) =>
read_ram 64 width
@@ -186,7 +189,8 @@ Definition __RISCV_read (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (r
B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
B0]
: mword 64) addr >>= fun w__3 : mword (8 * width) =>
- returnm ((Some w__3)
+ returnm ((Some
+ (w__3))
: option (mword (8 * width)))
| (true, false, true) =>
read_ram 64 width
@@ -195,7 +199,8 @@ Definition __RISCV_read (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (r
B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
B0]
: mword 64) addr >>= fun w__4 : mword (8 * width) =>
- returnm ((Some w__4)
+ returnm ((Some
+ (w__4))
: option (mword (8 * width)))
| (true, true, true) =>
read_ram 64 width
@@ -204,7 +209,8 @@ Definition __RISCV_read (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (r
B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
B0]
: mword 64) addr >>= fun w__5 : mword (8 * width) =>
- returnm ((Some w__5)
+ returnm ((Some
+ (w__5))
: option (mword (8 * width)))
| (false, true, false) => returnm (None : option (mword (8 * width)))
| (false, true, true) => returnm (None : option (mword (8 * width)))
@@ -213,16 +219,15 @@ Definition __RISCV_read (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (r
Axiom __TraceMemoryRead : forall {m : Z} (n : Z) (_ : bits m) (_ : bits (8 * n)) , unit.
-Definition ex_nat '((existT _ n _) : {n : Z & ArithFact (n >= 0)})
+Definition ex_nat '(existT _ n _ : {n : Z & ArithFact (n >= 0)})
: {syn_n : Z & ArithFact (syn_n >= 0)} :=
- build_ex
- n.
+ build_ex(n).
-Definition ex_int (n : Z) : {syn_n : Z & ArithFact (True)} := build_ex n.
+Definition ex_int (n : Z) : {syn_n : Z & ArithFact (True)} := build_ex(n).
Definition coerce_int_nat (x : Z)
: M ({n : Z & ArithFact (n >= 0)}) :=
- assert_exp' (x >=? 0) "" >>= fun _ => returnm ((build_ex x) : {n : Z & ArithFact (n >= 0)}).
+ assert_exp' (x >=? 0) "" >>= fun _ => returnm (build_ex (x : Z)).
Definition EXTS {n : Z} (m__tv : Z) (v : mword n) `{ArithFact (m__tv >= n)}
: mword m__tv :=
@@ -232,11 +237,11 @@ Definition EXTZ {n : Z} (m__tv : Z) (v : mword n) `{ArithFact (m__tv >= n)}
: mword m__tv :=
zero_extend v m__tv.
-Definition zopz0zI_s {n : Z} (x : mword n) (y : mword n) `{ArithFact (n >= 1)}
+Definition zopz0zI_s {n : Z} (x : mword n) (y : mword n) `{ArithFact (n >= (0 + 1))}
: bool :=
Z.ltb (projT1 (sint x)) (projT1 (sint y)).
-Definition zopz0zKzJ_s {n : Z} (x : mword n) (y : mword n) `{ArithFact (n >= 1)}
+Definition zopz0zKzJ_s {n : Z} (x : mword n) (y : mword n) `{ArithFact (n >= (0 + 1))}
: bool :=
Z.geb (projT1 (sint x)) (projT1 (sint y)).
@@ -286,31 +291,33 @@ Definition shift_right_arith32 (v : mword 32) (shift : mword 5)
let v64 : bits 64 := EXTS 64 v in
subrange_vec_dec (shift_bits_right v64 shift) 31 0.
-Let xlen := 64.
-Definition xlen_max_unsigned_spec := sub_range (build_ex (projT1 (pow2 xlen))) (build_ex 1).
-Let xlen_max_unsigned := projT1 xlen_max_unsigned_spec .
-
-Definition xlen_max_signed_spec :=
-sub_range (build_ex (projT1 (pow2 (projT1 (sub_range (build_ex xlen) (build_ex 1)))))) (build_ex 1).
-Let xlen_max_signed := projT1 xlen_max_signed_spec .
-
-Definition xlen_min_signed_spec :=
-sub_range (build_ex 0) (build_ex (projT1 (pow2 (projT1 (sub_range (build_ex xlen) (build_ex 1)))))).
-Let xlen_min_signed := projT1 xlen_min_signed_spec .
-
+Definition xlen := 64.
+Hint Unfold xlen : sail.
+Definition xlen_max_unsigned := projT1 (sub_range (build_ex (projT1 (pow2 xlen))) (build_ex 1)).
+Hint Unfold xlen_max_unsigned : sail.
+Definition xlen_max_signed :=
+projT1 (sub_range (build_ex (projT1 (pow2 (projT1 (sub_range (build_ex xlen) (build_ex 1))))))
+ (build_ex 1)).
+Hint Unfold xlen_max_signed : sail.
+Definition xlen_min_signed :=
+projT1 (sub_range (build_ex 0)
+ (build_ex (projT1 (pow2 (projT1 (sub_range (build_ex xlen) (build_ex 1))))))).
+Hint Unfold xlen_min_signed : sail.
Definition regbits_to_regno (b : mword 5)
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)} :=
- let '(existT _ r _) := uint b in
- build_ex
- r.
+ build_ex(let 'r := projT1 (uint b) in
+ r).
Definition creg2reg_bits (creg : mword 3)
: mword 5 :=
concat_vec (vec_of_bits [B0;B1] : mword 2) creg.
-Let zreg : regbits := (vec_of_bits [B0;B0;B0;B0;B0] : mword 5).
-Let ra : regbits := (vec_of_bits [B0;B0;B0;B0;B1] : mword 5).
-Let sp : regbits := (vec_of_bits [B0;B0;B0;B1;B0] : mword 5).
+Definition zreg : regbits := (vec_of_bits [B0;B0;B0;B0;B0] : mword 5).
+Hint Unfold zreg : sail.
+Definition ra : regbits := (vec_of_bits [B0;B0;B0;B0;B1] : mword 5).
+Hint Unfold ra : sail.
+Definition sp : regbits := (vec_of_bits [B0;B0;B0;B1;B0] : mword 5).
+Hint Unfold sp : sail.
Definition rX (r : Z) `{ArithFact (0 <= r /\ (r + 1) <= 32)}
: M (mword 64) :=
let p0_ := r in
@@ -321,37 +328,37 @@ Definition rX (r : Z) `{ArithFact (0 <= r /\ (r + 1) <= 32)}
B0]
: mword 64)
: mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 1)) then (read_reg x1_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 2)) then (read_reg x2_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 3)) then (read_reg x3_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 4)) then (read_reg x4_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 5)) then (read_reg x5_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 6)) then (read_reg x6_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 7)) then (read_reg x7_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 8)) then (read_reg x8_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 9)) then (read_reg x9_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 10)) then (read_reg x10_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 11)) then (read_reg x11_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 12)) then (read_reg x12_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 13)) then (read_reg x13_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 14)) then (read_reg x14_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 15)) then (read_reg x15_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 16)) then (read_reg x16_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 17)) then (read_reg x17_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 18)) then (read_reg x18_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 19)) then (read_reg x19_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 20)) then (read_reg x20_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 21)) then (read_reg x21_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 22)) then (read_reg x22_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 23)) then (read_reg x23_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 24)) then (read_reg x24_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 25)) then (read_reg x25_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 26)) then (read_reg x26_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 27)) then (read_reg x27_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 28)) then (read_reg x28_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 29)) then (read_reg x29_ref : M (mword 64)) : M (mword 64)
- else if sumbool_of_bool ((Z.eqb p0_ 30)) then (read_reg x30_ref : M (mword 64)) : M (mword 64)
- else (read_reg x31_ref : M (mword 64)) : M (mword 64))
+ else if sumbool_of_bool ((Z.eqb p0_ 1)) then (read_reg x1_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 2)) then (read_reg x2_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 3)) then (read_reg x3_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 4)) then (read_reg x4_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 5)) then (read_reg x5_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 6)) then (read_reg x6_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 7)) then (read_reg x7_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 8)) then (read_reg x8_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 9)) then (read_reg x9_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 10)) then (read_reg x10_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 11)) then (read_reg x11_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 12)) then (read_reg x12_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 13)) then (read_reg x13_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 14)) then (read_reg x14_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 15)) then (read_reg x15_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 16)) then (read_reg x16_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 17)) then (read_reg x17_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 18)) then (read_reg x18_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 19)) then (read_reg x19_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 20)) then (read_reg x20_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 21)) then (read_reg x21_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 22)) then (read_reg x22_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 23)) then (read_reg x23_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 24)) then (read_reg x24_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 25)) then (read_reg x25_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 26)) then (read_reg x26_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 27)) then (read_reg x27_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 28)) then (read_reg x28_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 29)) then (read_reg x29_ref : M (mword 64)) : M (xlenbits)
+ else if sumbool_of_bool ((Z.eqb p0_ 30)) then (read_reg x30_ref : M (mword 64)) : M (xlenbits)
+ else (read_reg x31_ref : M (mword 64)) : M (xlenbits))
: M (mword 64).
Definition wX (r : Z) (v : mword 64) `{ArithFact (0 <= r /\ (r + 1) <= 32)}
@@ -400,189 +407,250 @@ Definition reg_name_abi (r : mword 5)
: string :=
let b__0 := r in
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"zero"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"ra"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"sp"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"gp"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"tp"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"t0"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"t1"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"t2"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"fp"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s1"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a0"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a1"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a2"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a3"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a4"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a5"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a6"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a7"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s2"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s3"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s4"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s5"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s6"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s7"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s8"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s9"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s10"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s11"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"t3"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B1;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B1;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"t4"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno b__0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B1;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B1;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"t5"
else "t6".
@@ -596,14 +664,14 @@ Definition Architecture_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
Definition num_of_Architecture (arg_ : Architecture)
: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
- match arg_ with | RV32 => build_ex 0 | RV64 => build_ex 1 | RV128 => build_ex 2 end.
+ build_ex(match arg_ with | RV32 => 0 | RV64 => 1 | RV128 => 2 end).
Definition architecture (a : mword 2)
: option Architecture :=
let b__0 := a in
- if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then Some RV32
- else if ((eq_vec b__0 (vec_of_bits [B1;B0] : mword 2))) then Some RV64
- else if ((eq_vec b__0 (vec_of_bits [B1;B1] : mword 2))) then Some RV128
+ if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then Some (RV32)
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0] : mword 2))) then Some (RV64)
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1] : mword 2))) then Some (RV128)
else None.
Definition arch_to_bits (a : Architecture)
@@ -623,7 +691,7 @@ Definition Privilege_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
Definition num_of_Privilege (arg_ : Privilege)
: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
- match arg_ with | User => build_ex 0 | Supervisor => build_ex 1 | Machine => build_ex 2 end.
+ build_ex(match arg_ with | User => 0 | Supervisor => 1 | Machine => 2 end).
Definition privLevel_to_bits (p : Privilege)
: mword 2 :=
@@ -654,12 +722,7 @@ Definition AccessType_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 3)}
Definition num_of_AccessType (arg_ : AccessType)
: {e : Z & ArithFact (0 <= e /\ e <= 3)} :=
- match arg_ with
- | Read => build_ex 0
- | Write => build_ex 1
- | ReadWrite => build_ex 2
- | Execute => build_ex 3
- end.
+ build_ex(match arg_ with | Read => 0 | Write => 1 | ReadWrite => 2 | Execute => 3 end).
Definition accessType_to_str (a : AccessType)
: string :=
@@ -673,7 +736,7 @@ Definition ReadType_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 1)}
Definition num_of_ReadType (arg_ : ReadType)
: {e : Z & ArithFact (0 <= e /\ e <= 1)} :=
- match arg_ with | Instruction => build_ex 0 | Data => build_ex 1 end.
+ build_ex(match arg_ with | Instruction => 0 | Data => 1 end).
Definition readType_to_str (r : ReadType)
: string :=
@@ -689,12 +752,7 @@ Definition word_width_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 3)}
Definition num_of_word_width (arg_ : word_width)
: {e : Z & ArithFact (0 <= e /\ e <= 3)} :=
- match arg_ with
- | BYTE => build_ex 0
- | HALF => build_ex 1
- | WORD => build_ex 2
- | DOUBLE => build_ex 3
- end.
+ build_ex(match arg_ with | BYTE => 0 | HALF => 1 | WORD => 2 | DOUBLE => 3 end).
Definition InterruptType_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 8)}
: InterruptType :=
@@ -711,17 +769,17 @@ Definition InterruptType_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 8)}
Definition num_of_InterruptType (arg_ : InterruptType)
: {e : Z & ArithFact (0 <= e /\ e <= 8)} :=
- match arg_ with
- | I_U_Software => build_ex 0
- | I_S_Software => build_ex 1
- | I_M_Software => build_ex 2
- | I_U_Timer => build_ex 3
- | I_S_Timer => build_ex 4
- | I_M_Timer => build_ex 5
- | I_U_External => build_ex 6
- | I_S_External => build_ex 7
- | I_M_External => build_ex 8
- end.
+ build_ex(match arg_ with
+ | I_U_Software => 0
+ | I_S_Software => 1
+ | I_M_Software => 2
+ | I_U_Timer => 3
+ | I_S_Timer => 4
+ | I_M_Timer => 5
+ | I_U_External => 6
+ | I_S_External => 7
+ | I_M_External => 8
+ end).
Definition interruptType_to_bits (i : InterruptType)
: mword 4 :=
@@ -759,24 +817,24 @@ Definition ExceptionType_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 15)
Definition num_of_ExceptionType (arg_ : ExceptionType)
: {e : Z & ArithFact (0 <= e /\ e <= 15)} :=
- match arg_ with
- | E_Fetch_Addr_Align => build_ex 0
- | E_Fetch_Access_Fault => build_ex 1
- | E_Illegal_Instr => build_ex 2
- | E_Breakpoint => build_ex 3
- | E_Load_Addr_Align => build_ex 4
- | E_Load_Access_Fault => build_ex 5
- | E_SAMO_Addr_Align => build_ex 6
- | E_SAMO_Access_Fault => build_ex 7
- | E_U_EnvCall => build_ex 8
- | E_S_EnvCall => build_ex 9
- | E_Reserved_10 => build_ex 10
- | E_M_EnvCall => build_ex 11
- | E_Fetch_Page_Fault => build_ex 12
- | E_Load_Page_Fault => build_ex 13
- | E_Reserved_14 => build_ex 14
- | E_SAMO_Page_Fault => build_ex 15
- end.
+ build_ex(match arg_ with
+ | E_Fetch_Addr_Align => 0
+ | E_Fetch_Access_Fault => 1
+ | E_Illegal_Instr => 2
+ | E_Breakpoint => 3
+ | E_Load_Addr_Align => 4
+ | E_Load_Access_Fault => 5
+ | E_SAMO_Addr_Align => 6
+ | E_SAMO_Access_Fault => 7
+ | E_U_EnvCall => 8
+ | E_S_EnvCall => 9
+ | E_Reserved_10 => 10
+ | E_M_EnvCall => 11
+ | E_Fetch_Page_Fault => 12
+ | E_Load_Page_Fault => 13
+ | E_Reserved_14 => 14
+ | E_SAMO_Page_Fault => 15
+ end).
Definition exceptionType_to_bits (e : ExceptionType)
: mword 4 :=
@@ -822,12 +880,13 @@ Definition exceptionType_to_str (e : ExceptionType)
Definition not_implemented {a : Type} (message : string)
: M (a) :=
- (throw (Error_not_implemented message))
+ (throw (Error_not_implemented
+ (message)))
: M (a).
Definition internal_error {a : Type} (s : string)
: M (a) :=
- assert_exp' false s >>= fun _ => (throw (Error_internal_error tt)) : M (a).
+ assert_exp' false s >>= fun _ => (throw (Error_internal_error (tt))) : M (a).
Definition TrapVectorMode_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
: TrapVectorMode :=
@@ -838,11 +897,7 @@ Definition TrapVectorMode_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)
Definition num_of_TrapVectorMode (arg_ : TrapVectorMode)
: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
- match arg_ with
- | TV_Direct => build_ex 0
- | TV_Vector => build_ex 1
- | TV_Reserved => build_ex 2
- end.
+ build_ex(match arg_ with | TV_Direct => 0 | TV_Vector => 1 | TV_Reserved => 2 end).
Definition trapVectorMode_of_bits (m : mword 2)
: TrapVectorMode :=
@@ -861,12 +916,7 @@ Definition ExtStatus_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 3)}
Definition num_of_ExtStatus (arg_ : ExtStatus)
: {e : Z & ArithFact (0 <= e /\ e <= 3)} :=
- match arg_ with
- | Off => build_ex 0
- | Initial => build_ex 1
- | Clean => build_ex 2
- | Dirty => build_ex 3
- end.
+ build_ex(match arg_ with | Off => 0 | Initial => 1 | Clean => 2 | Dirty => 3 end).
Definition extStatus_to_bits (e : ExtStatus)
: mword 2 :=
@@ -894,22 +944,22 @@ Definition SATPMode_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
Definition num_of_SATPMode (arg_ : SATPMode)
: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
- match arg_ with | Sbare => build_ex 0 | Sv32 => build_ex 1 | Sv39 => build_ex 2 end.
+ build_ex(match arg_ with | Sbare => 0 | Sv32 => 1 | Sv39 => 2 end).
Definition satpMode_of_bits (a : Architecture) (m : mword 4)
: option SATPMode :=
match (a, m) with
- | (g__33, b__0) =>
- if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0] : mword 4))) then Some Sbare
+ | (g__19, b__0) =>
+ if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0] : mword 4))) then Some (Sbare)
else
- match (g__33, b__0) with
+ match (g__19, b__0) with
| (RV32, b__0) =>
- if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1] : mword 4))) then Some Sv32
- else match (RV32, b__0) with | (g__34, g__35) => None end
+ if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1] : mword 4))) then Some (Sv32)
+ else match (RV32, b__0) with | (g__20, g__21) => None end
| (RV64, b__0) =>
- if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B0] : mword 4))) then Some Sv39
- else match (RV64, b__0) with | (g__34, g__35) => None end
- | (g__34, g__35) => None
+ if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B0] : mword 4))) then Some (Sv39)
+ else match (RV64, b__0) with | (g__20, g__21) => None end
+ | (g__20, g__21) => None
end
end.
@@ -921,7 +971,7 @@ Definition uop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 1)}
Definition num_of_uop (arg_ : uop)
: {e : Z & ArithFact (0 <= e /\ e <= 1)} :=
- match arg_ with | RISCV_LUI => build_ex 0 | RISCV_AUIPC => build_ex 1 end.
+ build_ex(match arg_ with | RISCV_LUI => 0 | RISCV_AUIPC => 1 end).
Definition bop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 5)}
: bop :=
@@ -935,14 +985,14 @@ Definition bop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 5)}
Definition num_of_bop (arg_ : bop)
: {e : Z & ArithFact (0 <= e /\ e <= 5)} :=
- match arg_ with
- | RISCV_BEQ => build_ex 0
- | RISCV_BNE => build_ex 1
- | RISCV_BLT => build_ex 2
- | RISCV_BGE => build_ex 3
- | RISCV_BLTU => build_ex 4
- | RISCV_BGEU => build_ex 5
- end.
+ build_ex(match arg_ with
+ | RISCV_BEQ => 0
+ | RISCV_BNE => 1
+ | RISCV_BLT => 2
+ | RISCV_BGE => 3
+ | RISCV_BLTU => 4
+ | RISCV_BGEU => 5
+ end).
Definition iop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 5)}
: iop :=
@@ -956,14 +1006,14 @@ Definition iop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 5)}
Definition num_of_iop (arg_ : iop)
: {e : Z & ArithFact (0 <= e /\ e <= 5)} :=
- match arg_ with
- | RISCV_ADDI => build_ex 0
- | RISCV_SLTI => build_ex 1
- | RISCV_SLTIU => build_ex 2
- | RISCV_XORI => build_ex 3
- | RISCV_ORI => build_ex 4
- | RISCV_ANDI => build_ex 5
- end.
+ build_ex(match arg_ with
+ | RISCV_ADDI => 0
+ | RISCV_SLTI => 1
+ | RISCV_SLTIU => 2
+ | RISCV_XORI => 3
+ | RISCV_ORI => 4
+ | RISCV_ANDI => 5
+ end).
Definition sop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
: sop :=
@@ -974,11 +1024,7 @@ Definition sop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
Definition num_of_sop (arg_ : sop)
: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
- match arg_ with
- | RISCV_SLLI => build_ex 0
- | RISCV_SRLI => build_ex 1
- | RISCV_SRAI => build_ex 2
- end.
+ build_ex(match arg_ with | RISCV_SLLI => 0 | RISCV_SRLI => 1 | RISCV_SRAI => 2 end).
Definition rop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 9)}
: rop :=
@@ -996,18 +1042,18 @@ Definition rop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 9)}
Definition num_of_rop (arg_ : rop)
: {e : Z & ArithFact (0 <= e /\ e <= 9)} :=
- match arg_ with
- | RISCV_ADD => build_ex 0
- | RISCV_SUB => build_ex 1
- | RISCV_SLL => build_ex 2
- | RISCV_SLT => build_ex 3
- | RISCV_SLTU => build_ex 4
- | RISCV_XOR => build_ex 5
- | RISCV_SRL => build_ex 6
- | RISCV_SRA => build_ex 7
- | RISCV_OR => build_ex 8
- | RISCV_AND => build_ex 9
- end.
+ build_ex(match arg_ with
+ | RISCV_ADD => 0
+ | RISCV_SUB => 1
+ | RISCV_SLL => 2
+ | RISCV_SLT => 3
+ | RISCV_SLTU => 4
+ | RISCV_XOR => 5
+ | RISCV_SRL => 6
+ | RISCV_SRA => 7
+ | RISCV_OR => 8
+ | RISCV_AND => 9
+ end).
Definition ropw_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 4)}
: ropw :=
@@ -1020,13 +1066,13 @@ Definition ropw_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 4)}
Definition num_of_ropw (arg_ : ropw)
: {e : Z & ArithFact (0 <= e /\ e <= 4)} :=
- match arg_ with
- | RISCV_ADDW => build_ex 0
- | RISCV_SUBW => build_ex 1
- | RISCV_SLLW => build_ex 2
- | RISCV_SRLW => build_ex 3
- | RISCV_SRAW => build_ex 4
- end.
+ build_ex(match arg_ with
+ | RISCV_ADDW => 0
+ | RISCV_SUBW => 1
+ | RISCV_SLLW => 2
+ | RISCV_SRLW => 3
+ | RISCV_SRAW => 4
+ end).
Definition sopw_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
: sopw :=
@@ -1037,11 +1083,7 @@ Definition sopw_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
Definition num_of_sopw (arg_ : sopw)
: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
- match arg_ with
- | RISCV_SLLIW => build_ex 0
- | RISCV_SRLIW => build_ex 1
- | RISCV_SRAIW => build_ex 2
- end.
+ build_ex(match arg_ with | RISCV_SLLIW => 0 | RISCV_SRLIW => 1 | RISCV_SRAIW => 2 end).
Definition amoop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 8)}
: amoop :=
@@ -1058,17 +1100,17 @@ Definition amoop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 8)}
Definition num_of_amoop (arg_ : amoop)
: {e : Z & ArithFact (0 <= e /\ e <= 8)} :=
- match arg_ with
- | AMOSWAP => build_ex 0
- | AMOADD => build_ex 1
- | AMOXOR => build_ex 2
- | AMOAND => build_ex 3
- | AMOOR => build_ex 4
- | AMOMIN => build_ex 5
- | AMOMAX => build_ex 6
- | AMOMINU => build_ex 7
- | AMOMAXU => build_ex 8
- end.
+ build_ex(match arg_ with
+ | AMOSWAP => 0
+ | AMOADD => 1
+ | AMOXOR => 2
+ | AMOAND => 3
+ | AMOOR => 4
+ | AMOMIN => 5
+ | AMOMAX => 6
+ | AMOMINU => 7
+ | AMOMAXU => 8
+ end).
Definition csrop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
: csrop :=
@@ -1079,195 +1121,256 @@ Definition csrop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
Definition num_of_csrop (arg_ : csrop)
: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
- match arg_ with | CSRRW => build_ex 0 | CSRRS => build_ex 1 | CSRRC => build_ex 2 end.
+ build_ex(match arg_ with | CSRRW => 0 | CSRRS => 1 | CSRRC => 2 end).
Definition reg_name_forwards (arg_ : mword 5)
: string :=
- let p0_ := arg_ in
+ let b__0 := arg_ in
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"zero"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"ra"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"sp"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"gp"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"tp"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"t0"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"t1"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"t2"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"fp"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s1"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a0"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a1"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a2"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a3"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a4"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a5"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a6"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"a7"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s2"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s3"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s4"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s5"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s6"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s7"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s8"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s9"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s10"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"s11"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"t3"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B1;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B1;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"t4"
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B1;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B1;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
"t5"
else "t6".
@@ -1313,197 +1416,260 @@ Definition reg_name_backwards (arg_ : string)
Definition reg_name_forwards_matches (arg_ : mword 5)
: bool :=
- let p0_ := arg_ in
+ let b__0 := arg_ in
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B1;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B1;B0;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B1;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B1;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B1;B1;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B1;B1;B1]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else false.
@@ -1548,262 +1714,422 @@ Definition reg_name_backwards_matches (arg_ : string)
Definition reg_name_matches_prefix (arg_ : string)
: option ((mword 5 * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1871_ := arg_ in
- if ((andb (string_startswith _stringappend_1871_ "zero")
- (match (string_drop _stringappend_1871_ (string_length "zero")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "zero")) with
+ let _stringappend_1319_ := arg_ in
+ if ((andb (string_startswith _stringappend_1319_ "zero")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "zero")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "zero")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B0;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B0;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "ra")
- (match (string_drop _stringappend_1871_ (string_length "ra")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "ra")) with
+ else if ((andb (string_startswith _stringappend_1319_ "ra")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "ra")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "ra")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B0;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B0;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "sp")
- (match (string_drop _stringappend_1871_ (string_length "sp")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "sp")) with
+ else if ((andb (string_startswith _stringappend_1319_ "sp")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "sp")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "sp")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "gp")
- (match (string_drop _stringappend_1871_ (string_length "gp")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "gp")) with
+ else if ((andb (string_startswith _stringappend_1319_ "gp")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "gp")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "gp")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "tp")
- (match (string_drop _stringappend_1871_ (string_length "tp")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "tp")) with
+ else if ((andb (string_startswith _stringappend_1319_ "tp")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "tp")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "tp")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B0;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B0;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "t0")
- (match (string_drop _stringappend_1871_ (string_length "t0")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "t0")) with
+ else if ((andb (string_startswith _stringappend_1319_ "t0")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t0")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t0")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B0;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B0;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "t1")
- (match (string_drop _stringappend_1871_ (string_length "t1")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "t1")) with
+ else if ((andb (string_startswith _stringappend_1319_ "t1")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t1")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t1")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "t2")
- (match (string_drop _stringappend_1871_ (string_length "t2")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "t2")) with
+ else if ((andb (string_startswith _stringappend_1319_ "t2")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t2")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t2")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "fp")
- (match (string_drop _stringappend_1871_ (string_length "fp")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "fp")) with
+ else if ((andb (string_startswith _stringappend_1319_ "fp")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "fp")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "fp")))) with
| s_ =>
- Some ((vec_of_bits [B0;B1;B0;B0;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B1;B0;B0;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "s1")
- (match (string_drop _stringappend_1871_ (string_length "s1")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "s1")) with
+ else if ((andb (string_startswith _stringappend_1319_ "s1")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s1")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s1")))) with
| s_ =>
- Some ((vec_of_bits [B0;B1;B0;B0;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B1;B0;B0;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "a0")
- (match (string_drop _stringappend_1871_ (string_length "a0")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "a0")) with
+ else if ((andb (string_startswith _stringappend_1319_ "a0")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a0")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a0")))) with
| s_ =>
- Some ((vec_of_bits [B0;B1;B0;B1;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B1;B0;B1;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "a1")
- (match (string_drop _stringappend_1871_ (string_length "a1")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "a1")) with
+ else if ((andb (string_startswith _stringappend_1319_ "a1")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a1")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a1")))) with
| s_ =>
- Some ((vec_of_bits [B0;B1;B0;B1;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B1;B0;B1;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "a2")
- (match (string_drop _stringappend_1871_ (string_length "a2")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "a2")) with
+ else if ((andb (string_startswith _stringappend_1319_ "a2")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a2")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a2")))) with
| s_ =>
- Some ((vec_of_bits [B0;B1;B1;B0;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B1;B1;B0;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "a3")
- (match (string_drop _stringappend_1871_ (string_length "a3")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "a3")) with
+ else if ((andb (string_startswith _stringappend_1319_ "a3")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a3")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a3")))) with
| s_ =>
- Some ((vec_of_bits [B0;B1;B1;B0;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B1;B1;B0;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "a4")
- (match (string_drop _stringappend_1871_ (string_length "a4")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "a4")) with
+ else if ((andb (string_startswith _stringappend_1319_ "a4")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a4")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a4")))) with
| s_ =>
- Some ((vec_of_bits [B0;B1;B1;B1;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B1;B1;B1;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "a5")
- (match (string_drop _stringappend_1871_ (string_length "a5")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "a5")) with
+ else if ((andb (string_startswith _stringappend_1319_ "a5")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a5")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a5")))) with
| s_ =>
- Some ((vec_of_bits [B0;B1;B1;B1;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B1;B1;B1;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "a6")
- (match (string_drop _stringappend_1871_ (string_length "a6")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "a6")) with
+ else if ((andb (string_startswith _stringappend_1319_ "a6")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a6")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a6")))) with
| s_ =>
- Some ((vec_of_bits [B1;B0;B0;B0;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B0;B0;B0;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "a7")
- (match (string_drop _stringappend_1871_ (string_length "a7")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "a7")) with
+ else if ((andb (string_startswith _stringappend_1319_ "a7")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a7")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "a7")))) with
| s_ =>
- Some ((vec_of_bits [B1;B0;B0;B0;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B0;B0;B0;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "s2")
- (match (string_drop _stringappend_1871_ (string_length "s2")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "s2")) with
+ else if ((andb (string_startswith _stringappend_1319_ "s2")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s2")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s2")))) with
| s_ =>
- Some ((vec_of_bits [B1;B0;B0;B1;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B0;B0;B1;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "s3")
- (match (string_drop _stringappend_1871_ (string_length "s3")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "s3")) with
+ else if ((andb (string_startswith _stringappend_1319_ "s3")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s3")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s3")))) with
| s_ =>
- Some ((vec_of_bits [B1;B0;B0;B1;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B0;B0;B1;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "s4")
- (match (string_drop _stringappend_1871_ (string_length "s4")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "s4")) with
+ else if ((andb (string_startswith _stringappend_1319_ "s4")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s4")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s4")))) with
| s_ =>
- Some ((vec_of_bits [B1;B0;B1;B0;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B0;B1;B0;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "s5")
- (match (string_drop _stringappend_1871_ (string_length "s5")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "s5")) with
+ else if ((andb (string_startswith _stringappend_1319_ "s5")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s5")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s5")))) with
| s_ =>
- Some ((vec_of_bits [B1;B0;B1;B0;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B0;B1;B0;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "s6")
- (match (string_drop _stringappend_1871_ (string_length "s6")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "s6")) with
+ else if ((andb (string_startswith _stringappend_1319_ "s6")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s6")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s6")))) with
| s_ =>
- Some ((vec_of_bits [B1;B0;B1;B1;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B0;B1;B1;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "s7")
- (match (string_drop _stringappend_1871_ (string_length "s7")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "s7")) with
+ else if ((andb (string_startswith _stringappend_1319_ "s7")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s7")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s7")))) with
| s_ =>
- Some ((vec_of_bits [B1;B0;B1;B1;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B0;B1;B1;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "s8")
- (match (string_drop _stringappend_1871_ (string_length "s8")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "s8")) with
+ else if ((andb (string_startswith _stringappend_1319_ "s8")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s8")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s8")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B0;B0;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B0;B0;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "s9")
- (match (string_drop _stringappend_1871_ (string_length "s9")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "s9")) with
+ else if ((andb (string_startswith _stringappend_1319_ "s9")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s9")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s9")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B0;B0;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B0;B0;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "s10")
- (match (string_drop _stringappend_1871_ (string_length "s10")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "s10")) with
+ else if ((andb (string_startswith _stringappend_1319_ "s10")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s10")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s10")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B0;B1;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B0;B1;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "s11")
- (match (string_drop _stringappend_1871_ (string_length "s11")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "s11")) with
+ else if ((andb (string_startswith _stringappend_1319_ "s11")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s11")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "s11")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B0;B1;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B0;B1;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "t3")
- (match (string_drop _stringappend_1871_ (string_length "t3")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "t3")) with
+ else if ((andb (string_startswith _stringappend_1319_ "t3")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t3")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t3")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B1;B0;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B1;B0;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "t4")
- (match (string_drop _stringappend_1871_ (string_length "t4")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "t4")) with
+ else if ((andb (string_startswith _stringappend_1319_ "t4")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t4")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t4")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B1;B0;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B1;B0;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "t5")
- (match (string_drop _stringappend_1871_ (string_length "t5")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "t5")) with
+ else if ((andb (string_startswith _stringappend_1319_ "t5")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t5")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t5")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B1;B1;B0] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B1;B1;B0] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1871_ "t6")
- (match (string_drop _stringappend_1871_ (string_length "t6")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1871_ (string_length "t6")) with
+ else if ((andb (string_startswith _stringappend_1319_ "t6")
+ (match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t6")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1319_ (build_ex (projT1 (string_length "t6")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B1;B1;B1] : mword 5),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B1;B1;B1] : mword 5), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
else None.
@@ -1817,24 +2143,19 @@ Definition sep_forwards (arg_ : unit)
Definition sep_backwards (arg_ : string)
: M (unit) :=
- let _stringappend_1864_ := arg_ in
- match (opt_spc_matches_prefix _stringappend_1864_) with
- | Some (_stringappend_1865_,(existT _ _stringappend_1866_ _)) =>
- returnm ((_stringappend_1865_, build_ex _stringappend_1866_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__1 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1866_ _) := w__1 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1867_ := string_drop _stringappend_1864_ (build_ex _stringappend_1866_) in
- let _stringappend_1868_ := string_drop _stringappend_1867_ (string_length ",") in
- match (opt_spc_matches_prefix _stringappend_1868_) with
- | Some (_stringappend_1869_,(existT _ _stringappend_1870_ _)) =>
- returnm ((_stringappend_1869_, build_ex _stringappend_1870_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__3 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1870_ _) := w__3 : (unit * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1868_ (build_ex _stringappend_1870_)) with
+ let _stringappend_1314_ := arg_ in
+ (match (opt_spc_matches_prefix _stringappend_1314_) with
+ | Some (tt,(existT _ _stringappend_1315_ _)) => returnm (tt, build_ex _stringappend_1315_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_1315_ _) =>
+ let _stringappend_1316_ := string_drop _stringappend_1314_ (build_ex _stringappend_1315_) in
+ let _stringappend_1317_ :=
+ string_drop _stringappend_1316_ (build_ex (projT1 (string_length ","))) in
+ (match (opt_spc_matches_prefix _stringappend_1317_) with
+ | Some (tt,(existT _ _stringappend_1318_ _)) => returnm (tt, build_ex _stringappend_1318_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_1318_ _) =>
+ (match (string_drop _stringappend_1317_ (build_ex _stringappend_1318_)) with
| "" => returnm (tt : unit)
| _ => exit tt : M (unit)
end)
@@ -1844,16 +2165,17 @@ Definition sep_forwards_matches (arg_ : unit) : bool := match arg_ with | tt =>
Definition sep_backwards_matches (arg_ : string)
: M (bool) :=
- let _stringappend_1857_ := arg_ in
- (if ((match (opt_spc_matches_prefix _stringappend_1857_) with
- | Some (_stringappend_1858_,(existT _ _stringappend_1859_ _)) =>
- let _stringappend_1860_ :=
- string_drop _stringappend_1857_ (build_ex _stringappend_1859_) in
- if ((andb (string_startswith _stringappend_1860_ ",")
- (let _stringappend_1861_ := string_drop _stringappend_1860_ (string_length ",") in
- if ((match (opt_spc_matches_prefix _stringappend_1861_) with
- | Some (_stringappend_1862_,(existT _ _stringappend_1863_ _)) =>
- match (string_drop _stringappend_1861_ (build_ex _stringappend_1863_)) with
+ let _stringappend_1309_ := arg_ in
+ (if ((match (opt_spc_matches_prefix _stringappend_1309_) with
+ | Some (tt,(existT _ _stringappend_1310_ _)) =>
+ let _stringappend_1311_ :=
+ string_drop _stringappend_1309_ (build_ex _stringappend_1310_) in
+ if ((andb (string_startswith _stringappend_1311_ ",")
+ (let _stringappend_1312_ :=
+ string_drop _stringappend_1311_ (build_ex (projT1 (string_length ","))) in
+ if ((match (opt_spc_matches_prefix _stringappend_1312_) with
+ | Some (tt,(existT _ _stringappend_1313_ _)) =>
+ match (string_drop _stringappend_1312_ (build_ex _stringappend_1313_)) with
| "" => true
| _ => false
end
@@ -1865,23 +2187,18 @@ Definition sep_backwards_matches (arg_ : string)
else false
| None => false
end)) then
- match (opt_spc_matches_prefix _stringappend_1857_) with
- | Some (_stringappend_1858_,(existT _ _stringappend_1859_ _)) =>
- returnm ((_stringappend_1858_, build_ex _stringappend_1859_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__1 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1859_ _) := w__1 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1860_ := string_drop _stringappend_1857_ (build_ex _stringappend_1859_) in
- let _stringappend_1861_ := string_drop _stringappend_1860_ (string_length ",") in
- match (opt_spc_matches_prefix _stringappend_1861_) with
- | Some (_stringappend_1862_,(existT _ _stringappend_1863_ _)) =>
- returnm ((_stringappend_1862_, build_ex _stringappend_1863_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__3 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1863_ _) := w__3 : (unit * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1861_ (build_ex _stringappend_1863_)) with
+ (match (opt_spc_matches_prefix _stringappend_1309_) with
+ | Some (tt,(existT _ _stringappend_1310_ _)) => returnm (tt, build_ex _stringappend_1310_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_1310_ _) =>
+ let _stringappend_1311_ := string_drop _stringappend_1309_ (build_ex _stringappend_1310_) in
+ let _stringappend_1312_ :=
+ string_drop _stringappend_1311_ (build_ex (projT1 (string_length ","))) in
+ (match (opt_spc_matches_prefix _stringappend_1312_) with
+ | Some (tt,(existT _ _stringappend_1313_ _)) => returnm (tt, build_ex _stringappend_1313_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_1313_ _) =>
+ (match (string_drop _stringappend_1312_ (build_ex _stringappend_1313_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
@@ -1891,16 +2208,17 @@ Definition sep_backwards_matches (arg_ : string)
Definition sep_matches_prefix (arg_ : string)
: M (option ((unit * {n : Z & ArithFact (n >= 0)}))) :=
- let _stringappend_1850_ := arg_ in
- (if ((match (opt_spc_matches_prefix _stringappend_1850_) with
- | Some (_stringappend_1851_,(existT _ _stringappend_1852_ _)) =>
- let _stringappend_1853_ :=
- string_drop _stringappend_1850_ (build_ex _stringappend_1852_) in
- if ((andb (string_startswith _stringappend_1853_ ",")
- (let _stringappend_1854_ := string_drop _stringappend_1853_ (string_length ",") in
- if ((match (opt_spc_matches_prefix _stringappend_1854_) with
- | Some (_stringappend_1855_,(existT _ _stringappend_1856_ _)) =>
- match (string_drop _stringappend_1854_ (build_ex _stringappend_1856_)) with
+ let _stringappend_1304_ := arg_ in
+ (if ((match (opt_spc_matches_prefix _stringappend_1304_) with
+ | Some (tt,(existT _ _stringappend_1305_ _)) =>
+ let _stringappend_1306_ :=
+ string_drop _stringappend_1304_ (build_ex _stringappend_1305_) in
+ if ((andb (string_startswith _stringappend_1306_ ",")
+ (let _stringappend_1307_ :=
+ string_drop _stringappend_1306_ (build_ex (projT1 (string_length ","))) in
+ if ((match (opt_spc_matches_prefix _stringappend_1307_) with
+ | Some (tt,(existT _ _stringappend_1308_ _)) =>
+ match (string_drop _stringappend_1307_ (build_ex _stringappend_1308_)) with
| s_ => true
end
| None => false
@@ -1911,24 +2229,22 @@ Definition sep_matches_prefix (arg_ : string)
else false
| None => false
end)) then
- match (opt_spc_matches_prefix _stringappend_1850_) with
- | Some (_stringappend_1851_,(existT _ _stringappend_1852_ _)) =>
- returnm ((_stringappend_1851_, build_ex _stringappend_1852_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__1 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1852_ _) := w__1 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1853_ := string_drop _stringappend_1850_ (build_ex _stringappend_1852_) in
- let _stringappend_1854_ := string_drop _stringappend_1853_ (string_length ",") in
- match (opt_spc_matches_prefix _stringappend_1854_) with
- | Some (_stringappend_1855_,(existT _ _stringappend_1856_ _)) =>
- returnm ((_stringappend_1855_, build_ex _stringappend_1856_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__3 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1856_ _) := w__3 : (unit * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_1854_ (build_ex _stringappend_1856_)) with
- | s_ => Some (tt, sub_nat (string_length arg_) (string_length s_))
+ (match (opt_spc_matches_prefix _stringappend_1304_) with
+ | Some (tt,(existT _ _stringappend_1305_ _)) => returnm (tt, build_ex _stringappend_1305_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_1305_ _) =>
+ let _stringappend_1306_ := string_drop _stringappend_1304_ (build_ex _stringappend_1305_) in
+ let _stringappend_1307_ :=
+ string_drop _stringappend_1306_ (build_ex (projT1 (string_length ","))) in
+ (match (opt_spc_matches_prefix _stringappend_1307_) with
+ | Some (tt,(existT _ _stringappend_1308_ _)) => returnm (tt, build_ex _stringappend_1308_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_1308_ _) =>
+ returnm ((match (string_drop _stringappend_1307_ (build_ex _stringappend_1308_)) with
+ | s_ =>
+ Some
+ ((tt, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end)
: option ((unit * {n : Z & ArithFact (n >= 0)})))
else returnm (None : option ((unit * {n : Z & ArithFact (n >= 0)}))))
@@ -1943,8 +2259,8 @@ Definition bool_bits_forwards (arg_ : bool)
Definition bool_bits_backwards (arg_ : mword 1)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B1] : mword 1))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then true
else false.
Definition bool_bits_forwards_matches (arg_ : bool)
@@ -1953,9 +2269,9 @@ Definition bool_bits_forwards_matches (arg_ : bool)
Definition bool_bits_backwards_matches (arg_ : mword 1)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B1] : mword 1))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0] : mword 1))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then true
else false.
Definition bool_not_bits_forwards (arg_ : bool)
@@ -1967,8 +2283,8 @@ Definition bool_not_bits_forwards (arg_ : bool)
Definition bool_not_bits_backwards (arg_ : mword 1)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0] : mword 1))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then true
else false.
Definition bool_not_bits_forwards_matches (arg_ : bool)
@@ -1977,9 +2293,9 @@ Definition bool_not_bits_forwards_matches (arg_ : bool)
Definition bool_not_bits_backwards_matches (arg_ : mword 1)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0] : mword 1))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1] : mword 1))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then true
else false.
Definition size_bits_forwards (arg_ : word_width)
@@ -1993,10 +2309,10 @@ Definition size_bits_forwards (arg_ : word_width)
Definition size_bits_backwards (arg_ : mword 2)
: word_width :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B0] : mword 2))) then BYTE
- else if ((eq_vec p0_ (vec_of_bits [B0;B1] : mword 2))) then HALF
- else if ((eq_vec p0_ (vec_of_bits [B1;B0] : mword 2))) then WORD
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B0] : mword 2))) then BYTE
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then HALF
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0] : mword 2))) then WORD
else DOUBLE.
Definition size_bits_forwards_matches (arg_ : word_width)
@@ -2005,11 +2321,11 @@ Definition size_bits_forwards_matches (arg_ : word_width)
Definition size_bits_backwards_matches (arg_ : mword 2)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B0] : mword 2))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B1] : mword 2))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B0] : mword 2))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1] : mword 2))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B0] : mword 2))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0] : mword 2))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1] : mword 2))) then true
else false.
Definition size_mnemonic_forwards (arg_ : word_width)
@@ -2037,29 +2353,46 @@ Definition size_mnemonic_backwards_matches (arg_ : string)
Definition size_mnemonic_matches_prefix (arg_ : string)
: option ((word_width * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1846_ := arg_ in
- if ((andb (string_startswith _stringappend_1846_ "b")
- (match (string_drop _stringappend_1846_ (string_length "b")) with | s_ => true end))) then
- match (string_drop _stringappend_1846_ (string_length "b")) with
- | s_ => Some (BYTE, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1300_ := arg_ in
+ if ((andb (string_startswith _stringappend_1300_ "b")
+ (match (string_drop _stringappend_1300_ (build_ex (projT1 (string_length "b")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1300_ (build_ex (projT1 (string_length "b")))) with
+ | s_ =>
+ Some
+ ((BYTE, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1846_ "h")
- (match (string_drop _stringappend_1846_ (string_length "h")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1846_ (string_length "h")) with
- | s_ => Some (HALF, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1300_ "h")
+ (match (string_drop _stringappend_1300_ (build_ex (projT1 (string_length "h")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1300_ (build_ex (projT1 (string_length "h")))) with
+ | s_ =>
+ Some
+ ((HALF, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1846_ "w")
- (match (string_drop _stringappend_1846_ (string_length "w")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1846_ (string_length "w")) with
- | s_ => Some (WORD, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1300_ "w")
+ (match (string_drop _stringappend_1300_ (build_ex (projT1 (string_length "w")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1300_ (build_ex (projT1 (string_length "w")))) with
+ | s_ =>
+ Some
+ ((WORD, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1846_ "d")
- (match (string_drop _stringappend_1846_ (string_length "d")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1846_ (string_length "d")) with
- | s_ => Some (DOUBLE, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1300_ "d")
+ (match (string_drop _stringappend_1300_ (build_ex (projT1 (string_length "d")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1300_ (build_ex (projT1 (string_length "d")))) with
+ | s_ =>
+ Some
+ ((DOUBLE, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -2602,8 +2935,8 @@ Definition legalize_misa (m : Misa) (v : mword 64)
let v := Mk_Misa v in
and_boolM (returnm ((eq_vec (_get_Misa_C v) ((bool_to_bits false) : mword 1)) : bool))
((read_reg nextPC_ref : M (mword 64)) >>= fun w__0 : xlenbits =>
- cast_unit_vec (access_vec_dec w__0 1) >>= fun w__1 : mword 1 =>
- returnm ((eq_vec (w__1 : mword 1) ((bool_to_bits true) : mword 1))
+ bit_to_bool (access_vec_dec w__0 1) >>= fun w__1 : bool =>
+ returnm ((Bool.eqb (w__1 : bool) true)
: bool)) >>= fun w__2 : bool =>
returnm ((if (w__2) then m
else _update_Misa_C m (_get_Misa_C v))
@@ -3057,11 +3390,11 @@ Definition cur_Architecture '(tt : unit)
: M (Architecture) :=
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
match w__0 with
- | Machine => read_reg misa_ref >>= fun w__1 : Misa => returnm ((_get_Misa_MXL w__1) : arch_xlen)
+ | Machine => read_reg misa_ref >>= fun w__1 : Misa => returnm ((_get_Misa_MXL w__1) : mword 2)
| Supervisor =>
- read_reg mstatus_ref >>= fun w__2 : Mstatus => returnm ((_get_Mstatus_SXL w__2) : arch_xlen)
+ read_reg mstatus_ref >>= fun w__2 : Mstatus => returnm ((_get_Mstatus_SXL w__2) : mword 2)
| User =>
- read_reg mstatus_ref >>= fun w__3 : Mstatus => returnm ((_get_Mstatus_UXL w__3) : arch_xlen)
+ read_reg mstatus_ref >>= fun w__3 : Mstatus => returnm ((_get_Mstatus_UXL w__3) : mword 2)
end >>= fun a : arch_xlen =>
(match (architecture a) with
| Some (a) => returnm (a : Architecture)
@@ -3805,12 +4138,13 @@ Definition tvec_addr (m : Mtvec) (c : Mcause)
: option (mword 64) :=
let base : xlenbits := concat_vec (_get_Mtvec_Base m) (vec_of_bits [B0;B0] : mword 2) in
match (trapVectorMode_of_bits (_get_Mtvec_Mode m)) with
- | TV_Direct => Some base
+ | TV_Direct => Some (base)
| TV_Vector =>
if ((eq_vec (_get_Mcause_IsInterrupt c) ((bool_to_bits true) : mword 1))) then
- Some (add_vec base
- (shift_bits_left (EXTZ 64 (_get_Mcause_Cause c)) (vec_of_bits [B1;B0] : mword 2)))
- else Some base
+ Some
+ (add_vec base
+ (shift_bits_left (EXTZ 64 (_get_Mcause_Cause c)) (vec_of_bits [B1;B0] : mword 2)))
+ else Some (base)
| TV_Reserved => None
end.
@@ -3967,11 +4301,9 @@ Definition legalize_scounteren (c : Counteren) (v : mword 64)
Definition retire_instruction '(tt : unit)
: M (unit) :=
read_reg minstret_written_ref >>= fun w__0 : bool =>
- (if ((eq_vec ((bool_to_bits w__0) : mword 1) ((bool_to_bits true) : mword 1))) then
- write_reg minstret_written_ref false
- : M (unit)
+ (if ((Bool.eqb w__0 true)) then write_reg minstret_written_ref false : M (unit)
else
- (read_reg minstret_ref : M (mword 64)) >>= fun w__1 : mword 64 =>
+ (read_reg minstret_ref : M (mword 64)) >>= fun w__1 : xlenbits =>
write_reg minstret_ref (add_vec_int w__1 1)
: M (unit))
: M (unit).
@@ -4892,142 +5224,149 @@ Definition csr_name (csr : mword 12)
Definition csr_name_map_forwards (arg_ : mword 12)
: string :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then "ustatus"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then "uie"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then "ustatus"
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
+ "uie"
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
"utvec"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
"uscratch"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
"uepc"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
"ucause"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
"utval"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then "uip"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
+ "uip"
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
"fflags"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then "frm"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ "frm"
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
"fcsr"
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
"cycle"
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
"time"
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
"instret"
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
"cycleh"
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
"timeh"
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
"instreth"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
"sstatus"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
"sedeleg"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
"sideleg"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then "sie"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
+ "sie"
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
"stvec"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
"scounteren"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
"sscratch"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
"sepc"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
"scause"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
"stval"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then "sip"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
+ "sip"
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
"satp"
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B0;B1] : mword 12))) then
"mvendorid"
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B0] : mword 12))) then
"marchid"
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B1] : mword 12))) then
"mimpid"
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B1;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B1;B0;B0] : mword 12))) then
"mhartid"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
"mstatus"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
"misa"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
"medeleg"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
"mideleg"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then "mie"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
+ "mie"
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
"mtvec"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
"mcounteren"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
"mscratch"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
"mepc"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
"mcause"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
"mtval"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then "mip"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
+ "mip"
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
"pmpcfg0"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12))) then
"pmpcfg1"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then
"pmpcfg2"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12))) then
"pmpcfg3"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12))) then
"pmpaddr0"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B1] : mword 12))) then
"pmpaddr1"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B0] : mword 12))) then
"pmpaddr2"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B1] : mword 12))) then
"pmpaddr3"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B0] : mword 12))) then
"pmpaddr4"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B1] : mword 12))) then
"pmpaddr5"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B0] : mword 12))) then
"pmpaddr6"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B1] : mword 12))) then
"pmpaddr7"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B0] : mword 12))) then
"pmpaddr8"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B1] : mword 12))) then
"pmpaddr9"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B0] : mword 12))) then
"pmpaddr10"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B1] : mword 12))) then
"pmpaddr11"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B0] : mword 12))) then
"pmpaddr12"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B1] : mword 12))) then
"pmpaddr13"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B0] : mword 12))) then
"pmpaddr14"
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B1] : mword 12))) then
"pmpaddr15"
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
"mcycle"
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
"minstret"
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
"mcycleh"
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
"minstreth"
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
"tselect"
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12))) then
"tdata1"
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then
"tdata2"
else "tdata3".
@@ -5179,80 +5518,80 @@ Definition csr_name_map_backwards (arg_ : string)
Definition csr_name_map_forwards_matches (arg_ : mword 12)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B1;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12))) then true
else false.
Definition csr_name_map_backwards_matches (arg_ : string)
@@ -5336,648 +5675,980 @@ Definition csr_name_map_backwards_matches (arg_ : string)
Definition csr_name_map_matches_prefix (arg_ : string)
: option ((mword 12 * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1773_ := arg_ in
- if ((andb (string_startswith _stringappend_1773_ "ustatus")
- (match (string_drop _stringappend_1773_ (string_length "ustatus")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "ustatus")) with
+ let _stringappend_1227_ := arg_ in
+ if ((andb (string_startswith _stringappend_1227_ "ustatus")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "ustatus")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "ustatus")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "uie")
- (match (string_drop _stringappend_1773_ (string_length "uie")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "uie")) with
+ else if ((andb (string_startswith _stringappend_1227_ "uie")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "uie")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "uie")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "utvec")
- (match (string_drop _stringappend_1773_ (string_length "utvec")) with
+ else if ((andb (string_startswith _stringappend_1227_ "utvec")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "utvec")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "utvec")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "utvec")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "uscratch")
- (match (string_drop _stringappend_1773_ (string_length "uscratch")) with
+ else if ((andb (string_startswith _stringappend_1227_ "uscratch")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "uscratch")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "uscratch")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "uscratch")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "uepc")
- (match (string_drop _stringappend_1773_ (string_length "uepc")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "uepc")) with
+ else if ((andb (string_startswith _stringappend_1227_ "uepc")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "uepc")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "uepc")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "ucause")
- (match (string_drop _stringappend_1773_ (string_length "ucause")) with
+ else if ((andb (string_startswith _stringappend_1227_ "ucause")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "ucause")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "ucause")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "ucause")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "utval")
- (match (string_drop _stringappend_1773_ (string_length "utval")) with
+ else if ((andb (string_startswith _stringappend_1227_ "utval")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "utval")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "utval")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "utval")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "uip")
- (match (string_drop _stringappend_1773_ (string_length "uip")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "uip")) with
+ else if ((andb (string_startswith _stringappend_1227_ "uip")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "uip")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "uip")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "fflags")
- (match (string_drop _stringappend_1773_ (string_length "fflags")) with
+ else if ((andb (string_startswith _stringappend_1227_ "fflags")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "fflags")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "fflags")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "fflags")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "frm")
- (match (string_drop _stringappend_1773_ (string_length "frm")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "frm")) with
+ else if ((andb (string_startswith _stringappend_1227_ "frm")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "frm")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "frm")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "fcsr")
- (match (string_drop _stringappend_1773_ (string_length "fcsr")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "fcsr")) with
+ else if ((andb (string_startswith _stringappend_1227_ "fcsr")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "fcsr")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "fcsr")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "cycle")
- (match (string_drop _stringappend_1773_ (string_length "cycle")) with
+ else if ((andb (string_startswith _stringappend_1227_ "cycle")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "cycle")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "cycle")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "cycle")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "time")
- (match (string_drop _stringappend_1773_ (string_length "time")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "time")) with
+ else if ((andb (string_startswith _stringappend_1227_ "time")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "time")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "time")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "instret")
- (match (string_drop _stringappend_1773_ (string_length "instret")) with
+ else if ((andb (string_startswith _stringappend_1227_ "instret")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "instret")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "instret")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "instret")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "cycleh")
- (match (string_drop _stringappend_1773_ (string_length "cycleh")) with
+ else if ((andb (string_startswith _stringappend_1227_ "cycleh")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "cycleh")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "cycleh")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "cycleh")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "timeh")
- (match (string_drop _stringappend_1773_ (string_length "timeh")) with
+ else if ((andb (string_startswith _stringappend_1227_ "timeh")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "timeh")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "timeh")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "timeh")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "instreth")
- (match (string_drop _stringappend_1773_ (string_length "instreth")) with
+ else if ((andb (string_startswith _stringappend_1227_ "instreth")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "instreth")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "instreth")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "instreth")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "sstatus")
- (match (string_drop _stringappend_1773_ (string_length "sstatus")) with
+ else if ((andb (string_startswith _stringappend_1227_ "sstatus")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "sstatus")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "sstatus")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "sstatus")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "sedeleg")
- (match (string_drop _stringappend_1773_ (string_length "sedeleg")) with
+ else if ((andb (string_startswith _stringappend_1227_ "sedeleg")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "sedeleg")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "sedeleg")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "sedeleg")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "sideleg")
- (match (string_drop _stringappend_1773_ (string_length "sideleg")) with
+ else if ((andb (string_startswith _stringappend_1227_ "sideleg")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "sideleg")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "sideleg")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "sideleg")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "sie")
- (match (string_drop _stringappend_1773_ (string_length "sie")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "sie")) with
+ else if ((andb (string_startswith _stringappend_1227_ "sie")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "sie")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "sie")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "stvec")
- (match (string_drop _stringappend_1773_ (string_length "stvec")) with
+ else if ((andb (string_startswith _stringappend_1227_ "stvec")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "stvec")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "stvec")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "stvec")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "scounteren")
- (match (string_drop _stringappend_1773_ (string_length "scounteren")) with
+ else if ((andb (string_startswith _stringappend_1227_ "scounteren")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "scounteren")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "scounteren")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "scounteren")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "sscratch")
- (match (string_drop _stringappend_1773_ (string_length "sscratch")) with
+ else if ((andb (string_startswith _stringappend_1227_ "sscratch")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "sscratch")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "sscratch")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "sscratch")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "sepc")
- (match (string_drop _stringappend_1773_ (string_length "sepc")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "sepc")) with
+ else if ((andb (string_startswith _stringappend_1227_ "sepc")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "sepc")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "sepc")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "scause")
- (match (string_drop _stringappend_1773_ (string_length "scause")) with
+ else if ((andb (string_startswith _stringappend_1227_ "scause")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "scause")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "scause")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "scause")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "stval")
- (match (string_drop _stringappend_1773_ (string_length "stval")) with
+ else if ((andb (string_startswith _stringappend_1227_ "stval")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "stval")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "stval")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "stval")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "sip")
- (match (string_drop _stringappend_1773_ (string_length "sip")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "sip")) with
+ else if ((andb (string_startswith _stringappend_1227_ "sip")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "sip")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "sip")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "satp")
- (match (string_drop _stringappend_1773_ (string_length "satp")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "satp")) with
+ else if ((andb (string_startswith _stringappend_1227_ "satp")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "satp")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "satp")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mvendorid")
- (match (string_drop _stringappend_1773_ (string_length "mvendorid")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mvendorid")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "mvendorid")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "mvendorid")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mvendorid")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "marchid")
- (match (string_drop _stringappend_1773_ (string_length "marchid")) with
+ else if ((andb (string_startswith _stringappend_1227_ "marchid")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "marchid")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "marchid")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "marchid")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mimpid")
- (match (string_drop _stringappend_1773_ (string_length "mimpid")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mimpid")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mimpid")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "mimpid")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mimpid")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mhartid")
- (match (string_drop _stringappend_1773_ (string_length "mhartid")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mhartid")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mhartid")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "mhartid")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mhartid")))) with
| s_ =>
- Some ((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B1;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B1;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mstatus")
- (match (string_drop _stringappend_1773_ (string_length "mstatus")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mstatus")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mstatus")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "mstatus")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mstatus")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "misa")
- (match (string_drop _stringappend_1773_ (string_length "misa")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "misa")) with
+ else if ((andb (string_startswith _stringappend_1227_ "misa")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "misa")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "misa")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "medeleg")
- (match (string_drop _stringappend_1773_ (string_length "medeleg")) with
+ else if ((andb (string_startswith _stringappend_1227_ "medeleg")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "medeleg")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "medeleg")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "medeleg")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mideleg")
- (match (string_drop _stringappend_1773_ (string_length "mideleg")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mideleg")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mideleg")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "mideleg")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mideleg")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mie")
- (match (string_drop _stringappend_1773_ (string_length "mie")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "mie")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mie")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mie")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mie")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mtvec")
- (match (string_drop _stringappend_1773_ (string_length "mtvec")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mtvec")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mtvec")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "mtvec")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mtvec")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mcounteren")
- (match (string_drop _stringappend_1773_ (string_length "mcounteren")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mcounteren")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "mcounteren")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "mcounteren")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mcounteren")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mscratch")
- (match (string_drop _stringappend_1773_ (string_length "mscratch")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mscratch")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "mscratch")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "mscratch")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mscratch")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mepc")
- (match (string_drop _stringappend_1773_ (string_length "mepc")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "mepc")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mepc")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mepc")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mepc")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mcause")
- (match (string_drop _stringappend_1773_ (string_length "mcause")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mcause")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mcause")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "mcause")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mcause")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mtval")
- (match (string_drop _stringappend_1773_ (string_length "mtval")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mtval")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mtval")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "mtval")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mtval")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mip")
- (match (string_drop _stringappend_1773_ (string_length "mip")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1773_ (string_length "mip")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mip")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mip")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mip")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpcfg0")
- (match (string_drop _stringappend_1773_ (string_length "pmpcfg0")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpcfg0")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpcfg0")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpcfg0")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpcfg0")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpcfg1")
- (match (string_drop _stringappend_1773_ (string_length "pmpcfg1")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpcfg1")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpcfg1")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpcfg1")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpcfg1")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpcfg2")
- (match (string_drop _stringappend_1773_ (string_length "pmpcfg2")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpcfg2")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpcfg2")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpcfg2")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpcfg2")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpcfg3")
- (match (string_drop _stringappend_1773_ (string_length "pmpcfg3")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpcfg3")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpcfg3")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpcfg3")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpcfg3")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr0")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr0")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr0")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr0")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr0")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr0")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr1")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr1")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr1")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr1")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr1")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr1")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr2")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr2")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr2")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr2")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr2")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr2")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr3")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr3")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr3")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr3")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr3")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr3")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr4")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr4")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr4")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr4")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr4")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr4")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr5")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr5")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr5")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr5")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr5")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr5")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr6")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr6")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr6")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr6")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr6")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr6")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr7")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr7")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr7")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr7")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr7")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr7")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr8")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr8")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr8")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr8")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr8")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr8")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr9")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr9")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr9")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr9")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr9")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr9")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr10")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr10")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr10")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr10")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr10")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr10")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr11")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr11")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr11")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr11")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr11")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr11")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr12")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr12")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr12")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr12")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr12")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr12")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr13")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr13")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr13")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr13")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr13")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr13")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr14")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr14")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr14")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr14")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr14")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr14")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "pmpaddr15")
- (match (string_drop _stringappend_1773_ (string_length "pmpaddr15")) with
+ else if ((andb (string_startswith _stringappend_1227_ "pmpaddr15")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "pmpaddr15")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "pmpaddr15")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "pmpaddr15")))) with
| s_ =>
- Some ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mcycle")
- (match (string_drop _stringappend_1773_ (string_length "mcycle")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mcycle")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mcycle")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "mcycle")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mcycle")))) with
| s_ =>
- Some ((vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "minstret")
- (match (string_drop _stringappend_1773_ (string_length "minstret")) with
+ else if ((andb (string_startswith _stringappend_1227_ "minstret")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "minstret")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "minstret")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "minstret")))) with
| s_ =>
- Some ((vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "mcycleh")
- (match (string_drop _stringappend_1773_ (string_length "mcycleh")) with
+ else if ((andb (string_startswith _stringappend_1227_ "mcycleh")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mcycleh")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "mcycleh")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "mcycleh")))) with
| s_ =>
- Some ((vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "minstreth")
- (match (string_drop _stringappend_1773_ (string_length "minstreth")) with
+ else if ((andb (string_startswith _stringappend_1227_ "minstreth")
+ (match (string_drop _stringappend_1227_
+ (build_ex (projT1 (string_length "minstreth")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "minstreth")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "minstreth")))) with
| s_ =>
- Some ((vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "tselect")
- (match (string_drop _stringappend_1773_ (string_length "tselect")) with
+ else if ((andb (string_startswith _stringappend_1227_ "tselect")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "tselect")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "tselect")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "tselect")))) with
| s_ =>
- Some ((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "tdata1")
- (match (string_drop _stringappend_1773_ (string_length "tdata1")) with
+ else if ((andb (string_startswith _stringappend_1227_ "tdata1")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "tdata1")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "tdata1")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "tdata1")))) with
| s_ =>
- Some ((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "tdata2")
- (match (string_drop _stringappend_1773_ (string_length "tdata2")) with
+ else if ((andb (string_startswith _stringappend_1227_ "tdata2")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "tdata2")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "tdata2")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "tdata2")))) with
| s_ =>
- Some ((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
- else if ((andb (string_startswith _stringappend_1773_ "tdata3")
- (match (string_drop _stringappend_1773_ (string_length "tdata3")) with
+ else if ((andb (string_startswith _stringappend_1227_ "tdata3")
+ (match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "tdata3")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1773_ (string_length "tdata3")) with
+ match (string_drop _stringappend_1227_ (build_ex (projT1 (string_length "tdata3")))) with
| s_ =>
- Some ((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ (((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end
else None.
@@ -6066,10 +6737,7 @@ Definition is_CSR_defined (csr : mword 12) (p : Privilege)
Definition check_CSR_access (csrrw : mword 2) (csrpr : mword 2) (p : Privilege) (isWrite : bool)
: bool :=
- andb
- (negb
- (andb (eq_vec ((bool_to_bits isWrite) : mword 1) ((bool_to_bits true) : mword 1))
- (eq_vec csrrw (vec_of_bits [B1;B1] : mword 2))))
+ andb (negb (andb (Bool.eqb isWrite true) (eq_vec csrrw (vec_of_bits [B1;B1] : mword 2))))
(zopz0zKzJ_u (privLevel_to_bits p) csrpr).
Definition check_TVM_SATP (csr : mword 12) (p : Privilege)
@@ -6107,7 +6775,7 @@ Definition check_Counteren (csr : mword 12) (p : Privilege)
: bool)
else
returnm ((match (b__0, Supervisor) with
- | (g__31, g__32) =>
+ | (g__17, g__18) =>
if ((andb
(zopz0zIzJ_u
(vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12) csr)
@@ -6134,7 +6802,7 @@ Definition check_Counteren (csr : mword 12) (p : Privilege)
: bool)
else
returnm ((match (b__3, User) with
- | (g__31, g__32) =>
+ | (g__17, g__18) =>
if ((andb
(zopz0zIzJ_u
(vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12) csr)
@@ -6146,7 +6814,7 @@ Definition check_Counteren (csr : mword 12) (p : Privilege)
end)
: bool))
: M (bool)
- | (g__31, g__32) =>
+ | (g__17, g__18) =>
returnm ((if ((andb
(zopz0zIzJ_u (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12)
csr)
@@ -6175,7 +6843,7 @@ Axiom cancel_reservation : forall (_ : unit) , unit.
Definition exception_delegatee (e : ExceptionType) (p : Privilege)
: M (Privilege) :=
- let '(existT _ idx _) := num_of_ExceptionType e in
+ let 'idx := projT1 (num_of_ExceptionType e) in
read_reg medeleg_ref >>= fun w__0 : Medeleg =>
let super := access_vec_dec (_get_Medeleg_bits w__0) idx in
read_reg sedeleg_ref >>= fun w__1 : Sedeleg =>
@@ -6192,20 +6860,31 @@ Definition exception_delegatee (e : ExceptionType) (p : Privilege)
Definition findPendingInterrupt (ip : mword 64)
: option InterruptType :=
let ip := Mk_Minterrupts ip in
- if ((eq_vec (_get_Minterrupts_MEI ip) ((bool_to_bits true) : mword 1))) then Some I_M_External
+ if ((eq_vec (_get_Minterrupts_MEI ip) ((bool_to_bits true) : mword 1))) then Some (I_M_External)
else if ((eq_vec (_get_Minterrupts_MSI ip) ((bool_to_bits true) : mword 1))) then
- Some I_M_Software
- else if ((eq_vec (_get_Minterrupts_MTI ip) ((bool_to_bits true) : mword 1))) then Some I_M_Timer
+ Some
+ (I_M_Software)
+ else if ((eq_vec (_get_Minterrupts_MTI ip) ((bool_to_bits true) : mword 1))) then
+ Some
+ (I_M_Timer)
else if ((eq_vec (_get_Minterrupts_SEI ip) ((bool_to_bits true) : mword 1))) then
- Some I_S_External
+ Some
+ (I_S_External)
else if ((eq_vec (_get_Minterrupts_SSI ip) ((bool_to_bits true) : mword 1))) then
- Some I_S_Software
- else if ((eq_vec (_get_Minterrupts_STI ip) ((bool_to_bits true) : mword 1))) then Some I_S_Timer
+ Some
+ (I_S_Software)
+ else if ((eq_vec (_get_Minterrupts_STI ip) ((bool_to_bits true) : mword 1))) then
+ Some
+ (I_S_Timer)
else if ((eq_vec (_get_Minterrupts_UEI ip) ((bool_to_bits true) : mword 1))) then
- Some I_U_External
+ Some
+ (I_U_External)
else if ((eq_vec (_get_Minterrupts_USI ip) ((bool_to_bits true) : mword 1))) then
- Some I_U_Software
- else if ((eq_vec (_get_Minterrupts_UTI ip) ((bool_to_bits true) : mword 1))) then Some I_U_Timer
+ Some
+ (I_U_Software)
+ else if ((eq_vec (_get_Minterrupts_UTI ip) ((bool_to_bits true) : mword 1))) then
+ Some
+ (I_U_Timer)
else None.
Definition curInterrupt (priv : Privilege) (pend : Minterrupts) (enbl : Minterrupts) (delg : Minterrupts)
@@ -6249,7 +6928,8 @@ Definition curInterrupt (priv : Privilege) (pend : Minterrupts) (enbl : Minterru
(match (findPendingInterrupt eff_mip) with
| Some (i) =>
let r := (i, Machine) in
- returnm ((Some r)
+ returnm ((Some
+ (r))
: option ((InterruptType * Privilege)))
| None =>
(internal_error
@@ -6262,7 +6942,8 @@ Definition curInterrupt (priv : Privilege) (pend : Minterrupts) (enbl : Minterru
(match (findPendingInterrupt eff_sip) with
| Some (i) =>
let r := (i, Supervisor) in
- returnm ((Some r)
+ returnm ((Some
+ (r))
: option ((InterruptType * Privilege)))
| None =>
(internal_error
@@ -6410,7 +7091,7 @@ Definition handle_exception (cur_priv : Privilege) (ctl : ctl_result) (pc : mwor
(handle_trap del_priv false ((exceptionType_to_bits e.(sync_exception_trap)) : mword 4) pc
e.(sync_exception_excinfo))
: M (mword 64)
- | (_, CTL_MRET (_)) =>
+ | (_, CTL_MRET (tt)) =>
read_reg cur_privilege_ref >>= fun prev_priv =>
read_reg mstatus_ref >>= fun w__1 : Mstatus =>
_set_Mstatus_MIE mstatus_ref (_get_Mstatus_MPIE w__1) >>
@@ -6435,9 +7116,9 @@ Definition handle_exception (cur_priv : Privilege) (ctl : ctl_result) (pc : mwor
(String.append " to " ((privLevel_to_str w__5) : string)))))
: unit in
let '_ := (cancel_reservation tt) : unit in
- (read_reg mepc_ref : M (mword 64)) >>= fun w__6 : mword 64 =>
+ (read_reg mepc_ref : M (mword 64)) >>= fun w__6 : xlenbits =>
pc_alignment_mask tt >>= fun w__7 : mword 64 => returnm ((and_vec w__6 w__7) : mword 64)
- | (_, CTL_SRET (_)) =>
+ | (_, CTL_SRET (tt)) =>
read_reg cur_privilege_ref >>= fun prev_priv =>
read_reg mstatus_ref >>= fun w__8 : Mstatus =>
_set_Mstatus_SIE mstatus_ref (_get_Mstatus_SPIE w__8) >>
@@ -6465,17 +7146,19 @@ Definition handle_exception (cur_priv : Privilege) (ctl : ctl_result) (pc : mwor
(String.append " to " ((privLevel_to_str w__12) : string)))))
: unit in
let '_ := (cancel_reservation tt) : unit in
- (read_reg sepc_ref : M (mword 64)) >>= fun w__13 : mword 64 =>
+ (read_reg sepc_ref : M (mword 64)) >>= fun w__13 : xlenbits =>
pc_alignment_mask tt >>= fun w__14 : mword 64 => returnm ((and_vec w__13 w__14) : mword 64)
end)
: M (mword 64).
Definition handle_mem_exception (addr : mword 64) (e : ExceptionType)
: M (unit) :=
- let t : sync_exception := {| sync_exception_trap := e; sync_exception_excinfo := (Some addr) |} in
+ let t : sync_exception :=
+ {| sync_exception_trap := e;
+ sync_exception_excinfo := (Some (addr)) |} in
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
- (read_reg PC_ref : M (mword 64)) >>= fun w__1 : mword 64 =>
- handle_exception w__0 (CTL_TRAP t) w__1 >>= fun w__2 : xlenbits =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__1 : xlenbits =>
+ handle_exception w__0 (CTL_TRAP (t)) w__1 >>= fun w__2 : mword 64 =>
write_reg nextPC_ref w__2
: M (unit).
@@ -6483,17 +7166,17 @@ Definition handle_decode_exception (instbits : mword 64)
: M (unit) :=
let t : sync_exception :=
{| sync_exception_trap := E_Illegal_Instr;
- sync_exception_excinfo := (Some instbits) |} in
+ sync_exception_excinfo := (Some (instbits)) |} in
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
- (read_reg PC_ref : M (mword 64)) >>= fun w__1 : mword 64 =>
- handle_exception w__0 (CTL_TRAP t) w__1 >>= fun w__2 : xlenbits =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__1 : xlenbits =>
+ handle_exception w__0 (CTL_TRAP (t)) w__1 >>= fun w__2 : mword 64 =>
write_reg nextPC_ref w__2
: M (unit).
Definition handle_interrupt (i : InterruptType) (del_priv : Privilege)
: M (unit) :=
- (read_reg PC_ref : M (mword 64)) >>= fun w__0 : mword 64 =>
- handle_trap del_priv true ((interruptType_to_bits i) : mword 4) w__0 None >>= fun w__1 : xlenbits =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__0 : xlenbits =>
+ handle_trap del_priv true ((interruptType_to_bits i) : mword 4) w__0 None >>= fun w__1 : mword 64 =>
write_reg nextPC_ref w__1
: M (unit).
@@ -6503,8 +7186,8 @@ Definition handle_illegal '(tt : unit)
{| sync_exception_trap := E_Illegal_Instr;
sync_exception_excinfo := None |} in
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
- (read_reg PC_ref : M (mword 64)) >>= fun w__1 : mword 64 =>
- handle_exception w__0 (CTL_TRAP t) w__1 >>= fun w__2 : xlenbits =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__1 : xlenbits =>
+ handle_exception w__0 (CTL_TRAP (t)) w__1 >>= fun w__2 : mword 64 =>
write_reg nextPC_ref w__2
: M (unit).
@@ -6597,25 +7280,28 @@ Definition within_htif_readable (addr : mword 64) (width : Z)
Axiom plat_insns_per_tick : forall (_ : unit) , Z.
-Let MSIP_BASE : xlenbits :=
+Definition MSIP_BASE : xlenbits :=
(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]
: mword 64).
-Let MTIMECMP_BASE : xlenbits :=
+Hint Unfold MSIP_BASE : sail.
+Definition MTIMECMP_BASE : xlenbits :=
(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;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
B0]
: mword 64).
-Let MTIME_BASE : xlenbits :=
+Hint Unfold MTIMECMP_BASE : sail.
+Definition MTIME_BASE : xlenbits :=
(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;B1;B0;B1;B1;B1;B1;B1;B1;B1;B1;B1;B1;B1;B0;B0;
B0]
: mword 64).
-Definition clint_load (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
+Hint Unfold MTIME_BASE : sail.
+Definition clint_load (addr : mword 64) (width : Z) `{ArithFact (width >= (0 + 1))}
: M (MemoryOpResult (mword (8 * width))) :=
let addr := sub_vec addr (plat_clint_base tt) in
(if sumbool_of_bool ((andb (eq_vec addr MSIP_BASE) (orb (Z.eqb width 8) (Z.eqb width 4)))) then
@@ -6627,7 +7313,8 @@ Definition clint_load (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
(String.append "] -> " (string_of_bits (_get_Minterrupts_MSI w__0))))))
: unit in
read_reg mip_ref >>= fun w__1 : Minterrupts =>
- returnm ((MemValue (autocast (zero_extend (_get_Minterrupts_MSI w__1) (Z.mul 8 width))))
+ returnm ((MemValue
+ (autocast (zero_extend (_get_Minterrupts_MSI w__1) (Z.mul 8 width))))
: MemoryOpResult (mword (8 * width)))
else if sumbool_of_bool ((andb (eq_vec addr MTIMECMP_BASE) (Z.eqb width 8))) then
(read_reg mtimecmp_ref : M (mword 64)) >>= fun w__2 : xlenbits =>
@@ -6637,7 +7324,8 @@ Definition clint_load (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
(String.append (string_of_bits addr) (String.append "] -> " (string_of_bits w__2)))))
: unit in
(read_reg mtimecmp_ref : M (mword 64)) >>= fun w__3 : xlenbits =>
- returnm ((MemValue (autocast (zero_extend w__3 64)))
+ returnm ((MemValue
+ (autocast (autocast (zero_extend w__3 64))))
: MemoryOpResult (mword (8 * width)))
else if sumbool_of_bool ((andb (eq_vec addr MTIME_BASE) (Z.eqb width 8))) then
(read_reg mtime_ref : M (mword 64)) >>= fun w__4 : xlenbits =>
@@ -6647,14 +7335,16 @@ Definition clint_load (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
(String.append (string_of_bits addr) (String.append "] -> " (string_of_bits w__4)))))
: unit in
(read_reg mtime_ref : M (mword 64)) >>= fun w__5 : xlenbits =>
- returnm ((MemValue (autocast (zero_extend w__5 64)))
+ returnm ((MemValue
+ (autocast (autocast (zero_extend w__5 64))))
: MemoryOpResult (mword (8 * width)))
else
let '_ :=
(print_endline
(String.append "clint[" (String.append (string_of_bits addr) "] -> <not-mapped>")))
: unit in
- returnm ((MemException E_Load_Access_Fault)
+ returnm ((MemException
+ (E_Load_Access_Fault))
: MemoryOpResult (mword (8 * width))))
: M (MemoryOpResult (mword (8 * width))).
@@ -6664,7 +7354,7 @@ Definition clint_dispatch '(tt : unit)
let '_ := (print_endline (String.append "clint::tick mtime <- " (string_of_bits w__0))) : unit in
_set_Minterrupts_MTI mip_ref ((bool_to_bits false) : mword 1) >>
(read_reg mtimecmp_ref : M (mword 64)) >>= fun w__1 : xlenbits =>
- (read_reg mtime_ref : M (mword 64)) >>= fun w__2 : mword 64 =>
+ (read_reg mtime_ref : M (mword 64)) >>= fun w__2 : xlenbits =>
(if ((zopz0zIzJ_u w__1 w__2)) then
(read_reg mtime_ref : M (mword 64)) >>= fun w__3 : xlenbits =>
let '_ :=
@@ -6676,7 +7366,7 @@ Definition clint_dispatch '(tt : unit)
: M (unit).
Definition clint_store (addr : mword 64) (width : Z) (data : mword (8 * width)) `{ArithFact (width >=
- 1)}
+ (0 + 1))}
: M (MemoryOpResult bool) :=
let addr := sub_vec addr (plat_clint_base tt) in
(if sumbool_of_bool ((andb (eq_vec addr MSIP_BASE) (orb (Z.eqb width 8) (Z.eqb width 4)))) then
@@ -6694,7 +7384,7 @@ Definition clint_store (addr : mword 64) (width : Z) (data : mword (8 * width))
_set_Minterrupts_MSI mip_ref
((bool_to_bits (eq_vec (w__1 : mword 1) (vec_of_bits [B1] : mword 1)))
: mword 1) >>
- clint_dispatch tt >> returnm ((MemValue true) : MemoryOpResult bool)
+ clint_dispatch tt >> returnm ((MemValue (true)) : MemoryOpResult bool)
else if sumbool_of_bool ((andb (eq_vec addr MTIMECMP_BASE) (Z.eqb width 8))) then
let '_ :=
(print_endline
@@ -6703,7 +7393,7 @@ Definition clint_store (addr : mword 64) (width : Z) (data : mword (8 * width))
(String.append "] <- " (String.append (string_of_bits data) " (mtimecmp)")))))
: unit in
write_reg mtimecmp_ref (zero_extend data 64) >>
- clint_dispatch tt >> returnm ((MemValue true) : MemoryOpResult bool)
+ clint_dispatch tt >> returnm ((MemValue (true)) : MemoryOpResult bool)
else
let '_ :=
(print_endline
@@ -6711,15 +7401,16 @@ Definition clint_store (addr : mword 64) (width : Z) (data : mword (8 * width))
(String.append (string_of_bits addr)
(String.append "] <- " (String.append (string_of_bits data) " (<unmapped>)")))))
: unit in
- returnm ((MemException E_SAMO_Access_Fault)
+ returnm ((MemException
+ (E_SAMO_Access_Fault))
: MemoryOpResult bool))
: M (MemoryOpResult bool).
Definition tick_clock '(tt : unit)
: M (unit) :=
- (read_reg mcycle_ref : M (mword 64)) >>= fun w__0 : mword 64 =>
+ (read_reg mcycle_ref : M (mword 64)) >>= fun w__0 : xlenbits =>
write_reg mcycle_ref (add_vec_int w__0 1) >>
- (read_reg mtime_ref : M (mword 64)) >>= fun w__1 : mword 64 =>
+ (read_reg mtime_ref : M (mword 64)) >>= fun w__1 : xlenbits =>
write_reg mtime_ref (add_vec_int w__1 1) >> (clint_dispatch tt) : M (unit).
Axiom plat_term_write : forall (_ : bits 8) , unit.
@@ -6814,7 +7505,7 @@ Definition _update_htif_cmd_payload (v : htif_cmd) (x : mword 48)
htif_cmd_htif_cmd_chunk_0 :=
(update_subrange_vec_dec v.(htif_cmd_htif_cmd_chunk_0) 47 0 (subrange_vec_dec x 47 0)) ]}.
-Definition htif_load (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
+Definition htif_load (addr : mword 64) (width : Z) `{ArithFact (width >= (0 + 1))}
: M (MemoryOpResult (mword (8 * width))) :=
(read_reg htif_tohost_ref : M (mword 64)) >>= fun w__0 : xlenbits =>
let '_ :=
@@ -6824,12 +7515,13 @@ Definition htif_load (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
: unit in
(if sumbool_of_bool ((Z.eqb width 8)) then
(read_reg htif_tohost_ref : M (mword 64)) >>= fun w__1 : xlenbits =>
- returnm ((MemValue (autocast (zero_extend w__1 64)))
+ returnm ((MemValue
+ (autocast (autocast (zero_extend w__1 64))))
: MemoryOpResult (mword (8 * width)))
- else returnm ((MemException E_Load_Access_Fault) : MemoryOpResult (mword (8 * width))))
+ else returnm ((MemException (E_Load_Access_Fault)) : MemoryOpResult (mword (8 * width))))
: M (MemoryOpResult (mword (8 * width))).
-Definition htif_store (addr : mword 64) (width : Z) (data : mword (8 * width)) `{ArithFact (1 <=
+Definition htif_store (addr : mword 64) (width : Z) (data : mword (8 * width)) `{ArithFact ((0 + 1) <=
width /\
width <= 8)}
: M (MemoryOpResult bool) :=
@@ -6871,7 +7563,8 @@ Definition htif_store (addr : mword 64) (width : Z) (data : mword (8 * width)) `
else print_endline (String.append "Unknown term cmd: " (string_of_bits b__2))
else print_endline (String.append "htif-???? cmd: " (string_of_bits data)))
: unit)) >>
- returnm ((MemValue true)
+ returnm ((MemValue
+ (true))
: MemoryOpResult bool).
Definition htif_tick '(tt : unit)
@@ -6889,7 +7582,7 @@ Definition within_mmio_writable (addr : mword 64) (width : Z)
: bool :=
orb (within_clint addr width) (andb (within_htif_writable addr width) (Z.leb width 8)).
-Definition mmio_read (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
+Definition mmio_read (addr : mword 64) (width : Z) `{ArithFact (width >= (0 + 1))}
: M (MemoryOpResult (mword (8 * width))) :=
(if ((within_clint addr width)) then
(clint_load addr width)
@@ -6897,17 +7590,17 @@ Definition mmio_read (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
else if sumbool_of_bool ((andb (within_htif_readable addr width) (Z.leb 1 width))) then
(htif_load addr width)
: M (MemoryOpResult (mword (8 * width)))
- else returnm ((MemException E_Load_Access_Fault) : MemoryOpResult (mword (8 * width))))
+ else returnm ((MemException (E_Load_Access_Fault)) : MemoryOpResult (mword (8 * width))))
: M (MemoryOpResult (mword (8 * width))).
Definition mmio_write (addr : mword 64) (width : Z) (data : mword (8 * width)) `{ArithFact (width >=
- 1)}
+ (0 + 1))}
: M (MemoryOpResult bool) :=
(if ((within_clint addr width)) then (clint_store addr width data) : M (MemoryOpResult bool)
else if sumbool_of_bool ((andb (within_htif_writable addr width) (Z.leb width 8))) then
(htif_store addr width data)
: M (MemoryOpResult bool)
- else returnm ((MemException E_SAMO_Access_Fault) : MemoryOpResult bool))
+ else returnm ((MemException (E_SAMO_Access_Fault)) : MemoryOpResult bool))
: M (MemoryOpResult bool).
Definition init_platform '(tt : unit)
@@ -6925,16 +7618,17 @@ Definition tick_platform '(tt : unit)
Definition is_aligned_addr (addr : mword 64) (width : Z)
: bool :=
- Z.eqb (projT1 ((ex_int (modulus (projT1 (uint addr)) width)) : {syn_n : Z & ArithFact (True)}))
- 0.
+ Z.eqb
+ (projT1 ((build_ex (projT1 (ex_int (modulus (projT1 (uint addr)) width))))
+ : {syn_n : Z & ArithFact (True)})) 0.
Definition phys_mem_read (t : ReadType) (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (res : bool) `{ArithFact (width >=
0)}
: M (MemoryOpResult (mword (8 * width))) :=
__RISCV_read addr width aq rl res >>= fun w__0 : option (mword (8 * width)) =>
returnm ((match (t, w__0) with
- | (Instruction, None) => MemException E_Fetch_Access_Fault
- | (Data, None) => MemException E_Load_Access_Fault
+ | (Instruction, None) => MemException (E_Fetch_Access_Fault)
+ | (Data, None) => MemException (E_Load_Access_Fault)
| (_, Some (v)) =>
let '_ :=
(print_endline
@@ -6944,11 +7638,13 @@ Definition phys_mem_read (t : ReadType) (addr : mword 64) (width : Z) (aq : bool
(String.append (string_of_bits addr)
(String.append "] -> " (string_of_bits v)))))))
: unit in
- MemValue v
+ MemValue
+ (v)
end)
: MemoryOpResult (mword (8 * width))).
-Definition checked_mem_read (t : ReadType) (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
+Definition checked_mem_read (t : ReadType) (addr : mword 64) (width : Z) `{ArithFact (width >=
+ (0 + 1))}
: M (MemoryOpResult (mword (8 * width))) :=
(if ((andb (generic_eq ((readType_to_str t) : string) ((readType_to_str Data) : string))
(within_mmio_readable addr width))) then
@@ -6957,44 +7653,45 @@ Definition checked_mem_read (t : ReadType) (addr : mword 64) (width : Z) `{Arith
else if ((within_phys_mem addr width)) then
(phys_mem_read t addr width false false false)
: M (MemoryOpResult (mword (8 * width)))
- else returnm ((MemException E_Load_Access_Fault) : MemoryOpResult (mword (8 * width))))
+ else returnm ((MemException (E_Load_Access_Fault)) : MemoryOpResult (mword (8 * width))))
: M (MemoryOpResult (mword (8 * width))).
-Definition MEMr (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
+Definition MEMr (addr : mword 64) (width : Z) `{ArithFact (width >= (0 + 1))}
: M (MemoryOpResult (mword (8 * width))) :=
(checked_mem_read Data addr width)
: M (MemoryOpResult (mword (8 * width))).
-Definition MEMr_acquire (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
+Definition MEMr_acquire (addr : mword 64) (width : Z) `{ArithFact (width >= (0 + 1))}
: M (MemoryOpResult (mword (8 * width))) :=
(checked_mem_read Data addr width)
: M (MemoryOpResult (mword (8 * width))).
-Definition MEMr_strong_acquire (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
+Definition MEMr_strong_acquire (addr : mword 64) (width : Z) `{ArithFact (width >= (0 + 1))}
: M (MemoryOpResult (mword (8 * width))) :=
(checked_mem_read Data addr width)
: M (MemoryOpResult (mword (8 * width))).
-Definition MEMr_reserved (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
+Definition MEMr_reserved (addr : mword 64) (width : Z) `{ArithFact (width >= (0 + 1))}
: M (MemoryOpResult (mword (8 * width))) :=
(checked_mem_read Data addr width)
: M (MemoryOpResult (mword (8 * width))).
-Definition MEMr_reserved_acquire (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
+Definition MEMr_reserved_acquire (addr : mword 64) (width : Z) `{ArithFact (width >= (0 + 1))}
: M (MemoryOpResult (mword (8 * width))) :=
(checked_mem_read Data addr width)
: M (MemoryOpResult (mword (8 * width))).
-Definition MEMr_reserved_strong_acquire (addr : mword 64) (width : Z) `{ArithFact (width >= 1)}
+Definition MEMr_reserved_strong_acquire (addr : mword 64) (width : Z) `{ArithFact (width >= (0 + 1))}
: M (MemoryOpResult (mword (8 * width))) :=
(checked_mem_read Data addr width)
: M (MemoryOpResult (mword (8 * width))).
Definition mem_read (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (res : bool) `{ArithFact (width >=
- 1)}
+ (0 + 1))}
: M (MemoryOpResult (mword (8 * width))) :=
(if ((andb (orb aq res) (negb (is_aligned_addr addr width)))) then
- returnm ((MemException E_Load_Addr_Align)
+ returnm ((MemException
+ (E_Load_Addr_Align))
: MemoryOpResult (mword (8 * width)))
else
(match (aq, rl, res) with
@@ -7006,51 +7703,42 @@ Definition mem_read (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (res :
| (true, false, true) =>
(MEMr_reserved_acquire addr width) : M (MemoryOpResult (mword (8 * width)))
| (false, true, false) =>
- (throw (Error_not_implemented "load.rl")) : M (MemoryOpResult (mword (8 * width)))
+ (throw (Error_not_implemented ("load.rl"))) : M (MemoryOpResult (mword (8 * width)))
| (true, true, false) =>
(MEMr_strong_acquire addr width) : M (MemoryOpResult (mword (8 * width)))
| (false, true, true) =>
- (throw (Error_not_implemented "lr.rl")) : M (MemoryOpResult (mword (8 * width)))
+ (throw (Error_not_implemented ("lr.rl"))) : M (MemoryOpResult (mword (8 * width)))
| (true, true, true) =>
(MEMr_reserved_strong_acquire addr width) : M (MemoryOpResult (mword (8 * width)))
end)
: M (MemoryOpResult (mword (8 * width))))
: M (MemoryOpResult (mword (8 * width))).
-Axiom MEMea : forall (_ : xlenbits) (n : Z) , M (unit).
-
-Axiom MEMea_release : forall (_ : xlenbits) (n : Z) , M (unit).
-
-Axiom MEMea_strong_release : forall (_ : xlenbits) (n : Z) , M (unit).
-
-Axiom MEMea_conditional : forall (_ : xlenbits) (n : Z) , M (unit).
-
-Axiom MEMea_conditional_release : forall (_ : xlenbits) (n : Z) , M (unit).
-
-Axiom MEMea_conditional_strong_release : forall (_ : xlenbits) (n : Z) , M (unit).
-
Definition mem_write_ea (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (con : bool)
: M (MemoryOpResult unit) :=
(if ((andb (orb rl con) (negb (is_aligned_addr addr width)))) then
- returnm ((MemException E_SAMO_Addr_Align)
+ returnm ((MemException
+ (E_SAMO_Addr_Align))
: MemoryOpResult unit)
else
(match (aq, rl, con) with
- | (false, false, false) => MEMea addr width >> returnm ((MemValue tt) : MemoryOpResult unit)
+ | (false, false, false) =>
+ MEMea addr width >> returnm ((MemValue (tt)) : MemoryOpResult unit)
| (false, true, false) =>
- MEMea_release addr width >> returnm ((MemValue tt) : MemoryOpResult unit)
+ MEMea_release addr width >> returnm ((MemValue (tt)) : MemoryOpResult unit)
| (false, false, true) =>
- MEMea_conditional addr width >> returnm ((MemValue tt) : MemoryOpResult unit)
+ MEMea_conditional addr width >> returnm ((MemValue (tt)) : MemoryOpResult unit)
| (false, true, true) =>
- MEMea_conditional_release addr width >> returnm ((MemValue tt) : MemoryOpResult unit)
+ MEMea_conditional_release addr width >> returnm ((MemValue (tt)) : MemoryOpResult unit)
| (true, false, false) =>
- (throw (Error_not_implemented "store.aq")) : M (MemoryOpResult unit)
+ (throw (Error_not_implemented ("store.aq"))) : M (MemoryOpResult unit)
| (true, true, false) =>
- MEMea_strong_release addr width >> returnm ((MemValue tt) : MemoryOpResult unit)
- | (true, false, true) => (throw (Error_not_implemented "sc.aq")) : M (MemoryOpResult unit)
+ MEMea_strong_release addr width >> returnm ((MemValue (tt)) : MemoryOpResult unit)
+ | (true, false, true) => (throw (Error_not_implemented ("sc.aq"))) : M (MemoryOpResult unit)
| (true, true, true) =>
MEMea_conditional_strong_release addr width >>
- returnm ((MemValue tt)
+ returnm ((MemValue
+ (tt))
: MemoryOpResult unit)
end)
: M (MemoryOpResult unit))
@@ -7064,11 +7752,12 @@ Definition phys_mem_write (addr : mword 64) (width : Z) (data : mword (8 * width
(String.append (string_of_bits addr) (String.append "] <- " (string_of_bits data)))))
: unit in
__RISCV_write addr width data >>= fun w__0 : bool =>
- returnm ((MemValue w__0)
+ returnm ((MemValue
+ (w__0))
: MemoryOpResult bool).
Definition checked_mem_write (addr : mword 64) (width : Z) (data : mword (8 * width)) `{ArithFact (width >=
- 1)}
+ (0 + 1))}
: M (MemoryOpResult bool) :=
(if ((within_mmio_writable addr width)) then
(mmio_write addr width data)
@@ -7076,49 +7765,51 @@ Definition checked_mem_write (addr : mword 64) (width : Z) (data : mword (8 * wi
else if ((within_phys_mem addr width)) then
(phys_mem_write addr width data)
: M (MemoryOpResult bool)
- else returnm ((MemException E_SAMO_Access_Fault) : MemoryOpResult bool))
+ else returnm ((MemException (E_SAMO_Access_Fault)) : MemoryOpResult bool))
: M (MemoryOpResult bool).
-Definition MEMval (addr : mword 64) (width : Z) (data : mword (8 * width)) `{ArithFact (width >= 1)}
+Definition MEMval (addr : mword 64) (width : Z) (data : mword (8 * width)) `{ArithFact (width >=
+ (0 + 1))}
: M (MemoryOpResult bool) :=
(checked_mem_write addr width data)
: M (MemoryOpResult bool).
Definition MEMval_release (addr : mword 64) (width : Z) (data : mword (8 * width)) `{ArithFact (width >=
- 1)}
+ (0 + 1))}
: M (MemoryOpResult bool) :=
(checked_mem_write addr width data)
: M (MemoryOpResult bool).
Definition MEMval_strong_release (addr : mword 64) (width : Z) (data : mword (8 * width)) `{ArithFact (width >=
- 1)}
+ (0 + 1))}
: M (MemoryOpResult bool) :=
(checked_mem_write addr width data)
: M (MemoryOpResult bool).
Definition MEMval_conditional (addr : mword 64) (width : Z) (data : mword (8 * width)) `{ArithFact (width >=
- 1)}
+ (0 + 1))}
: M (MemoryOpResult bool) :=
(checked_mem_write addr width data)
: M (MemoryOpResult bool).
Definition MEMval_conditional_release (addr : mword 64) (width : Z) (data : mword (8 * width)) `{ArithFact (width >=
- 1)}
+ (0 + 1))}
: M (MemoryOpResult bool) :=
(checked_mem_write addr width data)
: M (MemoryOpResult bool).
Definition MEMval_conditional_strong_release (addr : mword 64) (width : Z) (data : mword (8 * width)) `{ArithFact (width >=
- 1)}
+ (0 + 1))}
: M (MemoryOpResult bool) :=
(checked_mem_write addr width data)
: M (MemoryOpResult bool).
Definition mem_write_value (addr : mword 64) (width : Z) (value : mword (8 * width)) (aq : bool) (rl : bool) (con : bool) `{ArithFact (width >=
- 1)}
+ (0 + 1))}
: M (MemoryOpResult bool) :=
(if ((andb (orb rl con) (negb (is_aligned_addr addr width)))) then
- returnm ((MemException E_SAMO_Addr_Align)
+ returnm ((MemException
+ (E_SAMO_Addr_Align))
: MemoryOpResult bool)
else
(match (aq, rl, con) with
@@ -7128,36 +7819,17 @@ Definition mem_write_value (addr : mword 64) (width : Z) (value : mword (8 * wid
| (false, true, true) =>
(MEMval_conditional_release addr width value) : M (MemoryOpResult bool)
| (true, false, false) =>
- (throw (Error_not_implemented "store.aq")) : M (MemoryOpResult bool)
+ (throw (Error_not_implemented ("store.aq"))) : M (MemoryOpResult bool)
| (true, true, false) => (MEMval_strong_release addr width value) : M (MemoryOpResult bool)
- | (true, false, true) => (throw (Error_not_implemented "sc.aq")) : M (MemoryOpResult bool)
+ | (true, false, true) => (throw (Error_not_implemented ("sc.aq"))) : M (MemoryOpResult bool)
| (true, true, true) =>
(MEMval_conditional_strong_release addr width value) : M (MemoryOpResult bool)
end)
: M (MemoryOpResult bool))
: M (MemoryOpResult bool).
-Axiom MEM_fence_rw_rw : forall (_ : unit) , M (unit).
-
-Axiom MEM_fence_r_rw : forall (_ : unit) , M (unit).
-
-Axiom MEM_fence_r_r : forall (_ : unit) , M (unit).
-
-Axiom MEM_fence_rw_w : forall (_ : unit) , M (unit).
-
-Axiom MEM_fence_w_w : forall (_ : unit) , M (unit).
-
-Axiom MEM_fence_w_rw : forall (_ : unit) , M (unit).
-
-Axiom MEM_fence_rw_r : forall (_ : unit) , M (unit).
-
-Axiom MEM_fence_r_w : forall (_ : unit) , M (unit).
-
-Axiom MEM_fence_w_r : forall (_ : unit) , M (unit).
-
-Axiom MEM_fence_i : forall (_ : unit) , M (unit).
-
-Let PAGESIZE_BITS := 12.
+Definition PAGESIZE_BITS := 12.
+Hint Unfold PAGESIZE_BITS : sail.
Definition Mk_PTE_Bits (v : mword 8)
: PTE_Bits :=
{| PTE_Bits_PTE_Bits_chunk_0 := (subrange_vec_dec v 7 0) |}.
@@ -7421,7 +8093,8 @@ Definition update_PTE_Bits (p : PTE_Bits) (a : AccessType)
if ((orb update_d update_a)) then
let np := _update_PTE_Bits_A p ((bool_to_bits true) : mword 1) in
let np := if (update_d) then _update_PTE_Bits_D p ((bool_to_bits true) : mword 1) else np in
- Some np
+ Some
+ (np)
else None.
Definition PTW_Error_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 4)}
@@ -7435,13 +8108,13 @@ Definition PTW_Error_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 4)}
Definition num_of_PTW_Error (arg_ : PTW_Error)
: {e : Z & ArithFact (0 <= e /\ e <= 4)} :=
- match arg_ with
- | PTW_Access => build_ex 0
- | PTW_Invalid_PTE => build_ex 1
- | PTW_No_Permission => build_ex 2
- | PTW_Misaligned => build_ex 3
- | PTW_PTE_Update => build_ex 4
- end.
+ build_ex(match arg_ with
+ | PTW_Access => 0
+ | PTW_Invalid_PTE => 1
+ | PTW_No_Permission => 2
+ | PTW_Misaligned => 3
+ | PTW_PTE_Update => 4
+ end).
Definition ptw_error_to_str (e : PTW_Error)
: string :=
@@ -7466,10 +8139,14 @@ Definition translationException (a : AccessType) (f : PTW_Error)
| (Fetch, _) => E_Fetch_Page_Fault
end.
-Let SV39_LEVEL_BITS := 9.
-Let SV39_LEVELS := 3.
-Let PTE39_LOG_SIZE := 3.
-Let PTE39_SIZE := 8.
+Definition SV39_LEVEL_BITS := 9.
+Hint Unfold SV39_LEVEL_BITS : sail.
+Definition SV39_LEVELS := 3.
+Hint Unfold SV39_LEVELS : sail.
+Definition PTE39_LOG_SIZE := 3.
+Hint Unfold PTE39_LOG_SIZE : sail.
+Definition PTE39_SIZE := 8.
+Hint Unfold PTE39_SIZE : sail.
Definition Mk_SV39_Vaddr (v : mword 39)
: SV39_Vaddr :=
{| SV39_Vaddr_SV39_Vaddr_chunk_0 := (subrange_vec_dec v 38 0) |}.
@@ -7694,18 +8371,21 @@ Definition _update_SV39_PTE_BITS (v : SV39_PTE) (x : mword 8)
Definition curAsid64 '(tt : unit)
: M (mword 16) :=
- (read_reg satp_ref : M (mword 64)) >>= fun w__0 : mword 64 =>
+ (read_reg satp_ref : M (mword 64)) >>= fun w__0 : xlenbits =>
let satp64 := Mk_Satp64 w__0 in
returnm ((_get_Satp64_Asid satp64)
: mword 16).
Definition curPTB39 '(tt : unit)
: M (mword 56) :=
- (read_reg satp_ref : M (mword 64)) >>= fun w__0 : mword 64 =>
+ (read_reg satp_ref : M (mword 64)) >>= fun w__0 : xlenbits =>
let satp64 := Mk_Satp64 w__0 in
returnm ((EXTZ 56 (shiftl (_get_Satp64_PPN satp64) PAGESIZE_BITS))
: mword 56).
+Hint Unfold PAGESIZE_BITS : sail.
+Hint Unfold SV39_LEVEL_BITS : sail.
+
Fixpoint walk39 (vaddr : mword 39) (ac : AccessType) (priv : Privilege) (mxr : bool) (do_sum : bool) (ptb : mword 56) (level : nat) (global : bool)
: M (PTW_Result) :=
let va := Mk_SV39_Vaddr vaddr in
@@ -7715,19 +8395,20 @@ Fixpoint walk39 (vaddr : mword 39) (ac : AccessType) (priv : Privilege) (mxr : b
(subrange_vec_dec (shiftr (_get_SV39_Vaddr_VPNi va) (Z.mul (Z.of_nat level) SV39_LEVEL_BITS))
(projT1 (sub_range (build_ex SV39_LEVEL_BITS) (build_ex 1))) 0)) PTE39_LOG_SIZE in
let pte_addr := add_vec ptb pt_ofs in
- phys_mem_read Data (EXTZ 64 pte_addr) 8 false false false >>= fun w__0 : MemoryOpResult (mword 64) =>
+ phys_mem_read Data (EXTZ 64 pte_addr) 8 false false false >>= fun w__0 : MemoryOpResult (mword (8 * 8)) =>
(match w__0 with
- | MemException (_) => returnm ((PTW_Failure PTW_Access) : PTW_Result)
+ | MemException (_) => returnm ((PTW_Failure (PTW_Access)) : PTW_Result )
| MemValue (v) =>
let pte := Mk_SV39_PTE v in
let pbits := _get_SV39_PTE_BITS pte in
let pattr := Mk_PTE_Bits pbits in
let is_global := orb global (eq_vec (_get_PTE_Bits_G pattr) ((bool_to_bits true) : mword 1)) in
- (if ((isInvalidPTE pbits)) then returnm ((PTW_Failure PTW_Invalid_PTE) : PTW_Result)
+ (if ((isInvalidPTE pbits)) then returnm ((PTW_Failure (PTW_Invalid_PTE)) : PTW_Result )
else if ((isPTEPtr pbits)) then
- (match level with | O =>
- returnm ((PTW_Failure PTW_Invalid_PTE)
- : PTW_Result)
+ (match level with O =>
+ returnm ((PTW_Failure
+ (PTW_Invalid_PTE))
+ : PTW_Result )
| S level' =>
(walk39 vaddr ac priv mxr do_sum
(EXTZ 56 (shiftl (_get_SV39_PTE_PPNi pte) PAGESIZE_BITS))
@@ -7736,7 +8417,7 @@ Fixpoint walk39 (vaddr : mword 39) (ac : AccessType) (priv : Privilege) (mxr : b
: M (PTW_Result)
else
checkPTEPermission ac priv mxr do_sum pattr >>= fun w__3 : bool =>
- returnm ((if ((negb w__3)) then PTW_Failure PTW_No_Permission
+ returnm ((if ((negb w__3)) then PTW_Failure (PTW_No_Permission)
else if sumbool_of_bool (Nat.ltb O level) then
let mask :=
sub_vec_int
@@ -7747,30 +8428,31 @@ Fixpoint walk39 (vaddr : mword 39) (ac : AccessType) (priv : Privilege) (mxr : b
(Z.mul (Z.of_nat level) SV39_LEVEL_BITS)) 1 in
if ((neq_vec (and_vec (_get_SV39_PTE_PPNi pte) mask)
(EXTZ 44 (vec_of_bits [B0] : mword 1)))) then
- PTW_Failure PTW_Misaligned
+ PTW_Failure
+ (PTW_Misaligned)
else
let ppn :=
or_vec (_get_SV39_PTE_PPNi pte)
(and_vec (EXTZ 44 (_get_SV39_Vaddr_VPNi va)) mask) in
- PTW_Success (concat_vec ppn (_get_SV39_Vaddr_PgOfs va),pte,pte_addr,build_ex
- (Z.of_nat level),is_global)
+ PTW_Success
+ (concat_vec ppn (_get_SV39_Vaddr_PgOfs va), pte, pte_addr, build_ex (Z.of_nat level),
+ is_global)
else
- PTW_Success (concat_vec (_get_SV39_PTE_PPNi pte) (_get_SV39_Vaddr_PgOfs va),pte,pte_addr,build_ex
- (Z.of_nat level),is_global))
+ PTW_Success
+ (concat_vec (_get_SV39_PTE_PPNi pte) (_get_SV39_Vaddr_PgOfs va), pte,
+ pte_addr, build_ex (Z.of_nat level), is_global))
: PTW_Result))
: M (PTW_Result)
end)
: M (PTW_Result).
-Hint Unfold PAGESIZE_BITS : sail.
-Hint Unfold SV39_LEVEL_BITS : sail.
-
-Definition make_TLB39_Entry (asid : mword 16) (global : bool) (vAddr : mword 39) (pAddr : mword 56) (pte : SV39_PTE) '((existT _ level _) : {n : Z & ArithFact (n >=
+Definition make_TLB39_Entry (asid : mword 16) (global : bool) (vAddr : mword 39) (pAddr : mword 56) (pte : SV39_PTE) '(existT _ level _ : {n : Z & ArithFact (n >=
0)}) (pteAddr : mword 56)
: M (TLB39_Entry) :=
- let '(existT _ shift _) :=
- (build_ex (projT1 (add_range (build_ex PAGESIZE_BITS) (build_ex (Z.mul level SV39_LEVEL_BITS)))))
- : {n : Z & ArithFact (n >= 0)} in
+ let 'shift :=
+ projT1 ((build_ex (projT1 (add_range (build_ex PAGESIZE_BITS)
+ (build_ex (Z.mul level SV39_LEVEL_BITS)))))
+ : {n : Z & ArithFact (n >= 0)}) in
let vAddrMask : vaddr39 :=
sub_vec_int
(shiftl (xor_vec vAddr (xor_vec vAddr (EXTZ 39 (vec_of_bits [B1] : mword 1)))) shift) 1 in
@@ -7796,21 +8478,22 @@ Definition lookupTLB39 (asid : mword 16) (vaddr : mword 39)
if ((andb (orb e.(TLB39_Entry_global) (eq_vec e.(TLB39_Entry_asid) asid))
(eq_vec e.(TLB39_Entry_vAddr) (and_vec e.(TLB39_Entry_vMatchMask) vaddr))))
then
- Some (0, e)
+ Some
+ ((0, e))
else None
end)
: option ((Z * TLB39_Entry))).
-Definition addToTLB39 (asid : mword 16) (vAddr : mword 39) (pAddr : mword 56) (pte : SV39_PTE) (pteAddr : mword 56) '((existT _ level _) : {n : Z & ArithFact (n >=
+Definition addToTLB39 (asid : mword 16) (vAddr : mword 39) (pAddr : mword 56) (pte : SV39_PTE) (pteAddr : mword 56) '(existT _ level _ : {n : Z & ArithFact (n >=
0)}) (global : bool)
: M (unit) :=
make_TLB39_Entry asid global vAddr pAddr pte (build_ex level) pteAddr >>= fun ent =>
- write_reg tlb39_ref (Some ent)
+ write_reg tlb39_ref (Some (ent))
: M (unit).
Definition writeTLB39 (idx : Z) (ent : TLB39_Entry)
: M (unit) :=
- write_reg tlb39_ref (Some ent)
+ write_reg tlb39_ref (Some (ent))
: M (unit).
Definition flushTLB (asid : option (mword 16)) (addr : option (mword 39))
@@ -7822,21 +8505,21 @@ Definition flushTLB (asid : option (mword 16)) (addr : option (mword 39))
| (Some (e), None, None) => None
| (Some (e), None, Some (a)) =>
if ((eq_vec e.(TLB39_Entry_vAddr) (and_vec e.(TLB39_Entry_vMatchMask) a))) then None
- else Some e
+ else Some (e)
| (Some (e), Some (i), None) =>
if ((andb (eq_vec e.(TLB39_Entry_asid) i) (negb e.(TLB39_Entry_global)))) then None
- else Some e
+ else Some (e)
| (Some (e), Some (i), Some (a)) =>
if ((andb (eq_vec e.(TLB39_Entry_asid) i)
(andb (eq_vec e.(TLB39_Entry_vAddr) (and_vec a e.(TLB39_Entry_vMatchMask)))
(negb e.(TLB39_Entry_global))))) then
None
- else Some e
+ else Some (e)
end in
write_reg tlb39_ref ent
: M (unit).
-Definition translate39 (vAddr : mword 39) (ac : AccessType) (priv : Privilege) (mxr : bool) (do_sum : bool) '((existT _ level _) : {n : Z & ArithFact (n >=
+Definition translate39 (vAddr : mword 39) (ac : AccessType) (priv : Privilege) (mxr : bool) (do_sum : bool) '(existT _ level _ : {n : Z & ArithFact (n >=
0)})
: M (TR39_Result) :=
curAsid64 tt >>= fun asid =>
@@ -7845,17 +8528,19 @@ Definition translate39 (vAddr : mword 39) (ac : AccessType) (priv : Privilege) (
| Some (idx,ent) =>
let pteBits := Mk_PTE_Bits (_get_SV39_PTE_BITS ent.(TLB39_Entry_pte)) in
checkPTEPermission ac priv mxr do_sum pteBits >>= fun w__1 : bool =>
- (if ((negb w__1)) then returnm ((TR39_Failure PTW_No_Permission) : TR39_Result)
+ (if ((negb w__1)) then returnm ((TR39_Failure (PTW_No_Permission)) : TR39_Result )
else
(match (update_PTE_Bits pteBits ac) with
| None =>
- returnm ((TR39_Address (or_vec ent.(TLB39_Entry_pAddr)
- (EXTZ 56 (and_vec vAddr ent.(TLB39_Entry_vAddrMask)))))
- : TR39_Result)
+ returnm ((TR39_Address
+ (or_vec ent.(TLB39_Entry_pAddr)
+ (EXTZ 56 (and_vec vAddr ent.(TLB39_Entry_vAddrMask)))))
+ : TR39_Result )
| Some (pbits) =>
(if ((negb (plat_enable_dirty_update tt))) then
- returnm ((TR39_Failure PTW_PTE_Update)
- : TR39_Result)
+ returnm ((TR39_Failure
+ (PTW_PTE_Update))
+ : TR39_Result )
else
let n_ent : TLB39_Entry := ent in
let n_ent :=
@@ -7871,8 +8556,9 @@ Definition translate39 (vAddr : mword 39) (ac : AccessType) (priv : Privilege) (
| MemException (e) =>
(internal_error "invalid physical address in TLB") : M (unit)
end >>
- returnm ((TR39_Address (or_vec ent.(TLB39_Entry_pAddr)
- (EXTZ 56 (and_vec vAddr ent.(TLB39_Entry_vAddrMask)))))
+ returnm ((TR39_Address
+ (or_vec ent.(TLB39_Entry_pAddr)
+ (EXTZ 56 (and_vec vAddr ent.(TLB39_Entry_vAddrMask)))))
: TR39_Result))
: M (TR39_Result)
end)
@@ -7882,26 +8568,29 @@ Definition translate39 (vAddr : mword 39) (ac : AccessType) (priv : Privilege) (
curPTB39 tt >>= fun w__6 : mword 56 =>
walk39 vAddr ac priv mxr do_sum w__6 (Z.to_nat level) false >>= fun w__7 : PTW_Result =>
(match w__7 with
- | PTW_Failure (f) => returnm ((TR39_Failure f) : TR39_Result)
+ | PTW_Failure (f) => returnm ((TR39_Failure (f)) : TR39_Result )
| PTW_Success (pAddr,pte,pteAddr,(existT _ level _),global) =>
(match (update_PTE_Bits (Mk_PTE_Bits (_get_SV39_PTE_BITS pte)) ac) with
| None =>
addToTLB39 asid vAddr pAddr pte pteAddr (build_ex level) global >>
- returnm ((TR39_Address pAddr)
- : TR39_Result)
+ returnm ((TR39_Address
+ (pAddr))
+ : TR39_Result )
| Some (pbits) =>
(if ((negb (plat_enable_dirty_update tt))) then
- returnm ((TR39_Failure PTW_PTE_Update)
- : TR39_Result)
+ returnm ((TR39_Failure
+ (PTW_PTE_Update))
+ : TR39_Result )
else
let w_pte : SV39_PTE := _update_SV39_PTE_BITS pte (_get_PTE_Bits_bits pbits) in
checked_mem_write (EXTZ 64 pteAddr) 8 (_get_SV39_PTE_bits w_pte) >>= fun w__8 : MemoryOpResult bool =>
(match w__8 with
| MemValue (_) =>
addToTLB39 asid vAddr pAddr w_pte pteAddr (build_ex level) global >>
- returnm ((TR39_Address pAddr)
- : TR39_Result)
- | MemException (e) => returnm ((TR39_Failure PTW_Access) : TR39_Result)
+ returnm ((TR39_Address
+ (pAddr))
+ : TR39_Result )
+ | MemException (e) => returnm ((TR39_Failure (PTW_Access)) : TR39_Result )
end)
: M (TR39_Result))
: M (TR39_Result)
@@ -7923,7 +8612,7 @@ Definition translationMode (priv : Privilege)
let arch := architecture (_get_Mstatus_SXL w__0) in
(match arch with
| Some (RV64) =>
- (read_reg satp_ref : M (mword 64)) >>= fun w__1 : mword 64 =>
+ (read_reg satp_ref : M (mword 64)) >>= fun w__1 : xlenbits =>
let mbits : satp_mode := _get_Satp64_Mode (Mk_Satp64 w__1) in
(match (satpMode_of_bits RV64 mbits) with
| Some (m) => returnm (m : SATPMode)
@@ -7954,10 +8643,10 @@ Definition translateAddr (vAddr : mword 64) (ac : AccessType) (rt : ReadType)
let do_sum : bool := eq_vec (_get_Mstatus_SUM w__6) ((bool_to_bits true) : mword 1) in
translationMode effPriv >>= fun mode : SATPMode =>
(match mode with
- | Sbare => returnm ((TR_Address vAddr) : TR_Result)
+ | Sbare => returnm ((TR_Address (vAddr)) : TR_Result )
| SV39 =>
translate39 (subrange_vec_dec vAddr 38 0) ac effPriv mxr do_sum
- (build_ex ( (Z.sub ( SV39_LEVELS) ( 1)))) >>= fun w__7 : TR39_Result =>
+ (build_ex (projT1 (sub_range (build_ex SV39_LEVELS) (build_ex 1)))) >>= fun w__7 : TR39_Result =>
returnm ((match w__7 with
| TR39_Address (pa) => TR_Address (EXTZ 64 pa)
| TR39_Failure (f) => TR_Failure (translationException ac f)
@@ -7975,8 +8664,8 @@ Definition encdec_uop_forwards (arg_ : uop)
Definition encdec_uop_backwards (arg_ : mword 7)
: uop :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B1;B1;B0;B1;B1;B1] : mword 7))) then RISCV_LUI
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B0;B1;B1;B1] : mword 7))) then RISCV_LUI
else RISCV_AUIPC.
Definition encdec_uop_forwards_matches (arg_ : uop)
@@ -7985,9 +8674,9 @@ Definition encdec_uop_forwards_matches (arg_ : uop)
Definition encdec_uop_backwards_matches (arg_ : mword 7)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B1;B1;B0;B1;B1;B1] : mword 7))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1;B0;B1;B1;B1] : mword 7))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B0;B1;B1;B1] : mword 7))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B0;B1;B1;B1] : mword 7))) then true
else false.
Definition utype_mnemonic_forwards (arg_ : uop)
@@ -8013,19 +8702,26 @@ Definition utype_mnemonic_backwards_matches (arg_ : string)
Definition utype_mnemonic_matches_prefix (arg_ : string)
: option ((uop * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1771_ := arg_ in
- if ((andb (string_startswith _stringappend_1771_ "lui")
- (match (string_drop _stringappend_1771_ (string_length "lui")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1771_ (string_length "lui")) with
- | s_ => Some (RISCV_LUI, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1225_ := arg_ in
+ if ((andb (string_startswith _stringappend_1225_ "lui")
+ (match (string_drop _stringappend_1225_ (build_ex (projT1 (string_length "lui")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1225_ (build_ex (projT1 (string_length "lui")))) with
+ | s_ =>
+ Some
+ ((RISCV_LUI, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1771_ "auipc")
- (match (string_drop _stringappend_1771_ (string_length "auipc")) with
+ else if ((andb (string_startswith _stringappend_1225_ "auipc")
+ (match (string_drop _stringappend_1225_ (build_ex (projT1 (string_length "auipc")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1771_ (string_length "auipc")) with
- | s_ => Some (RISCV_AUIPC, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1225_ (build_ex (projT1 (string_length "auipc")))) with
+ | s_ =>
+ Some
+ ((RISCV_AUIPC, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -8042,12 +8738,12 @@ Definition encdec_bop_forwards (arg_ : bop)
Definition encdec_bop_backwards (arg_ : mword 3)
: bop :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B0;B0] : mword 3))) then RISCV_BEQ
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1] : mword 3))) then RISCV_BNE
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B0] : mword 3))) then RISCV_BLT
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B1] : mword 3))) then RISCV_BGE
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0] : mword 3))) then RISCV_BLTU
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then RISCV_BEQ
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then RISCV_BNE
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0] : mword 3))) then RISCV_BLT
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1] : mword 3))) then RISCV_BGE
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0] : mword 3))) then RISCV_BLTU
else RISCV_BGEU.
Definition encdec_bop_forwards_matches (arg_ : bop)
@@ -8063,13 +8759,13 @@ Definition encdec_bop_forwards_matches (arg_ : bop)
Definition encdec_bop_backwards_matches (arg_ : mword 3)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B0;B0] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B0] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B1] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B1] : mword 3))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1] : mword 3))) then true
else false.
Definition btype_mnemonic_forwards (arg_ : bop)
@@ -8121,42 +8817,66 @@ Definition btype_mnemonic_backwards_matches (arg_ : string)
Definition btype_mnemonic_matches_prefix (arg_ : string)
: option ((bop * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1765_ := arg_ in
- if ((andb (string_startswith _stringappend_1765_ "beq")
- (match (string_drop _stringappend_1765_ (string_length "beq")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1765_ (string_length "beq")) with
- | s_ => Some (RISCV_BEQ, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1219_ := arg_ in
+ if ((andb (string_startswith _stringappend_1219_ "beq")
+ (match (string_drop _stringappend_1219_ (build_ex (projT1 (string_length "beq")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1219_ (build_ex (projT1 (string_length "beq")))) with
+ | s_ =>
+ Some
+ ((RISCV_BEQ, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1765_ "bne")
- (match (string_drop _stringappend_1765_ (string_length "bne")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1765_ (string_length "bne")) with
- | s_ => Some (RISCV_BNE, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1219_ "bne")
+ (match (string_drop _stringappend_1219_ (build_ex (projT1 (string_length "bne")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1219_ (build_ex (projT1 (string_length "bne")))) with
+ | s_ =>
+ Some
+ ((RISCV_BNE, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1765_ "blt")
- (match (string_drop _stringappend_1765_ (string_length "blt")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1765_ (string_length "blt")) with
- | s_ => Some (RISCV_BLT, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1219_ "blt")
+ (match (string_drop _stringappend_1219_ (build_ex (projT1 (string_length "blt")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1219_ (build_ex (projT1 (string_length "blt")))) with
+ | s_ =>
+ Some
+ ((RISCV_BLT, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1765_ "bge")
- (match (string_drop _stringappend_1765_ (string_length "bge")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1765_ (string_length "bge")) with
- | s_ => Some (RISCV_BGE, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1219_ "bge")
+ (match (string_drop _stringappend_1219_ (build_ex (projT1 (string_length "bge")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1219_ (build_ex (projT1 (string_length "bge")))) with
+ | s_ =>
+ Some
+ ((RISCV_BGE, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1765_ "bltu")
- (match (string_drop _stringappend_1765_ (string_length "bltu")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1765_ (string_length "bltu")) with
- | s_ => Some (RISCV_BLTU, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1219_ "bltu")
+ (match (string_drop _stringappend_1219_ (build_ex (projT1 (string_length "bltu")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1219_ (build_ex (projT1 (string_length "bltu")))) with
+ | s_ =>
+ Some
+ ((RISCV_BLTU, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1765_ "bgeu")
- (match (string_drop _stringappend_1765_ (string_length "bgeu")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1765_ (string_length "bgeu")) with
- | s_ => Some (RISCV_BGEU, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1219_ "bgeu")
+ (match (string_drop _stringappend_1219_ (build_ex (projT1 (string_length "bgeu")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1219_ (build_ex (projT1 (string_length "bgeu")))) with
+ | s_ =>
+ Some
+ ((RISCV_BGEU, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -8173,12 +8893,12 @@ Definition encdec_iop_forwards (arg_ : iop)
Definition encdec_iop_backwards (arg_ : mword 3)
: iop :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B0;B0] : mword 3))) then RISCV_ADDI
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B0] : mword 3))) then RISCV_SLTI
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B1] : mword 3))) then RISCV_SLTIU
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B0] : mword 3))) then RISCV_XORI
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0] : mword 3))) then RISCV_ORI
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then RISCV_ADDI
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0] : mword 3))) then RISCV_SLTI
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1] : mword 3))) then RISCV_SLTIU
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0] : mword 3))) then RISCV_XORI
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0] : mword 3))) then RISCV_ORI
else RISCV_ANDI.
Definition encdec_iop_forwards_matches (arg_ : iop)
@@ -8194,13 +8914,13 @@ Definition encdec_iop_forwards_matches (arg_ : iop)
Definition encdec_iop_backwards_matches (arg_ : mword 3)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B0;B0] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B0] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B1] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B0] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B0] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1;B1] : mword 3))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1] : mword 3))) then true
else false.
Definition itype_mnemonic_forwards (arg_ : iop)
@@ -8252,43 +8972,66 @@ Definition itype_mnemonic_backwards_matches (arg_ : string)
Definition itype_mnemonic_matches_prefix (arg_ : string)
: option ((iop * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1759_ := arg_ in
- if ((andb (string_startswith _stringappend_1759_ "addi")
- (match (string_drop _stringappend_1759_ (string_length "addi")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1759_ (string_length "addi")) with
- | s_ => Some (RISCV_ADDI, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1213_ := arg_ in
+ if ((andb (string_startswith _stringappend_1213_ "addi")
+ (match (string_drop _stringappend_1213_ (build_ex (projT1 (string_length "addi")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1213_ (build_ex (projT1 (string_length "addi")))) with
+ | s_ =>
+ Some
+ ((RISCV_ADDI, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1759_ "slti")
- (match (string_drop _stringappend_1759_ (string_length "slti")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1759_ (string_length "slti")) with
- | s_ => Some (RISCV_SLTI, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1213_ "slti")
+ (match (string_drop _stringappend_1213_ (build_ex (projT1 (string_length "slti")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1213_ (build_ex (projT1 (string_length "slti")))) with
+ | s_ =>
+ Some
+ ((RISCV_SLTI, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1759_ "sltiu")
- (match (string_drop _stringappend_1759_ (string_length "sltiu")) with
+ else if ((andb (string_startswith _stringappend_1213_ "sltiu")
+ (match (string_drop _stringappend_1213_ (build_ex (projT1 (string_length "sltiu")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1759_ (string_length "sltiu")) with
- | s_ => Some (RISCV_SLTIU, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1213_ (build_ex (projT1 (string_length "sltiu")))) with
+ | s_ =>
+ Some
+ ((RISCV_SLTIU, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1759_ "xori")
- (match (string_drop _stringappend_1759_ (string_length "xori")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1759_ (string_length "xori")) with
- | s_ => Some (RISCV_XORI, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1213_ "xori")
+ (match (string_drop _stringappend_1213_ (build_ex (projT1 (string_length "xori")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1213_ (build_ex (projT1 (string_length "xori")))) with
+ | s_ =>
+ Some
+ ((RISCV_XORI, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1759_ "ori")
- (match (string_drop _stringappend_1759_ (string_length "ori")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1759_ (string_length "ori")) with
- | s_ => Some (RISCV_ORI, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1213_ "ori")
+ (match (string_drop _stringappend_1213_ (build_ex (projT1 (string_length "ori")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1213_ (build_ex (projT1 (string_length "ori")))) with
+ | s_ =>
+ Some
+ ((RISCV_ORI, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1759_ "andi")
- (match (string_drop _stringappend_1759_ (string_length "andi")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1759_ (string_length "andi")) with
- | s_ => Some (RISCV_ANDI, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1213_ "andi")
+ (match (string_drop _stringappend_1213_ (build_ex (projT1 (string_length "andi")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1213_ (build_ex (projT1 (string_length "andi")))) with
+ | s_ =>
+ Some
+ ((RISCV_ANDI, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -8302,9 +9045,9 @@ Definition encdec_sop_forwards (arg_ : sop)
Definition encdec_sop_backwards (arg_ : mword 3)
: sop :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B0;B1] : mword 3))) then RISCV_SLLI
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B1] : mword 3))) then RISCV_SRLI
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then RISCV_SLLI
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1] : mword 3))) then RISCV_SRLI
else RISCV_SRAI.
Definition encdec_sop_forwards_matches (arg_ : sop)
@@ -8313,10 +9056,10 @@ Definition encdec_sop_forwards_matches (arg_ : sop)
Definition encdec_sop_backwards_matches (arg_ : mword 3)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B0;B1] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B1] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B0;B1] : mword 3))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1] : mword 3))) then true
else false.
Definition shiftiop_mnemonic_forwards (arg_ : sop)
@@ -8343,24 +9086,36 @@ Definition shiftiop_mnemonic_backwards_matches (arg_ : string)
Definition shiftiop_mnemonic_matches_prefix (arg_ : string)
: option ((sop * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1756_ := arg_ in
- if ((andb (string_startswith _stringappend_1756_ "slli")
- (match (string_drop _stringappend_1756_ (string_length "slli")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1756_ (string_length "slli")) with
- | s_ => Some (RISCV_SLLI, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1210_ := arg_ in
+ if ((andb (string_startswith _stringappend_1210_ "slli")
+ (match (string_drop _stringappend_1210_ (build_ex (projT1 (string_length "slli")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1210_ (build_ex (projT1 (string_length "slli")))) with
+ | s_ =>
+ Some
+ ((RISCV_SLLI, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1756_ "srli")
- (match (string_drop _stringappend_1756_ (string_length "srli")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1756_ (string_length "srli")) with
- | s_ => Some (RISCV_SRLI, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1210_ "srli")
+ (match (string_drop _stringappend_1210_ (build_ex (projT1 (string_length "srli")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1210_ (build_ex (projT1 (string_length "srli")))) with
+ | s_ =>
+ Some
+ ((RISCV_SRLI, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1756_ "srai")
- (match (string_drop _stringappend_1756_ (string_length "srai")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1756_ (string_length "srai")) with
- | s_ => Some (RISCV_SRAI, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1210_ "srai")
+ (match (string_drop _stringappend_1210_ (build_ex (projT1 (string_length "srai")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1210_ (build_ex (projT1 (string_length "srai")))) with
+ | s_ =>
+ Some
+ ((RISCV_SRAI, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -8429,86 +9184,126 @@ Definition rtype_mnemonic_backwards_matches (arg_ : string)
Definition rtype_mnemonic_matches_prefix (arg_ : string)
: option ((rop * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1746_ := arg_ in
- if ((andb (string_startswith _stringappend_1746_ "add")
- (match (string_drop _stringappend_1746_ (string_length "add")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1746_ (string_length "add")) with
- | s_ => Some (RISCV_ADD, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1200_ := arg_ in
+ if ((andb (string_startswith _stringappend_1200_ "add")
+ (match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "add")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "add")))) with
+ | s_ =>
+ Some
+ ((RISCV_ADD, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1746_ "sub")
- (match (string_drop _stringappend_1746_ (string_length "sub")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1746_ (string_length "sub")) with
- | s_ => Some (RISCV_SUB, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1200_ "sub")
+ (match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "sub")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "sub")))) with
+ | s_ =>
+ Some
+ ((RISCV_SUB, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1746_ "sll")
- (match (string_drop _stringappend_1746_ (string_length "sll")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1746_ (string_length "sll")) with
- | s_ => Some (RISCV_SLL, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1200_ "sll")
+ (match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "sll")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "sll")))) with
+ | s_ =>
+ Some
+ ((RISCV_SLL, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1746_ "slt")
- (match (string_drop _stringappend_1746_ (string_length "slt")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1746_ (string_length "slt")) with
- | s_ => Some (RISCV_SLT, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1200_ "slt")
+ (match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "slt")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "slt")))) with
+ | s_ =>
+ Some
+ ((RISCV_SLT, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1746_ "sltu")
- (match (string_drop _stringappend_1746_ (string_length "sltu")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1746_ (string_length "sltu")) with
- | s_ => Some (RISCV_SLTU, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1200_ "sltu")
+ (match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "sltu")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "sltu")))) with
+ | s_ =>
+ Some
+ ((RISCV_SLTU, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1746_ "xor")
- (match (string_drop _stringappend_1746_ (string_length "xor")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1746_ (string_length "xor")) with
- | s_ => Some (RISCV_XOR, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1200_ "xor")
+ (match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "xor")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "xor")))) with
+ | s_ =>
+ Some
+ ((RISCV_XOR, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1746_ "srl")
- (match (string_drop _stringappend_1746_ (string_length "srl")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1746_ (string_length "srl")) with
- | s_ => Some (RISCV_SRL, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1200_ "srl")
+ (match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "srl")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "srl")))) with
+ | s_ =>
+ Some
+ ((RISCV_SRL, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1746_ "sra")
- (match (string_drop _stringappend_1746_ (string_length "sra")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1746_ (string_length "sra")) with
- | s_ => Some (RISCV_SRA, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1200_ "sra")
+ (match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "sra")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "sra")))) with
+ | s_ =>
+ Some
+ ((RISCV_SRA, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1746_ "or")
- (match (string_drop _stringappend_1746_ (string_length "or")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1746_ (string_length "or")) with
- | s_ => Some (RISCV_OR, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1200_ "or")
+ (match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "or")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "or")))) with
+ | s_ =>
+ Some
+ ((RISCV_OR, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1746_ "and")
- (match (string_drop _stringappend_1746_ (string_length "and")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1746_ (string_length "and")) with
- | s_ => Some (RISCV_AND, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1200_ "and")
+ (match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "and")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1200_ (build_ex (projT1 (string_length "and")))) with
+ | s_ =>
+ Some
+ ((RISCV_AND, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
-Definition extend_value {n : Z} (is_unsigned : bool) (value : MemoryOpResult (mword (8 * n))) `{ArithFact (1 <=
+Definition extend_value {n : Z} (is_unsigned : bool) (value : MemoryOpResult (mword (8 * n))) `{ArithFact ((0 + 1) <=
n /\
n <= 8)}
: MemoryOpResult (mword 64) :=
match value with
| MemValue (v) => MemValue (if (is_unsigned) then EXTZ 64 v else (EXTS 64 v) : xlenbits)
- | MemException (e) => MemException e
+ | MemException (e) => MemException (e)
end.
-Definition process_load {n : Z} (rd : mword 5) (addr : mword 64) (value : MemoryOpResult (mword (8 * n))) (is_unsigned : bool) `{ArithFact (1 <=
+Definition process_load {n : Z} (rd : mword 5) (addr : mword 64) (value : MemoryOpResult (mword (8 * n))) (is_unsigned : bool) `{ArithFact ((0 + 1) <=
n /\
n <= 8)}
: M (bool) :=
(match (extend_value is_unsigned value) with
| MemValue (result) =>
wX
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {syn_n : Z & ArithFact (0 <= syn_n /\ (syn_n + 1) <= 32)})) result >>
returnm (true
: bool)
@@ -8523,29 +9318,29 @@ Definition check_misaligned (vaddr : mword 64) (width : word_width)
(match width with
| BYTE => returnm (false : bool)
| HALF =>
- cast_unit_vec (access_vec_dec vaddr 0) >>= fun w__0 : mword 1 =>
- returnm ((eq_vec (w__0 : mword 1) ((bool_to_bits true) : mword 1))
+ bit_to_bool (access_vec_dec vaddr 0) >>= fun w__0 : bool =>
+ returnm ((Bool.eqb (w__0 : bool) true)
: bool)
| WORD =>
(or_boolM
- (cast_unit_vec (access_vec_dec vaddr 0) >>= fun w__1 : mword 1 =>
- returnm ((eq_vec (w__1 : mword 1) ((bool_to_bits true) : mword 1))
+ (bit_to_bool (access_vec_dec vaddr 0) >>= fun w__1 : bool =>
+ returnm ((Bool.eqb (w__1 : bool) true)
: bool))
- (cast_unit_vec (access_vec_dec vaddr 1) >>= fun w__2 : mword 1 =>
- returnm ((eq_vec (w__2 : mword 1) ((bool_to_bits true) : mword 1))
+ (bit_to_bool (access_vec_dec vaddr 1) >>= fun w__2 : bool =>
+ returnm ((Bool.eqb (w__2 : bool) true)
: bool)))
: M (bool)
| DOUBLE =>
(or_boolM
- (cast_unit_vec (access_vec_dec vaddr 0) >>= fun w__4 : mword 1 =>
- returnm ((eq_vec (w__4 : mword 1) ((bool_to_bits true) : mword 1))
+ (bit_to_bool (access_vec_dec vaddr 0) >>= fun w__4 : bool =>
+ returnm ((Bool.eqb (w__4 : bool) true)
: bool))
((or_boolM
- (cast_unit_vec (access_vec_dec vaddr 1) >>= fun w__5 : mword 1 =>
- returnm ((eq_vec (w__5 : mword 1) ((bool_to_bits true) : mword 1))
+ (bit_to_bool (access_vec_dec vaddr 1) >>= fun w__5 : bool =>
+ returnm ((Bool.eqb (w__5 : bool) true)
: bool))
- (cast_unit_vec (access_vec_dec vaddr 2) >>= fun w__6 : mword 1 =>
- returnm ((eq_vec (w__6 : mword 1) ((bool_to_bits true) : mword 1))
+ (bit_to_bool (access_vec_dec vaddr 2) >>= fun w__6 : bool =>
+ returnm ((Bool.eqb (w__6 : bool) true)
: bool)))
: M (bool)))
: M (bool)
@@ -8576,18 +9371,26 @@ Definition maybe_aq_backwards_matches (arg_ : string)
Definition maybe_aq_matches_prefix (arg_ : string)
: option ((bool * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1744_ := arg_ in
- if ((andb (string_startswith _stringappend_1744_ ".aq")
- (match (string_drop _stringappend_1744_ (string_length ".aq")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1744_ (string_length ".aq")) with
- | s_ => Some (true, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1198_ := arg_ in
+ if ((andb (string_startswith _stringappend_1198_ ".aq")
+ (match (string_drop _stringappend_1198_ (build_ex (projT1 (string_length ".aq")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1198_ (build_ex (projT1 (string_length ".aq")))) with
+ | s_ =>
+ Some
+ ((true, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1744_ "")
- (match (string_drop _stringappend_1744_ (string_length "")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1744_ (string_length "")) with
- | s_ => Some (false, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1198_ "")
+ (match (string_drop _stringappend_1198_ (build_ex (projT1 (string_length "")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1198_ (build_ex (projT1 (string_length "")))) with
+ | s_ =>
+ Some
+ ((false, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -8614,18 +9417,26 @@ Definition maybe_rl_backwards_matches (arg_ : string)
Definition maybe_rl_matches_prefix (arg_ : string)
: option ((bool * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1742_ := arg_ in
- if ((andb (string_startswith _stringappend_1742_ ".rl")
- (match (string_drop _stringappend_1742_ (string_length ".rl")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1742_ (string_length ".rl")) with
- | s_ => Some (true, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1196_ := arg_ in
+ if ((andb (string_startswith _stringappend_1196_ ".rl")
+ (match (string_drop _stringappend_1196_ (build_ex (projT1 (string_length ".rl")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1196_ (build_ex (projT1 (string_length ".rl")))) with
+ | s_ =>
+ Some
+ ((true, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1742_ "")
- (match (string_drop _stringappend_1742_ (string_length "")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1742_ (string_length "")) with
- | s_ => Some (false, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1196_ "")
+ (match (string_drop _stringappend_1196_ (build_ex (projT1 (string_length "")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1196_ (build_ex (projT1 (string_length "")))) with
+ | s_ =>
+ Some
+ ((false, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -8652,17 +9463,26 @@ Definition maybe_u_backwards_matches (arg_ : string)
Definition maybe_u_matches_prefix (arg_ : string)
: option ((bool * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1740_ := arg_ in
- if ((andb (string_startswith _stringappend_1740_ "u")
- (match (string_drop _stringappend_1740_ (string_length "u")) with | s_ => true end))) then
- match (string_drop _stringappend_1740_ (string_length "u")) with
- | s_ => Some (true, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1194_ := arg_ in
+ if ((andb (string_startswith _stringappend_1194_ "u")
+ (match (string_drop _stringappend_1194_ (build_ex (projT1 (string_length "u")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1194_ (build_ex (projT1 (string_length "u")))) with
+ | s_ =>
+ Some
+ ((true, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1740_ "")
- (match (string_drop _stringappend_1740_ (string_length "")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1740_ (string_length "")) with
- | s_ => Some (false, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1194_ "")
+ (match (string_drop _stringappend_1194_ (build_ex (projT1 (string_length "")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1194_ (build_ex (projT1 (string_length "")))) with
+ | s_ =>
+ Some
+ ((false, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -8690,24 +9510,36 @@ Definition shiftw_mnemonic_backwards_matches (arg_ : string)
Definition shiftw_mnemonic_matches_prefix (arg_ : string)
: option ((sop * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1737_ := arg_ in
- if ((andb (string_startswith _stringappend_1737_ "slli")
- (match (string_drop _stringappend_1737_ (string_length "slli")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1737_ (string_length "slli")) with
- | s_ => Some (RISCV_SLLI, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1191_ := arg_ in
+ if ((andb (string_startswith _stringappend_1191_ "slli")
+ (match (string_drop _stringappend_1191_ (build_ex (projT1 (string_length "slli")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1191_ (build_ex (projT1 (string_length "slli")))) with
+ | s_ =>
+ Some
+ ((RISCV_SLLI, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1737_ "srli")
- (match (string_drop _stringappend_1737_ (string_length "srli")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1737_ (string_length "srli")) with
- | s_ => Some (RISCV_SRLI, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1191_ "srli")
+ (match (string_drop _stringappend_1191_ (build_ex (projT1 (string_length "srli")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1191_ (build_ex (projT1 (string_length "srli")))) with
+ | s_ =>
+ Some
+ ((RISCV_SRLI, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1737_ "srai")
- (match (string_drop _stringappend_1737_ (string_length "srai")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1737_ (string_length "srai")) with
- | s_ => Some (RISCV_SRAI, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1191_ "srai")
+ (match (string_drop _stringappend_1191_ (build_ex (projT1 (string_length "srai")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1191_ (build_ex (projT1 (string_length "srai")))) with
+ | s_ =>
+ Some
+ ((RISCV_SRAI, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -8756,36 +9588,56 @@ Definition rtypew_mnemonic_backwards_matches (arg_ : string)
Definition rtypew_mnemonic_matches_prefix (arg_ : string)
: option ((ropw * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1732_ := arg_ in
- if ((andb (string_startswith _stringappend_1732_ "addw")
- (match (string_drop _stringappend_1732_ (string_length "addw")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1732_ (string_length "addw")) with
- | s_ => Some (RISCV_ADDW, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1186_ := arg_ in
+ if ((andb (string_startswith _stringappend_1186_ "addw")
+ (match (string_drop _stringappend_1186_ (build_ex (projT1 (string_length "addw")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1186_ (build_ex (projT1 (string_length "addw")))) with
+ | s_ =>
+ Some
+ ((RISCV_ADDW, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1732_ "subw")
- (match (string_drop _stringappend_1732_ (string_length "subw")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1732_ (string_length "subw")) with
- | s_ => Some (RISCV_SUBW, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1186_ "subw")
+ (match (string_drop _stringappend_1186_ (build_ex (projT1 (string_length "subw")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1186_ (build_ex (projT1 (string_length "subw")))) with
+ | s_ =>
+ Some
+ ((RISCV_SUBW, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1732_ "sllw")
- (match (string_drop _stringappend_1732_ (string_length "sllw")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1732_ (string_length "sllw")) with
- | s_ => Some (RISCV_SLLW, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1186_ "sllw")
+ (match (string_drop _stringappend_1186_ (build_ex (projT1 (string_length "sllw")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1186_ (build_ex (projT1 (string_length "sllw")))) with
+ | s_ =>
+ Some
+ ((RISCV_SLLW, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1732_ "srlw")
- (match (string_drop _stringappend_1732_ (string_length "srlw")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1732_ (string_length "srlw")) with
- | s_ => Some (RISCV_SRLW, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1186_ "srlw")
+ (match (string_drop _stringappend_1186_ (build_ex (projT1 (string_length "srlw")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1186_ (build_ex (projT1 (string_length "srlw")))) with
+ | s_ =>
+ Some
+ ((RISCV_SRLW, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1732_ "sraw")
- (match (string_drop _stringappend_1732_ (string_length "sraw")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1732_ (string_length "sraw")) with
- | s_ => Some (RISCV_SRAW, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1186_ "sraw")
+ (match (string_drop _stringappend_1186_ (build_ex (projT1 (string_length "sraw")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1186_ (build_ex (projT1 (string_length "sraw")))) with
+ | s_ =>
+ Some
+ ((RISCV_SRAW, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -8813,26 +9665,36 @@ Definition shiftiwop_mnemonic_backwards_matches (arg_ : string)
Definition shiftiwop_mnemonic_matches_prefix (arg_ : string)
: option ((sopw * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1729_ := arg_ in
- if ((andb (string_startswith _stringappend_1729_ "slliw")
- (match (string_drop _stringappend_1729_ (string_length "slliw")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1729_ (string_length "slliw")) with
- | s_ => Some (RISCV_SLLIW, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1183_ := arg_ in
+ if ((andb (string_startswith _stringappend_1183_ "slliw")
+ (match (string_drop _stringappend_1183_ (build_ex (projT1 (string_length "slliw")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1183_ (build_ex (projT1 (string_length "slliw")))) with
+ | s_ =>
+ Some
+ ((RISCV_SLLIW, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1729_ "srliw")
- (match (string_drop _stringappend_1729_ (string_length "srliw")) with
+ else if ((andb (string_startswith _stringappend_1183_ "srliw")
+ (match (string_drop _stringappend_1183_ (build_ex (projT1 (string_length "srliw")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1729_ (string_length "srliw")) with
- | s_ => Some (RISCV_SRLIW, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1183_ (build_ex (projT1 (string_length "srliw")))) with
+ | s_ =>
+ Some
+ ((RISCV_SRLIW, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1729_ "sraiw")
- (match (string_drop _stringappend_1729_ (string_length "sraiw")) with
+ else if ((andb (string_startswith _stringappend_1183_ "sraiw")
+ (match (string_drop _stringappend_1183_ (build_ex (projT1 (string_length "sraiw")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1729_ (string_length "sraiw")) with
- | s_ => Some (RISCV_SRAIW, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1183_ (build_ex (projT1 (string_length "sraiw")))) with
+ | s_ =>
+ Some
+ ((RISCV_SRAIW, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -8850,10 +9712,10 @@ Definition encdec_mul_op_forwards (arg0 : bool) (arg1 : bool) (arg2 : bool)
Definition encdec_mul_op_backwards (arg_ : mword 3)
: (bool * bool * bool) :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B0;B0] : mword 3))) then (false, true, true)
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1] : mword 3))) then (true, true, true)
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B0] : mword 3))) then (true, true, false)
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then (false, true, true)
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then (true, true, true)
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0] : mword 3))) then (true, true, false)
else (true, false, false).
Definition encdec_mul_op_forwards_matches (arg0 : bool) (arg1 : bool) (arg2 : bool)
@@ -8869,11 +9731,11 @@ Definition encdec_mul_op_forwards_matches (arg0 : bool) (arg1 : bool) (arg2 : bo
Definition encdec_mul_op_backwards_matches (arg_ : mword 3)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B0;B0] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B0;B1] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B0] : mword 3))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0;B1;B1] : mword 3))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0] : mword 3))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1] : mword 3))) then true
else false.
Definition mul_mnemonic_forwards (arg0 : bool) (arg1 : bool) (arg2 : bool)
@@ -8891,10 +9753,10 @@ Definition mul_mnemonic_forwards (arg0 : bool) (arg1 : bool) (arg2 : bool)
Definition mul_mnemonic_backwards (arg_ : string)
: M ((bool * bool * bool)) :=
(match arg_ with
- | "mul" => returnm ((false, true, true) : (bool * bool * bool))
- | "mulh" => returnm ((true, true, true) : (bool * bool * bool))
- | "mulhsu" => returnm ((true, true, false) : (bool * bool * bool))
- | "mulhu" => returnm ((true, false, false) : (bool * bool * bool))
+ | "mul" => returnm (false, true, true)
+ | "mulh" => returnm (true, true, true)
+ | "mulhsu" => returnm (true, true, false)
+ | "mulhu" => returnm (true, false, false)
| _ => exit tt : M ((bool * bool * bool))
end)
: M ((bool * bool * bool)).
@@ -8922,32 +9784,46 @@ Definition mul_mnemonic_backwards_matches (arg_ : string)
Definition mul_mnemonic_matches_prefix (arg_ : string)
: option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1725_ := arg_ in
- if ((andb (string_startswith _stringappend_1725_ "mul")
- (match (string_drop _stringappend_1725_ (string_length "mul")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1725_ (string_length "mul")) with
- | s_ => Some ((false, true, true), sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1179_ := arg_ in
+ if ((andb (string_startswith _stringappend_1179_ "mul")
+ (match (string_drop _stringappend_1179_ (build_ex (projT1 (string_length "mul")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1179_ (build_ex (projT1 (string_length "mul")))) with
+ | s_ =>
+ Some
+ (((false, true, true), build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1725_ "mulh")
- (match (string_drop _stringappend_1725_ (string_length "mulh")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1725_ (string_length "mulh")) with
- | s_ => Some ((true, true, true), sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1179_ "mulh")
+ (match (string_drop _stringappend_1179_ (build_ex (projT1 (string_length "mulh")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1179_ (build_ex (projT1 (string_length "mulh")))) with
+ | s_ =>
+ Some
+ (((true, true, true), build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1725_ "mulhsu")
- (match (string_drop _stringappend_1725_ (string_length "mulhsu")) with
+ else if ((andb (string_startswith _stringappend_1179_ "mulhsu")
+ (match (string_drop _stringappend_1179_ (build_ex (projT1 (string_length "mulhsu")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1725_ (string_length "mulhsu")) with
- | s_ => Some ((true, true, false), sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1179_ (build_ex (projT1 (string_length "mulhsu")))) with
+ | s_ =>
+ Some
+ (((true, true, false), build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1725_ "mulhu")
- (match (string_drop _stringappend_1725_ (string_length "mulhu")) with
+ else if ((andb (string_startswith _stringappend_1179_ "mulhu")
+ (match (string_drop _stringappend_1179_ (build_ex (projT1 (string_length "mulhu")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1725_ (string_length "mulhu")) with
- | s_ => Some ((true, false, false), sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1179_ (build_ex (projT1 (string_length "mulhu")))) with
+ | s_ =>
+ Some
+ (((true, false, false), build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -8974,24 +9850,33 @@ Definition maybe_not_u_backwards_matches (arg_ : string)
Definition maybe_not_u_matches_prefix (arg_ : string)
: option ((bool * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1723_ := arg_ in
- if ((andb (string_startswith _stringappend_1723_ "u")
- (match (string_drop _stringappend_1723_ (string_length "u")) with | s_ => true end))) then
- match (string_drop _stringappend_1723_ (string_length "u")) with
- | s_ => Some (false, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1177_ := arg_ in
+ if ((andb (string_startswith _stringappend_1177_ "u")
+ (match (string_drop _stringappend_1177_ (build_ex (projT1 (string_length "u")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1177_ (build_ex (projT1 (string_length "u")))) with
+ | s_ =>
+ Some
+ ((false, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1723_ "")
- (match (string_drop _stringappend_1723_ (string_length "")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1723_ (string_length "")) with
- | s_ => Some (true, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1177_ "")
+ (match (string_drop _stringappend_1177_ (build_ex (projT1 (string_length "")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1177_ (build_ex (projT1 (string_length "")))) with
+ | s_ =>
+ Some
+ ((true, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
Definition bit_maybe_r_forwards (arg_ : mword 1)
: string :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B1] : mword 1))) then "r"
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then "r"
else "".
Definition bit_maybe_r_backwards (arg_ : string)
@@ -9005,9 +9890,9 @@ Definition bit_maybe_r_backwards (arg_ : string)
Definition bit_maybe_r_forwards_matches (arg_ : mword 1)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B1] : mword 1))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0] : mword 1))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then true
else false.
Definition bit_maybe_r_backwards_matches (arg_ : string)
@@ -9016,24 +9901,35 @@ Definition bit_maybe_r_backwards_matches (arg_ : string)
Definition bit_maybe_r_matches_prefix (arg_ : string)
: option ((mword 1 * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1721_ := arg_ in
- if ((andb (string_startswith _stringappend_1721_ "r")
- (match (string_drop _stringappend_1721_ (string_length "r")) with | s_ => true end))) then
- match (string_drop _stringappend_1721_ (string_length "r")) with
- | s_ => Some ((vec_of_bits [B1] : mword 1), sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1175_ := arg_ in
+ if ((andb (string_startswith _stringappend_1175_ "r")
+ (match (string_drop _stringappend_1175_ (build_ex (projT1 (string_length "r")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1175_ (build_ex (projT1 (string_length "r")))) with
+ | s_ =>
+ Some
+ (((vec_of_bits [B1] : mword 1), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1721_ "")
- (match (string_drop _stringappend_1721_ (string_length "")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1721_ (string_length "")) with
- | s_ => Some ((vec_of_bits [B0] : mword 1), sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1175_ "")
+ (match (string_drop _stringappend_1175_ (build_ex (projT1 (string_length "")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1175_ (build_ex (projT1 (string_length "")))) with
+ | s_ =>
+ Some
+ (((vec_of_bits [B0] : mword 1), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
Definition bit_maybe_w_forwards (arg_ : mword 1)
: string :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B1] : mword 1))) then "w"
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then "w"
else "".
Definition bit_maybe_w_backwards (arg_ : string)
@@ -9047,9 +9943,9 @@ Definition bit_maybe_w_backwards (arg_ : string)
Definition bit_maybe_w_forwards_matches (arg_ : mword 1)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B1] : mword 1))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0] : mword 1))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then true
else false.
Definition bit_maybe_w_backwards_matches (arg_ : string)
@@ -9058,24 +9954,35 @@ Definition bit_maybe_w_backwards_matches (arg_ : string)
Definition bit_maybe_w_matches_prefix (arg_ : string)
: option ((mword 1 * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1719_ := arg_ in
- if ((andb (string_startswith _stringappend_1719_ "w")
- (match (string_drop _stringappend_1719_ (string_length "w")) with | s_ => true end))) then
- match (string_drop _stringappend_1719_ (string_length "w")) with
- | s_ => Some ((vec_of_bits [B1] : mword 1), sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1173_ := arg_ in
+ if ((andb (string_startswith _stringappend_1173_ "w")
+ (match (string_drop _stringappend_1173_ (build_ex (projT1 (string_length "w")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1173_ (build_ex (projT1 (string_length "w")))) with
+ | s_ =>
+ Some
+ (((vec_of_bits [B1] : mword 1), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1719_ "")
- (match (string_drop _stringappend_1719_ (string_length "")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1719_ (string_length "")) with
- | s_ => Some ((vec_of_bits [B0] : mword 1), sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1173_ "")
+ (match (string_drop _stringappend_1173_ (build_ex (projT1 (string_length "")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1173_ (build_ex (projT1 (string_length "")))) with
+ | s_ =>
+ Some
+ (((vec_of_bits [B0] : mword 1), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
Definition bit_maybe_i_forwards (arg_ : mword 1)
: string :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B1] : mword 1))) then "i"
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then "i"
else "".
Definition bit_maybe_i_backwards (arg_ : string)
@@ -9089,9 +9996,9 @@ Definition bit_maybe_i_backwards (arg_ : string)
Definition bit_maybe_i_forwards_matches (arg_ : mword 1)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B1] : mword 1))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0] : mword 1))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then true
else false.
Definition bit_maybe_i_backwards_matches (arg_ : string)
@@ -9100,24 +10007,35 @@ Definition bit_maybe_i_backwards_matches (arg_ : string)
Definition bit_maybe_i_matches_prefix (arg_ : string)
: option ((mword 1 * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1717_ := arg_ in
- if ((andb (string_startswith _stringappend_1717_ "i")
- (match (string_drop _stringappend_1717_ (string_length "i")) with | s_ => true end))) then
- match (string_drop _stringappend_1717_ (string_length "i")) with
- | s_ => Some ((vec_of_bits [B1] : mword 1), sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1171_ := arg_ in
+ if ((andb (string_startswith _stringappend_1171_ "i")
+ (match (string_drop _stringappend_1171_ (build_ex (projT1 (string_length "i")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1171_ (build_ex (projT1 (string_length "i")))) with
+ | s_ =>
+ Some
+ (((vec_of_bits [B1] : mword 1), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1717_ "")
- (match (string_drop _stringappend_1717_ (string_length "")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1717_ (string_length "")) with
- | s_ => Some ((vec_of_bits [B0] : mword 1), sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1171_ "")
+ (match (string_drop _stringappend_1171_ (build_ex (projT1 (string_length "")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1171_ (build_ex (projT1 (string_length "")))) with
+ | s_ =>
+ Some
+ (((vec_of_bits [B0] : mword 1), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
Definition bit_maybe_o_forwards (arg_ : mword 1)
: string :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B1] : mword 1))) then "o"
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then "o"
else "".
Definition bit_maybe_o_backwards (arg_ : string)
@@ -9131,9 +10049,9 @@ Definition bit_maybe_o_backwards (arg_ : string)
Definition bit_maybe_o_forwards_matches (arg_ : mword 1)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B1] : mword 1))) then true
- else if ((eq_vec p0_ (vec_of_bits [B0] : mword 1))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then true
else false.
Definition bit_maybe_o_backwards_matches (arg_ : string)
@@ -9142,17 +10060,28 @@ Definition bit_maybe_o_backwards_matches (arg_ : string)
Definition bit_maybe_o_matches_prefix (arg_ : string)
: option ((mword 1 * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1715_ := arg_ in
- if ((andb (string_startswith _stringappend_1715_ "o")
- (match (string_drop _stringappend_1715_ (string_length "o")) with | s_ => true end))) then
- match (string_drop _stringappend_1715_ (string_length "o")) with
- | s_ => Some ((vec_of_bits [B1] : mword 1), sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1169_ := arg_ in
+ if ((andb (string_startswith _stringappend_1169_ "o")
+ (match (string_drop _stringappend_1169_ (build_ex (projT1 (string_length "o")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1169_ (build_ex (projT1 (string_length "o")))) with
+ | s_ =>
+ Some
+ (((vec_of_bits [B1] : mword 1), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1715_ "")
- (match (string_drop _stringappend_1715_ (string_length "")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1715_ (string_length "")) with
- | s_ => Some ((vec_of_bits [B0] : mword 1), sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1169_ "")
+ (match (string_drop _stringappend_1169_ (build_ex (projT1 (string_length "")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1169_ (build_ex (projT1 (string_length "")))) with
+ | s_ =>
+ Some
+ (((vec_of_bits [B0] : mword 1), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -9160,54 +10089,43 @@ Definition fence_bits_forwards (arg_ : mword 4)
: string :=
match arg_ with
| v__0 =>
- let r : bits 1 := subrange_vec_dec v__0 3 3 in
- let w : bits 1 := subrange_vec_dec v__0 2 2 in
- let i : bits 1 := subrange_vec_dec v__0 1 1 in
- let o : bits 1 := subrange_vec_dec v__0 0 0 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 o) "")))
+ let i : bits 1 := subrange_vec_dec v__0 3 3 in
+ let w : bits 1 := subrange_vec_dec v__0 0 0 in
+ let r : bits 1 := subrange_vec_dec v__0 1 1 in
+ let o : bits 1 := subrange_vec_dec v__0 2 2 in
+ let i : bits 1 := subrange_vec_dec v__0 3 3 in
+ string_append (bit_maybe_i_forwards i)
+ (string_append (bit_maybe_o_forwards o)
+ (string_append (bit_maybe_r_forwards r) (string_append (bit_maybe_w_forwards w) "")))
end.
Definition fence_bits_backwards (arg_ : string)
: M (mword 4) :=
- let _stringappend_1703_ := arg_ in
- match (bit_maybe_r_matches_prefix _stringappend_1703_) with
- | Some (_stringappend_1704_,(existT _ _stringappend_1705_ _)) =>
- returnm ((_stringappend_1704_, build_ex _stringappend_1705_)
- : (mword 1 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__1 : (mword 1 * {n : Z & ArithFact (n >= 0)}) =>
- let '(r, existT _ _stringappend_1705_ _) := w__1 : (mword 1 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1706_ := string_drop _stringappend_1703_ (build_ex _stringappend_1705_) in
- match (bit_maybe_w_matches_prefix _stringappend_1706_) with
- | Some (_stringappend_1707_,(existT _ _stringappend_1708_ _)) =>
- returnm ((_stringappend_1707_, build_ex _stringappend_1708_)
- : (mword 1 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__3 : (mword 1 * {n : Z & ArithFact (n >= 0)}) =>
- let '(w, existT _ _stringappend_1708_ _) := w__3 : (mword 1 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1709_ := string_drop _stringappend_1706_ (build_ex _stringappend_1708_) in
- match (bit_maybe_i_matches_prefix _stringappend_1709_) with
- | Some (_stringappend_1710_,(existT _ _stringappend_1711_ _)) =>
- returnm ((_stringappend_1710_, build_ex _stringappend_1711_)
- : (mword 1 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__5 : (mword 1 * {n : Z & ArithFact (n >= 0)}) =>
- let '(i, existT _ _stringappend_1711_ _) := w__5 : (mword 1 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1712_ := string_drop _stringappend_1709_ (build_ex _stringappend_1711_) in
- match (bit_maybe_o_matches_prefix _stringappend_1712_) with
- | Some (_stringappend_1713_,(existT _ _stringappend_1714_ _)) =>
- returnm ((_stringappend_1713_, build_ex _stringappend_1714_)
- : (mword 1 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__7 : (mword 1 * {n : Z & ArithFact (n >= 0)}) =>
- let '(o, existT _ _stringappend_1714_ _) := w__7 : (mword 1 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1712_ (build_ex _stringappend_1714_)) with
+ let _stringappend_1161_ := arg_ in
+ (match (bit_maybe_i_matches_prefix _stringappend_1161_) with
+ | Some (i,(existT _ _stringappend_1162_ _)) => returnm (i, build_ex _stringappend_1162_)
+ | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(i, existT _ _stringappend_1162_ _) =>
+ let _stringappend_1163_ := string_drop _stringappend_1161_ (build_ex _stringappend_1162_) in
+ (match (bit_maybe_o_matches_prefix _stringappend_1163_) with
+ | Some (o,(existT _ _stringappend_1164_ _)) => returnm (o, build_ex _stringappend_1164_)
+ | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(o, existT _ _stringappend_1164_ _) =>
+ let _stringappend_1165_ := string_drop _stringappend_1163_ (build_ex _stringappend_1164_) in
+ (match (bit_maybe_r_matches_prefix _stringappend_1165_) with
+ | Some (r,(existT _ _stringappend_1166_ _)) => returnm (r, build_ex _stringappend_1166_)
+ | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(r, existT _ _stringappend_1166_ _) =>
+ let _stringappend_1167_ := string_drop _stringappend_1165_ (build_ex _stringappend_1166_) in
+ (match (bit_maybe_w_matches_prefix _stringappend_1167_) with
+ | Some (w,(existT _ _stringappend_1168_ _)) => returnm (w, build_ex _stringappend_1168_)
+ | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(w, existT _ _stringappend_1168_ _) =>
+ (match (string_drop _stringappend_1167_ (build_ex _stringappend_1168_)) with
| "" =>
- returnm ((concat_vec (r : bits 1)
- (concat_vec (w : bits 1) (concat_vec (i : bits 1) (o : bits 1))))
- : mword 4)
+ returnm ((concat_vec (i : bits 1)
+ (concat_vec (o : bits 1) (concat_vec (r : bits 1) (w : bits 1))))
+ : mword (1 + 3))
| _ => exit tt : M (mword 4)
end)
: M (mword 4).
@@ -9218,24 +10136,23 @@ Definition fence_bits_forwards_matches (arg_ : mword 4)
Definition fence_bits_backwards_matches (arg_ : string)
: M (bool) :=
- let _stringappend_1691_ := arg_ in
- (if ((match (bit_maybe_r_matches_prefix _stringappend_1691_) with
- | Some (_stringappend_1692_,(existT _ _stringappend_1693_ _)) =>
- let _stringappend_1694_ :=
- string_drop _stringappend_1691_ (build_ex _stringappend_1693_) in
- if ((match (bit_maybe_w_matches_prefix _stringappend_1694_) with
- | Some (_stringappend_1695_,(existT _ _stringappend_1696_ _)) =>
- let _stringappend_1697_ :=
- string_drop _stringappend_1694_ (build_ex _stringappend_1696_) in
- if ((match (bit_maybe_i_matches_prefix _stringappend_1697_) with
- | Some (_stringappend_1698_,(existT _ _stringappend_1699_ _)) =>
- let _stringappend_1700_ :=
- string_drop _stringappend_1697_ (build_ex _stringappend_1699_) in
- if ((match (bit_maybe_o_matches_prefix _stringappend_1700_) with
- | Some (_stringappend_1701_,(existT _ _stringappend_1702_ _)) =>
- match (string_drop _stringappend_1700_
- (build_ex
- _stringappend_1702_)) with
+ let _stringappend_1153_ := arg_ in
+ (if ((match (bit_maybe_i_matches_prefix _stringappend_1153_) with
+ | Some (i,(existT _ _stringappend_1154_ _)) =>
+ let _stringappend_1155_ :=
+ string_drop _stringappend_1153_ (build_ex _stringappend_1154_) in
+ if ((match (bit_maybe_o_matches_prefix _stringappend_1155_) with
+ | Some (o,(existT _ _stringappend_1156_ _)) =>
+ let _stringappend_1157_ :=
+ string_drop _stringappend_1155_ (build_ex _stringappend_1156_) in
+ if ((match (bit_maybe_r_matches_prefix _stringappend_1157_) with
+ | Some (r,(existT _ _stringappend_1158_ _)) =>
+ let _stringappend_1159_ :=
+ string_drop _stringappend_1157_ (build_ex _stringappend_1158_) in
+ if ((match (bit_maybe_w_matches_prefix _stringappend_1159_) with
+ | Some (w,(existT _ _stringappend_1160_ _)) =>
+ match (string_drop _stringappend_1159_
+ (build_ex _stringappend_1160_)) with
| "" => true
| _ => false
end
@@ -9253,38 +10170,26 @@ Definition fence_bits_backwards_matches (arg_ : string)
else false
| None => false
end)) then
- match (bit_maybe_r_matches_prefix _stringappend_1691_) with
- | Some (_stringappend_1692_,(existT _ _stringappend_1693_ _)) =>
- returnm ((_stringappend_1692_, build_ex _stringappend_1693_)
- : (mword 1 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__1 : (mword 1 * {n : Z & ArithFact (n >= 0)}) =>
- let '(r, existT _ _stringappend_1693_ _) := w__1 : (mword 1 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1694_ := string_drop _stringappend_1691_ (build_ex _stringappend_1693_) in
- match (bit_maybe_w_matches_prefix _stringappend_1694_) with
- | Some (_stringappend_1695_,(existT _ _stringappend_1696_ _)) =>
- returnm ((_stringappend_1695_, build_ex _stringappend_1696_)
- : (mword 1 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__3 : (mword 1 * {n : Z & ArithFact (n >= 0)}) =>
- let '(w, existT _ _stringappend_1696_ _) := w__3 : (mword 1 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1697_ := string_drop _stringappend_1694_ (build_ex _stringappend_1696_) in
- match (bit_maybe_i_matches_prefix _stringappend_1697_) with
- | Some (_stringappend_1698_,(existT _ _stringappend_1699_ _)) =>
- returnm ((_stringappend_1698_, build_ex _stringappend_1699_)
- : (mword 1 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__5 : (mword 1 * {n : Z & ArithFact (n >= 0)}) =>
- let '(i, existT _ _stringappend_1699_ _) := w__5 : (mword 1 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1700_ := string_drop _stringappend_1697_ (build_ex _stringappend_1699_) in
- match (bit_maybe_o_matches_prefix _stringappend_1700_) with
- | Some (_stringappend_1701_,(existT _ _stringappend_1702_ _)) =>
- returnm ((_stringappend_1701_, build_ex _stringappend_1702_)
- : (mword 1 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__7 : (mword 1 * {n : Z & ArithFact (n >= 0)}) =>
- let '(o, existT _ _stringappend_1702_ _) := w__7 : (mword 1 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1700_ (build_ex _stringappend_1702_)) with
+ (match (bit_maybe_i_matches_prefix _stringappend_1153_) with
+ | Some (i,(existT _ _stringappend_1154_ _)) => returnm (i, build_ex _stringappend_1154_)
+ | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(i, existT _ _stringappend_1154_ _) =>
+ let _stringappend_1155_ := string_drop _stringappend_1153_ (build_ex _stringappend_1154_) in
+ (match (bit_maybe_o_matches_prefix _stringappend_1155_) with
+ | Some (o,(existT _ _stringappend_1156_ _)) => returnm (o, build_ex _stringappend_1156_)
+ | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(o, existT _ _stringappend_1156_ _) =>
+ let _stringappend_1157_ := string_drop _stringappend_1155_ (build_ex _stringappend_1156_) in
+ (match (bit_maybe_r_matches_prefix _stringappend_1157_) with
+ | Some (r,(existT _ _stringappend_1158_ _)) => returnm (r, build_ex _stringappend_1158_)
+ | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(r, existT _ _stringappend_1158_ _) =>
+ let _stringappend_1159_ := string_drop _stringappend_1157_ (build_ex _stringappend_1158_) in
+ (match (bit_maybe_w_matches_prefix _stringappend_1159_) with
+ | Some (w,(existT _ _stringappend_1160_ _)) => returnm (w, build_ex _stringappend_1160_)
+ | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(w, existT _ _stringappend_1160_ _) =>
+ (match (string_drop _stringappend_1159_ (build_ex _stringappend_1160_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
@@ -9294,24 +10199,23 @@ Definition fence_bits_backwards_matches (arg_ : string)
Definition fence_bits_matches_prefix (arg_ : string)
: M (option ((mword 4 * {n : Z & ArithFact (n >= 0)}))) :=
- let _stringappend_1679_ := arg_ in
- (if ((match (bit_maybe_r_matches_prefix _stringappend_1679_) with
- | Some (_stringappend_1680_,(existT _ _stringappend_1681_ _)) =>
- let _stringappend_1682_ :=
- string_drop _stringappend_1679_ (build_ex _stringappend_1681_) in
- if ((match (bit_maybe_w_matches_prefix _stringappend_1682_) with
- | Some (_stringappend_1683_,(existT _ _stringappend_1684_ _)) =>
- let _stringappend_1685_ :=
- string_drop _stringappend_1682_ (build_ex _stringappend_1684_) in
- if ((match (bit_maybe_i_matches_prefix _stringappend_1685_) with
- | Some (_stringappend_1686_,(existT _ _stringappend_1687_ _)) =>
- let _stringappend_1688_ :=
- string_drop _stringappend_1685_ (build_ex _stringappend_1687_) in
- if ((match (bit_maybe_o_matches_prefix _stringappend_1688_) with
- | Some (_stringappend_1689_,(existT _ _stringappend_1690_ _)) =>
- match (string_drop _stringappend_1688_
- (build_ex
- _stringappend_1690_)) with
+ let _stringappend_1145_ := arg_ in
+ (if ((match (bit_maybe_i_matches_prefix _stringappend_1145_) with
+ | Some (i,(existT _ _stringappend_1146_ _)) =>
+ let _stringappend_1147_ :=
+ string_drop _stringappend_1145_ (build_ex _stringappend_1146_) in
+ if ((match (bit_maybe_o_matches_prefix _stringappend_1147_) with
+ | Some (o,(existT _ _stringappend_1148_ _)) =>
+ let _stringappend_1149_ :=
+ string_drop _stringappend_1147_ (build_ex _stringappend_1148_) in
+ if ((match (bit_maybe_r_matches_prefix _stringappend_1149_) with
+ | Some (r,(existT _ _stringappend_1150_ _)) =>
+ let _stringappend_1151_ :=
+ string_drop _stringappend_1149_ (build_ex _stringappend_1150_) in
+ if ((match (bit_maybe_w_matches_prefix _stringappend_1151_) with
+ | Some (w,(existT _ _stringappend_1152_ _)) =>
+ match (string_drop _stringappend_1151_
+ (build_ex _stringappend_1152_)) with
| s_ => true
end
| None => false
@@ -9328,42 +10232,34 @@ Definition fence_bits_matches_prefix (arg_ : string)
else false
| None => false
end)) then
- match (bit_maybe_r_matches_prefix _stringappend_1679_) with
- | Some (_stringappend_1680_,(existT _ _stringappend_1681_ _)) =>
- returnm ((_stringappend_1680_, build_ex _stringappend_1681_)
- : (mword 1 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__1 : (mword 1 * {n : Z & ArithFact (n >= 0)}) =>
- let '(r, existT _ _stringappend_1681_ _) := w__1 : (mword 1 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1682_ := string_drop _stringappend_1679_ (build_ex _stringappend_1681_) in
- match (bit_maybe_w_matches_prefix _stringappend_1682_) with
- | Some (_stringappend_1683_,(existT _ _stringappend_1684_ _)) =>
- returnm ((_stringappend_1683_, build_ex _stringappend_1684_)
- : (mword 1 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__3 : (mword 1 * {n : Z & ArithFact (n >= 0)}) =>
- let '(w, existT _ _stringappend_1684_ _) := w__3 : (mword 1 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1685_ := string_drop _stringappend_1682_ (build_ex _stringappend_1684_) in
- match (bit_maybe_i_matches_prefix _stringappend_1685_) with
- | Some (_stringappend_1686_,(existT _ _stringappend_1687_ _)) =>
- returnm ((_stringappend_1686_, build_ex _stringappend_1687_)
- : (mword 1 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__5 : (mword 1 * {n : Z & ArithFact (n >= 0)}) =>
- let '(i, existT _ _stringappend_1687_ _) := w__5 : (mword 1 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1688_ := string_drop _stringappend_1685_ (build_ex _stringappend_1687_) in
- match (bit_maybe_o_matches_prefix _stringappend_1688_) with
- | Some (_stringappend_1689_,(existT _ _stringappend_1690_ _)) =>
- returnm ((_stringappend_1689_, build_ex _stringappend_1690_)
- : (mword 1 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__7 : (mword 1 * {n : Z & ArithFact (n >= 0)}) =>
- let '(o, existT _ _stringappend_1690_ _) := w__7 : (mword 1 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_1688_ (build_ex _stringappend_1690_)) with
+ (match (bit_maybe_i_matches_prefix _stringappend_1145_) with
+ | Some (i,(existT _ _stringappend_1146_ _)) => returnm (i, build_ex _stringappend_1146_)
+ | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(i, existT _ _stringappend_1146_ _) =>
+ let _stringappend_1147_ := string_drop _stringappend_1145_ (build_ex _stringappend_1146_) in
+ (match (bit_maybe_o_matches_prefix _stringappend_1147_) with
+ | Some (o,(existT _ _stringappend_1148_ _)) => returnm (o, build_ex _stringappend_1148_)
+ | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(o, existT _ _stringappend_1148_ _) =>
+ let _stringappend_1149_ := string_drop _stringappend_1147_ (build_ex _stringappend_1148_) in
+ (match (bit_maybe_r_matches_prefix _stringappend_1149_) with
+ | Some (r,(existT _ _stringappend_1150_ _)) => returnm (r, build_ex _stringappend_1150_)
+ | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(r, existT _ _stringappend_1150_ _) =>
+ let _stringappend_1151_ := string_drop _stringappend_1149_ (build_ex _stringappend_1150_) in
+ (match (bit_maybe_w_matches_prefix _stringappend_1151_) with
+ | Some (w,(existT _ _stringappend_1152_ _)) => returnm (w, build_ex _stringappend_1152_)
+ | _ => exit tt : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 1 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(w, existT _ _stringappend_1152_ _) =>
+ returnm ((match (string_drop _stringappend_1151_ (build_ex _stringappend_1152_)) with
| s_ =>
- Some (concat_vec (r : bits 1)
- (concat_vec (w : bits 1) (concat_vec (i : bits 1) (o : bits 1))),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ ((concat_vec (i : bits 1)
+ (concat_vec (o : bits 1) (concat_vec (r : bits 1) (w : bits 1))), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((mword 4 * {n : Z & ArithFact (n >= 0)})))
else returnm (None : option ((mword 4 * {n : Z & ArithFact (n >= 0)}))))
@@ -9382,7 +10278,7 @@ Definition lrsc_width_str (width : word_width)
: string :=
match width with | BYTE => ".b" | HALF => ".h" | WORD => ".w" | DOUBLE => ".d" end.
-Definition process_loadres {n : Z} (rd : mword 5) (addr : mword 64) (value : MemoryOpResult (mword (8 * n))) (is_unsigned : bool) `{ArithFact (1 <=
+Definition process_loadres {n : Z} (rd : mword 5) (addr : mword 64) (value : MemoryOpResult (mword (8 * n))) (is_unsigned : bool) `{ArithFact ((0 + 1) <=
n /\
n <= 8)}
: M (bool) :=
@@ -9390,7 +10286,7 @@ Definition process_loadres {n : Z} (rd : mword 5) (addr : mword 64) (value : Mem
| MemValue (result) =>
let '_ := (load_reservation addr) : unit in
wX
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {syn_n : Z & ArithFact (0 <= syn_n /\ (syn_n + 1) <= 32)})) result >>
returnm (true
: bool)
@@ -9414,53 +10310,68 @@ Definition encdec_amoop_forwards (arg_ : amoop)
Definition encdec_amoop_backwards (arg_ : mword 5)
: amoop :=
- let p0_ := arg_ in
+ let b__0 := arg_ in
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B1] : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
AMOSWAP
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
AMOADD
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
AMOXOR
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
AMOAND
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
AMOOR
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
AMOMIN
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
AMOMAX
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
AMOMINU
else AMOMAXU.
@@ -9481,59 +10392,76 @@ Definition encdec_amoop_forwards_matches (arg_ : amoop)
Definition encdec_amoop_backwards_matches (arg_ : mword 5)
: bool :=
- let p0_ := arg_ in
+ let b__0 := arg_ in
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B1] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B1] : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B1;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B1;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B0;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B0;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno p0_)
+ (projT1 ((build_ex (projT1 (regbits_to_regno b__0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B1;B1;B1;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B1;B1;B1;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
true
else false.
@@ -9599,68 +10527,96 @@ Definition amo_mnemonic_backwards_matches (arg_ : string)
Definition amo_mnemonic_matches_prefix (arg_ : string)
: option ((amoop * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1670_ := arg_ in
- if ((andb (string_startswith _stringappend_1670_ "amoswap")
- (match (string_drop _stringappend_1670_ (string_length "amoswap")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1670_ (string_length "amoswap")) with
- | s_ => Some (AMOSWAP, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1136_ := arg_ in
+ if ((andb (string_startswith _stringappend_1136_ "amoswap")
+ (match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amoswap")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amoswap")))) with
+ | s_ =>
+ Some
+ ((AMOSWAP, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1670_ "amoadd")
- (match (string_drop _stringappend_1670_ (string_length "amoadd")) with
+ else if ((andb (string_startswith _stringappend_1136_ "amoadd")
+ (match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amoadd")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1670_ (string_length "amoadd")) with
- | s_ => Some (AMOADD, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amoadd")))) with
+ | s_ =>
+ Some
+ ((AMOADD, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1670_ "amoxor")
- (match (string_drop _stringappend_1670_ (string_length "amoxor")) with
+ else if ((andb (string_startswith _stringappend_1136_ "amoxor")
+ (match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amoxor")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1670_ (string_length "amoxor")) with
- | s_ => Some (AMOXOR, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amoxor")))) with
+ | s_ =>
+ Some
+ ((AMOXOR, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1670_ "amoand")
- (match (string_drop _stringappend_1670_ (string_length "amoand")) with
+ else if ((andb (string_startswith _stringappend_1136_ "amoand")
+ (match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amoand")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1670_ (string_length "amoand")) with
- | s_ => Some (AMOAND, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amoand")))) with
+ | s_ =>
+ Some
+ ((AMOAND, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1670_ "amoor")
- (match (string_drop _stringappend_1670_ (string_length "amoor")) with
+ else if ((andb (string_startswith _stringappend_1136_ "amoor")
+ (match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amoor")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1670_ (string_length "amoor")) with
- | s_ => Some (AMOOR, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amoor")))) with
+ | s_ =>
+ Some
+ ((AMOOR, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1670_ "amomin")
- (match (string_drop _stringappend_1670_ (string_length "amomin")) with
+ else if ((andb (string_startswith _stringappend_1136_ "amomin")
+ (match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amomin")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1670_ (string_length "amomin")) with
- | s_ => Some (AMOMIN, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amomin")))) with
+ | s_ =>
+ Some
+ ((AMOMIN, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1670_ "amomax")
- (match (string_drop _stringappend_1670_ (string_length "amomax")) with
+ else if ((andb (string_startswith _stringappend_1136_ "amomax")
+ (match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amomax")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1670_ (string_length "amomax")) with
- | s_ => Some (AMOMAX, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amomax")))) with
+ | s_ =>
+ Some
+ ((AMOMAX, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1670_ "amominu")
- (match (string_drop _stringappend_1670_ (string_length "amominu")) with
+ else if ((andb (string_startswith _stringappend_1136_ "amominu")
+ (match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amominu")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1670_ (string_length "amominu")) with
- | s_ => Some (AMOMINU, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amominu")))) with
+ | s_ =>
+ Some
+ ((AMOMINU, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1670_ "amomaxu")
- (match (string_drop _stringappend_1670_ (string_length "amomaxu")) with
+ else if ((andb (string_startswith _stringappend_1136_ "amomaxu")
+ (match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amomaxu")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1670_ (string_length "amomaxu")) with
- | s_ => Some (AMOMAXU, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1136_ (build_ex (projT1 (string_length "amomaxu")))) with
+ | s_ =>
+ Some
+ ((AMOMAXU, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -9674,9 +10630,9 @@ Definition encdec_csrop_forwards (arg_ : csrop)
Definition encdec_csrop_backwards (arg_ : mword 2)
: csrop :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B1] : mword 2))) then CSRRW
- else if ((eq_vec p0_ (vec_of_bits [B1;B0] : mword 2))) then CSRRS
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then CSRRW
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0] : mword 2))) then CSRRS
else CSRRC.
Definition encdec_csrop_forwards_matches (arg_ : csrop)
@@ -9685,10 +10641,10 @@ Definition encdec_csrop_forwards_matches (arg_ : csrop)
Definition encdec_csrop_backwards_matches (arg_ : mword 2)
: bool :=
- let p0_ := arg_ in
- if ((eq_vec p0_ (vec_of_bits [B0;B1] : mword 2))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B0] : mword 2))) then true
- else if ((eq_vec p0_ (vec_of_bits [B1;B1] : mword 2))) then true
+ let b__0 := arg_ in
+ if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B0] : mword 2))) then true
+ else if ((eq_vec b__0 (vec_of_bits [B1;B1] : mword 2))) then true
else false.
Definition readCSR (csr : mword 12)
@@ -9707,40 +10663,40 @@ Definition readCSR (csr : mword 12)
(read_reg mhartid_ref : M (mword 64))
: M (xlenbits)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- read_reg mstatus_ref >>= fun w__4 : Mstatus => returnm ((_get_Mstatus_bits w__4) : xlenbits)
+ read_reg mstatus_ref >>= fun w__4 : Mstatus => returnm ((_get_Mstatus_bits w__4) : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- read_reg misa_ref >>= fun w__5 : Misa => returnm ((_get_Misa_bits w__5) : xlenbits)
+ read_reg misa_ref >>= fun w__5 : Misa => returnm ((_get_Misa_bits w__5) : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- read_reg medeleg_ref >>= fun w__6 : Medeleg => returnm ((_get_Medeleg_bits w__6) : xlenbits)
+ read_reg medeleg_ref >>= fun w__6 : Medeleg => returnm ((_get_Medeleg_bits w__6) : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
read_reg mideleg_ref >>= fun w__7 : Minterrupts =>
returnm ((_get_Minterrupts_bits w__7)
- : xlenbits)
+ : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
read_reg mie_ref >>= fun w__8 : Minterrupts =>
returnm ((_get_Minterrupts_bits w__8)
- : xlenbits)
+ : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- read_reg mtvec_ref >>= fun w__9 : Mtvec => returnm ((_get_Mtvec_bits w__9) : xlenbits)
+ read_reg mtvec_ref >>= fun w__9 : Mtvec => returnm ((_get_Mtvec_bits w__9) : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
read_reg mcounteren_ref >>= fun w__10 : Counteren =>
returnm ((EXTZ 64 (_get_Counteren_bits w__10))
- : xlenbits)
+ : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
(read_reg mscratch_ref : M (mword 64))
: M (xlenbits)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- (read_reg mepc_ref : M (mword 64)) >>= fun w__12 : mword 64 =>
- pc_alignment_mask tt >>= fun w__13 : mword 64 => returnm ((and_vec w__12 w__13) : xlenbits)
+ (read_reg mepc_ref : M (mword 64)) >>= fun w__12 : xlenbits =>
+ pc_alignment_mask tt >>= fun w__13 : mword 64 => returnm ((and_vec w__12 w__13) : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
- read_reg mcause_ref >>= fun w__14 : Mcause => returnm ((_get_Mcause_bits w__14) : xlenbits)
+ read_reg mcause_ref >>= fun w__14 : Mcause => returnm ((_get_Mcause_bits w__14) : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
(read_reg mtval_ref : M (mword 64))
: M (xlenbits)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
read_reg mip_ref >>= fun w__16 : Minterrupts =>
returnm ((_get_Minterrupts_bits w__16)
- : xlenbits)
+ : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
(read_reg pmpcfg0_ref : M (mword 64))
: M (xlenbits)
@@ -9750,34 +10706,34 @@ Definition readCSR (csr : mword 12)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
read_reg mstatus_ref >>= fun w__19 : Mstatus =>
returnm ((_get_Sstatus_bits (lower_mstatus w__19))
- : xlenbits)
+ : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
read_reg sedeleg_ref >>= fun w__20 : Sedeleg =>
returnm ((_get_Sedeleg_bits w__20)
- : xlenbits)
+ : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
read_reg sideleg_ref >>= fun w__21 : Sinterrupts =>
returnm ((_get_Sinterrupts_bits w__21)
- : xlenbits)
+ : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
read_reg mie_ref >>= fun w__22 : Minterrupts =>
read_reg mideleg_ref >>= fun w__23 : Minterrupts =>
returnm ((_get_Sinterrupts_bits (lower_mie w__22 w__23))
- : xlenbits)
+ : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- read_reg stvec_ref >>= fun w__24 : Mtvec => returnm ((_get_Mtvec_bits w__24) : xlenbits)
+ read_reg stvec_ref >>= fun w__24 : Mtvec => returnm ((_get_Mtvec_bits w__24) : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
read_reg scounteren_ref >>= fun w__25 : Counteren =>
returnm ((EXTZ 64 (_get_Counteren_bits w__25))
- : xlenbits)
+ : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
(read_reg sscratch_ref : M (mword 64))
: M (xlenbits)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- (read_reg sepc_ref : M (mword 64)) >>= fun w__27 : mword 64 =>
- pc_alignment_mask tt >>= fun w__28 : mword 64 => returnm ((and_vec w__27 w__28) : xlenbits)
+ (read_reg sepc_ref : M (mword 64)) >>= fun w__27 : xlenbits =>
+ pc_alignment_mask tt >>= fun w__28 : mword 64 => returnm ((and_vec w__27 w__28) : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
- read_reg scause_ref >>= fun w__29 : Mcause => returnm ((_get_Mcause_bits w__29) : xlenbits)
+ read_reg scause_ref >>= fun w__29 : Mcause => returnm ((_get_Mcause_bits w__29) : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
(read_reg stval_ref : M (mword 64))
: M (xlenbits)
@@ -9785,7 +10741,7 @@ Definition readCSR (csr : mword 12)
read_reg mip_ref >>= fun w__31 : Minterrupts =>
read_reg mideleg_ref >>= fun w__32 : Minterrupts =>
returnm ((_get_Sinterrupts_bits (lower_mip w__31 w__32))
- : xlenbits)
+ : mword 64)
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
(read_reg satp_ref : M (mword 64))
: M (xlenbits)
@@ -9799,9 +10755,9 @@ Definition readCSR (csr : mword 12)
(read_reg minstret_ref : M (mword 64))
: M (xlenbits)
else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
- (read_reg tselect_ref : M (mword 64)) >>= fun w__37 : mword 64 =>
+ (read_reg tselect_ref : M (mword 64)) >>= fun w__37 : xlenbits =>
returnm ((not_vec w__37)
- : xlenbits)
+ : mword 64)
else
let '_ := (print_bits "unhandled read to CSR " csr) : unit in
returnm ((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
@@ -9825,172 +10781,201 @@ Definition writeCSR (csr : mword 12) (value : mword 64)
read_reg mstatus_ref >>= fun w__0 : Mstatus =>
write_reg mstatus_ref (legalize_mstatus w__0 value) >>
read_reg mstatus_ref >>= fun w__1 : Mstatus =>
- returnm ((Some (_get_Mstatus_bits w__1))
- : option xlenbits)
+ returnm ((Some
+ (_get_Mstatus_bits w__1))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
read_reg misa_ref >>= fun w__2 : Misa =>
legalize_misa w__2 value >>= fun w__3 : Misa =>
write_reg misa_ref w__3 >>
read_reg misa_ref >>= fun w__4 : Misa =>
- returnm ((Some (_get_Misa_bits w__4))
- : option xlenbits)
+ returnm ((Some
+ (_get_Misa_bits w__4))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
read_reg medeleg_ref >>= fun w__5 : Medeleg =>
write_reg medeleg_ref (legalize_medeleg w__5 value) >>
read_reg medeleg_ref >>= fun w__6 : Medeleg =>
- returnm ((Some (_get_Medeleg_bits w__6))
- : option xlenbits)
+ returnm ((Some
+ (_get_Medeleg_bits w__6))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
read_reg mideleg_ref >>= fun w__7 : Minterrupts =>
write_reg mideleg_ref (legalize_mideleg w__7 value) >>
read_reg mideleg_ref >>= fun w__8 : Minterrupts =>
- returnm ((Some (_get_Minterrupts_bits w__8))
- : option xlenbits)
+ returnm ((Some
+ (_get_Minterrupts_bits w__8))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
read_reg mie_ref >>= fun w__9 : Minterrupts =>
write_reg mie_ref (legalize_mie w__9 value) >>
read_reg mie_ref >>= fun w__10 : Minterrupts =>
- returnm ((Some (_get_Minterrupts_bits w__10))
- : option xlenbits)
+ returnm ((Some
+ (_get_Minterrupts_bits w__10))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
read_reg mtvec_ref >>= fun w__11 : Mtvec =>
write_reg mtvec_ref (legalize_tvec w__11 value) >>
read_reg mtvec_ref >>= fun w__12 : Mtvec =>
- returnm ((Some (_get_Mtvec_bits w__12))
- : option xlenbits)
+ returnm ((Some
+ (_get_Mtvec_bits w__12))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
read_reg mcounteren_ref >>= fun w__13 : Counteren =>
legalize_mcounteren w__13 value >>= fun w__14 : Counteren =>
write_reg mcounteren_ref w__14 >>
read_reg mcounteren_ref >>= fun w__15 : Counteren =>
- returnm ((Some (EXTZ 64 (_get_Counteren_bits w__15)))
- : option xlenbits)
+ returnm ((Some
+ (EXTZ 64 (_get_Counteren_bits w__15)))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
write_reg mscratch_ref value >>
- (read_reg mscratch_ref : M (mword 64)) >>= fun w__16 : mword 64 =>
- returnm ((Some w__16)
- : option xlenbits)
+ (read_reg mscratch_ref : M (mword 64)) >>= fun w__16 : xlenbits =>
+ returnm ((Some
+ (w__16))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- legalize_xepc value >>= fun w__17 : xlenbits =>
+ legalize_xepc value >>= fun w__17 : mword 64 =>
write_reg mepc_ref w__17 >>
- (read_reg mepc_ref : M (mword 64)) >>= fun w__18 : mword 64 =>
- returnm ((Some w__18)
- : option xlenbits)
+ (read_reg mepc_ref : M (mword 64)) >>= fun w__18 : xlenbits =>
+ returnm ((Some
+ (w__18))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
_set_Mcause_bits mcause_ref value >>
read_reg mcause_ref >>= fun w__19 : Mcause =>
- returnm ((Some (_get_Mcause_bits w__19))
- : option xlenbits)
+ returnm ((Some
+ (_get_Mcause_bits w__19))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
write_reg mtval_ref value >>
- (read_reg mtval_ref : M (mword 64)) >>= fun w__20 : mword 64 =>
- returnm ((Some w__20)
- : option xlenbits)
+ (read_reg mtval_ref : M (mword 64)) >>= fun w__20 : xlenbits =>
+ returnm ((Some
+ (w__20))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
read_reg mip_ref >>= fun w__21 : Minterrupts =>
write_reg mip_ref (legalize_mip w__21 value) >>
read_reg mip_ref >>= fun w__22 : Minterrupts =>
- returnm ((Some (_get_Minterrupts_bits w__22))
- : option xlenbits)
+ returnm ((Some
+ (_get_Minterrupts_bits w__22))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
write_reg pmpcfg0_ref value >>
- (read_reg pmpcfg0_ref : M (mword 64)) >>= fun w__23 : mword 64 =>
- returnm ((Some w__23)
- : option xlenbits)
+ (read_reg pmpcfg0_ref : M (mword 64)) >>= fun w__23 : xlenbits =>
+ returnm ((Some
+ (w__23))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12))) then
write_reg pmpaddr0_ref value >>
- (read_reg pmpaddr0_ref : M (mword 64)) >>= fun w__24 : mword 64 =>
- returnm ((Some w__24)
- : option xlenbits)
+ (read_reg pmpaddr0_ref : M (mword 64)) >>= fun w__24 : xlenbits =>
+ returnm ((Some
+ (w__24))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
read_reg mstatus_ref >>= fun w__25 : Mstatus =>
write_reg mstatus_ref (legalize_sstatus w__25 value) >>
read_reg mstatus_ref >>= fun w__26 : Mstatus =>
- returnm ((Some (_get_Mstatus_bits w__26))
- : option xlenbits)
+ returnm ((Some
+ (_get_Mstatus_bits w__26))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
read_reg sedeleg_ref >>= fun w__27 : Sedeleg =>
write_reg sedeleg_ref (legalize_sedeleg w__27 value) >>
read_reg sedeleg_ref >>= fun w__28 : Sedeleg =>
- returnm ((Some (_get_Sedeleg_bits w__28))
- : option xlenbits)
+ returnm ((Some
+ (_get_Sedeleg_bits w__28))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
_set_Sinterrupts_bits sideleg_ref value >>
read_reg sideleg_ref >>= fun w__29 : Sinterrupts =>
- returnm ((Some (_get_Sinterrupts_bits w__29))
- : option xlenbits)
+ returnm ((Some
+ (_get_Sinterrupts_bits w__29))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
read_reg mie_ref >>= fun w__30 : Minterrupts =>
read_reg mideleg_ref >>= fun w__31 : Minterrupts =>
write_reg mie_ref (legalize_sie w__30 w__31 value) >>
read_reg mie_ref >>= fun w__32 : Minterrupts =>
- returnm ((Some (_get_Minterrupts_bits w__32))
- : option xlenbits)
+ returnm ((Some
+ (_get_Minterrupts_bits w__32))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
read_reg stvec_ref >>= fun w__33 : Mtvec =>
write_reg stvec_ref (legalize_tvec w__33 value) >>
read_reg stvec_ref >>= fun w__34 : Mtvec =>
- returnm ((Some (_get_Mtvec_bits w__34))
- : option xlenbits)
+ returnm ((Some
+ (_get_Mtvec_bits w__34))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
read_reg scounteren_ref >>= fun w__35 : Counteren =>
legalize_scounteren w__35 value >>= fun w__36 : Counteren =>
write_reg scounteren_ref w__36 >>
read_reg scounteren_ref >>= fun w__37 : Counteren =>
- returnm ((Some (EXTZ 64 (_get_Counteren_bits w__37)))
- : option xlenbits)
+ returnm ((Some
+ (EXTZ 64 (_get_Counteren_bits w__37)))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
write_reg sscratch_ref value >>
- (read_reg sscratch_ref : M (mword 64)) >>= fun w__38 : mword 64 =>
- returnm ((Some w__38)
- : option xlenbits)
+ (read_reg sscratch_ref : M (mword 64)) >>= fun w__38 : xlenbits =>
+ returnm ((Some
+ (w__38))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- legalize_xepc value >>= fun w__39 : xlenbits =>
+ legalize_xepc value >>= fun w__39 : mword 64 =>
write_reg sepc_ref w__39 >>
- (read_reg sepc_ref : M (mword 64)) >>= fun w__40 : mword 64 =>
- returnm ((Some w__40)
- : option xlenbits)
+ (read_reg sepc_ref : M (mword 64)) >>= fun w__40 : xlenbits =>
+ returnm ((Some
+ (w__40))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
_set_Mcause_bits scause_ref value >>
read_reg scause_ref >>= fun w__41 : Mcause =>
- returnm ((Some (_get_Mcause_bits w__41))
- : option xlenbits)
+ returnm ((Some
+ (_get_Mcause_bits w__41))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
write_reg stval_ref value >>
- (read_reg stval_ref : M (mword 64)) >>= fun w__42 : mword 64 =>
- returnm ((Some w__42)
- : option xlenbits)
+ (read_reg stval_ref : M (mword 64)) >>= fun w__42 : xlenbits =>
+ returnm ((Some
+ (w__42))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
read_reg mip_ref >>= fun w__43 : Minterrupts =>
read_reg mideleg_ref >>= fun w__44 : Minterrupts =>
write_reg mip_ref (legalize_sip w__43 w__44 value) >>
read_reg mip_ref >>= fun w__45 : Minterrupts =>
- returnm ((Some (_get_Minterrupts_bits w__45))
- : option xlenbits)
+ returnm ((Some
+ (_get_Minterrupts_bits w__45))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
cur_Architecture tt >>= fun w__46 : Architecture =>
- (read_reg satp_ref : M (mword 64)) >>= fun w__47 : mword 64 =>
+ (read_reg satp_ref : M (mword 64)) >>= fun w__47 : xlenbits =>
write_reg satp_ref (legalize_satp w__46 w__47 value) >>
- (read_reg satp_ref : M (mword 64)) >>= fun w__48 : mword 64 =>
- returnm ((Some w__48)
- : option xlenbits)
+ (read_reg satp_ref : M (mword 64)) >>= fun w__48 : xlenbits =>
+ returnm ((Some
+ (w__48))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
write_reg tselect_ref value >>
- (read_reg tselect_ref : M (mword 64)) >>= fun w__49 : mword 64 =>
- returnm ((Some w__49)
- : option xlenbits)
+ (read_reg tselect_ref : M (mword 64)) >>= fun w__49 : xlenbits =>
+ returnm ((Some
+ (w__49))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
write_reg mcycle_ref value >>
- (read_reg mcycle_ref : M (mword 64)) >>= fun w__50 : mword 64 =>
- returnm ((Some w__50)
- : option xlenbits)
+ (read_reg mcycle_ref : M (mword 64)) >>= fun w__50 : xlenbits =>
+ returnm ((Some
+ (w__50))
+ : option (mword 64))
else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
write_reg minstret_ref value >>
write_reg minstret_written_ref true >>
- (read_reg minstret_ref : M (mword 64)) >>= fun w__51 : mword 64 =>
- returnm ((Some w__51)
- : option xlenbits)
- else returnm (None : option xlenbits)) >>= fun res : option xlenbits =>
+ (read_reg minstret_ref : M (mword 64)) >>= fun w__51 : xlenbits =>
+ returnm ((Some
+ (w__51))
+ : option (mword 64))
+ else returnm (None : option (mword 64))) >>= fun res : option xlenbits =>
returnm ((match res with
| Some (v) =>
print_endline
@@ -10026,17 +11011,26 @@ Definition maybe_i_backwards_matches (arg_ : string)
Definition maybe_i_matches_prefix (arg_ : string)
: option ((bool * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1668_ := arg_ in
- if ((andb (string_startswith _stringappend_1668_ "i")
- (match (string_drop _stringappend_1668_ (string_length "i")) with | s_ => true end))) then
- match (string_drop _stringappend_1668_ (string_length "i")) with
- | s_ => Some (true, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1134_ := arg_ in
+ if ((andb (string_startswith _stringappend_1134_ "i")
+ (match (string_drop _stringappend_1134_ (build_ex (projT1 (string_length "i")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1134_ (build_ex (projT1 (string_length "i")))) with
+ | s_ =>
+ Some
+ ((true, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1668_ "")
- (match (string_drop _stringappend_1668_ (string_length "")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1668_ (string_length "")) with
- | s_ => Some (false, sub_nat (string_length arg_) (string_length s_))
+ else if ((andb (string_startswith _stringappend_1134_ "")
+ (match (string_drop _stringappend_1134_ (build_ex (projT1 (string_length "")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1134_ (build_ex (projT1 (string_length "")))) with
+ | s_ =>
+ Some
+ ((false, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -10064,26 +11058,36 @@ Definition csr_mnemonic_backwards_matches (arg_ : string)
Definition csr_mnemonic_matches_prefix (arg_ : string)
: option ((csrop * {n : Z & ArithFact (n >= 0)})) :=
- let _stringappend_1665_ := arg_ in
- if ((andb (string_startswith _stringappend_1665_ "csrrw")
- (match (string_drop _stringappend_1665_ (string_length "csrrw")) with | s_ => true end)))
- then
- match (string_drop _stringappend_1665_ (string_length "csrrw")) with
- | s_ => Some (CSRRW, sub_nat (string_length arg_) (string_length s_))
+ let _stringappend_1131_ := arg_ in
+ if ((andb (string_startswith _stringappend_1131_ "csrrw")
+ (match (string_drop _stringappend_1131_ (build_ex (projT1 (string_length "csrrw")))) with
+ | s_ => true
+ end))) then
+ match (string_drop _stringappend_1131_ (build_ex (projT1 (string_length "csrrw")))) with
+ | s_ =>
+ Some
+ ((CSRRW, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1665_ "csrrs")
- (match (string_drop _stringappend_1665_ (string_length "csrrs")) with
+ else if ((andb (string_startswith _stringappend_1131_ "csrrs")
+ (match (string_drop _stringappend_1131_ (build_ex (projT1 (string_length "csrrs")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1665_ (string_length "csrrs")) with
- | s_ => Some (CSRRS, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1131_ (build_ex (projT1 (string_length "csrrs")))) with
+ | s_ =>
+ Some
+ ((CSRRS, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
- else if ((andb (string_startswith _stringappend_1665_ "csrrc")
- (match (string_drop _stringappend_1665_ (string_length "csrrc")) with
+ else if ((andb (string_startswith _stringappend_1131_ "csrrc")
+ (match (string_drop _stringappend_1131_ (build_ex (projT1 (string_length "csrrc")))) with
| s_ => true
end))) then
- match (string_drop _stringappend_1665_ (string_length "csrrc")) with
- | s_ => Some (CSRRC, sub_nat (string_length arg_) (string_length s_))
+ match (string_drop _stringappend_1131_ (build_ex (projT1 (string_length "csrrc")))) with
+ | s_ =>
+ Some
+ ((CSRRC, build_ex (projT1 (sub_nat (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end
else None.
@@ -10093,9 +11097,12 @@ Definition decodeCompressed (v__2 : mword 16)
(eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B0;B0;B0] : mword 3))
(andb
(Z.eqb
- (projT1 ((regbits_to_regno (subrange_vec_dec v__2 11 7))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (subrange_vec_dec v__2 11 7))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2)))))
then
@@ -10103,343 +11110,384 @@ Definition decodeCompressed (v__2 : mword 16)
let nzi0 : bits 5 := subrange_vec_dec v__2 6 2 in
if sumbool_of_bool ((andb (eq_vec nzi1 (vec_of_bits [B0] : mword 1))
(Z.eqb
- (projT1 ((regbits_to_regno nzi0)
+ (projT1 ((build_ex (projT1 (regbits_to_regno nzi0)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))) then
- Some (NOP tt)
+ Some
+ (NOP
+ (tt))
else None
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B0;B0;B0] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B0] : mword 2)))) then
- let nz54 : bits 2 := subrange_vec_dec v__2 12 11 in
+ let rd : cregbits := subrange_vec_dec v__2 4 2 in
let nz96 : bits 4 := subrange_vec_dec v__2 10 7 in
- let nz2 : bits 1 := subrange_vec_dec v__2 6 6 in
+ let nz54 : bits 2 := subrange_vec_dec v__2 12 11 in
let nz3 : bits 1 := subrange_vec_dec v__2 5 5 in
- let rd : cregbits := subrange_vec_dec v__2 4 2 in
+ let nz2 : bits 1 := subrange_vec_dec v__2 6 6 in
let nzimm := (concat_vec nz96 (concat_vec nz54 (concat_vec nz3 nz2))) : bits 8 in
if ((eq_vec nzimm (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword 8))) then None
- else Some (C_ADDI4SPN (rd,nzimm))
+ else Some (C_ADDI4SPN (rd, nzimm))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B0;B1;B0] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B0] : mword 2)))) then
+ let ui6 : bits 1 := subrange_vec_dec v__2 5 5 in
let ui53 : bits 3 := subrange_vec_dec v__2 12 10 in
- let rs1 : cregbits := subrange_vec_dec v__2 9 7 in
let ui2 : bits 1 := subrange_vec_dec v__2 6 6 in
- let ui6 : bits 1 := subrange_vec_dec v__2 5 5 in
+ let rs1 : cregbits := subrange_vec_dec v__2 9 7 in
let rd : cregbits := subrange_vec_dec v__2 4 2 in
let uimm := (concat_vec ui6 (concat_vec ui53 ui2)) : bits 5 in
- Some (C_LW (uimm,rs1,rd))
+ Some
+ (C_LW
+ (uimm, rs1, rd))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B0;B1;B1] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B0] : mword 2)))) then
+ let ui76 : bits 2 := subrange_vec_dec v__2 6 5 in
let ui53 : bits 3 := subrange_vec_dec v__2 12 10 in
let rs1 : cregbits := subrange_vec_dec v__2 9 7 in
- let ui76 : bits 2 := subrange_vec_dec v__2 6 5 in
let rd : cregbits := subrange_vec_dec v__2 4 2 in
let uimm := (concat_vec ui76 ui53) : bits 5 in
- Some (C_LD (uimm,rs1,rd))
+ Some
+ (C_LD
+ (uimm, rs1, rd))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B1;B1;B0] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B0] : mword 2)))) then
+ let ui6 : bits 1 := subrange_vec_dec v__2 5 5 in
let ui53 : bits 3 := subrange_vec_dec v__2 12 10 in
- let rs1 : cregbits := subrange_vec_dec v__2 9 7 in
let ui2 : bits 1 := subrange_vec_dec v__2 6 6 in
- let ui6 : bits 1 := subrange_vec_dec v__2 5 5 in
let rs2 : cregbits := subrange_vec_dec v__2 4 2 in
+ let rs1 : cregbits := subrange_vec_dec v__2 9 7 in
let uimm := (concat_vec ui6 (concat_vec ui53 ui2)) : bits 5 in
- Some (C_SW (uimm,rs1,rs2))
+ Some
+ (C_SW
+ (uimm, rs1, rs2))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B1;B1;B1] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B0] : mword 2)))) then
- let ui53 : bits 3 := subrange_vec_dec v__2 12 10 in
- let rs1 : bits 3 := subrange_vec_dec v__2 9 7 in
let ui76 : bits 2 := subrange_vec_dec v__2 6 5 in
+ let ui53 : bits 3 := subrange_vec_dec v__2 12 10 in
let rs2 : bits 3 := subrange_vec_dec v__2 4 2 in
+ let rs1 : bits 3 := subrange_vec_dec v__2 9 7 in
let uimm := (concat_vec ui76 ui53) : bits 5 in
- Some (C_SD (uimm,rs1,rs2))
+ Some
+ (C_SD
+ (uimm, rs1, rs2))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B0;B0;B0] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2)))) then
- let nzi5 : bits 1 := subrange_vec_dec v__2 12 12 in
let rsd : regbits := subrange_vec_dec v__2 11 7 in
+ let nzi5 : bits 1 := subrange_vec_dec v__2 12 12 in
let nzi40 : bits 5 := subrange_vec_dec v__2 6 2 in
let nzi := (concat_vec nzi5 nzi40) : bits 6 in
if sumbool_of_bool ((orb (eq_vec nzi (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))
(Z.eqb
- (projT1 ((regbits_to_regno rsd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rsd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno zreg)
+ (projT1 ((build_ex (projT1 (regbits_to_regno zreg)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))) then
None
- else Some (C_ADDI (nzi,rsd))
+ else Some (C_ADDI (nzi, rsd))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B0;B0;B1] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2)))) then
- let imm5 : bits 1 := subrange_vec_dec v__2 12 12 in
let rsd : regbits := subrange_vec_dec v__2 11 7 in
+ let imm5 : bits 1 := subrange_vec_dec v__2 12 12 in
let imm40 : bits 5 := subrange_vec_dec v__2 6 2 in
- Some (C_ADDIW (concat_vec imm5 imm40,rsd))
+ Some
+ (C_ADDIW
+ (concat_vec imm5 imm40, rsd))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B0;B1;B0] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2)))) then
- let imm5 : bits 1 := subrange_vec_dec v__2 12 12 in
let rd : regbits := subrange_vec_dec v__2 11 7 in
+ let imm5 : bits 1 := subrange_vec_dec v__2 12 12 in
let imm40 : bits 5 := subrange_vec_dec v__2 6 2 in
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno zreg)
+ (projT1 ((build_ex (projT1 (regbits_to_regno zreg)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
None
- else Some (C_LI (concat_vec imm5 imm40,rd))
+ else Some (C_LI (concat_vec imm5 imm40, rd))
else if sumbool_of_bool ((andb
(eq_vec (subrange_vec_dec v__2 15 13)
(vec_of_bits [B0;B1;B1] : mword 3))
(andb
(Z.eqb
- (projT1 ((regbits_to_regno (subrange_vec_dec v__2 11 7))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (subrange_vec_dec v__2 11 7))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno
- (vec_of_bits [B0;B0;B0;B1;B0] : mword 5))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B1;B0]
+ : mword 5))))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))
(eq_vec (subrange_vec_dec v__2 1 0)
(vec_of_bits [B0;B1] : mword 2))))) then
let nzi9 : bits 1 := subrange_vec_dec v__2 12 12 in
- let nzi4 : bits 1 := subrange_vec_dec v__2 6 6 in
- let nzi6 : bits 1 := subrange_vec_dec v__2 5 5 in
let nzi87 : bits 2 := subrange_vec_dec v__2 4 3 in
+ let nzi6 : bits 1 := subrange_vec_dec v__2 5 5 in
let nzi5 : bits 1 := subrange_vec_dec v__2 2 2 in
+ let nzi4 : bits 1 := subrange_vec_dec v__2 6 6 in
let nzimm := concat_vec nzi9 (concat_vec nzi87 (concat_vec nzi6 (concat_vec nzi5 nzi4))) in
if ((eq_vec nzimm (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))) then None
- else Some (C_ADDI16SP nzimm)
+ else Some (C_ADDI16SP (nzimm))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B0;B1;B1] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2)))) then
- let imm17 : bits 1 := subrange_vec_dec v__2 12 12 in
let rd : regbits := subrange_vec_dec v__2 11 7 in
+ let imm17 : bits 1 := subrange_vec_dec v__2 12 12 in
let imm1612 : bits 5 := subrange_vec_dec v__2 6 2 in
if sumbool_of_bool ((orb
(Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno zreg)
+ (projT1 ((build_ex (projT1 (regbits_to_regno zreg)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))
(Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno sp)
+ (projT1 ((build_ex (projT1 (regbits_to_regno sp)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))) then
None
- else Some (C_LUI (concat_vec imm17 imm1612,rd))
+ else Some (C_LUI (concat_vec imm17 imm1612, rd))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B1;B0;B0] : mword 3))
(andb (eq_vec (subrange_vec_dec v__2 11 10) (vec_of_bits [B0;B0] : mword 2))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2))))) then
- let nzui5 : bits 1 := subrange_vec_dec v__2 12 12 in
let rsd : cregbits := subrange_vec_dec v__2 9 7 in
+ let nzui5 : bits 1 := subrange_vec_dec v__2 12 12 in
let nzui40 : bits 5 := subrange_vec_dec v__2 6 2 in
let shamt : bits 6 := concat_vec nzui5 nzui40 in
if ((eq_vec shamt (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))) then None
- else Some (C_SRLI (shamt,rsd))
+ else Some (C_SRLI (shamt, rsd))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B1;B0;B0] : mword 3))
(andb (eq_vec (subrange_vec_dec v__2 11 10) (vec_of_bits [B0;B1] : mword 2))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2))))) then
- let nzui5 : bits 1 := subrange_vec_dec v__2 12 12 in
let rsd : cregbits := subrange_vec_dec v__2 9 7 in
+ let nzui5 : bits 1 := subrange_vec_dec v__2 12 12 in
let nzui40 : bits 5 := subrange_vec_dec v__2 6 2 in
let shamt : bits 6 := concat_vec nzui5 nzui40 in
if ((eq_vec shamt (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))) then None
- else Some (C_SRAI (shamt,rsd))
+ else Some (C_SRAI (shamt, rsd))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B1;B0;B0] : mword 3))
(andb (eq_vec (subrange_vec_dec v__2 11 10) (vec_of_bits [B1;B0] : mword 2))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2))))) then
- let i5 : bits 1 := subrange_vec_dec v__2 12 12 in
let rsd : cregbits := subrange_vec_dec v__2 9 7 in
+ let i5 : bits 1 := subrange_vec_dec v__2 12 12 in
let i40 : bits 5 := subrange_vec_dec v__2 6 2 in
- Some (C_ANDI (concat_vec i5 i40,rsd))
+ Some
+ (C_ANDI
+ (concat_vec i5 i40, rsd))
else if ((andb
(eq_vec (subrange_vec_dec v__2 15 10) (vec_of_bits [B1;B0;B0;B0;B1;B1] : mword 6))
(andb (eq_vec (subrange_vec_dec v__2 6 5) (vec_of_bits [B0;B0] : mword 2))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2))))) then
let rsd : cregbits := subrange_vec_dec v__2 9 7 in
let rs2 : cregbits := subrange_vec_dec v__2 4 2 in
- Some (C_SUB (rsd,rs2))
+ Some
+ (C_SUB
+ (rsd, rs2))
else if ((andb
(eq_vec (subrange_vec_dec v__2 15 10) (vec_of_bits [B1;B0;B0;B0;B1;B1] : mword 6))
(andb (eq_vec (subrange_vec_dec v__2 6 5) (vec_of_bits [B0;B1] : mword 2))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2))))) then
let rsd : cregbits := subrange_vec_dec v__2 9 7 in
let rs2 : cregbits := subrange_vec_dec v__2 4 2 in
- Some (C_XOR (rsd,rs2))
+ Some
+ (C_XOR
+ (rsd, rs2))
else if ((andb
(eq_vec (subrange_vec_dec v__2 15 10) (vec_of_bits [B1;B0;B0;B0;B1;B1] : mword 6))
(andb (eq_vec (subrange_vec_dec v__2 6 5) (vec_of_bits [B1;B0] : mword 2))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2))))) then
let rsd : cregbits := subrange_vec_dec v__2 9 7 in
let rs2 : cregbits := subrange_vec_dec v__2 4 2 in
- Some (C_OR (rsd,rs2))
+ Some
+ (C_OR
+ (rsd, rs2))
else if ((andb
(eq_vec (subrange_vec_dec v__2 15 10) (vec_of_bits [B1;B0;B0;B0;B1;B1] : mword 6))
(andb (eq_vec (subrange_vec_dec v__2 6 5) (vec_of_bits [B1;B1] : mword 2))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2))))) then
let rsd : cregbits := subrange_vec_dec v__2 9 7 in
let rs2 : cregbits := subrange_vec_dec v__2 4 2 in
- Some (C_AND (rsd,rs2))
+ Some
+ (C_AND
+ (rsd, rs2))
else if ((andb
(eq_vec (subrange_vec_dec v__2 15 10) (vec_of_bits [B1;B0;B0;B1;B1;B1] : mword 6))
(andb (eq_vec (subrange_vec_dec v__2 6 5) (vec_of_bits [B0;B0] : mword 2))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2))))) then
let rsd : cregbits := subrange_vec_dec v__2 9 7 in
let rs2 : cregbits := subrange_vec_dec v__2 4 2 in
- Some (C_SUBW (rsd,rs2))
+ Some
+ (C_SUBW
+ (rsd, rs2))
else if ((andb
(eq_vec (subrange_vec_dec v__2 15 10) (vec_of_bits [B1;B0;B0;B1;B1;B1] : mword 6))
(andb (eq_vec (subrange_vec_dec v__2 6 5) (vec_of_bits [B0;B1] : mword 2))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2))))) then
let rsd : cregbits := subrange_vec_dec v__2 9 7 in
let rs2 : cregbits := subrange_vec_dec v__2 4 2 in
- Some (C_ADDW (rsd,rs2))
+ Some
+ (C_ADDW
+ (rsd, rs2))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B1;B0;B1] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2)))) then
- let i11 : bits 1 := subrange_vec_dec v__2 12 12 in
- let i4 : bits 1 := subrange_vec_dec v__2 11 11 in
let i98 : bits 2 := subrange_vec_dec v__2 10 9 in
- let i10 : bits 1 := subrange_vec_dec v__2 8 8 in
- let i6 : bits 1 := subrange_vec_dec v__2 7 7 in
let i7 : bits 1 := subrange_vec_dec v__2 6 6 in
- let i31 : bits 3 := subrange_vec_dec v__2 5 3 in
+ let i6 : bits 1 := subrange_vec_dec v__2 7 7 in
let i5 : bits 1 := subrange_vec_dec v__2 2 2 in
- Some (C_J (concat_vec i11
- (concat_vec i10
- (concat_vec i98
- (concat_vec i7 (concat_vec i6 (concat_vec i5 (concat_vec i4 i31))))))))
+ let i4 : bits 1 := subrange_vec_dec v__2 11 11 in
+ let i31 : bits 3 := subrange_vec_dec v__2 5 3 in
+ let i11 : bits 1 := subrange_vec_dec v__2 12 12 in
+ let i10 : bits 1 := subrange_vec_dec v__2 8 8 in
+ Some
+ (C_J
+ (concat_vec i11
+ (concat_vec i10
+ (concat_vec i98 (concat_vec i7 (concat_vec i6 (concat_vec i5 (concat_vec i4 i31))))))))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B1;B1;B0] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2)))) then
- let i8 : bits 1 := subrange_vec_dec v__2 12 12 in
- let i43 : bits 2 := subrange_vec_dec v__2 11 10 in
let rs : cregbits := subrange_vec_dec v__2 9 7 in
+ let i8 : bits 1 := subrange_vec_dec v__2 12 12 in
let i76 : bits 2 := subrange_vec_dec v__2 6 5 in
- let i21 : bits 2 := subrange_vec_dec v__2 4 3 in
let i5 : bits 1 := subrange_vec_dec v__2 2 2 in
- Some (C_BEQZ (concat_vec i8 (concat_vec i76 (concat_vec i5 (concat_vec i43 i21))),rs))
+ let i43 : bits 2 := subrange_vec_dec v__2 11 10 in
+ let i21 : bits 2 := subrange_vec_dec v__2 4 3 in
+ Some
+ (C_BEQZ
+ (concat_vec i8 (concat_vec i76 (concat_vec i5 (concat_vec i43 i21))), rs))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B1;B1;B1] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B0;B1] : mword 2)))) then
- let i8 : bits 1 := subrange_vec_dec v__2 12 12 in
- let i43 : bits 2 := subrange_vec_dec v__2 11 10 in
let rs : cregbits := subrange_vec_dec v__2 9 7 in
+ let i8 : bits 1 := subrange_vec_dec v__2 12 12 in
let i76 : bits 2 := subrange_vec_dec v__2 6 5 in
- let i21 : bits 2 := subrange_vec_dec v__2 4 3 in
let i5 : bits 1 := subrange_vec_dec v__2 2 2 in
- Some (C_BNEZ (concat_vec i8 (concat_vec i76 (concat_vec i5 (concat_vec i43 i21))),rs))
+ let i43 : bits 2 := subrange_vec_dec v__2 11 10 in
+ let i21 : bits 2 := subrange_vec_dec v__2 4 3 in
+ Some
+ (C_BNEZ
+ (concat_vec i8 (concat_vec i76 (concat_vec i5 (concat_vec i43 i21))), rs))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B0;B0;B0] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B1;B0] : mword 2)))) then
- let nzui5 : bits 1 := subrange_vec_dec v__2 12 12 in
let rsd : regbits := subrange_vec_dec v__2 11 7 in
+ let nzui5 : bits 1 := subrange_vec_dec v__2 12 12 in
let nzui40 : bits 5 := subrange_vec_dec v__2 6 2 in
let shamt : bits 6 := concat_vec nzui5 nzui40 in
if sumbool_of_bool ((orb (eq_vec shamt (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))
(Z.eqb
- (projT1 ((regbits_to_regno rsd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rsd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno zreg)
+ (projT1 ((build_ex (projT1 (regbits_to_regno zreg)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))) then
None
- else Some (C_SLLI (shamt,rsd))
+ else Some (C_SLLI (shamt, rsd))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B0;B1;B0] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B1;B0] : mword 2)))) then
+ let ui76 : bits 2 := subrange_vec_dec v__2 3 2 in
let ui5 : bits 1 := subrange_vec_dec v__2 12 12 in
- let rd : regbits := subrange_vec_dec v__2 11 7 in
let ui42 : bits 3 := subrange_vec_dec v__2 6 4 in
- let ui76 : bits 2 := subrange_vec_dec v__2 3 2 in
+ let rd : regbits := subrange_vec_dec v__2 11 7 in
let uimm : bits 6 := concat_vec ui76 (concat_vec ui5 ui42) in
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno zreg)
+ (projT1 ((build_ex (projT1 (regbits_to_regno zreg)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
None
- else Some (C_LWSP (uimm,rd))
+ else Some (C_LWSP (uimm, rd))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B0;B1;B1] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B1;B0] : mword 2)))) then
+ let ui86 : bits 3 := subrange_vec_dec v__2 4 2 in
let ui5 : bits 1 := subrange_vec_dec v__2 12 12 in
- let rd : regbits := subrange_vec_dec v__2 11 7 in
let ui43 : bits 2 := subrange_vec_dec v__2 6 5 in
- let ui86 : bits 3 := subrange_vec_dec v__2 4 2 in
+ let rd : regbits := subrange_vec_dec v__2 11 7 in
let uimm : bits 6 := concat_vec ui86 (concat_vec ui5 ui43) in
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno zreg)
+ (projT1 ((build_ex (projT1 (regbits_to_regno zreg)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
None
- else Some (C_LDSP (uimm,rd))
+ else Some (C_LDSP (uimm, rd))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B1;B1;B0] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B1;B0] : mword 2)))) then
- let ui52 : bits 4 := subrange_vec_dec v__2 12 9 in
let ui76 : bits 2 := subrange_vec_dec v__2 8 7 in
+ let ui52 : bits 4 := subrange_vec_dec v__2 12 9 in
let rs2 : regbits := subrange_vec_dec v__2 6 2 in
let uimm : bits 6 := concat_vec ui76 ui52 in
- Some (C_SWSP (uimm,rs2))
+ Some
+ (C_SWSP
+ (uimm, rs2))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 13) (vec_of_bits [B1;B1;B1] : mword 3))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B1;B0] : mword 2)))) then
- let ui53 : bits 3 := subrange_vec_dec v__2 12 10 in
let ui86 : bits 3 := subrange_vec_dec v__2 9 7 in
+ let ui53 : bits 3 := subrange_vec_dec v__2 12 10 in
let rs2 : regbits := subrange_vec_dec v__2 6 2 in
let uimm : bits 6 := concat_vec ui86 ui53 in
- Some (C_SDSP (uimm,rs2))
+ Some
+ (C_SDSP
+ (uimm, rs2))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 12) (vec_of_bits [B1;B0;B0;B0] : mword 4))
(eq_vec (subrange_vec_dec v__2 6 0) (vec_of_bits [B0;B0;B0;B0;B0;B1;B0] : mword 7))))
then
let rs1 : regbits := subrange_vec_dec v__2 11 7 in
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno zreg)
+ (projT1 ((build_ex (projT1 (regbits_to_regno zreg)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
None
- else Some (C_JR rs1)
+ else Some (C_JR (rs1))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 12) (vec_of_bits [B1;B0;B0;B1] : mword 4))
(eq_vec (subrange_vec_dec v__2 6 0) (vec_of_bits [B0;B0;B0;B0;B0;B1;B0] : mword 7))))
then
let rs1 : regbits := subrange_vec_dec v__2 11 7 in
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno zreg)
+ (projT1 ((build_ex (projT1 (regbits_to_regno zreg)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) then
None
- else Some (C_JALR rs1)
+ else Some (C_JALR (rs1))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 12) (vec_of_bits [B1;B0;B0;B0] : mword 4))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B1;B0] : mword 2)))) then
- let rd : regbits := subrange_vec_dec v__2 11 7 in
let rs2 : regbits := subrange_vec_dec v__2 6 2 in
+ let rd : regbits := subrange_vec_dec v__2 11 7 in
if sumbool_of_bool ((orb
(Z.eqb
- (projT1 ((regbits_to_regno rs2)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno zreg)
+ (projT1 ((build_ex (projT1 (regbits_to_regno zreg)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))
(Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno zreg)
+ (projT1 ((build_ex (projT1 (regbits_to_regno zreg)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))) then
None
- else Some (C_MV (rd,rs2))
+ else Some (C_MV (rd, rs2))
else if ((andb (eq_vec (subrange_vec_dec v__2 15 12) (vec_of_bits [B1;B0;B0;B1] : mword 4))
(eq_vec (subrange_vec_dec v__2 1 0) (vec_of_bits [B1;B0] : mword 2)))) then
let rsd : regbits := subrange_vec_dec v__2 11 7 in
let rs2 : regbits := subrange_vec_dec v__2 6 2 in
if sumbool_of_bool ((orb
(Z.eqb
- (projT1 ((regbits_to_regno rsd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rsd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno zreg)
+ (projT1 ((build_ex (projT1 (regbits_to_regno zreg)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))
(Z.eqb
- (projT1 ((regbits_to_regno rs2)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno zreg)
+ (projT1 ((build_ex (projT1 (regbits_to_regno zreg)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))) then
None
- else Some (C_ADD (rsd,rs2))
+ else Some (C_ADD (rsd, rs2))
else if ((eq_vec v__2 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 16)))
then
- Some (C_ILLEGAL tt)
+ Some
+ (C_ILLEGAL
+ (tt))
else None.
-Fixpoint execute_WFI (g__26 : unit)
+Definition execute_WFI '(tt : unit)
: M (bool) :=
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
(match w__0 with
@@ -10454,26 +11502,30 @@ Fixpoint execute_WFI (g__26 : unit)
end)
: M (bool).
-Fixpoint execute_UTYPE (imm : mword 20) (rd : mword 5) (op : uop)
+Definition execute_UTYPE (imm : mword 20) (rd : mword 5) (op : uop)
: M (bool) :=
let off : xlenbits :=
EXTS 64 (concat_vec imm (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)) in
match op with
| RISCV_LUI => returnm (off : xlenbits)
| RISCV_AUIPC =>
- (read_reg PC_ref : M (mword 64)) >>= fun w__0 : mword 64 =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__0 : xlenbits =>
returnm ((add_vec w__0 off)
- : xlenbits)
+ : mword 64)
end >>= fun ret : xlenbits =>
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) ret >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) ret >>
returnm (true
: bool).
-Fixpoint execute_THREAD_START (g__29 : unit) : bool := true.
+Definition execute_THREAD_START '(tt : unit) : bool := true.
-Fixpoint execute_STORECON (aq : bool) (rl : bool) (rs2 : mword 5) (rs1 : mword 5) (width : word_width) (rd : mword 5)
+Definition execute_STORECON (aq : bool) (rl : bool) (rs2 : mword 5) (rs1 : mword 5) (width : word_width) (rd : mword 5)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun vaddr : xlenbits =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun vaddr : xlenbits =>
match width with
| BYTE => returnm (true : bool)
| HALF =>
@@ -10488,8 +11540,10 @@ Fixpoint execute_STORECON (aq : bool) (rl : bool) (rs2 : mword 5) (rs1 : mword 5
(if ((negb aligned)) then handle_mem_exception vaddr E_SAMO_Addr_Align >> returnm (false : bool)
else
match_reservation vaddr >>= fun w__1 : bool =>
- (if ((eq_vec ((bool_to_bits w__1) : mword 1) ((bool_to_bits false) : mword 1))) then
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
+ (if ((Bool.eqb w__1 false)) then
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
(EXTZ 64 (vec_of_bits [B1] : mword 1)) >>
returnm (true
: bool)
@@ -10507,7 +11561,7 @@ Fixpoint execute_STORECON (aq : bool) (rl : bool) (rs2 : mword 5) (rs1 : mword 5
| MemException (e) => handle_mem_exception addr e >> returnm (false : bool)
| MemValue (_) =>
rX
- (projT1 ((regbits_to_regno rs2)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val =>
match width with
| WORD =>
@@ -10520,7 +11574,7 @@ Fixpoint execute_STORECON (aq : bool) (rl : bool) (rs2 : mword 5) (rs1 : mword 5
(match res with
| MemValue (true) =>
wX
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
(EXTZ 64 (vec_of_bits [B0] : mword 1)) >>
let '_ := (cancel_reservation tt) : unit in
@@ -10528,7 +11582,7 @@ Fixpoint execute_STORECON (aq : bool) (rl : bool) (rs2 : mword 5) (rs1 : mword 5
: bool)
| MemValue (false) =>
wX
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
(EXTZ 64 (vec_of_bits [B1] : mword 1)) >>
let '_ := (cancel_reservation tt) : unit in
@@ -10544,9 +11598,11 @@ Fixpoint execute_STORECON (aq : bool) (rl : bool) (rs2 : mword 5) (rs1 : mword 5
: M (bool))
: M (bool).
-Fixpoint execute_STORE (imm : mword 12) (rs2 : mword 5) (rs1 : mword 5) (width : word_width) (aq : bool) (rl : bool)
+Definition execute_STORE (imm : mword 12) (rs2 : mword 5) (rs1 : mword 5) (width : word_width) (aq : bool) (rl : bool)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
let vaddr : xlenbits := add_vec w__0 (EXTS 64 imm) in
check_misaligned vaddr width >>= fun w__1 : bool =>
(if (w__1) then handle_mem_exception vaddr E_SAMO_Addr_Align >> returnm (false : bool)
@@ -10564,7 +11620,9 @@ Fixpoint execute_STORE (imm : mword 12) (rs2 : mword 5) (rs1 : mword 5) (width :
(match eares with
| MemException (e) => handle_mem_exception addr e >> returnm (false : bool)
| MemValue (_) =>
- rX (projT1 ((regbits_to_regno rs2) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val =>
match width with
| BYTE =>
(mem_write_value addr 1 (subrange_vec_dec rs2_val 7 0) aq rl false)
@@ -10590,9 +11648,9 @@ Fixpoint execute_STORE (imm : mword 12) (rs2 : mword 5) (rs1 : mword 5) (width :
: M (bool))
: M (bool).
-Fixpoint execute_STOP_FETCHING (g__28 : unit) : bool := true.
+Definition execute_STOP_FETCHING '(tt : unit) : bool := true.
-Fixpoint execute_SRET (g__24 : unit)
+Definition execute_SRET '(tt : unit)
: M (bool) :=
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
match w__0 with
@@ -10604,24 +11662,26 @@ Fixpoint execute_SRET (g__24 : unit)
: M (unit)
else
read_reg cur_privilege_ref >>= fun w__2 : Privilege =>
- (read_reg PC_ref : M (mword 64)) >>= fun w__3 : mword 64 =>
- handle_exception w__2 (CTL_SRET tt) w__3 >>= fun w__4 : xlenbits =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__3 : xlenbits =>
+ handle_exception w__2 (CTL_SRET (tt)) w__3 >>= fun w__4 : mword 64 =>
write_reg nextPC_ref w__4
: M (unit))
: M (unit)
| Machine =>
read_reg cur_privilege_ref >>= fun w__5 : Privilege =>
- (read_reg PC_ref : M (mword 64)) >>= fun w__6 : mword 64 =>
- handle_exception w__5 (CTL_SRET tt) w__6 >>= fun w__7 : xlenbits =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__6 : xlenbits =>
+ handle_exception w__5 (CTL_SRET (tt)) w__6 >>= fun w__7 : mword 64 =>
write_reg nextPC_ref w__7
: M (unit)
end >>
returnm (false
: bool).
-Fixpoint execute_SHIFTW (shamt : mword 5) (rs1 : mword 5) (rd : mword 5) (op : sop)
+Definition execute_SHIFTW (shamt : mword 5) (rs1 : mword 5) (rd : mword 5) (op : sop)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
let rs1_val := subrange_vec_dec w__0 31 0 in
let result : bits 32 :=
match op with
@@ -10629,38 +11689,47 @@ Fixpoint execute_SHIFTW (shamt : mword 5) (rs1 : mword 5) (rd : mword 5) (op : s
| RISCV_SRLI => shift_bits_right rs1_val shamt
| RISCV_SRAI => shift_right_arith32 rs1_val shamt
end in
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (EXTS 64 result) >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) (EXTS 64 result) >>
returnm (true
: bool).
-Fixpoint execute_SHIFTIWOP (shamt : mword 5) (rs1 : mword 5) (rd : mword 5) (op : sopw)
+Definition execute_SHIFTIWOP (shamt : mword 5) (rs1 : mword 5) (rd : mword 5) (op : sopw)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
let result : xlenbits :=
match op with
| RISCV_SLLIW => EXTS 64 (shift_bits_left (subrange_vec_dec rs1_val 31 0) shamt)
| RISCV_SRLIW => EXTS 64 (shift_bits_right (subrange_vec_dec rs1_val 31 0) shamt)
| RISCV_SRAIW => EXTS 64 (shift_right_arith32 (subrange_vec_dec rs1_val 31 0) shamt)
end in
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) result >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) result >>
returnm (true
: bool).
-Fixpoint execute_SHIFTIOP (shamt : mword 6) (rs1 : mword 5) (rd : mword 5) (op : sop)
+Definition execute_SHIFTIOP (shamt : mword 6) (rs1 : mword 5) (rd : mword 5) (op : sop)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
let result : xlenbits :=
match op with
| RISCV_SLLI => shift_bits_left rs1_val shamt
| RISCV_SRLI => shift_bits_right rs1_val shamt
| RISCV_SRAI => shift_right_arith64 rs1_val shamt
end in
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) result >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) result >>
returnm (true
: bool).
-Fixpoint execute_SFENCE_VMA (rs1 : mword 5) (rs2 : mword 5)
+Definition execute_SFENCE_VMA (rs1 : mword 5) (rs2 : mword 5)
: M (bool) :=
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
(if ((eq_vec ((privLevel_to_bits w__0) : mword 2) ((privLevel_to_bits User) : mword 2))) then
@@ -10668,51 +11737,57 @@ Fixpoint execute_SFENCE_VMA (rs1 : mword 5) (rs2 : mword 5)
else
read_reg mstatus_ref >>= fun w__1 : Mstatus =>
read_reg mstatus_ref >>= fun w__2 : Mstatus =>
- let p__20 := (architecture (_get_Mstatus_SXL w__1), _get_Mstatus_TVM w__2) in
- (match p__20 with
+ let p__16 := (architecture (_get_Mstatus_SXL w__1), _get_Mstatus_TVM w__2) in
+ (match p__16 with
| (Some (RV64), v_0) =>
(if ((eq_vec v_0 ((bool_to_bits true) : mword 1))) then
handle_illegal tt >> returnm (false : bool)
else if ((eq_vec v_0 ((bool_to_bits false) : mword 1))) then
(if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
returnm (None
- : option vaddr39)
+ : option (mword 39))
else
rX
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__3 : mword 64 =>
- returnm ((Some (subrange_vec_dec w__3 38 0))
- : option vaddr39)) >>= fun addr : option vaddr39 =>
+ returnm ((Some
+ (subrange_vec_dec w__3 38 0))
+ : option (mword 39))) >>= fun addr : option vaddr39 =>
(if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs2)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
returnm (None
- : option asid64)
+ : option (mword 16))
else
rX
- (projT1 ((regbits_to_regno rs2)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__4 : mword 64 =>
- returnm ((Some (subrange_vec_dec w__4 15 0))
- : option asid64)) >>= fun asid : option asid64 =>
+ returnm ((Some
+ (subrange_vec_dec w__4 15 0))
+ : option (mword 16))) >>= fun asid : option asid64 =>
flushTLB asid addr >> returnm (true : bool)
else
- (match (Some RV64, v_0) with
- | (g__18, g__19) => (internal_error "unimplemented sfence architecture") : M (bool)
+ (match (Some (RV64), v_0) with
+ | (g__14, g__15) => (internal_error "unimplemented sfence architecture") : M (bool)
end)
: M (bool))
: M (bool)
- | (g__18, g__19) => (internal_error "unimplemented sfence architecture") : M (bool)
+ | (g__14, g__15) => (internal_error "unimplemented sfence architecture") : M (bool)
end)
: M (bool))
: M (bool).
-Fixpoint execute_RTYPEW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (op : ropw)
+Definition execute_RTYPEW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (op : ropw)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
let rs1_val := subrange_vec_dec w__0 31 0 in
- rX (projT1 ((regbits_to_regno rs2) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__1 : mword 64 =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__1 : mword 64 =>
let rs2_val := subrange_vec_dec w__1 31 0 in
let result : bits 32 :=
match op with
@@ -10722,15 +11797,20 @@ Fixpoint execute_RTYPEW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (op : rop
| RISCV_SRLW => shift_bits_right rs1_val (subrange_vec_dec rs2_val 4 0)
| RISCV_SRAW => shift_right_arith32 rs1_val (subrange_vec_dec rs2_val 4 0)
end in
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (EXTS 64 result) >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) (EXTS 64 result) >>
returnm (true
: bool).
-Fixpoint execute_RTYPE (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (op : rop)
+Definition execute_RTYPE (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (op : rop)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
- rX (projT1 ((regbits_to_regno rs2) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val =>
let result : xlenbits :=
match op with
| RISCV_ADD => add_vec rs1_val rs2_val
@@ -10744,105 +11824,137 @@ Fixpoint execute_RTYPE (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (op : rop)
| RISCV_OR => or_vec rs1_val rs2_val
| RISCV_AND => and_vec rs1_val rs2_val
end in
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) result >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) result >>
returnm (true
: bool).
-Fixpoint execute_RISCV_JALR (imm : mword 12) (rs1 : mword 5) (rd : mword 5)
+Definition execute_RISCV_JALR (imm : mword 12) (rs1 : mword 5) (rd : mword 5)
: M (bool) :=
- (read_reg nextPC_ref : M (mword 64)) >>= fun w__0 : mword 64 =>
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) w__0 >>
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__1 : mword 64 =>
+ (read_reg nextPC_ref : M (mword 64)) >>= fun w__0 : xlenbits =>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) w__0 >>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__1 : mword 64 =>
let newPC : xlenbits := add_vec w__1 (EXTS 64 imm) in
write_reg nextPC_ref (concat_vec (subrange_vec_dec newPC 63 1) (vec_of_bits [B0] : mword 1)) >>
returnm (true
: bool).
-Fixpoint execute_RISCV_JAL (imm : mword 21) (rd : mword 5)
+Definition execute_RISCV_JAL (imm : mword 21) (rd : mword 5)
: M (bool) :=
(read_reg PC_ref : M (mword 64)) >>= fun pc : xlenbits =>
- (read_reg nextPC_ref : M (mword 64)) >>= fun w__0 : mword 64 =>
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) w__0 >>
+ (read_reg nextPC_ref : M (mword 64)) >>= fun w__0 : xlenbits =>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) w__0 >>
let offset : xlenbits := EXTS 64 imm in
write_reg nextPC_ref (add_vec pc offset) >> returnm (true : bool).
-Fixpoint execute_REMW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool)
+Definition execute_REMW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
let rs1_val := subrange_vec_dec w__0 31 0 in
- rX (projT1 ((regbits_to_regno rs2) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__1 : mword 64 =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__1 : mword 64 =>
let rs2_val := subrange_vec_dec w__1 31 0 in
let rs1_int : Z := if (s) then projT1 (sint rs1_val) else projT1 (uint rs1_val) in
let rs2_int : Z := if (s) then projT1 (sint rs2_val) else projT1 (uint rs2_val) in
let r : Z :=
if sumbool_of_bool ((Z.eqb rs2_int 0)) then rs1_int
else rem_round_zero rs1_int rs2_int in
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (EXTS 64 (to_bits 32 r)) >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) (EXTS 64 (to_bits 32 r)) >>
returnm (true
: bool).
-Fixpoint execute_REM (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool)
+Definition execute_REM (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
- rX (projT1 ((regbits_to_regno rs2) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val =>
let rs1_int : Z := if (s) then projT1 (sint rs1_val) else projT1 (uint rs1_val) in
let rs2_int : Z := if (s) then projT1 (sint rs2_val) else projT1 (uint rs2_val) in
let r : Z :=
if sumbool_of_bool ((Z.eqb rs2_int 0)) then rs1_int
else rem_round_zero rs1_int rs2_int in
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (to_bits xlen r) >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) (to_bits xlen r) >>
returnm (true
: bool).
-Fixpoint execute_NOP (g__27 : unit) : bool := true.
+Definition execute_NOP '(tt : unit) : bool := true.
-Fixpoint execute_MULW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5)
+Definition execute_MULW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
let rs1_val := subrange_vec_dec w__0 31 0 in
- rX (projT1 ((regbits_to_regno rs2) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__1 : mword 64 =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__1 : mword 64 =>
let rs2_val := subrange_vec_dec w__1 31 0 in
let rs1_int : Z := projT1 (sint rs1_val) in
let rs2_int : Z := projT1 (sint rs2_val) in
let result32 := subrange_vec_dec (to_bits 64 (Z.mul rs1_int rs2_int)) 31 0 in
let result : xlenbits := EXTS 64 result32 in
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) result >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) result >>
returnm (true
: bool).
-Fixpoint execute_MUL (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (high : bool) (signed1 : bool) (signed2 : bool)
+Definition execute_MUL (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (high : bool) (signed1 : bool) (signed2 : bool)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
- rX (projT1 ((regbits_to_regno rs2) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val =>
let rs1_int : Z := if (signed1) then projT1 (sint rs1_val) else projT1 (uint rs1_val) in
let rs2_int : Z := if (signed2) then projT1 (sint rs2_val) else projT1 (uint rs2_val) in
let result128 := to_bits 128 (Z.mul rs1_int rs2_int) in
let result :=
if (high) then subrange_vec_dec result128 127 64
else subrange_vec_dec result128 63 0 in
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) result >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) result >>
returnm (true
: bool).
-Fixpoint execute_MRET (g__23 : unit)
+Definition execute_MRET '(tt : unit)
: M (bool) :=
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
(if ((eq_vec ((privLevel_to_bits w__0) : mword 2) ((privLevel_to_bits Machine) : mword 2)))
then
read_reg cur_privilege_ref >>= fun w__1 : Privilege =>
- (read_reg PC_ref : M (mword 64)) >>= fun w__2 : mword 64 =>
- handle_exception w__1 (CTL_MRET tt) w__2 >>= fun w__3 : xlenbits =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__2 : xlenbits =>
+ handle_exception w__1 (CTL_MRET (tt)) w__2 >>= fun w__3 : mword 64 =>
write_reg nextPC_ref w__3
: M (unit)
else (handle_illegal tt) : M (unit)) >>
returnm (false
: bool).
-Fixpoint execute_LOADRES (aq : bool) (rl : bool) (rs1 : mword 5) (width : word_width) (rd : mword 5)
+Definition execute_LOADRES (aq : bool) (rl : bool) (rs1 : mword 5) (width : word_width) (rd : mword 5)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun vaddr : xlenbits =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun vaddr : xlenbits =>
match width with
| BYTE => returnm (true : bool)
| HALF =>
@@ -10862,12 +11974,12 @@ Fixpoint execute_LOADRES (aq : bool) (rl : bool) (rs1 : mword 5) (width : word_w
| TR_Address (addr) =>
(match width with
| WORD =>
- mem_read addr 4 aq rl true >>= fun w__2 : MemoryOpResult (mword 32) =>
- (process_loadres (n := 4) rd vaddr w__2 false)
+ mem_read addr 4 aq rl true >>= fun w__2 : MemoryOpResult (mword (8 * 4)) =>
+ (process_loadres rd vaddr w__2 false)
: M (bool)
| DOUBLE =>
- mem_read addr 8 aq rl true >>= fun w__4 : MemoryOpResult (mword 64) =>
- (process_loadres (n := 8) rd vaddr w__4 false)
+ mem_read addr 8 aq rl true >>= fun w__4 : MemoryOpResult (mword (8 * 8)) =>
+ (process_loadres rd vaddr w__4 false)
: M (bool)
| _ => (internal_error "LOADRES expected WORD or DOUBLE") : M (bool)
end)
@@ -10876,9 +11988,11 @@ Fixpoint execute_LOADRES (aq : bool) (rl : bool) (rs1 : mword 5) (width : word_w
: M (bool))
: M (bool).
-Fixpoint execute_LOAD (imm : mword 12) (rs1 : mword 5) (rd : mword 5) (is_unsigned : bool) (width : word_width) (aq : bool) (rl : bool)
+Definition execute_LOAD (imm : mword 12) (rs1 : mword 5) (rd : mword 5) (is_unsigned : bool) (width : word_width) (aq : bool) (rl : bool)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
let vaddr : xlenbits := add_vec w__0 (EXTS 64 imm) in
check_misaligned vaddr width >>= fun w__1 : bool =>
(if (w__1) then handle_mem_exception vaddr E_Load_Addr_Align >> returnm (false : bool)
@@ -10889,20 +12003,20 @@ Fixpoint execute_LOAD (imm : mword 12) (rs1 : mword 5) (rd : mword 5) (is_unsign
| TR_Address (addr) =>
(match width with
| BYTE =>
- mem_read addr 1 aq rl false >>= fun w__3 : MemoryOpResult (mword 8) =>
- (process_load (n := 1) rd vaddr w__3 is_unsigned)
+ mem_read addr 1 aq rl false >>= fun w__3 : MemoryOpResult (mword (8 * 1)) =>
+ (process_load rd vaddr w__3 is_unsigned)
: M (bool)
| HALF =>
- mem_read addr 2 aq rl false >>= fun w__5 : MemoryOpResult (mword 16) =>
- (process_load (n := 2) rd vaddr w__5 is_unsigned)
+ mem_read addr 2 aq rl false >>= fun w__5 : MemoryOpResult (mword (8 * 2)) =>
+ (process_load rd vaddr w__5 is_unsigned)
: M (bool)
| WORD =>
- mem_read addr 4 aq rl false >>= fun w__7 : MemoryOpResult (mword 32) =>
- (process_load (n := 4) rd vaddr w__7 is_unsigned)
+ mem_read addr 4 aq rl false >>= fun w__7 : MemoryOpResult (mword (8 * 4)) =>
+ (process_load rd vaddr w__7 is_unsigned)
: M (bool)
| DOUBLE =>
- mem_read addr 8 aq rl false >>= fun w__9 : MemoryOpResult (mword 64) =>
- (process_load (n := 8) rd vaddr w__9 is_unsigned)
+ mem_read addr 8 aq rl false >>= fun w__9 : MemoryOpResult (mword (8 * 8)) =>
+ (process_load rd vaddr w__9 is_unsigned)
: M (bool)
end)
: M (bool)
@@ -10910,9 +12024,11 @@ Fixpoint execute_LOAD (imm : mword 12) (rs1 : mword 5) (rd : mword 5) (is_unsign
: M (bool))
: M (bool).
-Fixpoint execute_ITYPE (imm : mword 12) (rs1 : mword 5) (rd : mword 5) (op : iop)
+Definition execute_ITYPE (imm : mword 12) (rs1 : mword 5) (rd : mword 5) (op : iop)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
let immext : xlenbits := EXTS 64 imm in
let result : xlenbits :=
match op with
@@ -10923,15 +12039,19 @@ Fixpoint execute_ITYPE (imm : mword 12) (rs1 : mword 5) (rd : mword 5) (op : iop
| RISCV_ORI => or_vec rs1_val immext
| RISCV_ANDI => and_vec rs1_val immext
end in
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) result >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) result >>
returnm (true
: bool).
-Fixpoint execute_ILLEGAL (s : mword 32) : M (bool) := handle_illegal tt >> returnm (false : bool).
+Definition execute_ILLEGAL (s : mword 32)
+: M (bool) :=
+ handle_illegal tt >> returnm (false : bool).
-Fixpoint execute_FENCEI (g__21 : unit) : bool := true.
+Definition execute_FENCEI '(tt : unit) : bool := true.
-Fixpoint execute_FENCE (pred : mword 4) (succ : mword 4)
+Definition execute_FENCE (pred : mword 4) (succ : mword 4)
: M (bool) :=
match (pred, succ) with
| (v__132, v__133) =>
@@ -10985,7 +12105,7 @@ Fixpoint execute_FENCE (pred : mword 4) (succ : mword 4)
returnm (true
: bool).
-Fixpoint execute_ECALL (g__22 : unit)
+Definition execute_ECALL '(tt : unit)
: M (bool) :=
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
let t : sync_exception :=
@@ -10997,20 +12117,24 @@ Fixpoint execute_ECALL (g__22 : unit)
end);
sync_exception_excinfo := (None : option xlenbits) |} in
read_reg cur_privilege_ref >>= fun w__1 : Privilege =>
- (read_reg PC_ref : M (mword 64)) >>= fun w__2 : mword 64 =>
- handle_exception w__1 (CTL_TRAP t) w__2 >>= fun w__3 : xlenbits =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__2 : xlenbits =>
+ handle_exception w__1 (CTL_TRAP (t)) w__2 >>= fun w__3 : mword 64 =>
write_reg nextPC_ref w__3 >> returnm (false : bool).
-Fixpoint execute_EBREAK (g__25 : unit)
+Definition execute_EBREAK '(tt : unit)
: M (bool) :=
- (read_reg PC_ref : M (mword 64)) >>= fun w__0 : mword 64 =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__0 : xlenbits =>
handle_mem_exception w__0 E_Breakpoint >> returnm (false : bool).
-Fixpoint execute_DIVW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool)
+Definition execute_DIVW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
let rs1_val := subrange_vec_dec w__0 31 0 in
- rX (projT1 ((regbits_to_regno rs2) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__1 : mword 64 =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__1 : mword 64 =>
let rs2_val := subrange_vec_dec w__1 31 0 in
let rs1_int : Z := if (s) then projT1 (sint rs1_val) else projT1 (uint rs1_val) in
let rs2_int : Z := if (s) then projT1 (sint rs2_val) else projT1 (uint rs2_val) in
@@ -11019,50 +12143,65 @@ Fixpoint execute_DIVW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool)
if sumbool_of_bool ((andb s
(Z.gtb q (projT1 (sub_range (build_ex (projT1 (pow2 31))) (build_ex 1))))))
then
- (Z.sub ( 0) ((pow 2 31)))
+ projT1 (sub_range (build_ex 0)
+ (build_ex (projT1 ((build_ex (projT1 (ex_int (pow 2 31))))
+ : {n : Z & ArithFact (True)}))))
else q in
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (EXTS 64 (to_bits 32 q')) >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) (EXTS 64 (to_bits 32 q')) >>
returnm (true
: bool).
-Fixpoint execute_DIV (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool)
+Definition execute_DIV (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
- rX (projT1 ((regbits_to_regno rs2) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val =>
let rs1_int : Z := if (s) then projT1 (sint rs1_val) else projT1 (uint rs1_val) in
let rs2_int : Z := if (s) then projT1 (sint rs2_val) else projT1 (uint rs2_val) in
let q : Z := if sumbool_of_bool ((Z.eqb rs2_int 0)) then -1 else quot_round_zero rs1_int rs2_int in
let q' : Z := if sumbool_of_bool ((andb s (Z.gtb q xlen_max_signed))) then xlen_min_signed else q in
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (to_bits xlen q') >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) (to_bits xlen q') >>
returnm (true
: bool).
-Fixpoint execute_C_ILLEGAL (g__30 : unit)
+Definition execute_C_ILLEGAL '(tt : unit)
: M (bool) :=
handle_illegal tt >> returnm (false : bool).
-Fixpoint execute_C_ADDIW (imm : mword 6) (rsd : mword 5)
+Definition execute_C_ADDIW (imm : mword 6) (rsd : mword 5)
: M (bool) :=
let imm : bits 32 := EXTS 32 imm in
- rX (projT1 ((regbits_to_regno rsd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rsd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs_val =>
let res : bits 32 := add_vec (subrange_vec_dec rs_val 31 0) imm in
- wX (projT1 ((regbits_to_regno rsd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (EXTS 64 res) >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rsd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) (EXTS 64 res) >>
returnm (true
: bool).
-Fixpoint execute_CSR (csr : mword 12) (rs1 : mword 5) (rd : mword 5) (is_imm : bool) (op : csrop)
+Definition execute_CSR (csr : mword 12) (rs1 : mword 5) (rd : mword 5) (is_imm : bool) (op : csrop)
: M (bool) :=
- (if (is_imm) then returnm ((EXTZ 64 rs1) : xlenbits)
+ (if (is_imm) then returnm ((EXTZ 64 rs1) : mword 64)
else
- (rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))
- : M (xlenbits)) >>= fun rs1_val : xlenbits =>
+ (rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))
+ : M (mword 64)) >>= fun rs1_val : xlenbits =>
let isWrite : bool :=
match op with
| CSRRW => true
- | _ => if (is_imm) then neq_int (projT1 (uint rs1_val)) 0 else neq_int (projT1 (uint rs1)) 0
+ | _ =>
+ if (is_imm) then neq_range (build_ex (projT1 (uint rs1_val))) (build_ex 0)
+ else neq_range (build_ex (projT1 (uint rs1))) (build_ex 0)
end in
read_reg cur_privilege_ref >>= fun w__1 : Privilege =>
check_CSR csr w__1 isWrite >>= fun w__2 : bool =>
@@ -11079,15 +12218,21 @@ Fixpoint execute_CSR (csr : mword 12) (rs1 : mword 5) (rd : mword 5) (is_imm : b
(writeCSR csr new_val)
: M (unit)
else returnm (tt : unit)) >>
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) csr_val >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) csr_val >>
returnm (true
: bool))
: M (bool).
-Fixpoint execute_BTYPE (imm : mword 13) (rs2 : mword 5) (rs1 : mword 5) (op : bop)
+Definition execute_BTYPE (imm : mword 13) (rs2 : mword 5) (rs1 : mword 5) (op : bop)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
- rX (projT1 ((regbits_to_regno rs2) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs1_val =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val =>
let taken : bool :=
match op with
| RISCV_BEQ => eq_vec rs1_val rs2_val
@@ -11098,16 +12243,18 @@ Fixpoint execute_BTYPE (imm : mword 13) (rs2 : mword 5) (rs1 : mword 5) (op : bo
| RISCV_BGEU => zopz0zKzJ_u rs1_val rs2_val
end in
(if (taken) then
- (read_reg PC_ref : M (mword 64)) >>= fun w__0 : mword 64 =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__0 : xlenbits =>
write_reg nextPC_ref (add_vec w__0 (EXTS 64 imm))
: M (unit)
else returnm (tt : unit)) >>
returnm (true
: bool).
-Fixpoint execute_AMO (op : amoop) (aq : bool) (rl : bool) (rs2 : mword 5) (rs1 : mword 5) (width : word_width) (rd : mword 5)
+Definition execute_AMO (op : amoop) (aq : bool) (rl : bool) (rs2 : mword 5) (rs1 : mword 5) (width : word_width) (rd : mword 5)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun vaddr : xlenbits =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun vaddr : xlenbits =>
translateAddr vaddr ReadWrite Data >>= fun w__0 : TR_Result =>
(match w__0 with
| TR_Failure (e) => handle_mem_exception vaddr e >> returnm (false : bool)
@@ -11122,19 +12269,21 @@ Fixpoint execute_AMO (op : amoop) (aq : bool) (rl : bool) (rs2 : mword 5) (rs1 :
| MemValue (_) =>
match width with
| WORD =>
- mem_read addr 4 aq (andb aq rl) true >>= fun w__4 : MemoryOpResult (mword 32) =>
- returnm ((extend_value (n := 4) false w__4)
- : MemoryOpResult xlenbits)
+ mem_read addr 4 aq (andb aq rl) true >>= fun w__4 : MemoryOpResult (mword (8 * 4)) =>
+ returnm ((extend_value false w__4)
+ : MemoryOpResult (mword 64))
| DOUBLE =>
- mem_read addr 8 aq (andb aq rl) true >>= fun w__5 : MemoryOpResult (mword 64) =>
- returnm ((extend_value (n := 8) false w__5)
- : MemoryOpResult xlenbits)
- | _ => (internal_error "AMO expected WORD or DOUBLE") : M (MemoryOpResult xlenbits)
+ mem_read addr 8 aq (andb aq rl) true >>= fun w__5 : MemoryOpResult (mword (8 * 8)) =>
+ returnm ((extend_value false w__5)
+ : MemoryOpResult (mword 64))
+ | _ => (internal_error "AMO expected WORD or DOUBLE") : M (MemoryOpResult (mword 64))
end >>= fun rval : MemoryOpResult xlenbits =>
(match rval with
| MemException (e) => handle_mem_exception addr e >> returnm (false : bool)
| MemValue (loaded) =>
- rX (projT1 ((regbits_to_regno rs2) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val : xlenbits =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun rs2_val : xlenbits =>
let result : xlenbits :=
match op with
| AMOSWAP => rs2_val
@@ -11164,7 +12313,7 @@ Fixpoint execute_AMO (op : amoop) (aq : bool) (rl : bool) (rs2 : mword 5) (rs1 :
(match wval with
| MemValue (true) =>
wX
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) loaded >>
returnm (true
: bool)
@@ -11180,12 +12329,15 @@ Fixpoint execute_AMO (op : amoop) (aq : bool) (rl : bool) (rs2 : mword 5) (rs1 :
end)
: M (bool).
-Fixpoint execute_ADDIW (imm : mword 12) (rs1 : mword 5) (rd : mword 5)
+Definition execute_ADDIW (imm : mword 12) (rs1 : mword 5) (rd : mword 5)
: M (bool) :=
- rX (projT1 ((regbits_to_regno rs1) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
+ rX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) >>= fun w__0 : mword 64 =>
let result : xlenbits := add_vec (EXTS 64 imm) w__0 in
- wX (projT1 ((regbits_to_regno rd) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (EXTS 64 (subrange_vec_dec result 31 0)) >>
+ wX
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) (EXTS 64 (subrange_vec_dec result 31 0)) >>
returnm (true
: bool).
@@ -11196,127 +12348,131 @@ match i with
concat_vec (vec_of_bits [B0;B0] : mword 2)
(concat_vec nzimm (vec_of_bits [B0;B0] : mword 2)) in
let rd := creg2reg_bits rdc in
- ( (ITYPE (imm,sp,rd,RISCV_ADDI)))
+ ( (ITYPE (imm, sp, rd, RISCV_ADDI)))
| C_LW (uimm,rsc,rdc) =>
let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0] : mword 2)) in
let rd := creg2reg_bits rdc in
let rs := creg2reg_bits rsc in
- ( (LOAD (imm,rs,rd,false,WORD,false,false)))
+ ( (LOAD (imm, rs, rd, false, WORD, false, false)))
| C_LD (uimm,rsc,rdc) =>
let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0;B0] : mword 3)) in
let rd := creg2reg_bits rdc in
let rs := creg2reg_bits rsc in
- ( (LOAD (imm,rs,rd,false,DOUBLE,false,false)))
+ ( (LOAD (imm, rs, rd, false, DOUBLE, false, false)))
| C_SW (uimm,rsc1,rsc2) =>
let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0] : mword 2)) in
let rs1 := creg2reg_bits rsc1 in
let rs2 := creg2reg_bits rsc2 in
- ( (STORE (imm,rs2,rs1,WORD,false,false)))
+ ( (STORE (imm, rs2, rs1, WORD, false, false)))
| C_SD (uimm,rsc1,rsc2) =>
let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0;B0] : mword 3)) in
let rs1 := creg2reg_bits rsc1 in
let rs2 := creg2reg_bits rsc2 in
- ( (STORE (imm,rs2,rs1,DOUBLE,false,false)))
+ ( (STORE (imm, rs2, rs1, DOUBLE, false, false)))
| C_ADDI (nzi,rsd) =>
let imm : bits 12 := EXTS 12 nzi in
- ( (ITYPE (imm,rsd,rsd,RISCV_ADDI)))
+ ( (ITYPE (imm, rsd, rsd, RISCV_ADDI)))
| C_JAL (imm) =>
- ( (RISCV_JAL (EXTS 21 (concat_vec imm (vec_of_bits [B0] : mword 1)),ra)))
+ ( (RISCV_JAL (EXTS 21 (concat_vec imm (vec_of_bits [B0] : mword 1)), ra)))
| C_LI (imm,rd) =>
let imm : bits 12 := EXTS 12 imm in
- ( (ITYPE (imm,zreg,rd,RISCV_ADDI)))
+ ( (ITYPE (imm, zreg, rd, RISCV_ADDI)))
| C_ADDI16SP (imm) =>
let imm : bits 12 := EXTS 12 (concat_vec imm (vec_of_bits [B0;B0;B0;B0] : mword 4)) in
- ( (ITYPE (imm,sp,sp,RISCV_ADDI)))
+ ( (ITYPE (imm, sp, sp, RISCV_ADDI)))
| C_LUI (imm,rd) =>
let res : bits 20 := EXTS 20 imm in
- ( (UTYPE (res,rd,RISCV_LUI)))
+ ( (UTYPE (res, rd, RISCV_LUI)))
| C_SRLI (shamt,rsd) =>
let rsd := creg2reg_bits rsd in
- ( (SHIFTIOP (shamt,rsd,rsd,RISCV_SRLI)))
+ ( (SHIFTIOP (shamt, rsd, rsd, RISCV_SRLI)))
| C_SRAI (shamt,rsd) =>
let rsd := creg2reg_bits rsd in
- ( (SHIFTIOP (shamt,rsd,rsd,RISCV_SRAI)))
+ ( (SHIFTIOP (shamt, rsd, rsd, RISCV_SRAI)))
| C_ANDI (imm,rsd) =>
let rsd := creg2reg_bits rsd in
- ( (ITYPE (EXTS 12 imm,rsd,rsd,RISCV_ANDI)))
+ ( (ITYPE (EXTS 12 imm, rsd, rsd, RISCV_ANDI)))
| C_SUB (rsd,rs2) =>
let rsd := creg2reg_bits rsd in
let rs2 := creg2reg_bits rs2 in
- ( (RTYPE (rs2,rsd,rsd,RISCV_SUB)))
+ ( (RTYPE (rs2, rsd, rsd, RISCV_SUB)))
| C_XOR (rsd,rs2) =>
let rsd := creg2reg_bits rsd in
let rs2 := creg2reg_bits rs2 in
- ( (RTYPE (rs2,rsd,rsd,RISCV_XOR)))
+ ( (RTYPE (rs2, rsd, rsd, RISCV_XOR)))
| C_OR (rsd,rs2) =>
let rsd := creg2reg_bits rsd in
let rs2 := creg2reg_bits rs2 in
- ( (RTYPE (rs2,rsd,rsd,RISCV_OR)))
+ ( (RTYPE (rs2, rsd, rsd, RISCV_OR)))
| C_AND (rsd,rs2) =>
let rsd := creg2reg_bits rsd in
let rs2 := creg2reg_bits rs2 in
- ( (RTYPE (rs2,rsd,rsd,RISCV_AND)))
+ ( (RTYPE (rs2, rsd, rsd, RISCV_AND)))
| C_SUBW (rsd,rs2) =>
let rsd := creg2reg_bits rsd in
let rs2 := creg2reg_bits rs2 in
- ( (RTYPEW (rs2,rsd,rsd,RISCV_SUBW)))
+ ( (RTYPEW (rs2, rsd, rsd, RISCV_SUBW)))
| C_ADDW (rsd,rs2) =>
let rsd := creg2reg_bits rsd in
let rs2 := creg2reg_bits rs2 in
- ( (RTYPEW (rs2,rsd,rsd,RISCV_ADDW)))
+ ( (RTYPEW (rs2, rsd, rsd, RISCV_ADDW)))
| C_J (imm) =>
- ( (RISCV_JAL (EXTS 21 (concat_vec imm (vec_of_bits [B0] : mword 1)),zreg)))
+ ( (RISCV_JAL (EXTS 21 (concat_vec imm (vec_of_bits [B0] : mword 1)), zreg)))
| C_BEQZ (imm,rs) =>
(
- (BTYPE (EXTS 13 (concat_vec imm (vec_of_bits [B0] : mword 1)),zreg,creg2reg_bits rs,RISCV_BEQ)))
+ (BTYPE
+ (EXTS 13 (concat_vec imm (vec_of_bits [B0] : mword 1)), zreg, creg2reg_bits rs,
+ RISCV_BEQ)))
| C_BNEZ (imm,rs) =>
(
- (BTYPE (EXTS 13 (concat_vec imm (vec_of_bits [B0] : mword 1)),zreg,creg2reg_bits rs,RISCV_BNE)))
+ (BTYPE
+ (EXTS 13 (concat_vec imm (vec_of_bits [B0] : mword 1)), zreg, creg2reg_bits rs,
+ RISCV_BNE)))
- | C_SLLI (shamt,rsd) => ( (SHIFTIOP (shamt,rsd,rsd,RISCV_SLLI)))
+ | C_SLLI (shamt,rsd) => ( (SHIFTIOP (shamt, rsd, rsd, RISCV_SLLI)))
| C_LWSP (uimm,rd) =>
let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0] : mword 2)) in
- ( (LOAD (imm,sp,rd,false,WORD,false,false)))
+ ( (LOAD (imm, sp, rd, false, WORD, false, false)))
| C_LDSP (uimm,rd) =>
let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0;B0] : mword 3)) in
- ( (LOAD (imm,sp,rd,false,DOUBLE,false,false)))
+ ( (LOAD (imm, sp, rd, false, DOUBLE, false, false)))
| C_SWSP (uimm,rs2) =>
let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0] : mword 2)) in
- ( (STORE (imm,rs2,sp,WORD,false,false)))
+ ( (STORE (imm, rs2, sp, WORD, false, false)))
| C_SDSP (uimm,rs2) =>
let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0;B0] : mword 3)) in
- ( (STORE (imm,rs2,sp,DOUBLE,false,false)))
+ ( (STORE (imm, rs2, sp, DOUBLE, false, false)))
| C_JR (rs1) =>
- ( (RISCV_JALR (EXTZ 12 (vec_of_bits [B0] : mword 1),rs1,zreg)))
+ ( (RISCV_JALR (EXTZ 12 (vec_of_bits [B0] : mword 1), rs1, zreg)))
| C_JALR (rs1) =>
- ( (RISCV_JALR (EXTZ 12 (vec_of_bits [B0] : mword 1),rs1,ra)))
- | C_MV (rd,rs2) => ( (RTYPE (rs2,zreg,rd,RISCV_ADD)))
- | C_ADD (rsd,rs2) => ( (RTYPE (rs2,rsd,rsd,RISCV_ADD)))
-| _ => i
+ ( (RISCV_JALR (EXTZ 12 (vec_of_bits [B0] : mword 1), rs1, ra)))
+ | C_MV (rd,rs2) => ( (RTYPE (rs2, zreg, rd, RISCV_ADD)))
+ | C_ADD (rsd,rs2) => ( (RTYPE (rs2, rsd, rsd, RISCV_ADD)))
+| i => i
end.
Fixpoint execute (merge_var : ast)
@@ -11345,23 +12501,23 @@ let merge_var := expand_ast merge_var in
| DIVW (rs2,rs1,rd,s) => (execute_DIVW rs2 rs1 rd s) : M (bool)
| REMW (rs2,rs1,rd,s) => (execute_REMW rs2 rs1 rd s) : M (bool)
| FENCE (pred,succ) => (execute_FENCE pred succ) : M (bool)
- | FENCEI (g__21) => returnm ((execute_FENCEI g__21) : bool)
- | ECALL (g__22) => (execute_ECALL g__22) : M (bool)
- | MRET (g__23) => (execute_MRET g__23) : M (bool)
- | SRET (g__24) => (execute_SRET g__24) : M (bool)
- | EBREAK (g__25) => (execute_EBREAK g__25) : M (bool)
- | WFI (g__26) => (execute_WFI g__26) : M (bool)
+ | FENCEI (arg0) => returnm ((execute_FENCEI arg0) : bool)
+ | ECALL (arg0) => (execute_ECALL arg0) : M (bool)
+ | MRET (arg0) => (execute_MRET arg0) : M (bool)
+ | SRET (arg0) => (execute_SRET arg0) : M (bool)
+ | EBREAK (arg0) => (execute_EBREAK arg0) : M (bool)
+ | WFI (arg0) => (execute_WFI arg0) : M (bool)
| SFENCE_VMA (rs1,rs2) => (execute_SFENCE_VMA rs1 rs2) : M (bool)
| LOADRES (aq,rl,rs1,width,rd) => (execute_LOADRES aq rl rs1 width rd) : M (bool)
| STORECON (aq,rl,rs2,rs1,width,rd) => (execute_STORECON aq rl rs2 rs1 width rd) : M (bool)
| AMO (op,aq,rl,rs2,rs1,width,rd) => (execute_AMO op aq rl rs2 rs1 width rd) : M (bool)
| CSR (csr,rs1,rd,is_imm,op) => (execute_CSR csr rs1 rd is_imm op) : M (bool)
- | NOP (g__27) => returnm ((execute_NOP g__27) : bool)
+ | NOP (arg0) => returnm ((execute_NOP arg0) : bool)
| C_ADDIW (imm,rsd) => (execute_C_ADDIW imm rsd) : M (bool)
- | STOP_FETCHING (g__28) => returnm ((execute_STOP_FETCHING g__28) : bool)
- | THREAD_START (g__29) => returnm ((execute_THREAD_START g__29) : bool)
+ | STOP_FETCHING (arg0) => returnm ((execute_STOP_FETCHING arg0) : bool)
+ | THREAD_START (arg0) => returnm ((execute_THREAD_START arg0) : bool)
| ILLEGAL (s) => (execute_ILLEGAL s) : M (bool)
- | C_ILLEGAL (g__30) => (execute_C_ILLEGAL g__30) : M (bool)
+ | C_ILLEGAL (arg0) => (execute_C_ILLEGAL arg0) : M (bool)
| _ => Fail "Unexpanded instruction"
end.
@@ -11628,28 +12784,26 @@ Definition assembly_forwards (arg_ : ast)
(*
Definition assembly_backwards (arg_ : string)
: M (ast) :=
- let _stringappend_1112_ := arg_ in
- match (utype_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1113_,(existT _ _stringappend_1114_ _)) =>
- let _stringappend_1115_ := string_drop _stringappend_1112_ (build_ex _stringappend_1114_) in
- match (spc_matches_prefix _stringappend_1115_) with
- | Some (_stringappend_1116_,(existT _ _stringappend_1117_ _)) =>
- let _stringappend_1118_ := string_drop _stringappend_1115_ (build_ex _stringappend_1117_) in
- match (reg_name_matches_prefix _stringappend_1118_) with
- | Some (_stringappend_1119_,(existT _ _stringappend_1120_ _)) =>
- let _stringappend_1121_ :=
- string_drop _stringappend_1118_ (build_ex _stringappend_1120_) in
- sep_matches_prefix _stringappend_1121_ >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_756_ := arg_ in
+ match (utype_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_757_ _)) =>
+ let _stringappend_758_ := string_drop _stringappend_756_ (build_ex _stringappend_757_) in
+ match (spc_matches_prefix _stringappend_758_) with
+ | Some (tt,(existT _ _stringappend_759_ _)) =>
+ let _stringappend_760_ := string_drop _stringappend_758_ (build_ex _stringappend_759_) in
+ match (reg_name_matches_prefix _stringappend_760_) with
+ | Some (rd,(existT _ _stringappend_761_ _)) =>
+ let _stringappend_762_ := string_drop _stringappend_760_ (build_ex _stringappend_761_) in
+ sep_matches_prefix _stringappend_762_ >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__0 with
- | Some (_stringappend_1122_,(existT _ _stringappend_1123_ _)) =>
- let _stringappend_1124_ :=
- string_drop _stringappend_1121_ (build_ex _stringappend_1123_) in
- if ((match (hex_bits_20_matches_prefix _stringappend_1124_) with
- | Some (_stringappend_1125_,(existT _ _stringappend_1126_ _)) =>
- match (string_drop _stringappend_1124_
- (build_ex
- _stringappend_1126_)) with
+ | Some (tt,(existT _ _stringappend_763_ _)) =>
+ let _stringappend_764_ :=
+ string_drop _stringappend_762_ (build_ex _stringappend_763_) in
+ if ((match (hex_bits_20_matches_prefix _stringappend_764_) with
+ | Some (imm,(existT _ _stringappend_765_ _)) =>
+ match (string_drop _stringappend_764_
+ (build_ex _stringappend_765_)) with
| "" => true
| _ => false
end
@@ -11675,75 +12829,57 @@ Definition assembly_backwards (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__3 : bool =>
(if (w__3) then
- match (utype_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1113_,(existT _ _stringappend_1114_ _)) =>
- returnm ((_stringappend_1113_, build_ex _stringappend_1114_)
- : (uop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((uop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__5 : (uop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_1114_ _) := w__5 : (uop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1115_ := string_drop _stringappend_1112_ (build_ex _stringappend_1114_) in
- match (spc_matches_prefix _stringappend_1115_) with
- | Some (_stringappend_1116_,(existT _ _stringappend_1117_ _)) =>
- returnm ((_stringappend_1116_, build_ex _stringappend_1117_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__7 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1117_ _) := w__7 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1118_ := string_drop _stringappend_1115_ (build_ex _stringappend_1117_) in
- match (reg_name_matches_prefix _stringappend_1118_) with
- | Some (_stringappend_1119_,(existT _ _stringappend_1120_ _)) =>
- returnm ((_stringappend_1119_, build_ex _stringappend_1120_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__9 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_1120_ _) := w__9 : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1121_ := string_drop _stringappend_1118_ (build_ex _stringappend_1120_) in
- sep_matches_prefix _stringappend_1121_ >>= fun w__10 : option ((unit * {n : Z & ArithFact (n >=
+ (match (utype_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_757_ _)) => returnm (op, build_ex _stringappend_757_)
+ | _ => exit tt : M ((uop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((uop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_757_ _) =>
+ let _stringappend_758_ := string_drop _stringappend_756_ (build_ex _stringappend_757_) in
+ (match (spc_matches_prefix _stringappend_758_) with
+ | Some (tt,(existT _ _stringappend_759_ _)) => returnm (tt, build_ex _stringappend_759_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_759_ _) =>
+ let _stringappend_760_ := string_drop _stringappend_758_ (build_ex _stringappend_759_) in
+ (match (reg_name_matches_prefix _stringappend_760_) with
+ | Some (rd,(existT _ _stringappend_761_ _)) => returnm (rd, build_ex _stringappend_761_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_761_ _) =>
+ let _stringappend_762_ := string_drop _stringappend_760_ (build_ex _stringappend_761_) in
+ sep_matches_prefix _stringappend_762_ >>= fun w__10 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__10 with
- | Some (_stringappend_1122_,(existT _ _stringappend_1123_ _)) =>
- returnm ((_stringappend_1122_, build_ex _stringappend_1123_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__12 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1123_ _) := w__12 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1124_ := string_drop _stringappend_1121_ (build_ex _stringappend_1123_) in
- match (hex_bits_20_matches_prefix _stringappend_1124_) with
- | Some (_stringappend_1125_,(existT _ _stringappend_1126_ _)) =>
- returnm ((_stringappend_1125_, build_ex _stringappend_1126_)
- : (mword 20 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 20 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__14 : (mword 20 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_1126_ _) :=
- w__14
- : (mword 20 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1124_ (build_ex _stringappend_1126_)) with
- | "" => returnm ((UTYPE (imm,rd,op)) : ast)
+ (match w__10 with
+ | Some (tt,(existT _ _stringappend_763_ _)) => returnm (tt, build_ex _stringappend_763_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_763_ _) =>
+ let _stringappend_764_ := string_drop _stringappend_762_ (build_ex _stringappend_763_) in
+ (match (hex_bits_20_matches_prefix _stringappend_764_) with
+ | Some (imm,(existT _ _stringappend_765_ _)) => returnm (imm, build_ex _stringappend_765_)
+ | _ => exit tt : M ((mword 20 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 20 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_765_ _) =>
+ (match (string_drop _stringappend_764_ (build_ex _stringappend_765_)) with
+ | "" => returnm ((UTYPE (imm, rd, op)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
- and_boolM (returnm ((string_startswith _stringappend_1112_ "jal") : bool))
- (let _stringappend_1128_ := string_drop _stringappend_1112_ (string_length "jal") in
- match (spc_matches_prefix _stringappend_1128_) with
- | Some (_stringappend_1129_,(existT _ _stringappend_1130_ _)) =>
- let _stringappend_1131_ := string_drop _stringappend_1128_ (build_ex _stringappend_1130_) in
- match (reg_name_matches_prefix _stringappend_1131_) with
- | Some (_stringappend_1132_,(existT _ _stringappend_1133_ _)) =>
- let _stringappend_1134_ :=
- string_drop _stringappend_1131_ (build_ex _stringappend_1133_) in
- sep_matches_prefix _stringappend_1134_ >>= fun w__17 : option ((unit * {n : Z & ArithFact (n >=
+ and_boolM (returnm ((string_startswith _stringappend_756_ "jal") : bool))
+ (let _stringappend_767_ :=
+ string_drop _stringappend_756_ (build_ex (projT1 (string_length "jal"))) in
+ match (spc_matches_prefix _stringappend_767_) with
+ | Some (tt,(existT _ _stringappend_768_ _)) =>
+ let _stringappend_769_ := string_drop _stringappend_767_ (build_ex _stringappend_768_) in
+ match (reg_name_matches_prefix _stringappend_769_) with
+ | Some (rd,(existT _ _stringappend_770_ _)) =>
+ let _stringappend_771_ := string_drop _stringappend_769_ (build_ex _stringappend_770_) in
+ sep_matches_prefix _stringappend_771_ >>= fun w__17 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__17 with
- | Some (_stringappend_1135_,(existT _ _stringappend_1136_ _)) =>
- let _stringappend_1137_ :=
- string_drop _stringappend_1134_ (build_ex _stringappend_1136_) in
- if ((match (hex_bits_21_matches_prefix _stringappend_1137_) with
- | Some (_stringappend_1138_,(existT _ _stringappend_1139_ _)) =>
- match (string_drop _stringappend_1137_
- (build_ex
- _stringappend_1139_)) with
+ | Some (tt,(existT _ _stringappend_772_ _)) =>
+ let _stringappend_773_ :=
+ string_drop _stringappend_771_ (build_ex _stringappend_772_) in
+ if ((match (hex_bits_21_matches_prefix _stringappend_773_) with
+ | Some (imm,(existT _ _stringappend_774_ _)) =>
+ match (string_drop _stringappend_773_
+ (build_ex _stringappend_774_)) with
| "" => true
| _ => false
end
@@ -11767,84 +12903,66 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__20 : bool =>
(if (w__20) then
- let _stringappend_1128_ := string_drop _stringappend_1112_ (string_length "jal") in
- match (spc_matches_prefix _stringappend_1128_) with
- | Some (_stringappend_1129_,(existT _ _stringappend_1130_ _)) =>
- returnm ((_stringappend_1129_, build_ex _stringappend_1130_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__22 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1130_ _) := w__22 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1131_ := string_drop _stringappend_1128_ (build_ex _stringappend_1130_) in
- match (reg_name_matches_prefix _stringappend_1131_) with
- | Some (_stringappend_1132_,(existT _ _stringappend_1133_ _)) =>
- returnm ((_stringappend_1132_, build_ex _stringappend_1133_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__24 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_1133_ _) :=
- w__24
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1134_ := string_drop _stringappend_1131_ (build_ex _stringappend_1133_) in
- sep_matches_prefix _stringappend_1134_ >>= fun w__25 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_767_ :=
+ string_drop _stringappend_756_ (build_ex (projT1 (string_length "jal"))) in
+ (match (spc_matches_prefix _stringappend_767_) with
+ | Some (tt,(existT _ _stringappend_768_ _)) => returnm (tt, build_ex _stringappend_768_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_768_ _) =>
+ let _stringappend_769_ := string_drop _stringappend_767_ (build_ex _stringappend_768_) in
+ (match (reg_name_matches_prefix _stringappend_769_) with
+ | Some (rd,(existT _ _stringappend_770_ _)) => returnm (rd, build_ex _stringappend_770_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_770_ _) =>
+ let _stringappend_771_ := string_drop _stringappend_769_ (build_ex _stringappend_770_) in
+ sep_matches_prefix _stringappend_771_ >>= fun w__25 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__25 with
- | Some (_stringappend_1135_,(existT _ _stringappend_1136_ _)) =>
- returnm ((_stringappend_1135_, build_ex _stringappend_1136_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__27 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1136_ _) := w__27 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1137_ := string_drop _stringappend_1134_ (build_ex _stringappend_1136_) in
- match (hex_bits_21_matches_prefix _stringappend_1137_) with
- | Some (_stringappend_1138_,(existT _ _stringappend_1139_ _)) =>
- returnm ((_stringappend_1138_, build_ex _stringappend_1139_)
- : (mword 21 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 21 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__29 : (mword 21 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_1139_ _) :=
- w__29
- : (mword 21 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1137_ (build_ex _stringappend_1139_)) with
- | "" => returnm ((RISCV_JAL (imm,rd)) : ast)
+ (match w__25 with
+ | Some (tt,(existT _ _stringappend_772_ _)) => returnm (tt, build_ex _stringappend_772_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_772_ _) =>
+ let _stringappend_773_ := string_drop _stringappend_771_ (build_ex _stringappend_772_) in
+ (match (hex_bits_21_matches_prefix _stringappend_773_) with
+ | Some (imm,(existT _ _stringappend_774_ _)) => returnm (imm, build_ex _stringappend_774_)
+ | _ => exit tt : M ((mword 21 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 21 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_774_ _) =>
+ (match (string_drop _stringappend_773_ (build_ex _stringappend_774_)) with
+ | "" => returnm ((RISCV_JAL (imm, rd)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
- and_boolM (returnm ((string_startswith _stringappend_1112_ "jalr") : bool))
- (let _stringappend_1141_ := string_drop _stringappend_1112_ (string_length "jalr") in
- match (spc_matches_prefix _stringappend_1141_) with
- | Some (_stringappend_1142_,(existT _ _stringappend_1143_ _)) =>
- let _stringappend_1144_ :=
- string_drop _stringappend_1141_ (build_ex _stringappend_1143_) in
- match (reg_name_matches_prefix _stringappend_1144_) with
- | Some (_stringappend_1145_,(existT _ _stringappend_1146_ _)) =>
- let _stringappend_1147_ :=
- string_drop _stringappend_1144_ (build_ex _stringappend_1146_) in
- sep_matches_prefix _stringappend_1147_ >>= fun w__32 : option ((unit * {n : Z & ArithFact (n >=
+ and_boolM (returnm ((string_startswith _stringappend_756_ "jalr") : bool))
+ (let _stringappend_776_ :=
+ string_drop _stringappend_756_ (build_ex (projT1 (string_length "jalr"))) in
+ match (spc_matches_prefix _stringappend_776_) with
+ | Some (tt,(existT _ _stringappend_777_ _)) =>
+ let _stringappend_778_ := string_drop _stringappend_776_ (build_ex _stringappend_777_) in
+ match (reg_name_matches_prefix _stringappend_778_) with
+ | Some (rd,(existT _ _stringappend_779_ _)) =>
+ let _stringappend_780_ :=
+ string_drop _stringappend_778_ (build_ex _stringappend_779_) in
+ sep_matches_prefix _stringappend_780_ >>= fun w__32 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__32 with
- | Some (_stringappend_1148_,(existT _ _stringappend_1149_ _)) =>
- let _stringappend_1150_ :=
- string_drop _stringappend_1147_ (build_ex _stringappend_1149_) in
- match (reg_name_matches_prefix _stringappend_1150_) with
- | Some (_stringappend_1151_,(existT _ _stringappend_1152_ _)) =>
- let _stringappend_1153_ :=
- string_drop _stringappend_1150_ (build_ex _stringappend_1152_) in
- sep_matches_prefix _stringappend_1153_ >>= fun w__33 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_781_ _)) =>
+ let _stringappend_782_ :=
+ string_drop _stringappend_780_ (build_ex _stringappend_781_) in
+ match (reg_name_matches_prefix _stringappend_782_) with
+ | Some (rs1,(existT _ _stringappend_783_ _)) =>
+ let _stringappend_784_ :=
+ string_drop _stringappend_782_ (build_ex _stringappend_783_) in
+ sep_matches_prefix _stringappend_784_ >>= fun w__33 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__33 with
- | Some (_stringappend_1154_,(existT _ _stringappend_1155_ _)) =>
- let _stringappend_1156_ :=
- string_drop _stringappend_1153_
- (build_ex
- _stringappend_1155_) in
- if ((match (hex_bits_12_matches_prefix _stringappend_1156_) with
- | Some
- (_stringappend_1157_,(existT _ _stringappend_1158_ _)) =>
- match (string_drop _stringappend_1156_
- (build_ex
- _stringappend_1158_)) with
+ | Some (tt,(existT _ _stringappend_785_ _)) =>
+ let _stringappend_786_ :=
+ string_drop _stringappend_784_
+ (build_ex _stringappend_785_) in
+ if ((match (hex_bits_12_matches_prefix _stringappend_786_) with
+ | Some (imm,(existT _ _stringappend_787_ _)) =>
+ match (string_drop _stringappend_786_
+ (build_ex _stringappend_787_)) with
| "" => true
| _ => false
end
@@ -11878,119 +12996,87 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__38 : bool =>
(if (w__38) then
- let _stringappend_1141_ := string_drop _stringappend_1112_ (string_length "jalr") in
- match (spc_matches_prefix _stringappend_1141_) with
- | Some (_stringappend_1142_,(existT _ _stringappend_1143_ _)) =>
- returnm ((_stringappend_1142_, build_ex _stringappend_1143_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__40 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1143_ _) :=
- w__40
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1144_ :=
- string_drop _stringappend_1141_ (build_ex _stringappend_1143_) in
- match (reg_name_matches_prefix _stringappend_1144_) with
- | Some (_stringappend_1145_,(existT _ _stringappend_1146_ _)) =>
- returnm ((_stringappend_1145_, build_ex _stringappend_1146_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__42 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_1146_ _) :=
- w__42
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1147_ :=
- string_drop _stringappend_1144_ (build_ex _stringappend_1146_) in
- sep_matches_prefix _stringappend_1147_ >>= fun w__43 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_776_ :=
+ string_drop _stringappend_756_ (build_ex (projT1 (string_length "jalr"))) in
+ (match (spc_matches_prefix _stringappend_776_) with
+ | Some (tt,(existT _ _stringappend_777_ _)) =>
+ returnm (tt, build_ex _stringappend_777_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_777_ _) =>
+ let _stringappend_778_ := string_drop _stringappend_776_ (build_ex _stringappend_777_) in
+ (match (reg_name_matches_prefix _stringappend_778_) with
+ | Some (rd,(existT _ _stringappend_779_ _)) =>
+ returnm (rd, build_ex _stringappend_779_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_779_ _) =>
+ let _stringappend_780_ := string_drop _stringappend_778_ (build_ex _stringappend_779_) in
+ sep_matches_prefix _stringappend_780_ >>= fun w__43 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__43 with
- | Some (_stringappend_1148_,(existT _ _stringappend_1149_ _)) =>
- returnm ((_stringappend_1148_, build_ex _stringappend_1149_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__45 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1149_ _) :=
- w__45
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1150_ :=
- string_drop _stringappend_1147_ (build_ex _stringappend_1149_) in
- match (reg_name_matches_prefix _stringappend_1150_) with
- | Some (_stringappend_1151_,(existT _ _stringappend_1152_ _)) =>
- returnm ((_stringappend_1151_, build_ex _stringappend_1152_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__47 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_1152_ _) :=
- w__47
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1153_ :=
- string_drop _stringappend_1150_ (build_ex _stringappend_1152_) in
- sep_matches_prefix _stringappend_1153_ >>= fun w__48 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__43 with
+ | Some (tt,(existT _ _stringappend_781_ _)) =>
+ returnm (tt, build_ex _stringappend_781_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_781_ _) =>
+ let _stringappend_782_ := string_drop _stringappend_780_ (build_ex _stringappend_781_) in
+ (match (reg_name_matches_prefix _stringappend_782_) with
+ | Some (rs1,(existT _ _stringappend_783_ _)) =>
+ returnm (rs1, build_ex _stringappend_783_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_783_ _) =>
+ let _stringappend_784_ := string_drop _stringappend_782_ (build_ex _stringappend_783_) in
+ sep_matches_prefix _stringappend_784_ >>= fun w__48 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__48 with
- | Some (_stringappend_1154_,(existT _ _stringappend_1155_ _)) =>
- returnm ((_stringappend_1154_, build_ex _stringappend_1155_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__50 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1155_ _) :=
- w__50
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1156_ :=
- string_drop _stringappend_1153_ (build_ex _stringappend_1155_) in
- match (hex_bits_12_matches_prefix _stringappend_1156_) with
- | Some (_stringappend_1157_,(existT _ _stringappend_1158_ _)) =>
- returnm ((_stringappend_1157_, build_ex _stringappend_1158_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__52 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_1158_ _) :=
- w__52
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1156_ (build_ex _stringappend_1158_)) with
- | "" => returnm ((RISCV_JALR (imm,rs1,rd)) : ast)
+ (match w__48 with
+ | Some (tt,(existT _ _stringappend_785_ _)) =>
+ returnm (tt, build_ex _stringappend_785_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_785_ _) =>
+ let _stringappend_786_ := string_drop _stringappend_784_ (build_ex _stringappend_785_) in
+ (match (hex_bits_12_matches_prefix _stringappend_786_) with
+ | Some (imm,(existT _ _stringappend_787_ _)) =>
+ returnm (imm, build_ex _stringappend_787_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_787_ _) =>
+ (match (string_drop _stringappend_786_ (build_ex _stringappend_787_)) with
+ | "" => returnm ((RISCV_JALR (imm, rs1, rd)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
- match (btype_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1160_,(existT _ _stringappend_1161_ _)) =>
- let _stringappend_1162_ :=
- string_drop _stringappend_1112_ (build_ex _stringappend_1161_) in
- match (spc_matches_prefix _stringappend_1162_) with
- | Some (_stringappend_1163_,(existT _ _stringappend_1164_ _)) =>
- let _stringappend_1165_ :=
- string_drop _stringappend_1162_ (build_ex _stringappend_1164_) in
- match (reg_name_matches_prefix _stringappend_1165_) with
- | Some (_stringappend_1166_,(existT _ _stringappend_1167_ _)) =>
- let _stringappend_1168_ :=
- string_drop _stringappend_1165_ (build_ex _stringappend_1167_) in
- sep_matches_prefix _stringappend_1168_ >>= fun w__55 : option ((unit * {n : Z & ArithFact (n >=
+ match (btype_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_789_ _)) =>
+ let _stringappend_790_ :=
+ string_drop _stringappend_756_ (build_ex _stringappend_789_) in
+ match (spc_matches_prefix _stringappend_790_) with
+ | Some (tt,(existT _ _stringappend_791_ _)) =>
+ let _stringappend_792_ :=
+ string_drop _stringappend_790_ (build_ex _stringappend_791_) in
+ match (reg_name_matches_prefix _stringappend_792_) with
+ | Some (rs1,(existT _ _stringappend_793_ _)) =>
+ let _stringappend_794_ :=
+ string_drop _stringappend_792_ (build_ex _stringappend_793_) in
+ sep_matches_prefix _stringappend_794_ >>= fun w__55 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__55 with
- | Some (_stringappend_1169_,(existT _ _stringappend_1170_ _)) =>
- let _stringappend_1171_ :=
- string_drop _stringappend_1168_ (build_ex _stringappend_1170_) in
- match (reg_name_matches_prefix _stringappend_1171_) with
- | Some (_stringappend_1172_,(existT _ _stringappend_1173_ _)) =>
- let _stringappend_1174_ :=
- string_drop _stringappend_1171_ (build_ex _stringappend_1173_) in
- sep_matches_prefix _stringappend_1174_ >>= fun w__56 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_795_ _)) =>
+ let _stringappend_796_ :=
+ string_drop _stringappend_794_ (build_ex _stringappend_795_) in
+ match (reg_name_matches_prefix _stringappend_796_) with
+ | Some (rs2,(existT _ _stringappend_797_ _)) =>
+ let _stringappend_798_ :=
+ string_drop _stringappend_796_ (build_ex _stringappend_797_) in
+ sep_matches_prefix _stringappend_798_ >>= fun w__56 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__56 with
- | Some
- (_stringappend_1175_,(existT _ _stringappend_1176_ _)) =>
- let _stringappend_1177_ :=
- string_drop _stringappend_1174_
- (build_ex
- _stringappend_1176_) in
+ | Some (tt,(existT _ _stringappend_799_ _)) =>
+ let _stringappend_800_ :=
+ string_drop _stringappend_798_
+ (build_ex _stringappend_799_) in
if ((match (hex_bits_13_matches_prefix
- _stringappend_1177_) with
- | Some
- (_stringappend_1178_,(existT _ _stringappend_1179_ _)) =>
- match (string_drop _stringappend_1177_
- (build_ex
- _stringappend_1179_)) with
+ _stringappend_800_) with
+ | Some (imm,(existT _ _stringappend_801_ _)) =>
+ match (string_drop _stringappend_800_
+ (build_ex _stringappend_801_)) with
| "" => true
| _ => false
end
@@ -12026,129 +13112,97 @@ Definition assembly_backwards (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__61 : bool =>
(if (w__61) then
- match (btype_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1160_,(existT _ _stringappend_1161_ _)) =>
- returnm ((_stringappend_1160_, build_ex _stringappend_1161_)
- : (bop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__63 : (bop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_1161_ _) :=
- w__63
- : (bop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1162_ :=
- string_drop _stringappend_1112_ (build_ex _stringappend_1161_) in
- match (spc_matches_prefix _stringappend_1162_) with
- | Some (_stringappend_1163_,(existT _ _stringappend_1164_ _)) =>
- returnm ((_stringappend_1163_, build_ex _stringappend_1164_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__65 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1164_ _) :=
- w__65
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1165_ :=
- string_drop _stringappend_1162_ (build_ex _stringappend_1164_) in
- match (reg_name_matches_prefix _stringappend_1165_) with
- | Some (_stringappend_1166_,(existT _ _stringappend_1167_ _)) =>
- returnm ((_stringappend_1166_, build_ex _stringappend_1167_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__67 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_1167_ _) :=
- w__67
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1168_ :=
- string_drop _stringappend_1165_ (build_ex _stringappend_1167_) in
- sep_matches_prefix _stringappend_1168_ >>= fun w__68 : option ((unit * {n : Z & ArithFact (n >=
+ (match (btype_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_789_ _)) =>
+ returnm (op, build_ex _stringappend_789_)
+ | _ => exit tt : M ((bop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_789_ _) =>
+ let _stringappend_790_ :=
+ string_drop _stringappend_756_ (build_ex _stringappend_789_) in
+ (match (spc_matches_prefix _stringappend_790_) with
+ | Some (tt,(existT _ _stringappend_791_ _)) =>
+ returnm (tt, build_ex _stringappend_791_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_791_ _) =>
+ let _stringappend_792_ :=
+ string_drop _stringappend_790_ (build_ex _stringappend_791_) in
+ (match (reg_name_matches_prefix _stringappend_792_) with
+ | Some (rs1,(existT _ _stringappend_793_ _)) =>
+ returnm (rs1, build_ex _stringappend_793_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_793_ _) =>
+ let _stringappend_794_ :=
+ string_drop _stringappend_792_ (build_ex _stringappend_793_) in
+ sep_matches_prefix _stringappend_794_ >>= fun w__68 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__68 with
- | Some (_stringappend_1169_,(existT _ _stringappend_1170_ _)) =>
- returnm ((_stringappend_1169_, build_ex _stringappend_1170_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__70 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1170_ _) :=
- w__70
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1171_ :=
- string_drop _stringappend_1168_ (build_ex _stringappend_1170_) in
- match (reg_name_matches_prefix _stringappend_1171_) with
- | Some (_stringappend_1172_,(existT _ _stringappend_1173_ _)) =>
- returnm ((_stringappend_1172_, build_ex _stringappend_1173_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__72 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs2, existT _ _stringappend_1173_ _) :=
- w__72
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1174_ :=
- string_drop _stringappend_1171_ (build_ex _stringappend_1173_) in
- sep_matches_prefix _stringappend_1174_ >>= fun w__73 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__68 with
+ | Some (tt,(existT _ _stringappend_795_ _)) =>
+ returnm (tt, build_ex _stringappend_795_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_795_ _) =>
+ let _stringappend_796_ :=
+ string_drop _stringappend_794_ (build_ex _stringappend_795_) in
+ (match (reg_name_matches_prefix _stringappend_796_) with
+ | Some (rs2,(existT _ _stringappend_797_ _)) =>
+ returnm (rs2, build_ex _stringappend_797_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_797_ _) =>
+ let _stringappend_798_ :=
+ string_drop _stringappend_796_ (build_ex _stringappend_797_) in
+ sep_matches_prefix _stringappend_798_ >>= fun w__73 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__73 with
- | Some (_stringappend_1175_,(existT _ _stringappend_1176_ _)) =>
- returnm ((_stringappend_1175_, build_ex _stringappend_1176_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__75 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1176_ _) :=
- w__75
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1177_ :=
- string_drop _stringappend_1174_ (build_ex _stringappend_1176_) in
- match (hex_bits_13_matches_prefix _stringappend_1177_) with
- | Some (_stringappend_1178_,(existT _ _stringappend_1179_ _)) =>
- returnm ((_stringappend_1178_, build_ex _stringappend_1179_)
- : (mword 13 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 13 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__77 : (mword 13 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_1179_ _) :=
- w__77
- : (mword 13 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1177_ (build_ex _stringappend_1179_)) with
- | "" => returnm ((BTYPE (imm,rs2,rs1,op)) : ast)
+ (match w__73 with
+ | Some (tt,(existT _ _stringappend_799_ _)) =>
+ returnm (tt, build_ex _stringappend_799_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_799_ _) =>
+ let _stringappend_800_ :=
+ string_drop _stringappend_798_ (build_ex _stringappend_799_) in
+ (match (hex_bits_13_matches_prefix _stringappend_800_) with
+ | Some (imm,(existT _ _stringappend_801_ _)) =>
+ returnm (imm, build_ex _stringappend_801_)
+ | _ => exit tt : M ((mword 13 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 13 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_801_ _) =>
+ (match (string_drop _stringappend_800_ (build_ex _stringappend_801_)) with
+ | "" => returnm ((BTYPE (imm, rs2, rs1, op)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
- match (itype_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1181_,(existT _ _stringappend_1182_ _)) =>
- let _stringappend_1183_ :=
- string_drop _stringappend_1112_ (build_ex _stringappend_1182_) in
- match (spc_matches_prefix _stringappend_1183_) with
- | Some (_stringappend_1184_,(existT _ _stringappend_1185_ _)) =>
- let _stringappend_1186_ :=
- string_drop _stringappend_1183_ (build_ex _stringappend_1185_) in
- match (reg_name_matches_prefix _stringappend_1186_) with
- | Some (_stringappend_1187_,(existT _ _stringappend_1188_ _)) =>
- let _stringappend_1189_ :=
- string_drop _stringappend_1186_ (build_ex _stringappend_1188_) in
- sep_matches_prefix _stringappend_1189_ >>= fun w__80 : option ((unit * {n : Z & ArithFact (n >=
+ match (itype_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_803_ _)) =>
+ let _stringappend_804_ :=
+ string_drop _stringappend_756_ (build_ex _stringappend_803_) in
+ match (spc_matches_prefix _stringappend_804_) with
+ | Some (tt,(existT _ _stringappend_805_ _)) =>
+ let _stringappend_806_ :=
+ string_drop _stringappend_804_ (build_ex _stringappend_805_) in
+ match (reg_name_matches_prefix _stringappend_806_) with
+ | Some (rd,(existT _ _stringappend_807_ _)) =>
+ let _stringappend_808_ :=
+ string_drop _stringappend_806_ (build_ex _stringappend_807_) in
+ sep_matches_prefix _stringappend_808_ >>= fun w__80 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__80 with
- | Some (_stringappend_1190_,(existT _ _stringappend_1191_ _)) =>
- let _stringappend_1192_ :=
- string_drop _stringappend_1189_ (build_ex _stringappend_1191_) in
- match (reg_name_matches_prefix _stringappend_1192_) with
- | Some (_stringappend_1193_,(existT _ _stringappend_1194_ _)) =>
- let _stringappend_1195_ :=
- string_drop _stringappend_1192_ (build_ex _stringappend_1194_) in
- sep_matches_prefix _stringappend_1195_ >>= fun w__81 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_809_ _)) =>
+ let _stringappend_810_ :=
+ string_drop _stringappend_808_ (build_ex _stringappend_809_) in
+ match (reg_name_matches_prefix _stringappend_810_) with
+ | Some (rs1,(existT _ _stringappend_811_ _)) =>
+ let _stringappend_812_ :=
+ string_drop _stringappend_810_ (build_ex _stringappend_811_) in
+ sep_matches_prefix _stringappend_812_ >>= fun w__81 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__81 with
- | Some
- (_stringappend_1196_,(existT _ _stringappend_1197_ _)) =>
- let _stringappend_1198_ :=
- string_drop _stringappend_1195_
- (build_ex
- _stringappend_1197_) in
+ | Some (tt,(existT _ _stringappend_813_ _)) =>
+ let _stringappend_814_ :=
+ string_drop _stringappend_812_
+ (build_ex _stringappend_813_) in
if ((match (hex_bits_12_matches_prefix
- _stringappend_1198_) with
- | Some
- (_stringappend_1199_,(existT _ _stringappend_1200_ _)) =>
- match (string_drop _stringappend_1198_
- (build_ex
- _stringappend_1200_)) with
+ _stringappend_814_) with
+ | Some (imm,(existT _ _stringappend_815_ _)) =>
+ match (string_drop _stringappend_814_
+ (build_ex _stringappend_815_)) with
| "" => true
| _ => false
end
@@ -12184,126 +13238,93 @@ Definition assembly_backwards (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__86 : bool =>
(if (w__86) then
- match (itype_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1181_,(existT _ _stringappend_1182_ _)) =>
- returnm ((_stringappend_1181_, build_ex _stringappend_1182_)
- : (iop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((iop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__88 : (iop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_1182_ _) :=
- w__88
- : (iop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1183_ :=
- string_drop _stringappend_1112_ (build_ex _stringappend_1182_) in
- match (spc_matches_prefix _stringappend_1183_) with
- | Some (_stringappend_1184_,(existT _ _stringappend_1185_ _)) =>
- returnm ((_stringappend_1184_, build_ex _stringappend_1185_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__90 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1185_ _) :=
- w__90
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1186_ :=
- string_drop _stringappend_1183_ (build_ex _stringappend_1185_) in
- match (reg_name_matches_prefix _stringappend_1186_) with
- | Some (_stringappend_1187_,(existT _ _stringappend_1188_ _)) =>
- returnm ((_stringappend_1187_, build_ex _stringappend_1188_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__92 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_1188_ _) :=
- w__92
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1189_ :=
- string_drop _stringappend_1186_ (build_ex _stringappend_1188_) in
- sep_matches_prefix _stringappend_1189_ >>= fun w__93 : option ((unit * {n : Z & ArithFact (n >=
+ (match (itype_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_803_ _)) =>
+ returnm (op, build_ex _stringappend_803_)
+ | _ => exit tt : M ((iop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((iop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_803_ _) =>
+ let _stringappend_804_ :=
+ string_drop _stringappend_756_ (build_ex _stringappend_803_) in
+ (match (spc_matches_prefix _stringappend_804_) with
+ | Some (tt,(existT _ _stringappend_805_ _)) =>
+ returnm (tt, build_ex _stringappend_805_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_805_ _) =>
+ let _stringappend_806_ :=
+ string_drop _stringappend_804_ (build_ex _stringappend_805_) in
+ (match (reg_name_matches_prefix _stringappend_806_) with
+ | Some (rd,(existT _ _stringappend_807_ _)) =>
+ returnm (rd, build_ex _stringappend_807_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_807_ _) =>
+ let _stringappend_808_ :=
+ string_drop _stringappend_806_ (build_ex _stringappend_807_) in
+ sep_matches_prefix _stringappend_808_ >>= fun w__93 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__93 with
- | Some (_stringappend_1190_,(existT _ _stringappend_1191_ _)) =>
- returnm ((_stringappend_1190_, build_ex _stringappend_1191_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__95 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1191_ _) :=
- w__95
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1192_ :=
- string_drop _stringappend_1189_ (build_ex _stringappend_1191_) in
- match (reg_name_matches_prefix _stringappend_1192_) with
- | Some (_stringappend_1193_,(existT _ _stringappend_1194_ _)) =>
- returnm ((_stringappend_1193_, build_ex _stringappend_1194_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__97 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_1194_ _) :=
- w__97
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1195_ :=
- string_drop _stringappend_1192_ (build_ex _stringappend_1194_) in
- sep_matches_prefix _stringappend_1195_ >>= fun w__98 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__93 with
+ | Some (tt,(existT _ _stringappend_809_ _)) =>
+ returnm (tt, build_ex _stringappend_809_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_809_ _) =>
+ let _stringappend_810_ :=
+ string_drop _stringappend_808_ (build_ex _stringappend_809_) in
+ (match (reg_name_matches_prefix _stringappend_810_) with
+ | Some (rs1,(existT _ _stringappend_811_ _)) =>
+ returnm (rs1, build_ex _stringappend_811_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_811_ _) =>
+ let _stringappend_812_ :=
+ string_drop _stringappend_810_ (build_ex _stringappend_811_) in
+ sep_matches_prefix _stringappend_812_ >>= fun w__98 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__98 with
- | Some (_stringappend_1196_,(existT _ _stringappend_1197_ _)) =>
- returnm ((_stringappend_1196_, build_ex _stringappend_1197_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__100 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1197_ _) :=
- w__100
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1198_ :=
- string_drop _stringappend_1195_ (build_ex _stringappend_1197_) in
- match (hex_bits_12_matches_prefix _stringappend_1198_) with
- | Some (_stringappend_1199_,(existT _ _stringappend_1200_ _)) =>
- returnm ((_stringappend_1199_, build_ex _stringappend_1200_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__102 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_1200_ _) :=
- w__102
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1198_ (build_ex _stringappend_1200_)) with
- | "" => returnm ((ITYPE (imm,rs1,rd,op)) : ast)
+ (match w__98 with
+ | Some (tt,(existT _ _stringappend_813_ _)) =>
+ returnm (tt, build_ex _stringappend_813_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_813_ _) =>
+ let _stringappend_814_ :=
+ string_drop _stringappend_812_ (build_ex _stringappend_813_) in
+ (match (hex_bits_12_matches_prefix _stringappend_814_) with
+ | Some (imm,(existT _ _stringappend_815_ _)) =>
+ returnm (imm, build_ex _stringappend_815_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_815_ _) =>
+ (match (string_drop _stringappend_814_ (build_ex _stringappend_815_)) with
+ | "" => returnm ((ITYPE (imm, rs1, rd, op)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
- match (shiftiop_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1202_,(existT _ _stringappend_1203_ _)) =>
- let _stringappend_1204_ :=
- string_drop _stringappend_1112_ (build_ex _stringappend_1203_) in
- match (spc_matches_prefix _stringappend_1204_) with
- | Some (_stringappend_1205_,(existT _ _stringappend_1206_ _)) =>
- let _stringappend_1207_ :=
- string_drop _stringappend_1204_ (build_ex _stringappend_1206_) in
- match (reg_name_matches_prefix _stringappend_1207_) with
- | Some (_stringappend_1208_,(existT _ _stringappend_1209_ _)) =>
- let _stringappend_1210_ :=
- string_drop _stringappend_1207_ (build_ex _stringappend_1209_) in
- sep_matches_prefix _stringappend_1210_ >>= fun w__105 : option ((unit * {n : Z & ArithFact (n >=
+ match (shiftiop_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_817_ _)) =>
+ let _stringappend_818_ :=
+ string_drop _stringappend_756_ (build_ex _stringappend_817_) in
+ match (spc_matches_prefix _stringappend_818_) with
+ | Some (tt,(existT _ _stringappend_819_ _)) =>
+ let _stringappend_820_ :=
+ string_drop _stringappend_818_ (build_ex _stringappend_819_) in
+ match (reg_name_matches_prefix _stringappend_820_) with
+ | Some (rd,(existT _ _stringappend_821_ _)) =>
+ let _stringappend_822_ :=
+ string_drop _stringappend_820_ (build_ex _stringappend_821_) in
+ sep_matches_prefix _stringappend_822_ >>= fun w__105 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__105 with
- | Some
- (_stringappend_1211_,(existT _ _stringappend_1212_ _)) =>
- let _stringappend_1213_ :=
- string_drop _stringappend_1210_
- (build_ex
- _stringappend_1212_) in
- if ((match (reg_name_matches_prefix _stringappend_1213_) with
- | Some
- (_stringappend_1214_,(existT _ _stringappend_1215_ _)) =>
- let _stringappend_1216_ :=
- string_drop _stringappend_1213_
- (build_ex
- _stringappend_1215_) in
+ | Some (tt,(existT _ _stringappend_823_ _)) =>
+ let _stringappend_824_ :=
+ string_drop _stringappend_822_
+ (build_ex _stringappend_823_) in
+ if ((match (reg_name_matches_prefix _stringappend_824_) with
+ | Some (rs1,(existT _ _stringappend_825_ _)) =>
+ let _stringappend_826_ :=
+ string_drop _stringappend_824_
+ (build_ex _stringappend_825_) in
if ((match (hex_bits_6_matches_prefix
- _stringappend_1216_) with
+ _stringappend_826_) with
| Some
- (_stringappend_1217_,(existT _ _stringappend_1218_ _)) =>
- match (string_drop _stringappend_1216_
- (build_ex
- _stringappend_1218_)) with
+ (shamt,(existT _ _stringappend_827_ _)) =>
+ match (string_drop _stringappend_826_
+ (build_ex _stringappend_827_)) with
| "" => true
| _ => false
end
@@ -12333,116 +13354,89 @@ Definition assembly_backwards (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__108 : bool =>
(if (w__108) then
- match (shiftiop_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1202_,(existT _ _stringappend_1203_ _)) =>
- returnm ((_stringappend_1202_, build_ex _stringappend_1203_)
- : (sop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((sop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__110 : (sop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_1203_ _) :=
- w__110
- : (sop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1204_ :=
- string_drop _stringappend_1112_ (build_ex _stringappend_1203_) in
- match (spc_matches_prefix _stringappend_1204_) with
- | Some (_stringappend_1205_,(existT _ _stringappend_1206_ _)) =>
- returnm ((_stringappend_1205_, build_ex _stringappend_1206_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__112 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1206_ _) :=
- w__112
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1207_ :=
- string_drop _stringappend_1204_ (build_ex _stringappend_1206_) in
- match (reg_name_matches_prefix _stringappend_1207_) with
- | Some (_stringappend_1208_,(existT _ _stringappend_1209_ _)) =>
- returnm ((_stringappend_1208_, build_ex _stringappend_1209_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__114 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_1209_ _) :=
- w__114
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1210_ :=
- string_drop _stringappend_1207_ (build_ex _stringappend_1209_) in
- sep_matches_prefix _stringappend_1210_ >>= fun w__115 : option ((unit * {n : Z & ArithFact (n >=
+ (match (shiftiop_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_817_ _)) =>
+ returnm (op, build_ex _stringappend_817_)
+ | _ => exit tt : M ((sop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((sop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_817_ _) =>
+ let _stringappend_818_ :=
+ string_drop _stringappend_756_ (build_ex _stringappend_817_) in
+ (match (spc_matches_prefix _stringappend_818_) with
+ | Some (tt,(existT _ _stringappend_819_ _)) =>
+ returnm (tt, build_ex _stringappend_819_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_819_ _) =>
+ let _stringappend_820_ :=
+ string_drop _stringappend_818_ (build_ex _stringappend_819_) in
+ (match (reg_name_matches_prefix _stringappend_820_) with
+ | Some (rd,(existT _ _stringappend_821_ _)) =>
+ returnm (rd, build_ex _stringappend_821_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_821_ _) =>
+ let _stringappend_822_ :=
+ string_drop _stringappend_820_ (build_ex _stringappend_821_) in
+ sep_matches_prefix _stringappend_822_ >>= fun w__115 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__115 with
- | Some (_stringappend_1211_,(existT _ _stringappend_1212_ _)) =>
- returnm ((_stringappend_1211_, build_ex _stringappend_1212_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__117 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1212_ _) :=
- w__117
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1213_ :=
- string_drop _stringappend_1210_ (build_ex _stringappend_1212_) in
- match (reg_name_matches_prefix _stringappend_1213_) with
- | Some (_stringappend_1214_,(existT _ _stringappend_1215_ _)) =>
- returnm ((_stringappend_1214_, build_ex _stringappend_1215_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__119 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_1215_ _) :=
- w__119
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1216_ :=
- string_drop _stringappend_1213_ (build_ex _stringappend_1215_) in
- match (hex_bits_6_matches_prefix _stringappend_1216_) with
- | Some (_stringappend_1217_,(existT _ _stringappend_1218_ _)) =>
- returnm ((_stringappend_1217_, build_ex _stringappend_1218_)
- : (mword 6 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 6 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__121 : (mword 6 * {n : Z & ArithFact (n >= 0)}) =>
- let '(shamt, existT _ _stringappend_1218_ _) :=
- w__121
- : (mword 6 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1216_ (build_ex _stringappend_1218_)) with
- | "" => returnm ((SHIFTIOP (shamt,rs1,rd,op)) : ast)
+ (match w__115 with
+ | Some (tt,(existT _ _stringappend_823_ _)) =>
+ returnm (tt, build_ex _stringappend_823_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_823_ _) =>
+ let _stringappend_824_ :=
+ string_drop _stringappend_822_ (build_ex _stringappend_823_) in
+ (match (reg_name_matches_prefix _stringappend_824_) with
+ | Some (rs1,(existT _ _stringappend_825_ _)) =>
+ returnm (rs1, build_ex _stringappend_825_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_825_ _) =>
+ let _stringappend_826_ :=
+ string_drop _stringappend_824_ (build_ex _stringappend_825_) in
+ (match (hex_bits_6_matches_prefix _stringappend_826_) with
+ | Some (shamt,(existT _ _stringappend_827_ _)) =>
+ returnm (shamt, build_ex _stringappend_827_)
+ | _ => exit tt : M ((mword 6 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 6 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(shamt, existT _ _stringappend_827_ _) =>
+ (match (string_drop _stringappend_826_ (build_ex _stringappend_827_)) with
+ | "" => returnm ((SHIFTIOP (shamt, rs1, rd, op)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
- match (rtype_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1220_,(existT _ _stringappend_1221_ _)) =>
- let _stringappend_1222_ :=
- string_drop _stringappend_1112_ (build_ex _stringappend_1221_) in
- match (spc_matches_prefix _stringappend_1222_) with
- | Some (_stringappend_1223_,(existT _ _stringappend_1224_ _)) =>
- let _stringappend_1225_ :=
- string_drop _stringappend_1222_ (build_ex _stringappend_1224_) in
- match (reg_name_matches_prefix _stringappend_1225_) with
- | Some (_stringappend_1226_,(existT _ _stringappend_1227_ _)) =>
- let _stringappend_1228_ :=
- string_drop _stringappend_1225_ (build_ex _stringappend_1227_) in
- sep_matches_prefix _stringappend_1228_ >>= fun w__124 : option ((unit * {n : Z & ArithFact (n >=
+ match (rtype_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_829_ _)) =>
+ let _stringappend_830_ :=
+ string_drop _stringappend_756_ (build_ex _stringappend_829_) in
+ match (spc_matches_prefix _stringappend_830_) with
+ | Some (tt,(existT _ _stringappend_831_ _)) =>
+ let _stringappend_832_ :=
+ string_drop _stringappend_830_ (build_ex _stringappend_831_) in
+ match (reg_name_matches_prefix _stringappend_832_) with
+ | Some (rd,(existT _ _stringappend_833_ _)) =>
+ let _stringappend_834_ :=
+ string_drop _stringappend_832_ (build_ex _stringappend_833_) in
+ sep_matches_prefix _stringappend_834_ >>= fun w__124 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__124 with
- | Some (_stringappend_1229_,(existT _ _stringappend_1230_ _)) =>
- let _stringappend_1231_ :=
- string_drop _stringappend_1228_ (build_ex _stringappend_1230_) in
- match (reg_name_matches_prefix _stringappend_1231_) with
- | Some (_stringappend_1232_,(existT _ _stringappend_1233_ _)) =>
- let _stringappend_1234_ :=
- string_drop _stringappend_1231_ (build_ex _stringappend_1233_) in
- sep_matches_prefix _stringappend_1234_ >>= fun w__125 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_835_ _)) =>
+ let _stringappend_836_ :=
+ string_drop _stringappend_834_ (build_ex _stringappend_835_) in
+ match (reg_name_matches_prefix _stringappend_836_) with
+ | Some (rs1,(existT _ _stringappend_837_ _)) =>
+ let _stringappend_838_ :=
+ string_drop _stringappend_836_ (build_ex _stringappend_837_) in
+ sep_matches_prefix _stringappend_838_ >>= fun w__125 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__125 with
- | Some
- (_stringappend_1235_,(existT _ _stringappend_1236_ _)) =>
- let _stringappend_1237_ :=
- string_drop _stringappend_1234_
- (build_ex
- _stringappend_1236_) in
+ | Some (tt,(existT _ _stringappend_839_ _)) =>
+ let _stringappend_840_ :=
+ string_drop _stringappend_838_
+ (build_ex _stringappend_839_) in
if ((match (reg_name_matches_prefix
- _stringappend_1237_) with
+ _stringappend_840_) with
| Some
- (_stringappend_1238_,(existT _ _stringappend_1239_ _)) =>
- match (string_drop _stringappend_1237_
- (build_ex
- _stringappend_1239_)) with
+ (rs2,(existT _ _stringappend_841_ _)) =>
+ match (string_drop _stringappend_840_
+ (build_ex _stringappend_841_)) with
| "" => true
| _ => false
end
@@ -12478,159 +13472,119 @@ Definition assembly_backwards (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__130 : bool =>
(if (w__130) then
- match (rtype_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1220_,(existT _ _stringappend_1221_ _)) =>
- returnm ((_stringappend_1220_, build_ex _stringappend_1221_)
- : (rop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((rop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__132 : (rop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_1221_ _) :=
- w__132
- : (rop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1222_ :=
- string_drop _stringappend_1112_ (build_ex _stringappend_1221_) in
- match (spc_matches_prefix _stringappend_1222_) with
- | Some (_stringappend_1223_,(existT _ _stringappend_1224_ _)) =>
- returnm ((_stringappend_1223_, build_ex _stringappend_1224_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__134 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1224_ _) :=
- w__134
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1225_ :=
- string_drop _stringappend_1222_ (build_ex _stringappend_1224_) in
- match (reg_name_matches_prefix _stringappend_1225_) with
- | Some (_stringappend_1226_,(existT _ _stringappend_1227_ _)) =>
- returnm ((_stringappend_1226_, build_ex _stringappend_1227_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__136 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_1227_ _) :=
- w__136
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1228_ :=
- string_drop _stringappend_1225_ (build_ex _stringappend_1227_) in
- sep_matches_prefix _stringappend_1228_ >>= fun w__137 : option ((unit * {n : Z & ArithFact (n >=
+ (match (rtype_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_829_ _)) =>
+ returnm (op, build_ex _stringappend_829_)
+ | _ => exit tt : M ((rop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((rop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_829_ _) =>
+ let _stringappend_830_ :=
+ string_drop _stringappend_756_ (build_ex _stringappend_829_) in
+ (match (spc_matches_prefix _stringappend_830_) with
+ | Some (tt,(existT _ _stringappend_831_ _)) =>
+ returnm (tt, build_ex _stringappend_831_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_831_ _) =>
+ let _stringappend_832_ :=
+ string_drop _stringappend_830_ (build_ex _stringappend_831_) in
+ (match (reg_name_matches_prefix _stringappend_832_) with
+ | Some (rd,(existT _ _stringappend_833_ _)) =>
+ returnm (rd, build_ex _stringappend_833_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_833_ _) =>
+ let _stringappend_834_ :=
+ string_drop _stringappend_832_ (build_ex _stringappend_833_) in
+ sep_matches_prefix _stringappend_834_ >>= fun w__137 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__137 with
- | Some (_stringappend_1229_,(existT _ _stringappend_1230_ _)) =>
- returnm ((_stringappend_1229_, build_ex _stringappend_1230_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__139 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1230_ _) :=
- w__139
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1231_ :=
- string_drop _stringappend_1228_ (build_ex _stringappend_1230_) in
- match (reg_name_matches_prefix _stringappend_1231_) with
- | Some (_stringappend_1232_,(existT _ _stringappend_1233_ _)) =>
- returnm ((_stringappend_1232_, build_ex _stringappend_1233_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__141 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_1233_ _) :=
- w__141
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1234_ :=
- string_drop _stringappend_1231_ (build_ex _stringappend_1233_) in
- sep_matches_prefix _stringappend_1234_ >>= fun w__142 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__137 with
+ | Some (tt,(existT _ _stringappend_835_ _)) =>
+ returnm (tt, build_ex _stringappend_835_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_835_ _) =>
+ let _stringappend_836_ :=
+ string_drop _stringappend_834_ (build_ex _stringappend_835_) in
+ (match (reg_name_matches_prefix _stringappend_836_) with
+ | Some (rs1,(existT _ _stringappend_837_ _)) =>
+ returnm (rs1, build_ex _stringappend_837_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_837_ _) =>
+ let _stringappend_838_ :=
+ string_drop _stringappend_836_ (build_ex _stringappend_837_) in
+ sep_matches_prefix _stringappend_838_ >>= fun w__142 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__142 with
- | Some (_stringappend_1235_,(existT _ _stringappend_1236_ _)) =>
- returnm ((_stringappend_1235_, build_ex _stringappend_1236_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__144 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1236_ _) :=
- w__144
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1237_ :=
- string_drop _stringappend_1234_ (build_ex _stringappend_1236_) in
- match (reg_name_matches_prefix _stringappend_1237_) with
- | Some (_stringappend_1238_,(existT _ _stringappend_1239_ _)) =>
- returnm ((_stringappend_1238_, build_ex _stringappend_1239_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__146 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs2, existT _ _stringappend_1239_ _) :=
- w__146
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1237_ (build_ex _stringappend_1239_)) with
- | "" => returnm ((RTYPE (rs2,rs1,rd,op)) : ast)
+ (match w__142 with
+ | Some (tt,(existT _ _stringappend_839_ _)) =>
+ returnm (tt, build_ex _stringappend_839_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_839_ _) =>
+ let _stringappend_840_ :=
+ string_drop _stringappend_838_ (build_ex _stringappend_839_) in
+ (match (reg_name_matches_prefix _stringappend_840_) with
+ | Some (rs2,(existT _ _stringappend_841_ _)) =>
+ returnm (rs2, build_ex _stringappend_841_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_841_ _) =>
+ (match (string_drop _stringappend_840_ (build_ex _stringappend_841_)) with
+ | "" => returnm ((RTYPE (rs2, rs1, rd, op)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
- and_boolM (returnm ((string_startswith _stringappend_1112_ "l") : bool))
- (let _stringappend_1241_ :=
- string_drop _stringappend_1112_ (string_length "l") in
- match (size_mnemonic_matches_prefix _stringappend_1241_) with
- | Some (_stringappend_1242_,(existT _ _stringappend_1243_ _)) =>
- let _stringappend_1244_ :=
- string_drop _stringappend_1241_ (build_ex _stringappend_1243_) in
- match (maybe_u_matches_prefix _stringappend_1244_) with
- | Some (_stringappend_1245_,(existT _ _stringappend_1246_ _)) =>
- let _stringappend_1247_ :=
- string_drop _stringappend_1244_ (build_ex _stringappend_1246_) in
- match (maybe_aq_matches_prefix _stringappend_1247_) with
- | Some (_stringappend_1248_,(existT _ _stringappend_1249_ _)) =>
- let _stringappend_1250_ :=
- string_drop _stringappend_1247_ (build_ex _stringappend_1249_) in
- match (maybe_rl_matches_prefix _stringappend_1250_) with
- | Some (_stringappend_1251_,(existT _ _stringappend_1252_ _)) =>
- let _stringappend_1253_ :=
- string_drop _stringappend_1250_
- (build_ex
- _stringappend_1252_) in
- match (spc_matches_prefix _stringappend_1253_) with
- | Some (_stringappend_1254_,(existT _ _stringappend_1255_ _)) =>
- let _stringappend_1256_ :=
- string_drop _stringappend_1253_
- (build_ex
- _stringappend_1255_) in
- match (reg_name_matches_prefix _stringappend_1256_) with
- | Some
- (_stringappend_1257_,(existT _ _stringappend_1258_ _)) =>
- let _stringappend_1259_ :=
- string_drop _stringappend_1256_
- (build_ex
- _stringappend_1258_) in
- sep_matches_prefix _stringappend_1259_ >>= fun w__149 : option ((unit * {n : Z & ArithFact (n >=
+ and_boolM (returnm ((string_startswith _stringappend_756_ "l") : bool))
+ (let _stringappend_843_ :=
+ string_drop _stringappend_756_ (build_ex (projT1 (string_length "l"))) in
+ match (size_mnemonic_matches_prefix _stringappend_843_) with
+ | Some (size,(existT _ _stringappend_844_ _)) =>
+ let _stringappend_845_ :=
+ string_drop _stringappend_843_ (build_ex _stringappend_844_) in
+ match (maybe_u_matches_prefix _stringappend_845_) with
+ | Some (is_unsigned,(existT _ _stringappend_846_ _)) =>
+ let _stringappend_847_ :=
+ string_drop _stringappend_845_ (build_ex _stringappend_846_) in
+ match (maybe_aq_matches_prefix _stringappend_847_) with
+ | Some (aq,(existT _ _stringappend_848_ _)) =>
+ let _stringappend_849_ :=
+ string_drop _stringappend_847_ (build_ex _stringappend_848_) in
+ match (maybe_rl_matches_prefix _stringappend_849_) with
+ | Some (rl,(existT _ _stringappend_850_ _)) =>
+ let _stringappend_851_ :=
+ string_drop _stringappend_849_ (build_ex _stringappend_850_) in
+ match (spc_matches_prefix _stringappend_851_) with
+ | Some (tt,(existT _ _stringappend_852_ _)) =>
+ let _stringappend_853_ :=
+ string_drop _stringappend_851_
+ (build_ex _stringappend_852_) in
+ match (reg_name_matches_prefix _stringappend_853_) with
+ | Some (rd,(existT _ _stringappend_854_ _)) =>
+ let _stringappend_855_ :=
+ string_drop _stringappend_853_
+ (build_ex _stringappend_854_) in
+ sep_matches_prefix _stringappend_855_ >>= fun w__149 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__149 with
- | Some
- (_stringappend_1260_,(existT _ _stringappend_1261_ _)) =>
- let _stringappend_1262_ :=
- string_drop _stringappend_1259_
- (build_ex
- _stringappend_1261_) in
- match (reg_name_matches_prefix _stringappend_1262_) with
- | Some
- (_stringappend_1263_,(existT _ _stringappend_1264_ _)) =>
- let _stringappend_1265_ :=
- string_drop _stringappend_1262_
- (build_ex
- _stringappend_1264_) in
- sep_matches_prefix _stringappend_1265_ >>= fun w__150 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_856_ _)) =>
+ let _stringappend_857_ :=
+ string_drop _stringappend_855_
+ (build_ex _stringappend_856_) in
+ match (reg_name_matches_prefix _stringappend_857_) with
+ | Some (rs1,(existT _ _stringappend_858_ _)) =>
+ let _stringappend_859_ :=
+ string_drop _stringappend_857_
+ (build_ex _stringappend_858_) in
+ sep_matches_prefix _stringappend_859_ >>= fun w__150 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__150 with
| Some
- (_stringappend_1266_,(existT _ _stringappend_1267_ _)) =>
- let _stringappend_1268_ :=
- string_drop
- _stringappend_1265_
- (build_ex
- _stringappend_1267_) in
+ (tt,(existT _ _stringappend_860_ _)) =>
+ let _stringappend_861_ :=
+ string_drop _stringappend_859_
+ (build_ex _stringappend_860_) in
if ((match (hex_bits_12_matches_prefix
- _stringappend_1268_) with
+ _stringappend_861_) with
| Some
- (_stringappend_1269_,(existT _ _stringappend_1270_ _)) =>
+ (imm,(existT _ _stringappend_862_ _)) =>
match (string_drop
- _stringappend_1268_
- (build_ex
- _stringappend_1270_)) with
+ _stringappend_861_
+ (build_ex _stringappend_862_)) with
| "" => true
| _ => false
end
@@ -12684,190 +13638,140 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__159 : bool =>
(if (w__159) then
- let _stringappend_1241_ :=
- string_drop _stringappend_1112_ (string_length "l") in
- match (size_mnemonic_matches_prefix _stringappend_1241_) with
- | Some (_stringappend_1242_,(existT _ _stringappend_1243_ _)) =>
- returnm ((_stringappend_1242_, build_ex _stringappend_1243_)
- : (word_width * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((word_width * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__161 : (word_width * {n : Z & ArithFact (n >= 0)}) =>
- let '(size, existT _ _stringappend_1243_ _) :=
- w__161
- : (word_width * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1244_ :=
- string_drop _stringappend_1241_ (build_ex _stringappend_1243_) in
- match (maybe_u_matches_prefix _stringappend_1244_) with
- | Some (_stringappend_1245_,(existT _ _stringappend_1246_ _)) =>
- returnm ((_stringappend_1245_, build_ex _stringappend_1246_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__163 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(is_unsigned, existT _ _stringappend_1246_ _) :=
- w__163
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1247_ :=
- string_drop _stringappend_1244_ (build_ex _stringappend_1246_) in
- match (maybe_aq_matches_prefix _stringappend_1247_) with
- | Some (_stringappend_1248_,(existT _ _stringappend_1249_ _)) =>
- returnm ((_stringappend_1248_, build_ex _stringappend_1249_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__165 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(aq, existT _ _stringappend_1249_ _) :=
- w__165
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1250_ :=
- string_drop _stringappend_1247_ (build_ex _stringappend_1249_) in
- match (maybe_rl_matches_prefix _stringappend_1250_) with
- | Some (_stringappend_1251_,(existT _ _stringappend_1252_ _)) =>
- returnm ((_stringappend_1251_, build_ex _stringappend_1252_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__167 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(rl, existT _ _stringappend_1252_ _) :=
- w__167
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1253_ :=
- string_drop _stringappend_1250_ (build_ex _stringappend_1252_) in
- match (spc_matches_prefix _stringappend_1253_) with
- | Some (_stringappend_1254_,(existT _ _stringappend_1255_ _)) =>
- returnm ((_stringappend_1254_, build_ex _stringappend_1255_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__169 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1255_ _) :=
- w__169
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1256_ :=
- string_drop _stringappend_1253_ (build_ex _stringappend_1255_) in
- match (reg_name_matches_prefix _stringappend_1256_) with
- | Some (_stringappend_1257_,(existT _ _stringappend_1258_ _)) =>
- returnm ((_stringappend_1257_, build_ex _stringappend_1258_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__171 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_1258_ _) :=
- w__171
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1259_ :=
- string_drop _stringappend_1256_ (build_ex _stringappend_1258_) in
- sep_matches_prefix _stringappend_1259_ >>= fun w__172 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_843_ :=
+ string_drop _stringappend_756_ (build_ex (projT1 (string_length "l"))) in
+ (match (size_mnemonic_matches_prefix _stringappend_843_) with
+ | Some (size,(existT _ _stringappend_844_ _)) =>
+ returnm (size, build_ex _stringappend_844_)
+ | _ => exit tt : M ((word_width * {n : Z & ArithFact (n >= 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >= 0)}))) >>= fun '(size, existT _ _stringappend_844_ _) =>
+ let _stringappend_845_ :=
+ string_drop _stringappend_843_ (build_ex _stringappend_844_) in
+ (match (maybe_u_matches_prefix _stringappend_845_) with
+ | Some (is_unsigned,(existT _ _stringappend_846_ _)) =>
+ returnm (is_unsigned, build_ex _stringappend_846_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(is_unsigned, existT _ _stringappend_846_ _) =>
+ let _stringappend_847_ :=
+ string_drop _stringappend_845_ (build_ex _stringappend_846_) in
+ (match (maybe_aq_matches_prefix _stringappend_847_) with
+ | Some (aq,(existT _ _stringappend_848_ _)) =>
+ returnm (aq, build_ex _stringappend_848_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(aq, existT _ _stringappend_848_ _) =>
+ let _stringappend_849_ :=
+ string_drop _stringappend_847_ (build_ex _stringappend_848_) in
+ (match (maybe_rl_matches_prefix _stringappend_849_) with
+ | Some (rl,(existT _ _stringappend_850_ _)) =>
+ returnm (rl, build_ex _stringappend_850_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rl, existT _ _stringappend_850_ _) =>
+ let _stringappend_851_ :=
+ string_drop _stringappend_849_ (build_ex _stringappend_850_) in
+ (match (spc_matches_prefix _stringappend_851_) with
+ | Some (tt,(existT _ _stringappend_852_ _)) =>
+ returnm (tt, build_ex _stringappend_852_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_852_ _) =>
+ let _stringappend_853_ :=
+ string_drop _stringappend_851_ (build_ex _stringappend_852_) in
+ (match (reg_name_matches_prefix _stringappend_853_) with
+ | Some (rd,(existT _ _stringappend_854_ _)) =>
+ returnm (rd, build_ex _stringappend_854_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_854_ _) =>
+ let _stringappend_855_ :=
+ string_drop _stringappend_853_ (build_ex _stringappend_854_) in
+ sep_matches_prefix _stringappend_855_ >>= fun w__172 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__172 with
- | Some (_stringappend_1260_,(existT _ _stringappend_1261_ _)) =>
- returnm ((_stringappend_1260_, build_ex _stringappend_1261_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__174 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1261_ _) :=
- w__174
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1262_ :=
- string_drop _stringappend_1259_ (build_ex _stringappend_1261_) in
- match (reg_name_matches_prefix _stringappend_1262_) with
- | Some (_stringappend_1263_,(existT _ _stringappend_1264_ _)) =>
- returnm ((_stringappend_1263_, build_ex _stringappend_1264_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__176 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_1264_ _) :=
- w__176
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1265_ :=
- string_drop _stringappend_1262_ (build_ex _stringappend_1264_) in
- sep_matches_prefix _stringappend_1265_ >>= fun w__177 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__172 with
+ | Some (tt,(existT _ _stringappend_856_ _)) =>
+ returnm (tt, build_ex _stringappend_856_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_856_ _) =>
+ let _stringappend_857_ :=
+ string_drop _stringappend_855_ (build_ex _stringappend_856_) in
+ (match (reg_name_matches_prefix _stringappend_857_) with
+ | Some (rs1,(existT _ _stringappend_858_ _)) =>
+ returnm (rs1, build_ex _stringappend_858_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_858_ _) =>
+ let _stringappend_859_ :=
+ string_drop _stringappend_857_ (build_ex _stringappend_858_) in
+ sep_matches_prefix _stringappend_859_ >>= fun w__177 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__177 with
- | Some (_stringappend_1266_,(existT _ _stringappend_1267_ _)) =>
- returnm ((_stringappend_1266_, build_ex _stringappend_1267_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__179 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1267_ _) :=
- w__179
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1268_ :=
- string_drop _stringappend_1265_ (build_ex _stringappend_1267_) in
- match (hex_bits_12_matches_prefix _stringappend_1268_) with
- | Some (_stringappend_1269_,(existT _ _stringappend_1270_ _)) =>
- returnm ((_stringappend_1269_, build_ex _stringappend_1270_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__181 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_1270_ _) :=
- w__181
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1268_ (build_ex _stringappend_1270_)) with
- | "" => returnm ((LOAD (imm,rs1,rd,is_unsigned,size,aq,rl)) : ast)
+ (match w__177 with
+ | Some (tt,(existT _ _stringappend_860_ _)) =>
+ returnm (tt, build_ex _stringappend_860_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_860_ _) =>
+ let _stringappend_861_ :=
+ string_drop _stringappend_859_ (build_ex _stringappend_860_) in
+ (match (hex_bits_12_matches_prefix _stringappend_861_) with
+ | Some (imm,(existT _ _stringappend_862_ _)) =>
+ returnm (imm, build_ex _stringappend_862_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_862_ _) =>
+ (match (string_drop _stringappend_861_ (build_ex _stringappend_862_)) with
+ | "" =>
+ returnm ((LOAD (imm, rs1, rd, is_unsigned, size, aq, rl)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
- and_boolM (returnm ((string_startswith _stringappend_1112_ "s") : bool))
- (let _stringappend_1272_ :=
- string_drop _stringappend_1112_ (string_length "s") in
- match (size_mnemonic_matches_prefix _stringappend_1272_) with
- | Some (_stringappend_1273_,(existT _ _stringappend_1274_ _)) =>
- let _stringappend_1275_ :=
- string_drop _stringappend_1272_ (build_ex _stringappend_1274_) in
- match (maybe_aq_matches_prefix _stringappend_1275_) with
- | Some (_stringappend_1276_,(existT _ _stringappend_1277_ _)) =>
- let _stringappend_1278_ :=
- string_drop _stringappend_1275_ (build_ex _stringappend_1277_) in
- match (maybe_rl_matches_prefix _stringappend_1278_) with
- | Some (_stringappend_1279_,(existT _ _stringappend_1280_ _)) =>
- let _stringappend_1281_ :=
- string_drop _stringappend_1278_
- (build_ex
- _stringappend_1280_) in
- match (spc_matches_prefix _stringappend_1281_) with
- | Some (_stringappend_1282_,(existT _ _stringappend_1283_ _)) =>
- let _stringappend_1284_ :=
- string_drop _stringappend_1281_
- (build_ex
- _stringappend_1283_) in
- match (reg_name_matches_prefix _stringappend_1284_) with
- | Some
- (_stringappend_1285_,(existT _ _stringappend_1286_ _)) =>
- let _stringappend_1287_ :=
- string_drop _stringappend_1284_
- (build_ex
- _stringappend_1286_) in
- sep_matches_prefix _stringappend_1287_ >>= fun w__184 : option ((unit * {n : Z & ArithFact (n >=
+ and_boolM (returnm ((string_startswith _stringappend_756_ "s") : bool))
+ (let _stringappend_864_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length "s"))) in
+ match (size_mnemonic_matches_prefix _stringappend_864_) with
+ | Some (size,(existT _ _stringappend_865_ _)) =>
+ let _stringappend_866_ :=
+ string_drop _stringappend_864_ (build_ex _stringappend_865_) in
+ match (maybe_aq_matches_prefix _stringappend_866_) with
+ | Some (aq,(existT _ _stringappend_867_ _)) =>
+ let _stringappend_868_ :=
+ string_drop _stringappend_866_ (build_ex _stringappend_867_) in
+ match (maybe_rl_matches_prefix _stringappend_868_) with
+ | Some (rl,(existT _ _stringappend_869_ _)) =>
+ let _stringappend_870_ :=
+ string_drop _stringappend_868_ (build_ex _stringappend_869_) in
+ match (spc_matches_prefix _stringappend_870_) with
+ | Some (tt,(existT _ _stringappend_871_ _)) =>
+ let _stringappend_872_ :=
+ string_drop _stringappend_870_
+ (build_ex _stringappend_871_) in
+ match (reg_name_matches_prefix _stringappend_872_) with
+ | Some (rd,(existT _ _stringappend_873_ _)) =>
+ let _stringappend_874_ :=
+ string_drop _stringappend_872_
+ (build_ex _stringappend_873_) in
+ sep_matches_prefix _stringappend_874_ >>= fun w__184 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__184 with
- | Some
- (_stringappend_1288_,(existT _ _stringappend_1289_ _)) =>
- let _stringappend_1290_ :=
- string_drop _stringappend_1287_
- (build_ex
- _stringappend_1289_) in
- match (reg_name_matches_prefix _stringappend_1290_) with
- | Some
- (_stringappend_1291_,(existT _ _stringappend_1292_ _)) =>
- let _stringappend_1293_ :=
- string_drop _stringappend_1290_
- (build_ex
- _stringappend_1292_) in
- sep_matches_prefix _stringappend_1293_ >>= fun w__185 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_875_ _)) =>
+ let _stringappend_876_ :=
+ string_drop _stringappend_874_
+ (build_ex _stringappend_875_) in
+ match (reg_name_matches_prefix _stringappend_876_) with
+ | Some (rs1,(existT _ _stringappend_877_ _)) =>
+ let _stringappend_878_ :=
+ string_drop _stringappend_876_
+ (build_ex _stringappend_877_) in
+ sep_matches_prefix _stringappend_878_ >>= fun w__185 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__185 with
| Some
- (_stringappend_1294_,(existT _ _stringappend_1295_ _)) =>
- let _stringappend_1296_ :=
- string_drop
- _stringappend_1293_
- (build_ex
- _stringappend_1295_) in
+ (tt,(existT _ _stringappend_879_ _)) =>
+ let _stringappend_880_ :=
+ string_drop _stringappend_878_
+ (build_ex _stringappend_879_) in
if ((match (hex_bits_12_matches_prefix
- _stringappend_1296_) with
+ _stringappend_880_) with
| Some
- (_stringappend_1297_,(existT _ _stringappend_1298_ _)) =>
+ (imm,(existT _ _stringappend_881_ _)) =>
match (string_drop
- _stringappend_1296_
- (build_ex
- _stringappend_1298_)) with
+ _stringappend_880_
+ (build_ex _stringappend_881_)) with
| "" => true
| _ => false
end
@@ -12916,162 +13820,121 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__193 : bool =>
(if (w__193) then
- let _stringappend_1272_ :=
- string_drop _stringappend_1112_ (string_length "s") in
- match (size_mnemonic_matches_prefix _stringappend_1272_) with
- | Some (_stringappend_1273_,(existT _ _stringappend_1274_ _)) =>
- returnm ((_stringappend_1273_, build_ex _stringappend_1274_)
- : (word_width * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((word_width * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__195 : (word_width * {n : Z & ArithFact (n >= 0)}) =>
- let '(size, existT _ _stringappend_1274_ _) :=
- w__195
- : (word_width * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1275_ :=
- string_drop _stringappend_1272_ (build_ex _stringappend_1274_) in
- match (maybe_aq_matches_prefix _stringappend_1275_) with
- | Some (_stringappend_1276_,(existT _ _stringappend_1277_ _)) =>
- returnm ((_stringappend_1276_, build_ex _stringappend_1277_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__197 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(aq, existT _ _stringappend_1277_ _) :=
- w__197
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1278_ :=
- string_drop _stringappend_1275_ (build_ex _stringappend_1277_) in
- match (maybe_rl_matches_prefix _stringappend_1278_) with
- | Some (_stringappend_1279_,(existT _ _stringappend_1280_ _)) =>
- returnm ((_stringappend_1279_, build_ex _stringappend_1280_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__199 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(rl, existT _ _stringappend_1280_ _) :=
- w__199
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1281_ :=
- string_drop _stringappend_1278_ (build_ex _stringappend_1280_) in
- match (spc_matches_prefix _stringappend_1281_) with
- | Some (_stringappend_1282_,(existT _ _stringappend_1283_ _)) =>
- returnm ((_stringappend_1282_, build_ex _stringappend_1283_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__201 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1283_ _) :=
- w__201
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1284_ :=
- string_drop _stringappend_1281_ (build_ex _stringappend_1283_) in
- match (reg_name_matches_prefix _stringappend_1284_) with
- | Some (_stringappend_1285_,(existT _ _stringappend_1286_ _)) =>
- returnm ((_stringappend_1285_, build_ex _stringappend_1286_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__203 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_1286_ _) :=
- w__203
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1287_ :=
- string_drop _stringappend_1284_ (build_ex _stringappend_1286_) in
- sep_matches_prefix _stringappend_1287_ >>= fun w__204 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_864_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length "s"))) in
+ (match (size_mnemonic_matches_prefix _stringappend_864_) with
+ | Some (size,(existT _ _stringappend_865_ _)) =>
+ returnm (size, build_ex _stringappend_865_)
+ | _ => exit tt : M ((word_width * {n : Z & ArithFact (n >= 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >= 0)}))) >>= fun '(size, existT _ _stringappend_865_ _) =>
+ let _stringappend_866_ :=
+ string_drop _stringappend_864_ (build_ex _stringappend_865_) in
+ (match (maybe_aq_matches_prefix _stringappend_866_) with
+ | Some (aq,(existT _ _stringappend_867_ _)) =>
+ returnm (aq, build_ex _stringappend_867_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(aq, existT _ _stringappend_867_ _) =>
+ let _stringappend_868_ :=
+ string_drop _stringappend_866_ (build_ex _stringappend_867_) in
+ (match (maybe_rl_matches_prefix _stringappend_868_) with
+ | Some (rl,(existT _ _stringappend_869_ _)) =>
+ returnm (rl, build_ex _stringappend_869_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rl, existT _ _stringappend_869_ _) =>
+ let _stringappend_870_ :=
+ string_drop _stringappend_868_ (build_ex _stringappend_869_) in
+ (match (spc_matches_prefix _stringappend_870_) with
+ | Some (tt,(existT _ _stringappend_871_ _)) =>
+ returnm (tt, build_ex _stringappend_871_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_871_ _) =>
+ let _stringappend_872_ :=
+ string_drop _stringappend_870_ (build_ex _stringappend_871_) in
+ (match (reg_name_matches_prefix _stringappend_872_) with
+ | Some (rd,(existT _ _stringappend_873_ _)) =>
+ returnm (rd, build_ex _stringappend_873_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_873_ _) =>
+ let _stringappend_874_ :=
+ string_drop _stringappend_872_ (build_ex _stringappend_873_) in
+ sep_matches_prefix _stringappend_874_ >>= fun w__204 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__204 with
- | Some (_stringappend_1288_,(existT _ _stringappend_1289_ _)) =>
- returnm ((_stringappend_1288_, build_ex _stringappend_1289_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__206 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1289_ _) :=
- w__206
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1290_ :=
- string_drop _stringappend_1287_ (build_ex _stringappend_1289_) in
- match (reg_name_matches_prefix _stringappend_1290_) with
- | Some (_stringappend_1291_,(existT _ _stringappend_1292_ _)) =>
- returnm ((_stringappend_1291_, build_ex _stringappend_1292_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__208 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_1292_ _) :=
- w__208
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1293_ :=
- string_drop _stringappend_1290_ (build_ex _stringappend_1292_) in
- sep_matches_prefix _stringappend_1293_ >>= fun w__209 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__204 with
+ | Some (tt,(existT _ _stringappend_875_ _)) =>
+ returnm (tt, build_ex _stringappend_875_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_875_ _) =>
+ let _stringappend_876_ :=
+ string_drop _stringappend_874_ (build_ex _stringappend_875_) in
+ (match (reg_name_matches_prefix _stringappend_876_) with
+ | Some (rs1,(existT _ _stringappend_877_ _)) =>
+ returnm (rs1, build_ex _stringappend_877_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_877_ _) =>
+ let _stringappend_878_ :=
+ string_drop _stringappend_876_ (build_ex _stringappend_877_) in
+ sep_matches_prefix _stringappend_878_ >>= fun w__209 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__209 with
- | Some (_stringappend_1294_,(existT _ _stringappend_1295_ _)) =>
- returnm ((_stringappend_1294_, build_ex _stringappend_1295_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__211 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1295_ _) :=
- w__211
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1296_ :=
- string_drop _stringappend_1293_ (build_ex _stringappend_1295_) in
- match (hex_bits_12_matches_prefix _stringappend_1296_) with
- | Some (_stringappend_1297_,(existT _ _stringappend_1298_ _)) =>
- returnm ((_stringappend_1297_, build_ex _stringappend_1298_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__213 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_1298_ _) :=
- w__213
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1296_ (build_ex _stringappend_1298_)) with
- | "" => returnm ((STORE (imm,rs1,rd,size,aq,rl)) : ast)
+ (match w__209 with
+ | Some (tt,(existT _ _stringappend_879_ _)) =>
+ returnm (tt, build_ex _stringappend_879_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_879_ _) =>
+ let _stringappend_880_ :=
+ string_drop _stringappend_878_ (build_ex _stringappend_879_) in
+ (match (hex_bits_12_matches_prefix _stringappend_880_) with
+ | Some (imm,(existT _ _stringappend_881_ _)) =>
+ returnm (imm, build_ex _stringappend_881_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_881_ _) =>
+ (match (string_drop _stringappend_880_ (build_ex _stringappend_881_)) with
+ | "" => returnm ((STORE (imm, rs1, rd, size, aq, rl)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
and_boolM
- (returnm ((string_startswith _stringappend_1112_ "addiw")
+ (returnm ((string_startswith _stringappend_756_ "addiw")
: bool))
- (let _stringappend_1300_ :=
- string_drop _stringappend_1112_ (string_length "addiw") in
- match (spc_matches_prefix _stringappend_1300_) with
- | Some (_stringappend_1301_,(existT _ _stringappend_1302_ _)) =>
- let _stringappend_1303_ :=
- string_drop _stringappend_1300_ (build_ex _stringappend_1302_) in
- match (reg_name_matches_prefix _stringappend_1303_) with
- | Some (_stringappend_1304_,(existT _ _stringappend_1305_ _)) =>
- let _stringappend_1306_ :=
- string_drop _stringappend_1303_
- (build_ex
- _stringappend_1305_) in
- sep_matches_prefix _stringappend_1306_ >>= fun w__216 : option ((unit * {n : Z & ArithFact (n >=
+ (let _stringappend_883_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length "addiw"))) in
+ match (spc_matches_prefix _stringappend_883_) with
+ | Some (tt,(existT _ _stringappend_884_ _)) =>
+ let _stringappend_885_ :=
+ string_drop _stringappend_883_ (build_ex _stringappend_884_) in
+ match (reg_name_matches_prefix _stringappend_885_) with
+ | Some (rd,(existT _ _stringappend_886_ _)) =>
+ let _stringappend_887_ :=
+ string_drop _stringappend_885_ (build_ex _stringappend_886_) in
+ sep_matches_prefix _stringappend_887_ >>= fun w__216 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__216 with
- | Some (_stringappend_1307_,(existT _ _stringappend_1308_ _)) =>
- let _stringappend_1309_ :=
- string_drop _stringappend_1306_
- (build_ex
- _stringappend_1308_) in
- match (reg_name_matches_prefix _stringappend_1309_) with
- | Some
- (_stringappend_1310_,(existT _ _stringappend_1311_ _)) =>
- let _stringappend_1312_ :=
- string_drop _stringappend_1309_
- (build_ex
- _stringappend_1311_) in
- sep_matches_prefix _stringappend_1312_ >>= fun w__217 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_888_ _)) =>
+ let _stringappend_889_ :=
+ string_drop _stringappend_887_
+ (build_ex _stringappend_888_) in
+ match (reg_name_matches_prefix _stringappend_889_) with
+ | Some (rs1,(existT _ _stringappend_890_ _)) =>
+ let _stringappend_891_ :=
+ string_drop _stringappend_889_
+ (build_ex _stringappend_890_) in
+ sep_matches_prefix _stringappend_891_ >>= fun w__217 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__217 with
| Some
- (_stringappend_1313_,(existT _ _stringappend_1314_ _)) =>
- let _stringappend_1315_ :=
- string_drop _stringappend_1312_
- (build_ex
- _stringappend_1314_) in
+ (tt,(existT _ _stringappend_892_ _)) =>
+ let _stringappend_893_ :=
+ string_drop _stringappend_891_
+ (build_ex _stringappend_892_) in
if ((match (hex_bits_12_matches_prefix
- _stringappend_1315_) with
+ _stringappend_893_) with
| Some
- (_stringappend_1316_,(existT _ _stringappend_1317_ _)) =>
+ (imm,(existT _ _stringappend_894_ _)) =>
match (string_drop
- _stringappend_1315_
- (build_ex
- _stringappend_1317_)) with
+ _stringappend_893_
+ (build_ex _stringappend_894_)) with
| "" => true
| _ => false
end
@@ -13105,133 +13968,101 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__222 : bool =>
(if (w__222) then
- let _stringappend_1300_ :=
- string_drop _stringappend_1112_ (string_length "addiw") in
- match (spc_matches_prefix _stringappend_1300_) with
- | Some (_stringappend_1301_,(existT _ _stringappend_1302_ _)) =>
- returnm ((_stringappend_1301_, build_ex _stringappend_1302_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__224 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1302_ _) :=
- w__224
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1303_ :=
- string_drop _stringappend_1300_ (build_ex _stringappend_1302_) in
- match (reg_name_matches_prefix _stringappend_1303_) with
- | Some (_stringappend_1304_,(existT _ _stringappend_1305_ _)) =>
- returnm ((_stringappend_1304_, build_ex _stringappend_1305_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__226 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_1305_ _) :=
- w__226
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1306_ :=
- string_drop _stringappend_1303_ (build_ex _stringappend_1305_) in
- sep_matches_prefix _stringappend_1306_ >>= fun w__227 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_883_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length "addiw"))) in
+ (match (spc_matches_prefix _stringappend_883_) with
+ | Some (tt,(existT _ _stringappend_884_ _)) =>
+ returnm (tt, build_ex _stringappend_884_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_884_ _) =>
+ let _stringappend_885_ :=
+ string_drop _stringappend_883_ (build_ex _stringappend_884_) in
+ (match (reg_name_matches_prefix _stringappend_885_) with
+ | Some (rd,(existT _ _stringappend_886_ _)) =>
+ returnm (rd, build_ex _stringappend_886_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_886_ _) =>
+ let _stringappend_887_ :=
+ string_drop _stringappend_885_ (build_ex _stringappend_886_) in
+ sep_matches_prefix _stringappend_887_ >>= fun w__227 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__227 with
- | Some (_stringappend_1307_,(existT _ _stringappend_1308_ _)) =>
- returnm ((_stringappend_1307_, build_ex _stringappend_1308_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__229 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1308_ _) :=
- w__229
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1309_ :=
- string_drop _stringappend_1306_ (build_ex _stringappend_1308_) in
- match (reg_name_matches_prefix _stringappend_1309_) with
- | Some (_stringappend_1310_,(existT _ _stringappend_1311_ _)) =>
- returnm ((_stringappend_1310_, build_ex _stringappend_1311_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__231 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_1311_ _) :=
- w__231
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1312_ :=
- string_drop _stringappend_1309_ (build_ex _stringappend_1311_) in
- sep_matches_prefix _stringappend_1312_ >>= fun w__232 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__227 with
+ | Some (tt,(existT _ _stringappend_888_ _)) =>
+ returnm (tt, build_ex _stringappend_888_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_888_ _) =>
+ let _stringappend_889_ :=
+ string_drop _stringappend_887_ (build_ex _stringappend_888_) in
+ (match (reg_name_matches_prefix _stringappend_889_) with
+ | Some (rs1,(existT _ _stringappend_890_ _)) =>
+ returnm (rs1, build_ex _stringappend_890_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_890_ _) =>
+ let _stringappend_891_ :=
+ string_drop _stringappend_889_ (build_ex _stringappend_890_) in
+ sep_matches_prefix _stringappend_891_ >>= fun w__232 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__232 with
- | Some (_stringappend_1313_,(existT _ _stringappend_1314_ _)) =>
- returnm ((_stringappend_1313_, build_ex _stringappend_1314_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__234 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1314_ _) :=
- w__234
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1315_ :=
- string_drop _stringappend_1312_ (build_ex _stringappend_1314_) in
- match (hex_bits_12_matches_prefix _stringappend_1315_) with
- | Some (_stringappend_1316_,(existT _ _stringappend_1317_ _)) =>
- returnm ((_stringappend_1316_, build_ex _stringappend_1317_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__236 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_1317_ _) :=
- w__236
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1315_
- (build_ex
- _stringappend_1317_)) with
- | "" => returnm ((ADDIW (imm,rs1,rd)) : ast)
+ (match w__232 with
+ | Some (tt,(existT _ _stringappend_892_ _)) =>
+ returnm (tt, build_ex _stringappend_892_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_892_ _) =>
+ let _stringappend_893_ :=
+ string_drop _stringappend_891_ (build_ex _stringappend_892_) in
+ (match (hex_bits_12_matches_prefix _stringappend_893_) with
+ | Some (imm,(existT _ _stringappend_894_ _)) =>
+ returnm (imm, build_ex _stringappend_894_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_894_ _) =>
+ (match (string_drop _stringappend_893_
+ (build_ex _stringappend_894_)) with
+ | "" => returnm ((ADDIW (imm, rs1, rd)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
- match (shiftw_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1319_,(existT _ _stringappend_1320_ _)) =>
- let _stringappend_1321_ :=
- string_drop _stringappend_1112_ (build_ex _stringappend_1320_) in
- match (spc_matches_prefix _stringappend_1321_) with
- | Some (_stringappend_1322_,(existT _ _stringappend_1323_ _)) =>
- let _stringappend_1324_ :=
- string_drop _stringappend_1321_
- (build_ex
- _stringappend_1323_) in
- match (reg_name_matches_prefix _stringappend_1324_) with
- | Some (_stringappend_1325_,(existT _ _stringappend_1326_ _)) =>
- let _stringappend_1327_ :=
- string_drop _stringappend_1324_
- (build_ex
- _stringappend_1326_) in
- sep_matches_prefix _stringappend_1327_ >>= fun w__239 : option ((unit * {n : Z & ArithFact (n >=
+ match (shiftw_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_896_ _)) =>
+ let _stringappend_897_ :=
+ string_drop _stringappend_756_ (build_ex _stringappend_896_) in
+ match (spc_matches_prefix _stringappend_897_) with
+ | Some (tt,(existT _ _stringappend_898_ _)) =>
+ let _stringappend_899_ :=
+ string_drop _stringappend_897_
+ (build_ex _stringappend_898_) in
+ match (reg_name_matches_prefix _stringappend_899_) with
+ | Some (rd,(existT _ _stringappend_900_ _)) =>
+ let _stringappend_901_ :=
+ string_drop _stringappend_899_
+ (build_ex _stringappend_900_) in
+ sep_matches_prefix _stringappend_901_ >>= fun w__239 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__239 with
- | Some
- (_stringappend_1328_,(existT _ _stringappend_1329_ _)) =>
- let _stringappend_1330_ :=
- string_drop _stringappend_1327_
- (build_ex
- _stringappend_1329_) in
- match (reg_name_matches_prefix _stringappend_1330_) with
- | Some
- (_stringappend_1331_,(existT _ _stringappend_1332_ _)) =>
- let _stringappend_1333_ :=
- string_drop _stringappend_1330_
- (build_ex
- _stringappend_1332_) in
- sep_matches_prefix _stringappend_1333_ >>= fun w__240 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_902_ _)) =>
+ let _stringappend_903_ :=
+ string_drop _stringappend_901_
+ (build_ex _stringappend_902_) in
+ match (reg_name_matches_prefix _stringappend_903_) with
+ | Some (rs1,(existT _ _stringappend_904_ _)) =>
+ let _stringappend_905_ :=
+ string_drop _stringappend_903_
+ (build_ex _stringappend_904_) in
+ sep_matches_prefix _stringappend_905_ >>= fun w__240 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__240 with
| Some
- (_stringappend_1334_,(existT _ _stringappend_1335_ _)) =>
- let _stringappend_1336_ :=
- string_drop _stringappend_1333_
- (build_ex
- _stringappend_1335_) in
+ (tt,(existT _ _stringappend_906_ _)) =>
+ let _stringappend_907_ :=
+ string_drop _stringappend_905_
+ (build_ex _stringappend_906_) in
if ((match (hex_bits_5_matches_prefix
- _stringappend_1336_) with
+ _stringappend_907_) with
| Some
- (_stringappend_1337_,(existT _ _stringappend_1338_ _)) =>
+ (shamt,(existT _ _stringappend_908_ _)) =>
match (string_drop
- _stringappend_1336_
- (build_ex
- _stringappend_1338_)) with
+ _stringappend_907_
+ (build_ex _stringappend_908_)) with
| "" => true
| _ => false
end
@@ -13267,146 +14098,107 @@ Definition assembly_backwards (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__245 : bool =>
(if (w__245) then
- match (shiftw_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1319_,(existT _ _stringappend_1320_ _)) =>
- returnm ((_stringappend_1319_, build_ex _stringappend_1320_)
- : (sop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((sop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__247 : (sop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_1320_ _) :=
- w__247
- : (sop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1321_ :=
- string_drop _stringappend_1112_ (build_ex _stringappend_1320_) in
- match (spc_matches_prefix _stringappend_1321_) with
- | Some (_stringappend_1322_,(existT _ _stringappend_1323_ _)) =>
- returnm ((_stringappend_1322_, build_ex _stringappend_1323_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__249 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1323_ _) :=
- w__249
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1324_ :=
- string_drop _stringappend_1321_ (build_ex _stringappend_1323_) in
- match (reg_name_matches_prefix _stringappend_1324_) with
- | Some (_stringappend_1325_,(existT _ _stringappend_1326_ _)) =>
- returnm ((_stringappend_1325_, build_ex _stringappend_1326_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__251 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_1326_ _) :=
- w__251
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1327_ :=
- string_drop _stringappend_1324_ (build_ex _stringappend_1326_) in
- sep_matches_prefix _stringappend_1327_ >>= fun w__252 : option ((unit * {n : Z & ArithFact (n >=
+ (match (shiftw_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_896_ _)) =>
+ returnm (op, build_ex _stringappend_896_)
+ | _ => exit tt : M ((sop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((sop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_896_ _) =>
+ let _stringappend_897_ :=
+ string_drop _stringappend_756_ (build_ex _stringappend_896_) in
+ (match (spc_matches_prefix _stringappend_897_) with
+ | Some (tt,(existT _ _stringappend_898_ _)) =>
+ returnm (tt, build_ex _stringappend_898_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_898_ _) =>
+ let _stringappend_899_ :=
+ string_drop _stringappend_897_ (build_ex _stringappend_898_) in
+ (match (reg_name_matches_prefix _stringappend_899_) with
+ | Some (rd,(existT _ _stringappend_900_ _)) =>
+ returnm (rd, build_ex _stringappend_900_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_900_ _) =>
+ let _stringappend_901_ :=
+ string_drop _stringappend_899_ (build_ex _stringappend_900_) in
+ sep_matches_prefix _stringappend_901_ >>= fun w__252 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__252 with
- | Some (_stringappend_1328_,(existT _ _stringappend_1329_ _)) =>
- returnm ((_stringappend_1328_, build_ex _stringappend_1329_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__254 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1329_ _) :=
- w__254
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1330_ :=
- string_drop _stringappend_1327_ (build_ex _stringappend_1329_) in
- match (reg_name_matches_prefix _stringappend_1330_) with
- | Some (_stringappend_1331_,(existT _ _stringappend_1332_ _)) =>
- returnm ((_stringappend_1331_, build_ex _stringappend_1332_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__256 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_1332_ _) :=
- w__256
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1333_ :=
- string_drop _stringappend_1330_ (build_ex _stringappend_1332_) in
- sep_matches_prefix _stringappend_1333_ >>= fun w__257 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__252 with
+ | Some (tt,(existT _ _stringappend_902_ _)) =>
+ returnm (tt, build_ex _stringappend_902_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_902_ _) =>
+ let _stringappend_903_ :=
+ string_drop _stringappend_901_ (build_ex _stringappend_902_) in
+ (match (reg_name_matches_prefix _stringappend_903_) with
+ | Some (rs1,(existT _ _stringappend_904_ _)) =>
+ returnm (rs1, build_ex _stringappend_904_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_904_ _) =>
+ let _stringappend_905_ :=
+ string_drop _stringappend_903_ (build_ex _stringappend_904_) in
+ sep_matches_prefix _stringappend_905_ >>= fun w__257 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__257 with
- | Some (_stringappend_1334_,(existT _ _stringappend_1335_ _)) =>
- returnm ((_stringappend_1334_, build_ex _stringappend_1335_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__259 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1335_ _) :=
- w__259
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1336_ :=
- string_drop _stringappend_1333_ (build_ex _stringappend_1335_) in
- match (hex_bits_5_matches_prefix _stringappend_1336_) with
- | Some (_stringappend_1337_,(existT _ _stringappend_1338_ _)) =>
- returnm ((_stringappend_1337_, build_ex _stringappend_1338_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__261 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(shamt, existT _ _stringappend_1338_ _) :=
- w__261
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1336_
- (build_ex
- _stringappend_1338_)) with
- | "" => returnm ((SHIFTW (shamt,rs1,rd,op)) : ast)
+ (match w__257 with
+ | Some (tt,(existT _ _stringappend_906_ _)) =>
+ returnm (tt, build_ex _stringappend_906_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_906_ _) =>
+ let _stringappend_907_ :=
+ string_drop _stringappend_905_ (build_ex _stringappend_906_) in
+ (match (hex_bits_5_matches_prefix _stringappend_907_) with
+ | Some (shamt,(existT _ _stringappend_908_ _)) =>
+ returnm (shamt, build_ex _stringappend_908_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(shamt, existT _ _stringappend_908_ _) =>
+ (match (string_drop _stringappend_907_
+ (build_ex _stringappend_908_)) with
+ | "" => returnm ((SHIFTW (shamt, rs1, rd, op)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
- match (rtypew_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1340_,(existT _ _stringappend_1341_ _)) =>
- let _stringappend_1342_ :=
- string_drop _stringappend_1112_
- (build_ex
- _stringappend_1341_) in
- match (spc_matches_prefix _stringappend_1342_) with
- | Some (_stringappend_1343_,(existT _ _stringappend_1344_ _)) =>
- let _stringappend_1345_ :=
- string_drop _stringappend_1342_
- (build_ex
- _stringappend_1344_) in
- match (reg_name_matches_prefix _stringappend_1345_) with
- | Some
- (_stringappend_1346_,(existT _ _stringappend_1347_ _)) =>
- let _stringappend_1348_ :=
- string_drop _stringappend_1345_
- (build_ex
- _stringappend_1347_) in
- sep_matches_prefix _stringappend_1348_ >>= fun w__264 : option ((unit * {n : Z & ArithFact (n >=
+ match (rtypew_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_910_ _)) =>
+ let _stringappend_911_ :=
+ string_drop _stringappend_756_
+ (build_ex _stringappend_910_) in
+ match (spc_matches_prefix _stringappend_911_) with
+ | Some (tt,(existT _ _stringappend_912_ _)) =>
+ let _stringappend_913_ :=
+ string_drop _stringappend_911_
+ (build_ex _stringappend_912_) in
+ match (reg_name_matches_prefix _stringappend_913_) with
+ | Some (rd,(existT _ _stringappend_914_ _)) =>
+ let _stringappend_915_ :=
+ string_drop _stringappend_913_
+ (build_ex _stringappend_914_) in
+ sep_matches_prefix _stringappend_915_ >>= fun w__264 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__264 with
- | Some
- (_stringappend_1349_,(existT _ _stringappend_1350_ _)) =>
- let _stringappend_1351_ :=
- string_drop _stringappend_1348_
- (build_ex
- _stringappend_1350_) in
- match (reg_name_matches_prefix _stringappend_1351_) with
- | Some
- (_stringappend_1352_,(existT _ _stringappend_1353_ _)) =>
- let _stringappend_1354_ :=
- string_drop _stringappend_1351_
- (build_ex
- _stringappend_1353_) in
- sep_matches_prefix _stringappend_1354_ >>= fun w__265 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_916_ _)) =>
+ let _stringappend_917_ :=
+ string_drop _stringappend_915_
+ (build_ex _stringappend_916_) in
+ match (reg_name_matches_prefix _stringappend_917_) with
+ | Some (rs1,(existT _ _stringappend_918_ _)) =>
+ let _stringappend_919_ :=
+ string_drop _stringappend_917_
+ (build_ex _stringappend_918_) in
+ sep_matches_prefix _stringappend_919_ >>= fun w__265 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__265 with
| Some
- (_stringappend_1355_,(existT _ _stringappend_1356_ _)) =>
- let _stringappend_1357_ :=
+ (tt,(existT _ _stringappend_920_ _)) =>
+ let _stringappend_921_ :=
string_drop
- _stringappend_1354_
- (build_ex
- _stringappend_1356_) in
+ _stringappend_919_
+ (build_ex _stringappend_920_) in
if ((match (reg_name_matches_prefix
- _stringappend_1357_) with
+ _stringappend_921_) with
| Some
- (_stringappend_1358_,(existT _ _stringappend_1359_ _)) =>
+ (rs2,(existT _ _stringappend_922_ _)) =>
match (string_drop
- _stringappend_1357_
- (build_ex
- _stringappend_1359_)) with
+ _stringappend_921_
+ (build_ex _stringappend_922_)) with
| "" => true
| _ => false
end
@@ -13442,168 +14234,111 @@ Definition assembly_backwards (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__270 : bool =>
(if (w__270) then
- match (rtypew_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1340_,(existT _ _stringappend_1341_ _)) =>
- returnm ((_stringappend_1340_,
- build_ex
- _stringappend_1341_)
- : (ropw * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((ropw * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__272 : (ropw * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_1341_ _) :=
- w__272
- : (ropw * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1342_ :=
- string_drop _stringappend_1112_
- (build_ex
- _stringappend_1341_) in
- match (spc_matches_prefix _stringappend_1342_) with
- | Some (_stringappend_1343_,(existT _ _stringappend_1344_ _)) =>
- returnm ((_stringappend_1343_,
- build_ex
- _stringappend_1344_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__274 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1344_ _) :=
- w__274
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1345_ :=
- string_drop _stringappend_1342_
- (build_ex
- _stringappend_1344_) in
- match (reg_name_matches_prefix _stringappend_1345_) with
- | Some (_stringappend_1346_,(existT _ _stringappend_1347_ _)) =>
- returnm ((_stringappend_1346_,
- build_ex
- _stringappend_1347_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__276 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_1347_ _) :=
- w__276
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1348_ :=
- string_drop _stringappend_1345_
- (build_ex
- _stringappend_1347_) in
- sep_matches_prefix _stringappend_1348_ >>= fun w__277 : option ((unit * {n : Z & ArithFact (n >=
+ (match (rtypew_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_910_ _)) =>
+ returnm (op, build_ex _stringappend_910_)
+ | _ => exit tt : M ((ropw * {n : Z & ArithFact (n >= 0)}))
+ end : M ((ropw * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_910_ _) =>
+ let _stringappend_911_ :=
+ string_drop _stringappend_756_
+ (build_ex _stringappend_910_) in
+ (match (spc_matches_prefix _stringappend_911_) with
+ | Some (tt,(existT _ _stringappend_912_ _)) =>
+ returnm (tt, build_ex _stringappend_912_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_912_ _) =>
+ let _stringappend_913_ :=
+ string_drop _stringappend_911_
+ (build_ex _stringappend_912_) in
+ (match (reg_name_matches_prefix _stringappend_913_) with
+ | Some (rd,(existT _ _stringappend_914_ _)) =>
+ returnm (rd, build_ex _stringappend_914_)
+ | _ =>
+ exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_914_ _) =>
+ let _stringappend_915_ :=
+ string_drop _stringappend_913_
+ (build_ex _stringappend_914_) in
+ sep_matches_prefix _stringappend_915_ >>= fun w__277 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__277 with
- | Some (_stringappend_1349_,(existT _ _stringappend_1350_ _)) =>
- returnm ((_stringappend_1349_,
- build_ex
- _stringappend_1350_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__279 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1350_ _) :=
- w__279
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1351_ :=
- string_drop _stringappend_1348_
- (build_ex
- _stringappend_1350_) in
- match (reg_name_matches_prefix _stringappend_1351_) with
- | Some (_stringappend_1352_,(existT _ _stringappend_1353_ _)) =>
- returnm ((_stringappend_1352_,
- build_ex
- _stringappend_1353_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__281 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_1353_ _) :=
- w__281
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1354_ :=
- string_drop _stringappend_1351_
- (build_ex
- _stringappend_1353_) in
- sep_matches_prefix _stringappend_1354_ >>= fun w__282 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__277 with
+ | Some (tt,(existT _ _stringappend_916_ _)) =>
+ returnm (tt, build_ex _stringappend_916_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_916_ _) =>
+ let _stringappend_917_ :=
+ string_drop _stringappend_915_
+ (build_ex _stringappend_916_) in
+ (match (reg_name_matches_prefix _stringappend_917_) with
+ | Some (rs1,(existT _ _stringappend_918_ _)) =>
+ returnm (rs1, build_ex _stringappend_918_)
+ | _ =>
+ exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_918_ _) =>
+ let _stringappend_919_ :=
+ string_drop _stringappend_917_
+ (build_ex _stringappend_918_) in
+ sep_matches_prefix _stringappend_919_ >>= fun w__282 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__282 with
- | Some (_stringappend_1355_,(existT _ _stringappend_1356_ _)) =>
- returnm ((_stringappend_1355_,
- build_ex
- _stringappend_1356_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__284 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1356_ _) :=
- w__284
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1357_ :=
- string_drop _stringappend_1354_
- (build_ex
- _stringappend_1356_) in
- match (reg_name_matches_prefix _stringappend_1357_) with
- | Some (_stringappend_1358_,(existT _ _stringappend_1359_ _)) =>
- returnm ((_stringappend_1358_,
- build_ex
- _stringappend_1359_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__286 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs2, existT _ _stringappend_1359_ _) :=
- w__286
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1357_
- (build_ex
- _stringappend_1359_)) with
- | "" => returnm ((RTYPEW (rs2,rs1,rd,op)) : ast)
+ (match w__282 with
+ | Some (tt,(existT _ _stringappend_920_ _)) =>
+ returnm (tt, build_ex _stringappend_920_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_920_ _) =>
+ let _stringappend_921_ :=
+ string_drop _stringappend_919_
+ (build_ex _stringappend_920_) in
+ (match (reg_name_matches_prefix _stringappend_921_) with
+ | Some (rs2,(existT _ _stringappend_922_ _)) =>
+ returnm (rs2, build_ex _stringappend_922_)
+ | _ =>
+ exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_922_ _) =>
+ (match (string_drop _stringappend_921_
+ (build_ex _stringappend_922_)) with
+ | "" => returnm ((RTYPEW (rs2, rs1, rd, op)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
- match (shiftiwop_mnemonic_matches_prefix _stringappend_1112_) with
- | Some (_stringappend_1361_,(existT _ _stringappend_1362_ _)) =>
- let _stringappend_1363_ :=
- string_drop _stringappend_1112_
- (build_ex
- _stringappend_1362_) in
- match (spc_matches_prefix _stringappend_1363_) with
- | Some
- (_stringappend_1364_,(existT _ _stringappend_1365_ _)) =>
- let _stringappend_1366_ :=
- string_drop _stringappend_1363_
- (build_ex
- _stringappend_1365_) in
- match (reg_name_matches_prefix _stringappend_1366_) with
- | Some
- (_stringappend_1367_,(existT _ _stringappend_1368_ _)) =>
- let _stringappend_1369_ :=
- string_drop _stringappend_1366_
- (build_ex
- _stringappend_1368_) in
- sep_matches_prefix _stringappend_1369_ >>= fun w__289 : option ((unit * {n : Z & ArithFact (n >=
+ match (shiftiwop_mnemonic_matches_prefix _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_924_ _)) =>
+ let _stringappend_925_ :=
+ string_drop _stringappend_756_
+ (build_ex _stringappend_924_) in
+ match (spc_matches_prefix _stringappend_925_) with
+ | Some (tt,(existT _ _stringappend_926_ _)) =>
+ let _stringappend_927_ :=
+ string_drop _stringappend_925_
+ (build_ex _stringappend_926_) in
+ match (reg_name_matches_prefix _stringappend_927_) with
+ | Some (rd,(existT _ _stringappend_928_ _)) =>
+ let _stringappend_929_ :=
+ string_drop _stringappend_927_
+ (build_ex _stringappend_928_) in
+ sep_matches_prefix _stringappend_929_ >>= fun w__289 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__289 with
| Some
- (_stringappend_1370_,(existT _ _stringappend_1371_ _)) =>
- let _stringappend_1372_ :=
- string_drop _stringappend_1369_
- (build_ex
- _stringappend_1371_) in
+ (tt,(existT _ _stringappend_930_ _)) =>
+ let _stringappend_931_ :=
+ string_drop _stringappend_929_
+ (build_ex _stringappend_930_) in
if ((match (reg_name_matches_prefix
- _stringappend_1372_) with
+ _stringappend_931_) with
| Some
- (_stringappend_1373_,(existT _ _stringappend_1374_ _)) =>
- let _stringappend_1375_ :=
+ (rs1,(existT _ _stringappend_932_ _)) =>
+ let _stringappend_933_ :=
string_drop
- _stringappend_1372_
- (build_ex
- _stringappend_1374_) in
+ _stringappend_931_
+ (build_ex _stringappend_932_) in
if ((match (hex_bits_5_matches_prefix
- _stringappend_1375_) with
+ _stringappend_933_) with
| Some
- (_stringappend_1376_,(existT _ _stringappend_1377_ _)) =>
+ (shamt,(existT _ _stringappend_934_ _)) =>
match (string_drop
- _stringappend_1375_
- (build_ex
- _stringappend_1377_)) with
+ _stringappend_933_
+ (build_ex _stringappend_934_)) with
| "" => true
| _ => false
end
@@ -13633,171 +14368,116 @@ Definition assembly_backwards (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__292 : bool =>
(if (w__292) then
- match (shiftiwop_mnemonic_matches_prefix
- _stringappend_1112_) with
- | Some
- (_stringappend_1361_,(existT _ _stringappend_1362_ _)) =>
- returnm ((_stringappend_1361_,
- build_ex
- _stringappend_1362_)
- : (sopw * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((sopw * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__294 : (sopw * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_1362_ _) :=
- w__294
- : (sopw * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1363_ :=
- string_drop _stringappend_1112_
- (build_ex
- _stringappend_1362_) in
- match (spc_matches_prefix _stringappend_1363_) with
- | Some
- (_stringappend_1364_,(existT _ _stringappend_1365_ _)) =>
- returnm ((_stringappend_1364_,
- build_ex
- _stringappend_1365_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__296 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1365_ _) :=
- w__296
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1366_ :=
- string_drop _stringappend_1363_
- (build_ex
- _stringappend_1365_) in
- match (reg_name_matches_prefix _stringappend_1366_) with
- | Some
- (_stringappend_1367_,(existT _ _stringappend_1368_ _)) =>
- returnm ((_stringappend_1367_,
- build_ex
- _stringappend_1368_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__298 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_1368_ _) :=
- w__298
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1369_ :=
- string_drop _stringappend_1366_
- (build_ex
- _stringappend_1368_) in
- sep_matches_prefix _stringappend_1369_ >>= fun w__299 : option ((unit * {n : Z & ArithFact (n >=
+ (match (shiftiwop_mnemonic_matches_prefix
+ _stringappend_756_) with
+ | Some (op,(existT _ _stringappend_924_ _)) =>
+ returnm (op, build_ex _stringappend_924_)
+ | _ =>
+ exit tt : M ((sopw * {n : Z & ArithFact (n >= 0)}))
+ end : M ((sopw * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_924_ _) =>
+ let _stringappend_925_ :=
+ string_drop _stringappend_756_
+ (build_ex _stringappend_924_) in
+ (match (spc_matches_prefix _stringappend_925_) with
+ | Some (tt,(existT _ _stringappend_926_ _)) =>
+ returnm (tt, build_ex _stringappend_926_)
+ | _ =>
+ exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_926_ _) =>
+ let _stringappend_927_ :=
+ string_drop _stringappend_925_
+ (build_ex _stringappend_926_) in
+ (match (reg_name_matches_prefix _stringappend_927_) with
+ | Some (rd,(existT _ _stringappend_928_ _)) =>
+ returnm (rd, build_ex _stringappend_928_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_928_ _) =>
+ let _stringappend_929_ :=
+ string_drop _stringappend_927_
+ (build_ex _stringappend_928_) in
+ sep_matches_prefix _stringappend_929_ >>= fun w__299 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__299 with
- | Some
- (_stringappend_1370_,(existT _ _stringappend_1371_ _)) =>
- returnm ((_stringappend_1370_,
- build_ex
- _stringappend_1371_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__301 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1371_ _) :=
- w__301
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1372_ :=
- string_drop _stringappend_1369_
- (build_ex
- _stringappend_1371_) in
- match (reg_name_matches_prefix _stringappend_1372_) with
- | Some
- (_stringappend_1373_,(existT _ _stringappend_1374_ _)) =>
- returnm ((_stringappend_1373_,
- build_ex
- _stringappend_1374_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__303 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_1374_ _) :=
- w__303
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1375_ :=
- string_drop _stringappend_1372_
- (build_ex
- _stringappend_1374_) in
- match (hex_bits_5_matches_prefix _stringappend_1375_) with
- | Some
- (_stringappend_1376_,(existT _ _stringappend_1377_ _)) =>
- returnm ((_stringappend_1376_,
- build_ex
- _stringappend_1377_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__305 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(shamt, existT _ _stringappend_1377_ _) :=
- w__305
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1375_
- (build_ex
- _stringappend_1377_)) with
- | "" => returnm ((SHIFTIWOP (shamt,rs1,rd,op)) : ast)
+ (match w__299 with
+ | Some (tt,(existT _ _stringappend_930_ _)) =>
+ returnm (tt, build_ex _stringappend_930_)
+ | _ =>
+ exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_930_ _) =>
+ let _stringappend_931_ :=
+ string_drop _stringappend_929_
+ (build_ex _stringappend_930_) in
+ (match (reg_name_matches_prefix _stringappend_931_) with
+ | Some (rs1,(existT _ _stringappend_932_ _)) =>
+ returnm (rs1, build_ex _stringappend_932_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_932_ _) =>
+ let _stringappend_933_ :=
+ string_drop _stringappend_931_
+ (build_ex _stringappend_932_) in
+ (match (hex_bits_5_matches_prefix _stringappend_933_) with
+ | Some (shamt,(existT _ _stringappend_934_ _)) =>
+ returnm (shamt, build_ex _stringappend_934_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(shamt, existT _ _stringappend_934_ _) =>
+ (match (string_drop _stringappend_933_
+ (build_ex _stringappend_934_)) with
+ | "" =>
+ returnm ((SHIFTIWOP (shamt, rs1, rd, op)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
- match (mul_mnemonic_matches_prefix _stringappend_1112_) with
+ match (mul_mnemonic_matches_prefix _stringappend_756_) with
| Some
- (_stringappend_1379_,(existT _ _stringappend_1380_ _)) =>
- let _stringappend_1381_ :=
- string_drop _stringappend_1112_
- (build_ex
- _stringappend_1380_) in
- match (spc_matches_prefix _stringappend_1381_) with
- | Some
- (_stringappend_1382_,(existT _ _stringappend_1383_ _)) =>
- let _stringappend_1384_ :=
- string_drop _stringappend_1381_
- (build_ex
- _stringappend_1383_) in
- match (reg_name_matches_prefix _stringappend_1384_) with
- | Some
- (_stringappend_1385_,(existT _ _stringappend_1386_ _)) =>
- let _stringappend_1387_ :=
- string_drop _stringappend_1384_
- (build_ex
- _stringappend_1386_) in
- sep_matches_prefix _stringappend_1387_ >>= fun w__308 : option ((unit * {n : Z & ArithFact (n >=
+ ((high, signed1, signed2),(existT _ _stringappend_936_ _)) =>
+ let _stringappend_937_ :=
+ string_drop _stringappend_756_
+ (build_ex _stringappend_936_) in
+ match (spc_matches_prefix _stringappend_937_) with
+ | Some (tt,(existT _ _stringappend_938_ _)) =>
+ let _stringappend_939_ :=
+ string_drop _stringappend_937_
+ (build_ex _stringappend_938_) in
+ match (reg_name_matches_prefix _stringappend_939_) with
+ | Some (rd,(existT _ _stringappend_940_ _)) =>
+ let _stringappend_941_ :=
+ string_drop _stringappend_939_
+ (build_ex _stringappend_940_) in
+ sep_matches_prefix _stringappend_941_ >>= fun w__308 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__308 with
- | Some
- (_stringappend_1388_,(existT _ _stringappend_1389_ _)) =>
- let _stringappend_1390_ :=
- string_drop _stringappend_1387_
- (build_ex
- _stringappend_1389_) in
+ | Some (tt,(existT _ _stringappend_942_ _)) =>
+ let _stringappend_943_ :=
+ string_drop _stringappend_941_
+ (build_ex _stringappend_942_) in
match (reg_name_matches_prefix
- _stringappend_1390_) with
- | Some
- (_stringappend_1391_,(existT _ _stringappend_1392_ _)) =>
- let _stringappend_1393_ :=
- string_drop _stringappend_1390_
- (build_ex
- _stringappend_1392_) in
- sep_matches_prefix _stringappend_1393_ >>= fun w__309 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_943_) with
+ | Some (rs1,(existT _ _stringappend_944_ _)) =>
+ let _stringappend_945_ :=
+ string_drop _stringappend_943_
+ (build_ex _stringappend_944_) in
+ sep_matches_prefix _stringappend_945_ >>= fun w__309 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__309 with
| Some
- (_stringappend_1394_,(existT _ _stringappend_1395_ _)) =>
- let _stringappend_1396_ :=
+ (tt,(existT _ _stringappend_946_ _)) =>
+ let _stringappend_947_ :=
string_drop
- _stringappend_1393_
- (build_ex
- _stringappend_1395_) in
+ _stringappend_945_
+ (build_ex _stringappend_946_) in
if ((match (reg_name_matches_prefix
- _stringappend_1396_) with
+ _stringappend_947_) with
| Some
- (_stringappend_1397_,(existT _ _stringappend_1398_ _)) =>
+ (rs2,(existT _ _stringappend_948_ _)) =>
match (string_drop
- _stringappend_1396_
- (build_ex
- _stringappend_1398_)) with
+ _stringappend_947_
+ (build_ex _stringappend_948_)) with
| "" => true
| _ => false
end
@@ -13833,209 +14513,142 @@ Definition assembly_backwards (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__314 : bool =>
(if (w__314) then
- match (mul_mnemonic_matches_prefix _stringappend_1112_) with
- | Some
- (_stringappend_1379_,(existT _ _stringappend_1380_ _)) =>
- returnm ((_stringappend_1379_,
- build_ex
- _stringappend_1380_)
- : ((bool * bool * bool) * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M (((bool * bool * bool) * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__316 : ((bool * bool * bool) * {n : Z & ArithFact (n >=
- 0)}) =>
- let '((high, signed1, signed2), existT _ _stringappend_1380_ _) :=
- w__316
- : ((bool * bool * bool) * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1381_ :=
- string_drop _stringappend_1112_
- (build_ex
- _stringappend_1380_) in
- match (spc_matches_prefix _stringappend_1381_) with
- | Some
- (_stringappend_1382_,(existT _ _stringappend_1383_ _)) =>
- returnm ((_stringappend_1382_,
- build_ex
- _stringappend_1383_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__318 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1383_ _) :=
- w__318
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1384_ :=
- string_drop _stringappend_1381_
- (build_ex
- _stringappend_1383_) in
- match (reg_name_matches_prefix _stringappend_1384_) with
- | Some
- (_stringappend_1385_,(existT _ _stringappend_1386_ _)) =>
- returnm ((_stringappend_1385_,
- build_ex
- _stringappend_1386_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__320 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1386_ _) :=
- w__320
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1387_ :=
- string_drop _stringappend_1384_
- (build_ex
- _stringappend_1386_) in
- sep_matches_prefix _stringappend_1387_ >>= fun w__321 : option ((unit * {n : Z & ArithFact (n >=
+ (match (mul_mnemonic_matches_prefix _stringappend_756_) with
+ | Some
+ ((high, signed1, signed2),(existT _ _stringappend_936_ _)) =>
+ returnm ((high, signed1, signed2), build_ex _stringappend_936_)
+ | _ =>
+ exit tt
+ : M (((bool * bool * bool) * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M (((bool * bool * bool) * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '((high, signed1, signed2), existT _ _stringappend_936_ _) =>
+ let _stringappend_937_ :=
+ string_drop _stringappend_756_
+ (build_ex _stringappend_936_) in
+ (match (spc_matches_prefix _stringappend_937_) with
+ | Some (tt,(existT _ _stringappend_938_ _)) =>
+ returnm (tt, build_ex _stringappend_938_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_938_ _) =>
+ let _stringappend_939_ :=
+ string_drop _stringappend_937_
+ (build_ex _stringappend_938_) in
+ (match (reg_name_matches_prefix _stringappend_939_) with
+ | Some (rd,(existT _ _stringappend_940_ _)) =>
+ returnm (rd, build_ex _stringappend_940_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_940_ _) =>
+ let _stringappend_941_ :=
+ string_drop _stringappend_939_
+ (build_ex _stringappend_940_) in
+ sep_matches_prefix _stringappend_941_ >>= fun w__321 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__321 with
- | Some
- (_stringappend_1388_,(existT _ _stringappend_1389_ _)) =>
- returnm ((_stringappend_1388_,
- build_ex
- _stringappend_1389_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__323 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1389_ _) :=
- w__323
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1390_ :=
- string_drop _stringappend_1387_
- (build_ex
- _stringappend_1389_) in
- match (reg_name_matches_prefix _stringappend_1390_) with
- | Some
- (_stringappend_1391_,(existT _ _stringappend_1392_ _)) =>
- returnm ((_stringappend_1391_,
- build_ex
- _stringappend_1392_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__325 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1392_ _) :=
- w__325
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1393_ :=
- string_drop _stringappend_1390_
- (build_ex
- _stringappend_1392_) in
- sep_matches_prefix _stringappend_1393_ >>= fun w__326 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__321 with
+ | Some (tt,(existT _ _stringappend_942_ _)) =>
+ returnm (tt, build_ex _stringappend_942_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_942_ _) =>
+ let _stringappend_943_ :=
+ string_drop _stringappend_941_
+ (build_ex _stringappend_942_) in
+ (match (reg_name_matches_prefix _stringappend_943_) with
+ | Some (rs1,(existT _ _stringappend_944_ _)) =>
+ returnm (rs1, build_ex _stringappend_944_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_944_ _) =>
+ let _stringappend_945_ :=
+ string_drop _stringappend_943_
+ (build_ex _stringappend_944_) in
+ sep_matches_prefix _stringappend_945_ >>= fun w__326 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__326 with
- | Some
- (_stringappend_1394_,(existT _ _stringappend_1395_ _)) =>
- returnm ((_stringappend_1394_,
- build_ex
- _stringappend_1395_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__328 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_1395_ _) :=
- w__328
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1396_ :=
- string_drop _stringappend_1393_
- (build_ex
- _stringappend_1395_) in
- match (reg_name_matches_prefix _stringappend_1396_) with
- | Some
- (_stringappend_1397_,(existT _ _stringappend_1398_ _)) =>
- returnm ((_stringappend_1397_,
- build_ex
- _stringappend_1398_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__330 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_1398_ _) :=
- w__330
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1396_
- (build_ex
- _stringappend_1398_)) with
+ (match w__326 with
+ | Some (tt,(existT _ _stringappend_946_ _)) =>
+ returnm (tt, build_ex _stringappend_946_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_946_ _) =>
+ let _stringappend_947_ :=
+ string_drop _stringappend_945_
+ (build_ex _stringappend_946_) in
+ (match (reg_name_matches_prefix _stringappend_947_) with
+ | Some (rs2,(existT _ _stringappend_948_ _)) =>
+ returnm (rs2, build_ex _stringappend_948_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_948_ _) =>
+ (match (string_drop _stringappend_947_
+ (build_ex _stringappend_948_)) with
| "" =>
- returnm ((MUL (rs2,rs1,rd,high,signed1,signed2))
- : ast)
+ returnm ((MUL
+ (rs2, rs1, rd, high, signed1, signed2))
+ : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
and_boolM
- (returnm ((string_startswith _stringappend_1112_
- "div")
+ (returnm ((string_startswith _stringappend_756_ "div")
: bool))
- (let _stringappend_1400_ :=
- string_drop _stringappend_1112_
- (string_length "div") in
- match (maybe_not_u_matches_prefix _stringappend_1400_) with
- | Some
- (_stringappend_1401_,(existT _ _stringappend_1402_ _)) =>
- let _stringappend_1403_ :=
- string_drop _stringappend_1400_
- (build_ex
- _stringappend_1402_) in
- match (spc_matches_prefix _stringappend_1403_) with
- | Some
- (_stringappend_1404_,(existT _ _stringappend_1405_ _)) =>
- let _stringappend_1406_ :=
- string_drop _stringappend_1403_
- (build_ex
- _stringappend_1405_) in
+ (let _stringappend_950_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length "div"))) in
+ match (maybe_not_u_matches_prefix _stringappend_950_) with
+ | Some (s,(existT _ _stringappend_951_ _)) =>
+ let _stringappend_952_ :=
+ string_drop _stringappend_950_
+ (build_ex _stringappend_951_) in
+ match (spc_matches_prefix _stringappend_952_) with
+ | Some (tt,(existT _ _stringappend_953_ _)) =>
+ let _stringappend_954_ :=
+ string_drop _stringappend_952_
+ (build_ex _stringappend_953_) in
match (reg_name_matches_prefix
- _stringappend_1406_) with
- | Some
- (_stringappend_1407_,(existT _ _stringappend_1408_ _)) =>
- let _stringappend_1409_ :=
- string_drop _stringappend_1406_
- (build_ex
- _stringappend_1408_) in
- sep_matches_prefix _stringappend_1409_ >>= fun w__333 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_954_) with
+ | Some (rd,(existT _ _stringappend_955_ _)) =>
+ let _stringappend_956_ :=
+ string_drop _stringappend_954_
+ (build_ex _stringappend_955_) in
+ sep_matches_prefix _stringappend_956_ >>= fun w__333 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__333 with
- | Some
- (_stringappend_1410_,(existT _ _stringappend_1411_ _)) =>
- let _stringappend_1412_ :=
- string_drop _stringappend_1409_
- (build_ex
- _stringappend_1411_) in
+ | Some (tt,(existT _ _stringappend_957_ _)) =>
+ let _stringappend_958_ :=
+ string_drop _stringappend_956_
+ (build_ex _stringappend_957_) in
match (reg_name_matches_prefix
- _stringappend_1412_) with
+ _stringappend_958_) with
| Some
- (_stringappend_1413_,(existT _ _stringappend_1414_ _)) =>
- let _stringappend_1415_ :=
- string_drop _stringappend_1412_
- (build_ex
- _stringappend_1414_) in
- sep_matches_prefix _stringappend_1415_ >>= fun w__334 : option ((unit * {n : Z & ArithFact (n >=
+ (rs1,(existT _ _stringappend_959_ _)) =>
+ let _stringappend_960_ :=
+ string_drop _stringappend_958_
+ (build_ex _stringappend_959_) in
+ sep_matches_prefix _stringappend_960_ >>= fun w__334 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__334 with
| Some
- (_stringappend_1416_,(existT _ _stringappend_1417_ _)) =>
- let _stringappend_1418_ :=
+ (tt,(existT _ _stringappend_961_ _)) =>
+ let _stringappend_962_ :=
string_drop
- _stringappend_1415_
- (build_ex
- _stringappend_1417_) in
+ _stringappend_960_
+ (build_ex _stringappend_961_) in
if ((match (reg_name_matches_prefix
- _stringappend_1418_) with
+ _stringappend_962_) with
| Some
- (_stringappend_1419_,(existT _ _stringappend_1420_ _)) =>
+ (rs2,(existT _ _stringappend_963_ _)) =>
match (string_drop
- _stringappend_1418_
- (build_ex
- _stringappend_1420_)) with
+ _stringappend_962_
+ (build_ex _stringappend_963_)) with
| "" => true
| _ => false
end
@@ -14074,217 +14687,144 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__340 : bool =>
(if (w__340) then
- let _stringappend_1400_ :=
- string_drop _stringappend_1112_
- (string_length "div") in
- match (maybe_not_u_matches_prefix
- _stringappend_1400_) with
- | Some
- (_stringappend_1401_,(existT _ _stringappend_1402_ _)) =>
- returnm ((_stringappend_1401_,
- build_ex
- _stringappend_1402_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__342 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_1402_ _) :=
- w__342
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1403_ :=
- string_drop _stringappend_1400_
- (build_ex
- _stringappend_1402_) in
- match (spc_matches_prefix _stringappend_1403_) with
- | Some
- (_stringappend_1404_,(existT _ _stringappend_1405_ _)) =>
- returnm ((_stringappend_1404_,
- build_ex
- _stringappend_1405_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__344 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1405_ _) :=
- w__344
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1406_ :=
- string_drop _stringappend_1403_
- (build_ex
- _stringappend_1405_) in
- match (reg_name_matches_prefix _stringappend_1406_) with
- | Some
- (_stringappend_1407_,(existT _ _stringappend_1408_ _)) =>
- returnm ((_stringappend_1407_,
- build_ex
- _stringappend_1408_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__346 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1408_ _) :=
- w__346
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1409_ :=
- string_drop _stringappend_1406_
- (build_ex
- _stringappend_1408_) in
- sep_matches_prefix _stringappend_1409_ >>= fun w__347 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_950_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length "div"))) in
+ (match (maybe_not_u_matches_prefix
+ _stringappend_950_) with
+ | Some (s,(existT _ _stringappend_951_ _)) =>
+ returnm (s, build_ex _stringappend_951_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(s, existT _ _stringappend_951_ _) =>
+ let _stringappend_952_ :=
+ string_drop _stringappend_950_
+ (build_ex _stringappend_951_) in
+ (match (spc_matches_prefix _stringappend_952_) with
+ | Some (tt,(existT _ _stringappend_953_ _)) =>
+ returnm (tt, build_ex _stringappend_953_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_953_ _) =>
+ let _stringappend_954_ :=
+ string_drop _stringappend_952_
+ (build_ex _stringappend_953_) in
+ (match (reg_name_matches_prefix _stringappend_954_) with
+ | Some (rd,(existT _ _stringappend_955_ _)) =>
+ returnm (rd, build_ex _stringappend_955_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_955_ _) =>
+ let _stringappend_956_ :=
+ string_drop _stringappend_954_
+ (build_ex _stringappend_955_) in
+ sep_matches_prefix _stringappend_956_ >>= fun w__347 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__347 with
- | Some
- (_stringappend_1410_,(existT _ _stringappend_1411_ _)) =>
- returnm ((_stringappend_1410_,
- build_ex
- _stringappend_1411_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__349 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1411_ _) :=
- w__349
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1412_ :=
- string_drop _stringappend_1409_
- (build_ex
- _stringappend_1411_) in
- match (reg_name_matches_prefix _stringappend_1412_) with
- | Some
- (_stringappend_1413_,(existT _ _stringappend_1414_ _)) =>
- returnm ((_stringappend_1413_,
- build_ex
- _stringappend_1414_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__351 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1414_ _) :=
- w__351
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1415_ :=
- string_drop _stringappend_1412_
- (build_ex
- _stringappend_1414_) in
- sep_matches_prefix _stringappend_1415_ >>= fun w__352 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__347 with
+ | Some (tt,(existT _ _stringappend_957_ _)) =>
+ returnm (tt, build_ex _stringappend_957_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_957_ _) =>
+ let _stringappend_958_ :=
+ string_drop _stringappend_956_
+ (build_ex _stringappend_957_) in
+ (match (reg_name_matches_prefix _stringappend_958_) with
+ | Some (rs1,(existT _ _stringappend_959_ _)) =>
+ returnm (rs1, build_ex _stringappend_959_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_959_ _) =>
+ let _stringappend_960_ :=
+ string_drop _stringappend_958_
+ (build_ex _stringappend_959_) in
+ sep_matches_prefix _stringappend_960_ >>= fun w__352 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__352 with
- | Some
- (_stringappend_1416_,(existT _ _stringappend_1417_ _)) =>
- returnm ((_stringappend_1416_,
- build_ex
- _stringappend_1417_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__354 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1417_ _) :=
- w__354
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1418_ :=
- string_drop _stringappend_1415_
- (build_ex
- _stringappend_1417_) in
- match (reg_name_matches_prefix _stringappend_1418_) with
- | Some
- (_stringappend_1419_,(existT _ _stringappend_1420_ _)) =>
- returnm ((_stringappend_1419_,
- build_ex
- _stringappend_1420_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__356 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_1420_ _) :=
- w__356
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1418_
- (build_ex
- _stringappend_1420_)) with
- | "" => returnm ((DIV (rs2,rs1,rd,s)) : ast)
+ (match w__352 with
+ | Some (tt,(existT _ _stringappend_961_ _)) =>
+ returnm (tt, build_ex _stringappend_961_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_961_ _) =>
+ let _stringappend_962_ :=
+ string_drop _stringappend_960_
+ (build_ex _stringappend_961_) in
+ (match (reg_name_matches_prefix _stringappend_962_) with
+ | Some (rs2,(existT _ _stringappend_963_ _)) =>
+ returnm (rs2, build_ex _stringappend_963_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_963_ _) =>
+ (match (string_drop _stringappend_962_
+ (build_ex _stringappend_963_)) with
+ | "" => returnm ((DIV (rs2, rs1, rd, s)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
and_boolM
- (returnm ((string_startswith _stringappend_1112_
+ (returnm ((string_startswith _stringappend_756_
"rem")
: bool))
- (let _stringappend_1422_ :=
- string_drop _stringappend_1112_
- (string_length "rem") in
+ (let _stringappend_965_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length "rem"))) in
match (maybe_not_u_matches_prefix
- _stringappend_1422_) with
- | Some
- (_stringappend_1423_,(existT _ _stringappend_1424_ _)) =>
- let _stringappend_1425_ :=
- string_drop _stringappend_1422_
- (build_ex
- _stringappend_1424_) in
- match (spc_matches_prefix _stringappend_1425_) with
- | Some
- (_stringappend_1426_,(existT _ _stringappend_1427_ _)) =>
- let _stringappend_1428_ :=
- string_drop _stringappend_1425_
- (build_ex
- _stringappend_1427_) in
+ _stringappend_965_) with
+ | Some (s,(existT _ _stringappend_966_ _)) =>
+ let _stringappend_967_ :=
+ string_drop _stringappend_965_
+ (build_ex _stringappend_966_) in
+ match (spc_matches_prefix _stringappend_967_) with
+ | Some (tt,(existT _ _stringappend_968_ _)) =>
+ let _stringappend_969_ :=
+ string_drop _stringappend_967_
+ (build_ex _stringappend_968_) in
match (reg_name_matches_prefix
- _stringappend_1428_) with
- | Some
- (_stringappend_1429_,(existT _ _stringappend_1430_ _)) =>
- let _stringappend_1431_ :=
- string_drop _stringappend_1428_
- (build_ex
- _stringappend_1430_) in
- sep_matches_prefix _stringappend_1431_ >>= fun w__359 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_969_) with
+ | Some (rd,(existT _ _stringappend_970_ _)) =>
+ let _stringappend_971_ :=
+ string_drop _stringappend_969_
+ (build_ex _stringappend_970_) in
+ sep_matches_prefix _stringappend_971_ >>= fun w__359 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__359 with
| Some
- (_stringappend_1432_,(existT _ _stringappend_1433_ _)) =>
- let _stringappend_1434_ :=
- string_drop _stringappend_1431_
- (build_ex
- _stringappend_1433_) in
+ (tt,(existT _ _stringappend_972_ _)) =>
+ let _stringappend_973_ :=
+ string_drop _stringappend_971_
+ (build_ex _stringappend_972_) in
match (reg_name_matches_prefix
- _stringappend_1434_) with
+ _stringappend_973_) with
| Some
- (_stringappend_1435_,(existT _ _stringappend_1436_ _)) =>
- let _stringappend_1437_ :=
- string_drop _stringappend_1434_
- (build_ex
- _stringappend_1436_) in
+ (rs1,(existT _ _stringappend_974_ _)) =>
+ let _stringappend_975_ :=
+ string_drop _stringappend_973_
+ (build_ex _stringappend_974_) in
sep_matches_prefix
- _stringappend_1437_ >>= fun w__360 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_975_ >>= fun w__360 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__360 with
| Some
- (_stringappend_1438_,(existT _ _stringappend_1439_ _)) =>
- let _stringappend_1440_ :=
+ (tt,(existT _ _stringappend_976_ _)) =>
+ let _stringappend_977_ :=
string_drop
- _stringappend_1437_
- (build_ex
- _stringappend_1439_) in
+ _stringappend_975_
+ (build_ex _stringappend_976_) in
if ((match (reg_name_matches_prefix
- _stringappend_1440_) with
+ _stringappend_977_) with
| Some
- (_stringappend_1441_,(existT _ _stringappend_1442_ _)) =>
+ (rs2,(existT _ _stringappend_978_ _)) =>
match (string_drop
- _stringappend_1440_
- (build_ex
- _stringappend_1442_)) with
+ _stringappend_977_
+ (build_ex _stringappend_978_)) with
| "" =>
true
| _ =>
@@ -14326,212 +14866,142 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__366 : bool =>
(if (w__366) then
- let _stringappend_1422_ :=
- string_drop _stringappend_1112_
- (string_length "rem") in
- match (maybe_not_u_matches_prefix
- _stringappend_1422_) with
- | Some
- (_stringappend_1423_,(existT _ _stringappend_1424_ _)) =>
- returnm ((_stringappend_1423_,
- build_ex
- _stringappend_1424_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__368 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_1424_ _) :=
- w__368
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1425_ :=
- string_drop _stringappend_1422_
- (build_ex
- _stringappend_1424_) in
- match (spc_matches_prefix _stringappend_1425_) with
- | Some
- (_stringappend_1426_,(existT _ _stringappend_1427_ _)) =>
- returnm ((_stringappend_1426_,
- build_ex
- _stringappend_1427_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__370 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1427_ _) :=
- w__370
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1428_ :=
- string_drop _stringappend_1425_
- (build_ex
- _stringappend_1427_) in
- match (reg_name_matches_prefix
- _stringappend_1428_) with
- | Some
- (_stringappend_1429_,(existT _ _stringappend_1430_ _)) =>
- returnm ((_stringappend_1429_,
- build_ex
- _stringappend_1430_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__372 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1430_ _) :=
- w__372
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1431_ :=
- string_drop _stringappend_1428_
- (build_ex
- _stringappend_1430_) in
- sep_matches_prefix _stringappend_1431_ >>= fun w__373 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_965_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length "rem"))) in
+ (match (maybe_not_u_matches_prefix
+ _stringappend_965_) with
+ | Some (s,(existT _ _stringappend_966_ _)) =>
+ returnm (s, build_ex _stringappend_966_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(s, existT _ _stringappend_966_ _) =>
+ let _stringappend_967_ :=
+ string_drop _stringappend_965_
+ (build_ex _stringappend_966_) in
+ (match (spc_matches_prefix _stringappend_967_) with
+ | Some (tt,(existT _ _stringappend_968_ _)) =>
+ returnm (tt, build_ex _stringappend_968_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_968_ _) =>
+ let _stringappend_969_ :=
+ string_drop _stringappend_967_
+ (build_ex _stringappend_968_) in
+ (match (reg_name_matches_prefix
+ _stringappend_969_) with
+ | Some (rd,(existT _ _stringappend_970_ _)) =>
+ returnm (rd, build_ex _stringappend_970_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_970_ _) =>
+ let _stringappend_971_ :=
+ string_drop _stringappend_969_
+ (build_ex _stringappend_970_) in
+ sep_matches_prefix _stringappend_971_ >>= fun w__373 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__373 with
- | Some
- (_stringappend_1432_,(existT _ _stringappend_1433_ _)) =>
- returnm ((_stringappend_1432_,
- build_ex
- _stringappend_1433_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__375 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1433_ _) :=
- w__375
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1434_ :=
- string_drop _stringappend_1431_
- (build_ex
- _stringappend_1433_) in
- match (reg_name_matches_prefix
- _stringappend_1434_) with
- | Some
- (_stringappend_1435_,(existT _ _stringappend_1436_ _)) =>
- returnm ((_stringappend_1435_,
- build_ex
- _stringappend_1436_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__377 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1436_ _) :=
- w__377
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1437_ :=
- string_drop _stringappend_1434_
- (build_ex
- _stringappend_1436_) in
- sep_matches_prefix _stringappend_1437_ >>= fun w__378 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__373 with
+ | Some (tt,(existT _ _stringappend_972_ _)) =>
+ returnm (tt, build_ex _stringappend_972_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_972_ _) =>
+ let _stringappend_973_ :=
+ string_drop _stringappend_971_
+ (build_ex _stringappend_972_) in
+ (match (reg_name_matches_prefix
+ _stringappend_973_) with
+ | Some (rs1,(existT _ _stringappend_974_ _)) =>
+ returnm (rs1, build_ex _stringappend_974_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_974_ _) =>
+ let _stringappend_975_ :=
+ string_drop _stringappend_973_
+ (build_ex _stringappend_974_) in
+ sep_matches_prefix _stringappend_975_ >>= fun w__378 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__378 with
- | Some
- (_stringappend_1438_,(existT _ _stringappend_1439_ _)) =>
- returnm ((_stringappend_1438_,
- build_ex
- _stringappend_1439_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__380 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1439_ _) :=
- w__380
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1440_ :=
- string_drop _stringappend_1437_
- (build_ex
- _stringappend_1439_) in
- match (reg_name_matches_prefix
- _stringappend_1440_) with
- | Some
- (_stringappend_1441_,(existT _ _stringappend_1442_ _)) =>
- returnm ((_stringappend_1441_,
- build_ex
- _stringappend_1442_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__382 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_1442_ _) :=
- w__382
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1440_
- (build_ex
- _stringappend_1442_)) with
- | "" => returnm ((REM (rs2,rs1,rd,s)) : ast)
+ (match w__378 with
+ | Some (tt,(existT _ _stringappend_976_ _)) =>
+ returnm (tt, build_ex _stringappend_976_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_976_ _) =>
+ let _stringappend_977_ :=
+ string_drop _stringappend_975_
+ (build_ex _stringappend_976_) in
+ (match (reg_name_matches_prefix
+ _stringappend_977_) with
+ | Some (rs2,(existT _ _stringappend_978_ _)) =>
+ returnm (rs2, build_ex _stringappend_978_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_978_ _) =>
+ (match (string_drop _stringappend_977_
+ (build_ex _stringappend_978_)) with
+ | "" =>
+ returnm ((REM (rs2, rs1, rd, s)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
and_boolM
- (returnm ((string_startswith
- _stringappend_1112_ "mulw")
+ (returnm ((string_startswith _stringappend_756_
+ "mulw")
: bool))
- (let _stringappend_1444_ :=
- string_drop _stringappend_1112_
- (string_length "mulw") in
- match (spc_matches_prefix _stringappend_1444_) with
- | Some
- (_stringappend_1445_,(existT _ _stringappend_1446_ _)) =>
- let _stringappend_1447_ :=
- string_drop _stringappend_1444_
- (build_ex
- _stringappend_1446_) in
+ (let _stringappend_980_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length "mulw"))) in
+ match (spc_matches_prefix _stringappend_980_) with
+ | Some (tt,(existT _ _stringappend_981_ _)) =>
+ let _stringappend_982_ :=
+ string_drop _stringappend_980_
+ (build_ex _stringappend_981_) in
match (reg_name_matches_prefix
- _stringappend_1447_) with
- | Some
- (_stringappend_1448_,(existT _ _stringappend_1449_ _)) =>
- let _stringappend_1450_ :=
- string_drop _stringappend_1447_
- (build_ex
- _stringappend_1449_) in
- sep_matches_prefix _stringappend_1450_ >>= fun w__385 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_982_) with
+ | Some (rd,(existT _ _stringappend_983_ _)) =>
+ let _stringappend_984_ :=
+ string_drop _stringappend_982_
+ (build_ex _stringappend_983_) in
+ sep_matches_prefix _stringappend_984_ >>= fun w__385 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__385 with
| Some
- (_stringappend_1451_,(existT _ _stringappend_1452_ _)) =>
- let _stringappend_1453_ :=
- string_drop _stringappend_1450_
- (build_ex
- _stringappend_1452_) in
+ (tt,(existT _ _stringappend_985_ _)) =>
+ let _stringappend_986_ :=
+ string_drop _stringappend_984_
+ (build_ex _stringappend_985_) in
match (reg_name_matches_prefix
- _stringappend_1453_) with
+ _stringappend_986_) with
| Some
- (_stringappend_1454_,(existT _ _stringappend_1455_ _)) =>
- let _stringappend_1456_ :=
- string_drop _stringappend_1453_
- (build_ex
- _stringappend_1455_) in
+ (rs1,(existT _ _stringappend_987_ _)) =>
+ let _stringappend_988_ :=
+ string_drop _stringappend_986_
+ (build_ex _stringappend_987_) in
sep_matches_prefix
- _stringappend_1456_ >>= fun w__386 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_988_ >>= fun w__386 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__386 with
| Some
- (_stringappend_1457_,(existT _ _stringappend_1458_ _)) =>
- let _stringappend_1459_ :=
+ (tt,(existT _ _stringappend_989_ _)) =>
+ let _stringappend_990_ :=
string_drop
- _stringappend_1456_
- (build_ex
- _stringappend_1458_) in
+ _stringappend_988_
+ (build_ex _stringappend_989_) in
if ((match (reg_name_matches_prefix
- _stringappend_1459_) with
+ _stringappend_990_) with
| Some
- (_stringappend_1460_,(existT _ _stringappend_1461_ _)) =>
+ (rs2,(existT _ _stringappend_991_ _)) =>
match (string_drop
- _stringappend_1459_
- (build_ex
- _stringappend_1461_)) with
+ _stringappend_990_
+ (build_ex _stringappend_991_)) with
| "" =>
true
| _ =>
@@ -14568,211 +15038,157 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__391 : bool =>
(if (w__391) then
- let _stringappend_1444_ :=
- string_drop _stringappend_1112_
- (string_length "mulw") in
- match (spc_matches_prefix _stringappend_1444_) with
- | Some
- (_stringappend_1445_,(existT _ _stringappend_1446_ _)) =>
- returnm ((_stringappend_1445_,
- build_ex
- _stringappend_1446_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__393 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1446_ _) :=
- w__393
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1447_ :=
- string_drop _stringappend_1444_
- (build_ex
- _stringappend_1446_) in
- match (reg_name_matches_prefix
- _stringappend_1447_) with
- | Some
- (_stringappend_1448_,(existT _ _stringappend_1449_ _)) =>
- returnm ((_stringappend_1448_,
- build_ex
- _stringappend_1449_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__395 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1449_ _) :=
- w__395
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1450_ :=
- string_drop _stringappend_1447_
- (build_ex
- _stringappend_1449_) in
- sep_matches_prefix _stringappend_1450_ >>= fun w__396 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_980_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length "mulw"))) in
+ (match (spc_matches_prefix _stringappend_980_) with
+ | Some (tt,(existT _ _stringappend_981_ _)) =>
+ returnm (tt, build_ex _stringappend_981_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_981_ _) =>
+ let _stringappend_982_ :=
+ string_drop _stringappend_980_
+ (build_ex _stringappend_981_) in
+ (match (reg_name_matches_prefix
+ _stringappend_982_) with
+ | Some (rd,(existT _ _stringappend_983_ _)) =>
+ returnm (rd, build_ex _stringappend_983_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_983_ _) =>
+ let _stringappend_984_ :=
+ string_drop _stringappend_982_
+ (build_ex _stringappend_983_) in
+ sep_matches_prefix _stringappend_984_ >>= fun w__396 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__396 with
- | Some
- (_stringappend_1451_,(existT _ _stringappend_1452_ _)) =>
- returnm ((_stringappend_1451_,
- build_ex
- _stringappend_1452_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__398 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1452_ _) :=
- w__398
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1453_ :=
- string_drop _stringappend_1450_
- (build_ex
- _stringappend_1452_) in
- match (reg_name_matches_prefix
- _stringappend_1453_) with
- | Some
- (_stringappend_1454_,(existT _ _stringappend_1455_ _)) =>
- returnm ((_stringappend_1454_,
- build_ex
- _stringappend_1455_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__400 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1455_ _) :=
- w__400
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1456_ :=
- string_drop _stringappend_1453_
- (build_ex
- _stringappend_1455_) in
- sep_matches_prefix _stringappend_1456_ >>= fun w__401 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__396 with
+ | Some (tt,(existT _ _stringappend_985_ _)) =>
+ returnm (tt, build_ex _stringappend_985_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_985_ _) =>
+ let _stringappend_986_ :=
+ string_drop _stringappend_984_
+ (build_ex _stringappend_985_) in
+ (match (reg_name_matches_prefix
+ _stringappend_986_) with
+ | Some (rs1,(existT _ _stringappend_987_ _)) =>
+ returnm (rs1, build_ex _stringappend_987_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_987_ _) =>
+ let _stringappend_988_ :=
+ string_drop _stringappend_986_
+ (build_ex _stringappend_987_) in
+ sep_matches_prefix _stringappend_988_ >>= fun w__401 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__401 with
- | Some
- (_stringappend_1457_,(existT _ _stringappend_1458_ _)) =>
- returnm ((_stringappend_1457_,
- build_ex
- _stringappend_1458_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__403 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1458_ _) :=
- w__403
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1459_ :=
- string_drop _stringappend_1456_
- (build_ex
- _stringappend_1458_) in
- match (reg_name_matches_prefix
- _stringappend_1459_) with
- | Some
- (_stringappend_1460_,(existT _ _stringappend_1461_ _)) =>
- returnm ((_stringappend_1460_,
- build_ex
- _stringappend_1461_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__405 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_1461_ _) :=
- w__405
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1459_
- (build_ex
- _stringappend_1461_)) with
- | "" => returnm ((MULW (rs2,rs1,rd)) : ast)
+ (match w__401 with
+ | Some (tt,(existT _ _stringappend_989_ _)) =>
+ returnm (tt, build_ex _stringappend_989_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_989_ _) =>
+ let _stringappend_990_ :=
+ string_drop _stringappend_988_
+ (build_ex _stringappend_989_) in
+ (match (reg_name_matches_prefix
+ _stringappend_990_) with
+ | Some (rs2,(existT _ _stringappend_991_ _)) =>
+ returnm (rs2, build_ex _stringappend_991_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_991_ _) =>
+ (match (string_drop _stringappend_990_
+ (build_ex _stringappend_991_)) with
+ | "" =>
+ returnm ((MULW (rs2, rs1, rd)) : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
and_boolM
(returnm ((string_startswith
- _stringappend_1112_ "div")
+ _stringappend_756_ "div")
: bool))
- (let _stringappend_1463_ :=
- string_drop _stringappend_1112_
- (string_length "div") in
+ (let _stringappend_993_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length "div"))) in
match (maybe_not_u_matches_prefix
- _stringappend_1463_) with
- | Some
- (_stringappend_1464_,(existT _ _stringappend_1465_ _)) =>
- let _stringappend_1466_ :=
- string_drop _stringappend_1463_
- (build_ex
- _stringappend_1465_) in
+ _stringappend_993_) with
+ | Some (s,(existT _ _stringappend_994_ _)) =>
+ let _stringappend_995_ :=
+ string_drop _stringappend_993_
+ (build_ex _stringappend_994_) in
and_boolM
(returnm ((string_startswith
- _stringappend_1466_ "w")
+ _stringappend_995_ "w")
: bool))
- (let _stringappend_1467_ :=
- string_drop _stringappend_1466_
- (string_length "w") in
+ (let _stringappend_996_ :=
+ string_drop _stringappend_995_
+ (build_ex (projT1 (string_length
+ "w"))) in
match (spc_matches_prefix
- _stringappend_1467_) with
+ _stringappend_996_) with
| Some
- (_stringappend_1468_,(existT _ _stringappend_1469_ _)) =>
- let _stringappend_1470_ :=
- string_drop _stringappend_1467_
- (build_ex
- _stringappend_1469_) in
+ (tt,(existT _ _stringappend_997_ _)) =>
+ let _stringappend_998_ :=
+ string_drop _stringappend_996_
+ (build_ex _stringappend_997_) in
match (reg_name_matches_prefix
- _stringappend_1470_) with
+ _stringappend_998_) with
| Some
- (_stringappend_1471_,(existT _ _stringappend_1472_ _)) =>
- let _stringappend_1473_ :=
- string_drop _stringappend_1470_
- (build_ex
- _stringappend_1472_) in
+ (rd,(existT _ _stringappend_999_ _)) =>
+ let _stringappend_1000_ :=
+ string_drop _stringappend_998_
+ (build_ex _stringappend_999_) in
sep_matches_prefix
- _stringappend_1473_ >>= fun w__408 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1000_ >>= fun w__408 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__408 with
| Some
- (_stringappend_1474_,(existT _ _stringappend_1475_ _)) =>
- let _stringappend_1476_ :=
+ (tt,(existT _ _stringappend_1001_ _)) =>
+ let _stringappend_1002_ :=
string_drop
- _stringappend_1473_
- (build_ex
- _stringappend_1475_) in
+ _stringappend_1000_
+ (build_ex _stringappend_1001_) in
match (reg_name_matches_prefix
- _stringappend_1476_) with
+ _stringappend_1002_) with
| Some
- (_stringappend_1477_,(existT _ _stringappend_1478_ _)) =>
- let _stringappend_1479_ :=
+ (rs1,(existT _ _stringappend_1003_ _)) =>
+ let _stringappend_1004_ :=
string_drop
- _stringappend_1476_
- (build_ex
- _stringappend_1478_) in
+ _stringappend_1002_
+ (build_ex _stringappend_1003_) in
sep_matches_prefix
- _stringappend_1479_ >>= fun w__409 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1004_ >>= fun w__409 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__409 with
| Some
- (_stringappend_1480_,(existT _ _stringappend_1481_ _)) =>
- let _stringappend_1482_ :=
+ (tt,(existT _ _stringappend_1005_ _)) =>
+ let _stringappend_1006_ :=
string_drop
- _stringappend_1479_
- (build_ex
- _stringappend_1481_) in
+ _stringappend_1004_
+ (build_ex _stringappend_1005_) in
if ((match (reg_name_matches_prefix
- _stringappend_1482_) with
+ _stringappend_1006_) with
| Some
- (_stringappend_1483_,(existT _ _stringappend_1484_ _)) =>
+ (rs2,(existT _ _stringappend_1007_ _)) =>
match (string_drop
- _stringappend_1482_
- (build_ex
- _stringappend_1484_)) with
+ _stringappend_1006_
+ (build_ex _stringappend_1007_)) with
| "" =>
true
| _ =>
@@ -14822,240 +15238,191 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__416 : bool =>
(if (w__416) then
- let _stringappend_1463_ :=
- string_drop _stringappend_1112_
- (string_length "div") in
- match (maybe_not_u_matches_prefix
- _stringappend_1463_) with
- | Some
- (_stringappend_1464_,(existT _ _stringappend_1465_ _)) =>
- returnm ((_stringappend_1464_,
- build_ex
- _stringappend_1465_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__418 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_1465_ _) :=
- w__418
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1466_ :=
- string_drop _stringappend_1463_
- (build_ex
- _stringappend_1465_) in
- let _stringappend_1467_ :=
- string_drop _stringappend_1466_
- (string_length "w") in
- match (spc_matches_prefix
- _stringappend_1467_) with
- | Some
- (_stringappend_1468_,(existT _ _stringappend_1469_ _)) =>
- returnm ((_stringappend_1468_,
- build_ex
- _stringappend_1469_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__420 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1469_ _) :=
- w__420
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1470_ :=
- string_drop _stringappend_1467_
- (build_ex
- _stringappend_1469_) in
- match (reg_name_matches_prefix
- _stringappend_1470_) with
- | Some
- (_stringappend_1471_,(existT _ _stringappend_1472_ _)) =>
- returnm ((_stringappend_1471_,
- build_ex
- _stringappend_1472_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__422 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1472_ _) :=
- w__422
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1473_ :=
- string_drop _stringappend_1470_
- (build_ex
- _stringappend_1472_) in
- sep_matches_prefix _stringappend_1473_ >>= fun w__423 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_993_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length "div"))) in
+ (match (maybe_not_u_matches_prefix
+ _stringappend_993_) with
+ | Some (s,(existT _ _stringappend_994_ _)) =>
+ returnm (s, build_ex _stringappend_994_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(s, existT _ _stringappend_994_ _) =>
+ let _stringappend_995_ :=
+ string_drop _stringappend_993_
+ (build_ex _stringappend_994_) in
+ let _stringappend_996_ :=
+ string_drop _stringappend_995_
+ (build_ex (projT1 (string_length "w"))) in
+ (match (spc_matches_prefix
+ _stringappend_996_) with
+ | Some
+ (tt,(existT _ _stringappend_997_ _)) =>
+ returnm (tt, build_ex _stringappend_997_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_997_ _) =>
+ let _stringappend_998_ :=
+ string_drop _stringappend_996_
+ (build_ex _stringappend_997_) in
+ (match (reg_name_matches_prefix
+ _stringappend_998_) with
+ | Some
+ (rd,(existT _ _stringappend_999_ _)) =>
+ returnm (rd, build_ex _stringappend_999_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_999_ _) =>
+ let _stringappend_1000_ :=
+ string_drop _stringappend_998_
+ (build_ex _stringappend_999_) in
+ sep_matches_prefix _stringappend_1000_ >>= fun w__423 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__423 with
- | Some
- (_stringappend_1474_,(existT _ _stringappend_1475_ _)) =>
- returnm ((_stringappend_1474_,
- build_ex
- _stringappend_1475_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__425 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1475_ _) :=
- w__425
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1476_ :=
- string_drop _stringappend_1473_
- (build_ex
- _stringappend_1475_) in
- match (reg_name_matches_prefix
- _stringappend_1476_) with
- | Some
- (_stringappend_1477_,(existT _ _stringappend_1478_ _)) =>
- returnm ((_stringappend_1477_,
- build_ex
- _stringappend_1478_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__427 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1478_ _) :=
- w__427
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1479_ :=
- string_drop _stringappend_1476_
- (build_ex
- _stringappend_1478_) in
- sep_matches_prefix _stringappend_1479_ >>= fun w__428 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__423 with
+ | Some
+ (tt,(existT _ _stringappend_1001_ _)) =>
+ returnm (tt, build_ex _stringappend_1001_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1001_ _) =>
+ let _stringappend_1002_ :=
+ string_drop _stringappend_1000_
+ (build_ex _stringappend_1001_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1002_) with
+ | Some
+ (rs1,(existT _ _stringappend_1003_ _)) =>
+ returnm (rs1, build_ex _stringappend_1003_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_1003_ _) =>
+ let _stringappend_1004_ :=
+ string_drop _stringappend_1002_
+ (build_ex _stringappend_1003_) in
+ sep_matches_prefix _stringappend_1004_ >>= fun w__428 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__428 with
- | Some
- (_stringappend_1480_,(existT _ _stringappend_1481_ _)) =>
- returnm ((_stringappend_1480_,
- build_ex
- _stringappend_1481_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__430 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1481_ _) :=
- w__430
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1482_ :=
- string_drop _stringappend_1479_
- (build_ex
- _stringappend_1481_) in
- match (reg_name_matches_prefix
- _stringappend_1482_) with
- | Some
- (_stringappend_1483_,(existT _ _stringappend_1484_ _)) =>
- returnm ((_stringappend_1483_,
- build_ex
- _stringappend_1484_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__432 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_1484_ _) :=
- w__432
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_1482_
- (build_ex
- _stringappend_1484_)) with
+ (match w__428 with
+ | Some
+ (tt,(existT _ _stringappend_1005_ _)) =>
+ returnm (tt, build_ex _stringappend_1005_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1005_ _) =>
+ let _stringappend_1006_ :=
+ string_drop _stringappend_1004_
+ (build_ex _stringappend_1005_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1006_) with
+ | Some
+ (rs2,(existT _ _stringappend_1007_ _)) =>
+ returnm (rs2, build_ex _stringappend_1007_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_1007_ _) =>
+ (match (string_drop _stringappend_1006_
+ (build_ex _stringappend_1007_)) with
| "" =>
- returnm ((DIVW (rs2,rs1,rd,s)) : ast)
+ returnm ((DIVW
+ (rs2, rs1, rd, s))
+ : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
and_boolM
(returnm ((string_startswith
- _stringappend_1112_ "rem")
+ _stringappend_756_ "rem")
: bool))
- (let _stringappend_1486_ :=
- string_drop _stringappend_1112_
- (string_length "rem") in
+ (let _stringappend_1009_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length
+ "rem"))) in
match (maybe_not_u_matches_prefix
- _stringappend_1486_) with
+ _stringappend_1009_) with
| Some
- (_stringappend_1487_,(existT _ _stringappend_1488_ _)) =>
- let _stringappend_1489_ :=
- string_drop _stringappend_1486_
- (build_ex
- _stringappend_1488_) in
+ (s,(existT _ _stringappend_1010_ _)) =>
+ let _stringappend_1011_ :=
+ string_drop _stringappend_1009_
+ (build_ex _stringappend_1010_) in
and_boolM
(returnm ((string_startswith
- _stringappend_1489_ "w")
+ _stringappend_1011_ "w")
: bool))
- (let _stringappend_1490_ :=
- string_drop _stringappend_1489_
- (string_length "w") in
+ (let _stringappend_1012_ :=
+ string_drop _stringappend_1011_
+ (build_ex (projT1 (string_length
+ "w"))) in
match (spc_matches_prefix
- _stringappend_1490_) with
+ _stringappend_1012_) with
| Some
- (_stringappend_1491_,(existT _ _stringappend_1492_ _)) =>
- let _stringappend_1493_ :=
- string_drop _stringappend_1490_
- (build_ex
- _stringappend_1492_) in
+ (tt,(existT _ _stringappend_1013_ _)) =>
+ let _stringappend_1014_ :=
+ string_drop _stringappend_1012_
+ (build_ex _stringappend_1013_) in
match (reg_name_matches_prefix
- _stringappend_1493_) with
+ _stringappend_1014_) with
| Some
- (_stringappend_1494_,(existT _ _stringappend_1495_ _)) =>
- let _stringappend_1496_ :=
+ (rd,(existT _ _stringappend_1015_ _)) =>
+ let _stringappend_1016_ :=
string_drop
- _stringappend_1493_
- (build_ex
- _stringappend_1495_) in
+ _stringappend_1014_
+ (build_ex _stringappend_1015_) in
sep_matches_prefix
- _stringappend_1496_ >>= fun w__435 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1016_ >>= fun w__435 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__435 with
| Some
- (_stringappend_1497_,(existT _ _stringappend_1498_ _)) =>
- let _stringappend_1499_ :=
+ (tt,(existT _ _stringappend_1017_ _)) =>
+ let _stringappend_1018_ :=
string_drop
- _stringappend_1496_
- (build_ex
- _stringappend_1498_) in
+ _stringappend_1016_
+ (build_ex _stringappend_1017_) in
match (reg_name_matches_prefix
- _stringappend_1499_) with
+ _stringappend_1018_) with
| Some
- (_stringappend_1500_,(existT _ _stringappend_1501_ _)) =>
- let _stringappend_1502_ :=
+ (rs1,(existT _ _stringappend_1019_ _)) =>
+ let _stringappend_1020_ :=
string_drop
- _stringappend_1499_
- (build_ex
- _stringappend_1501_) in
+ _stringappend_1018_
+ (build_ex _stringappend_1019_) in
sep_matches_prefix
- _stringappend_1502_ >>= fun w__436 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1020_ >>= fun w__436 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__436 with
| Some
- (_stringappend_1503_,(existT _ _stringappend_1504_ _)) =>
- let _stringappend_1505_ :=
+ (tt,(existT _ _stringappend_1021_ _)) =>
+ let _stringappend_1022_ :=
string_drop
- _stringappend_1502_
- (build_ex
- _stringappend_1504_) in
+ _stringappend_1020_
+ (build_ex _stringappend_1021_) in
if ((match (reg_name_matches_prefix
- _stringappend_1505_) with
+ _stringappend_1022_) with
| Some
- (_stringappend_1506_,(existT _ _stringappend_1507_ _)) =>
+ (rs2,(existT _ _stringappend_1023_ _)) =>
match (string_drop
- _stringappend_1505_
- (build_ex
- _stringappend_1507_)) with
+ _stringappend_1022_
+ (build_ex _stringappend_1023_)) with
| "" =>
true
| _ =>
@@ -15108,221 +15475,165 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__443 : bool =>
(if (w__443) then
- let _stringappend_1486_ :=
- string_drop _stringappend_1112_
- (string_length "rem") in
- match (maybe_not_u_matches_prefix
- _stringappend_1486_) with
- | Some
- (_stringappend_1487_,(existT _ _stringappend_1488_ _)) =>
- returnm ((_stringappend_1487_,
- build_ex
- _stringappend_1488_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__445 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_1488_ _) :=
- w__445
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1489_ :=
- string_drop _stringappend_1486_
- (build_ex
- _stringappend_1488_) in
- let _stringappend_1490_ :=
- string_drop _stringappend_1489_
- (string_length "w") in
- match (spc_matches_prefix
- _stringappend_1490_) with
- | Some
- (_stringappend_1491_,(existT _ _stringappend_1492_ _)) =>
- returnm ((_stringappend_1491_,
- build_ex
- _stringappend_1492_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__447 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1492_ _) :=
- w__447
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1493_ :=
- string_drop _stringappend_1490_
- (build_ex
- _stringappend_1492_) in
- match (reg_name_matches_prefix
- _stringappend_1493_) with
- | Some
- (_stringappend_1494_,(existT _ _stringappend_1495_ _)) =>
- returnm ((_stringappend_1494_,
- build_ex
- _stringappend_1495_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__449 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1495_ _) :=
- w__449
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1496_ :=
- string_drop _stringappend_1493_
- (build_ex
- _stringappend_1495_) in
- sep_matches_prefix _stringappend_1496_ >>= fun w__450 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_1009_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length
+ "rem"))) in
+ (match (maybe_not_u_matches_prefix
+ _stringappend_1009_) with
+ | Some
+ (s,(existT _ _stringappend_1010_ _)) =>
+ returnm (s, build_ex _stringappend_1010_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(s, existT _ _stringappend_1010_ _) =>
+ let _stringappend_1011_ :=
+ string_drop _stringappend_1009_
+ (build_ex _stringappend_1010_) in
+ let _stringappend_1012_ :=
+ string_drop _stringappend_1011_
+ (build_ex (projT1 (string_length "w"))) in
+ (match (spc_matches_prefix
+ _stringappend_1012_) with
+ | Some
+ (tt,(existT _ _stringappend_1013_ _)) =>
+ returnm (tt, build_ex _stringappend_1013_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1013_ _) =>
+ let _stringappend_1014_ :=
+ string_drop _stringappend_1012_
+ (build_ex _stringappend_1013_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1014_) with
+ | Some
+ (rd,(existT _ _stringappend_1015_ _)) =>
+ returnm (rd, build_ex _stringappend_1015_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_1015_ _) =>
+ let _stringappend_1016_ :=
+ string_drop _stringappend_1014_
+ (build_ex _stringappend_1015_) in
+ sep_matches_prefix _stringappend_1016_ >>= fun w__450 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__450 with
- | Some
- (_stringappend_1497_,(existT _ _stringappend_1498_ _)) =>
- returnm ((_stringappend_1497_,
- build_ex
- _stringappend_1498_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__452 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1498_ _) :=
- w__452
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1499_ :=
- string_drop _stringappend_1496_
- (build_ex
- _stringappend_1498_) in
- match (reg_name_matches_prefix
- _stringappend_1499_) with
- | Some
- (_stringappend_1500_,(existT _ _stringappend_1501_ _)) =>
- returnm ((_stringappend_1500_,
- build_ex
- _stringappend_1501_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__454 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1501_ _) :=
- w__454
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1502_ :=
- string_drop _stringappend_1499_
- (build_ex
- _stringappend_1501_) in
- sep_matches_prefix _stringappend_1502_ >>= fun w__455 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__450 with
+ | Some
+ (tt,(existT _ _stringappend_1017_ _)) =>
+ returnm (tt, build_ex _stringappend_1017_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1017_ _) =>
+ let _stringappend_1018_ :=
+ string_drop _stringappend_1016_
+ (build_ex _stringappend_1017_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1018_) with
+ | Some
+ (rs1,(existT _ _stringappend_1019_ _)) =>
+ returnm (rs1, build_ex _stringappend_1019_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_1019_ _) =>
+ let _stringappend_1020_ :=
+ string_drop _stringappend_1018_
+ (build_ex _stringappend_1019_) in
+ sep_matches_prefix _stringappend_1020_ >>= fun w__455 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__455 with
- | Some
- (_stringappend_1503_,(existT _ _stringappend_1504_ _)) =>
- returnm ((_stringappend_1503_,
- build_ex
- _stringappend_1504_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__457 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1504_ _) :=
- w__457
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_1505_ :=
- string_drop _stringappend_1502_
- (build_ex
- _stringappend_1504_) in
- match (reg_name_matches_prefix
- _stringappend_1505_) with
- | Some
- (_stringappend_1506_,(existT _ _stringappend_1507_ _)) =>
- returnm ((_stringappend_1506_,
- build_ex
- _stringappend_1507_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__459 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_1507_ _) :=
- w__459
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- (match (string_drop _stringappend_1505_
- (build_ex
- _stringappend_1507_)) with
+ (match w__455 with
+ | Some
+ (tt,(existT _ _stringappend_1021_ _)) =>
+ returnm (tt, build_ex _stringappend_1021_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1021_ _) =>
+ let _stringappend_1022_ :=
+ string_drop _stringappend_1020_
+ (build_ex _stringappend_1021_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1022_) with
+ | Some
+ (rs2,(existT _ _stringappend_1023_ _)) =>
+ returnm (rs2, build_ex _stringappend_1023_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_1023_ _) =>
+ (match (string_drop _stringappend_1022_
+ (build_ex _stringappend_1023_)) with
| "" =>
- returnm ((REMW (rs2,rs1,rd,s))
- : ast)
+ returnm ((REMW
+ (rs2, rs1, rd, s))
+ : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
and_boolM
(returnm ((string_startswith
- _stringappend_1112_
+ _stringappend_756_
"fence")
: bool))
- (let _stringappend_1509_ :=
- string_drop _stringappend_1112_
- (string_length "fence") in
+ (let _stringappend_1025_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length
+ "fence"))) in
match (spc_matches_prefix
- _stringappend_1509_) with
+ _stringappend_1025_) with
| Some
- (_stringappend_1510_,(existT _ _stringappend_1511_ _)) =>
- let _stringappend_1512_ :=
- string_drop _stringappend_1509_
- (build_ex
- _stringappend_1511_) in
+ (tt,(existT _ _stringappend_1026_ _)) =>
+ let _stringappend_1027_ :=
+ string_drop _stringappend_1025_
+ (build_ex _stringappend_1026_) in
fence_bits_matches_prefix
- _stringappend_1512_ >>= fun w__462 : option ((mword 4 * {n : Z & ArithFact (n >=
+ _stringappend_1027_ >>= fun w__462 : option ((mword 4 * {n : Z & ArithFact (n >=
0)})) =>
match w__462 with
| Some
- (_stringappend_1513_,(existT _ _stringappend_1514_ _)) =>
- let _stringappend_1515_ :=
+ (pred,(existT _ _stringappend_1028_ _)) =>
+ let _stringappend_1029_ :=
string_drop
- _stringappend_1512_
- (build_ex
- _stringappend_1514_) in
+ _stringappend_1027_
+ (build_ex _stringappend_1028_) in
sep_matches_prefix
- _stringappend_1515_ >>= fun w__463 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1029_ >>= fun w__463 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__463 with
| Some
- (_stringappend_1516_,(existT _ _stringappend_1517_ _)) =>
- let _stringappend_1518_ :=
+ (tt,(existT _ _stringappend_1030_ _)) =>
+ let _stringappend_1031_ :=
string_drop
- _stringappend_1515_
- (build_ex
- _stringappend_1517_) in
+ _stringappend_1029_
+ (build_ex _stringappend_1030_) in
fence_bits_matches_prefix
- _stringappend_1518_ >>= fun w__464 : option ((mword 4 * {n : Z & ArithFact (n >=
+ _stringappend_1031_ >>= fun w__464 : option ((mword 4 * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__464 with
| Some
- (_stringappend_1519_,(existT _ _stringappend_1520_ _)) =>
+ (succ,(existT _ _stringappend_1032_ _)) =>
match (string_drop
- _stringappend_1518_
- (build_ex
- _stringappend_1520_)) with
+ _stringappend_1031_
+ (build_ex _stringappend_1032_)) with
| "" =>
true
| _ =>
@@ -15351,173 +15662,137 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__468 : bool =>
(if (w__468) then
- let _stringappend_1509_ :=
- string_drop _stringappend_1112_
- (string_length "fence") in
- match (spc_matches_prefix
- _stringappend_1509_) with
- | Some
- (_stringappend_1510_,(existT _ _stringappend_1511_ _)) =>
- returnm ((_stringappend_1510_,
- build_ex
- _stringappend_1511_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__470 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1511_ _) :=
- w__470
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1512_ :=
- string_drop _stringappend_1509_
- (build_ex
- _stringappend_1511_) in
+ let _stringappend_1025_ :=
+ string_drop _stringappend_756_
+ (build_ex (projT1 (string_length
+ "fence"))) in
+ (match (spc_matches_prefix
+ _stringappend_1025_) with
+ | Some
+ (tt,(existT _ _stringappend_1026_ _)) =>
+ returnm (tt, build_ex _stringappend_1026_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1026_ _) =>
+ let _stringappend_1027_ :=
+ string_drop _stringappend_1025_
+ (build_ex _stringappend_1026_) in
fence_bits_matches_prefix
- _stringappend_1512_ >>= fun w__471 : option ((mword 4 * {n : Z & ArithFact (n >=
+ _stringappend_1027_ >>= fun w__471 : option ((mword 4 * {n : Z & ArithFact (n >=
0)})) =>
- match w__471 with
- | Some
- (_stringappend_1513_,(existT _ _stringappend_1514_ _)) =>
- returnm ((_stringappend_1513_,
- build_ex
- _stringappend_1514_)
- : (mword 4 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 4 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__473 : (mword 4 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(pred, existT _ _stringappend_1514_ _) :=
- w__473
- : (mword 4 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1515_ :=
- string_drop _stringappend_1512_
- (build_ex
- _stringappend_1514_) in
+ (match w__471 with
+ | Some
+ (pred,(existT _ _stringappend_1028_ _)) =>
+ returnm (pred, build_ex _stringappend_1028_)
+ | _ =>
+ exit tt
+ : M ((mword 4 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 4 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(pred, existT _ _stringappend_1028_ _) =>
+ let _stringappend_1029_ :=
+ string_drop _stringappend_1027_
+ (build_ex _stringappend_1028_) in
sep_matches_prefix
- _stringappend_1515_ >>= fun w__474 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1029_ >>= fun w__474 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__474 with
- | Some
- (_stringappend_1516_,(existT _ _stringappend_1517_ _)) =>
- returnm ((_stringappend_1516_,
- build_ex
- _stringappend_1517_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__476 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1517_ _) :=
- w__476
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1518_ :=
- string_drop _stringappend_1515_
- (build_ex
- _stringappend_1517_) in
+ (match w__474 with
+ | Some
+ (tt,(existT _ _stringappend_1030_ _)) =>
+ returnm (tt, build_ex _stringappend_1030_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1030_ _) =>
+ let _stringappend_1031_ :=
+ string_drop _stringappend_1029_
+ (build_ex _stringappend_1030_) in
fence_bits_matches_prefix
- _stringappend_1518_ >>= fun w__477 : option ((mword 4 * {n : Z & ArithFact (n >=
+ _stringappend_1031_ >>= fun w__477 : option ((mword 4 * {n : Z & ArithFact (n >=
0)})) =>
- match w__477 with
- | Some
- (_stringappend_1519_,(existT _ _stringappend_1520_ _)) =>
- returnm ((_stringappend_1519_,
- build_ex
- _stringappend_1520_)
- : (mword 4 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 4 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__479 : (mword 4 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(succ, existT _ _stringappend_1520_ _) :=
- w__479
- : (mword 4 * {n : Z & ArithFact (n >=
- 0)}) in
+ (match w__477 with
+ | Some
+ (succ,(existT _ _stringappend_1032_ _)) =>
+ returnm (succ, build_ex _stringappend_1032_)
+ | _ =>
+ exit tt
+ : M ((mword 4 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 4 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(succ, existT _ _stringappend_1032_ _) =>
(match (string_drop
- _stringappend_1518_
- (build_ex
- _stringappend_1520_)) with
+ _stringappend_1031_
+ (build_ex _stringappend_1032_)) with
| "" =>
- returnm ((FENCE (pred,succ))
- : ast)
+ returnm ((FENCE
+ (pred, succ))
+ : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (match _stringappend_1112_ with
+ (match _stringappend_756_ with
| "fence.i" =>
- returnm ((FENCEI tt) : ast)
+ returnm ((FENCEI (tt)) : ast )
| "ecall" =>
- returnm ((ECALL tt) : ast)
+ returnm ((ECALL (tt)) : ast )
| "mret" =>
- returnm ((MRET tt) : ast)
+ returnm ((MRET (tt)) : ast )
| "sret" =>
- returnm ((SRET tt) : ast)
+ returnm ((SRET (tt)) : ast )
| "ebreak" =>
- returnm ((EBREAK tt) : ast)
- | "wfi" => returnm ((WFI tt) : ast)
- | _stringappend_1112_ =>
+ returnm ((EBREAK (tt)) : ast )
+ | "wfi" =>
+ returnm ((WFI (tt)) : ast )
+ | _stringappend_756_ =>
and_boolM
(returnm ((string_startswith
- _stringappend_1112_
+ _stringappend_756_
"sfence.vma")
: bool))
- (let _stringappend_1522_ :=
+ (let _stringappend_1034_ :=
string_drop
- _stringappend_1112_
- (string_length "sfence.vma") in
+ _stringappend_756_
+ (build_ex (projT1 (string_length
+ "sfence.vma"))) in
match (spc_matches_prefix
- _stringappend_1522_) with
+ _stringappend_1034_) with
| Some
- (_stringappend_1523_,(existT _ _stringappend_1524_ _)) =>
- let _stringappend_1525_ :=
+ (tt,(existT _ _stringappend_1035_ _)) =>
+ let _stringappend_1036_ :=
string_drop
- _stringappend_1522_
- (build_ex
- _stringappend_1524_) in
+ _stringappend_1034_
+ (build_ex _stringappend_1035_) in
match (reg_name_matches_prefix
- _stringappend_1525_) with
+ _stringappend_1036_) with
| Some
- (_stringappend_1526_,(existT _ _stringappend_1527_ _)) =>
- let _stringappend_1528_ :=
+ (rs1,(existT _ _stringappend_1037_ _)) =>
+ let _stringappend_1038_ :=
string_drop
- _stringappend_1525_
- (build_ex
- _stringappend_1527_) in
+ _stringappend_1036_
+ (build_ex _stringappend_1037_) in
sep_matches_prefix
- _stringappend_1528_ >>= fun w__482 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1038_ >>= fun w__482 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__482 with
| Some
- (_stringappend_1529_,(existT _ _stringappend_1530_ _)) =>
- let _stringappend_1531_ :=
+ (tt,(existT _ _stringappend_1039_ _)) =>
+ let _stringappend_1040_ :=
string_drop
- _stringappend_1528_
- (build_ex
- _stringappend_1530_) in
+ _stringappend_1038_
+ (build_ex _stringappend_1039_) in
if ((match (reg_name_matches_prefix
- _stringappend_1531_) with
+ _stringappend_1040_) with
| Some
- (_stringappend_1532_,(existT _ _stringappend_1533_ _)) =>
+ (rs2,(existT _ _stringappend_1041_ _)) =>
match (string_drop
- _stringappend_1531_
- (build_ex
- _stringappend_1533_)) with
+ _stringappend_1040_
+ (build_ex _stringappend_1041_)) with
| "" =>
true
| _ =>
@@ -15550,189 +15825,148 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__485 : bool =>
(if (w__485) then
- let _stringappend_1522_ :=
+ let _stringappend_1034_ :=
string_drop
- _stringappend_1112_
- (string_length
- "sfence.vma") in
- match (spc_matches_prefix
- _stringappend_1522_) with
- | Some
- (_stringappend_1523_,(existT _ _stringappend_1524_ _)) =>
- returnm ((_stringappend_1523_,
- build_ex
- _stringappend_1524_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__487 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1524_ _) :=
- w__487
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1525_ :=
+ _stringappend_756_
+ (build_ex (projT1 (string_length
+ "sfence.vma"))) in
+ (match (spc_matches_prefix
+ _stringappend_1034_) with
+ | Some
+ (tt,(existT _ _stringappend_1035_ _)) =>
+ returnm (tt, build_ex _stringappend_1035_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1035_ _) =>
+ let _stringappend_1036_ :=
string_drop
- _stringappend_1522_
- (build_ex
- _stringappend_1524_) in
- match (reg_name_matches_prefix
- _stringappend_1525_) with
- | Some
- (_stringappend_1526_,(existT _ _stringappend_1527_ _)) =>
- returnm ((_stringappend_1526_,
- build_ex
- _stringappend_1527_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__489 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1527_ _) :=
- w__489
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1528_ :=
+ _stringappend_1034_
+ (build_ex _stringappend_1035_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1036_) with
+ | Some
+ (rs1,(existT _ _stringappend_1037_ _)) =>
+ returnm (rs1, build_ex _stringappend_1037_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_1037_ _) =>
+ let _stringappend_1038_ :=
string_drop
- _stringappend_1525_
- (build_ex
- _stringappend_1527_) in
+ _stringappend_1036_
+ (build_ex _stringappend_1037_) in
sep_matches_prefix
- _stringappend_1528_ >>= fun w__490 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1038_ >>= fun w__490 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__490 with
- | Some
- (_stringappend_1529_,(existT _ _stringappend_1530_ _)) =>
- returnm ((_stringappend_1529_,
- build_ex
- _stringappend_1530_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__492 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1530_ _) :=
- w__492
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1531_ :=
+ (match w__490 with
+ | Some
+ (tt,(existT _ _stringappend_1039_ _)) =>
+ returnm (tt, build_ex _stringappend_1039_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1039_ _) =>
+ let _stringappend_1040_ :=
string_drop
- _stringappend_1528_
- (build_ex
- _stringappend_1530_) in
- match (reg_name_matches_prefix
- _stringappend_1531_) with
- | Some
- (_stringappend_1532_,(existT _ _stringappend_1533_ _)) =>
- returnm ((_stringappend_1532_,
- build_ex
- _stringappend_1533_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__494 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_1533_ _) :=
- w__494
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_1038_
+ (build_ex _stringappend_1039_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1040_) with
+ | Some
+ (rs2,(existT _ _stringappend_1041_ _)) =>
+ returnm (rs2, build_ex _stringappend_1041_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_1041_ _) =>
(match (string_drop
- _stringappend_1531_
- (build_ex
- _stringappend_1533_)) with
+ _stringappend_1040_
+ (build_ex _stringappend_1041_)) with
| "" =>
- returnm ((SFENCE_VMA (rs1,rs2))
- : ast)
+ returnm ((SFENCE_VMA
+ (rs1, rs2))
+ : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
and_boolM
(returnm ((string_startswith
- _stringappend_1112_
+ _stringappend_756_
"lr.")
: bool))
- (let _stringappend_1535_ :=
+ (let _stringappend_1043_ :=
string_drop
- _stringappend_1112_
- (string_length "lr.") in
+ _stringappend_756_
+ (build_ex (projT1 (string_length
+ "lr."))) in
match (maybe_aq_matches_prefix
- _stringappend_1535_) with
+ _stringappend_1043_) with
| Some
- (_stringappend_1536_,(existT _ _stringappend_1537_ _)) =>
- let _stringappend_1538_ :=
+ (aq,(existT _ _stringappend_1044_ _)) =>
+ let _stringappend_1045_ :=
string_drop
- _stringappend_1535_
- (build_ex
- _stringappend_1537_) in
+ _stringappend_1043_
+ (build_ex _stringappend_1044_) in
match (maybe_rl_matches_prefix
- _stringappend_1538_) with
+ _stringappend_1045_) with
| Some
- (_stringappend_1539_,(existT _ _stringappend_1540_ _)) =>
- let _stringappend_1541_ :=
+ (rl,(existT _ _stringappend_1046_ _)) =>
+ let _stringappend_1047_ :=
string_drop
- _stringappend_1538_
- (build_ex
- _stringappend_1540_) in
+ _stringappend_1045_
+ (build_ex _stringappend_1046_) in
match (size_mnemonic_matches_prefix
- _stringappend_1541_) with
+ _stringappend_1047_) with
| Some
- (_stringappend_1542_,(existT _ _stringappend_1543_ _)) =>
- let _stringappend_1544_ :=
+ (size,(existT _ _stringappend_1048_ _)) =>
+ let _stringappend_1049_ :=
string_drop
- _stringappend_1541_
- (build_ex
- _stringappend_1543_) in
+ _stringappend_1047_
+ (build_ex _stringappend_1048_) in
match (spc_matches_prefix
- _stringappend_1544_) with
+ _stringappend_1049_) with
| Some
- (_stringappend_1545_,(existT _ _stringappend_1546_ _)) =>
- let _stringappend_1547_ :=
+ (tt,(existT _ _stringappend_1050_ _)) =>
+ let _stringappend_1051_ :=
string_drop
- _stringappend_1544_
- (build_ex
- _stringappend_1546_) in
+ _stringappend_1049_
+ (build_ex _stringappend_1050_) in
match (reg_name_matches_prefix
- _stringappend_1547_) with
+ _stringappend_1051_) with
| Some
- (_stringappend_1548_,(existT _ _stringappend_1549_ _)) =>
- let _stringappend_1550_ :=
+ (rd,(existT _ _stringappend_1052_ _)) =>
+ let _stringappend_1053_ :=
string_drop
- _stringappend_1547_
- (build_ex
- _stringappend_1549_) in
+ _stringappend_1051_
+ (build_ex _stringappend_1052_) in
sep_matches_prefix
- _stringappend_1550_ >>= fun w__497 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1053_ >>= fun w__497 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__497 with
| Some
- (_stringappend_1551_,(existT _ _stringappend_1552_ _)) =>
- let _stringappend_1553_ :=
+ (tt,(existT _ _stringappend_1054_ _)) =>
+ let _stringappend_1055_ :=
string_drop
- _stringappend_1550_
- (build_ex
- _stringappend_1552_) in
+ _stringappend_1053_
+ (build_ex _stringappend_1054_) in
if
((match (reg_name_matches_prefix
- _stringappend_1553_) with
+ _stringappend_1055_) with
| Some
- (_stringappend_1554_,(existT _ _stringappend_1555_ _)) =>
+ (rs1,(existT _ _stringappend_1056_ _)) =>
match (string_drop
- _stringappend_1553_
- (build_ex
- _stringappend_1555_)) with
+ _stringappend_1055_
+ (build_ex _stringappend_1056_)) with
| "" =>
true
| _ =>
@@ -15799,281 +16033,214 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__503 : bool =>
(if (w__503) then
- let _stringappend_1535_ :=
+ let _stringappend_1043_ :=
string_drop
- _stringappend_1112_
- (string_length "lr.") in
- match (maybe_aq_matches_prefix
- _stringappend_1535_) with
- | Some
- (_stringappend_1536_,(existT _ _stringappend_1537_ _)) =>
- returnm ((_stringappend_1536_,
- build_ex
- _stringappend_1537_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__505 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(aq, existT _ _stringappend_1537_ _) :=
- w__505
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1538_ :=
+ _stringappend_756_
+ (build_ex (projT1 (string_length
+ "lr."))) in
+ (match (maybe_aq_matches_prefix
+ _stringappend_1043_) with
+ | Some
+ (aq,(existT _ _stringappend_1044_ _)) =>
+ returnm (aq, build_ex _stringappend_1044_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(aq, existT _ _stringappend_1044_ _) =>
+ let _stringappend_1045_ :=
string_drop
- _stringappend_1535_
- (build_ex
- _stringappend_1537_) in
- match (maybe_rl_matches_prefix
- _stringappend_1538_) with
- | Some
- (_stringappend_1539_,(existT _ _stringappend_1540_ _)) =>
- returnm ((_stringappend_1539_,
- build_ex
- _stringappend_1540_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__507 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rl, existT _ _stringappend_1540_ _) :=
- w__507
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1541_ :=
+ _stringappend_1043_
+ (build_ex _stringappend_1044_) in
+ (match (maybe_rl_matches_prefix
+ _stringappend_1045_) with
+ | Some
+ (rl,(existT _ _stringappend_1046_ _)) =>
+ returnm (rl, build_ex _stringappend_1046_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rl, existT _ _stringappend_1046_ _) =>
+ let _stringappend_1047_ :=
string_drop
- _stringappend_1538_
- (build_ex
- _stringappend_1540_) in
- match (size_mnemonic_matches_prefix
- _stringappend_1541_) with
- | Some
- (_stringappend_1542_,(existT _ _stringappend_1543_ _)) =>
- returnm ((_stringappend_1542_,
- build_ex
- _stringappend_1543_)
- : (word_width * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((word_width * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__509 : (word_width * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(size, existT _ _stringappend_1543_ _) :=
- w__509
- : (word_width * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1544_ :=
+ _stringappend_1045_
+ (build_ex _stringappend_1046_) in
+ (match (size_mnemonic_matches_prefix
+ _stringappend_1047_) with
+ | Some
+ (size,(existT _ _stringappend_1048_ _)) =>
+ returnm (size, build_ex _stringappend_1048_)
+ | _ =>
+ exit tt
+ : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(size, existT _ _stringappend_1048_ _) =>
+ let _stringappend_1049_ :=
string_drop
- _stringappend_1541_
- (build_ex
- _stringappend_1543_) in
- match (spc_matches_prefix
- _stringappend_1544_) with
- | Some
- (_stringappend_1545_,(existT _ _stringappend_1546_ _)) =>
- returnm ((_stringappend_1545_,
- build_ex
- _stringappend_1546_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__511 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1546_ _) :=
- w__511
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1547_ :=
+ _stringappend_1047_
+ (build_ex _stringappend_1048_) in
+ (match (spc_matches_prefix
+ _stringappend_1049_) with
+ | Some
+ (tt,(existT _ _stringappend_1050_ _)) =>
+ returnm (tt, build_ex _stringappend_1050_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1050_ _) =>
+ let _stringappend_1051_ :=
string_drop
- _stringappend_1544_
- (build_ex
- _stringappend_1546_) in
- match (reg_name_matches_prefix
- _stringappend_1547_) with
- | Some
- (_stringappend_1548_,(existT _ _stringappend_1549_ _)) =>
- returnm ((_stringappend_1548_,
- build_ex
- _stringappend_1549_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__513 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1549_ _) :=
- w__513
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1550_ :=
+ _stringappend_1049_
+ (build_ex _stringappend_1050_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1051_) with
+ | Some
+ (rd,(existT _ _stringappend_1052_ _)) =>
+ returnm (rd, build_ex _stringappend_1052_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_1052_ _) =>
+ let _stringappend_1053_ :=
string_drop
- _stringappend_1547_
- (build_ex
- _stringappend_1549_) in
+ _stringappend_1051_
+ (build_ex _stringappend_1052_) in
sep_matches_prefix
- _stringappend_1550_ >>= fun w__514 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1053_ >>= fun w__514 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__514 with
- | Some
- (_stringappend_1551_,(existT _ _stringappend_1552_ _)) =>
- returnm ((_stringappend_1551_,
- build_ex
- _stringappend_1552_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__516 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1552_ _) :=
- w__516
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1553_ :=
+ (match w__514 with
+ | Some
+ (tt,(existT _ _stringappend_1054_ _)) =>
+ returnm (tt, build_ex _stringappend_1054_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1054_ _) =>
+ let _stringappend_1055_ :=
string_drop
- _stringappend_1550_
- (build_ex
- _stringappend_1552_) in
- match (reg_name_matches_prefix
- _stringappend_1553_) with
- | Some
- (_stringappend_1554_,(existT _ _stringappend_1555_ _)) =>
- returnm ((_stringappend_1554_,
- build_ex
- _stringappend_1555_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__518 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1555_ _) :=
- w__518
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_1053_
+ (build_ex _stringappend_1054_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1055_) with
+ | Some
+ (rs1,(existT _ _stringappend_1056_ _)) =>
+ returnm (rs1, build_ex _stringappend_1056_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_1056_ _) =>
(match (string_drop
- _stringappend_1553_
- (build_ex
- _stringappend_1555_)) with
+ _stringappend_1055_
+ (build_ex _stringappend_1056_)) with
| "" =>
- returnm ((LOADRES (aq,rl,rs1,size,rd))
- : ast)
+ returnm ((LOADRES
+ (aq, rl,
+ rs1, size,
+ rd))
+ : ast )
| _ => exit tt : M (ast)
end)
: M (ast)
else
and_boolM
(returnm ((string_startswith
- _stringappend_1112_
+ _stringappend_756_
"sc.")
: bool))
- (let _stringappend_1557_ :=
+ (let _stringappend_1058_ :=
string_drop
- _stringappend_1112_
- (string_length "sc.") in
+ _stringappend_756_
+ (build_ex (projT1 (string_length
+ "sc."))) in
match (maybe_aq_matches_prefix
- _stringappend_1557_) with
+ _stringappend_1058_) with
| Some
- (_stringappend_1558_,(existT _ _stringappend_1559_ _)) =>
- let _stringappend_1560_ :=
+ (aq,(existT _ _stringappend_1059_ _)) =>
+ let _stringappend_1060_ :=
string_drop
- _stringappend_1557_
- (build_ex
- _stringappend_1559_) in
+ _stringappend_1058_
+ (build_ex _stringappend_1059_) in
match (maybe_rl_matches_prefix
- _stringappend_1560_) with
+ _stringappend_1060_) with
| Some
- (_stringappend_1561_,(existT _ _stringappend_1562_ _)) =>
- let _stringappend_1563_ :=
+ (rl,(existT _ _stringappend_1061_ _)) =>
+ let _stringappend_1062_ :=
string_drop
- _stringappend_1560_
- (build_ex
- _stringappend_1562_) in
+ _stringappend_1060_
+ (build_ex _stringappend_1061_) in
match (size_mnemonic_matches_prefix
- _stringappend_1563_) with
+ _stringappend_1062_) with
| Some
- (_stringappend_1564_,(existT _ _stringappend_1565_ _)) =>
- let _stringappend_1566_ :=
+ (size,(existT _ _stringappend_1063_ _)) =>
+ let _stringappend_1064_ :=
string_drop
- _stringappend_1563_
- (build_ex
- _stringappend_1565_) in
+ _stringappend_1062_
+ (build_ex _stringappend_1063_) in
match (spc_matches_prefix
- _stringappend_1566_) with
+ _stringappend_1064_) with
| Some
- (_stringappend_1567_,(existT _ _stringappend_1568_ _)) =>
- let _stringappend_1569_ :=
+ (tt,(existT _ _stringappend_1065_ _)) =>
+ let _stringappend_1066_ :=
string_drop
- _stringappend_1566_
- (build_ex
- _stringappend_1568_) in
+ _stringappend_1064_
+ (build_ex _stringappend_1065_) in
match (reg_name_matches_prefix
- _stringappend_1569_) with
+ _stringappend_1066_) with
| Some
- (_stringappend_1570_,(existT _ _stringappend_1571_ _)) =>
- let _stringappend_1572_ :=
+ (rd,(existT _ _stringappend_1067_ _)) =>
+ let _stringappend_1068_ :=
string_drop
- _stringappend_1569_
- (build_ex
- _stringappend_1571_) in
+ _stringappend_1066_
+ (build_ex _stringappend_1067_) in
sep_matches_prefix
- _stringappend_1572_ >>= fun w__521 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1068_ >>= fun w__521 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__521 with
| Some
- (_stringappend_1573_,(existT _ _stringappend_1574_ _)) =>
- let _stringappend_1575_ :=
+ (tt,(existT _ _stringappend_1069_ _)) =>
+ let _stringappend_1070_ :=
string_drop
- _stringappend_1572_
- (build_ex
- _stringappend_1574_) in
+ _stringappend_1068_
+ (build_ex _stringappend_1069_) in
match (reg_name_matches_prefix
- _stringappend_1575_) with
+ _stringappend_1070_) with
| Some
- (_stringappend_1576_,(existT _ _stringappend_1577_ _)) =>
- let _stringappend_1578_ :=
+ (rs1,(existT _ _stringappend_1071_ _)) =>
+ let _stringappend_1072_ :=
string_drop
- _stringappend_1575_
- (build_ex
- _stringappend_1577_) in
+ _stringappend_1070_
+ (build_ex _stringappend_1071_) in
sep_matches_prefix
- _stringappend_1578_ >>= fun w__522 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1072_ >>= fun w__522 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if
((match w__522 with
| Some
- (_stringappend_1579_,(existT _ _stringappend_1580_ _)) =>
- let _stringappend_1581_ :=
+ (tt,(existT _ _stringappend_1073_ _)) =>
+ let _stringappend_1074_ :=
string_drop
- _stringappend_1578_
- (build_ex
- _stringappend_1580_) in
+ _stringappend_1072_
+ (build_ex _stringappend_1073_) in
if
((match (reg_name_matches_prefix
- _stringappend_1581_) with
+ _stringappend_1074_) with
| Some
- (_stringappend_1582_,(existT _ _stringappend_1583_ _)) =>
+ (rs2,(existT _ _stringappend_1075_ _)) =>
match (string_drop
- _stringappend_1581_
- (build_ex
- _stringappend_1583_)) with
+ _stringappend_1074_
+ (build_ex _stringappend_1075_)) with
| "" =>
true
| _ =>
@@ -16165,342 +16332,255 @@ Definition assembly_backwards (arg_ : string)
else false)
: bool)) >>= fun w__530 : bool =>
(if (w__530) then
- let _stringappend_1557_ :=
+ let _stringappend_1058_ :=
string_drop
- _stringappend_1112_
- (string_length "sc.") in
- match (maybe_aq_matches_prefix
- _stringappend_1557_) with
- | Some
- (_stringappend_1558_,(existT _ _stringappend_1559_ _)) =>
- returnm ((_stringappend_1558_,
- build_ex
- _stringappend_1559_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__532 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(aq, existT _ _stringappend_1559_ _) :=
- w__532
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1560_ :=
+ _stringappend_756_
+ (build_ex (projT1 (string_length
+ "sc."))) in
+ (match (maybe_aq_matches_prefix
+ _stringappend_1058_) with
+ | Some
+ (aq,(existT _ _stringappend_1059_ _)) =>
+ returnm (aq, build_ex _stringappend_1059_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(aq, existT _ _stringappend_1059_ _) =>
+ let _stringappend_1060_ :=
string_drop
- _stringappend_1557_
- (build_ex
- _stringappend_1559_) in
- match (maybe_rl_matches_prefix
- _stringappend_1560_) with
- | Some
- (_stringappend_1561_,(existT _ _stringappend_1562_ _)) =>
- returnm ((_stringappend_1561_,
- build_ex
- _stringappend_1562_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__534 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rl, existT _ _stringappend_1562_ _) :=
- w__534
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1563_ :=
+ _stringappend_1058_
+ (build_ex _stringappend_1059_) in
+ (match (maybe_rl_matches_prefix
+ _stringappend_1060_) with
+ | Some
+ (rl,(existT _ _stringappend_1061_ _)) =>
+ returnm (rl, build_ex _stringappend_1061_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rl, existT _ _stringappend_1061_ _) =>
+ let _stringappend_1062_ :=
string_drop
- _stringappend_1560_
- (build_ex
- _stringappend_1562_) in
- match (size_mnemonic_matches_prefix
- _stringappend_1563_) with
- | Some
- (_stringappend_1564_,(existT _ _stringappend_1565_ _)) =>
- returnm ((_stringappend_1564_,
- build_ex
- _stringappend_1565_)
- : (word_width * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((word_width * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__536 : (word_width * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(size, existT _ _stringappend_1565_ _) :=
- w__536
- : (word_width * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1566_ :=
+ _stringappend_1060_
+ (build_ex _stringappend_1061_) in
+ (match (size_mnemonic_matches_prefix
+ _stringappend_1062_) with
+ | Some
+ (size,(existT _ _stringappend_1063_ _)) =>
+ returnm (size, build_ex _stringappend_1063_)
+ | _ =>
+ exit tt
+ : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(size, existT _ _stringappend_1063_ _) =>
+ let _stringappend_1064_ :=
string_drop
- _stringappend_1563_
- (build_ex
- _stringappend_1565_) in
- match (spc_matches_prefix
- _stringappend_1566_) with
- | Some
- (_stringappend_1567_,(existT _ _stringappend_1568_ _)) =>
- returnm ((_stringappend_1567_,
- build_ex
- _stringappend_1568_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__538 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1568_ _) :=
- w__538
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1569_ :=
+ _stringappend_1062_
+ (build_ex _stringappend_1063_) in
+ (match (spc_matches_prefix
+ _stringappend_1064_) with
+ | Some
+ (tt,(existT _ _stringappend_1065_ _)) =>
+ returnm (tt, build_ex _stringappend_1065_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1065_ _) =>
+ let _stringappend_1066_ :=
string_drop
- _stringappend_1566_
- (build_ex
- _stringappend_1568_) in
- match (reg_name_matches_prefix
- _stringappend_1569_) with
- | Some
- (_stringappend_1570_,(existT _ _stringappend_1571_ _)) =>
- returnm ((_stringappend_1570_,
- build_ex
- _stringappend_1571_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__540 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1571_ _) :=
- w__540
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1572_ :=
+ _stringappend_1064_
+ (build_ex _stringappend_1065_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1066_) with
+ | Some
+ (rd,(existT _ _stringappend_1067_ _)) =>
+ returnm (rd, build_ex _stringappend_1067_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_1067_ _) =>
+ let _stringappend_1068_ :=
string_drop
- _stringappend_1569_
- (build_ex
- _stringappend_1571_) in
+ _stringappend_1066_
+ (build_ex _stringappend_1067_) in
sep_matches_prefix
- _stringappend_1572_ >>= fun w__541 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1068_ >>= fun w__541 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__541 with
- | Some
- (_stringappend_1573_,(existT _ _stringappend_1574_ _)) =>
- returnm ((_stringappend_1573_,
- build_ex
- _stringappend_1574_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__543 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1574_ _) :=
- w__543
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1575_ :=
+ (match w__541 with
+ | Some
+ (tt,(existT _ _stringappend_1069_ _)) =>
+ returnm (tt, build_ex _stringappend_1069_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1069_ _) =>
+ let _stringappend_1070_ :=
string_drop
- _stringappend_1572_
- (build_ex
- _stringappend_1574_) in
- match (reg_name_matches_prefix
- _stringappend_1575_) with
- | Some
- (_stringappend_1576_,(existT _ _stringappend_1577_ _)) =>
- returnm ((_stringappend_1576_,
- build_ex
- _stringappend_1577_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__545 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1577_ _) :=
- w__545
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1578_ :=
+ _stringappend_1068_
+ (build_ex _stringappend_1069_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1070_) with
+ | Some
+ (rs1,(existT _ _stringappend_1071_ _)) =>
+ returnm (rs1, build_ex _stringappend_1071_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_1071_ _) =>
+ let _stringappend_1072_ :=
string_drop
- _stringappend_1575_
- (build_ex
- _stringappend_1577_) in
+ _stringappend_1070_
+ (build_ex _stringappend_1071_) in
sep_matches_prefix
- _stringappend_1578_ >>= fun w__546 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1072_ >>= fun w__546 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__546 with
- | Some
- (_stringappend_1579_,(existT _ _stringappend_1580_ _)) =>
- returnm ((_stringappend_1579_,
- build_ex
- _stringappend_1580_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__548 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1580_ _) :=
- w__548
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1581_ :=
+ (match w__546 with
+ | Some
+ (tt,(existT _ _stringappend_1073_ _)) =>
+ returnm (tt, build_ex _stringappend_1073_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1073_ _) =>
+ let _stringappend_1074_ :=
string_drop
- _stringappend_1578_
- (build_ex
- _stringappend_1580_) in
- match (reg_name_matches_prefix
- _stringappend_1581_) with
- | Some
- (_stringappend_1582_,(existT _ _stringappend_1583_ _)) =>
- returnm ((_stringappend_1582_,
- build_ex
- _stringappend_1583_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__550 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_1583_ _) :=
- w__550
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_1072_
+ (build_ex _stringappend_1073_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1074_) with
+ | Some
+ (rs2,(existT _ _stringappend_1075_ _)) =>
+ returnm (rs2, build_ex _stringappend_1075_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_1075_ _) =>
(match (string_drop
- _stringappend_1581_
- (build_ex
- _stringappend_1583_)) with
+ _stringappend_1074_
+ (build_ex _stringappend_1075_)) with
| "" =>
- returnm ((STORECON (aq,rl,rs2,rs1,size,rd))
- : ast)
+ returnm ((STORECON
+ (aq, rl,
+ rs2, rs1,
+ size, rd))
+ : ast )
| _ =>
exit tt : M (ast)
end)
: M (ast)
else
match (amo_mnemonic_matches_prefix
- _stringappend_1112_) with
+ _stringappend_756_) with
| Some
- (_stringappend_1585_,(existT _ _stringappend_1586_ _)) =>
- let _stringappend_1587_ :=
+ (op,(existT _ _stringappend_1077_ _)) =>
+ let _stringappend_1078_ :=
string_drop
- _stringappend_1112_
- (build_ex
- _stringappend_1586_) in
+ _stringappend_756_
+ (build_ex _stringappend_1077_) in
and_boolM
(returnm ((string_startswith
- _stringappend_1587_
+ _stringappend_1078_
".")
: bool))
- (let _stringappend_1588_ :=
+ (let _stringappend_1079_ :=
string_drop
- _stringappend_1587_
- (string_length
- ".") in
+ _stringappend_1078_
+ (build_ex (projT1 (string_length
+ "."))) in
match (size_mnemonic_matches_prefix
- _stringappend_1588_) with
+ _stringappend_1079_) with
| Some
- (_stringappend_1589_,(existT _ _stringappend_1590_ _)) =>
- let _stringappend_1591_ :=
+ (width,(existT _ _stringappend_1080_ _)) =>
+ let _stringappend_1081_ :=
string_drop
- _stringappend_1588_
- (build_ex
- _stringappend_1590_) in
+ _stringappend_1079_
+ (build_ex _stringappend_1080_) in
match (maybe_aq_matches_prefix
- _stringappend_1591_) with
+ _stringappend_1081_) with
| Some
- (_stringappend_1592_,(existT _ _stringappend_1593_ _)) =>
- let _stringappend_1594_ :=
+ (aq,(existT _ _stringappend_1082_ _)) =>
+ let _stringappend_1083_ :=
string_drop
- _stringappend_1591_
- (build_ex
- _stringappend_1593_) in
+ _stringappend_1081_
+ (build_ex _stringappend_1082_) in
match (maybe_rl_matches_prefix
- _stringappend_1594_) with
+ _stringappend_1083_) with
| Some
- (_stringappend_1595_,(existT _ _stringappend_1596_ _)) =>
- let _stringappend_1597_ :=
+ (rl,(existT _ _stringappend_1084_ _)) =>
+ let _stringappend_1085_ :=
string_drop
- _stringappend_1594_
- (build_ex
- _stringappend_1596_) in
+ _stringappend_1083_
+ (build_ex _stringappend_1084_) in
match (spc_matches_prefix
- _stringappend_1597_) with
+ _stringappend_1085_) with
| Some
- (_stringappend_1598_,(existT _ _stringappend_1599_ _)) =>
- let _stringappend_1600_ :=
+ (tt,(existT _ _stringappend_1086_ _)) =>
+ let _stringappend_1087_ :=
string_drop
- _stringappend_1597_
- (build_ex
- _stringappend_1599_) in
+ _stringappend_1085_
+ (build_ex _stringappend_1086_) in
match (reg_name_matches_prefix
- _stringappend_1600_) with
+ _stringappend_1087_) with
| Some
- (_stringappend_1601_,(existT _ _stringappend_1602_ _)) =>
- let _stringappend_1603_ :=
+ (rd,(existT _ _stringappend_1088_ _)) =>
+ let _stringappend_1089_ :=
string_drop
- _stringappend_1600_
- (build_ex
- _stringappend_1602_) in
+ _stringappend_1087_
+ (build_ex _stringappend_1088_) in
sep_matches_prefix
- _stringappend_1603_ >>= fun w__553 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1089_ >>= fun w__553 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__553 with
| Some
- (_stringappend_1604_,(existT _ _stringappend_1605_ _)) =>
- let _stringappend_1606_ :=
+ (tt,(existT _ _stringappend_1090_ _)) =>
+ let _stringappend_1091_ :=
string_drop
- _stringappend_1603_
- (build_ex
- _stringappend_1605_) in
+ _stringappend_1089_
+ (build_ex _stringappend_1090_) in
match (reg_name_matches_prefix
- _stringappend_1606_) with
+ _stringappend_1091_) with
| Some
- (_stringappend_1607_,(existT _ _stringappend_1608_ _)) =>
- let _stringappend_1609_ :=
+ (rs1,(existT _ _stringappend_1092_ _)) =>
+ let _stringappend_1093_ :=
string_drop
- _stringappend_1606_
- (build_ex
- _stringappend_1608_) in
+ _stringappend_1091_
+ (build_ex _stringappend_1092_) in
sep_matches_prefix
- _stringappend_1609_ >>= fun w__554 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1093_ >>= fun w__554 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if
((match w__554 with
| Some
- (_stringappend_1610_,(existT _ _stringappend_1611_ _)) =>
- let _stringappend_1612_ :=
+ (tt,(existT _ _stringappend_1094_ _)) =>
+ let _stringappend_1095_ :=
string_drop
- _stringappend_1609_
- (build_ex
- _stringappend_1611_) in
+ _stringappend_1093_
+ (build_ex _stringappend_1094_) in
if
((match (reg_name_matches_prefix
- _stringappend_1612_) with
+ _stringappend_1095_) with
| Some
- (_stringappend_1613_,(existT _ _stringappend_1614_ _)) =>
+ (rs2,(existT _ _stringappend_1096_ _)) =>
match (string_drop
- _stringappend_1612_
- (build_ex
- _stringappend_1614_)) with
+ _stringappend_1095_
+ (build_ex _stringappend_1096_)) with
| "" =>
true
| _ =>
@@ -16605,257 +16685,174 @@ Definition assembly_backwards (arg_ : string)
: bool)
end >>= fun w__563 : bool =>
(if (w__563) then
- match (amo_mnemonic_matches_prefix
- _stringappend_1112_) with
- | Some
- (_stringappend_1585_,(existT _ _stringappend_1586_ _)) =>
- returnm ((_stringappend_1585_,
- build_ex
- _stringappend_1586_)
- : (amoop * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((amoop * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__565 : (amoop * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(op, existT _ _stringappend_1586_ _) :=
- w__565
- : (amoop * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1587_ :=
+ (match (amo_mnemonic_matches_prefix
+ _stringappend_756_) with
+ | Some
+ (op,(existT _ _stringappend_1077_ _)) =>
+ returnm (op, build_ex _stringappend_1077_)
+ | _ =>
+ exit tt
+ : M ((amoop * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((amoop * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(op, existT _ _stringappend_1077_ _) =>
+ let _stringappend_1078_ :=
string_drop
- _stringappend_1112_
- (build_ex
- _stringappend_1586_) in
- let _stringappend_1588_ :=
+ _stringappend_756_
+ (build_ex _stringappend_1077_) in
+ let _stringappend_1079_ :=
string_drop
- _stringappend_1587_
- (string_length
- ".") in
- match (size_mnemonic_matches_prefix
- _stringappend_1588_) with
- | Some
- (_stringappend_1589_,(existT _ _stringappend_1590_ _)) =>
- returnm ((_stringappend_1589_,
- build_ex
- _stringappend_1590_)
- : (word_width * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((word_width * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__567 : (word_width * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(width, existT _ _stringappend_1590_ _) :=
- w__567
- : (word_width * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1591_ :=
+ _stringappend_1078_
+ (build_ex (projT1 (string_length
+ "."))) in
+ (match (size_mnemonic_matches_prefix
+ _stringappend_1079_) with
+ | Some
+ (width,(existT _ _stringappend_1080_ _)) =>
+ returnm (width, build_ex _stringappend_1080_)
+ | _ =>
+ exit tt
+ : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(width, existT _ _stringappend_1080_ _) =>
+ let _stringappend_1081_ :=
string_drop
- _stringappend_1588_
- (build_ex
- _stringappend_1590_) in
- match (maybe_aq_matches_prefix
- _stringappend_1591_) with
- | Some
- (_stringappend_1592_,(existT _ _stringappend_1593_ _)) =>
- returnm ((_stringappend_1592_,
- build_ex
- _stringappend_1593_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__569 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(aq, existT _ _stringappend_1593_ _) :=
- w__569
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1594_ :=
+ _stringappend_1079_
+ (build_ex _stringappend_1080_) in
+ (match (maybe_aq_matches_prefix
+ _stringappend_1081_) with
+ | Some
+ (aq,(existT _ _stringappend_1082_ _)) =>
+ returnm (aq, build_ex _stringappend_1082_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(aq, existT _ _stringappend_1082_ _) =>
+ let _stringappend_1083_ :=
string_drop
- _stringappend_1591_
- (build_ex
- _stringappend_1593_) in
- match (maybe_rl_matches_prefix
- _stringappend_1594_) with
- | Some
- (_stringappend_1595_,(existT _ _stringappend_1596_ _)) =>
- returnm ((_stringappend_1595_,
- build_ex
- _stringappend_1596_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__571 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rl, existT _ _stringappend_1596_ _) :=
- w__571
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1597_ :=
+ _stringappend_1081_
+ (build_ex _stringappend_1082_) in
+ (match (maybe_rl_matches_prefix
+ _stringappend_1083_) with
+ | Some
+ (rl,(existT _ _stringappend_1084_ _)) =>
+ returnm (rl, build_ex _stringappend_1084_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rl, existT _ _stringappend_1084_ _) =>
+ let _stringappend_1085_ :=
string_drop
- _stringappend_1594_
- (build_ex
- _stringappend_1596_) in
- match (spc_matches_prefix
- _stringappend_1597_) with
- | Some
- (_stringappend_1598_,(existT _ _stringappend_1599_ _)) =>
- returnm ((_stringappend_1598_,
- build_ex
- _stringappend_1599_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__573 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1599_ _) :=
- w__573
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1600_ :=
+ _stringappend_1083_
+ (build_ex _stringappend_1084_) in
+ (match (spc_matches_prefix
+ _stringappend_1085_) with
+ | Some
+ (tt,(existT _ _stringappend_1086_ _)) =>
+ returnm (tt, build_ex _stringappend_1086_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1086_ _) =>
+ let _stringappend_1087_ :=
string_drop
- _stringappend_1597_
- (build_ex
- _stringappend_1599_) in
- match (reg_name_matches_prefix
- _stringappend_1600_) with
- | Some
- (_stringappend_1601_,(existT _ _stringappend_1602_ _)) =>
- returnm ((_stringappend_1601_,
- build_ex
- _stringappend_1602_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__575 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1602_ _) :=
- w__575
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1603_ :=
+ _stringappend_1085_
+ (build_ex _stringappend_1086_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1087_) with
+ | Some
+ (rd,(existT _ _stringappend_1088_ _)) =>
+ returnm (rd, build_ex _stringappend_1088_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_1088_ _) =>
+ let _stringappend_1089_ :=
string_drop
- _stringappend_1600_
- (build_ex
- _stringappend_1602_) in
+ _stringappend_1087_
+ (build_ex _stringappend_1088_) in
sep_matches_prefix
- _stringappend_1603_ >>= fun w__576 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1089_ >>= fun w__576 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__576 with
- | Some
- (_stringappend_1604_,(existT _ _stringappend_1605_ _)) =>
- returnm ((_stringappend_1604_,
- build_ex
- _stringappend_1605_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__578 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1605_ _) :=
- w__578
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1606_ :=
+ (match w__576 with
+ | Some
+ (tt,(existT _ _stringappend_1090_ _)) =>
+ returnm (tt, build_ex _stringappend_1090_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1090_ _) =>
+ let _stringappend_1091_ :=
string_drop
- _stringappend_1603_
- (build_ex
- _stringappend_1605_) in
- match (reg_name_matches_prefix
- _stringappend_1606_) with
- | Some
- (_stringappend_1607_,(existT _ _stringappend_1608_ _)) =>
- returnm ((_stringappend_1607_,
- build_ex
- _stringappend_1608_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__580 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1608_ _) :=
- w__580
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1609_ :=
+ _stringappend_1089_
+ (build_ex _stringappend_1090_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1091_) with
+ | Some
+ (rs1,(existT _ _stringappend_1092_ _)) =>
+ returnm (rs1, build_ex _stringappend_1092_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_1092_ _) =>
+ let _stringappend_1093_ :=
string_drop
- _stringappend_1606_
- (build_ex
- _stringappend_1608_) in
+ _stringappend_1091_
+ (build_ex _stringappend_1092_) in
sep_matches_prefix
- _stringappend_1609_ >>= fun w__581 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1093_ >>= fun w__581 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__581 with
- | Some
- (_stringappend_1610_,(existT _ _stringappend_1611_ _)) =>
- returnm ((_stringappend_1610_,
- build_ex
- _stringappend_1611_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__583 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1611_ _) :=
- w__583
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1612_ :=
+ (match w__581 with
+ | Some
+ (tt,(existT _ _stringappend_1094_ _)) =>
+ returnm (tt, build_ex _stringappend_1094_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1094_ _) =>
+ let _stringappend_1095_ :=
string_drop
- _stringappend_1609_
- (build_ex
- _stringappend_1611_) in
- match (reg_name_matches_prefix
- _stringappend_1612_) with
- | Some
- (_stringappend_1613_,(existT _ _stringappend_1614_ _)) =>
- returnm ((_stringappend_1613_,
- build_ex
- _stringappend_1614_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__585 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_1614_ _) :=
- w__585
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_1093_
+ (build_ex _stringappend_1094_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1095_) with
+ | Some
+ (rs2,(existT _ _stringappend_1096_ _)) =>
+ returnm (rs2, build_ex _stringappend_1096_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_1096_ _) =>
(match (string_drop
- _stringappend_1612_
- (build_ex
- _stringappend_1614_)) with
+ _stringappend_1095_
+ (build_ex _stringappend_1096_)) with
| "" =>
- returnm ((AMO (op,aq,rl,rs2,rs1,width,rd))
- : ast)
+ returnm ((AMO
+ (op,
+ aq,
+ rl,
+ rs2,
+ rs1,
+ width,
+ rd))
+ : ast )
| _ =>
exit tt
: M (ast)
@@ -16863,83 +16860,76 @@ Definition assembly_backwards (arg_ : string)
: M (ast)
else
match (csr_mnemonic_matches_prefix
- _stringappend_1112_) with
+ _stringappend_756_) with
| Some
- (_stringappend_1616_,(existT _ _stringappend_1617_ _)) =>
- let _stringappend_1618_ :=
+ (op,(existT _ _stringappend_1098_ _)) =>
+ let _stringappend_1099_ :=
string_drop
- _stringappend_1112_
- (build_ex
- _stringappend_1617_) in
+ _stringappend_756_
+ (build_ex _stringappend_1098_) in
and_boolM
(returnm ((string_startswith
- _stringappend_1618_
+ _stringappend_1099_
"i")
: bool))
- (let _stringappend_1619_ :=
+ (let _stringappend_1100_ :=
string_drop
- _stringappend_1618_
- (string_length
- "i") in
+ _stringappend_1099_
+ (build_ex (projT1 (string_length
+ "i"))) in
match (spc_matches_prefix
- _stringappend_1619_) with
+ _stringappend_1100_) with
| Some
- (_stringappend_1620_,(existT _ _stringappend_1621_ _)) =>
- let _stringappend_1622_ :=
+ (tt,(existT _ _stringappend_1101_ _)) =>
+ let _stringappend_1102_ :=
string_drop
- _stringappend_1619_
- (build_ex
- _stringappend_1621_) in
+ _stringappend_1100_
+ (build_ex _stringappend_1101_) in
match (reg_name_matches_prefix
- _stringappend_1622_) with
+ _stringappend_1102_) with
| Some
- (_stringappend_1623_,(existT _ _stringappend_1624_ _)) =>
- let _stringappend_1625_ :=
+ (rd,(existT _ _stringappend_1103_ _)) =>
+ let _stringappend_1104_ :=
string_drop
- _stringappend_1622_
- (build_ex
- _stringappend_1624_) in
+ _stringappend_1102_
+ (build_ex _stringappend_1103_) in
sep_matches_prefix
- _stringappend_1625_ >>= fun w__588 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1104_ >>= fun w__588 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__588 with
| Some
- (_stringappend_1626_,(existT _ _stringappend_1627_ _)) =>
- let _stringappend_1628_ :=
+ (tt,(existT _ _stringappend_1105_ _)) =>
+ let _stringappend_1106_ :=
string_drop
- _stringappend_1625_
- (build_ex
- _stringappend_1627_) in
+ _stringappend_1104_
+ (build_ex _stringappend_1105_) in
match (hex_bits_5_matches_prefix
- _stringappend_1628_) with
+ _stringappend_1106_) with
| Some
- (_stringappend_1629_,(existT _ _stringappend_1630_ _)) =>
- let _stringappend_1631_ :=
+ (rs1,(existT _ _stringappend_1107_ _)) =>
+ let _stringappend_1108_ :=
string_drop
- _stringappend_1628_
- (build_ex
- _stringappend_1630_) in
+ _stringappend_1106_
+ (build_ex _stringappend_1107_) in
sep_matches_prefix
- _stringappend_1631_ >>= fun w__589 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1108_ >>= fun w__589 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if
((match w__589 with
| Some
- (_stringappend_1632_,(existT _ _stringappend_1633_ _)) =>
- let _stringappend_1634_ :=
+ (tt,(existT _ _stringappend_1109_ _)) =>
+ let _stringappend_1110_ :=
string_drop
- _stringappend_1631_
- (build_ex
- _stringappend_1633_) in
+ _stringappend_1108_
+ (build_ex _stringappend_1109_) in
if
((match (csr_name_map_matches_prefix
- _stringappend_1634_) with
+ _stringappend_1110_) with
| Some
- (_stringappend_1635_,(existT _ _stringappend_1636_ _)) =>
+ (csr,(existT _ _stringappend_1111_ _)) =>
match (string_drop
- _stringappend_1634_
- (build_ex
- _stringappend_1636_)) with
+ _stringappend_1110_
+ (build_ex _stringappend_1111_)) with
| "" =>
true
| _ =>
@@ -17013,185 +17003,127 @@ Definition assembly_backwards (arg_ : string)
: bool)
end >>= fun w__595 : bool =>
(if (w__595) then
- match (csr_mnemonic_matches_prefix
- _stringappend_1112_) with
- | Some
- (_stringappend_1616_,(existT _ _stringappend_1617_ _)) =>
- returnm ((_stringappend_1616_,
- build_ex
- _stringappend_1617_)
- : (csrop * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((csrop * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__597 : (csrop * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(op, existT _ _stringappend_1617_ _) :=
- w__597
- : (csrop * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1618_ :=
+ (match (csr_mnemonic_matches_prefix
+ _stringappend_756_) with
+ | Some
+ (op,(existT _ _stringappend_1098_ _)) =>
+ returnm (op, build_ex _stringappend_1098_)
+ | _ =>
+ exit tt
+ : M ((csrop * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((csrop * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(op, existT _ _stringappend_1098_ _) =>
+ let _stringappend_1099_ :=
string_drop
- _stringappend_1112_
- (build_ex
- _stringappend_1617_) in
- let _stringappend_1619_ :=
+ _stringappend_756_
+ (build_ex _stringappend_1098_) in
+ let _stringappend_1100_ :=
string_drop
- _stringappend_1618_
- (string_length
- "i") in
- match (spc_matches_prefix
- _stringappend_1619_) with
- | Some
- (_stringappend_1620_,(existT _ _stringappend_1621_ _)) =>
- returnm ((_stringappend_1620_,
- build_ex
- _stringappend_1621_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__599 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1621_ _) :=
- w__599
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1622_ :=
+ _stringappend_1099_
+ (build_ex (projT1 (string_length
+ "i"))) in
+ (match (spc_matches_prefix
+ _stringappend_1100_) with
+ | Some
+ (tt,(existT _ _stringappend_1101_ _)) =>
+ returnm (tt, build_ex _stringappend_1101_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1101_ _) =>
+ let _stringappend_1102_ :=
string_drop
- _stringappend_1619_
- (build_ex
- _stringappend_1621_) in
- match (reg_name_matches_prefix
- _stringappend_1622_) with
- | Some
- (_stringappend_1623_,(existT _ _stringappend_1624_ _)) =>
- returnm ((_stringappend_1623_,
- build_ex
- _stringappend_1624_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__601 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1624_ _) :=
- w__601
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1625_ :=
+ _stringappend_1100_
+ (build_ex _stringappend_1101_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1102_) with
+ | Some
+ (rd,(existT _ _stringappend_1103_ _)) =>
+ returnm (rd, build_ex _stringappend_1103_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_1103_ _) =>
+ let _stringappend_1104_ :=
string_drop
- _stringappend_1622_
- (build_ex
- _stringappend_1624_) in
+ _stringappend_1102_
+ (build_ex _stringappend_1103_) in
sep_matches_prefix
- _stringappend_1625_ >>= fun w__602 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1104_ >>= fun w__602 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__602 with
- | Some
- (_stringappend_1626_,(existT _ _stringappend_1627_ _)) =>
- returnm ((_stringappend_1626_,
- build_ex
- _stringappend_1627_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__604 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1627_ _) :=
- w__604
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1628_ :=
+ (match w__602 with
+ | Some
+ (tt,(existT _ _stringappend_1105_ _)) =>
+ returnm (tt, build_ex _stringappend_1105_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1105_ _) =>
+ let _stringappend_1106_ :=
string_drop
- _stringappend_1625_
- (build_ex
- _stringappend_1627_) in
- match (hex_bits_5_matches_prefix
- _stringappend_1628_) with
- | Some
- (_stringappend_1629_,(existT _ _stringappend_1630_ _)) =>
- returnm ((_stringappend_1629_,
- build_ex
- _stringappend_1630_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__606 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1630_ _) :=
- w__606
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1631_ :=
+ _stringappend_1104_
+ (build_ex _stringappend_1105_) in
+ (match (hex_bits_5_matches_prefix
+ _stringappend_1106_) with
+ | Some
+ (rs1,(existT _ _stringappend_1107_ _)) =>
+ returnm (rs1, build_ex _stringappend_1107_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_1107_ _) =>
+ let _stringappend_1108_ :=
string_drop
- _stringappend_1628_
- (build_ex
- _stringappend_1630_) in
+ _stringappend_1106_
+ (build_ex _stringappend_1107_) in
sep_matches_prefix
- _stringappend_1631_ >>= fun w__607 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1108_ >>= fun w__607 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__607 with
- | Some
- (_stringappend_1632_,(existT _ _stringappend_1633_ _)) =>
- returnm ((_stringappend_1632_,
- build_ex
- _stringappend_1633_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__609 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1633_ _) :=
- w__609
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1634_ :=
+ (match w__607 with
+ | Some
+ (tt,(existT _ _stringappend_1109_ _)) =>
+ returnm (tt, build_ex _stringappend_1109_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1109_ _) =>
+ let _stringappend_1110_ :=
string_drop
- _stringappend_1631_
- (build_ex
- _stringappend_1633_) in
- match (csr_name_map_matches_prefix
- _stringappend_1634_) with
- | Some
- (_stringappend_1635_,(existT _ _stringappend_1636_ _)) =>
- returnm ((_stringappend_1635_,
- build_ex
- _stringappend_1636_)
- : (mword 12 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 12 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__611 : (mword 12 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(csr, existT _ _stringappend_1636_ _) :=
- w__611
- : (mword 12 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_1108_
+ (build_ex _stringappend_1109_) in
+ (match (csr_name_map_matches_prefix
+ _stringappend_1110_) with
+ | Some
+ (csr,(existT _ _stringappend_1111_ _)) =>
+ returnm (csr, build_ex _stringappend_1111_)
+ | _ =>
+ exit tt
+ : M ((mword 12 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(csr, existT _ _stringappend_1111_ _) =>
(match (string_drop
- _stringappend_1634_
- (build_ex
- _stringappend_1636_)) with
+ _stringappend_1110_
+ (build_ex _stringappend_1111_)) with
| "" =>
- returnm ((CSR (csr,rs1,rd,true,op))
- : ast)
+ returnm ((CSR
+ (csr,
+ rs1,
+ rd,
+ true,
+ op))
+ : ast )
| _ =>
exit tt
: M (ast)
@@ -17199,73 +17131,66 @@ Definition assembly_backwards (arg_ : string)
: M (ast)
else
match (csr_mnemonic_matches_prefix
- _stringappend_1112_) with
+ _stringappend_756_) with
| Some
- (_stringappend_1638_,(existT _ _stringappend_1639_ _)) =>
- let _stringappend_1640_ :=
+ (op,(existT _ _stringappend_1113_ _)) =>
+ let _stringappend_1114_ :=
string_drop
- _stringappend_1112_
- (build_ex
- _stringappend_1639_) in
+ _stringappend_756_
+ (build_ex _stringappend_1113_) in
match (spc_matches_prefix
- _stringappend_1640_) with
+ _stringappend_1114_) with
| Some
- (_stringappend_1641_,(existT _ _stringappend_1642_ _)) =>
- let _stringappend_1643_ :=
+ (tt,(existT _ _stringappend_1115_ _)) =>
+ let _stringappend_1116_ :=
string_drop
- _stringappend_1640_
- (build_ex
- _stringappend_1642_) in
+ _stringappend_1114_
+ (build_ex _stringappend_1115_) in
match (reg_name_matches_prefix
- _stringappend_1643_) with
+ _stringappend_1116_) with
| Some
- (_stringappend_1644_,(existT _ _stringappend_1645_ _)) =>
- let _stringappend_1646_ :=
+ (rd,(existT _ _stringappend_1117_ _)) =>
+ let _stringappend_1118_ :=
string_drop
- _stringappend_1643_
- (build_ex
- _stringappend_1645_) in
+ _stringappend_1116_
+ (build_ex _stringappend_1117_) in
sep_matches_prefix
- _stringappend_1646_ >>= fun w__614 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1118_ >>= fun w__614 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__614 with
| Some
- (_stringappend_1647_,(existT _ _stringappend_1648_ _)) =>
- let _stringappend_1649_ :=
+ (tt,(existT _ _stringappend_1119_ _)) =>
+ let _stringappend_1120_ :=
string_drop
- _stringappend_1646_
- (build_ex
- _stringappend_1648_) in
+ _stringappend_1118_
+ (build_ex _stringappend_1119_) in
match (reg_name_matches_prefix
- _stringappend_1649_) with
+ _stringappend_1120_) with
| Some
- (_stringappend_1650_,(existT _ _stringappend_1651_ _)) =>
- let _stringappend_1652_ :=
+ (rs1,(existT _ _stringappend_1121_ _)) =>
+ let _stringappend_1122_ :=
string_drop
- _stringappend_1649_
- (build_ex
- _stringappend_1651_) in
+ _stringappend_1120_
+ (build_ex _stringappend_1121_) in
sep_matches_prefix
- _stringappend_1652_ >>= fun w__615 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1122_ >>= fun w__615 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if
((match w__615 with
| Some
- (_stringappend_1653_,(existT _ _stringappend_1654_ _)) =>
- let _stringappend_1655_ :=
+ (tt,(existT _ _stringappend_1123_ _)) =>
+ let _stringappend_1124_ :=
string_drop
- _stringappend_1652_
- (build_ex
- _stringappend_1654_) in
+ _stringappend_1122_
+ (build_ex _stringappend_1123_) in
if
((match (csr_name_map_matches_prefix
- _stringappend_1655_) with
+ _stringappend_1124_) with
| Some
- (_stringappend_1656_,(existT _ _stringappend_1657_ _)) =>
+ (csr,(existT _ _stringappend_1125_ _)) =>
match (string_drop
- _stringappend_1655_
- (build_ex
- _stringappend_1657_)) with
+ _stringappend_1124_
+ (build_ex _stringappend_1125_)) with
| "" =>
true
| _ =>
@@ -17334,241 +17259,166 @@ Definition assembly_backwards (arg_ : string)
: bool)
end >>= fun w__620 : bool =>
(if (w__620) then
- match (csr_mnemonic_matches_prefix
- _stringappend_1112_) with
- | Some
- (_stringappend_1638_,(existT _ _stringappend_1639_ _)) =>
- returnm ((_stringappend_1638_,
- build_ex
- _stringappend_1639_)
- : (csrop * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((csrop * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__622 : (csrop * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(op, existT _ _stringappend_1639_ _) :=
- w__622
- : (csrop * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1640_ :=
+ (match (csr_mnemonic_matches_prefix
+ _stringappend_756_) with
+ | Some
+ (op,(existT _ _stringappend_1113_ _)) =>
+ returnm (op, build_ex _stringappend_1113_)
+ | _ =>
+ exit tt
+ : M ((csrop * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((csrop * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(op, existT _ _stringappend_1113_ _) =>
+ let _stringappend_1114_ :=
string_drop
- _stringappend_1112_
- (build_ex
- _stringappend_1639_) in
- match (spc_matches_prefix
- _stringappend_1640_) with
- | Some
- (_stringappend_1641_,(existT _ _stringappend_1642_ _)) =>
- returnm ((_stringappend_1641_,
- build_ex
- _stringappend_1642_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__624 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1642_ _) :=
- w__624
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1643_ :=
+ _stringappend_756_
+ (build_ex _stringappend_1113_) in
+ (match (spc_matches_prefix
+ _stringappend_1114_) with
+ | Some
+ (tt,(existT _ _stringappend_1115_ _)) =>
+ returnm (tt, build_ex _stringappend_1115_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1115_ _) =>
+ let _stringappend_1116_ :=
string_drop
- _stringappend_1640_
- (build_ex
- _stringappend_1642_) in
- match (reg_name_matches_prefix
- _stringappend_1643_) with
- | Some
- (_stringappend_1644_,(existT _ _stringappend_1645_ _)) =>
- returnm ((_stringappend_1644_,
- build_ex
- _stringappend_1645_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__626 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1645_ _) :=
- w__626
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1646_ :=
+ _stringappend_1114_
+ (build_ex _stringappend_1115_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1116_) with
+ | Some
+ (rd,(existT _ _stringappend_1117_ _)) =>
+ returnm (rd, build_ex _stringappend_1117_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_1117_ _) =>
+ let _stringappend_1118_ :=
string_drop
- _stringappend_1643_
- (build_ex
- _stringappend_1645_) in
+ _stringappend_1116_
+ (build_ex _stringappend_1117_) in
sep_matches_prefix
- _stringappend_1646_ >>= fun w__627 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1118_ >>= fun w__627 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__627 with
- | Some
- (_stringappend_1647_,(existT _ _stringappend_1648_ _)) =>
- returnm ((_stringappend_1647_,
- build_ex
- _stringappend_1648_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__629 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1648_ _) :=
- w__629
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1649_ :=
+ (match w__627 with
+ | Some
+ (tt,(existT _ _stringappend_1119_ _)) =>
+ returnm (tt, build_ex _stringappend_1119_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1119_ _) =>
+ let _stringappend_1120_ :=
string_drop
- _stringappend_1646_
- (build_ex
- _stringappend_1648_) in
- match (reg_name_matches_prefix
- _stringappend_1649_) with
- | Some
- (_stringappend_1650_,(existT _ _stringappend_1651_ _)) =>
- returnm ((_stringappend_1650_,
- build_ex
- _stringappend_1651_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__631 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1651_ _) :=
- w__631
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1652_ :=
+ _stringappend_1118_
+ (build_ex _stringappend_1119_) in
+ (match (reg_name_matches_prefix
+ _stringappend_1120_) with
+ | Some
+ (rs1,(existT _ _stringappend_1121_ _)) =>
+ returnm (rs1, build_ex _stringappend_1121_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_1121_ _) =>
+ let _stringappend_1122_ :=
string_drop
- _stringappend_1649_
- (build_ex
- _stringappend_1651_) in
+ _stringappend_1120_
+ (build_ex _stringappend_1121_) in
sep_matches_prefix
- _stringappend_1652_ >>= fun w__632 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_1122_ >>= fun w__632 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__632 with
- | Some
- (_stringappend_1653_,(existT _ _stringappend_1654_ _)) =>
- returnm ((_stringappend_1653_,
- build_ex
- _stringappend_1654_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__634 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1654_ _) :=
- w__634
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1655_ :=
+ (match w__632 with
+ | Some
+ (tt,(existT _ _stringappend_1123_ _)) =>
+ returnm (tt, build_ex _stringappend_1123_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1123_ _) =>
+ let _stringappend_1124_ :=
string_drop
- _stringappend_1652_
- (build_ex
- _stringappend_1654_) in
- match (csr_name_map_matches_prefix
- _stringappend_1655_) with
- | Some
- (_stringappend_1656_,(existT _ _stringappend_1657_ _)) =>
- returnm ((_stringappend_1656_,
- build_ex
- _stringappend_1657_)
- : (mword 12 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 12 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__636 : (mword 12 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(csr, existT _ _stringappend_1657_ _) :=
- w__636
- : (mword 12 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_1122_
+ (build_ex _stringappend_1123_) in
+ (match (csr_name_map_matches_prefix
+ _stringappend_1124_) with
+ | Some
+ (csr,(existT _ _stringappend_1125_ _)) =>
+ returnm (csr, build_ex _stringappend_1125_)
+ | _ =>
+ exit tt
+ : M ((mword 12 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(csr, existT _ _stringappend_1125_ _) =>
(match (string_drop
- _stringappend_1655_
- (build_ex
- _stringappend_1657_)) with
+ _stringappend_1124_
+ (build_ex _stringappend_1125_)) with
| "" =>
- returnm ((CSR (csr,rs1,rd,false,op))
- : ast)
+ returnm ((CSR
+ (csr,
+ rs1,
+ rd,
+ false,
+ op))
+ : ast )
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- let _stringappend_1659_ :=
+ let _stringappend_1127_ :=
string_drop
- _stringappend_1112_
- (string_length
- "illegal") in
- match (spc_matches_prefix
- _stringappend_1659_) with
- | Some
- (_stringappend_1660_,(existT _ _stringappend_1661_ _)) =>
- returnm ((_stringappend_1660_,
- build_ex
- _stringappend_1661_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__640 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1661_ _) :=
- w__640
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1662_ :=
+ _stringappend_756_
+ (build_ex (projT1 (string_length
+ "illegal"))) in
+ (match (spc_matches_prefix
+ _stringappend_1127_) with
+ | Some
+ (tt,(existT _ _stringappend_1128_ _)) =>
+ returnm (tt, build_ex _stringappend_1128_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_1128_ _) =>
+ let _stringappend_1129_ :=
string_drop
- _stringappend_1659_
- (build_ex
- _stringappend_1661_) in
- match (hex_bits_32_matches_prefix
- _stringappend_1662_) with
- | Some
- (_stringappend_1663_,(existT _ _stringappend_1664_ _)) =>
- returnm ((_stringappend_1663_,
- build_ex
- _stringappend_1664_)
- : (mword 32 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 32 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__642 : (mword 32 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_1664_ _) :=
- w__642
- : (mword 32 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_1127_
+ (build_ex _stringappend_1128_) in
+ (match (hex_bits_32_matches_prefix
+ _stringappend_1129_) with
+ | Some
+ (s,(existT _ _stringappend_1130_ _)) =>
+ returnm (s, build_ex _stringappend_1130_)
+ | _ =>
+ exit tt
+ : M ((mword 32 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 32 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(s, existT _ _stringappend_1130_ _) =>
(match (string_drop
- _stringappend_1662_
- (build_ex
- _stringappend_1664_)) with
+ _stringappend_1129_
+ (build_ex _stringappend_1130_)) with
| "" =>
- returnm ((ILLEGAL s)
- : ast)
+ returnm ((ILLEGAL
+ (s))
+ : ast )
| _ =>
exit tt
: M (ast)
@@ -17644,27 +17494,26 @@ Definition assembly_forwards_matches (arg_ : ast)
Definition assembly_backwards_matches (arg_ : string)
: M (bool) :=
- let _stringappend_559_ := arg_ in
- match (utype_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_560_,(existT _ _stringappend_561_ _)) =>
- let _stringappend_562_ := string_drop _stringappend_559_ (build_ex _stringappend_561_) in
- match (spc_matches_prefix _stringappend_562_) with
- | Some (_stringappend_563_,(existT _ _stringappend_564_ _)) =>
- let _stringappend_565_ := string_drop _stringappend_562_ (build_ex _stringappend_564_) in
- match (reg_name_matches_prefix _stringappend_565_) with
- | Some (_stringappend_566_,(existT _ _stringappend_567_ _)) =>
- let _stringappend_568_ := string_drop _stringappend_565_ (build_ex _stringappend_567_) in
- sep_matches_prefix _stringappend_568_ >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_381_ := arg_ in
+ match (utype_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_382_ _)) =>
+ let _stringappend_383_ := string_drop _stringappend_381_ (build_ex _stringappend_382_) in
+ match (spc_matches_prefix _stringappend_383_) with
+ | Some (tt,(existT _ _stringappend_384_ _)) =>
+ let _stringappend_385_ := string_drop _stringappend_383_ (build_ex _stringappend_384_) in
+ match (reg_name_matches_prefix _stringappend_385_) with
+ | Some (rd,(existT _ _stringappend_386_ _)) =>
+ let _stringappend_387_ := string_drop _stringappend_385_ (build_ex _stringappend_386_) in
+ sep_matches_prefix _stringappend_387_ >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__0 with
- | Some (_stringappend_569_,(existT _ _stringappend_570_ _)) =>
- let _stringappend_571_ :=
- string_drop _stringappend_568_ (build_ex _stringappend_570_) in
- if ((match (hex_bits_20_matches_prefix _stringappend_571_) with
- | Some (_stringappend_572_,(existT _ _stringappend_573_ _)) =>
- match (string_drop _stringappend_571_
- (build_ex
- _stringappend_573_)) with
+ | Some (tt,(existT _ _stringappend_388_ _)) =>
+ let _stringappend_389_ :=
+ string_drop _stringappend_387_ (build_ex _stringappend_388_) in
+ if ((match (hex_bits_20_matches_prefix _stringappend_389_) with
+ | Some (imm,(existT _ _stringappend_390_ _)) =>
+ match (string_drop _stringappend_389_
+ (build_ex _stringappend_390_)) with
| "" => true
| _ => false
end
@@ -17690,74 +17539,57 @@ Definition assembly_backwards_matches (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__3 : bool =>
(if (w__3) then
- match (utype_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_560_,(existT _ _stringappend_561_ _)) =>
- returnm ((_stringappend_560_, build_ex _stringappend_561_)
- : (uop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((uop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__5 : (uop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_561_ _) := w__5 : (uop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_562_ := string_drop _stringappend_559_ (build_ex _stringappend_561_) in
- match (spc_matches_prefix _stringappend_562_) with
- | Some (_stringappend_563_,(existT _ _stringappend_564_ _)) =>
- returnm ((_stringappend_563_, build_ex _stringappend_564_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__7 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_564_ _) := w__7 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_565_ := string_drop _stringappend_562_ (build_ex _stringappend_564_) in
- match (reg_name_matches_prefix _stringappend_565_) with
- | Some (_stringappend_566_,(existT _ _stringappend_567_ _)) =>
- returnm ((_stringappend_566_, build_ex _stringappend_567_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__9 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_567_ _) := w__9 : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_568_ := string_drop _stringappend_565_ (build_ex _stringappend_567_) in
- sep_matches_prefix _stringappend_568_ >>= fun w__10 : option ((unit * {n : Z & ArithFact (n >=
+ (match (utype_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_382_ _)) => returnm (op, build_ex _stringappend_382_)
+ | _ => exit tt : M ((uop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((uop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_382_ _) =>
+ let _stringappend_383_ := string_drop _stringappend_381_ (build_ex _stringappend_382_) in
+ (match (spc_matches_prefix _stringappend_383_) with
+ | Some (tt,(existT _ _stringappend_384_ _)) => returnm (tt, build_ex _stringappend_384_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_384_ _) =>
+ let _stringappend_385_ := string_drop _stringappend_383_ (build_ex _stringappend_384_) in
+ (match (reg_name_matches_prefix _stringappend_385_) with
+ | Some (rd,(existT _ _stringappend_386_ _)) => returnm (rd, build_ex _stringappend_386_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_386_ _) =>
+ let _stringappend_387_ := string_drop _stringappend_385_ (build_ex _stringappend_386_) in
+ sep_matches_prefix _stringappend_387_ >>= fun w__10 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__10 with
- | Some (_stringappend_569_,(existT _ _stringappend_570_ _)) =>
- returnm ((_stringappend_569_, build_ex _stringappend_570_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__12 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_570_ _) := w__12 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_571_ := string_drop _stringappend_568_ (build_ex _stringappend_570_) in
- match (hex_bits_20_matches_prefix _stringappend_571_) with
- | Some (_stringappend_572_,(existT _ _stringappend_573_ _)) =>
- returnm ((_stringappend_572_, build_ex _stringappend_573_)
- : (mword 20 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 20 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__14 : (mword 20 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_573_ _) :=
- w__14
- : (mword 20 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_571_ (build_ex _stringappend_573_)) with
+ (match w__10 with
+ | Some (tt,(existT _ _stringappend_388_ _)) => returnm (tt, build_ex _stringappend_388_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_388_ _) =>
+ let _stringappend_389_ := string_drop _stringappend_387_ (build_ex _stringappend_388_) in
+ (match (hex_bits_20_matches_prefix _stringappend_389_) with
+ | Some (imm,(existT _ _stringappend_390_ _)) => returnm (imm, build_ex _stringappend_390_)
+ | _ => exit tt : M ((mword 20 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 20 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_390_ _) =>
+ (match (string_drop _stringappend_389_ (build_ex _stringappend_390_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
- and_boolM (returnm ((string_startswith _stringappend_559_ "jal") : bool))
- (let _stringappend_575_ := string_drop _stringappend_559_ (string_length "jal") in
- match (spc_matches_prefix _stringappend_575_) with
- | Some (_stringappend_576_,(existT _ _stringappend_577_ _)) =>
- let _stringappend_578_ := string_drop _stringappend_575_ (build_ex _stringappend_577_) in
- match (reg_name_matches_prefix _stringappend_578_) with
- | Some (_stringappend_579_,(existT _ _stringappend_580_ _)) =>
- let _stringappend_581_ := string_drop _stringappend_578_ (build_ex _stringappend_580_) in
- sep_matches_prefix _stringappend_581_ >>= fun w__17 : option ((unit * {n : Z & ArithFact (n >=
+ and_boolM (returnm ((string_startswith _stringappend_381_ "jal") : bool))
+ (let _stringappend_392_ :=
+ string_drop _stringappend_381_ (build_ex (projT1 (string_length "jal"))) in
+ match (spc_matches_prefix _stringappend_392_) with
+ | Some (tt,(existT _ _stringappend_393_ _)) =>
+ let _stringappend_394_ := string_drop _stringappend_392_ (build_ex _stringappend_393_) in
+ match (reg_name_matches_prefix _stringappend_394_) with
+ | Some (rd,(existT _ _stringappend_395_ _)) =>
+ let _stringappend_396_ := string_drop _stringappend_394_ (build_ex _stringappend_395_) in
+ sep_matches_prefix _stringappend_396_ >>= fun w__17 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__17 with
- | Some (_stringappend_582_,(existT _ _stringappend_583_ _)) =>
- let _stringappend_584_ :=
- string_drop _stringappend_581_ (build_ex _stringappend_583_) in
- if ((match (hex_bits_21_matches_prefix _stringappend_584_) with
- | Some (_stringappend_585_,(existT _ _stringappend_586_ _)) =>
- match (string_drop _stringappend_584_
- (build_ex
- _stringappend_586_)) with
+ | Some (tt,(existT _ _stringappend_397_ _)) =>
+ let _stringappend_398_ :=
+ string_drop _stringappend_396_ (build_ex _stringappend_397_) in
+ if ((match (hex_bits_21_matches_prefix _stringappend_398_) with
+ | Some (imm,(existT _ _stringappend_399_ _)) =>
+ match (string_drop _stringappend_398_
+ (build_ex _stringappend_399_)) with
| "" => true
| _ => false
end
@@ -17781,83 +17613,66 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__20 : bool =>
(if (w__20) then
- let _stringappend_575_ := string_drop _stringappend_559_ (string_length "jal") in
- match (spc_matches_prefix _stringappend_575_) with
- | Some (_stringappend_576_,(existT _ _stringappend_577_ _)) =>
- returnm ((_stringappend_576_, build_ex _stringappend_577_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__22 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_577_ _) := w__22 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_578_ := string_drop _stringappend_575_ (build_ex _stringappend_577_) in
- match (reg_name_matches_prefix _stringappend_578_) with
- | Some (_stringappend_579_,(existT _ _stringappend_580_ _)) =>
- returnm ((_stringappend_579_, build_ex _stringappend_580_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__24 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_580_ _) :=
- w__24
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_581_ := string_drop _stringappend_578_ (build_ex _stringappend_580_) in
- sep_matches_prefix _stringappend_581_ >>= fun w__25 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_392_ :=
+ string_drop _stringappend_381_ (build_ex (projT1 (string_length "jal"))) in
+ (match (spc_matches_prefix _stringappend_392_) with
+ | Some (tt,(existT _ _stringappend_393_ _)) => returnm (tt, build_ex _stringappend_393_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_393_ _) =>
+ let _stringappend_394_ := string_drop _stringappend_392_ (build_ex _stringappend_393_) in
+ (match (reg_name_matches_prefix _stringappend_394_) with
+ | Some (rd,(existT _ _stringappend_395_ _)) => returnm (rd, build_ex _stringappend_395_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_395_ _) =>
+ let _stringappend_396_ := string_drop _stringappend_394_ (build_ex _stringappend_395_) in
+ sep_matches_prefix _stringappend_396_ >>= fun w__25 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__25 with
- | Some (_stringappend_582_,(existT _ _stringappend_583_ _)) =>
- returnm ((_stringappend_582_, build_ex _stringappend_583_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__27 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_583_ _) := w__27 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_584_ := string_drop _stringappend_581_ (build_ex _stringappend_583_) in
- match (hex_bits_21_matches_prefix _stringappend_584_) with
- | Some (_stringappend_585_,(existT _ _stringappend_586_ _)) =>
- returnm ((_stringappend_585_, build_ex _stringappend_586_)
- : (mword 21 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 21 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__29 : (mword 21 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_586_ _) :=
- w__29
- : (mword 21 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_584_ (build_ex _stringappend_586_)) with
+ (match w__25 with
+ | Some (tt,(existT _ _stringappend_397_ _)) => returnm (tt, build_ex _stringappend_397_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_397_ _) =>
+ let _stringappend_398_ := string_drop _stringappend_396_ (build_ex _stringappend_397_) in
+ (match (hex_bits_21_matches_prefix _stringappend_398_) with
+ | Some (imm,(existT _ _stringappend_399_ _)) => returnm (imm, build_ex _stringappend_399_)
+ | _ => exit tt : M ((mword 21 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 21 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_399_ _) =>
+ (match (string_drop _stringappend_398_ (build_ex _stringappend_399_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
- and_boolM (returnm ((string_startswith _stringappend_559_ "jalr") : bool))
- (let _stringappend_588_ := string_drop _stringappend_559_ (string_length "jalr") in
- match (spc_matches_prefix _stringappend_588_) with
- | Some (_stringappend_589_,(existT _ _stringappend_590_ _)) =>
- let _stringappend_591_ := string_drop _stringappend_588_ (build_ex _stringappend_590_) in
- match (reg_name_matches_prefix _stringappend_591_) with
- | Some (_stringappend_592_,(existT _ _stringappend_593_ _)) =>
- let _stringappend_594_ :=
- string_drop _stringappend_591_ (build_ex _stringappend_593_) in
- sep_matches_prefix _stringappend_594_ >>= fun w__32 : option ((unit * {n : Z & ArithFact (n >=
+ and_boolM (returnm ((string_startswith _stringappend_381_ "jalr") : bool))
+ (let _stringappend_401_ :=
+ string_drop _stringappend_381_ (build_ex (projT1 (string_length "jalr"))) in
+ match (spc_matches_prefix _stringappend_401_) with
+ | Some (tt,(existT _ _stringappend_402_ _)) =>
+ let _stringappend_403_ := string_drop _stringappend_401_ (build_ex _stringappend_402_) in
+ match (reg_name_matches_prefix _stringappend_403_) with
+ | Some (rd,(existT _ _stringappend_404_ _)) =>
+ let _stringappend_405_ :=
+ string_drop _stringappend_403_ (build_ex _stringappend_404_) in
+ sep_matches_prefix _stringappend_405_ >>= fun w__32 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__32 with
- | Some (_stringappend_595_,(existT _ _stringappend_596_ _)) =>
- let _stringappend_597_ :=
- string_drop _stringappend_594_ (build_ex _stringappend_596_) in
- match (reg_name_matches_prefix _stringappend_597_) with
- | Some (_stringappend_598_,(existT _ _stringappend_599_ _)) =>
- let _stringappend_600_ :=
- string_drop _stringappend_597_ (build_ex _stringappend_599_) in
- sep_matches_prefix _stringappend_600_ >>= fun w__33 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_406_ _)) =>
+ let _stringappend_407_ :=
+ string_drop _stringappend_405_ (build_ex _stringappend_406_) in
+ match (reg_name_matches_prefix _stringappend_407_) with
+ | Some (rs1,(existT _ _stringappend_408_ _)) =>
+ let _stringappend_409_ :=
+ string_drop _stringappend_407_ (build_ex _stringappend_408_) in
+ sep_matches_prefix _stringappend_409_ >>= fun w__33 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__33 with
- | Some (_stringappend_601_,(existT _ _stringappend_602_ _)) =>
- let _stringappend_603_ :=
- string_drop _stringappend_600_
- (build_ex
- _stringappend_602_) in
- if ((match (hex_bits_12_matches_prefix _stringappend_603_) with
- | Some
- (_stringappend_604_,(existT _ _stringappend_605_ _)) =>
- match (string_drop _stringappend_603_
- (build_ex
- _stringappend_605_)) with
+ | Some (tt,(existT _ _stringappend_410_ _)) =>
+ let _stringappend_411_ :=
+ string_drop _stringappend_409_
+ (build_ex _stringappend_410_) in
+ if ((match (hex_bits_12_matches_prefix _stringappend_411_) with
+ | Some (imm,(existT _ _stringappend_412_ _)) =>
+ match (string_drop _stringappend_411_
+ (build_ex _stringappend_412_)) with
| "" => true
| _ => false
end
@@ -17891,114 +17706,87 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__38 : bool =>
(if (w__38) then
- let _stringappend_588_ := string_drop _stringappend_559_ (string_length "jalr") in
- match (spc_matches_prefix _stringappend_588_) with
- | Some (_stringappend_589_,(existT _ _stringappend_590_ _)) =>
- returnm ((_stringappend_589_, build_ex _stringappend_590_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__40 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_590_ _) :=
- w__40
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_591_ := string_drop _stringappend_588_ (build_ex _stringappend_590_) in
- match (reg_name_matches_prefix _stringappend_591_) with
- | Some (_stringappend_592_,(existT _ _stringappend_593_ _)) =>
- returnm ((_stringappend_592_, build_ex _stringappend_593_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__42 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_593_ _) :=
- w__42
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_594_ := string_drop _stringappend_591_ (build_ex _stringappend_593_) in
- sep_matches_prefix _stringappend_594_ >>= fun w__43 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_401_ :=
+ string_drop _stringappend_381_ (build_ex (projT1 (string_length "jalr"))) in
+ (match (spc_matches_prefix _stringappend_401_) with
+ | Some (tt,(existT _ _stringappend_402_ _)) =>
+ returnm (tt, build_ex _stringappend_402_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_402_ _) =>
+ let _stringappend_403_ := string_drop _stringappend_401_ (build_ex _stringappend_402_) in
+ (match (reg_name_matches_prefix _stringappend_403_) with
+ | Some (rd,(existT _ _stringappend_404_ _)) =>
+ returnm (rd, build_ex _stringappend_404_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_404_ _) =>
+ let _stringappend_405_ := string_drop _stringappend_403_ (build_ex _stringappend_404_) in
+ sep_matches_prefix _stringappend_405_ >>= fun w__43 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__43 with
- | Some (_stringappend_595_,(existT _ _stringappend_596_ _)) =>
- returnm ((_stringappend_595_, build_ex _stringappend_596_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__45 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_596_ _) :=
- w__45
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_597_ := string_drop _stringappend_594_ (build_ex _stringappend_596_) in
- match (reg_name_matches_prefix _stringappend_597_) with
- | Some (_stringappend_598_,(existT _ _stringappend_599_ _)) =>
- returnm ((_stringappend_598_, build_ex _stringappend_599_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__47 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_599_ _) :=
- w__47
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_600_ := string_drop _stringappend_597_ (build_ex _stringappend_599_) in
- sep_matches_prefix _stringappend_600_ >>= fun w__48 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__43 with
+ | Some (tt,(existT _ _stringappend_406_ _)) =>
+ returnm (tt, build_ex _stringappend_406_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_406_ _) =>
+ let _stringappend_407_ := string_drop _stringappend_405_ (build_ex _stringappend_406_) in
+ (match (reg_name_matches_prefix _stringappend_407_) with
+ | Some (rs1,(existT _ _stringappend_408_ _)) =>
+ returnm (rs1, build_ex _stringappend_408_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_408_ _) =>
+ let _stringappend_409_ := string_drop _stringappend_407_ (build_ex _stringappend_408_) in
+ sep_matches_prefix _stringappend_409_ >>= fun w__48 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__48 with
- | Some (_stringappend_601_,(existT _ _stringappend_602_ _)) =>
- returnm ((_stringappend_601_, build_ex _stringappend_602_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__50 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_602_ _) :=
- w__50
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_603_ := string_drop _stringappend_600_ (build_ex _stringappend_602_) in
- match (hex_bits_12_matches_prefix _stringappend_603_) with
- | Some (_stringappend_604_,(existT _ _stringappend_605_ _)) =>
- returnm ((_stringappend_604_, build_ex _stringappend_605_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__52 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_605_ _) :=
- w__52
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_603_ (build_ex _stringappend_605_)) with
+ (match w__48 with
+ | Some (tt,(existT _ _stringappend_410_ _)) =>
+ returnm (tt, build_ex _stringappend_410_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_410_ _) =>
+ let _stringappend_411_ := string_drop _stringappend_409_ (build_ex _stringappend_410_) in
+ (match (hex_bits_12_matches_prefix _stringappend_411_) with
+ | Some (imm,(existT _ _stringappend_412_ _)) =>
+ returnm (imm, build_ex _stringappend_412_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_412_ _) =>
+ (match (string_drop _stringappend_411_ (build_ex _stringappend_412_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
- match (btype_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_607_,(existT _ _stringappend_608_ _)) =>
- let _stringappend_609_ :=
- string_drop _stringappend_559_ (build_ex _stringappend_608_) in
- match (spc_matches_prefix _stringappend_609_) with
- | Some (_stringappend_610_,(existT _ _stringappend_611_ _)) =>
- let _stringappend_612_ :=
- string_drop _stringappend_609_ (build_ex _stringappend_611_) in
- match (reg_name_matches_prefix _stringappend_612_) with
- | Some (_stringappend_613_,(existT _ _stringappend_614_ _)) =>
- let _stringappend_615_ :=
- string_drop _stringappend_612_ (build_ex _stringappend_614_) in
- sep_matches_prefix _stringappend_615_ >>= fun w__55 : option ((unit * {n : Z & ArithFact (n >=
+ match (btype_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_414_ _)) =>
+ let _stringappend_415_ :=
+ string_drop _stringappend_381_ (build_ex _stringappend_414_) in
+ match (spc_matches_prefix _stringappend_415_) with
+ | Some (tt,(existT _ _stringappend_416_ _)) =>
+ let _stringappend_417_ :=
+ string_drop _stringappend_415_ (build_ex _stringappend_416_) in
+ match (reg_name_matches_prefix _stringappend_417_) with
+ | Some (rs1,(existT _ _stringappend_418_ _)) =>
+ let _stringappend_419_ :=
+ string_drop _stringappend_417_ (build_ex _stringappend_418_) in
+ sep_matches_prefix _stringappend_419_ >>= fun w__55 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__55 with
- | Some (_stringappend_616_,(existT _ _stringappend_617_ _)) =>
- let _stringappend_618_ :=
- string_drop _stringappend_615_ (build_ex _stringappend_617_) in
- match (reg_name_matches_prefix _stringappend_618_) with
- | Some (_stringappend_619_,(existT _ _stringappend_620_ _)) =>
- let _stringappend_621_ :=
- string_drop _stringappend_618_ (build_ex _stringappend_620_) in
- sep_matches_prefix _stringappend_621_ >>= fun w__56 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_420_ _)) =>
+ let _stringappend_421_ :=
+ string_drop _stringappend_419_ (build_ex _stringappend_420_) in
+ match (reg_name_matches_prefix _stringappend_421_) with
+ | Some (rs2,(existT _ _stringappend_422_ _)) =>
+ let _stringappend_423_ :=
+ string_drop _stringappend_421_ (build_ex _stringappend_422_) in
+ sep_matches_prefix _stringappend_423_ >>= fun w__56 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__56 with
- | Some
- (_stringappend_622_,(existT _ _stringappend_623_ _)) =>
- let _stringappend_624_ :=
- string_drop _stringappend_621_
- (build_ex
- _stringappend_623_) in
+ | Some (tt,(existT _ _stringappend_424_ _)) =>
+ let _stringappend_425_ :=
+ string_drop _stringappend_423_
+ (build_ex _stringappend_424_) in
if ((match (hex_bits_13_matches_prefix
- _stringappend_624_) with
- | Some
- (_stringappend_625_,(existT _ _stringappend_626_ _)) =>
- match (string_drop _stringappend_624_
- (build_ex
- _stringappend_626_)) with
+ _stringappend_425_) with
+ | Some (imm,(existT _ _stringappend_426_ _)) =>
+ match (string_drop _stringappend_425_
+ (build_ex _stringappend_426_)) with
| "" => true
| _ => false
end
@@ -18034,129 +17822,97 @@ Definition assembly_backwards_matches (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__61 : bool =>
(if (w__61) then
- match (btype_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_607_,(existT _ _stringappend_608_ _)) =>
- returnm ((_stringappend_607_, build_ex _stringappend_608_)
- : (bop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__63 : (bop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_608_ _) :=
- w__63
- : (bop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_609_ :=
- string_drop _stringappend_559_ (build_ex _stringappend_608_) in
- match (spc_matches_prefix _stringappend_609_) with
- | Some (_stringappend_610_,(existT _ _stringappend_611_ _)) =>
- returnm ((_stringappend_610_, build_ex _stringappend_611_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__65 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_611_ _) :=
- w__65
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_612_ :=
- string_drop _stringappend_609_ (build_ex _stringappend_611_) in
- match (reg_name_matches_prefix _stringappend_612_) with
- | Some (_stringappend_613_,(existT _ _stringappend_614_ _)) =>
- returnm ((_stringappend_613_, build_ex _stringappend_614_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__67 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_614_ _) :=
- w__67
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_615_ :=
- string_drop _stringappend_612_ (build_ex _stringappend_614_) in
- sep_matches_prefix _stringappend_615_ >>= fun w__68 : option ((unit * {n : Z & ArithFact (n >=
+ (match (btype_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_414_ _)) =>
+ returnm (op, build_ex _stringappend_414_)
+ | _ => exit tt : M ((bop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_414_ _) =>
+ let _stringappend_415_ :=
+ string_drop _stringappend_381_ (build_ex _stringappend_414_) in
+ (match (spc_matches_prefix _stringappend_415_) with
+ | Some (tt,(existT _ _stringappend_416_ _)) =>
+ returnm (tt, build_ex _stringappend_416_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_416_ _) =>
+ let _stringappend_417_ :=
+ string_drop _stringappend_415_ (build_ex _stringappend_416_) in
+ (match (reg_name_matches_prefix _stringappend_417_) with
+ | Some (rs1,(existT _ _stringappend_418_ _)) =>
+ returnm (rs1, build_ex _stringappend_418_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_418_ _) =>
+ let _stringappend_419_ :=
+ string_drop _stringappend_417_ (build_ex _stringappend_418_) in
+ sep_matches_prefix _stringappend_419_ >>= fun w__68 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__68 with
- | Some (_stringappend_616_,(existT _ _stringappend_617_ _)) =>
- returnm ((_stringappend_616_, build_ex _stringappend_617_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__70 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_617_ _) :=
- w__70
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_618_ :=
- string_drop _stringappend_615_ (build_ex _stringappend_617_) in
- match (reg_name_matches_prefix _stringappend_618_) with
- | Some (_stringappend_619_,(existT _ _stringappend_620_ _)) =>
- returnm ((_stringappend_619_, build_ex _stringappend_620_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__72 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs2, existT _ _stringappend_620_ _) :=
- w__72
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_621_ :=
- string_drop _stringappend_618_ (build_ex _stringappend_620_) in
- sep_matches_prefix _stringappend_621_ >>= fun w__73 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__68 with
+ | Some (tt,(existT _ _stringappend_420_ _)) =>
+ returnm (tt, build_ex _stringappend_420_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_420_ _) =>
+ let _stringappend_421_ :=
+ string_drop _stringappend_419_ (build_ex _stringappend_420_) in
+ (match (reg_name_matches_prefix _stringappend_421_) with
+ | Some (rs2,(existT _ _stringappend_422_ _)) =>
+ returnm (rs2, build_ex _stringappend_422_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_422_ _) =>
+ let _stringappend_423_ :=
+ string_drop _stringappend_421_ (build_ex _stringappend_422_) in
+ sep_matches_prefix _stringappend_423_ >>= fun w__73 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__73 with
- | Some (_stringappend_622_,(existT _ _stringappend_623_ _)) =>
- returnm ((_stringappend_622_, build_ex _stringappend_623_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__75 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_623_ _) :=
- w__75
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_624_ :=
- string_drop _stringappend_621_ (build_ex _stringappend_623_) in
- match (hex_bits_13_matches_prefix _stringappend_624_) with
- | Some (_stringappend_625_,(existT _ _stringappend_626_ _)) =>
- returnm ((_stringappend_625_, build_ex _stringappend_626_)
- : (mword 13 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 13 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__77 : (mword 13 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_626_ _) :=
- w__77
- : (mword 13 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_624_ (build_ex _stringappend_626_)) with
+ (match w__73 with
+ | Some (tt,(existT _ _stringappend_424_ _)) =>
+ returnm (tt, build_ex _stringappend_424_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_424_ _) =>
+ let _stringappend_425_ :=
+ string_drop _stringappend_423_ (build_ex _stringappend_424_) in
+ (match (hex_bits_13_matches_prefix _stringappend_425_) with
+ | Some (imm,(existT _ _stringappend_426_ _)) =>
+ returnm (imm, build_ex _stringappend_426_)
+ | _ => exit tt : M ((mword 13 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 13 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_426_ _) =>
+ (match (string_drop _stringappend_425_ (build_ex _stringappend_426_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
- match (itype_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_628_,(existT _ _stringappend_629_ _)) =>
- let _stringappend_630_ :=
- string_drop _stringappend_559_ (build_ex _stringappend_629_) in
- match (spc_matches_prefix _stringappend_630_) with
- | Some (_stringappend_631_,(existT _ _stringappend_632_ _)) =>
- let _stringappend_633_ :=
- string_drop _stringappend_630_ (build_ex _stringappend_632_) in
- match (reg_name_matches_prefix _stringappend_633_) with
- | Some (_stringappend_634_,(existT _ _stringappend_635_ _)) =>
- let _stringappend_636_ :=
- string_drop _stringappend_633_ (build_ex _stringappend_635_) in
- sep_matches_prefix _stringappend_636_ >>= fun w__80 : option ((unit * {n : Z & ArithFact (n >=
+ match (itype_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_428_ _)) =>
+ let _stringappend_429_ :=
+ string_drop _stringappend_381_ (build_ex _stringappend_428_) in
+ match (spc_matches_prefix _stringappend_429_) with
+ | Some (tt,(existT _ _stringappend_430_ _)) =>
+ let _stringappend_431_ :=
+ string_drop _stringappend_429_ (build_ex _stringappend_430_) in
+ match (reg_name_matches_prefix _stringappend_431_) with
+ | Some (rd,(existT _ _stringappend_432_ _)) =>
+ let _stringappend_433_ :=
+ string_drop _stringappend_431_ (build_ex _stringappend_432_) in
+ sep_matches_prefix _stringappend_433_ >>= fun w__80 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__80 with
- | Some (_stringappend_637_,(existT _ _stringappend_638_ _)) =>
- let _stringappend_639_ :=
- string_drop _stringappend_636_ (build_ex _stringappend_638_) in
- match (reg_name_matches_prefix _stringappend_639_) with
- | Some (_stringappend_640_,(existT _ _stringappend_641_ _)) =>
- let _stringappend_642_ :=
- string_drop _stringappend_639_ (build_ex _stringappend_641_) in
- sep_matches_prefix _stringappend_642_ >>= fun w__81 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_434_ _)) =>
+ let _stringappend_435_ :=
+ string_drop _stringappend_433_ (build_ex _stringappend_434_) in
+ match (reg_name_matches_prefix _stringappend_435_) with
+ | Some (rs1,(existT _ _stringappend_436_ _)) =>
+ let _stringappend_437_ :=
+ string_drop _stringappend_435_ (build_ex _stringappend_436_) in
+ sep_matches_prefix _stringappend_437_ >>= fun w__81 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__81 with
- | Some
- (_stringappend_643_,(existT _ _stringappend_644_ _)) =>
- let _stringappend_645_ :=
- string_drop _stringappend_642_
- (build_ex
- _stringappend_644_) in
+ | Some (tt,(existT _ _stringappend_438_ _)) =>
+ let _stringappend_439_ :=
+ string_drop _stringappend_437_
+ (build_ex _stringappend_438_) in
if ((match (hex_bits_12_matches_prefix
- _stringappend_645_) with
- | Some
- (_stringappend_646_,(existT _ _stringappend_647_ _)) =>
- match (string_drop _stringappend_645_
- (build_ex
- _stringappend_647_)) with
+ _stringappend_439_) with
+ | Some (imm,(existT _ _stringappend_440_ _)) =>
+ match (string_drop _stringappend_439_
+ (build_ex _stringappend_440_)) with
| "" => true
| _ => false
end
@@ -18192,126 +17948,93 @@ Definition assembly_backwards_matches (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__86 : bool =>
(if (w__86) then
- match (itype_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_628_,(existT _ _stringappend_629_ _)) =>
- returnm ((_stringappend_628_, build_ex _stringappend_629_)
- : (iop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((iop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__88 : (iop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_629_ _) :=
- w__88
- : (iop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_630_ :=
- string_drop _stringappend_559_ (build_ex _stringappend_629_) in
- match (spc_matches_prefix _stringappend_630_) with
- | Some (_stringappend_631_,(existT _ _stringappend_632_ _)) =>
- returnm ((_stringappend_631_, build_ex _stringappend_632_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__90 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_632_ _) :=
- w__90
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_633_ :=
- string_drop _stringappend_630_ (build_ex _stringappend_632_) in
- match (reg_name_matches_prefix _stringappend_633_) with
- | Some (_stringappend_634_,(existT _ _stringappend_635_ _)) =>
- returnm ((_stringappend_634_, build_ex _stringappend_635_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__92 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_635_ _) :=
- w__92
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_636_ :=
- string_drop _stringappend_633_ (build_ex _stringappend_635_) in
- sep_matches_prefix _stringappend_636_ >>= fun w__93 : option ((unit * {n : Z & ArithFact (n >=
+ (match (itype_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_428_ _)) =>
+ returnm (op, build_ex _stringappend_428_)
+ | _ => exit tt : M ((iop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((iop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_428_ _) =>
+ let _stringappend_429_ :=
+ string_drop _stringappend_381_ (build_ex _stringappend_428_) in
+ (match (spc_matches_prefix _stringappend_429_) with
+ | Some (tt,(existT _ _stringappend_430_ _)) =>
+ returnm (tt, build_ex _stringappend_430_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_430_ _) =>
+ let _stringappend_431_ :=
+ string_drop _stringappend_429_ (build_ex _stringappend_430_) in
+ (match (reg_name_matches_prefix _stringappend_431_) with
+ | Some (rd,(existT _ _stringappend_432_ _)) =>
+ returnm (rd, build_ex _stringappend_432_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_432_ _) =>
+ let _stringappend_433_ :=
+ string_drop _stringappend_431_ (build_ex _stringappend_432_) in
+ sep_matches_prefix _stringappend_433_ >>= fun w__93 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__93 with
- | Some (_stringappend_637_,(existT _ _stringappend_638_ _)) =>
- returnm ((_stringappend_637_, build_ex _stringappend_638_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__95 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_638_ _) :=
- w__95
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_639_ :=
- string_drop _stringappend_636_ (build_ex _stringappend_638_) in
- match (reg_name_matches_prefix _stringappend_639_) with
- | Some (_stringappend_640_,(existT _ _stringappend_641_ _)) =>
- returnm ((_stringappend_640_, build_ex _stringappend_641_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__97 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_641_ _) :=
- w__97
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_642_ :=
- string_drop _stringappend_639_ (build_ex _stringappend_641_) in
- sep_matches_prefix _stringappend_642_ >>= fun w__98 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__93 with
+ | Some (tt,(existT _ _stringappend_434_ _)) =>
+ returnm (tt, build_ex _stringappend_434_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_434_ _) =>
+ let _stringappend_435_ :=
+ string_drop _stringappend_433_ (build_ex _stringappend_434_) in
+ (match (reg_name_matches_prefix _stringappend_435_) with
+ | Some (rs1,(existT _ _stringappend_436_ _)) =>
+ returnm (rs1, build_ex _stringappend_436_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_436_ _) =>
+ let _stringappend_437_ :=
+ string_drop _stringappend_435_ (build_ex _stringappend_436_) in
+ sep_matches_prefix _stringappend_437_ >>= fun w__98 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__98 with
- | Some (_stringappend_643_,(existT _ _stringappend_644_ _)) =>
- returnm ((_stringappend_643_, build_ex _stringappend_644_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__100 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_644_ _) :=
- w__100
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_645_ :=
- string_drop _stringappend_642_ (build_ex _stringappend_644_) in
- match (hex_bits_12_matches_prefix _stringappend_645_) with
- | Some (_stringappend_646_,(existT _ _stringappend_647_ _)) =>
- returnm ((_stringappend_646_, build_ex _stringappend_647_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__102 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_647_ _) :=
- w__102
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_645_ (build_ex _stringappend_647_)) with
+ (match w__98 with
+ | Some (tt,(existT _ _stringappend_438_ _)) =>
+ returnm (tt, build_ex _stringappend_438_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_438_ _) =>
+ let _stringappend_439_ :=
+ string_drop _stringappend_437_ (build_ex _stringappend_438_) in
+ (match (hex_bits_12_matches_prefix _stringappend_439_) with
+ | Some (imm,(existT _ _stringappend_440_ _)) =>
+ returnm (imm, build_ex _stringappend_440_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_440_ _) =>
+ (match (string_drop _stringappend_439_ (build_ex _stringappend_440_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
- match (shiftiop_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_649_,(existT _ _stringappend_650_ _)) =>
- let _stringappend_651_ :=
- string_drop _stringappend_559_ (build_ex _stringappend_650_) in
- match (spc_matches_prefix _stringappend_651_) with
- | Some (_stringappend_652_,(existT _ _stringappend_653_ _)) =>
- let _stringappend_654_ :=
- string_drop _stringappend_651_ (build_ex _stringappend_653_) in
- match (reg_name_matches_prefix _stringappend_654_) with
- | Some (_stringappend_655_,(existT _ _stringappend_656_ _)) =>
- let _stringappend_657_ :=
- string_drop _stringappend_654_ (build_ex _stringappend_656_) in
- sep_matches_prefix _stringappend_657_ >>= fun w__105 : option ((unit * {n : Z & ArithFact (n >=
+ match (shiftiop_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_442_ _)) =>
+ let _stringappend_443_ :=
+ string_drop _stringappend_381_ (build_ex _stringappend_442_) in
+ match (spc_matches_prefix _stringappend_443_) with
+ | Some (tt,(existT _ _stringappend_444_ _)) =>
+ let _stringappend_445_ :=
+ string_drop _stringappend_443_ (build_ex _stringappend_444_) in
+ match (reg_name_matches_prefix _stringappend_445_) with
+ | Some (rd,(existT _ _stringappend_446_ _)) =>
+ let _stringappend_447_ :=
+ string_drop _stringappend_445_ (build_ex _stringappend_446_) in
+ sep_matches_prefix _stringappend_447_ >>= fun w__105 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__105 with
- | Some
- (_stringappend_658_,(existT _ _stringappend_659_ _)) =>
- let _stringappend_660_ :=
- string_drop _stringappend_657_
- (build_ex
- _stringappend_659_) in
- if ((match (reg_name_matches_prefix _stringappend_660_) with
- | Some
- (_stringappend_661_,(existT _ _stringappend_662_ _)) =>
- let _stringappend_663_ :=
- string_drop _stringappend_660_
- (build_ex
- _stringappend_662_) in
+ | Some (tt,(existT _ _stringappend_448_ _)) =>
+ let _stringappend_449_ :=
+ string_drop _stringappend_447_
+ (build_ex _stringappend_448_) in
+ if ((match (reg_name_matches_prefix _stringappend_449_) with
+ | Some (rs1,(existT _ _stringappend_450_ _)) =>
+ let _stringappend_451_ :=
+ string_drop _stringappend_449_
+ (build_ex _stringappend_450_) in
if ((match (hex_bits_6_matches_prefix
- _stringappend_663_) with
+ _stringappend_451_) with
| Some
- (_stringappend_664_,(existT _ _stringappend_665_ _)) =>
- match (string_drop _stringappend_663_
- (build_ex
- _stringappend_665_)) with
+ (shamt,(existT _ _stringappend_452_ _)) =>
+ match (string_drop _stringappend_451_
+ (build_ex _stringappend_452_)) with
| "" => true
| _ => false
end
@@ -18341,116 +18064,89 @@ Definition assembly_backwards_matches (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__108 : bool =>
(if (w__108) then
- match (shiftiop_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_649_,(existT _ _stringappend_650_ _)) =>
- returnm ((_stringappend_649_, build_ex _stringappend_650_)
- : (sop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((sop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__110 : (sop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_650_ _) :=
- w__110
- : (sop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_651_ :=
- string_drop _stringappend_559_ (build_ex _stringappend_650_) in
- match (spc_matches_prefix _stringappend_651_) with
- | Some (_stringappend_652_,(existT _ _stringappend_653_ _)) =>
- returnm ((_stringappend_652_, build_ex _stringappend_653_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__112 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_653_ _) :=
- w__112
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_654_ :=
- string_drop _stringappend_651_ (build_ex _stringappend_653_) in
- match (reg_name_matches_prefix _stringappend_654_) with
- | Some (_stringappend_655_,(existT _ _stringappend_656_ _)) =>
- returnm ((_stringappend_655_, build_ex _stringappend_656_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__114 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_656_ _) :=
- w__114
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_657_ :=
- string_drop _stringappend_654_ (build_ex _stringappend_656_) in
- sep_matches_prefix _stringappend_657_ >>= fun w__115 : option ((unit * {n : Z & ArithFact (n >=
+ (match (shiftiop_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_442_ _)) =>
+ returnm (op, build_ex _stringappend_442_)
+ | _ => exit tt : M ((sop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((sop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_442_ _) =>
+ let _stringappend_443_ :=
+ string_drop _stringappend_381_ (build_ex _stringappend_442_) in
+ (match (spc_matches_prefix _stringappend_443_) with
+ | Some (tt,(existT _ _stringappend_444_ _)) =>
+ returnm (tt, build_ex _stringappend_444_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_444_ _) =>
+ let _stringappend_445_ :=
+ string_drop _stringappend_443_ (build_ex _stringappend_444_) in
+ (match (reg_name_matches_prefix _stringappend_445_) with
+ | Some (rd,(existT _ _stringappend_446_ _)) =>
+ returnm (rd, build_ex _stringappend_446_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_446_ _) =>
+ let _stringappend_447_ :=
+ string_drop _stringappend_445_ (build_ex _stringappend_446_) in
+ sep_matches_prefix _stringappend_447_ >>= fun w__115 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__115 with
- | Some (_stringappend_658_,(existT _ _stringappend_659_ _)) =>
- returnm ((_stringappend_658_, build_ex _stringappend_659_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__117 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_659_ _) :=
- w__117
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_660_ :=
- string_drop _stringappend_657_ (build_ex _stringappend_659_) in
- match (reg_name_matches_prefix _stringappend_660_) with
- | Some (_stringappend_661_,(existT _ _stringappend_662_ _)) =>
- returnm ((_stringappend_661_, build_ex _stringappend_662_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__119 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_662_ _) :=
- w__119
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_663_ :=
- string_drop _stringappend_660_ (build_ex _stringappend_662_) in
- match (hex_bits_6_matches_prefix _stringappend_663_) with
- | Some (_stringappend_664_,(existT _ _stringappend_665_ _)) =>
- returnm ((_stringappend_664_, build_ex _stringappend_665_)
- : (mword 6 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 6 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__121 : (mword 6 * {n : Z & ArithFact (n >= 0)}) =>
- let '(shamt, existT _ _stringappend_665_ _) :=
- w__121
- : (mword 6 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_663_ (build_ex _stringappend_665_)) with
+ (match w__115 with
+ | Some (tt,(existT _ _stringappend_448_ _)) =>
+ returnm (tt, build_ex _stringappend_448_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_448_ _) =>
+ let _stringappend_449_ :=
+ string_drop _stringappend_447_ (build_ex _stringappend_448_) in
+ (match (reg_name_matches_prefix _stringappend_449_) with
+ | Some (rs1,(existT _ _stringappend_450_ _)) =>
+ returnm (rs1, build_ex _stringappend_450_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_450_ _) =>
+ let _stringappend_451_ :=
+ string_drop _stringappend_449_ (build_ex _stringappend_450_) in
+ (match (hex_bits_6_matches_prefix _stringappend_451_) with
+ | Some (shamt,(existT _ _stringappend_452_ _)) =>
+ returnm (shamt, build_ex _stringappend_452_)
+ | _ => exit tt : M ((mword 6 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 6 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(shamt, existT _ _stringappend_452_ _) =>
+ (match (string_drop _stringappend_451_ (build_ex _stringappend_452_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
- match (rtype_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_667_,(existT _ _stringappend_668_ _)) =>
- let _stringappend_669_ :=
- string_drop _stringappend_559_ (build_ex _stringappend_668_) in
- match (spc_matches_prefix _stringappend_669_) with
- | Some (_stringappend_670_,(existT _ _stringappend_671_ _)) =>
- let _stringappend_672_ :=
- string_drop _stringappend_669_ (build_ex _stringappend_671_) in
- match (reg_name_matches_prefix _stringappend_672_) with
- | Some (_stringappend_673_,(existT _ _stringappend_674_ _)) =>
- let _stringappend_675_ :=
- string_drop _stringappend_672_ (build_ex _stringappend_674_) in
- sep_matches_prefix _stringappend_675_ >>= fun w__124 : option ((unit * {n : Z & ArithFact (n >=
+ match (rtype_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_454_ _)) =>
+ let _stringappend_455_ :=
+ string_drop _stringappend_381_ (build_ex _stringappend_454_) in
+ match (spc_matches_prefix _stringappend_455_) with
+ | Some (tt,(existT _ _stringappend_456_ _)) =>
+ let _stringappend_457_ :=
+ string_drop _stringappend_455_ (build_ex _stringappend_456_) in
+ match (reg_name_matches_prefix _stringappend_457_) with
+ | Some (rd,(existT _ _stringappend_458_ _)) =>
+ let _stringappend_459_ :=
+ string_drop _stringappend_457_ (build_ex _stringappend_458_) in
+ sep_matches_prefix _stringappend_459_ >>= fun w__124 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__124 with
- | Some (_stringappend_676_,(existT _ _stringappend_677_ _)) =>
- let _stringappend_678_ :=
- string_drop _stringappend_675_ (build_ex _stringappend_677_) in
- match (reg_name_matches_prefix _stringappend_678_) with
- | Some (_stringappend_679_,(existT _ _stringappend_680_ _)) =>
- let _stringappend_681_ :=
- string_drop _stringappend_678_ (build_ex _stringappend_680_) in
- sep_matches_prefix _stringappend_681_ >>= fun w__125 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_460_ _)) =>
+ let _stringappend_461_ :=
+ string_drop _stringappend_459_ (build_ex _stringappend_460_) in
+ match (reg_name_matches_prefix _stringappend_461_) with
+ | Some (rs1,(existT _ _stringappend_462_ _)) =>
+ let _stringappend_463_ :=
+ string_drop _stringappend_461_ (build_ex _stringappend_462_) in
+ sep_matches_prefix _stringappend_463_ >>= fun w__125 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__125 with
- | Some
- (_stringappend_682_,(existT _ _stringappend_683_ _)) =>
- let _stringappend_684_ :=
- string_drop _stringappend_681_
- (build_ex
- _stringappend_683_) in
+ | Some (tt,(existT _ _stringappend_464_ _)) =>
+ let _stringappend_465_ :=
+ string_drop _stringappend_463_
+ (build_ex _stringappend_464_) in
if ((match (reg_name_matches_prefix
- _stringappend_684_) with
+ _stringappend_465_) with
| Some
- (_stringappend_685_,(existT _ _stringappend_686_ _)) =>
- match (string_drop _stringappend_684_
- (build_ex
- _stringappend_686_)) with
+ (rs2,(existT _ _stringappend_466_ _)) =>
+ match (string_drop _stringappend_465_
+ (build_ex _stringappend_466_)) with
| "" => true
| _ => false
end
@@ -18486,155 +18182,119 @@ Definition assembly_backwards_matches (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__130 : bool =>
(if (w__130) then
- match (rtype_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_667_,(existT _ _stringappend_668_ _)) =>
- returnm ((_stringappend_667_, build_ex _stringappend_668_)
- : (rop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((rop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__132 : (rop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_668_ _) :=
- w__132
- : (rop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_669_ :=
- string_drop _stringappend_559_ (build_ex _stringappend_668_) in
- match (spc_matches_prefix _stringappend_669_) with
- | Some (_stringappend_670_,(existT _ _stringappend_671_ _)) =>
- returnm ((_stringappend_670_, build_ex _stringappend_671_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__134 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_671_ _) :=
- w__134
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_672_ :=
- string_drop _stringappend_669_ (build_ex _stringappend_671_) in
- match (reg_name_matches_prefix _stringappend_672_) with
- | Some (_stringappend_673_,(existT _ _stringappend_674_ _)) =>
- returnm ((_stringappend_673_, build_ex _stringappend_674_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__136 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_674_ _) :=
- w__136
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_675_ :=
- string_drop _stringappend_672_ (build_ex _stringappend_674_) in
- sep_matches_prefix _stringappend_675_ >>= fun w__137 : option ((unit * {n : Z & ArithFact (n >=
+ (match (rtype_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_454_ _)) =>
+ returnm (op, build_ex _stringappend_454_)
+ | _ => exit tt : M ((rop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((rop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_454_ _) =>
+ let _stringappend_455_ :=
+ string_drop _stringappend_381_ (build_ex _stringappend_454_) in
+ (match (spc_matches_prefix _stringappend_455_) with
+ | Some (tt,(existT _ _stringappend_456_ _)) =>
+ returnm (tt, build_ex _stringappend_456_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_456_ _) =>
+ let _stringappend_457_ :=
+ string_drop _stringappend_455_ (build_ex _stringappend_456_) in
+ (match (reg_name_matches_prefix _stringappend_457_) with
+ | Some (rd,(existT _ _stringappend_458_ _)) =>
+ returnm (rd, build_ex _stringappend_458_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_458_ _) =>
+ let _stringappend_459_ :=
+ string_drop _stringappend_457_ (build_ex _stringappend_458_) in
+ sep_matches_prefix _stringappend_459_ >>= fun w__137 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__137 with
- | Some (_stringappend_676_,(existT _ _stringappend_677_ _)) =>
- returnm ((_stringappend_676_, build_ex _stringappend_677_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__139 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_677_ _) :=
- w__139
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_678_ :=
- string_drop _stringappend_675_ (build_ex _stringappend_677_) in
- match (reg_name_matches_prefix _stringappend_678_) with
- | Some (_stringappend_679_,(existT _ _stringappend_680_ _)) =>
- returnm ((_stringappend_679_, build_ex _stringappend_680_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__141 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_680_ _) :=
- w__141
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_681_ :=
- string_drop _stringappend_678_ (build_ex _stringappend_680_) in
- sep_matches_prefix _stringappend_681_ >>= fun w__142 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__137 with
+ | Some (tt,(existT _ _stringappend_460_ _)) =>
+ returnm (tt, build_ex _stringappend_460_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_460_ _) =>
+ let _stringappend_461_ :=
+ string_drop _stringappend_459_ (build_ex _stringappend_460_) in
+ (match (reg_name_matches_prefix _stringappend_461_) with
+ | Some (rs1,(existT _ _stringappend_462_ _)) =>
+ returnm (rs1, build_ex _stringappend_462_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_462_ _) =>
+ let _stringappend_463_ :=
+ string_drop _stringappend_461_ (build_ex _stringappend_462_) in
+ sep_matches_prefix _stringappend_463_ >>= fun w__142 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__142 with
- | Some (_stringappend_682_,(existT _ _stringappend_683_ _)) =>
- returnm ((_stringappend_682_, build_ex _stringappend_683_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__144 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_683_ _) :=
- w__144
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_684_ :=
- string_drop _stringappend_681_ (build_ex _stringappend_683_) in
- match (reg_name_matches_prefix _stringappend_684_) with
- | Some (_stringappend_685_,(existT _ _stringappend_686_ _)) =>
- returnm ((_stringappend_685_, build_ex _stringappend_686_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__146 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs2, existT _ _stringappend_686_ _) :=
- w__146
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_684_ (build_ex _stringappend_686_)) with
+ (match w__142 with
+ | Some (tt,(existT _ _stringappend_464_ _)) =>
+ returnm (tt, build_ex _stringappend_464_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_464_ _) =>
+ let _stringappend_465_ :=
+ string_drop _stringappend_463_ (build_ex _stringappend_464_) in
+ (match (reg_name_matches_prefix _stringappend_465_) with
+ | Some (rs2,(existT _ _stringappend_466_ _)) =>
+ returnm (rs2, build_ex _stringappend_466_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_466_ _) =>
+ (match (string_drop _stringappend_465_ (build_ex _stringappend_466_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
- and_boolM (returnm ((string_startswith _stringappend_559_ "l") : bool))
- (let _stringappend_688_ :=
- string_drop _stringappend_559_ (string_length "l") in
- match (size_mnemonic_matches_prefix _stringappend_688_) with
- | Some (_stringappend_689_,(existT _ _stringappend_690_ _)) =>
- let _stringappend_691_ :=
- string_drop _stringappend_688_ (build_ex _stringappend_690_) in
- match (maybe_u_matches_prefix _stringappend_691_) with
- | Some (_stringappend_692_,(existT _ _stringappend_693_ _)) =>
- let _stringappend_694_ :=
- string_drop _stringappend_691_ (build_ex _stringappend_693_) in
- match (maybe_aq_matches_prefix _stringappend_694_) with
- | Some (_stringappend_695_,(existT _ _stringappend_696_ _)) =>
- let _stringappend_697_ :=
- string_drop _stringappend_694_ (build_ex _stringappend_696_) in
- match (maybe_rl_matches_prefix _stringappend_697_) with
- | Some (_stringappend_698_,(existT _ _stringappend_699_ _)) =>
- let _stringappend_700_ :=
- string_drop _stringappend_697_ (build_ex _stringappend_699_) in
- match (spc_matches_prefix _stringappend_700_) with
- | Some (_stringappend_701_,(existT _ _stringappend_702_ _)) =>
- let _stringappend_703_ :=
- string_drop _stringappend_700_
- (build_ex
- _stringappend_702_) in
- match (reg_name_matches_prefix _stringappend_703_) with
- | Some (_stringappend_704_,(existT _ _stringappend_705_ _)) =>
- let _stringappend_706_ :=
- string_drop _stringappend_703_
- (build_ex
- _stringappend_705_) in
- sep_matches_prefix _stringappend_706_ >>= fun w__149 : option ((unit * {n : Z & ArithFact (n >=
+ and_boolM (returnm ((string_startswith _stringappend_381_ "l") : bool))
+ (let _stringappend_468_ :=
+ string_drop _stringappend_381_ (build_ex (projT1 (string_length "l"))) in
+ match (size_mnemonic_matches_prefix _stringappend_468_) with
+ | Some (size,(existT _ _stringappend_469_ _)) =>
+ let _stringappend_470_ :=
+ string_drop _stringappend_468_ (build_ex _stringappend_469_) in
+ match (maybe_u_matches_prefix _stringappend_470_) with
+ | Some (is_unsigned,(existT _ _stringappend_471_ _)) =>
+ let _stringappend_472_ :=
+ string_drop _stringappend_470_ (build_ex _stringappend_471_) in
+ match (maybe_aq_matches_prefix _stringappend_472_) with
+ | Some (aq,(existT _ _stringappend_473_ _)) =>
+ let _stringappend_474_ :=
+ string_drop _stringappend_472_ (build_ex _stringappend_473_) in
+ match (maybe_rl_matches_prefix _stringappend_474_) with
+ | Some (rl,(existT _ _stringappend_475_ _)) =>
+ let _stringappend_476_ :=
+ string_drop _stringappend_474_ (build_ex _stringappend_475_) in
+ match (spc_matches_prefix _stringappend_476_) with
+ | Some (tt,(existT _ _stringappend_477_ _)) =>
+ let _stringappend_478_ :=
+ string_drop _stringappend_476_
+ (build_ex _stringappend_477_) in
+ match (reg_name_matches_prefix _stringappend_478_) with
+ | Some (rd,(existT _ _stringappend_479_ _)) =>
+ let _stringappend_480_ :=
+ string_drop _stringappend_478_
+ (build_ex _stringappend_479_) in
+ sep_matches_prefix _stringappend_480_ >>= fun w__149 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__149 with
- | Some
- (_stringappend_707_,(existT _ _stringappend_708_ _)) =>
- let _stringappend_709_ :=
- string_drop _stringappend_706_
- (build_ex
- _stringappend_708_) in
- match (reg_name_matches_prefix _stringappend_709_) with
- | Some
- (_stringappend_710_,(existT _ _stringappend_711_ _)) =>
- let _stringappend_712_ :=
- string_drop _stringappend_709_
- (build_ex
- _stringappend_711_) in
- sep_matches_prefix _stringappend_712_ >>= fun w__150 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_481_ _)) =>
+ let _stringappend_482_ :=
+ string_drop _stringappend_480_
+ (build_ex _stringappend_481_) in
+ match (reg_name_matches_prefix _stringappend_482_) with
+ | Some (rs1,(existT _ _stringappend_483_ _)) =>
+ let _stringappend_484_ :=
+ string_drop _stringappend_482_
+ (build_ex _stringappend_483_) in
+ sep_matches_prefix _stringappend_484_ >>= fun w__150 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__150 with
| Some
- (_stringappend_713_,(existT _ _stringappend_714_ _)) =>
- let _stringappend_715_ :=
- string_drop _stringappend_712_
- (build_ex
- _stringappend_714_) in
+ (tt,(existT _ _stringappend_485_ _)) =>
+ let _stringappend_486_ :=
+ string_drop _stringappend_484_
+ (build_ex _stringappend_485_) in
if ((match (hex_bits_12_matches_prefix
- _stringappend_715_) with
+ _stringappend_486_) with
| Some
- (_stringappend_716_,(existT _ _stringappend_717_ _)) =>
+ (imm,(existT _ _stringappend_487_ _)) =>
match (string_drop
- _stringappend_715_
- (build_ex
- _stringappend_717_)) with
+ _stringappend_486_
+ (build_ex _stringappend_487_)) with
| "" => true
| _ => false
end
@@ -18688,186 +18348,139 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__159 : bool =>
(if (w__159) then
- let _stringappend_688_ :=
- string_drop _stringappend_559_ (string_length "l") in
- match (size_mnemonic_matches_prefix _stringappend_688_) with
- | Some (_stringappend_689_,(existT _ _stringappend_690_ _)) =>
- returnm ((_stringappend_689_, build_ex _stringappend_690_)
- : (word_width * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((word_width * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__161 : (word_width * {n : Z & ArithFact (n >= 0)}) =>
- let '(size, existT _ _stringappend_690_ _) :=
- w__161
- : (word_width * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_691_ :=
- string_drop _stringappend_688_ (build_ex _stringappend_690_) in
- match (maybe_u_matches_prefix _stringappend_691_) with
- | Some (_stringappend_692_,(existT _ _stringappend_693_ _)) =>
- returnm ((_stringappend_692_, build_ex _stringappend_693_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__163 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(is_unsigned, existT _ _stringappend_693_ _) :=
- w__163
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_694_ :=
- string_drop _stringappend_691_ (build_ex _stringappend_693_) in
- match (maybe_aq_matches_prefix _stringappend_694_) with
- | Some (_stringappend_695_,(existT _ _stringappend_696_ _)) =>
- returnm ((_stringappend_695_, build_ex _stringappend_696_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__165 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(aq, existT _ _stringappend_696_ _) :=
- w__165
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_697_ :=
- string_drop _stringappend_694_ (build_ex _stringappend_696_) in
- match (maybe_rl_matches_prefix _stringappend_697_) with
- | Some (_stringappend_698_,(existT _ _stringappend_699_ _)) =>
- returnm ((_stringappend_698_, build_ex _stringappend_699_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__167 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(rl, existT _ _stringappend_699_ _) :=
- w__167
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_700_ :=
- string_drop _stringappend_697_ (build_ex _stringappend_699_) in
- match (spc_matches_prefix _stringappend_700_) with
- | Some (_stringappend_701_,(existT _ _stringappend_702_ _)) =>
- returnm ((_stringappend_701_, build_ex _stringappend_702_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__169 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_702_ _) :=
- w__169
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_703_ :=
- string_drop _stringappend_700_ (build_ex _stringappend_702_) in
- match (reg_name_matches_prefix _stringappend_703_) with
- | Some (_stringappend_704_,(existT _ _stringappend_705_ _)) =>
- returnm ((_stringappend_704_, build_ex _stringappend_705_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__171 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_705_ _) :=
- w__171
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_706_ :=
- string_drop _stringappend_703_ (build_ex _stringappend_705_) in
- sep_matches_prefix _stringappend_706_ >>= fun w__172 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_468_ :=
+ string_drop _stringappend_381_ (build_ex (projT1 (string_length "l"))) in
+ (match (size_mnemonic_matches_prefix _stringappend_468_) with
+ | Some (size,(existT _ _stringappend_469_ _)) =>
+ returnm (size, build_ex _stringappend_469_)
+ | _ => exit tt : M ((word_width * {n : Z & ArithFact (n >= 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >= 0)}))) >>= fun '(size, existT _ _stringappend_469_ _) =>
+ let _stringappend_470_ :=
+ string_drop _stringappend_468_ (build_ex _stringappend_469_) in
+ (match (maybe_u_matches_prefix _stringappend_470_) with
+ | Some (is_unsigned,(existT _ _stringappend_471_ _)) =>
+ returnm (is_unsigned, build_ex _stringappend_471_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(is_unsigned, existT _ _stringappend_471_ _) =>
+ let _stringappend_472_ :=
+ string_drop _stringappend_470_ (build_ex _stringappend_471_) in
+ (match (maybe_aq_matches_prefix _stringappend_472_) with
+ | Some (aq,(existT _ _stringappend_473_ _)) =>
+ returnm (aq, build_ex _stringappend_473_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(aq, existT _ _stringappend_473_ _) =>
+ let _stringappend_474_ :=
+ string_drop _stringappend_472_ (build_ex _stringappend_473_) in
+ (match (maybe_rl_matches_prefix _stringappend_474_) with
+ | Some (rl,(existT _ _stringappend_475_ _)) =>
+ returnm (rl, build_ex _stringappend_475_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rl, existT _ _stringappend_475_ _) =>
+ let _stringappend_476_ :=
+ string_drop _stringappend_474_ (build_ex _stringappend_475_) in
+ (match (spc_matches_prefix _stringappend_476_) with
+ | Some (tt,(existT _ _stringappend_477_ _)) =>
+ returnm (tt, build_ex _stringappend_477_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_477_ _) =>
+ let _stringappend_478_ :=
+ string_drop _stringappend_476_ (build_ex _stringappend_477_) in
+ (match (reg_name_matches_prefix _stringappend_478_) with
+ | Some (rd,(existT _ _stringappend_479_ _)) =>
+ returnm (rd, build_ex _stringappend_479_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_479_ _) =>
+ let _stringappend_480_ :=
+ string_drop _stringappend_478_ (build_ex _stringappend_479_) in
+ sep_matches_prefix _stringappend_480_ >>= fun w__172 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__172 with
- | Some (_stringappend_707_,(existT _ _stringappend_708_ _)) =>
- returnm ((_stringappend_707_, build_ex _stringappend_708_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__174 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_708_ _) :=
- w__174
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_709_ :=
- string_drop _stringappend_706_ (build_ex _stringappend_708_) in
- match (reg_name_matches_prefix _stringappend_709_) with
- | Some (_stringappend_710_,(existT _ _stringappend_711_ _)) =>
- returnm ((_stringappend_710_, build_ex _stringappend_711_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__176 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_711_ _) :=
- w__176
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_712_ :=
- string_drop _stringappend_709_ (build_ex _stringappend_711_) in
- sep_matches_prefix _stringappend_712_ >>= fun w__177 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__172 with
+ | Some (tt,(existT _ _stringappend_481_ _)) =>
+ returnm (tt, build_ex _stringappend_481_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_481_ _) =>
+ let _stringappend_482_ :=
+ string_drop _stringappend_480_ (build_ex _stringappend_481_) in
+ (match (reg_name_matches_prefix _stringappend_482_) with
+ | Some (rs1,(existT _ _stringappend_483_ _)) =>
+ returnm (rs1, build_ex _stringappend_483_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_483_ _) =>
+ let _stringappend_484_ :=
+ string_drop _stringappend_482_ (build_ex _stringappend_483_) in
+ sep_matches_prefix _stringappend_484_ >>= fun w__177 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__177 with
- | Some (_stringappend_713_,(existT _ _stringappend_714_ _)) =>
- returnm ((_stringappend_713_, build_ex _stringappend_714_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__179 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_714_ _) :=
- w__179
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_715_ :=
- string_drop _stringappend_712_ (build_ex _stringappend_714_) in
- match (hex_bits_12_matches_prefix _stringappend_715_) with
- | Some (_stringappend_716_,(existT _ _stringappend_717_ _)) =>
- returnm ((_stringappend_716_, build_ex _stringappend_717_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__181 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_717_ _) :=
- w__181
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_715_ (build_ex _stringappend_717_)) with
+ (match w__177 with
+ | Some (tt,(existT _ _stringappend_485_ _)) =>
+ returnm (tt, build_ex _stringappend_485_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_485_ _) =>
+ let _stringappend_486_ :=
+ string_drop _stringappend_484_ (build_ex _stringappend_485_) in
+ (match (hex_bits_12_matches_prefix _stringappend_486_) with
+ | Some (imm,(existT _ _stringappend_487_ _)) =>
+ returnm (imm, build_ex _stringappend_487_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_487_ _) =>
+ (match (string_drop _stringappend_486_ (build_ex _stringappend_487_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
- and_boolM (returnm ((string_startswith _stringappend_559_ "s") : bool))
- (let _stringappend_719_ :=
- string_drop _stringappend_559_ (string_length "s") in
- match (size_mnemonic_matches_prefix _stringappend_719_) with
- | Some (_stringappend_720_,(existT _ _stringappend_721_ _)) =>
- let _stringappend_722_ :=
- string_drop _stringappend_719_ (build_ex _stringappend_721_) in
- match (maybe_aq_matches_prefix _stringappend_722_) with
- | Some (_stringappend_723_,(existT _ _stringappend_724_ _)) =>
- let _stringappend_725_ :=
- string_drop _stringappend_722_ (build_ex _stringappend_724_) in
- match (maybe_rl_matches_prefix _stringappend_725_) with
- | Some (_stringappend_726_,(existT _ _stringappend_727_ _)) =>
- let _stringappend_728_ :=
- string_drop _stringappend_725_ (build_ex _stringappend_727_) in
- match (spc_matches_prefix _stringappend_728_) with
- | Some (_stringappend_729_,(existT _ _stringappend_730_ _)) =>
- let _stringappend_731_ :=
- string_drop _stringappend_728_
- (build_ex
- _stringappend_730_) in
- match (reg_name_matches_prefix _stringappend_731_) with
- | Some (_stringappend_732_,(existT _ _stringappend_733_ _)) =>
- let _stringappend_734_ :=
- string_drop _stringappend_731_
- (build_ex
- _stringappend_733_) in
- sep_matches_prefix _stringappend_734_ >>= fun w__184 : option ((unit * {n : Z & ArithFact (n >=
+ and_boolM (returnm ((string_startswith _stringappend_381_ "s") : bool))
+ (let _stringappend_489_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length "s"))) in
+ match (size_mnemonic_matches_prefix _stringappend_489_) with
+ | Some (size,(existT _ _stringappend_490_ _)) =>
+ let _stringappend_491_ :=
+ string_drop _stringappend_489_ (build_ex _stringappend_490_) in
+ match (maybe_aq_matches_prefix _stringappend_491_) with
+ | Some (aq,(existT _ _stringappend_492_ _)) =>
+ let _stringappend_493_ :=
+ string_drop _stringappend_491_ (build_ex _stringappend_492_) in
+ match (maybe_rl_matches_prefix _stringappend_493_) with
+ | Some (rl,(existT _ _stringappend_494_ _)) =>
+ let _stringappend_495_ :=
+ string_drop _stringappend_493_ (build_ex _stringappend_494_) in
+ match (spc_matches_prefix _stringappend_495_) with
+ | Some (tt,(existT _ _stringappend_496_ _)) =>
+ let _stringappend_497_ :=
+ string_drop _stringappend_495_
+ (build_ex _stringappend_496_) in
+ match (reg_name_matches_prefix _stringappend_497_) with
+ | Some (rd,(existT _ _stringappend_498_ _)) =>
+ let _stringappend_499_ :=
+ string_drop _stringappend_497_
+ (build_ex _stringappend_498_) in
+ sep_matches_prefix _stringappend_499_ >>= fun w__184 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__184 with
- | Some
- (_stringappend_735_,(existT _ _stringappend_736_ _)) =>
- let _stringappend_737_ :=
- string_drop _stringappend_734_
- (build_ex
- _stringappend_736_) in
- match (reg_name_matches_prefix _stringappend_737_) with
- | Some
- (_stringappend_738_,(existT _ _stringappend_739_ _)) =>
- let _stringappend_740_ :=
- string_drop _stringappend_737_
- (build_ex
- _stringappend_739_) in
- sep_matches_prefix _stringappend_740_ >>= fun w__185 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_500_ _)) =>
+ let _stringappend_501_ :=
+ string_drop _stringappend_499_
+ (build_ex _stringappend_500_) in
+ match (reg_name_matches_prefix _stringappend_501_) with
+ | Some (rs1,(existT _ _stringappend_502_ _)) =>
+ let _stringappend_503_ :=
+ string_drop _stringappend_501_
+ (build_ex _stringappend_502_) in
+ sep_matches_prefix _stringappend_503_ >>= fun w__185 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__185 with
| Some
- (_stringappend_741_,(existT _ _stringappend_742_ _)) =>
- let _stringappend_743_ :=
- string_drop _stringappend_740_
- (build_ex
- _stringappend_742_) in
+ (tt,(existT _ _stringappend_504_ _)) =>
+ let _stringappend_505_ :=
+ string_drop _stringappend_503_
+ (build_ex _stringappend_504_) in
if ((match (hex_bits_12_matches_prefix
- _stringappend_743_) with
+ _stringappend_505_) with
| Some
- (_stringappend_744_,(existT _ _stringappend_745_ _)) =>
+ (imm,(existT _ _stringappend_506_ _)) =>
match (string_drop
- _stringappend_743_
- (build_ex
- _stringappend_745_)) with
+ _stringappend_505_
+ (build_ex _stringappend_506_)) with
| "" => true
| _ => false
end
@@ -18916,159 +18529,121 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__193 : bool =>
(if (w__193) then
- let _stringappend_719_ :=
- string_drop _stringappend_559_ (string_length "s") in
- match (size_mnemonic_matches_prefix _stringappend_719_) with
- | Some (_stringappend_720_,(existT _ _stringappend_721_ _)) =>
- returnm ((_stringappend_720_, build_ex _stringappend_721_)
- : (word_width * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((word_width * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__195 : (word_width * {n : Z & ArithFact (n >= 0)}) =>
- let '(size, existT _ _stringappend_721_ _) :=
- w__195
- : (word_width * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_722_ :=
- string_drop _stringappend_719_ (build_ex _stringappend_721_) in
- match (maybe_aq_matches_prefix _stringappend_722_) with
- | Some (_stringappend_723_,(existT _ _stringappend_724_ _)) =>
- returnm ((_stringappend_723_, build_ex _stringappend_724_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__197 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(aq, existT _ _stringappend_724_ _) :=
- w__197
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_725_ :=
- string_drop _stringappend_722_ (build_ex _stringappend_724_) in
- match (maybe_rl_matches_prefix _stringappend_725_) with
- | Some (_stringappend_726_,(existT _ _stringappend_727_ _)) =>
- returnm ((_stringappend_726_, build_ex _stringappend_727_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__199 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(rl, existT _ _stringappend_727_ _) :=
- w__199
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_728_ :=
- string_drop _stringappend_725_ (build_ex _stringappend_727_) in
- match (spc_matches_prefix _stringappend_728_) with
- | Some (_stringappend_729_,(existT _ _stringappend_730_ _)) =>
- returnm ((_stringappend_729_, build_ex _stringappend_730_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__201 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_730_ _) :=
- w__201
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_731_ :=
- string_drop _stringappend_728_ (build_ex _stringappend_730_) in
- match (reg_name_matches_prefix _stringappend_731_) with
- | Some (_stringappend_732_,(existT _ _stringappend_733_ _)) =>
- returnm ((_stringappend_732_, build_ex _stringappend_733_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__203 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_733_ _) :=
- w__203
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_734_ :=
- string_drop _stringappend_731_ (build_ex _stringappend_733_) in
- sep_matches_prefix _stringappend_734_ >>= fun w__204 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_489_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length "s"))) in
+ (match (size_mnemonic_matches_prefix _stringappend_489_) with
+ | Some (size,(existT _ _stringappend_490_ _)) =>
+ returnm (size, build_ex _stringappend_490_)
+ | _ => exit tt : M ((word_width * {n : Z & ArithFact (n >= 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >= 0)}))) >>= fun '(size, existT _ _stringappend_490_ _) =>
+ let _stringappend_491_ :=
+ string_drop _stringappend_489_ (build_ex _stringappend_490_) in
+ (match (maybe_aq_matches_prefix _stringappend_491_) with
+ | Some (aq,(existT _ _stringappend_492_ _)) =>
+ returnm (aq, build_ex _stringappend_492_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(aq, existT _ _stringappend_492_ _) =>
+ let _stringappend_493_ :=
+ string_drop _stringappend_491_ (build_ex _stringappend_492_) in
+ (match (maybe_rl_matches_prefix _stringappend_493_) with
+ | Some (rl,(existT _ _stringappend_494_ _)) =>
+ returnm (rl, build_ex _stringappend_494_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rl, existT _ _stringappend_494_ _) =>
+ let _stringappend_495_ :=
+ string_drop _stringappend_493_ (build_ex _stringappend_494_) in
+ (match (spc_matches_prefix _stringappend_495_) with
+ | Some (tt,(existT _ _stringappend_496_ _)) =>
+ returnm (tt, build_ex _stringappend_496_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_496_ _) =>
+ let _stringappend_497_ :=
+ string_drop _stringappend_495_ (build_ex _stringappend_496_) in
+ (match (reg_name_matches_prefix _stringappend_497_) with
+ | Some (rd,(existT _ _stringappend_498_ _)) =>
+ returnm (rd, build_ex _stringappend_498_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_498_ _) =>
+ let _stringappend_499_ :=
+ string_drop _stringappend_497_ (build_ex _stringappend_498_) in
+ sep_matches_prefix _stringappend_499_ >>= fun w__204 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__204 with
- | Some (_stringappend_735_,(existT _ _stringappend_736_ _)) =>
- returnm ((_stringappend_735_, build_ex _stringappend_736_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__206 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_736_ _) :=
- w__206
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_737_ :=
- string_drop _stringappend_734_ (build_ex _stringappend_736_) in
- match (reg_name_matches_prefix _stringappend_737_) with
- | Some (_stringappend_738_,(existT _ _stringappend_739_ _)) =>
- returnm ((_stringappend_738_, build_ex _stringappend_739_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__208 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_739_ _) :=
- w__208
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_740_ :=
- string_drop _stringappend_737_ (build_ex _stringappend_739_) in
- sep_matches_prefix _stringappend_740_ >>= fun w__209 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__204 with
+ | Some (tt,(existT _ _stringappend_500_ _)) =>
+ returnm (tt, build_ex _stringappend_500_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_500_ _) =>
+ let _stringappend_501_ :=
+ string_drop _stringappend_499_ (build_ex _stringappend_500_) in
+ (match (reg_name_matches_prefix _stringappend_501_) with
+ | Some (rs1,(existT _ _stringappend_502_ _)) =>
+ returnm (rs1, build_ex _stringappend_502_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_502_ _) =>
+ let _stringappend_503_ :=
+ string_drop _stringappend_501_ (build_ex _stringappend_502_) in
+ sep_matches_prefix _stringappend_503_ >>= fun w__209 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__209 with
- | Some (_stringappend_741_,(existT _ _stringappend_742_ _)) =>
- returnm ((_stringappend_741_, build_ex _stringappend_742_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__211 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_742_ _) :=
- w__211
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_743_ :=
- string_drop _stringappend_740_ (build_ex _stringappend_742_) in
- match (hex_bits_12_matches_prefix _stringappend_743_) with
- | Some (_stringappend_744_,(existT _ _stringappend_745_ _)) =>
- returnm ((_stringappend_744_, build_ex _stringappend_745_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__213 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_745_ _) :=
- w__213
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_743_ (build_ex _stringappend_745_)) with
+ (match w__209 with
+ | Some (tt,(existT _ _stringappend_504_ _)) =>
+ returnm (tt, build_ex _stringappend_504_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_504_ _) =>
+ let _stringappend_505_ :=
+ string_drop _stringappend_503_ (build_ex _stringappend_504_) in
+ (match (hex_bits_12_matches_prefix _stringappend_505_) with
+ | Some (imm,(existT _ _stringappend_506_ _)) =>
+ returnm (imm, build_ex _stringappend_506_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_506_ _) =>
+ (match (string_drop _stringappend_505_ (build_ex _stringappend_506_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
and_boolM
- (returnm ((string_startswith _stringappend_559_ "addiw")
+ (returnm ((string_startswith _stringappend_381_ "addiw")
: bool))
- (let _stringappend_747_ :=
- string_drop _stringappend_559_ (string_length "addiw") in
- match (spc_matches_prefix _stringappend_747_) with
- | Some (_stringappend_748_,(existT _ _stringappend_749_ _)) =>
- let _stringappend_750_ :=
- string_drop _stringappend_747_ (build_ex _stringappend_749_) in
- match (reg_name_matches_prefix _stringappend_750_) with
- | Some (_stringappend_751_,(existT _ _stringappend_752_ _)) =>
- let _stringappend_753_ :=
- string_drop _stringappend_750_ (build_ex _stringappend_752_) in
- sep_matches_prefix _stringappend_753_ >>= fun w__216 : option ((unit * {n : Z & ArithFact (n >=
+ (let _stringappend_508_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length "addiw"))) in
+ match (spc_matches_prefix _stringappend_508_) with
+ | Some (tt,(existT _ _stringappend_509_ _)) =>
+ let _stringappend_510_ :=
+ string_drop _stringappend_508_ (build_ex _stringappend_509_) in
+ match (reg_name_matches_prefix _stringappend_510_) with
+ | Some (rd,(existT _ _stringappend_511_ _)) =>
+ let _stringappend_512_ :=
+ string_drop _stringappend_510_ (build_ex _stringappend_511_) in
+ sep_matches_prefix _stringappend_512_ >>= fun w__216 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__216 with
- | Some (_stringappend_754_,(existT _ _stringappend_755_ _)) =>
- let _stringappend_756_ :=
- string_drop _stringappend_753_
- (build_ex
- _stringappend_755_) in
- match (reg_name_matches_prefix _stringappend_756_) with
- | Some (_stringappend_757_,(existT _ _stringappend_758_ _)) =>
- let _stringappend_759_ :=
- string_drop _stringappend_756_
- (build_ex
- _stringappend_758_) in
- sep_matches_prefix _stringappend_759_ >>= fun w__217 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_513_ _)) =>
+ let _stringappend_514_ :=
+ string_drop _stringappend_512_
+ (build_ex _stringappend_513_) in
+ match (reg_name_matches_prefix _stringappend_514_) with
+ | Some (rs1,(existT _ _stringappend_515_ _)) =>
+ let _stringappend_516_ :=
+ string_drop _stringappend_514_
+ (build_ex _stringappend_515_) in
+ sep_matches_prefix _stringappend_516_ >>= fun w__217 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__217 with
| Some
- (_stringappend_760_,(existT _ _stringappend_761_ _)) =>
- let _stringappend_762_ :=
- string_drop _stringappend_759_
- (build_ex
- _stringappend_761_) in
+ (tt,(existT _ _stringappend_517_ _)) =>
+ let _stringappend_518_ :=
+ string_drop _stringappend_516_
+ (build_ex _stringappend_517_) in
if ((match (hex_bits_12_matches_prefix
- _stringappend_762_) with
+ _stringappend_518_) with
| Some
- (_stringappend_763_,(existT _ _stringappend_764_ _)) =>
+ (imm,(existT _ _stringappend_519_ _)) =>
match (string_drop
- _stringappend_762_
- (build_ex
- _stringappend_764_)) with
+ _stringappend_518_
+ (build_ex _stringappend_519_)) with
| "" => true
| _ => false
end
@@ -19102,133 +18677,101 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__222 : bool =>
(if (w__222) then
- let _stringappend_747_ :=
- string_drop _stringappend_559_ (string_length "addiw") in
- match (spc_matches_prefix _stringappend_747_) with
- | Some (_stringappend_748_,(existT _ _stringappend_749_ _)) =>
- returnm ((_stringappend_748_, build_ex _stringappend_749_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__224 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_749_ _) :=
- w__224
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_750_ :=
- string_drop _stringappend_747_ (build_ex _stringappend_749_) in
- match (reg_name_matches_prefix _stringappend_750_) with
- | Some (_stringappend_751_,(existT _ _stringappend_752_ _)) =>
- returnm ((_stringappend_751_, build_ex _stringappend_752_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__226 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_752_ _) :=
- w__226
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_753_ :=
- string_drop _stringappend_750_ (build_ex _stringappend_752_) in
- sep_matches_prefix _stringappend_753_ >>= fun w__227 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_508_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length "addiw"))) in
+ (match (spc_matches_prefix _stringappend_508_) with
+ | Some (tt,(existT _ _stringappend_509_ _)) =>
+ returnm (tt, build_ex _stringappend_509_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_509_ _) =>
+ let _stringappend_510_ :=
+ string_drop _stringappend_508_ (build_ex _stringappend_509_) in
+ (match (reg_name_matches_prefix _stringappend_510_) with
+ | Some (rd,(existT _ _stringappend_511_ _)) =>
+ returnm (rd, build_ex _stringappend_511_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_511_ _) =>
+ let _stringappend_512_ :=
+ string_drop _stringappend_510_ (build_ex _stringappend_511_) in
+ sep_matches_prefix _stringappend_512_ >>= fun w__227 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__227 with
- | Some (_stringappend_754_,(existT _ _stringappend_755_ _)) =>
- returnm ((_stringappend_754_, build_ex _stringappend_755_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__229 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_755_ _) :=
- w__229
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_756_ :=
- string_drop _stringappend_753_ (build_ex _stringappend_755_) in
- match (reg_name_matches_prefix _stringappend_756_) with
- | Some (_stringappend_757_,(existT _ _stringappend_758_ _)) =>
- returnm ((_stringappend_757_, build_ex _stringappend_758_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__231 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_758_ _) :=
- w__231
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_759_ :=
- string_drop _stringappend_756_ (build_ex _stringappend_758_) in
- sep_matches_prefix _stringappend_759_ >>= fun w__232 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__227 with
+ | Some (tt,(existT _ _stringappend_513_ _)) =>
+ returnm (tt, build_ex _stringappend_513_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_513_ _) =>
+ let _stringappend_514_ :=
+ string_drop _stringappend_512_ (build_ex _stringappend_513_) in
+ (match (reg_name_matches_prefix _stringappend_514_) with
+ | Some (rs1,(existT _ _stringappend_515_ _)) =>
+ returnm (rs1, build_ex _stringappend_515_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_515_ _) =>
+ let _stringappend_516_ :=
+ string_drop _stringappend_514_ (build_ex _stringappend_515_) in
+ sep_matches_prefix _stringappend_516_ >>= fun w__232 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__232 with
- | Some (_stringappend_760_,(existT _ _stringappend_761_ _)) =>
- returnm ((_stringappend_760_, build_ex _stringappend_761_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__234 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_761_ _) :=
- w__234
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_762_ :=
- string_drop _stringappend_759_ (build_ex _stringappend_761_) in
- match (hex_bits_12_matches_prefix _stringappend_762_) with
- | Some (_stringappend_763_,(existT _ _stringappend_764_ _)) =>
- returnm ((_stringappend_763_, build_ex _stringappend_764_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__236 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_764_ _) :=
- w__236
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_762_
- (build_ex
- _stringappend_764_)) with
+ (match w__232 with
+ | Some (tt,(existT _ _stringappend_517_ _)) =>
+ returnm (tt, build_ex _stringappend_517_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_517_ _) =>
+ let _stringappend_518_ :=
+ string_drop _stringappend_516_ (build_ex _stringappend_517_) in
+ (match (hex_bits_12_matches_prefix _stringappend_518_) with
+ | Some (imm,(existT _ _stringappend_519_ _)) =>
+ returnm (imm, build_ex _stringappend_519_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_519_ _) =>
+ (match (string_drop _stringappend_518_
+ (build_ex _stringappend_519_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
- match (shiftw_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_766_,(existT _ _stringappend_767_ _)) =>
- let _stringappend_768_ :=
- string_drop _stringappend_559_ (build_ex _stringappend_767_) in
- match (spc_matches_prefix _stringappend_768_) with
- | Some (_stringappend_769_,(existT _ _stringappend_770_ _)) =>
- let _stringappend_771_ :=
- string_drop _stringappend_768_
- (build_ex
- _stringappend_770_) in
- match (reg_name_matches_prefix _stringappend_771_) with
- | Some (_stringappend_772_,(existT _ _stringappend_773_ _)) =>
- let _stringappend_774_ :=
- string_drop _stringappend_771_
- (build_ex
- _stringappend_773_) in
- sep_matches_prefix _stringappend_774_ >>= fun w__239 : option ((unit * {n : Z & ArithFact (n >=
+ match (shiftw_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_521_ _)) =>
+ let _stringappend_522_ :=
+ string_drop _stringappend_381_ (build_ex _stringappend_521_) in
+ match (spc_matches_prefix _stringappend_522_) with
+ | Some (tt,(existT _ _stringappend_523_ _)) =>
+ let _stringappend_524_ :=
+ string_drop _stringappend_522_
+ (build_ex _stringappend_523_) in
+ match (reg_name_matches_prefix _stringappend_524_) with
+ | Some (rd,(existT _ _stringappend_525_ _)) =>
+ let _stringappend_526_ :=
+ string_drop _stringappend_524_
+ (build_ex _stringappend_525_) in
+ sep_matches_prefix _stringappend_526_ >>= fun w__239 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__239 with
- | Some
- (_stringappend_775_,(existT _ _stringappend_776_ _)) =>
- let _stringappend_777_ :=
- string_drop _stringappend_774_
- (build_ex
- _stringappend_776_) in
- match (reg_name_matches_prefix _stringappend_777_) with
- | Some
- (_stringappend_778_,(existT _ _stringappend_779_ _)) =>
- let _stringappend_780_ :=
- string_drop _stringappend_777_
- (build_ex
- _stringappend_779_) in
- sep_matches_prefix _stringappend_780_ >>= fun w__240 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_527_ _)) =>
+ let _stringappend_528_ :=
+ string_drop _stringappend_526_
+ (build_ex _stringappend_527_) in
+ match (reg_name_matches_prefix _stringappend_528_) with
+ | Some (rs1,(existT _ _stringappend_529_ _)) =>
+ let _stringappend_530_ :=
+ string_drop _stringappend_528_
+ (build_ex _stringappend_529_) in
+ sep_matches_prefix _stringappend_530_ >>= fun w__240 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__240 with
| Some
- (_stringappend_781_,(existT _ _stringappend_782_ _)) =>
- let _stringappend_783_ :=
- string_drop _stringappend_780_
- (build_ex
- _stringappend_782_) in
+ (tt,(existT _ _stringappend_531_ _)) =>
+ let _stringappend_532_ :=
+ string_drop _stringappend_530_
+ (build_ex _stringappend_531_) in
if ((match (hex_bits_5_matches_prefix
- _stringappend_783_) with
+ _stringappend_532_) with
| Some
- (_stringappend_784_,(existT _ _stringappend_785_ _)) =>
+ (shamt,(existT _ _stringappend_533_ _)) =>
match (string_drop
- _stringappend_783_
- (build_ex
- _stringappend_785_)) with
+ _stringappend_532_
+ (build_ex _stringappend_533_)) with
| "" => true
| _ => false
end
@@ -19264,146 +18807,107 @@ Definition assembly_backwards_matches (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__245 : bool =>
(if (w__245) then
- match (shiftw_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_766_,(existT _ _stringappend_767_ _)) =>
- returnm ((_stringappend_766_, build_ex _stringappend_767_)
- : (sop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((sop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__247 : (sop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_767_ _) :=
- w__247
- : (sop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_768_ :=
- string_drop _stringappend_559_ (build_ex _stringappend_767_) in
- match (spc_matches_prefix _stringappend_768_) with
- | Some (_stringappend_769_,(existT _ _stringappend_770_ _)) =>
- returnm ((_stringappend_769_, build_ex _stringappend_770_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__249 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_770_ _) :=
- w__249
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_771_ :=
- string_drop _stringappend_768_ (build_ex _stringappend_770_) in
- match (reg_name_matches_prefix _stringappend_771_) with
- | Some (_stringappend_772_,(existT _ _stringappend_773_ _)) =>
- returnm ((_stringappend_772_, build_ex _stringappend_773_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__251 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_773_ _) :=
- w__251
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_774_ :=
- string_drop _stringappend_771_ (build_ex _stringappend_773_) in
- sep_matches_prefix _stringappend_774_ >>= fun w__252 : option ((unit * {n : Z & ArithFact (n >=
+ (match (shiftw_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_521_ _)) =>
+ returnm (op, build_ex _stringappend_521_)
+ | _ => exit tt : M ((sop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((sop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_521_ _) =>
+ let _stringappend_522_ :=
+ string_drop _stringappend_381_ (build_ex _stringappend_521_) in
+ (match (spc_matches_prefix _stringappend_522_) with
+ | Some (tt,(existT _ _stringappend_523_ _)) =>
+ returnm (tt, build_ex _stringappend_523_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_523_ _) =>
+ let _stringappend_524_ :=
+ string_drop _stringappend_522_ (build_ex _stringappend_523_) in
+ (match (reg_name_matches_prefix _stringappend_524_) with
+ | Some (rd,(existT _ _stringappend_525_ _)) =>
+ returnm (rd, build_ex _stringappend_525_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_525_ _) =>
+ let _stringappend_526_ :=
+ string_drop _stringappend_524_ (build_ex _stringappend_525_) in
+ sep_matches_prefix _stringappend_526_ >>= fun w__252 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__252 with
- | Some (_stringappend_775_,(existT _ _stringappend_776_ _)) =>
- returnm ((_stringappend_775_, build_ex _stringappend_776_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__254 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_776_ _) :=
- w__254
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_777_ :=
- string_drop _stringappend_774_ (build_ex _stringappend_776_) in
- match (reg_name_matches_prefix _stringappend_777_) with
- | Some (_stringappend_778_,(existT _ _stringappend_779_ _)) =>
- returnm ((_stringappend_778_, build_ex _stringappend_779_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__256 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_779_ _) :=
- w__256
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_780_ :=
- string_drop _stringappend_777_ (build_ex _stringappend_779_) in
- sep_matches_prefix _stringappend_780_ >>= fun w__257 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__252 with
+ | Some (tt,(existT _ _stringappend_527_ _)) =>
+ returnm (tt, build_ex _stringappend_527_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_527_ _) =>
+ let _stringappend_528_ :=
+ string_drop _stringappend_526_ (build_ex _stringappend_527_) in
+ (match (reg_name_matches_prefix _stringappend_528_) with
+ | Some (rs1,(existT _ _stringappend_529_ _)) =>
+ returnm (rs1, build_ex _stringappend_529_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_529_ _) =>
+ let _stringappend_530_ :=
+ string_drop _stringappend_528_ (build_ex _stringappend_529_) in
+ sep_matches_prefix _stringappend_530_ >>= fun w__257 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__257 with
- | Some (_stringappend_781_,(existT _ _stringappend_782_ _)) =>
- returnm ((_stringappend_781_, build_ex _stringappend_782_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__259 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_782_ _) :=
- w__259
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_783_ :=
- string_drop _stringappend_780_ (build_ex _stringappend_782_) in
- match (hex_bits_5_matches_prefix _stringappend_783_) with
- | Some (_stringappend_784_,(existT _ _stringappend_785_ _)) =>
- returnm ((_stringappend_784_, build_ex _stringappend_785_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__261 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(shamt, existT _ _stringappend_785_ _) :=
- w__261
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_783_
- (build_ex
- _stringappend_785_)) with
+ (match w__257 with
+ | Some (tt,(existT _ _stringappend_531_ _)) =>
+ returnm (tt, build_ex _stringappend_531_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_531_ _) =>
+ let _stringappend_532_ :=
+ string_drop _stringappend_530_ (build_ex _stringappend_531_) in
+ (match (hex_bits_5_matches_prefix _stringappend_532_) with
+ | Some (shamt,(existT _ _stringappend_533_ _)) =>
+ returnm (shamt, build_ex _stringappend_533_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(shamt, existT _ _stringappend_533_ _) =>
+ (match (string_drop _stringappend_532_
+ (build_ex _stringappend_533_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
- match (rtypew_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_787_,(existT _ _stringappend_788_ _)) =>
- let _stringappend_789_ :=
- string_drop _stringappend_559_
- (build_ex
- _stringappend_788_) in
- match (spc_matches_prefix _stringappend_789_) with
- | Some (_stringappend_790_,(existT _ _stringappend_791_ _)) =>
- let _stringappend_792_ :=
- string_drop _stringappend_789_
- (build_ex
- _stringappend_791_) in
- match (reg_name_matches_prefix _stringappend_792_) with
- | Some
- (_stringappend_793_,(existT _ _stringappend_794_ _)) =>
- let _stringappend_795_ :=
- string_drop _stringappend_792_
- (build_ex
- _stringappend_794_) in
- sep_matches_prefix _stringappend_795_ >>= fun w__264 : option ((unit * {n : Z & ArithFact (n >=
+ match (rtypew_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_535_ _)) =>
+ let _stringappend_536_ :=
+ string_drop _stringappend_381_
+ (build_ex _stringappend_535_) in
+ match (spc_matches_prefix _stringappend_536_) with
+ | Some (tt,(existT _ _stringappend_537_ _)) =>
+ let _stringappend_538_ :=
+ string_drop _stringappend_536_
+ (build_ex _stringappend_537_) in
+ match (reg_name_matches_prefix _stringappend_538_) with
+ | Some (rd,(existT _ _stringappend_539_ _)) =>
+ let _stringappend_540_ :=
+ string_drop _stringappend_538_
+ (build_ex _stringappend_539_) in
+ sep_matches_prefix _stringappend_540_ >>= fun w__264 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__264 with
- | Some
- (_stringappend_796_,(existT _ _stringappend_797_ _)) =>
- let _stringappend_798_ :=
- string_drop _stringappend_795_
- (build_ex
- _stringappend_797_) in
- match (reg_name_matches_prefix _stringappend_798_) with
- | Some
- (_stringappend_799_,(existT _ _stringappend_800_ _)) =>
- let _stringappend_801_ :=
- string_drop _stringappend_798_
- (build_ex
- _stringappend_800_) in
- sep_matches_prefix _stringappend_801_ >>= fun w__265 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_541_ _)) =>
+ let _stringappend_542_ :=
+ string_drop _stringappend_540_
+ (build_ex _stringappend_541_) in
+ match (reg_name_matches_prefix _stringappend_542_) with
+ | Some (rs1,(existT _ _stringappend_543_ _)) =>
+ let _stringappend_544_ :=
+ string_drop _stringappend_542_
+ (build_ex _stringappend_543_) in
+ sep_matches_prefix _stringappend_544_ >>= fun w__265 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__265 with
| Some
- (_stringappend_802_,(existT _ _stringappend_803_ _)) =>
- let _stringappend_804_ :=
+ (tt,(existT _ _stringappend_545_ _)) =>
+ let _stringappend_546_ :=
string_drop
- _stringappend_801_
- (build_ex
- _stringappend_803_) in
+ _stringappend_544_
+ (build_ex _stringappend_545_) in
if ((match (reg_name_matches_prefix
- _stringappend_804_) with
+ _stringappend_546_) with
| Some
- (_stringappend_805_,(existT _ _stringappend_806_ _)) =>
+ (rs2,(existT _ _stringappend_547_ _)) =>
match (string_drop
- _stringappend_804_
- (build_ex
- _stringappend_806_)) with
+ _stringappend_546_
+ (build_ex _stringappend_547_)) with
| "" => true
| _ => false
end
@@ -19439,154 +18943,111 @@ Definition assembly_backwards_matches (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__270 : bool =>
(if (w__270) then
- match (rtypew_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_787_,(existT _ _stringappend_788_ _)) =>
- returnm ((_stringappend_787_, build_ex _stringappend_788_)
- : (ropw * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((ropw * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__272 : (ropw * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_788_ _) :=
- w__272
- : (ropw * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_789_ :=
- string_drop _stringappend_559_
- (build_ex
- _stringappend_788_) in
- match (spc_matches_prefix _stringappend_789_) with
- | Some (_stringappend_790_,(existT _ _stringappend_791_ _)) =>
- returnm ((_stringappend_790_, build_ex _stringappend_791_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__274 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_791_ _) :=
- w__274
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_792_ :=
- string_drop _stringappend_789_
- (build_ex
- _stringappend_791_) in
- match (reg_name_matches_prefix _stringappend_792_) with
- | Some (_stringappend_793_,(existT _ _stringappend_794_ _)) =>
- returnm ((_stringappend_793_, build_ex _stringappend_794_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__276 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_794_ _) :=
- w__276
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_795_ :=
- string_drop _stringappend_792_
- (build_ex
- _stringappend_794_) in
- sep_matches_prefix _stringappend_795_ >>= fun w__277 : option ((unit * {n : Z & ArithFact (n >=
+ (match (rtypew_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_535_ _)) =>
+ returnm (op, build_ex _stringappend_535_)
+ | _ => exit tt : M ((ropw * {n : Z & ArithFact (n >= 0)}))
+ end : M ((ropw * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_535_ _) =>
+ let _stringappend_536_ :=
+ string_drop _stringappend_381_
+ (build_ex _stringappend_535_) in
+ (match (spc_matches_prefix _stringappend_536_) with
+ | Some (tt,(existT _ _stringappend_537_ _)) =>
+ returnm (tt, build_ex _stringappend_537_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_537_ _) =>
+ let _stringappend_538_ :=
+ string_drop _stringappend_536_
+ (build_ex _stringappend_537_) in
+ (match (reg_name_matches_prefix _stringappend_538_) with
+ | Some (rd,(existT _ _stringappend_539_ _)) =>
+ returnm (rd, build_ex _stringappend_539_)
+ | _ =>
+ exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_539_ _) =>
+ let _stringappend_540_ :=
+ string_drop _stringappend_538_
+ (build_ex _stringappend_539_) in
+ sep_matches_prefix _stringappend_540_ >>= fun w__277 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__277 with
- | Some (_stringappend_796_,(existT _ _stringappend_797_ _)) =>
- returnm ((_stringappend_796_, build_ex _stringappend_797_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__279 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_797_ _) :=
- w__279
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_798_ :=
- string_drop _stringappend_795_
- (build_ex
- _stringappend_797_) in
- match (reg_name_matches_prefix _stringappend_798_) with
- | Some (_stringappend_799_,(existT _ _stringappend_800_ _)) =>
- returnm ((_stringappend_799_, build_ex _stringappend_800_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__281 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_800_ _) :=
- w__281
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_801_ :=
- string_drop _stringappend_798_
- (build_ex
- _stringappend_800_) in
- sep_matches_prefix _stringappend_801_ >>= fun w__282 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__277 with
+ | Some (tt,(existT _ _stringappend_541_ _)) =>
+ returnm (tt, build_ex _stringappend_541_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_541_ _) =>
+ let _stringappend_542_ :=
+ string_drop _stringappend_540_
+ (build_ex _stringappend_541_) in
+ (match (reg_name_matches_prefix _stringappend_542_) with
+ | Some (rs1,(existT _ _stringappend_543_ _)) =>
+ returnm (rs1, build_ex _stringappend_543_)
+ | _ =>
+ exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_543_ _) =>
+ let _stringappend_544_ :=
+ string_drop _stringappend_542_
+ (build_ex _stringappend_543_) in
+ sep_matches_prefix _stringappend_544_ >>= fun w__282 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__282 with
- | Some (_stringappend_802_,(existT _ _stringappend_803_ _)) =>
- returnm ((_stringappend_802_, build_ex _stringappend_803_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__284 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_803_ _) :=
- w__284
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_804_ :=
- string_drop _stringappend_801_
- (build_ex
- _stringappend_803_) in
- match (reg_name_matches_prefix _stringappend_804_) with
- | Some (_stringappend_805_,(existT _ _stringappend_806_ _)) =>
- returnm ((_stringappend_805_, build_ex _stringappend_806_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__286 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs2, existT _ _stringappend_806_ _) :=
- w__286
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_804_
- (build_ex
- _stringappend_806_)) with
+ (match w__282 with
+ | Some (tt,(existT _ _stringappend_545_ _)) =>
+ returnm (tt, build_ex _stringappend_545_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_545_ _) =>
+ let _stringappend_546_ :=
+ string_drop _stringappend_544_
+ (build_ex _stringappend_545_) in
+ (match (reg_name_matches_prefix _stringappend_546_) with
+ | Some (rs2,(existT _ _stringappend_547_ _)) =>
+ returnm (rs2, build_ex _stringappend_547_)
+ | _ =>
+ exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_547_ _) =>
+ (match (string_drop _stringappend_546_
+ (build_ex _stringappend_547_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
- match (shiftiwop_mnemonic_matches_prefix _stringappend_559_) with
- | Some (_stringappend_808_,(existT _ _stringappend_809_ _)) =>
- let _stringappend_810_ :=
- string_drop _stringappend_559_
- (build_ex
- _stringappend_809_) in
- match (spc_matches_prefix _stringappend_810_) with
- | Some
- (_stringappend_811_,(existT _ _stringappend_812_ _)) =>
- let _stringappend_813_ :=
- string_drop _stringappend_810_
- (build_ex
- _stringappend_812_) in
- match (reg_name_matches_prefix _stringappend_813_) with
- | Some
- (_stringappend_814_,(existT _ _stringappend_815_ _)) =>
- let _stringappend_816_ :=
- string_drop _stringappend_813_
- (build_ex
- _stringappend_815_) in
- sep_matches_prefix _stringappend_816_ >>= fun w__289 : option ((unit * {n : Z & ArithFact (n >=
+ match (shiftiwop_mnemonic_matches_prefix _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_549_ _)) =>
+ let _stringappend_550_ :=
+ string_drop _stringappend_381_
+ (build_ex _stringappend_549_) in
+ match (spc_matches_prefix _stringappend_550_) with
+ | Some (tt,(existT _ _stringappend_551_ _)) =>
+ let _stringappend_552_ :=
+ string_drop _stringappend_550_
+ (build_ex _stringappend_551_) in
+ match (reg_name_matches_prefix _stringappend_552_) with
+ | Some (rd,(existT _ _stringappend_553_ _)) =>
+ let _stringappend_554_ :=
+ string_drop _stringappend_552_
+ (build_ex _stringappend_553_) in
+ sep_matches_prefix _stringappend_554_ >>= fun w__289 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__289 with
| Some
- (_stringappend_817_,(existT _ _stringappend_818_ _)) =>
- let _stringappend_819_ :=
- string_drop _stringappend_816_
- (build_ex
- _stringappend_818_) in
+ (tt,(existT _ _stringappend_555_ _)) =>
+ let _stringappend_556_ :=
+ string_drop _stringappend_554_
+ (build_ex _stringappend_555_) in
if ((match (reg_name_matches_prefix
- _stringappend_819_) with
+ _stringappend_556_) with
| Some
- (_stringappend_820_,(existT _ _stringappend_821_ _)) =>
- let _stringappend_822_ :=
+ (rs1,(existT _ _stringappend_557_ _)) =>
+ let _stringappend_558_ :=
string_drop
- _stringappend_819_
- (build_ex
- _stringappend_821_) in
+ _stringappend_556_
+ (build_ex _stringappend_557_) in
if ((match (hex_bits_5_matches_prefix
- _stringappend_822_) with
+ _stringappend_558_) with
| Some
- (_stringappend_823_,(existT _ _stringappend_824_ _)) =>
+ (shamt,(existT _ _stringappend_559_ _)) =>
match (string_drop
- _stringappend_822_
- (build_ex
- _stringappend_824_)) with
+ _stringappend_558_
+ (build_ex _stringappend_559_)) with
| "" => true
| _ => false
end
@@ -19616,171 +19077,115 @@ Definition assembly_backwards_matches (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__292 : bool =>
(if (w__292) then
- match (shiftiwop_mnemonic_matches_prefix
- _stringappend_559_) with
- | Some
- (_stringappend_808_,(existT _ _stringappend_809_ _)) =>
- returnm ((_stringappend_808_,
- build_ex
- _stringappend_809_)
- : (sopw * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((sopw * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__294 : (sopw * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_809_ _) :=
- w__294
- : (sopw * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_810_ :=
- string_drop _stringappend_559_
- (build_ex
- _stringappend_809_) in
- match (spc_matches_prefix _stringappend_810_) with
- | Some
- (_stringappend_811_,(existT _ _stringappend_812_ _)) =>
- returnm ((_stringappend_811_,
- build_ex
- _stringappend_812_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__296 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_812_ _) :=
- w__296
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_813_ :=
- string_drop _stringappend_810_
- (build_ex
- _stringappend_812_) in
- match (reg_name_matches_prefix _stringappend_813_) with
- | Some
- (_stringappend_814_,(existT _ _stringappend_815_ _)) =>
- returnm ((_stringappend_814_,
- build_ex
- _stringappend_815_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__298 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_815_ _) :=
- w__298
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_816_ :=
- string_drop _stringappend_813_
- (build_ex
- _stringappend_815_) in
- sep_matches_prefix _stringappend_816_ >>= fun w__299 : option ((unit * {n : Z & ArithFact (n >=
+ (match (shiftiwop_mnemonic_matches_prefix
+ _stringappend_381_) with
+ | Some (op,(existT _ _stringappend_549_ _)) =>
+ returnm (op, build_ex _stringappend_549_)
+ | _ =>
+ exit tt : M ((sopw * {n : Z & ArithFact (n >= 0)}))
+ end : M ((sopw * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_549_ _) =>
+ let _stringappend_550_ :=
+ string_drop _stringappend_381_
+ (build_ex _stringappend_549_) in
+ (match (spc_matches_prefix _stringappend_550_) with
+ | Some (tt,(existT _ _stringappend_551_ _)) =>
+ returnm (tt, build_ex _stringappend_551_)
+ | _ =>
+ exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_551_ _) =>
+ let _stringappend_552_ :=
+ string_drop _stringappend_550_
+ (build_ex _stringappend_551_) in
+ (match (reg_name_matches_prefix _stringappend_552_) with
+ | Some (rd,(existT _ _stringappend_553_ _)) =>
+ returnm (rd, build_ex _stringappend_553_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_553_ _) =>
+ let _stringappend_554_ :=
+ string_drop _stringappend_552_
+ (build_ex _stringappend_553_) in
+ sep_matches_prefix _stringappend_554_ >>= fun w__299 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__299 with
- | Some
- (_stringappend_817_,(existT _ _stringappend_818_ _)) =>
- returnm ((_stringappend_817_,
- build_ex
- _stringappend_818_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__301 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_818_ _) :=
- w__301
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_819_ :=
- string_drop _stringappend_816_
- (build_ex
- _stringappend_818_) in
- match (reg_name_matches_prefix _stringappend_819_) with
- | Some
- (_stringappend_820_,(existT _ _stringappend_821_ _)) =>
- returnm ((_stringappend_820_,
- build_ex
- _stringappend_821_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__303 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_821_ _) :=
- w__303
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_822_ :=
- string_drop _stringappend_819_
- (build_ex
- _stringappend_821_) in
- match (hex_bits_5_matches_prefix _stringappend_822_) with
- | Some
- (_stringappend_823_,(existT _ _stringappend_824_ _)) =>
- returnm ((_stringappend_823_,
- build_ex
- _stringappend_824_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__305 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(shamt, existT _ _stringappend_824_ _) :=
- w__305
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_822_
- (build_ex
- _stringappend_824_)) with
+ (match w__299 with
+ | Some (tt,(existT _ _stringappend_555_ _)) =>
+ returnm (tt, build_ex _stringappend_555_)
+ | _ =>
+ exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_555_ _) =>
+ let _stringappend_556_ :=
+ string_drop _stringappend_554_
+ (build_ex _stringappend_555_) in
+ (match (reg_name_matches_prefix _stringappend_556_) with
+ | Some (rs1,(existT _ _stringappend_557_ _)) =>
+ returnm (rs1, build_ex _stringappend_557_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_557_ _) =>
+ let _stringappend_558_ :=
+ string_drop _stringappend_556_
+ (build_ex _stringappend_557_) in
+ (match (hex_bits_5_matches_prefix _stringappend_558_) with
+ | Some (shamt,(existT _ _stringappend_559_ _)) =>
+ returnm (shamt, build_ex _stringappend_559_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(shamt, existT _ _stringappend_559_ _) =>
+ (match (string_drop _stringappend_558_
+ (build_ex _stringappend_559_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
- match (mul_mnemonic_matches_prefix _stringappend_559_) with
+ match (mul_mnemonic_matches_prefix _stringappend_381_) with
| Some
- (_stringappend_826_,(existT _ _stringappend_827_ _)) =>
- let _stringappend_828_ :=
- string_drop _stringappend_559_
- (build_ex
- _stringappend_827_) in
- match (spc_matches_prefix _stringappend_828_) with
- | Some
- (_stringappend_829_,(existT _ _stringappend_830_ _)) =>
- let _stringappend_831_ :=
- string_drop _stringappend_828_
- (build_ex
- _stringappend_830_) in
- match (reg_name_matches_prefix _stringappend_831_) with
- | Some
- (_stringappend_832_,(existT _ _stringappend_833_ _)) =>
- let _stringappend_834_ :=
- string_drop _stringappend_831_
- (build_ex
- _stringappend_833_) in
- sep_matches_prefix _stringappend_834_ >>= fun w__308 : option ((unit * {n : Z & ArithFact (n >=
+ ((high, signed1, signed2),(existT _ _stringappend_561_ _)) =>
+ let _stringappend_562_ :=
+ string_drop _stringappend_381_
+ (build_ex _stringappend_561_) in
+ match (spc_matches_prefix _stringappend_562_) with
+ | Some (tt,(existT _ _stringappend_563_ _)) =>
+ let _stringappend_564_ :=
+ string_drop _stringappend_562_
+ (build_ex _stringappend_563_) in
+ match (reg_name_matches_prefix _stringappend_564_) with
+ | Some (rd,(existT _ _stringappend_565_ _)) =>
+ let _stringappend_566_ :=
+ string_drop _stringappend_564_
+ (build_ex _stringappend_565_) in
+ sep_matches_prefix _stringappend_566_ >>= fun w__308 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__308 with
- | Some
- (_stringappend_835_,(existT _ _stringappend_836_ _)) =>
- let _stringappend_837_ :=
- string_drop _stringappend_834_
- (build_ex
- _stringappend_836_) in
+ | Some (tt,(existT _ _stringappend_567_ _)) =>
+ let _stringappend_568_ :=
+ string_drop _stringappend_566_
+ (build_ex _stringappend_567_) in
match (reg_name_matches_prefix
- _stringappend_837_) with
- | Some
- (_stringappend_838_,(existT _ _stringappend_839_ _)) =>
- let _stringappend_840_ :=
- string_drop _stringappend_837_
- (build_ex
- _stringappend_839_) in
- sep_matches_prefix _stringappend_840_ >>= fun w__309 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_568_) with
+ | Some (rs1,(existT _ _stringappend_569_ _)) =>
+ let _stringappend_570_ :=
+ string_drop _stringappend_568_
+ (build_ex _stringappend_569_) in
+ sep_matches_prefix _stringappend_570_ >>= fun w__309 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__309 with
| Some
- (_stringappend_841_,(existT _ _stringappend_842_ _)) =>
- let _stringappend_843_ :=
+ (tt,(existT _ _stringappend_571_ _)) =>
+ let _stringappend_572_ :=
string_drop
- _stringappend_840_
- (build_ex
- _stringappend_842_) in
+ _stringappend_570_
+ (build_ex _stringappend_571_) in
if ((match (reg_name_matches_prefix
- _stringappend_843_) with
+ _stringappend_572_) with
| Some
- (_stringappend_844_,(existT _ _stringappend_845_ _)) =>
+ (rs2,(existT _ _stringappend_573_ _)) =>
match (string_drop
- _stringappend_843_
- (build_ex
- _stringappend_845_)) with
+ _stringappend_572_
+ (build_ex _stringappend_573_)) with
| "" => true
| _ => false
end
@@ -19816,206 +19221,139 @@ Definition assembly_backwards_matches (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__314 : bool =>
(if (w__314) then
- match (mul_mnemonic_matches_prefix _stringappend_559_) with
- | Some
- (_stringappend_826_,(existT _ _stringappend_827_ _)) =>
- returnm ((_stringappend_826_,
- build_ex
- _stringappend_827_)
- : ((bool * bool * bool) * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M (((bool * bool * bool) * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__316 : ((bool * bool * bool) * {n : Z & ArithFact (n >=
- 0)}) =>
- let '((high, signed1, signed2), existT _ _stringappend_827_ _) :=
- w__316
- : ((bool * bool * bool) * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_828_ :=
- string_drop _stringappend_559_
- (build_ex
- _stringappend_827_) in
- match (spc_matches_prefix _stringappend_828_) with
- | Some
- (_stringappend_829_,(existT _ _stringappend_830_ _)) =>
- returnm ((_stringappend_829_,
- build_ex
- _stringappend_830_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__318 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_830_ _) :=
- w__318
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_831_ :=
- string_drop _stringappend_828_
- (build_ex
- _stringappend_830_) in
- match (reg_name_matches_prefix _stringappend_831_) with
- | Some
- (_stringappend_832_,(existT _ _stringappend_833_ _)) =>
- returnm ((_stringappend_832_,
- build_ex
- _stringappend_833_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__320 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_833_ _) :=
- w__320
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_834_ :=
- string_drop _stringappend_831_
- (build_ex
- _stringappend_833_) in
- sep_matches_prefix _stringappend_834_ >>= fun w__321 : option ((unit * {n : Z & ArithFact (n >=
+ (match (mul_mnemonic_matches_prefix _stringappend_381_) with
+ | Some
+ ((high, signed1, signed2),(existT _ _stringappend_561_ _)) =>
+ returnm ((high, signed1, signed2), build_ex _stringappend_561_)
+ | _ =>
+ exit tt
+ : M (((bool * bool * bool) * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M (((bool * bool * bool) * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '((high, signed1, signed2), existT _ _stringappend_561_ _) =>
+ let _stringappend_562_ :=
+ string_drop _stringappend_381_
+ (build_ex _stringappend_561_) in
+ (match (spc_matches_prefix _stringappend_562_) with
+ | Some (tt,(existT _ _stringappend_563_ _)) =>
+ returnm (tt, build_ex _stringappend_563_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_563_ _) =>
+ let _stringappend_564_ :=
+ string_drop _stringappend_562_
+ (build_ex _stringappend_563_) in
+ (match (reg_name_matches_prefix _stringappend_564_) with
+ | Some (rd,(existT _ _stringappend_565_ _)) =>
+ returnm (rd, build_ex _stringappend_565_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_565_ _) =>
+ let _stringappend_566_ :=
+ string_drop _stringappend_564_
+ (build_ex _stringappend_565_) in
+ sep_matches_prefix _stringappend_566_ >>= fun w__321 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__321 with
- | Some
- (_stringappend_835_,(existT _ _stringappend_836_ _)) =>
- returnm ((_stringappend_835_,
- build_ex
- _stringappend_836_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__323 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_836_ _) :=
- w__323
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_837_ :=
- string_drop _stringappend_834_
- (build_ex
- _stringappend_836_) in
- match (reg_name_matches_prefix _stringappend_837_) with
- | Some
- (_stringappend_838_,(existT _ _stringappend_839_ _)) =>
- returnm ((_stringappend_838_,
- build_ex
- _stringappend_839_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__325 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_839_ _) :=
- w__325
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_840_ :=
- string_drop _stringappend_837_
- (build_ex
- _stringappend_839_) in
- sep_matches_prefix _stringappend_840_ >>= fun w__326 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__321 with
+ | Some (tt,(existT _ _stringappend_567_ _)) =>
+ returnm (tt, build_ex _stringappend_567_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_567_ _) =>
+ let _stringappend_568_ :=
+ string_drop _stringappend_566_
+ (build_ex _stringappend_567_) in
+ (match (reg_name_matches_prefix _stringappend_568_) with
+ | Some (rs1,(existT _ _stringappend_569_ _)) =>
+ returnm (rs1, build_ex _stringappend_569_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_569_ _) =>
+ let _stringappend_570_ :=
+ string_drop _stringappend_568_
+ (build_ex _stringappend_569_) in
+ sep_matches_prefix _stringappend_570_ >>= fun w__326 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__326 with
- | Some
- (_stringappend_841_,(existT _ _stringappend_842_ _)) =>
- returnm ((_stringappend_841_,
- build_ex
- _stringappend_842_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__328 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_842_ _) :=
- w__328
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_843_ :=
- string_drop _stringappend_840_
- (build_ex
- _stringappend_842_) in
- match (reg_name_matches_prefix _stringappend_843_) with
- | Some
- (_stringappend_844_,(existT _ _stringappend_845_ _)) =>
- returnm ((_stringappend_844_,
- build_ex
- _stringappend_845_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__330 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_845_ _) :=
- w__330
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_843_
- (build_ex
- _stringappend_845_)) with
+ (match w__326 with
+ | Some (tt,(existT _ _stringappend_571_ _)) =>
+ returnm (tt, build_ex _stringappend_571_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_571_ _) =>
+ let _stringappend_572_ :=
+ string_drop _stringappend_570_
+ (build_ex _stringappend_571_) in
+ (match (reg_name_matches_prefix _stringappend_572_) with
+ | Some (rs2,(existT _ _stringappend_573_ _)) =>
+ returnm (rs2, build_ex _stringappend_573_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_573_ _) =>
+ (match (string_drop _stringappend_572_
+ (build_ex _stringappend_573_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
and_boolM
- (returnm ((string_startswith _stringappend_559_ "div")
+ (returnm ((string_startswith _stringappend_381_ "div")
: bool))
- (let _stringappend_847_ :=
- string_drop _stringappend_559_
- (string_length "div") in
- match (maybe_not_u_matches_prefix _stringappend_847_) with
- | Some
- (_stringappend_848_,(existT _ _stringappend_849_ _)) =>
- let _stringappend_850_ :=
- string_drop _stringappend_847_
- (build_ex
- _stringappend_849_) in
- match (spc_matches_prefix _stringappend_850_) with
- | Some
- (_stringappend_851_,(existT _ _stringappend_852_ _)) =>
- let _stringappend_853_ :=
- string_drop _stringappend_850_
- (build_ex
- _stringappend_852_) in
+ (let _stringappend_575_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length "div"))) in
+ match (maybe_not_u_matches_prefix _stringappend_575_) with
+ | Some (s,(existT _ _stringappend_576_ _)) =>
+ let _stringappend_577_ :=
+ string_drop _stringappend_575_
+ (build_ex _stringappend_576_) in
+ match (spc_matches_prefix _stringappend_577_) with
+ | Some (tt,(existT _ _stringappend_578_ _)) =>
+ let _stringappend_579_ :=
+ string_drop _stringappend_577_
+ (build_ex _stringappend_578_) in
match (reg_name_matches_prefix
- _stringappend_853_) with
- | Some
- (_stringappend_854_,(existT _ _stringappend_855_ _)) =>
- let _stringappend_856_ :=
- string_drop _stringappend_853_
- (build_ex
- _stringappend_855_) in
- sep_matches_prefix _stringappend_856_ >>= fun w__333 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_579_) with
+ | Some (rd,(existT _ _stringappend_580_ _)) =>
+ let _stringappend_581_ :=
+ string_drop _stringappend_579_
+ (build_ex _stringappend_580_) in
+ sep_matches_prefix _stringappend_581_ >>= fun w__333 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__333 with
- | Some
- (_stringappend_857_,(existT _ _stringappend_858_ _)) =>
- let _stringappend_859_ :=
- string_drop _stringappend_856_
- (build_ex
- _stringappend_858_) in
+ | Some (tt,(existT _ _stringappend_582_ _)) =>
+ let _stringappend_583_ :=
+ string_drop _stringappend_581_
+ (build_ex _stringappend_582_) in
match (reg_name_matches_prefix
- _stringappend_859_) with
+ _stringappend_583_) with
| Some
- (_stringappend_860_,(existT _ _stringappend_861_ _)) =>
- let _stringappend_862_ :=
- string_drop _stringappend_859_
- (build_ex
- _stringappend_861_) in
- sep_matches_prefix _stringappend_862_ >>= fun w__334 : option ((unit * {n : Z & ArithFact (n >=
+ (rs1,(existT _ _stringappend_584_ _)) =>
+ let _stringappend_585_ :=
+ string_drop _stringappend_583_
+ (build_ex _stringappend_584_) in
+ sep_matches_prefix _stringappend_585_ >>= fun w__334 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__334 with
| Some
- (_stringappend_863_,(existT _ _stringappend_864_ _)) =>
- let _stringappend_865_ :=
+ (tt,(existT _ _stringappend_586_ _)) =>
+ let _stringappend_587_ :=
string_drop
- _stringappend_862_
- (build_ex
- _stringappend_864_) in
+ _stringappend_585_
+ (build_ex _stringappend_586_) in
if ((match (reg_name_matches_prefix
- _stringappend_865_) with
+ _stringappend_587_) with
| Some
- (_stringappend_866_,(existT _ _stringappend_867_ _)) =>
+ (rs2,(existT _ _stringappend_588_ _)) =>
match (string_drop
- _stringappend_865_
- (build_ex
- _stringappend_867_)) with
+ _stringappend_587_
+ (build_ex _stringappend_588_)) with
| "" => true
| _ => false
end
@@ -20054,216 +19392,144 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__340 : bool =>
(if (w__340) then
- let _stringappend_847_ :=
- string_drop _stringappend_559_
- (string_length "div") in
- match (maybe_not_u_matches_prefix _stringappend_847_) with
- | Some
- (_stringappend_848_,(existT _ _stringappend_849_ _)) =>
- returnm ((_stringappend_848_,
- build_ex
- _stringappend_849_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__342 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_849_ _) :=
- w__342
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_850_ :=
- string_drop _stringappend_847_
- (build_ex
- _stringappend_849_) in
- match (spc_matches_prefix _stringappend_850_) with
- | Some
- (_stringappend_851_,(existT _ _stringappend_852_ _)) =>
- returnm ((_stringappend_851_,
- build_ex
- _stringappend_852_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__344 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_852_ _) :=
- w__344
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_853_ :=
- string_drop _stringappend_850_
- (build_ex
- _stringappend_852_) in
- match (reg_name_matches_prefix _stringappend_853_) with
- | Some
- (_stringappend_854_,(existT _ _stringappend_855_ _)) =>
- returnm ((_stringappend_854_,
- build_ex
- _stringappend_855_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__346 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_855_ _) :=
- w__346
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_856_ :=
- string_drop _stringappend_853_
- (build_ex
- _stringappend_855_) in
- sep_matches_prefix _stringappend_856_ >>= fun w__347 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_575_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length "div"))) in
+ (match (maybe_not_u_matches_prefix
+ _stringappend_575_) with
+ | Some (s,(existT _ _stringappend_576_ _)) =>
+ returnm (s, build_ex _stringappend_576_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(s, existT _ _stringappend_576_ _) =>
+ let _stringappend_577_ :=
+ string_drop _stringappend_575_
+ (build_ex _stringappend_576_) in
+ (match (spc_matches_prefix _stringappend_577_) with
+ | Some (tt,(existT _ _stringappend_578_ _)) =>
+ returnm (tt, build_ex _stringappend_578_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_578_ _) =>
+ let _stringappend_579_ :=
+ string_drop _stringappend_577_
+ (build_ex _stringappend_578_) in
+ (match (reg_name_matches_prefix _stringappend_579_) with
+ | Some (rd,(existT _ _stringappend_580_ _)) =>
+ returnm (rd, build_ex _stringappend_580_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_580_ _) =>
+ let _stringappend_581_ :=
+ string_drop _stringappend_579_
+ (build_ex _stringappend_580_) in
+ sep_matches_prefix _stringappend_581_ >>= fun w__347 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__347 with
- | Some
- (_stringappend_857_,(existT _ _stringappend_858_ _)) =>
- returnm ((_stringappend_857_,
- build_ex
- _stringappend_858_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__349 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_858_ _) :=
- w__349
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_859_ :=
- string_drop _stringappend_856_
- (build_ex
- _stringappend_858_) in
- match (reg_name_matches_prefix _stringappend_859_) with
- | Some
- (_stringappend_860_,(existT _ _stringappend_861_ _)) =>
- returnm ((_stringappend_860_,
- build_ex
- _stringappend_861_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__351 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_861_ _) :=
- w__351
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_862_ :=
- string_drop _stringappend_859_
- (build_ex
- _stringappend_861_) in
- sep_matches_prefix _stringappend_862_ >>= fun w__352 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__347 with
+ | Some (tt,(existT _ _stringappend_582_ _)) =>
+ returnm (tt, build_ex _stringappend_582_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_582_ _) =>
+ let _stringappend_583_ :=
+ string_drop _stringappend_581_
+ (build_ex _stringappend_582_) in
+ (match (reg_name_matches_prefix _stringappend_583_) with
+ | Some (rs1,(existT _ _stringappend_584_ _)) =>
+ returnm (rs1, build_ex _stringappend_584_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_584_ _) =>
+ let _stringappend_585_ :=
+ string_drop _stringappend_583_
+ (build_ex _stringappend_584_) in
+ sep_matches_prefix _stringappend_585_ >>= fun w__352 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__352 with
- | Some
- (_stringappend_863_,(existT _ _stringappend_864_ _)) =>
- returnm ((_stringappend_863_,
- build_ex
- _stringappend_864_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__354 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_864_ _) :=
- w__354
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_865_ :=
- string_drop _stringappend_862_
- (build_ex
- _stringappend_864_) in
- match (reg_name_matches_prefix _stringappend_865_) with
- | Some
- (_stringappend_866_,(existT _ _stringappend_867_ _)) =>
- returnm ((_stringappend_866_,
- build_ex
- _stringappend_867_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__356 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_867_ _) :=
- w__356
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_865_
- (build_ex
- _stringappend_867_)) with
+ (match w__352 with
+ | Some (tt,(existT _ _stringappend_586_ _)) =>
+ returnm (tt, build_ex _stringappend_586_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_586_ _) =>
+ let _stringappend_587_ :=
+ string_drop _stringappend_585_
+ (build_ex _stringappend_586_) in
+ (match (reg_name_matches_prefix _stringappend_587_) with
+ | Some (rs2,(existT _ _stringappend_588_ _)) =>
+ returnm (rs2, build_ex _stringappend_588_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_588_ _) =>
+ (match (string_drop _stringappend_587_
+ (build_ex _stringappend_588_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
and_boolM
- (returnm ((string_startswith _stringappend_559_
+ (returnm ((string_startswith _stringappend_381_
"rem")
: bool))
- (let _stringappend_869_ :=
- string_drop _stringappend_559_
- (string_length "rem") in
+ (let _stringappend_590_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length "rem"))) in
match (maybe_not_u_matches_prefix
- _stringappend_869_) with
- | Some
- (_stringappend_870_,(existT _ _stringappend_871_ _)) =>
- let _stringappend_872_ :=
- string_drop _stringappend_869_
- (build_ex
- _stringappend_871_) in
- match (spc_matches_prefix _stringappend_872_) with
- | Some
- (_stringappend_873_,(existT _ _stringappend_874_ _)) =>
- let _stringappend_875_ :=
- string_drop _stringappend_872_
- (build_ex
- _stringappend_874_) in
+ _stringappend_590_) with
+ | Some (s,(existT _ _stringappend_591_ _)) =>
+ let _stringappend_592_ :=
+ string_drop _stringappend_590_
+ (build_ex _stringappend_591_) in
+ match (spc_matches_prefix _stringappend_592_) with
+ | Some (tt,(existT _ _stringappend_593_ _)) =>
+ let _stringappend_594_ :=
+ string_drop _stringappend_592_
+ (build_ex _stringappend_593_) in
match (reg_name_matches_prefix
- _stringappend_875_) with
- | Some
- (_stringappend_876_,(existT _ _stringappend_877_ _)) =>
- let _stringappend_878_ :=
- string_drop _stringappend_875_
- (build_ex
- _stringappend_877_) in
- sep_matches_prefix _stringappend_878_ >>= fun w__359 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_594_) with
+ | Some (rd,(existT _ _stringappend_595_ _)) =>
+ let _stringappend_596_ :=
+ string_drop _stringappend_594_
+ (build_ex _stringappend_595_) in
+ sep_matches_prefix _stringappend_596_ >>= fun w__359 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__359 with
| Some
- (_stringappend_879_,(existT _ _stringappend_880_ _)) =>
- let _stringappend_881_ :=
- string_drop _stringappend_878_
- (build_ex
- _stringappend_880_) in
+ (tt,(existT _ _stringappend_597_ _)) =>
+ let _stringappend_598_ :=
+ string_drop _stringappend_596_
+ (build_ex _stringappend_597_) in
match (reg_name_matches_prefix
- _stringappend_881_) with
+ _stringappend_598_) with
| Some
- (_stringappend_882_,(existT _ _stringappend_883_ _)) =>
- let _stringappend_884_ :=
- string_drop _stringappend_881_
- (build_ex
- _stringappend_883_) in
+ (rs1,(existT _ _stringappend_599_ _)) =>
+ let _stringappend_600_ :=
+ string_drop _stringappend_598_
+ (build_ex _stringappend_599_) in
sep_matches_prefix
- _stringappend_884_ >>= fun w__360 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_600_ >>= fun w__360 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__360 with
| Some
- (_stringappend_885_,(existT _ _stringappend_886_ _)) =>
- let _stringappend_887_ :=
+ (tt,(existT _ _stringappend_601_ _)) =>
+ let _stringappend_602_ :=
string_drop
- _stringappend_884_
- (build_ex
- _stringappend_886_) in
+ _stringappend_600_
+ (build_ex _stringappend_601_) in
if ((match (reg_name_matches_prefix
- _stringappend_887_) with
+ _stringappend_602_) with
| Some
- (_stringappend_888_,(existT _ _stringappend_889_ _)) =>
+ (rs2,(existT _ _stringappend_603_ _)) =>
match (string_drop
- _stringappend_887_
- (build_ex
- _stringappend_889_)) with
+ _stringappend_602_
+ (build_ex _stringappend_603_)) with
| "" =>
true
| _ =>
@@ -20305,209 +19571,141 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__366 : bool =>
(if (w__366) then
- let _stringappend_869_ :=
- string_drop _stringappend_559_
- (string_length "rem") in
- match (maybe_not_u_matches_prefix
- _stringappend_869_) with
- | Some
- (_stringappend_870_,(existT _ _stringappend_871_ _)) =>
- returnm ((_stringappend_870_,
- build_ex
- _stringappend_871_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__368 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_871_ _) :=
- w__368
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_872_ :=
- string_drop _stringappend_869_
- (build_ex
- _stringappend_871_) in
- match (spc_matches_prefix _stringappend_872_) with
- | Some
- (_stringappend_873_,(existT _ _stringappend_874_ _)) =>
- returnm ((_stringappend_873_,
- build_ex
- _stringappend_874_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__370 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_874_ _) :=
- w__370
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_875_ :=
- string_drop _stringappend_872_
- (build_ex
- _stringappend_874_) in
- match (reg_name_matches_prefix _stringappend_875_) with
- | Some
- (_stringappend_876_,(existT _ _stringappend_877_ _)) =>
- returnm ((_stringappend_876_,
- build_ex
- _stringappend_877_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__372 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_877_ _) :=
- w__372
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_878_ :=
- string_drop _stringappend_875_
- (build_ex
- _stringappend_877_) in
- sep_matches_prefix _stringappend_878_ >>= fun w__373 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_590_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length "rem"))) in
+ (match (maybe_not_u_matches_prefix
+ _stringappend_590_) with
+ | Some (s,(existT _ _stringappend_591_ _)) =>
+ returnm (s, build_ex _stringappend_591_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(s, existT _ _stringappend_591_ _) =>
+ let _stringappend_592_ :=
+ string_drop _stringappend_590_
+ (build_ex _stringappend_591_) in
+ (match (spc_matches_prefix _stringappend_592_) with
+ | Some (tt,(existT _ _stringappend_593_ _)) =>
+ returnm (tt, build_ex _stringappend_593_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_593_ _) =>
+ let _stringappend_594_ :=
+ string_drop _stringappend_592_
+ (build_ex _stringappend_593_) in
+ (match (reg_name_matches_prefix
+ _stringappend_594_) with
+ | Some (rd,(existT _ _stringappend_595_ _)) =>
+ returnm (rd, build_ex _stringappend_595_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_595_ _) =>
+ let _stringappend_596_ :=
+ string_drop _stringappend_594_
+ (build_ex _stringappend_595_) in
+ sep_matches_prefix _stringappend_596_ >>= fun w__373 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__373 with
- | Some
- (_stringappend_879_,(existT _ _stringappend_880_ _)) =>
- returnm ((_stringappend_879_,
- build_ex
- _stringappend_880_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__375 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_880_ _) :=
- w__375
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_881_ :=
- string_drop _stringappend_878_
- (build_ex
- _stringappend_880_) in
- match (reg_name_matches_prefix _stringappend_881_) with
- | Some
- (_stringappend_882_,(existT _ _stringappend_883_ _)) =>
- returnm ((_stringappend_882_,
- build_ex
- _stringappend_883_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__377 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_883_ _) :=
- w__377
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_884_ :=
- string_drop _stringappend_881_
- (build_ex
- _stringappend_883_) in
- sep_matches_prefix _stringappend_884_ >>= fun w__378 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__373 with
+ | Some (tt,(existT _ _stringappend_597_ _)) =>
+ returnm (tt, build_ex _stringappend_597_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_597_ _) =>
+ let _stringappend_598_ :=
+ string_drop _stringappend_596_
+ (build_ex _stringappend_597_) in
+ (match (reg_name_matches_prefix
+ _stringappend_598_) with
+ | Some (rs1,(existT _ _stringappend_599_ _)) =>
+ returnm (rs1, build_ex _stringappend_599_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_599_ _) =>
+ let _stringappend_600_ :=
+ string_drop _stringappend_598_
+ (build_ex _stringappend_599_) in
+ sep_matches_prefix _stringappend_600_ >>= fun w__378 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__378 with
- | Some
- (_stringappend_885_,(existT _ _stringappend_886_ _)) =>
- returnm ((_stringappend_885_,
- build_ex
- _stringappend_886_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__380 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_886_ _) :=
- w__380
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_887_ :=
- string_drop _stringappend_884_
- (build_ex
- _stringappend_886_) in
- match (reg_name_matches_prefix _stringappend_887_) with
- | Some
- (_stringappend_888_,(existT _ _stringappend_889_ _)) =>
- returnm ((_stringappend_888_,
- build_ex
- _stringappend_889_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__382 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_889_ _) :=
- w__382
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_887_
- (build_ex
- _stringappend_889_)) with
+ (match w__378 with
+ | Some (tt,(existT _ _stringappend_601_ _)) =>
+ returnm (tt, build_ex _stringappend_601_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_601_ _) =>
+ let _stringappend_602_ :=
+ string_drop _stringappend_600_
+ (build_ex _stringappend_601_) in
+ (match (reg_name_matches_prefix
+ _stringappend_602_) with
+ | Some (rs2,(existT _ _stringappend_603_ _)) =>
+ returnm (rs2, build_ex _stringappend_603_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_603_ _) =>
+ (match (string_drop _stringappend_602_
+ (build_ex _stringappend_603_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
and_boolM
- (returnm ((string_startswith _stringappend_559_
+ (returnm ((string_startswith _stringappend_381_
"mulw")
: bool))
- (let _stringappend_891_ :=
- string_drop _stringappend_559_
- (string_length "mulw") in
- match (spc_matches_prefix _stringappend_891_) with
- | Some
- (_stringappend_892_,(existT _ _stringappend_893_ _)) =>
- let _stringappend_894_ :=
- string_drop _stringappend_891_
- (build_ex
- _stringappend_893_) in
+ (let _stringappend_605_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length "mulw"))) in
+ match (spc_matches_prefix _stringappend_605_) with
+ | Some (tt,(existT _ _stringappend_606_ _)) =>
+ let _stringappend_607_ :=
+ string_drop _stringappend_605_
+ (build_ex _stringappend_606_) in
match (reg_name_matches_prefix
- _stringappend_894_) with
- | Some
- (_stringappend_895_,(existT _ _stringappend_896_ _)) =>
- let _stringappend_897_ :=
- string_drop _stringappend_894_
- (build_ex
- _stringappend_896_) in
- sep_matches_prefix _stringappend_897_ >>= fun w__385 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_607_) with
+ | Some (rd,(existT _ _stringappend_608_ _)) =>
+ let _stringappend_609_ :=
+ string_drop _stringappend_607_
+ (build_ex _stringappend_608_) in
+ sep_matches_prefix _stringappend_609_ >>= fun w__385 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__385 with
| Some
- (_stringappend_898_,(existT _ _stringappend_899_ _)) =>
- let _stringappend_900_ :=
- string_drop _stringappend_897_
- (build_ex
- _stringappend_899_) in
+ (tt,(existT _ _stringappend_610_ _)) =>
+ let _stringappend_611_ :=
+ string_drop _stringappend_609_
+ (build_ex _stringappend_610_) in
match (reg_name_matches_prefix
- _stringappend_900_) with
+ _stringappend_611_) with
| Some
- (_stringappend_901_,(existT _ _stringappend_902_ _)) =>
- let _stringappend_903_ :=
- string_drop _stringappend_900_
- (build_ex
- _stringappend_902_) in
+ (rs1,(existT _ _stringappend_612_ _)) =>
+ let _stringappend_613_ :=
+ string_drop _stringappend_611_
+ (build_ex _stringappend_612_) in
sep_matches_prefix
- _stringappend_903_ >>= fun w__386 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_613_ >>= fun w__386 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__386 with
| Some
- (_stringappend_904_,(existT _ _stringappend_905_ _)) =>
- let _stringappend_906_ :=
+ (tt,(existT _ _stringappend_614_ _)) =>
+ let _stringappend_615_ :=
string_drop
- _stringappend_903_
- (build_ex
- _stringappend_905_) in
+ _stringappend_613_
+ (build_ex _stringappend_614_) in
if ((match (reg_name_matches_prefix
- _stringappend_906_) with
+ _stringappend_615_) with
| Some
- (_stringappend_907_,(existT _ _stringappend_908_ _)) =>
+ (rs2,(existT _ _stringappend_616_ _)) =>
match (string_drop
- _stringappend_906_
- (build_ex
- _stringappend_908_)) with
+ _stringappend_615_
+ (build_ex _stringappend_616_)) with
| "" =>
true
| _ =>
@@ -20544,129 +19742,81 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__391 : bool =>
(if (w__391) then
- let _stringappend_891_ :=
- string_drop _stringappend_559_
- (string_length "mulw") in
- match (spc_matches_prefix _stringappend_891_) with
- | Some
- (_stringappend_892_,(existT _ _stringappend_893_ _)) =>
- returnm ((_stringappend_892_,
- build_ex
- _stringappend_893_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__393 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_893_ _) :=
- w__393
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_894_ :=
- string_drop _stringappend_891_
- (build_ex
- _stringappend_893_) in
- match (reg_name_matches_prefix
- _stringappend_894_) with
- | Some
- (_stringappend_895_,(existT _ _stringappend_896_ _)) =>
- returnm ((_stringappend_895_,
- build_ex
- _stringappend_896_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__395 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_896_ _) :=
- w__395
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_897_ :=
- string_drop _stringappend_894_
- (build_ex
- _stringappend_896_) in
- sep_matches_prefix _stringappend_897_ >>= fun w__396 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_605_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length "mulw"))) in
+ (match (spc_matches_prefix _stringappend_605_) with
+ | Some (tt,(existT _ _stringappend_606_ _)) =>
+ returnm (tt, build_ex _stringappend_606_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_606_ _) =>
+ let _stringappend_607_ :=
+ string_drop _stringappend_605_
+ (build_ex _stringappend_606_) in
+ (match (reg_name_matches_prefix
+ _stringappend_607_) with
+ | Some (rd,(existT _ _stringappend_608_ _)) =>
+ returnm (rd, build_ex _stringappend_608_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_608_ _) =>
+ let _stringappend_609_ :=
+ string_drop _stringappend_607_
+ (build_ex _stringappend_608_) in
+ sep_matches_prefix _stringappend_609_ >>= fun w__396 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__396 with
- | Some
- (_stringappend_898_,(existT _ _stringappend_899_ _)) =>
- returnm ((_stringappend_898_,
- build_ex
- _stringappend_899_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__398 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_899_ _) :=
- w__398
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_900_ :=
- string_drop _stringappend_897_
- (build_ex
- _stringappend_899_) in
- match (reg_name_matches_prefix
- _stringappend_900_) with
- | Some
- (_stringappend_901_,(existT _ _stringappend_902_ _)) =>
- returnm ((_stringappend_901_,
- build_ex
- _stringappend_902_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__400 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_902_ _) :=
- w__400
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_903_ :=
- string_drop _stringappend_900_
- (build_ex
- _stringappend_902_) in
- sep_matches_prefix _stringappend_903_ >>= fun w__401 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__396 with
+ | Some (tt,(existT _ _stringappend_610_ _)) =>
+ returnm (tt, build_ex _stringappend_610_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_610_ _) =>
+ let _stringappend_611_ :=
+ string_drop _stringappend_609_
+ (build_ex _stringappend_610_) in
+ (match (reg_name_matches_prefix
+ _stringappend_611_) with
+ | Some (rs1,(existT _ _stringappend_612_ _)) =>
+ returnm (rs1, build_ex _stringappend_612_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_612_ _) =>
+ let _stringappend_613_ :=
+ string_drop _stringappend_611_
+ (build_ex _stringappend_612_) in
+ sep_matches_prefix _stringappend_613_ >>= fun w__401 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__401 with
- | Some
- (_stringappend_904_,(existT _ _stringappend_905_ _)) =>
- returnm ((_stringappend_904_,
- build_ex
- _stringappend_905_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__403 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_905_ _) :=
- w__403
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_906_ :=
- string_drop _stringappend_903_
- (build_ex
- _stringappend_905_) in
- match (reg_name_matches_prefix
- _stringappend_906_) with
- | Some
- (_stringappend_907_,(existT _ _stringappend_908_ _)) =>
- returnm ((_stringappend_907_,
- build_ex
- _stringappend_908_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__405 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_908_ _) :=
- w__405
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_906_
- (build_ex
- _stringappend_908_)) with
+ (match w__401 with
+ | Some (tt,(existT _ _stringappend_614_ _)) =>
+ returnm (tt, build_ex _stringappend_614_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_614_ _) =>
+ let _stringappend_615_ :=
+ string_drop _stringappend_613_
+ (build_ex _stringappend_614_) in
+ (match (reg_name_matches_prefix
+ _stringappend_615_) with
+ | Some (rs2,(existT _ _stringappend_616_ _)) =>
+ returnm (rs2, build_ex _stringappend_616_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_616_ _) =>
+ (match (string_drop _stringappend_615_
+ (build_ex _stringappend_616_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
@@ -20674,81 +19824,74 @@ Definition assembly_backwards_matches (arg_ : string)
else
and_boolM
(returnm ((string_startswith
- _stringappend_559_ "div")
+ _stringappend_381_ "div")
: bool))
- (let _stringappend_910_ :=
- string_drop _stringappend_559_
- (string_length "div") in
+ (let _stringappend_618_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length "div"))) in
match (maybe_not_u_matches_prefix
- _stringappend_910_) with
- | Some
- (_stringappend_911_,(existT _ _stringappend_912_ _)) =>
- let _stringappend_913_ :=
- string_drop _stringappend_910_
- (build_ex
- _stringappend_912_) in
+ _stringappend_618_) with
+ | Some (s,(existT _ _stringappend_619_ _)) =>
+ let _stringappend_620_ :=
+ string_drop _stringappend_618_
+ (build_ex _stringappend_619_) in
and_boolM
(returnm ((string_startswith
- _stringappend_913_ "w")
+ _stringappend_620_ "w")
: bool))
- (let _stringappend_914_ :=
- string_drop _stringappend_913_
- (string_length "w") in
+ (let _stringappend_621_ :=
+ string_drop _stringappend_620_
+ (build_ex (projT1 (string_length
+ "w"))) in
match (spc_matches_prefix
- _stringappend_914_) with
+ _stringappend_621_) with
| Some
- (_stringappend_915_,(existT _ _stringappend_916_ _)) =>
- let _stringappend_917_ :=
- string_drop _stringappend_914_
- (build_ex
- _stringappend_916_) in
+ (tt,(existT _ _stringappend_622_ _)) =>
+ let _stringappend_623_ :=
+ string_drop _stringappend_621_
+ (build_ex _stringappend_622_) in
match (reg_name_matches_prefix
- _stringappend_917_) with
+ _stringappend_623_) with
| Some
- (_stringappend_918_,(existT _ _stringappend_919_ _)) =>
- let _stringappend_920_ :=
- string_drop _stringappend_917_
- (build_ex
- _stringappend_919_) in
+ (rd,(existT _ _stringappend_624_ _)) =>
+ let _stringappend_625_ :=
+ string_drop _stringappend_623_
+ (build_ex _stringappend_624_) in
sep_matches_prefix
- _stringappend_920_ >>= fun w__408 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_625_ >>= fun w__408 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__408 with
| Some
- (_stringappend_921_,(existT _ _stringappend_922_ _)) =>
- let _stringappend_923_ :=
+ (tt,(existT _ _stringappend_626_ _)) =>
+ let _stringappend_627_ :=
string_drop
- _stringappend_920_
- (build_ex
- _stringappend_922_) in
+ _stringappend_625_
+ (build_ex _stringappend_626_) in
match (reg_name_matches_prefix
- _stringappend_923_) with
+ _stringappend_627_) with
| Some
- (_stringappend_924_,(existT _ _stringappend_925_ _)) =>
- let _stringappend_926_ :=
+ (rs1,(existT _ _stringappend_628_ _)) =>
+ let _stringappend_629_ :=
string_drop
- _stringappend_923_
- (build_ex
- _stringappend_925_) in
+ _stringappend_627_
+ (build_ex _stringappend_628_) in
sep_matches_prefix
- _stringappend_926_ >>= fun w__409 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_629_ >>= fun w__409 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__409 with
| Some
- (_stringappend_927_,(existT _ _stringappend_928_ _)) =>
- let _stringappend_929_ :=
+ (tt,(existT _ _stringappend_630_ _)) =>
+ let _stringappend_631_ :=
string_drop
- _stringappend_926_
- (build_ex
- _stringappend_928_) in
+ _stringappend_629_
+ (build_ex _stringappend_630_) in
if ((match (reg_name_matches_prefix
- _stringappend_929_) with
+ _stringappend_631_) with
| Some
- (_stringappend_930_,(existT _ _stringappend_931_ _)) =>
+ (rs2,(existT _ _stringappend_632_ _)) =>
match (string_drop
- _stringappend_929_
- (build_ex
- _stringappend_931_)) with
+ _stringappend_631_
+ (build_ex _stringappend_632_)) with
| "" =>
true
| _ =>
@@ -20798,156 +19941,110 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__416 : bool =>
(if (w__416) then
- let _stringappend_910_ :=
- string_drop _stringappend_559_
- (string_length "div") in
- match (maybe_not_u_matches_prefix
- _stringappend_910_) with
- | Some
- (_stringappend_911_,(existT _ _stringappend_912_ _)) =>
- returnm ((_stringappend_911_,
- build_ex
- _stringappend_912_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__418 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_912_ _) :=
- w__418
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_913_ :=
- string_drop _stringappend_910_
- (build_ex
- _stringappend_912_) in
- let _stringappend_914_ :=
- string_drop _stringappend_913_
- (string_length "w") in
- match (spc_matches_prefix
- _stringappend_914_) with
- | Some
- (_stringappend_915_,(existT _ _stringappend_916_ _)) =>
- returnm ((_stringappend_915_,
- build_ex
- _stringappend_916_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__420 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_916_ _) :=
- w__420
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_917_ :=
- string_drop _stringappend_914_
- (build_ex
- _stringappend_916_) in
- match (reg_name_matches_prefix
- _stringappend_917_) with
- | Some
- (_stringappend_918_,(existT _ _stringappend_919_ _)) =>
- returnm ((_stringappend_918_,
- build_ex
- _stringappend_919_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__422 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_919_ _) :=
- w__422
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_920_ :=
- string_drop _stringappend_917_
- (build_ex
- _stringappend_919_) in
- sep_matches_prefix _stringappend_920_ >>= fun w__423 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_618_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length "div"))) in
+ (match (maybe_not_u_matches_prefix
+ _stringappend_618_) with
+ | Some (s,(existT _ _stringappend_619_ _)) =>
+ returnm (s, build_ex _stringappend_619_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(s, existT _ _stringappend_619_ _) =>
+ let _stringappend_620_ :=
+ string_drop _stringappend_618_
+ (build_ex _stringappend_619_) in
+ let _stringappend_621_ :=
+ string_drop _stringappend_620_
+ (build_ex (projT1 (string_length "w"))) in
+ (match (spc_matches_prefix
+ _stringappend_621_) with
+ | Some
+ (tt,(existT _ _stringappend_622_ _)) =>
+ returnm (tt, build_ex _stringappend_622_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_622_ _) =>
+ let _stringappend_623_ :=
+ string_drop _stringappend_621_
+ (build_ex _stringappend_622_) in
+ (match (reg_name_matches_prefix
+ _stringappend_623_) with
+ | Some
+ (rd,(existT _ _stringappend_624_ _)) =>
+ returnm (rd, build_ex _stringappend_624_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_624_ _) =>
+ let _stringappend_625_ :=
+ string_drop _stringappend_623_
+ (build_ex _stringappend_624_) in
+ sep_matches_prefix _stringappend_625_ >>= fun w__423 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__423 with
- | Some
- (_stringappend_921_,(existT _ _stringappend_922_ _)) =>
- returnm ((_stringappend_921_,
- build_ex
- _stringappend_922_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__425 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_922_ _) :=
- w__425
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_923_ :=
- string_drop _stringappend_920_
- (build_ex
- _stringappend_922_) in
- match (reg_name_matches_prefix
- _stringappend_923_) with
- | Some
- (_stringappend_924_,(existT _ _stringappend_925_ _)) =>
- returnm ((_stringappend_924_,
- build_ex
- _stringappend_925_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__427 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_925_ _) :=
- w__427
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_926_ :=
- string_drop _stringappend_923_
- (build_ex
- _stringappend_925_) in
- sep_matches_prefix _stringappend_926_ >>= fun w__428 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__423 with
+ | Some
+ (tt,(existT _ _stringappend_626_ _)) =>
+ returnm (tt, build_ex _stringappend_626_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_626_ _) =>
+ let _stringappend_627_ :=
+ string_drop _stringappend_625_
+ (build_ex _stringappend_626_) in
+ (match (reg_name_matches_prefix
+ _stringappend_627_) with
+ | Some
+ (rs1,(existT _ _stringappend_628_ _)) =>
+ returnm (rs1, build_ex _stringappend_628_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_628_ _) =>
+ let _stringappend_629_ :=
+ string_drop _stringappend_627_
+ (build_ex _stringappend_628_) in
+ sep_matches_prefix _stringappend_629_ >>= fun w__428 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__428 with
- | Some
- (_stringappend_927_,(existT _ _stringappend_928_ _)) =>
- returnm ((_stringappend_927_,
- build_ex
- _stringappend_928_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__430 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_928_ _) :=
- w__430
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_929_ :=
- string_drop _stringappend_926_
- (build_ex
- _stringappend_928_) in
- match (reg_name_matches_prefix
- _stringappend_929_) with
- | Some
- (_stringappend_930_,(existT _ _stringappend_931_ _)) =>
- returnm ((_stringappend_930_,
- build_ex
- _stringappend_931_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__432 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_931_ _) :=
- w__432
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- (match (string_drop _stringappend_929_
- (build_ex
- _stringappend_931_)) with
+ (match w__428 with
+ | Some
+ (tt,(existT _ _stringappend_630_ _)) =>
+ returnm (tt, build_ex _stringappend_630_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_630_ _) =>
+ let _stringappend_631_ :=
+ string_drop _stringappend_629_
+ (build_ex _stringappend_630_) in
+ (match (reg_name_matches_prefix
+ _stringappend_631_) with
+ | Some
+ (rs2,(existT _ _stringappend_632_ _)) =>
+ returnm (rs2, build_ex _stringappend_632_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_632_ _) =>
+ (match (string_drop _stringappend_631_
+ (build_ex _stringappend_632_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
@@ -20955,82 +20052,77 @@ Definition assembly_backwards_matches (arg_ : string)
else
and_boolM
(returnm ((string_startswith
- _stringappend_559_ "rem")
+ _stringappend_381_ "rem")
: bool))
- (let _stringappend_933_ :=
- string_drop _stringappend_559_
- (string_length "rem") in
+ (let _stringappend_634_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length
+ "rem"))) in
match (maybe_not_u_matches_prefix
- _stringappend_933_) with
+ _stringappend_634_) with
| Some
- (_stringappend_934_,(existT _ _stringappend_935_ _)) =>
- let _stringappend_936_ :=
- string_drop _stringappend_933_
- (build_ex
- _stringappend_935_) in
+ (s,(existT _ _stringappend_635_ _)) =>
+ let _stringappend_636_ :=
+ string_drop _stringappend_634_
+ (build_ex _stringappend_635_) in
and_boolM
(returnm ((string_startswith
- _stringappend_936_ "w")
+ _stringappend_636_ "w")
: bool))
- (let _stringappend_937_ :=
- string_drop _stringappend_936_
- (string_length "w") in
+ (let _stringappend_637_ :=
+ string_drop _stringappend_636_
+ (build_ex (projT1 (string_length
+ "w"))) in
match (spc_matches_prefix
- _stringappend_937_) with
+ _stringappend_637_) with
| Some
- (_stringappend_938_,(existT _ _stringappend_939_ _)) =>
- let _stringappend_940_ :=
- string_drop _stringappend_937_
- (build_ex
- _stringappend_939_) in
+ (tt,(existT _ _stringappend_638_ _)) =>
+ let _stringappend_639_ :=
+ string_drop _stringappend_637_
+ (build_ex _stringappend_638_) in
match (reg_name_matches_prefix
- _stringappend_940_) with
+ _stringappend_639_) with
| Some
- (_stringappend_941_,(existT _ _stringappend_942_ _)) =>
- let _stringappend_943_ :=
+ (rd,(existT _ _stringappend_640_ _)) =>
+ let _stringappend_641_ :=
string_drop
- _stringappend_940_
- (build_ex
- _stringappend_942_) in
+ _stringappend_639_
+ (build_ex _stringappend_640_) in
sep_matches_prefix
- _stringappend_943_ >>= fun w__435 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_641_ >>= fun w__435 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__435 with
| Some
- (_stringappend_944_,(existT _ _stringappend_945_ _)) =>
- let _stringappend_946_ :=
+ (tt,(existT _ _stringappend_642_ _)) =>
+ let _stringappend_643_ :=
string_drop
- _stringappend_943_
- (build_ex
- _stringappend_945_) in
+ _stringappend_641_
+ (build_ex _stringappend_642_) in
match (reg_name_matches_prefix
- _stringappend_946_) with
+ _stringappend_643_) with
| Some
- (_stringappend_947_,(existT _ _stringappend_948_ _)) =>
- let _stringappend_949_ :=
+ (rs1,(existT _ _stringappend_644_ _)) =>
+ let _stringappend_645_ :=
string_drop
- _stringappend_946_
- (build_ex
- _stringappend_948_) in
+ _stringappend_643_
+ (build_ex _stringappend_644_) in
sep_matches_prefix
- _stringappend_949_ >>= fun w__436 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_645_ >>= fun w__436 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__436 with
| Some
- (_stringappend_950_,(existT _ _stringappend_951_ _)) =>
- let _stringappend_952_ :=
+ (tt,(existT _ _stringappend_646_ _)) =>
+ let _stringappend_647_ :=
string_drop
- _stringappend_949_
- (build_ex
- _stringappend_951_) in
+ _stringappend_645_
+ (build_ex _stringappend_646_) in
if ((match (reg_name_matches_prefix
- _stringappend_952_) with
+ _stringappend_647_) with
| Some
- (_stringappend_953_,(existT _ _stringappend_954_ _)) =>
+ (rs2,(existT _ _stringappend_648_ _)) =>
match (string_drop
- _stringappend_952_
- (build_ex
- _stringappend_954_)) with
+ _stringappend_647_
+ (build_ex _stringappend_648_)) with
| "" =>
true
| _ =>
@@ -21083,166 +20175,112 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__443 : bool =>
(if (w__443) then
- let _stringappend_933_ :=
- string_drop _stringappend_559_
- (string_length "rem") in
- match (maybe_not_u_matches_prefix
- _stringappend_933_) with
- | Some
- (_stringappend_934_,(existT _ _stringappend_935_ _)) =>
- returnm ((_stringappend_934_,
- build_ex
- _stringappend_935_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__445 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_935_ _) :=
- w__445
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_936_ :=
- string_drop _stringappend_933_
- (build_ex
- _stringappend_935_) in
- let _stringappend_937_ :=
- string_drop _stringappend_936_
- (string_length "w") in
- match (spc_matches_prefix
- _stringappend_937_) with
- | Some
- (_stringappend_938_,(existT _ _stringappend_939_ _)) =>
- returnm ((_stringappend_938_,
- build_ex
- _stringappend_939_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__447 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_939_ _) :=
- w__447
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_940_ :=
- string_drop _stringappend_937_
- (build_ex
- _stringappend_939_) in
- match (reg_name_matches_prefix
- _stringappend_940_) with
- | Some
- (_stringappend_941_,(existT _ _stringappend_942_ _)) =>
- returnm ((_stringappend_941_,
- build_ex
- _stringappend_942_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__449 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_942_ _) :=
- w__449
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_943_ :=
- string_drop _stringappend_940_
- (build_ex
- _stringappend_942_) in
- sep_matches_prefix _stringappend_943_ >>= fun w__450 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_634_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length
+ "rem"))) in
+ (match (maybe_not_u_matches_prefix
+ _stringappend_634_) with
+ | Some
+ (s,(existT _ _stringappend_635_ _)) =>
+ returnm (s, build_ex _stringappend_635_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(s, existT _ _stringappend_635_ _) =>
+ let _stringappend_636_ :=
+ string_drop _stringappend_634_
+ (build_ex _stringappend_635_) in
+ let _stringappend_637_ :=
+ string_drop _stringappend_636_
+ (build_ex (projT1 (string_length "w"))) in
+ (match (spc_matches_prefix
+ _stringappend_637_) with
+ | Some
+ (tt,(existT _ _stringappend_638_ _)) =>
+ returnm (tt, build_ex _stringappend_638_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_638_ _) =>
+ let _stringappend_639_ :=
+ string_drop _stringappend_637_
+ (build_ex _stringappend_638_) in
+ (match (reg_name_matches_prefix
+ _stringappend_639_) with
+ | Some
+ (rd,(existT _ _stringappend_640_ _)) =>
+ returnm (rd, build_ex _stringappend_640_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_640_ _) =>
+ let _stringappend_641_ :=
+ string_drop _stringappend_639_
+ (build_ex _stringappend_640_) in
+ sep_matches_prefix _stringappend_641_ >>= fun w__450 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__450 with
- | Some
- (_stringappend_944_,(existT _ _stringappend_945_ _)) =>
- returnm ((_stringappend_944_,
- build_ex
- _stringappend_945_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__452 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_945_ _) :=
- w__452
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_946_ :=
- string_drop _stringappend_943_
- (build_ex
- _stringappend_945_) in
- match (reg_name_matches_prefix
- _stringappend_946_) with
- | Some
- (_stringappend_947_,(existT _ _stringappend_948_ _)) =>
- returnm ((_stringappend_947_,
- build_ex
- _stringappend_948_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__454 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_948_ _) :=
- w__454
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_949_ :=
- string_drop _stringappend_946_
- (build_ex
- _stringappend_948_) in
- sep_matches_prefix _stringappend_949_ >>= fun w__455 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__450 with
+ | Some
+ (tt,(existT _ _stringappend_642_ _)) =>
+ returnm (tt, build_ex _stringappend_642_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_642_ _) =>
+ let _stringappend_643_ :=
+ string_drop _stringappend_641_
+ (build_ex _stringappend_642_) in
+ (match (reg_name_matches_prefix
+ _stringappend_643_) with
+ | Some
+ (rs1,(existT _ _stringappend_644_ _)) =>
+ returnm (rs1, build_ex _stringappend_644_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_644_ _) =>
+ let _stringappend_645_ :=
+ string_drop _stringappend_643_
+ (build_ex _stringappend_644_) in
+ sep_matches_prefix _stringappend_645_ >>= fun w__455 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__455 with
- | Some
- (_stringappend_950_,(existT _ _stringappend_951_ _)) =>
- returnm ((_stringappend_950_,
- build_ex
- _stringappend_951_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__457 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_951_ _) :=
- w__457
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_952_ :=
- string_drop _stringappend_949_
- (build_ex
- _stringappend_951_) in
- match (reg_name_matches_prefix
- _stringappend_952_) with
- | Some
- (_stringappend_953_,(existT _ _stringappend_954_ _)) =>
- returnm ((_stringappend_953_,
- build_ex
- _stringappend_954_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__459 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_954_ _) :=
- w__459
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- (match (string_drop _stringappend_952_
- (build_ex
- _stringappend_954_)) with
+ (match w__455 with
+ | Some
+ (tt,(existT _ _stringappend_646_ _)) =>
+ returnm (tt, build_ex _stringappend_646_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_646_ _) =>
+ let _stringappend_647_ :=
+ string_drop _stringappend_645_
+ (build_ex _stringappend_646_) in
+ (match (reg_name_matches_prefix
+ _stringappend_647_) with
+ | Some
+ (rs2,(existT _ _stringappend_648_ _)) =>
+ returnm (rs2, build_ex _stringappend_648_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_648_ _) =>
+ (match (string_drop _stringappend_647_
+ (build_ex _stringappend_648_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
@@ -21250,51 +20288,48 @@ Definition assembly_backwards_matches (arg_ : string)
else
and_boolM
(returnm ((string_startswith
- _stringappend_559_
+ _stringappend_381_
"fence")
: bool))
- (let _stringappend_956_ :=
- string_drop _stringappend_559_
- (string_length "fence") in
+ (let _stringappend_650_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length
+ "fence"))) in
match (spc_matches_prefix
- _stringappend_956_) with
+ _stringappend_650_) with
| Some
- (_stringappend_957_,(existT _ _stringappend_958_ _)) =>
- let _stringappend_959_ :=
- string_drop _stringappend_956_
- (build_ex
- _stringappend_958_) in
+ (tt,(existT _ _stringappend_651_ _)) =>
+ let _stringappend_652_ :=
+ string_drop _stringappend_650_
+ (build_ex _stringappend_651_) in
fence_bits_matches_prefix
- _stringappend_959_ >>= fun w__462 : option ((mword 4 * {n : Z & ArithFact (n >=
+ _stringappend_652_ >>= fun w__462 : option ((mword 4 * {n : Z & ArithFact (n >=
0)})) =>
match w__462 with
| Some
- (_stringappend_960_,(existT _ _stringappend_961_ _)) =>
- let _stringappend_962_ :=
- string_drop _stringappend_959_
- (build_ex
- _stringappend_961_) in
+ (pred,(existT _ _stringappend_653_ _)) =>
+ let _stringappend_654_ :=
+ string_drop _stringappend_652_
+ (build_ex _stringappend_653_) in
sep_matches_prefix
- _stringappend_962_ >>= fun w__463 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_654_ >>= fun w__463 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__463 with
| Some
- (_stringappend_963_,(existT _ _stringappend_964_ _)) =>
- let _stringappend_965_ :=
+ (tt,(existT _ _stringappend_655_ _)) =>
+ let _stringappend_656_ :=
string_drop
- _stringappend_962_
- (build_ex
- _stringappend_964_) in
+ _stringappend_654_
+ (build_ex _stringappend_655_) in
fence_bits_matches_prefix
- _stringappend_965_ >>= fun w__464 : option ((mword 4 * {n : Z & ArithFact (n >=
+ _stringappend_656_ >>= fun w__464 : option ((mword 4 * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__464 with
| Some
- (_stringappend_966_,(existT _ _stringappend_967_ _)) =>
+ (succ,(existT _ _stringappend_657_ _)) =>
match (string_drop
- _stringappend_965_
- (build_ex
- _stringappend_967_)) with
+ _stringappend_656_
+ (build_ex _stringappend_657_)) with
| "" =>
true
| _ =>
@@ -21323,112 +20358,77 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__468 : bool =>
(if (w__468) then
- let _stringappend_956_ :=
- string_drop _stringappend_559_
- (string_length "fence") in
- match (spc_matches_prefix
- _stringappend_956_) with
- | Some
- (_stringappend_957_,(existT _ _stringappend_958_ _)) =>
- returnm ((_stringappend_957_,
- build_ex
- _stringappend_958_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__470 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_958_ _) :=
- w__470
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_959_ :=
- string_drop _stringappend_956_
- (build_ex
- _stringappend_958_) in
+ let _stringappend_650_ :=
+ string_drop _stringappend_381_
+ (build_ex (projT1 (string_length
+ "fence"))) in
+ (match (spc_matches_prefix
+ _stringappend_650_) with
+ | Some
+ (tt,(existT _ _stringappend_651_ _)) =>
+ returnm (tt, build_ex _stringappend_651_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_651_ _) =>
+ let _stringappend_652_ :=
+ string_drop _stringappend_650_
+ (build_ex _stringappend_651_) in
fence_bits_matches_prefix
- _stringappend_959_ >>= fun w__471 : option ((mword 4 * {n : Z & ArithFact (n >=
+ _stringappend_652_ >>= fun w__471 : option ((mword 4 * {n : Z & ArithFact (n >=
0)})) =>
- match w__471 with
- | Some
- (_stringappend_960_,(existT _ _stringappend_961_ _)) =>
- returnm ((_stringappend_960_,
- build_ex
- _stringappend_961_)
- : (mword 4 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 4 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__473 : (mword 4 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(pred, existT _ _stringappend_961_ _) :=
- w__473
- : (mword 4 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_962_ :=
- string_drop _stringappend_959_
- (build_ex
- _stringappend_961_) in
- sep_matches_prefix _stringappend_962_ >>= fun w__474 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__471 with
+ | Some
+ (pred,(existT _ _stringappend_653_ _)) =>
+ returnm (pred, build_ex _stringappend_653_)
+ | _ =>
+ exit tt
+ : M ((mword 4 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 4 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(pred, existT _ _stringappend_653_ _) =>
+ let _stringappend_654_ :=
+ string_drop _stringappend_652_
+ (build_ex _stringappend_653_) in
+ sep_matches_prefix _stringappend_654_ >>= fun w__474 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__474 with
- | Some
- (_stringappend_963_,(existT _ _stringappend_964_ _)) =>
- returnm ((_stringappend_963_,
- build_ex
- _stringappend_964_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__476 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_964_ _) :=
- w__476
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_965_ :=
- string_drop _stringappend_962_
- (build_ex
- _stringappend_964_) in
+ (match w__474 with
+ | Some
+ (tt,(existT _ _stringappend_655_ _)) =>
+ returnm (tt, build_ex _stringappend_655_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_655_ _) =>
+ let _stringappend_656_ :=
+ string_drop _stringappend_654_
+ (build_ex _stringappend_655_) in
fence_bits_matches_prefix
- _stringappend_965_ >>= fun w__477 : option ((mword 4 * {n : Z & ArithFact (n >=
+ _stringappend_656_ >>= fun w__477 : option ((mword 4 * {n : Z & ArithFact (n >=
0)})) =>
- match w__477 with
- | Some
- (_stringappend_966_,(existT _ _stringappend_967_ _)) =>
- returnm ((_stringappend_966_,
- build_ex
- _stringappend_967_)
- : (mword 4 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 4 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__479 : (mword 4 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(succ, existT _ _stringappend_967_ _) :=
- w__479
- : (mword 4 * {n : Z & ArithFact (n >=
- 0)}) in
+ (match w__477 with
+ | Some
+ (succ,(existT _ _stringappend_657_ _)) =>
+ returnm (succ, build_ex _stringappend_657_)
+ | _ =>
+ exit tt
+ : M ((mword 4 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 4 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(succ, existT _ _stringappend_657_ _) =>
(match (string_drop
- _stringappend_965_
- (build_ex
- _stringappend_967_)) with
+ _stringappend_656_
+ (build_ex _stringappend_657_)) with
| "" => returnm (true : bool)
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (match _stringappend_559_ with
+ (match _stringappend_381_ with
| "fence.i" =>
returnm (true : bool)
| "ecall" => returnm (true : bool)
@@ -21436,53 +20436,50 @@ Definition assembly_backwards_matches (arg_ : string)
| "sret" => returnm (true : bool)
| "ebreak" => returnm (true : bool)
| "wfi" => returnm (true : bool)
- | _stringappend_559_ =>
+ | _stringappend_381_ =>
and_boolM
(returnm ((string_startswith
- _stringappend_559_
+ _stringappend_381_
"sfence.vma")
: bool))
- (let _stringappend_969_ :=
+ (let _stringappend_659_ :=
string_drop
- _stringappend_559_
- (string_length "sfence.vma") in
+ _stringappend_381_
+ (build_ex (projT1 (string_length
+ "sfence.vma"))) in
match (spc_matches_prefix
- _stringappend_969_) with
+ _stringappend_659_) with
| Some
- (_stringappend_970_,(existT _ _stringappend_971_ _)) =>
- let _stringappend_972_ :=
+ (tt,(existT _ _stringappend_660_ _)) =>
+ let _stringappend_661_ :=
string_drop
- _stringappend_969_
- (build_ex
- _stringappend_971_) in
+ _stringappend_659_
+ (build_ex _stringappend_660_) in
match (reg_name_matches_prefix
- _stringappend_972_) with
+ _stringappend_661_) with
| Some
- (_stringappend_973_,(existT _ _stringappend_974_ _)) =>
- let _stringappend_975_ :=
+ (rs1,(existT _ _stringappend_662_ _)) =>
+ let _stringappend_663_ :=
string_drop
- _stringappend_972_
- (build_ex
- _stringappend_974_) in
+ _stringappend_661_
+ (build_ex _stringappend_662_) in
sep_matches_prefix
- _stringappend_975_ >>= fun w__482 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_663_ >>= fun w__482 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__482 with
| Some
- (_stringappend_976_,(existT _ _stringappend_977_ _)) =>
- let _stringappend_978_ :=
+ (tt,(existT _ _stringappend_664_ _)) =>
+ let _stringappend_665_ :=
string_drop
- _stringappend_975_
- (build_ex
- _stringappend_977_) in
+ _stringappend_663_
+ (build_ex _stringappend_664_) in
if ((match (reg_name_matches_prefix
- _stringappend_978_) with
+ _stringappend_665_) with
| Some
- (_stringappend_979_,(existT _ _stringappend_980_ _)) =>
+ (rs2,(existT _ _stringappend_666_ _)) =>
match (string_drop
- _stringappend_978_
- (build_ex
- _stringappend_980_)) with
+ _stringappend_665_
+ (build_ex _stringappend_666_)) with
| "" =>
true
| _ =>
@@ -21515,108 +20512,72 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__485 : bool =>
(if (w__485) then
- let _stringappend_969_ :=
+ let _stringappend_659_ :=
string_drop
- _stringappend_559_
- (string_length
- "sfence.vma") in
- match (spc_matches_prefix
- _stringappend_969_) with
- | Some
- (_stringappend_970_,(existT _ _stringappend_971_ _)) =>
- returnm ((_stringappend_970_,
- build_ex
- _stringappend_971_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__487 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_971_ _) :=
- w__487
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_972_ :=
+ _stringappend_381_
+ (build_ex (projT1 (string_length
+ "sfence.vma"))) in
+ (match (spc_matches_prefix
+ _stringappend_659_) with
+ | Some
+ (tt,(existT _ _stringappend_660_ _)) =>
+ returnm (tt, build_ex _stringappend_660_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_660_ _) =>
+ let _stringappend_661_ :=
string_drop
- _stringappend_969_
- (build_ex
- _stringappend_971_) in
- match (reg_name_matches_prefix
- _stringappend_972_) with
- | Some
- (_stringappend_973_,(existT _ _stringappend_974_ _)) =>
- returnm ((_stringappend_973_,
- build_ex
- _stringappend_974_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__489 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_974_ _) :=
- w__489
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_975_ :=
+ _stringappend_659_
+ (build_ex _stringappend_660_) in
+ (match (reg_name_matches_prefix
+ _stringappend_661_) with
+ | Some
+ (rs1,(existT _ _stringappend_662_ _)) =>
+ returnm (rs1, build_ex _stringappend_662_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_662_ _) =>
+ let _stringappend_663_ :=
string_drop
- _stringappend_972_
- (build_ex
- _stringappend_974_) in
+ _stringappend_661_
+ (build_ex _stringappend_662_) in
sep_matches_prefix
- _stringappend_975_ >>= fun w__490 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_663_ >>= fun w__490 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__490 with
- | Some
- (_stringappend_976_,(existT _ _stringappend_977_ _)) =>
- returnm ((_stringappend_976_,
- build_ex
- _stringappend_977_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__492 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_977_ _) :=
- w__492
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_978_ :=
+ (match w__490 with
+ | Some
+ (tt,(existT _ _stringappend_664_ _)) =>
+ returnm (tt, build_ex _stringappend_664_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_664_ _) =>
+ let _stringappend_665_ :=
string_drop
- _stringappend_975_
- (build_ex
- _stringappend_977_) in
- match (reg_name_matches_prefix
- _stringappend_978_) with
- | Some
- (_stringappend_979_,(existT _ _stringappend_980_ _)) =>
- returnm ((_stringappend_979_,
- build_ex
- _stringappend_980_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__494 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_980_ _) :=
- w__494
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_663_
+ (build_ex _stringappend_664_) in
+ (match (reg_name_matches_prefix
+ _stringappend_665_) with
+ | Some
+ (rs2,(existT _ _stringappend_666_ _)) =>
+ returnm (rs2, build_ex _stringappend_666_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_666_ _) =>
(match (string_drop
- _stringappend_978_
- (build_ex
- _stringappend_980_)) with
+ _stringappend_665_
+ (build_ex _stringappend_666_)) with
| "" =>
returnm (true : bool)
| _ => exit tt : M (bool)
@@ -21625,78 +20586,72 @@ Definition assembly_backwards_matches (arg_ : string)
else
and_boolM
(returnm ((string_startswith
- _stringappend_559_
+ _stringappend_381_
"lr.")
: bool))
- (let _stringappend_982_ :=
+ (let _stringappend_668_ :=
string_drop
- _stringappend_559_
- (string_length "lr.") in
+ _stringappend_381_
+ (build_ex (projT1 (string_length
+ "lr."))) in
match (maybe_aq_matches_prefix
- _stringappend_982_) with
+ _stringappend_668_) with
| Some
- (_stringappend_983_,(existT _ _stringappend_984_ _)) =>
- let _stringappend_985_ :=
+ (aq,(existT _ _stringappend_669_ _)) =>
+ let _stringappend_670_ :=
string_drop
- _stringappend_982_
- (build_ex
- _stringappend_984_) in
+ _stringappend_668_
+ (build_ex _stringappend_669_) in
match (maybe_rl_matches_prefix
- _stringappend_985_) with
+ _stringappend_670_) with
| Some
- (_stringappend_986_,(existT _ _stringappend_987_ _)) =>
- let _stringappend_988_ :=
+ (rl,(existT _ _stringappend_671_ _)) =>
+ let _stringappend_672_ :=
string_drop
- _stringappend_985_
- (build_ex
- _stringappend_987_) in
+ _stringappend_670_
+ (build_ex _stringappend_671_) in
match (size_mnemonic_matches_prefix
- _stringappend_988_) with
+ _stringappend_672_) with
| Some
- (_stringappend_989_,(existT _ _stringappend_990_ _)) =>
- let _stringappend_991_ :=
+ (size,(existT _ _stringappend_673_ _)) =>
+ let _stringappend_674_ :=
string_drop
- _stringappend_988_
- (build_ex
- _stringappend_990_) in
+ _stringappend_672_
+ (build_ex _stringappend_673_) in
match (spc_matches_prefix
- _stringappend_991_) with
+ _stringappend_674_) with
| Some
- (_stringappend_992_,(existT _ _stringappend_993_ _)) =>
- let _stringappend_994_ :=
+ (tt,(existT _ _stringappend_675_ _)) =>
+ let _stringappend_676_ :=
string_drop
- _stringappend_991_
- (build_ex
- _stringappend_993_) in
+ _stringappend_674_
+ (build_ex _stringappend_675_) in
match (reg_name_matches_prefix
- _stringappend_994_) with
+ _stringappend_676_) with
| Some
- (_stringappend_995_,(existT _ _stringappend_996_ _)) =>
- let _stringappend_997_ :=
+ (rd,(existT _ _stringappend_677_ _)) =>
+ let _stringappend_678_ :=
string_drop
- _stringappend_994_
- (build_ex
- _stringappend_996_) in
+ _stringappend_676_
+ (build_ex _stringappend_677_) in
sep_matches_prefix
- _stringappend_997_ >>= fun w__497 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_678_ >>= fun w__497 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__497 with
| Some
- (_stringappend_998_,(existT _ _stringappend_999_ _)) =>
- let _stringappend_1000_ :=
+ (tt,(existT _ _stringappend_679_ _)) =>
+ let _stringappend_680_ :=
string_drop
- _stringappend_997_
- (build_ex
- _stringappend_999_) in
+ _stringappend_678_
+ (build_ex _stringappend_679_) in
if
((match (reg_name_matches_prefix
- _stringappend_1000_) with
+ _stringappend_680_) with
| Some
- (_stringappend_1001_,(existT _ _stringappend_1002_ _)) =>
+ (rs1,(existT _ _stringappend_681_ _)) =>
match (string_drop
- _stringappend_1000_
- (build_ex
- _stringappend_1002_)) with
+ _stringappend_680_
+ (build_ex _stringappend_681_)) with
| "" =>
true
| _ =>
@@ -21763,179 +20718,117 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__503 : bool =>
(if (w__503) then
- let _stringappend_982_ :=
+ let _stringappend_668_ :=
string_drop
- _stringappend_559_
- (string_length "lr.") in
- match (maybe_aq_matches_prefix
- _stringappend_982_) with
- | Some
- (_stringappend_983_,(existT _ _stringappend_984_ _)) =>
- returnm ((_stringappend_983_,
- build_ex
- _stringappend_984_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__505 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(aq, existT _ _stringappend_984_ _) :=
- w__505
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_985_ :=
+ _stringappend_381_
+ (build_ex (projT1 (string_length
+ "lr."))) in
+ (match (maybe_aq_matches_prefix
+ _stringappend_668_) with
+ | Some
+ (aq,(existT _ _stringappend_669_ _)) =>
+ returnm (aq, build_ex _stringappend_669_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(aq, existT _ _stringappend_669_ _) =>
+ let _stringappend_670_ :=
string_drop
- _stringappend_982_
- (build_ex
- _stringappend_984_) in
- match (maybe_rl_matches_prefix
- _stringappend_985_) with
- | Some
- (_stringappend_986_,(existT _ _stringappend_987_ _)) =>
- returnm ((_stringappend_986_,
- build_ex
- _stringappend_987_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__507 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rl, existT _ _stringappend_987_ _) :=
- w__507
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_988_ :=
+ _stringappend_668_
+ (build_ex _stringappend_669_) in
+ (match (maybe_rl_matches_prefix
+ _stringappend_670_) with
+ | Some
+ (rl,(existT _ _stringappend_671_ _)) =>
+ returnm (rl, build_ex _stringappend_671_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rl, existT _ _stringappend_671_ _) =>
+ let _stringappend_672_ :=
string_drop
- _stringappend_985_
- (build_ex
- _stringappend_987_) in
- match (size_mnemonic_matches_prefix
- _stringappend_988_) with
- | Some
- (_stringappend_989_,(existT _ _stringappend_990_ _)) =>
- returnm ((_stringappend_989_,
- build_ex
- _stringappend_990_)
- : (word_width * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((word_width * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__509 : (word_width * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(size, existT _ _stringappend_990_ _) :=
- w__509
- : (word_width * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_991_ :=
+ _stringappend_670_
+ (build_ex _stringappend_671_) in
+ (match (size_mnemonic_matches_prefix
+ _stringappend_672_) with
+ | Some
+ (size,(existT _ _stringappend_673_ _)) =>
+ returnm (size, build_ex _stringappend_673_)
+ | _ =>
+ exit tt
+ : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(size, existT _ _stringappend_673_ _) =>
+ let _stringappend_674_ :=
string_drop
- _stringappend_988_
- (build_ex
- _stringappend_990_) in
- match (spc_matches_prefix
- _stringappend_991_) with
- | Some
- (_stringappend_992_,(existT _ _stringappend_993_ _)) =>
- returnm ((_stringappend_992_,
- build_ex
- _stringappend_993_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__511 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_993_ _) :=
- w__511
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_994_ :=
+ _stringappend_672_
+ (build_ex _stringappend_673_) in
+ (match (spc_matches_prefix
+ _stringappend_674_) with
+ | Some
+ (tt,(existT _ _stringappend_675_ _)) =>
+ returnm (tt, build_ex _stringappend_675_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_675_ _) =>
+ let _stringappend_676_ :=
string_drop
- _stringappend_991_
- (build_ex
- _stringappend_993_) in
- match (reg_name_matches_prefix
- _stringappend_994_) with
- | Some
- (_stringappend_995_,(existT _ _stringappend_996_ _)) =>
- returnm ((_stringappend_995_,
- build_ex
- _stringappend_996_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__513 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_996_ _) :=
- w__513
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_997_ :=
+ _stringappend_674_
+ (build_ex _stringappend_675_) in
+ (match (reg_name_matches_prefix
+ _stringappend_676_) with
+ | Some
+ (rd,(existT _ _stringappend_677_ _)) =>
+ returnm (rd, build_ex _stringappend_677_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_677_ _) =>
+ let _stringappend_678_ :=
string_drop
- _stringappend_994_
- (build_ex
- _stringappend_996_) in
+ _stringappend_676_
+ (build_ex _stringappend_677_) in
sep_matches_prefix
- _stringappend_997_ >>= fun w__514 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_678_ >>= fun w__514 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__514 with
- | Some
- (_stringappend_998_,(existT _ _stringappend_999_ _)) =>
- returnm ((_stringappend_998_,
- build_ex
- _stringappend_999_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__516 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_999_ _) :=
- w__516
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1000_ :=
+ (match w__514 with
+ | Some
+ (tt,(existT _ _stringappend_679_ _)) =>
+ returnm (tt, build_ex _stringappend_679_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_679_ _) =>
+ let _stringappend_680_ :=
string_drop
- _stringappend_997_
- (build_ex
- _stringappend_999_) in
- match (reg_name_matches_prefix
- _stringappend_1000_) with
- | Some
- (_stringappend_1001_,(existT _ _stringappend_1002_ _)) =>
- returnm ((_stringappend_1001_,
- build_ex
- _stringappend_1002_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__518 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1002_ _) :=
- w__518
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_678_
+ (build_ex _stringappend_679_) in
+ (match (reg_name_matches_prefix
+ _stringappend_680_) with
+ | Some
+ (rs1,(existT _ _stringappend_681_ _)) =>
+ returnm (rs1, build_ex _stringappend_681_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_681_ _) =>
(match (string_drop
- _stringappend_1000_
- (build_ex
- _stringappend_1002_)) with
+ _stringappend_680_
+ (build_ex _stringappend_681_)) with
| "" =>
returnm (true : bool)
| _ => exit tt : M (bool)
@@ -21944,99 +20837,91 @@ Definition assembly_backwards_matches (arg_ : string)
else
and_boolM
(returnm ((string_startswith
- _stringappend_559_
+ _stringappend_381_
"sc.")
: bool))
- (let _stringappend_1004_ :=
+ (let _stringappend_683_ :=
string_drop
- _stringappend_559_
- (string_length "sc.") in
+ _stringappend_381_
+ (build_ex (projT1 (string_length
+ "sc."))) in
match (maybe_aq_matches_prefix
- _stringappend_1004_) with
+ _stringappend_683_) with
| Some
- (_stringappend_1005_,(existT _ _stringappend_1006_ _)) =>
- let _stringappend_1007_ :=
+ (aq,(existT _ _stringappend_684_ _)) =>
+ let _stringappend_685_ :=
string_drop
- _stringappend_1004_
- (build_ex
- _stringappend_1006_) in
+ _stringappend_683_
+ (build_ex _stringappend_684_) in
match (maybe_rl_matches_prefix
- _stringappend_1007_) with
+ _stringappend_685_) with
| Some
- (_stringappend_1008_,(existT _ _stringappend_1009_ _)) =>
- let _stringappend_1010_ :=
+ (rl,(existT _ _stringappend_686_ _)) =>
+ let _stringappend_687_ :=
string_drop
- _stringappend_1007_
- (build_ex
- _stringappend_1009_) in
+ _stringappend_685_
+ (build_ex _stringappend_686_) in
match (size_mnemonic_matches_prefix
- _stringappend_1010_) with
+ _stringappend_687_) with
| Some
- (_stringappend_1011_,(existT _ _stringappend_1012_ _)) =>
- let _stringappend_1013_ :=
+ (size,(existT _ _stringappend_688_ _)) =>
+ let _stringappend_689_ :=
string_drop
- _stringappend_1010_
- (build_ex
- _stringappend_1012_) in
+ _stringappend_687_
+ (build_ex _stringappend_688_) in
match (spc_matches_prefix
- _stringappend_1013_) with
+ _stringappend_689_) with
| Some
- (_stringappend_1014_,(existT _ _stringappend_1015_ _)) =>
- let _stringappend_1016_ :=
+ (tt,(existT _ _stringappend_690_ _)) =>
+ let _stringappend_691_ :=
string_drop
- _stringappend_1013_
- (build_ex
- _stringappend_1015_) in
+ _stringappend_689_
+ (build_ex _stringappend_690_) in
match (reg_name_matches_prefix
- _stringappend_1016_) with
+ _stringappend_691_) with
| Some
- (_stringappend_1017_,(existT _ _stringappend_1018_ _)) =>
- let _stringappend_1019_ :=
+ (rd,(existT _ _stringappend_692_ _)) =>
+ let _stringappend_693_ :=
string_drop
- _stringappend_1016_
- (build_ex
- _stringappend_1018_) in
+ _stringappend_691_
+ (build_ex _stringappend_692_) in
sep_matches_prefix
- _stringappend_1019_ >>= fun w__521 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_693_ >>= fun w__521 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__521 with
| Some
- (_stringappend_1020_,(existT _ _stringappend_1021_ _)) =>
- let _stringappend_1022_ :=
+ (tt,(existT _ _stringappend_694_ _)) =>
+ let _stringappend_695_ :=
string_drop
- _stringappend_1019_
- (build_ex
- _stringappend_1021_) in
+ _stringappend_693_
+ (build_ex _stringappend_694_) in
match (reg_name_matches_prefix
- _stringappend_1022_) with
+ _stringappend_695_) with
| Some
- (_stringappend_1023_,(existT _ _stringappend_1024_ _)) =>
- let _stringappend_1025_ :=
+ (rs1,(existT _ _stringappend_696_ _)) =>
+ let _stringappend_697_ :=
string_drop
- _stringappend_1022_
- (build_ex
- _stringappend_1024_) in
+ _stringappend_695_
+ (build_ex _stringappend_696_) in
sep_matches_prefix
- _stringappend_1025_ >>= fun w__522 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_697_ >>= fun w__522 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if
((match w__522 with
| Some
- (_stringappend_1026_,(existT _ _stringappend_1027_ _)) =>
- let _stringappend_1028_ :=
+ (tt,(existT _ _stringappend_698_ _)) =>
+ let _stringappend_699_ :=
string_drop
- _stringappend_1025_
- (build_ex
- _stringappend_1027_) in
+ _stringappend_697_
+ (build_ex _stringappend_698_) in
if
((match (reg_name_matches_prefix
- _stringappend_1028_) with
+ _stringappend_699_) with
| Some
- (_stringappend_1029_,(existT _ _stringappend_1030_ _)) =>
+ (rs2,(existT _ _stringappend_700_ _)) =>
match (string_drop
- _stringappend_1028_
- (build_ex
- _stringappend_1030_)) with
+ _stringappend_699_
+ (build_ex _stringappend_700_)) with
| "" =>
true
| _ =>
@@ -22128,229 +21013,149 @@ Definition assembly_backwards_matches (arg_ : string)
else false)
: bool)) >>= fun w__530 : bool =>
(if (w__530) then
- let _stringappend_1004_ :=
+ let _stringappend_683_ :=
string_drop
- _stringappend_559_
- (string_length "sc.") in
- match (maybe_aq_matches_prefix
- _stringappend_1004_) with
- | Some
- (_stringappend_1005_,(existT _ _stringappend_1006_ _)) =>
- returnm ((_stringappend_1005_,
- build_ex
- _stringappend_1006_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__532 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(aq, existT _ _stringappend_1006_ _) :=
- w__532
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1007_ :=
+ _stringappend_381_
+ (build_ex (projT1 (string_length
+ "sc."))) in
+ (match (maybe_aq_matches_prefix
+ _stringappend_683_) with
+ | Some
+ (aq,(existT _ _stringappend_684_ _)) =>
+ returnm (aq, build_ex _stringappend_684_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(aq, existT _ _stringappend_684_ _) =>
+ let _stringappend_685_ :=
string_drop
- _stringappend_1004_
- (build_ex
- _stringappend_1006_) in
- match (maybe_rl_matches_prefix
- _stringappend_1007_) with
- | Some
- (_stringappend_1008_,(existT _ _stringappend_1009_ _)) =>
- returnm ((_stringappend_1008_,
- build_ex
- _stringappend_1009_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__534 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rl, existT _ _stringappend_1009_ _) :=
- w__534
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1010_ :=
+ _stringappend_683_
+ (build_ex _stringappend_684_) in
+ (match (maybe_rl_matches_prefix
+ _stringappend_685_) with
+ | Some
+ (rl,(existT _ _stringappend_686_ _)) =>
+ returnm (rl, build_ex _stringappend_686_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rl, existT _ _stringappend_686_ _) =>
+ let _stringappend_687_ :=
string_drop
- _stringappend_1007_
- (build_ex
- _stringappend_1009_) in
- match (size_mnemonic_matches_prefix
- _stringappend_1010_) with
- | Some
- (_stringappend_1011_,(existT _ _stringappend_1012_ _)) =>
- returnm ((_stringappend_1011_,
- build_ex
- _stringappend_1012_)
- : (word_width * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((word_width * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__536 : (word_width * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(size, existT _ _stringappend_1012_ _) :=
- w__536
- : (word_width * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1013_ :=
+ _stringappend_685_
+ (build_ex _stringappend_686_) in
+ (match (size_mnemonic_matches_prefix
+ _stringappend_687_) with
+ | Some
+ (size,(existT _ _stringappend_688_ _)) =>
+ returnm (size, build_ex _stringappend_688_)
+ | _ =>
+ exit tt
+ : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(size, existT _ _stringappend_688_ _) =>
+ let _stringappend_689_ :=
string_drop
- _stringappend_1010_
- (build_ex
- _stringappend_1012_) in
- match (spc_matches_prefix
- _stringappend_1013_) with
- | Some
- (_stringappend_1014_,(existT _ _stringappend_1015_ _)) =>
- returnm ((_stringappend_1014_,
- build_ex
- _stringappend_1015_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__538 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1015_ _) :=
- w__538
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1016_ :=
+ _stringappend_687_
+ (build_ex _stringappend_688_) in
+ (match (spc_matches_prefix
+ _stringappend_689_) with
+ | Some
+ (tt,(existT _ _stringappend_690_ _)) =>
+ returnm (tt, build_ex _stringappend_690_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_690_ _) =>
+ let _stringappend_691_ :=
string_drop
- _stringappend_1013_
- (build_ex
- _stringappend_1015_) in
- match (reg_name_matches_prefix
- _stringappend_1016_) with
- | Some
- (_stringappend_1017_,(existT _ _stringappend_1018_ _)) =>
- returnm ((_stringappend_1017_,
- build_ex
- _stringappend_1018_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__540 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1018_ _) :=
- w__540
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1019_ :=
+ _stringappend_689_
+ (build_ex _stringappend_690_) in
+ (match (reg_name_matches_prefix
+ _stringappend_691_) with
+ | Some
+ (rd,(existT _ _stringappend_692_ _)) =>
+ returnm (rd, build_ex _stringappend_692_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_692_ _) =>
+ let _stringappend_693_ :=
string_drop
- _stringappend_1016_
- (build_ex
- _stringappend_1018_) in
+ _stringappend_691_
+ (build_ex _stringappend_692_) in
sep_matches_prefix
- _stringappend_1019_ >>= fun w__541 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_693_ >>= fun w__541 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__541 with
- | Some
- (_stringappend_1020_,(existT _ _stringappend_1021_ _)) =>
- returnm ((_stringappend_1020_,
- build_ex
- _stringappend_1021_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__543 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1021_ _) :=
- w__543
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1022_ :=
+ (match w__541 with
+ | Some
+ (tt,(existT _ _stringappend_694_ _)) =>
+ returnm (tt, build_ex _stringappend_694_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_694_ _) =>
+ let _stringappend_695_ :=
string_drop
- _stringappend_1019_
- (build_ex
- _stringappend_1021_) in
- match (reg_name_matches_prefix
- _stringappend_1022_) with
- | Some
- (_stringappend_1023_,(existT _ _stringappend_1024_ _)) =>
- returnm ((_stringappend_1023_,
- build_ex
- _stringappend_1024_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__545 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1024_ _) :=
- w__545
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1025_ :=
+ _stringappend_693_
+ (build_ex _stringappend_694_) in
+ (match (reg_name_matches_prefix
+ _stringappend_695_) with
+ | Some
+ (rs1,(existT _ _stringappend_696_ _)) =>
+ returnm (rs1, build_ex _stringappend_696_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_696_ _) =>
+ let _stringappend_697_ :=
string_drop
- _stringappend_1022_
- (build_ex
- _stringappend_1024_) in
+ _stringappend_695_
+ (build_ex _stringappend_696_) in
sep_matches_prefix
- _stringappend_1025_ >>= fun w__546 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_697_ >>= fun w__546 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__546 with
- | Some
- (_stringappend_1026_,(existT _ _stringappend_1027_ _)) =>
- returnm ((_stringappend_1026_,
- build_ex
- _stringappend_1027_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__548 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1027_ _) :=
- w__548
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1028_ :=
+ (match w__546 with
+ | Some
+ (tt,(existT _ _stringappend_698_ _)) =>
+ returnm (tt, build_ex _stringappend_698_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_698_ _) =>
+ let _stringappend_699_ :=
string_drop
- _stringappend_1025_
- (build_ex
- _stringappend_1027_) in
- match (reg_name_matches_prefix
- _stringappend_1028_) with
- | Some
- (_stringappend_1029_,(existT _ _stringappend_1030_ _)) =>
- returnm ((_stringappend_1029_,
- build_ex
- _stringappend_1030_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__550 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_1030_ _) :=
- w__550
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_697_
+ (build_ex _stringappend_698_) in
+ (match (reg_name_matches_prefix
+ _stringappend_699_) with
+ | Some
+ (rs2,(existT _ _stringappend_700_ _)) =>
+ returnm (rs2, build_ex _stringappend_700_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_700_ _) =>
(match (string_drop
- _stringappend_1028_
- (build_ex
- _stringappend_1030_)) with
+ _stringappend_699_
+ (build_ex _stringappend_700_)) with
| "" =>
returnm (true
: bool)
@@ -22360,110 +21165,100 @@ Definition assembly_backwards_matches (arg_ : string)
: M (bool)
else
match (amo_mnemonic_matches_prefix
- _stringappend_559_) with
+ _stringappend_381_) with
| Some
- (_stringappend_1032_,(existT _ _stringappend_1033_ _)) =>
- let _stringappend_1034_ :=
+ (op,(existT _ _stringappend_702_ _)) =>
+ let _stringappend_703_ :=
string_drop
- _stringappend_559_
- (build_ex
- _stringappend_1033_) in
+ _stringappend_381_
+ (build_ex _stringappend_702_) in
and_boolM
(returnm ((string_startswith
- _stringappend_1034_
+ _stringappend_703_
".")
: bool))
- (let _stringappend_1035_ :=
+ (let _stringappend_704_ :=
string_drop
- _stringappend_1034_
- (string_length
- ".") in
+ _stringappend_703_
+ (build_ex (projT1 (string_length
+ "."))) in
match (size_mnemonic_matches_prefix
- _stringappend_1035_) with
+ _stringappend_704_) with
| Some
- (_stringappend_1036_,(existT _ _stringappend_1037_ _)) =>
- let _stringappend_1038_ :=
+ (width,(existT _ _stringappend_705_ _)) =>
+ let _stringappend_706_ :=
string_drop
- _stringappend_1035_
- (build_ex
- _stringappend_1037_) in
+ _stringappend_704_
+ (build_ex _stringappend_705_) in
match (maybe_aq_matches_prefix
- _stringappend_1038_) with
+ _stringappend_706_) with
| Some
- (_stringappend_1039_,(existT _ _stringappend_1040_ _)) =>
- let _stringappend_1041_ :=
+ (aq,(existT _ _stringappend_707_ _)) =>
+ let _stringappend_708_ :=
string_drop
- _stringappend_1038_
- (build_ex
- _stringappend_1040_) in
+ _stringappend_706_
+ (build_ex _stringappend_707_) in
match (maybe_rl_matches_prefix
- _stringappend_1041_) with
+ _stringappend_708_) with
| Some
- (_stringappend_1042_,(existT _ _stringappend_1043_ _)) =>
- let _stringappend_1044_ :=
+ (rl,(existT _ _stringappend_709_ _)) =>
+ let _stringappend_710_ :=
string_drop
- _stringappend_1041_
- (build_ex
- _stringappend_1043_) in
+ _stringappend_708_
+ (build_ex _stringappend_709_) in
match (spc_matches_prefix
- _stringappend_1044_) with
+ _stringappend_710_) with
| Some
- (_stringappend_1045_,(existT _ _stringappend_1046_ _)) =>
- let _stringappend_1047_ :=
+ (tt,(existT _ _stringappend_711_ _)) =>
+ let _stringappend_712_ :=
string_drop
- _stringappend_1044_
- (build_ex
- _stringappend_1046_) in
+ _stringappend_710_
+ (build_ex _stringappend_711_) in
match (reg_name_matches_prefix
- _stringappend_1047_) with
+ _stringappend_712_) with
| Some
- (_stringappend_1048_,(existT _ _stringappend_1049_ _)) =>
- let _stringappend_1050_ :=
+ (rd,(existT _ _stringappend_713_ _)) =>
+ let _stringappend_714_ :=
string_drop
- _stringappend_1047_
- (build_ex
- _stringappend_1049_) in
+ _stringappend_712_
+ (build_ex _stringappend_713_) in
sep_matches_prefix
- _stringappend_1050_ >>= fun w__553 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_714_ >>= fun w__553 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__553 with
| Some
- (_stringappend_1051_,(existT _ _stringappend_1052_ _)) =>
- let _stringappend_1053_ :=
+ (tt,(existT _ _stringappend_715_ _)) =>
+ let _stringappend_716_ :=
string_drop
- _stringappend_1050_
- (build_ex
- _stringappend_1052_) in
+ _stringappend_714_
+ (build_ex _stringappend_715_) in
match (reg_name_matches_prefix
- _stringappend_1053_) with
+ _stringappend_716_) with
| Some
- (_stringappend_1054_,(existT _ _stringappend_1055_ _)) =>
- let _stringappend_1056_ :=
+ (rs1,(existT _ _stringappend_717_ _)) =>
+ let _stringappend_718_ :=
string_drop
- _stringappend_1053_
- (build_ex
- _stringappend_1055_) in
+ _stringappend_716_
+ (build_ex _stringappend_717_) in
sep_matches_prefix
- _stringappend_1056_ >>= fun w__554 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_718_ >>= fun w__554 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if
((match w__554 with
| Some
- (_stringappend_1057_,(existT _ _stringappend_1058_ _)) =>
- let _stringappend_1059_ :=
+ (tt,(existT _ _stringappend_719_ _)) =>
+ let _stringappend_720_ :=
string_drop
- _stringappend_1056_
- (build_ex
- _stringappend_1058_) in
+ _stringappend_718_
+ (build_ex _stringappend_719_) in
if
((match (reg_name_matches_prefix
- _stringappend_1059_) with
+ _stringappend_720_) with
| Some
- (_stringappend_1060_,(existT _ _stringappend_1061_ _)) =>
+ (rs2,(existT _ _stringappend_721_ _)) =>
match (string_drop
- _stringappend_1059_
- (build_ex
- _stringappend_1061_)) with
+ _stringappend_720_
+ (build_ex _stringappend_721_)) with
| "" =>
true
| _ =>
@@ -22568,254 +21363,164 @@ Definition assembly_backwards_matches (arg_ : string)
: bool)
end >>= fun w__563 : bool =>
(if (w__563) then
- match (amo_mnemonic_matches_prefix
- _stringappend_559_) with
- | Some
- (_stringappend_1032_,(existT _ _stringappend_1033_ _)) =>
- returnm ((_stringappend_1032_,
- build_ex
- _stringappend_1033_)
- : (amoop * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((amoop * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__565 : (amoop * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(op, existT _ _stringappend_1033_ _) :=
- w__565
- : (amoop * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1034_ :=
+ (match (amo_mnemonic_matches_prefix
+ _stringappend_381_) with
+ | Some
+ (op,(existT _ _stringappend_702_ _)) =>
+ returnm (op, build_ex _stringappend_702_)
+ | _ =>
+ exit tt
+ : M ((amoop * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((amoop * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(op, existT _ _stringappend_702_ _) =>
+ let _stringappend_703_ :=
string_drop
- _stringappend_559_
- (build_ex
- _stringappend_1033_) in
- let _stringappend_1035_ :=
+ _stringappend_381_
+ (build_ex _stringappend_702_) in
+ let _stringappend_704_ :=
string_drop
- _stringappend_1034_
- (string_length
- ".") in
- match (size_mnemonic_matches_prefix
- _stringappend_1035_) with
- | Some
- (_stringappend_1036_,(existT _ _stringappend_1037_ _)) =>
- returnm ((_stringappend_1036_,
- build_ex
- _stringappend_1037_)
- : (word_width * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((word_width * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__567 : (word_width * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(width, existT _ _stringappend_1037_ _) :=
- w__567
- : (word_width * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1038_ :=
+ _stringappend_703_
+ (build_ex (projT1 (string_length
+ "."))) in
+ (match (size_mnemonic_matches_prefix
+ _stringappend_704_) with
+ | Some
+ (width,(existT _ _stringappend_705_ _)) =>
+ returnm (width, build_ex _stringappend_705_)
+ | _ =>
+ exit tt
+ : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(width, existT _ _stringappend_705_ _) =>
+ let _stringappend_706_ :=
string_drop
- _stringappend_1035_
- (build_ex
- _stringappend_1037_) in
- match (maybe_aq_matches_prefix
- _stringappend_1038_) with
- | Some
- (_stringappend_1039_,(existT _ _stringappend_1040_ _)) =>
- returnm ((_stringappend_1039_,
- build_ex
- _stringappend_1040_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__569 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(aq, existT _ _stringappend_1040_ _) :=
- w__569
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1041_ :=
+ _stringappend_704_
+ (build_ex _stringappend_705_) in
+ (match (maybe_aq_matches_prefix
+ _stringappend_706_) with
+ | Some
+ (aq,(existT _ _stringappend_707_ _)) =>
+ returnm (aq, build_ex _stringappend_707_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(aq, existT _ _stringappend_707_ _) =>
+ let _stringappend_708_ :=
string_drop
- _stringappend_1038_
- (build_ex
- _stringappend_1040_) in
- match (maybe_rl_matches_prefix
- _stringappend_1041_) with
- | Some
- (_stringappend_1042_,(existT _ _stringappend_1043_ _)) =>
- returnm ((_stringappend_1042_,
- build_ex
- _stringappend_1043_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__571 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rl, existT _ _stringappend_1043_ _) :=
- w__571
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1044_ :=
+ _stringappend_706_
+ (build_ex _stringappend_707_) in
+ (match (maybe_rl_matches_prefix
+ _stringappend_708_) with
+ | Some
+ (rl,(existT _ _stringappend_709_ _)) =>
+ returnm (rl, build_ex _stringappend_709_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rl, existT _ _stringappend_709_ _) =>
+ let _stringappend_710_ :=
string_drop
- _stringappend_1041_
- (build_ex
- _stringappend_1043_) in
- match (spc_matches_prefix
- _stringappend_1044_) with
- | Some
- (_stringappend_1045_,(existT _ _stringappend_1046_ _)) =>
- returnm ((_stringappend_1045_,
- build_ex
- _stringappend_1046_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__573 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1046_ _) :=
- w__573
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1047_ :=
+ _stringappend_708_
+ (build_ex _stringappend_709_) in
+ (match (spc_matches_prefix
+ _stringappend_710_) with
+ | Some
+ (tt,(existT _ _stringappend_711_ _)) =>
+ returnm (tt, build_ex _stringappend_711_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_711_ _) =>
+ let _stringappend_712_ :=
string_drop
- _stringappend_1044_
- (build_ex
- _stringappend_1046_) in
- match (reg_name_matches_prefix
- _stringappend_1047_) with
- | Some
- (_stringappend_1048_,(existT _ _stringappend_1049_ _)) =>
- returnm ((_stringappend_1048_,
- build_ex
- _stringappend_1049_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__575 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1049_ _) :=
- w__575
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1050_ :=
+ _stringappend_710_
+ (build_ex _stringappend_711_) in
+ (match (reg_name_matches_prefix
+ _stringappend_712_) with
+ | Some
+ (rd,(existT _ _stringappend_713_ _)) =>
+ returnm (rd, build_ex _stringappend_713_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_713_ _) =>
+ let _stringappend_714_ :=
string_drop
- _stringappend_1047_
- (build_ex
- _stringappend_1049_) in
+ _stringappend_712_
+ (build_ex _stringappend_713_) in
sep_matches_prefix
- _stringappend_1050_ >>= fun w__576 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_714_ >>= fun w__576 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__576 with
- | Some
- (_stringappend_1051_,(existT _ _stringappend_1052_ _)) =>
- returnm ((_stringappend_1051_,
- build_ex
- _stringappend_1052_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__578 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1052_ _) :=
- w__578
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1053_ :=
+ (match w__576 with
+ | Some
+ (tt,(existT _ _stringappend_715_ _)) =>
+ returnm (tt, build_ex _stringappend_715_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_715_ _) =>
+ let _stringappend_716_ :=
string_drop
- _stringappend_1050_
- (build_ex
- _stringappend_1052_) in
- match (reg_name_matches_prefix
- _stringappend_1053_) with
- | Some
- (_stringappend_1054_,(existT _ _stringappend_1055_ _)) =>
- returnm ((_stringappend_1054_,
- build_ex
- _stringappend_1055_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__580 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1055_ _) :=
- w__580
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1056_ :=
+ _stringappend_714_
+ (build_ex _stringappend_715_) in
+ (match (reg_name_matches_prefix
+ _stringappend_716_) with
+ | Some
+ (rs1,(existT _ _stringappend_717_ _)) =>
+ returnm (rs1, build_ex _stringappend_717_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_717_ _) =>
+ let _stringappend_718_ :=
string_drop
- _stringappend_1053_
- (build_ex
- _stringappend_1055_) in
+ _stringappend_716_
+ (build_ex _stringappend_717_) in
sep_matches_prefix
- _stringappend_1056_ >>= fun w__581 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_718_ >>= fun w__581 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__581 with
- | Some
- (_stringappend_1057_,(existT _ _stringappend_1058_ _)) =>
- returnm ((_stringappend_1057_,
- build_ex
- _stringappend_1058_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__583 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1058_ _) :=
- w__583
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1059_ :=
+ (match w__581 with
+ | Some
+ (tt,(existT _ _stringappend_719_ _)) =>
+ returnm (tt, build_ex _stringappend_719_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_719_ _) =>
+ let _stringappend_720_ :=
string_drop
- _stringappend_1056_
- (build_ex
- _stringappend_1058_) in
- match (reg_name_matches_prefix
- _stringappend_1059_) with
- | Some
- (_stringappend_1060_,(existT _ _stringappend_1061_ _)) =>
- returnm ((_stringappend_1060_,
- build_ex
- _stringappend_1061_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__585 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_1061_ _) :=
- w__585
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_718_
+ (build_ex _stringappend_719_) in
+ (match (reg_name_matches_prefix
+ _stringappend_720_) with
+ | Some
+ (rs2,(existT _ _stringappend_721_ _)) =>
+ returnm (rs2, build_ex _stringappend_721_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_721_ _) =>
(match (string_drop
- _stringappend_1059_
- (build_ex
- _stringappend_1061_)) with
+ _stringappend_720_
+ (build_ex _stringappend_721_)) with
| "" =>
returnm (true
: bool)
@@ -22826,83 +21531,76 @@ Definition assembly_backwards_matches (arg_ : string)
: M (bool)
else
match (csr_mnemonic_matches_prefix
- _stringappend_559_) with
+ _stringappend_381_) with
| Some
- (_stringappend_1063_,(existT _ _stringappend_1064_ _)) =>
- let _stringappend_1065_ :=
+ (op,(existT _ _stringappend_723_ _)) =>
+ let _stringappend_724_ :=
string_drop
- _stringappend_559_
- (build_ex
- _stringappend_1064_) in
+ _stringappend_381_
+ (build_ex _stringappend_723_) in
and_boolM
(returnm ((string_startswith
- _stringappend_1065_
+ _stringappend_724_
"i")
: bool))
- (let _stringappend_1066_ :=
+ (let _stringappend_725_ :=
string_drop
- _stringappend_1065_
- (string_length
- "i") in
+ _stringappend_724_
+ (build_ex (projT1 (string_length
+ "i"))) in
match (spc_matches_prefix
- _stringappend_1066_) with
+ _stringappend_725_) with
| Some
- (_stringappend_1067_,(existT _ _stringappend_1068_ _)) =>
- let _stringappend_1069_ :=
+ (tt,(existT _ _stringappend_726_ _)) =>
+ let _stringappend_727_ :=
string_drop
- _stringappend_1066_
- (build_ex
- _stringappend_1068_) in
+ _stringappend_725_
+ (build_ex _stringappend_726_) in
match (reg_name_matches_prefix
- _stringappend_1069_) with
+ _stringappend_727_) with
| Some
- (_stringappend_1070_,(existT _ _stringappend_1071_ _)) =>
- let _stringappend_1072_ :=
+ (rd,(existT _ _stringappend_728_ _)) =>
+ let _stringappend_729_ :=
string_drop
- _stringappend_1069_
- (build_ex
- _stringappend_1071_) in
+ _stringappend_727_
+ (build_ex _stringappend_728_) in
sep_matches_prefix
- _stringappend_1072_ >>= fun w__588 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_729_ >>= fun w__588 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__588 with
| Some
- (_stringappend_1073_,(existT _ _stringappend_1074_ _)) =>
- let _stringappend_1075_ :=
+ (tt,(existT _ _stringappend_730_ _)) =>
+ let _stringappend_731_ :=
string_drop
- _stringappend_1072_
- (build_ex
- _stringappend_1074_) in
+ _stringappend_729_
+ (build_ex _stringappend_730_) in
match (hex_bits_5_matches_prefix
- _stringappend_1075_) with
+ _stringappend_731_) with
| Some
- (_stringappend_1076_,(existT _ _stringappend_1077_ _)) =>
- let _stringappend_1078_ :=
+ (rs1,(existT _ _stringappend_732_ _)) =>
+ let _stringappend_733_ :=
string_drop
- _stringappend_1075_
- (build_ex
- _stringappend_1077_) in
+ _stringappend_731_
+ (build_ex _stringappend_732_) in
sep_matches_prefix
- _stringappend_1078_ >>= fun w__589 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_733_ >>= fun w__589 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if
((match w__589 with
| Some
- (_stringappend_1079_,(existT _ _stringappend_1080_ _)) =>
- let _stringappend_1081_ :=
+ (tt,(existT _ _stringappend_734_ _)) =>
+ let _stringappend_735_ :=
string_drop
- _stringappend_1078_
- (build_ex
- _stringappend_1080_) in
+ _stringappend_733_
+ (build_ex _stringappend_734_) in
if
((match (csr_name_map_matches_prefix
- _stringappend_1081_) with
+ _stringappend_735_) with
| Some
- (_stringappend_1082_,(existT _ _stringappend_1083_ _)) =>
+ (csr,(existT _ _stringappend_736_ _)) =>
match (string_drop
- _stringappend_1081_
- (build_ex
- _stringappend_1083_)) with
+ _stringappend_735_
+ (build_ex _stringappend_736_)) with
| "" =>
true
| _ =>
@@ -22976,182 +21674,119 @@ Definition assembly_backwards_matches (arg_ : string)
: bool)
end >>= fun w__595 : bool =>
(if (w__595) then
- match (csr_mnemonic_matches_prefix
- _stringappend_559_) with
- | Some
- (_stringappend_1063_,(existT _ _stringappend_1064_ _)) =>
- returnm ((_stringappend_1063_,
- build_ex
- _stringappend_1064_)
- : (csrop * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((csrop * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__597 : (csrop * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(op, existT _ _stringappend_1064_ _) :=
- w__597
- : (csrop * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1065_ :=
+ (match (csr_mnemonic_matches_prefix
+ _stringappend_381_) with
+ | Some
+ (op,(existT _ _stringappend_723_ _)) =>
+ returnm (op, build_ex _stringappend_723_)
+ | _ =>
+ exit tt
+ : M ((csrop * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((csrop * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(op, existT _ _stringappend_723_ _) =>
+ let _stringappend_724_ :=
string_drop
- _stringappend_559_
- (build_ex
- _stringappend_1064_) in
- let _stringappend_1066_ :=
+ _stringappend_381_
+ (build_ex _stringappend_723_) in
+ let _stringappend_725_ :=
string_drop
- _stringappend_1065_
- (string_length
- "i") in
- match (spc_matches_prefix
- _stringappend_1066_) with
- | Some
- (_stringappend_1067_,(existT _ _stringappend_1068_ _)) =>
- returnm ((_stringappend_1067_,
- build_ex
- _stringappend_1068_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__599 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1068_ _) :=
- w__599
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1069_ :=
+ _stringappend_724_
+ (build_ex (projT1 (string_length
+ "i"))) in
+ (match (spc_matches_prefix
+ _stringappend_725_) with
+ | Some
+ (tt,(existT _ _stringappend_726_ _)) =>
+ returnm (tt, build_ex _stringappend_726_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_726_ _) =>
+ let _stringappend_727_ :=
string_drop
- _stringappend_1066_
- (build_ex
- _stringappend_1068_) in
- match (reg_name_matches_prefix
- _stringappend_1069_) with
- | Some
- (_stringappend_1070_,(existT _ _stringappend_1071_ _)) =>
- returnm ((_stringappend_1070_,
- build_ex
- _stringappend_1071_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__601 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1071_ _) :=
- w__601
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1072_ :=
+ _stringappend_725_
+ (build_ex _stringappend_726_) in
+ (match (reg_name_matches_prefix
+ _stringappend_727_) with
+ | Some
+ (rd,(existT _ _stringappend_728_ _)) =>
+ returnm (rd, build_ex _stringappend_728_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_728_ _) =>
+ let _stringappend_729_ :=
string_drop
- _stringappend_1069_
- (build_ex
- _stringappend_1071_) in
+ _stringappend_727_
+ (build_ex _stringappend_728_) in
sep_matches_prefix
- _stringappend_1072_ >>= fun w__602 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_729_ >>= fun w__602 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__602 with
- | Some
- (_stringappend_1073_,(existT _ _stringappend_1074_ _)) =>
- returnm ((_stringappend_1073_,
- build_ex
- _stringappend_1074_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__604 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1074_ _) :=
- w__604
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1075_ :=
+ (match w__602 with
+ | Some
+ (tt,(existT _ _stringappend_730_ _)) =>
+ returnm (tt, build_ex _stringappend_730_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_730_ _) =>
+ let _stringappend_731_ :=
string_drop
- _stringappend_1072_
- (build_ex
- _stringappend_1074_) in
- match (hex_bits_5_matches_prefix
- _stringappend_1075_) with
- | Some
- (_stringappend_1076_,(existT _ _stringappend_1077_ _)) =>
- returnm ((_stringappend_1076_,
- build_ex
- _stringappend_1077_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__606 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1077_ _) :=
- w__606
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1078_ :=
+ _stringappend_729_
+ (build_ex _stringappend_730_) in
+ (match (hex_bits_5_matches_prefix
+ _stringappend_731_) with
+ | Some
+ (rs1,(existT _ _stringappend_732_ _)) =>
+ returnm (rs1, build_ex _stringappend_732_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_732_ _) =>
+ let _stringappend_733_ :=
string_drop
- _stringappend_1075_
- (build_ex
- _stringappend_1077_) in
+ _stringappend_731_
+ (build_ex _stringappend_732_) in
sep_matches_prefix
- _stringappend_1078_ >>= fun w__607 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_733_ >>= fun w__607 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__607 with
- | Some
- (_stringappend_1079_,(existT _ _stringappend_1080_ _)) =>
- returnm ((_stringappend_1079_,
- build_ex
- _stringappend_1080_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__609 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1080_ _) :=
- w__609
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1081_ :=
+ (match w__607 with
+ | Some
+ (tt,(existT _ _stringappend_734_ _)) =>
+ returnm (tt, build_ex _stringappend_734_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_734_ _) =>
+ let _stringappend_735_ :=
string_drop
- _stringappend_1078_
- (build_ex
- _stringappend_1080_) in
- match (csr_name_map_matches_prefix
- _stringappend_1081_) with
- | Some
- (_stringappend_1082_,(existT _ _stringappend_1083_ _)) =>
- returnm ((_stringappend_1082_,
- build_ex
- _stringappend_1083_)
- : (mword 12 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 12 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__611 : (mword 12 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(csr, existT _ _stringappend_1083_ _) :=
- w__611
- : (mword 12 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_733_
+ (build_ex _stringappend_734_) in
+ (match (csr_name_map_matches_prefix
+ _stringappend_735_) with
+ | Some
+ (csr,(existT _ _stringappend_736_ _)) =>
+ returnm (csr, build_ex _stringappend_736_)
+ | _ =>
+ exit tt
+ : M ((mword 12 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(csr, existT _ _stringappend_736_ _) =>
(match (string_drop
- _stringappend_1081_
- (build_ex
- _stringappend_1083_)) with
+ _stringappend_735_
+ (build_ex _stringappend_736_)) with
| "" =>
returnm (true
: bool)
@@ -23162,73 +21797,66 @@ Definition assembly_backwards_matches (arg_ : string)
: M (bool)
else
match (csr_mnemonic_matches_prefix
- _stringappend_559_) with
+ _stringappend_381_) with
| Some
- (_stringappend_1085_,(existT _ _stringappend_1086_ _)) =>
- let _stringappend_1087_ :=
+ (op,(existT _ _stringappend_738_ _)) =>
+ let _stringappend_739_ :=
string_drop
- _stringappend_559_
- (build_ex
- _stringappend_1086_) in
+ _stringappend_381_
+ (build_ex _stringappend_738_) in
match (spc_matches_prefix
- _stringappend_1087_) with
+ _stringappend_739_) with
| Some
- (_stringappend_1088_,(existT _ _stringappend_1089_ _)) =>
- let _stringappend_1090_ :=
+ (tt,(existT _ _stringappend_740_ _)) =>
+ let _stringappend_741_ :=
string_drop
- _stringappend_1087_
- (build_ex
- _stringappend_1089_) in
+ _stringappend_739_
+ (build_ex _stringappend_740_) in
match (reg_name_matches_prefix
- _stringappend_1090_) with
+ _stringappend_741_) with
| Some
- (_stringappend_1091_,(existT _ _stringappend_1092_ _)) =>
- let _stringappend_1093_ :=
+ (rd,(existT _ _stringappend_742_ _)) =>
+ let _stringappend_743_ :=
string_drop
- _stringappend_1090_
- (build_ex
- _stringappend_1092_) in
+ _stringappend_741_
+ (build_ex _stringappend_742_) in
sep_matches_prefix
- _stringappend_1093_ >>= fun w__614 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_743_ >>= fun w__614 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__614 with
| Some
- (_stringappend_1094_,(existT _ _stringappend_1095_ _)) =>
- let _stringappend_1096_ :=
+ (tt,(existT _ _stringappend_744_ _)) =>
+ let _stringappend_745_ :=
string_drop
- _stringappend_1093_
- (build_ex
- _stringappend_1095_) in
+ _stringappend_743_
+ (build_ex _stringappend_744_) in
match (reg_name_matches_prefix
- _stringappend_1096_) with
+ _stringappend_745_) with
| Some
- (_stringappend_1097_,(existT _ _stringappend_1098_ _)) =>
- let _stringappend_1099_ :=
+ (rs1,(existT _ _stringappend_746_ _)) =>
+ let _stringappend_747_ :=
string_drop
- _stringappend_1096_
- (build_ex
- _stringappend_1098_) in
+ _stringappend_745_
+ (build_ex _stringappend_746_) in
sep_matches_prefix
- _stringappend_1099_ >>= fun w__615 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_747_ >>= fun w__615 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if
((match w__615 with
| Some
- (_stringappend_1100_,(existT _ _stringappend_1101_ _)) =>
- let _stringappend_1102_ :=
+ (tt,(existT _ _stringappend_748_ _)) =>
+ let _stringappend_749_ :=
string_drop
- _stringappend_1099_
- (build_ex
- _stringappend_1101_) in
+ _stringappend_747_
+ (build_ex _stringappend_748_) in
if
((match (csr_name_map_matches_prefix
- _stringappend_1102_) with
+ _stringappend_749_) with
| Some
- (_stringappend_1103_,(existT _ _stringappend_1104_ _)) =>
+ (csr,(existT _ _stringappend_750_ _)) =>
match (string_drop
- _stringappend_1102_
- (build_ex
- _stringappend_1104_)) with
+ _stringappend_749_
+ (build_ex _stringappend_750_)) with
| "" =>
true
| _ =>
@@ -23297,177 +21925,114 @@ Definition assembly_backwards_matches (arg_ : string)
: bool)
end >>= fun w__620 : bool =>
(if (w__620) then
- match (csr_mnemonic_matches_prefix
- _stringappend_559_) with
- | Some
- (_stringappend_1085_,(existT _ _stringappend_1086_ _)) =>
- returnm ((_stringappend_1085_,
- build_ex
- _stringappend_1086_)
- : (csrop * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((csrop * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__622 : (csrop * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(op, existT _ _stringappend_1086_ _) :=
- w__622
- : (csrop * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1087_ :=
+ (match (csr_mnemonic_matches_prefix
+ _stringappend_381_) with
+ | Some
+ (op,(existT _ _stringappend_738_ _)) =>
+ returnm (op, build_ex _stringappend_738_)
+ | _ =>
+ exit tt
+ : M ((csrop * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((csrop * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(op, existT _ _stringappend_738_ _) =>
+ let _stringappend_739_ :=
string_drop
- _stringappend_559_
- (build_ex
- _stringappend_1086_) in
- match (spc_matches_prefix
- _stringappend_1087_) with
- | Some
- (_stringappend_1088_,(existT _ _stringappend_1089_ _)) =>
- returnm ((_stringappend_1088_,
- build_ex
- _stringappend_1089_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__624 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1089_ _) :=
- w__624
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1090_ :=
+ _stringappend_381_
+ (build_ex _stringappend_738_) in
+ (match (spc_matches_prefix
+ _stringappend_739_) with
+ | Some
+ (tt,(existT _ _stringappend_740_ _)) =>
+ returnm (tt, build_ex _stringappend_740_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_740_ _) =>
+ let _stringappend_741_ :=
string_drop
- _stringappend_1087_
- (build_ex
- _stringappend_1089_) in
- match (reg_name_matches_prefix
- _stringappend_1090_) with
- | Some
- (_stringappend_1091_,(existT _ _stringappend_1092_ _)) =>
- returnm ((_stringappend_1091_,
- build_ex
- _stringappend_1092_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__626 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_1092_ _) :=
- w__626
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1093_ :=
+ _stringappend_739_
+ (build_ex _stringappend_740_) in
+ (match (reg_name_matches_prefix
+ _stringappend_741_) with
+ | Some
+ (rd,(existT _ _stringappend_742_ _)) =>
+ returnm (rd, build_ex _stringappend_742_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_742_ _) =>
+ let _stringappend_743_ :=
string_drop
- _stringappend_1090_
- (build_ex
- _stringappend_1092_) in
+ _stringappend_741_
+ (build_ex _stringappend_742_) in
sep_matches_prefix
- _stringappend_1093_ >>= fun w__627 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_743_ >>= fun w__627 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__627 with
- | Some
- (_stringappend_1094_,(existT _ _stringappend_1095_ _)) =>
- returnm ((_stringappend_1094_,
- build_ex
- _stringappend_1095_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__629 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1095_ _) :=
- w__629
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1096_ :=
+ (match w__627 with
+ | Some
+ (tt,(existT _ _stringappend_744_ _)) =>
+ returnm (tt, build_ex _stringappend_744_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_744_ _) =>
+ let _stringappend_745_ :=
string_drop
- _stringappend_1093_
- (build_ex
- _stringappend_1095_) in
- match (reg_name_matches_prefix
- _stringappend_1096_) with
- | Some
- (_stringappend_1097_,(existT _ _stringappend_1098_ _)) =>
- returnm ((_stringappend_1097_,
- build_ex
- _stringappend_1098_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__631 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_1098_ _) :=
- w__631
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1099_ :=
+ _stringappend_743_
+ (build_ex _stringappend_744_) in
+ (match (reg_name_matches_prefix
+ _stringappend_745_) with
+ | Some
+ (rs1,(existT _ _stringappend_746_ _)) =>
+ returnm (rs1, build_ex _stringappend_746_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_746_ _) =>
+ let _stringappend_747_ :=
string_drop
- _stringappend_1096_
- (build_ex
- _stringappend_1098_) in
+ _stringappend_745_
+ (build_ex _stringappend_746_) in
sep_matches_prefix
- _stringappend_1099_ >>= fun w__632 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_747_ >>= fun w__632 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__632 with
- | Some
- (_stringappend_1100_,(existT _ _stringappend_1101_ _)) =>
- returnm ((_stringappend_1100_,
- build_ex
- _stringappend_1101_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__634 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1101_ _) :=
- w__634
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1102_ :=
+ (match w__632 with
+ | Some
+ (tt,(existT _ _stringappend_748_ _)) =>
+ returnm (tt, build_ex _stringappend_748_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_748_ _) =>
+ let _stringappend_749_ :=
string_drop
- _stringappend_1099_
- (build_ex
- _stringappend_1101_) in
- match (csr_name_map_matches_prefix
- _stringappend_1102_) with
- | Some
- (_stringappend_1103_,(existT _ _stringappend_1104_ _)) =>
- returnm ((_stringappend_1103_,
- build_ex
- _stringappend_1104_)
- : (mword 12 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 12 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__636 : (mword 12 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(csr, existT _ _stringappend_1104_ _) :=
- w__636
- : (mword 12 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_747_
+ (build_ex _stringappend_748_) in
+ (match (csr_name_map_matches_prefix
+ _stringappend_749_) with
+ | Some
+ (csr,(existT _ _stringappend_750_ _)) =>
+ returnm (csr, build_ex _stringappend_750_)
+ | _ =>
+ exit tt
+ : M ((mword 12 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(csr, existT _ _stringappend_750_ _) =>
(match (string_drop
- _stringappend_1102_
- (build_ex
- _stringappend_1104_)) with
+ _stringappend_749_
+ (build_ex _stringappend_750_)) with
| "" =>
returnm (true
: bool)
@@ -23478,31 +22043,29 @@ Definition assembly_backwards_matches (arg_ : string)
: M (bool)
else if ((andb
(string_startswith
- _stringappend_559_
+ _stringappend_381_
"illegal")
- (let _stringappend_1106_ :=
+ (let _stringappend_752_ :=
string_drop
- _stringappend_559_
- (string_length
- "illegal") in
+ _stringappend_381_
+ (build_ex (projT1 (string_length
+ "illegal"))) in
if ((match (spc_matches_prefix
- _stringappend_1106_) with
+ _stringappend_752_) with
| Some
- (_stringappend_1107_,(existT _ _stringappend_1108_ _)) =>
- let _stringappend_1109_ :=
+ (tt,(existT _ _stringappend_753_ _)) =>
+ let _stringappend_754_ :=
string_drop
- _stringappend_1106_
- (build_ex
- _stringappend_1108_) in
+ _stringappend_752_
+ (build_ex _stringappend_753_) in
if
((match (hex_bits_32_matches_prefix
- _stringappend_1109_) with
+ _stringappend_754_) with
| Some
- (_stringappend_1110_,(existT _ _stringappend_1111_ _)) =>
+ (s,(existT _ _stringappend_755_ _)) =>
match (string_drop
- _stringappend_1109_
- (build_ex
- _stringappend_1111_)) with
+ _stringappend_754_
+ (build_ex _stringappend_755_)) with
| "" =>
true
| _ =>
@@ -23523,58 +22086,40 @@ Definition assembly_backwards_matches (arg_ : string)
else
false)))
then
- let _stringappend_1106_ :=
+ let _stringappend_752_ :=
string_drop
- _stringappend_559_
- (string_length
- "illegal") in
- match (spc_matches_prefix
- _stringappend_1106_) with
- | Some
- (_stringappend_1107_,(existT _ _stringappend_1108_ _)) =>
- returnm ((_stringappend_1107_,
- build_ex
- _stringappend_1108_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__640 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_1108_ _) :=
- w__640
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_1109_ :=
+ _stringappend_381_
+ (build_ex (projT1 (string_length
+ "illegal"))) in
+ (match (spc_matches_prefix
+ _stringappend_752_) with
+ | Some
+ (tt,(existT _ _stringappend_753_ _)) =>
+ returnm (tt, build_ex _stringappend_753_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_753_ _) =>
+ let _stringappend_754_ :=
string_drop
- _stringappend_1106_
- (build_ex
- _stringappend_1108_) in
- match (hex_bits_32_matches_prefix
- _stringappend_1109_) with
- | Some
- (_stringappend_1110_,(existT _ _stringappend_1111_ _)) =>
- returnm ((_stringappend_1110_,
- build_ex
- _stringappend_1111_)
- : (mword 32 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 32 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__642 : (mword 32 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_1111_ _) :=
- w__642
- : (mword 32 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_752_
+ (build_ex _stringappend_753_) in
+ (match (hex_bits_32_matches_prefix
+ _stringappend_754_) with
+ | Some
+ (s,(existT _ _stringappend_755_ _)) =>
+ returnm (s, build_ex _stringappend_755_)
+ | _ =>
+ exit tt
+ : M ((mword 32 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 32 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(s, existT _ _stringappend_755_ _) =>
(match (string_drop
- _stringappend_1109_
- (build_ex
- _stringappend_1111_)) with
+ _stringappend_754_
+ (build_ex _stringappend_755_)) with
| "" =>
returnm (true
: bool)
@@ -23619,25 +22164,24 @@ Definition assembly_matches_prefix (arg_ : string)
: M (option ((ast * {n : Z & ArithFact (n >= 0)}))) :=
let _stringappend_0_ := arg_ in
match (utype_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_1_,(existT _ _stringappend_2_ _)) =>
- let _stringappend_3_ := string_drop _stringappend_0_ (build_ex _stringappend_2_) in
- match (spc_matches_prefix _stringappend_3_) with
- | Some (_stringappend_4_,(existT _ _stringappend_5_ _)) =>
- let _stringappend_6_ := string_drop _stringappend_3_ (build_ex _stringappend_5_) in
- match (reg_name_matches_prefix _stringappend_6_) with
- | Some (_stringappend_7_,(existT _ _stringappend_8_ _)) =>
- let _stringappend_9_ := string_drop _stringappend_6_ (build_ex _stringappend_8_) in
- sep_matches_prefix _stringappend_9_ >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op,(existT _ _stringappend_1_ _)) =>
+ let _stringappend_2_ := string_drop _stringappend_0_ (build_ex _stringappend_1_) in
+ match (spc_matches_prefix _stringappend_2_) with
+ | Some (tt,(existT _ _stringappend_3_ _)) =>
+ let _stringappend_4_ := string_drop _stringappend_2_ (build_ex _stringappend_3_) in
+ match (reg_name_matches_prefix _stringappend_4_) with
+ | Some (rd,(existT _ _stringappend_5_ _)) =>
+ let _stringappend_6_ := string_drop _stringappend_4_ (build_ex _stringappend_5_) in
+ sep_matches_prefix _stringappend_6_ >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__0 with
- | Some (_stringappend_10_,(existT _ _stringappend_11_ _)) =>
- let _stringappend_12_ :=
- string_drop _stringappend_9_ (build_ex _stringappend_11_) in
- if ((match (hex_bits_20_matches_prefix _stringappend_12_) with
- | Some (_stringappend_13_,(existT _ _stringappend_14_ _)) =>
- match (string_drop _stringappend_12_
- (build_ex
- _stringappend_14_)) with
+ | Some (tt,(existT _ _stringappend_7_ _)) =>
+ let _stringappend_8_ :=
+ string_drop _stringappend_6_ (build_ex _stringappend_7_) in
+ if ((match (hex_bits_20_matches_prefix _stringappend_8_) with
+ | Some (imm,(existT _ _stringappend_9_ _)) =>
+ match (string_drop _stringappend_8_
+ (build_ex _stringappend_9_)) with
| s_ => true
end
| None => false
@@ -23662,70 +22206,60 @@ Definition assembly_matches_prefix (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__3 : bool =>
(if (w__3) then
- match (utype_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_1_,(existT _ _stringappend_2_ _)) =>
- returnm ((_stringappend_1_, build_ex _stringappend_2_)
- : (uop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((uop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__5 : (uop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_2_ _) := w__5 : (uop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_3_ := string_drop _stringappend_0_ (build_ex _stringappend_2_) in
- match (spc_matches_prefix _stringappend_3_) with
- | Some (_stringappend_4_,(existT _ _stringappend_5_ _)) =>
- returnm ((_stringappend_4_, build_ex _stringappend_5_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__7 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_5_ _) := w__7 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_6_ := string_drop _stringappend_3_ (build_ex _stringappend_5_) in
- match (reg_name_matches_prefix _stringappend_6_) with
- | Some (_stringappend_7_,(existT _ _stringappend_8_ _)) =>
- returnm ((_stringappend_7_, build_ex _stringappend_8_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__9 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_8_ _) := w__9 : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_9_ := string_drop _stringappend_6_ (build_ex _stringappend_8_) in
- sep_matches_prefix _stringappend_9_ >>= fun w__10 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- match w__10 with
- | Some (_stringappend_10_,(existT _ _stringappend_11_ _)) =>
- returnm ((_stringappend_10_, build_ex _stringappend_11_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__12 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_11_ _) := w__12 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_12_ := string_drop _stringappend_9_ (build_ex _stringappend_11_) in
- match (hex_bits_20_matches_prefix _stringappend_12_) with
- | Some (_stringappend_13_,(existT _ _stringappend_14_ _)) =>
- returnm ((_stringappend_13_, build_ex _stringappend_14_)
- : (mword 20 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 20 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__14 : (mword 20 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_14_ _) := w__14 : (mword 20 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_12_ (build_ex _stringappend_14_)) with
- | s_ => Some (UTYPE (imm,rd,op), sub_nat (string_length arg_) (string_length s_))
+ (match (utype_mnemonic_matches_prefix _stringappend_0_) with
+ | Some (op,(existT _ _stringappend_1_ _)) => returnm (op, build_ex _stringappend_1_)
+ | _ => exit tt : M ((uop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((uop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_1_ _) =>
+ let _stringappend_2_ := string_drop _stringappend_0_ (build_ex _stringappend_1_) in
+ (match (spc_matches_prefix _stringappend_2_) with
+ | Some (tt,(existT _ _stringappend_3_ _)) => returnm (tt, build_ex _stringappend_3_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_3_ _) =>
+ let _stringappend_4_ := string_drop _stringappend_2_ (build_ex _stringappend_3_) in
+ (match (reg_name_matches_prefix _stringappend_4_) with
+ | Some (rd,(existT _ _stringappend_5_ _)) => returnm (rd, build_ex _stringappend_5_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_5_ _) =>
+ let _stringappend_6_ := string_drop _stringappend_4_ (build_ex _stringappend_5_) in
+ sep_matches_prefix _stringappend_6_ >>= fun w__10 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+ (match w__10 with
+ | Some (tt,(existT _ _stringappend_7_ _)) => returnm (tt, build_ex _stringappend_7_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_7_ _) =>
+ let _stringappend_8_ := string_drop _stringappend_6_ (build_ex _stringappend_7_) in
+ (match (hex_bits_20_matches_prefix _stringappend_8_) with
+ | Some (imm,(existT _ _stringappend_9_ _)) => returnm (imm, build_ex _stringappend_9_)
+ | _ => exit tt : M ((mword 20 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 20 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_9_ _) =>
+ returnm ((match (string_drop _stringappend_8_ (build_ex _stringappend_9_)) with
+ | s_ =>
+ Some
+ ((UTYPE
+ (imm, rd, op), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
and_boolM (returnm ((string_startswith _stringappend_0_ "jal") : bool))
- (let _stringappend_16_ := string_drop _stringappend_0_ (string_length "jal") in
- match (spc_matches_prefix _stringappend_16_) with
- | Some (_stringappend_17_,(existT _ _stringappend_18_ _)) =>
- let _stringappend_19_ := string_drop _stringappend_16_ (build_ex _stringappend_18_) in
- match (reg_name_matches_prefix _stringappend_19_) with
- | Some (_stringappend_20_,(existT _ _stringappend_21_ _)) =>
- let _stringappend_22_ := string_drop _stringappend_19_ (build_ex _stringappend_21_) in
- sep_matches_prefix _stringappend_22_ >>= fun w__15 : option ((unit * {n : Z & ArithFact (n >=
+ (let _stringappend_11_ :=
+ string_drop _stringappend_0_ (build_ex (projT1 (string_length "jal"))) in
+ match (spc_matches_prefix _stringappend_11_) with
+ | Some (tt,(existT _ _stringappend_12_ _)) =>
+ let _stringappend_13_ := string_drop _stringappend_11_ (build_ex _stringappend_12_) in
+ match (reg_name_matches_prefix _stringappend_13_) with
+ | Some (rd,(existT _ _stringappend_14_ _)) =>
+ let _stringappend_15_ := string_drop _stringappend_13_ (build_ex _stringappend_14_) in
+ sep_matches_prefix _stringappend_15_ >>= fun w__15 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__15 with
- | Some (_stringappend_23_,(existT _ _stringappend_24_ _)) =>
- let _stringappend_25_ :=
- string_drop _stringappend_22_ (build_ex _stringappend_24_) in
- if ((match (hex_bits_21_matches_prefix _stringappend_25_) with
- | Some (_stringappend_26_,(existT _ _stringappend_27_ _)) =>
- match (string_drop _stringappend_25_
- (build_ex
- _stringappend_27_)) with
+ | Some (tt,(existT _ _stringappend_16_ _)) =>
+ let _stringappend_17_ :=
+ string_drop _stringappend_15_ (build_ex _stringappend_16_) in
+ if ((match (hex_bits_21_matches_prefix _stringappend_17_) with
+ | Some (imm,(existT _ _stringappend_18_ _)) =>
+ match (string_drop _stringappend_17_
+ (build_ex _stringappend_18_)) with
| s_ => true
end
| None => false
@@ -23748,82 +22282,69 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__18 : bool =>
(if (w__18) then
- let _stringappend_16_ := string_drop _stringappend_0_ (string_length "jal") in
- match (spc_matches_prefix _stringappend_16_) with
- | Some (_stringappend_17_,(existT _ _stringappend_18_ _)) =>
- returnm ((_stringappend_17_, build_ex _stringappend_18_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__20 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_18_ _) := w__20 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_19_ := string_drop _stringappend_16_ (build_ex _stringappend_18_) in
- match (reg_name_matches_prefix _stringappend_19_) with
- | Some (_stringappend_20_,(existT _ _stringappend_21_ _)) =>
- returnm ((_stringappend_20_, build_ex _stringappend_21_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__22 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_21_ _) :=
- w__22
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_22_ := string_drop _stringappend_19_ (build_ex _stringappend_21_) in
- sep_matches_prefix _stringappend_22_ >>= fun w__23 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_11_ :=
+ string_drop _stringappend_0_ (build_ex (projT1 (string_length "jal"))) in
+ (match (spc_matches_prefix _stringappend_11_) with
+ | Some (tt,(existT _ _stringappend_12_ _)) => returnm (tt, build_ex _stringappend_12_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_12_ _) =>
+ let _stringappend_13_ := string_drop _stringappend_11_ (build_ex _stringappend_12_) in
+ (match (reg_name_matches_prefix _stringappend_13_) with
+ | Some (rd,(existT _ _stringappend_14_ _)) => returnm (rd, build_ex _stringappend_14_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_14_ _) =>
+ let _stringappend_15_ := string_drop _stringappend_13_ (build_ex _stringappend_14_) in
+ sep_matches_prefix _stringappend_15_ >>= fun w__23 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__23 with
- | Some (_stringappend_23_,(existT _ _stringappend_24_ _)) =>
- returnm ((_stringappend_23_, build_ex _stringappend_24_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__25 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_24_ _) := w__25 : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_25_ := string_drop _stringappend_22_ (build_ex _stringappend_24_) in
- match (hex_bits_21_matches_prefix _stringappend_25_) with
- | Some (_stringappend_26_,(existT _ _stringappend_27_ _)) =>
- returnm ((_stringappend_26_, build_ex _stringappend_27_)
- : (mword 21 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 21 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__27 : (mword 21 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_27_ _) :=
- w__27
- : (mword 21 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_25_ (build_ex _stringappend_27_)) with
+ (match w__23 with
+ | Some (tt,(existT _ _stringappend_16_ _)) => returnm (tt, build_ex _stringappend_16_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_16_ _) =>
+ let _stringappend_17_ := string_drop _stringappend_15_ (build_ex _stringappend_16_) in
+ (match (hex_bits_21_matches_prefix _stringappend_17_) with
+ | Some (imm,(existT _ _stringappend_18_ _)) => returnm (imm, build_ex _stringappend_18_)
+ | _ => exit tt : M ((mword 21 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 21 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_18_ _) =>
+ returnm ((match (string_drop _stringappend_17_ (build_ex _stringappend_18_)) with
| s_ =>
- Some (RISCV_JAL (imm,rd), sub_nat (string_length arg_) (string_length s_))
+ Some
+ ((RISCV_JAL
+ (imm, rd), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length arg_)))
+ (build_ex (projT1 (string_length s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
and_boolM (returnm ((string_startswith _stringappend_0_ "jalr") : bool))
- (let _stringappend_29_ := string_drop _stringappend_0_ (string_length "jalr") in
- match (spc_matches_prefix _stringappend_29_) with
- | Some (_stringappend_30_,(existT _ _stringappend_31_ _)) =>
- let _stringappend_32_ := string_drop _stringappend_29_ (build_ex _stringappend_31_) in
- match (reg_name_matches_prefix _stringappend_32_) with
- | Some (_stringappend_33_,(existT _ _stringappend_34_ _)) =>
- let _stringappend_35_ := string_drop _stringappend_32_ (build_ex _stringappend_34_) in
- sep_matches_prefix _stringappend_35_ >>= fun w__28 : option ((unit * {n : Z & ArithFact (n >=
+ (let _stringappend_20_ :=
+ string_drop _stringappend_0_ (build_ex (projT1 (string_length "jalr"))) in
+ match (spc_matches_prefix _stringappend_20_) with
+ | Some (tt,(existT _ _stringappend_21_ _)) =>
+ let _stringappend_22_ := string_drop _stringappend_20_ (build_ex _stringappend_21_) in
+ match (reg_name_matches_prefix _stringappend_22_) with
+ | Some (rd,(existT _ _stringappend_23_ _)) =>
+ let _stringappend_24_ := string_drop _stringappend_22_ (build_ex _stringappend_23_) in
+ sep_matches_prefix _stringappend_24_ >>= fun w__28 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__28 with
- | Some (_stringappend_36_,(existT _ _stringappend_37_ _)) =>
- let _stringappend_38_ :=
- string_drop _stringappend_35_ (build_ex _stringappend_37_) in
- match (reg_name_matches_prefix _stringappend_38_) with
- | Some (_stringappend_39_,(existT _ _stringappend_40_ _)) =>
- let _stringappend_41_ :=
- string_drop _stringappend_38_ (build_ex _stringappend_40_) in
- sep_matches_prefix _stringappend_41_ >>= fun w__29 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_25_ _)) =>
+ let _stringappend_26_ :=
+ string_drop _stringappend_24_ (build_ex _stringappend_25_) in
+ match (reg_name_matches_prefix _stringappend_26_) with
+ | Some (rs1,(existT _ _stringappend_27_ _)) =>
+ let _stringappend_28_ :=
+ string_drop _stringappend_26_ (build_ex _stringappend_27_) in
+ sep_matches_prefix _stringappend_28_ >>= fun w__29 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__29 with
- | Some (_stringappend_42_,(existT _ _stringappend_43_ _)) =>
- let _stringappend_44_ :=
- string_drop _stringappend_41_
- (build_ex
- _stringappend_43_) in
- if ((match (hex_bits_12_matches_prefix _stringappend_44_) with
- | Some
- (_stringappend_45_,(existT _ _stringappend_46_ _)) =>
- match (string_drop _stringappend_44_
- (build_ex
- _stringappend_46_)) with
+ | Some (tt,(existT _ _stringappend_29_ _)) =>
+ let _stringappend_30_ :=
+ string_drop _stringappend_28_
+ (build_ex _stringappend_29_) in
+ if ((match (hex_bits_12_matches_prefix _stringappend_30_) with
+ | Some (imm,(existT _ _stringappend_31_ _)) =>
+ match (string_drop _stringappend_30_
+ (build_ex _stringappend_31_)) with
| s_ => true
end
| None => false
@@ -23856,113 +22377,88 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__34 : bool =>
(if (w__34) then
- let _stringappend_29_ := string_drop _stringappend_0_ (string_length "jalr") in
- match (spc_matches_prefix _stringappend_29_) with
- | Some (_stringappend_30_,(existT _ _stringappend_31_ _)) =>
- returnm ((_stringappend_30_, build_ex _stringappend_31_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__36 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_31_ _) :=
- w__36
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_32_ := string_drop _stringappend_29_ (build_ex _stringappend_31_) in
- match (reg_name_matches_prefix _stringappend_32_) with
- | Some (_stringappend_33_,(existT _ _stringappend_34_ _)) =>
- returnm ((_stringappend_33_, build_ex _stringappend_34_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__38 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_34_ _) :=
- w__38
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_35_ := string_drop _stringappend_32_ (build_ex _stringappend_34_) in
- sep_matches_prefix _stringappend_35_ >>= fun w__39 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_20_ :=
+ string_drop _stringappend_0_ (build_ex (projT1 (string_length "jalr"))) in
+ (match (spc_matches_prefix _stringappend_20_) with
+ | Some (tt,(existT _ _stringappend_21_ _)) => returnm (tt, build_ex _stringappend_21_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_21_ _) =>
+ let _stringappend_22_ := string_drop _stringappend_20_ (build_ex _stringappend_21_) in
+ (match (reg_name_matches_prefix _stringappend_22_) with
+ | Some (rd,(existT _ _stringappend_23_ _)) => returnm (rd, build_ex _stringappend_23_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_23_ _) =>
+ let _stringappend_24_ := string_drop _stringappend_22_ (build_ex _stringappend_23_) in
+ sep_matches_prefix _stringappend_24_ >>= fun w__39 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__39 with
- | Some (_stringappend_36_,(existT _ _stringappend_37_ _)) =>
- returnm ((_stringappend_36_, build_ex _stringappend_37_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__41 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_37_ _) :=
- w__41
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_38_ := string_drop _stringappend_35_ (build_ex _stringappend_37_) in
- match (reg_name_matches_prefix _stringappend_38_) with
- | Some (_stringappend_39_,(existT _ _stringappend_40_ _)) =>
- returnm ((_stringappend_39_, build_ex _stringappend_40_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__43 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_40_ _) :=
- w__43
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_41_ := string_drop _stringappend_38_ (build_ex _stringappend_40_) in
- sep_matches_prefix _stringappend_41_ >>= fun w__44 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__39 with
+ | Some (tt,(existT _ _stringappend_25_ _)) => returnm (tt, build_ex _stringappend_25_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_25_ _) =>
+ let _stringappend_26_ := string_drop _stringappend_24_ (build_ex _stringappend_25_) in
+ (match (reg_name_matches_prefix _stringappend_26_) with
+ | Some (rs1,(existT _ _stringappend_27_ _)) =>
+ returnm (rs1, build_ex _stringappend_27_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_27_ _) =>
+ let _stringappend_28_ := string_drop _stringappend_26_ (build_ex _stringappend_27_) in
+ sep_matches_prefix _stringappend_28_ >>= fun w__44 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__44 with
- | Some (_stringappend_42_,(existT _ _stringappend_43_ _)) =>
- returnm ((_stringappend_42_, build_ex _stringappend_43_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__46 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_43_ _) :=
- w__46
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_44_ := string_drop _stringappend_41_ (build_ex _stringappend_43_) in
- match (hex_bits_12_matches_prefix _stringappend_44_) with
- | Some (_stringappend_45_,(existT _ _stringappend_46_ _)) =>
- returnm ((_stringappend_45_, build_ex _stringappend_46_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__48 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_46_ _) :=
- w__48
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_44_ (build_ex _stringappend_46_)) with
+ (match w__44 with
+ | Some (tt,(existT _ _stringappend_29_ _)) => returnm (tt, build_ex _stringappend_29_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_29_ _) =>
+ let _stringappend_30_ := string_drop _stringappend_28_ (build_ex _stringappend_29_) in
+ (match (hex_bits_12_matches_prefix _stringappend_30_) with
+ | Some (imm,(existT _ _stringappend_31_ _)) =>
+ returnm (imm, build_ex _stringappend_31_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_31_ _) =>
+ returnm ((match (string_drop _stringappend_30_ (build_ex _stringappend_31_)) with
| s_ =>
- Some (RISCV_JALR (imm,rs1,rd),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ ((RISCV_JALR
+ (imm, rs1, rd), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
match (btype_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_48_,(existT _ _stringappend_49_ _)) =>
- let _stringappend_50_ := string_drop _stringappend_0_ (build_ex _stringappend_49_) in
- match (spc_matches_prefix _stringappend_50_) with
- | Some (_stringappend_51_,(existT _ _stringappend_52_ _)) =>
- let _stringappend_53_ :=
- string_drop _stringappend_50_ (build_ex _stringappend_52_) in
- match (reg_name_matches_prefix _stringappend_53_) with
- | Some (_stringappend_54_,(existT _ _stringappend_55_ _)) =>
- let _stringappend_56_ :=
- string_drop _stringappend_53_ (build_ex _stringappend_55_) in
- sep_matches_prefix _stringappend_56_ >>= fun w__49 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op,(existT _ _stringappend_33_ _)) =>
+ let _stringappend_34_ := string_drop _stringappend_0_ (build_ex _stringappend_33_) in
+ match (spc_matches_prefix _stringappend_34_) with
+ | Some (tt,(existT _ _stringappend_35_ _)) =>
+ let _stringappend_36_ :=
+ string_drop _stringappend_34_ (build_ex _stringappend_35_) in
+ match (reg_name_matches_prefix _stringappend_36_) with
+ | Some (rs1,(existT _ _stringappend_37_ _)) =>
+ let _stringappend_38_ :=
+ string_drop _stringappend_36_ (build_ex _stringappend_37_) in
+ sep_matches_prefix _stringappend_38_ >>= fun w__49 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__49 with
- | Some (_stringappend_57_,(existT _ _stringappend_58_ _)) =>
- let _stringappend_59_ :=
- string_drop _stringappend_56_ (build_ex _stringappend_58_) in
- match (reg_name_matches_prefix _stringappend_59_) with
- | Some (_stringappend_60_,(existT _ _stringappend_61_ _)) =>
- let _stringappend_62_ :=
- string_drop _stringappend_59_ (build_ex _stringappend_61_) in
- sep_matches_prefix _stringappend_62_ >>= fun w__50 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_39_ _)) =>
+ let _stringappend_40_ :=
+ string_drop _stringappend_38_ (build_ex _stringappend_39_) in
+ match (reg_name_matches_prefix _stringappend_40_) with
+ | Some (rs2,(existT _ _stringappend_41_ _)) =>
+ let _stringappend_42_ :=
+ string_drop _stringappend_40_ (build_ex _stringappend_41_) in
+ sep_matches_prefix _stringappend_42_ >>= fun w__50 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__50 with
- | Some (_stringappend_63_,(existT _ _stringappend_64_ _)) =>
- let _stringappend_65_ :=
- string_drop _stringappend_62_
- (build_ex
- _stringappend_64_) in
+ | Some (tt,(existT _ _stringappend_43_ _)) =>
+ let _stringappend_44_ :=
+ string_drop _stringappend_42_
+ (build_ex _stringappend_43_) in
if ((match (hex_bits_13_matches_prefix
- _stringappend_65_) with
- | Some
- (_stringappend_66_,(existT _ _stringappend_67_ _)) =>
- match (string_drop _stringappend_65_
- (build_ex
- _stringappend_67_)) with
+ _stringappend_44_) with
+ | Some (imm,(existT _ _stringappend_45_ _)) =>
+ match (string_drop _stringappend_44_
+ (build_ex _stringappend_45_)) with
| s_ => true
end
| None => false
@@ -23997,123 +22493,96 @@ Definition assembly_matches_prefix (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__55 : bool =>
(if (w__55) then
- match (btype_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_48_,(existT _ _stringappend_49_ _)) =>
- returnm ((_stringappend_48_, build_ex _stringappend_49_)
- : (bop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__57 : (bop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_49_ _) :=
- w__57
- : (bop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_50_ := string_drop _stringappend_0_ (build_ex _stringappend_49_) in
- match (spc_matches_prefix _stringappend_50_) with
- | Some (_stringappend_51_,(existT _ _stringappend_52_ _)) =>
- returnm ((_stringappend_51_, build_ex _stringappend_52_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__59 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_52_ _) :=
- w__59
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_53_ := string_drop _stringappend_50_ (build_ex _stringappend_52_) in
- match (reg_name_matches_prefix _stringappend_53_) with
- | Some (_stringappend_54_,(existT _ _stringappend_55_ _)) =>
- returnm ((_stringappend_54_, build_ex _stringappend_55_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__61 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_55_ _) :=
- w__61
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_56_ := string_drop _stringappend_53_ (build_ex _stringappend_55_) in
- sep_matches_prefix _stringappend_56_ >>= fun w__62 : option ((unit * {n : Z & ArithFact (n >=
+ (match (btype_mnemonic_matches_prefix _stringappend_0_) with
+ | Some (op,(existT _ _stringappend_33_ _)) =>
+ returnm (op, build_ex _stringappend_33_)
+ | _ => exit tt : M ((bop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_33_ _) =>
+ let _stringappend_34_ := string_drop _stringappend_0_ (build_ex _stringappend_33_) in
+ (match (spc_matches_prefix _stringappend_34_) with
+ | Some (tt,(existT _ _stringappend_35_ _)) =>
+ returnm (tt, build_ex _stringappend_35_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_35_ _) =>
+ let _stringappend_36_ := string_drop _stringappend_34_ (build_ex _stringappend_35_) in
+ (match (reg_name_matches_prefix _stringappend_36_) with
+ | Some (rs1,(existT _ _stringappend_37_ _)) =>
+ returnm (rs1, build_ex _stringappend_37_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_37_ _) =>
+ let _stringappend_38_ := string_drop _stringappend_36_ (build_ex _stringappend_37_) in
+ sep_matches_prefix _stringappend_38_ >>= fun w__62 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__62 with
- | Some (_stringappend_57_,(existT _ _stringappend_58_ _)) =>
- returnm ((_stringappend_57_, build_ex _stringappend_58_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__64 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_58_ _) :=
- w__64
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_59_ := string_drop _stringappend_56_ (build_ex _stringappend_58_) in
- match (reg_name_matches_prefix _stringappend_59_) with
- | Some (_stringappend_60_,(existT _ _stringappend_61_ _)) =>
- returnm ((_stringappend_60_, build_ex _stringappend_61_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__66 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs2, existT _ _stringappend_61_ _) :=
- w__66
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_62_ := string_drop _stringappend_59_ (build_ex _stringappend_61_) in
- sep_matches_prefix _stringappend_62_ >>= fun w__67 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__62 with
+ | Some (tt,(existT _ _stringappend_39_ _)) =>
+ returnm (tt, build_ex _stringappend_39_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_39_ _) =>
+ let _stringappend_40_ := string_drop _stringappend_38_ (build_ex _stringappend_39_) in
+ (match (reg_name_matches_prefix _stringappend_40_) with
+ | Some (rs2,(existT _ _stringappend_41_ _)) =>
+ returnm (rs2, build_ex _stringappend_41_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_41_ _) =>
+ let _stringappend_42_ := string_drop _stringappend_40_ (build_ex _stringappend_41_) in
+ sep_matches_prefix _stringappend_42_ >>= fun w__67 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__67 with
- | Some (_stringappend_63_,(existT _ _stringappend_64_ _)) =>
- returnm ((_stringappend_63_, build_ex _stringappend_64_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__69 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_64_ _) :=
- w__69
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_65_ := string_drop _stringappend_62_ (build_ex _stringappend_64_) in
- match (hex_bits_13_matches_prefix _stringappend_65_) with
- | Some (_stringappend_66_,(existT _ _stringappend_67_ _)) =>
- returnm ((_stringappend_66_, build_ex _stringappend_67_)
- : (mword 13 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 13 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__71 : (mword 13 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_67_ _) :=
- w__71
- : (mword 13 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_65_ (build_ex _stringappend_67_)) with
+ (match w__67 with
+ | Some (tt,(existT _ _stringappend_43_ _)) =>
+ returnm (tt, build_ex _stringappend_43_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_43_ _) =>
+ let _stringappend_44_ := string_drop _stringappend_42_ (build_ex _stringappend_43_) in
+ (match (hex_bits_13_matches_prefix _stringappend_44_) with
+ | Some (imm,(existT _ _stringappend_45_ _)) =>
+ returnm (imm, build_ex _stringappend_45_)
+ | _ => exit tt : M ((mword 13 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 13 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_45_ _) =>
+ returnm ((match (string_drop _stringappend_44_ (build_ex _stringappend_45_)) with
| s_ =>
- Some (BTYPE (imm,rs2,rs1,op),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ ((BTYPE
+ (imm, rs2, rs1, op), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
match (itype_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_69_,(existT _ _stringappend_70_ _)) =>
- let _stringappend_71_ := string_drop _stringappend_0_ (build_ex _stringappend_70_) in
- match (spc_matches_prefix _stringappend_71_) with
- | Some (_stringappend_72_,(existT _ _stringappend_73_ _)) =>
- let _stringappend_74_ :=
- string_drop _stringappend_71_ (build_ex _stringappend_73_) in
- match (reg_name_matches_prefix _stringappend_74_) with
- | Some (_stringappend_75_,(existT _ _stringappend_76_ _)) =>
- let _stringappend_77_ :=
- string_drop _stringappend_74_ (build_ex _stringappend_76_) in
- sep_matches_prefix _stringappend_77_ >>= fun w__72 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op,(existT _ _stringappend_47_ _)) =>
+ let _stringappend_48_ := string_drop _stringappend_0_ (build_ex _stringappend_47_) in
+ match (spc_matches_prefix _stringappend_48_) with
+ | Some (tt,(existT _ _stringappend_49_ _)) =>
+ let _stringappend_50_ :=
+ string_drop _stringappend_48_ (build_ex _stringappend_49_) in
+ match (reg_name_matches_prefix _stringappend_50_) with
+ | Some (rd,(existT _ _stringappend_51_ _)) =>
+ let _stringappend_52_ :=
+ string_drop _stringappend_50_ (build_ex _stringappend_51_) in
+ sep_matches_prefix _stringappend_52_ >>= fun w__72 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__72 with
- | Some (_stringappend_78_,(existT _ _stringappend_79_ _)) =>
- let _stringappend_80_ :=
- string_drop _stringappend_77_ (build_ex _stringappend_79_) in
- match (reg_name_matches_prefix _stringappend_80_) with
- | Some (_stringappend_81_,(existT _ _stringappend_82_ _)) =>
- let _stringappend_83_ :=
- string_drop _stringappend_80_ (build_ex _stringappend_82_) in
- sep_matches_prefix _stringappend_83_ >>= fun w__73 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_53_ _)) =>
+ let _stringappend_54_ :=
+ string_drop _stringappend_52_ (build_ex _stringappend_53_) in
+ match (reg_name_matches_prefix _stringappend_54_) with
+ | Some (rs1,(existT _ _stringappend_55_ _)) =>
+ let _stringappend_56_ :=
+ string_drop _stringappend_54_ (build_ex _stringappend_55_) in
+ sep_matches_prefix _stringappend_56_ >>= fun w__73 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__73 with
- | Some
- (_stringappend_84_,(existT _ _stringappend_85_ _)) =>
- let _stringappend_86_ :=
- string_drop _stringappend_83_
- (build_ex
- _stringappend_85_) in
+ | Some (tt,(existT _ _stringappend_57_ _)) =>
+ let _stringappend_58_ :=
+ string_drop _stringappend_56_
+ (build_ex _stringappend_57_) in
if ((match (hex_bits_12_matches_prefix
- _stringappend_86_) with
- | Some
- (_stringappend_87_,(existT _ _stringappend_88_ _)) =>
- match (string_drop _stringappend_86_
- (build_ex
- _stringappend_88_)) with
+ _stringappend_58_) with
+ | Some (imm,(existT _ _stringappend_59_ _)) =>
+ match (string_drop _stringappend_58_
+ (build_ex _stringappend_59_)) with
| s_ => true
end
| None => false
@@ -24148,125 +22617,98 @@ Definition assembly_matches_prefix (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__78 : bool =>
(if (w__78) then
- match (itype_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_69_,(existT _ _stringappend_70_ _)) =>
- returnm ((_stringappend_69_, build_ex _stringappend_70_)
- : (iop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((iop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__80 : (iop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_70_ _) :=
- w__80
- : (iop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_71_ := string_drop _stringappend_0_ (build_ex _stringappend_70_) in
- match (spc_matches_prefix _stringappend_71_) with
- | Some (_stringappend_72_,(existT _ _stringappend_73_ _)) =>
- returnm ((_stringappend_72_, build_ex _stringappend_73_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__82 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_73_ _) :=
- w__82
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_74_ :=
- string_drop _stringappend_71_ (build_ex _stringappend_73_) in
- match (reg_name_matches_prefix _stringappend_74_) with
- | Some (_stringappend_75_,(existT _ _stringappend_76_ _)) =>
- returnm ((_stringappend_75_, build_ex _stringappend_76_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__84 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_76_ _) :=
- w__84
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_77_ :=
- string_drop _stringappend_74_ (build_ex _stringappend_76_) in
- sep_matches_prefix _stringappend_77_ >>= fun w__85 : option ((unit * {n : Z & ArithFact (n >=
+ (match (itype_mnemonic_matches_prefix _stringappend_0_) with
+ | Some (op,(existT _ _stringappend_47_ _)) =>
+ returnm (op, build_ex _stringappend_47_)
+ | _ => exit tt : M ((iop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((iop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_47_ _) =>
+ let _stringappend_48_ := string_drop _stringappend_0_ (build_ex _stringappend_47_) in
+ (match (spc_matches_prefix _stringappend_48_) with
+ | Some (tt,(existT _ _stringappend_49_ _)) =>
+ returnm (tt, build_ex _stringappend_49_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_49_ _) =>
+ let _stringappend_50_ :=
+ string_drop _stringappend_48_ (build_ex _stringappend_49_) in
+ (match (reg_name_matches_prefix _stringappend_50_) with
+ | Some (rd,(existT _ _stringappend_51_ _)) =>
+ returnm (rd, build_ex _stringappend_51_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_51_ _) =>
+ let _stringappend_52_ :=
+ string_drop _stringappend_50_ (build_ex _stringappend_51_) in
+ sep_matches_prefix _stringappend_52_ >>= fun w__85 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__85 with
- | Some (_stringappend_78_,(existT _ _stringappend_79_ _)) =>
- returnm ((_stringappend_78_, build_ex _stringappend_79_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__87 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_79_ _) :=
- w__87
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_80_ :=
- string_drop _stringappend_77_ (build_ex _stringappend_79_) in
- match (reg_name_matches_prefix _stringappend_80_) with
- | Some (_stringappend_81_,(existT _ _stringappend_82_ _)) =>
- returnm ((_stringappend_81_, build_ex _stringappend_82_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__89 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_82_ _) :=
- w__89
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_83_ :=
- string_drop _stringappend_80_ (build_ex _stringappend_82_) in
- sep_matches_prefix _stringappend_83_ >>= fun w__90 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__85 with
+ | Some (tt,(existT _ _stringappend_53_ _)) =>
+ returnm (tt, build_ex _stringappend_53_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_53_ _) =>
+ let _stringappend_54_ :=
+ string_drop _stringappend_52_ (build_ex _stringappend_53_) in
+ (match (reg_name_matches_prefix _stringappend_54_) with
+ | Some (rs1,(existT _ _stringappend_55_ _)) =>
+ returnm (rs1, build_ex _stringappend_55_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_55_ _) =>
+ let _stringappend_56_ :=
+ string_drop _stringappend_54_ (build_ex _stringappend_55_) in
+ sep_matches_prefix _stringappend_56_ >>= fun w__90 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__90 with
- | Some (_stringappend_84_,(existT _ _stringappend_85_ _)) =>
- returnm ((_stringappend_84_, build_ex _stringappend_85_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__92 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_85_ _) :=
- w__92
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_86_ :=
- string_drop _stringappend_83_ (build_ex _stringappend_85_) in
- match (hex_bits_12_matches_prefix _stringappend_86_) with
- | Some (_stringappend_87_,(existT _ _stringappend_88_ _)) =>
- returnm ((_stringappend_87_, build_ex _stringappend_88_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__94 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_88_ _) :=
- w__94
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_86_ (build_ex _stringappend_88_)) with
+ (match w__90 with
+ | Some (tt,(existT _ _stringappend_57_ _)) =>
+ returnm (tt, build_ex _stringappend_57_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_57_ _) =>
+ let _stringappend_58_ :=
+ string_drop _stringappend_56_ (build_ex _stringappend_57_) in
+ (match (hex_bits_12_matches_prefix _stringappend_58_) with
+ | Some (imm,(existT _ _stringappend_59_ _)) =>
+ returnm (imm, build_ex _stringappend_59_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_59_ _) =>
+ returnm ((match (string_drop _stringappend_58_ (build_ex _stringappend_59_)) with
| s_ =>
- Some (ITYPE (imm,rs1,rd,op),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ ((ITYPE
+ (imm, rs1, rd, op), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
match (shiftiop_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_90_,(existT _ _stringappend_91_ _)) =>
- let _stringappend_92_ :=
- string_drop _stringappend_0_ (build_ex _stringappend_91_) in
- match (spc_matches_prefix _stringappend_92_) with
- | Some (_stringappend_93_,(existT _ _stringappend_94_ _)) =>
- let _stringappend_95_ :=
- string_drop _stringappend_92_ (build_ex _stringappend_94_) in
- match (reg_name_matches_prefix _stringappend_95_) with
- | Some (_stringappend_96_,(existT _ _stringappend_97_ _)) =>
- let _stringappend_98_ :=
- string_drop _stringappend_95_ (build_ex _stringappend_97_) in
- sep_matches_prefix _stringappend_98_ >>= fun w__95 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op,(existT _ _stringappend_61_ _)) =>
+ let _stringappend_62_ :=
+ string_drop _stringappend_0_ (build_ex _stringappend_61_) in
+ match (spc_matches_prefix _stringappend_62_) with
+ | Some (tt,(existT _ _stringappend_63_ _)) =>
+ let _stringappend_64_ :=
+ string_drop _stringappend_62_ (build_ex _stringappend_63_) in
+ match (reg_name_matches_prefix _stringappend_64_) with
+ | Some (rd,(existT _ _stringappend_65_ _)) =>
+ let _stringappend_66_ :=
+ string_drop _stringappend_64_ (build_ex _stringappend_65_) in
+ sep_matches_prefix _stringappend_66_ >>= fun w__95 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__95 with
- | Some (_stringappend_99_,(existT _ _stringappend_100_ _)) =>
- let _stringappend_101_ :=
- string_drop _stringappend_98_
- (build_ex
- _stringappend_100_) in
- if ((match (reg_name_matches_prefix _stringappend_101_) with
- | Some
- (_stringappend_102_,(existT _ _stringappend_103_ _)) =>
- let _stringappend_104_ :=
- string_drop _stringappend_101_
- (build_ex
- _stringappend_103_) in
+ | Some (tt,(existT _ _stringappend_67_ _)) =>
+ let _stringappend_68_ :=
+ string_drop _stringappend_66_
+ (build_ex _stringappend_67_) in
+ if ((match (reg_name_matches_prefix _stringappend_68_) with
+ | Some (rs1,(existT _ _stringappend_69_ _)) =>
+ let _stringappend_70_ :=
+ string_drop _stringappend_68_
+ (build_ex _stringappend_69_) in
if ((match (hex_bits_6_matches_prefix
- _stringappend_104_) with
+ _stringappend_70_) with
| Some
- (_stringappend_105_,(existT _ _stringappend_106_ _)) =>
- match (string_drop _stringappend_104_
- (build_ex
- _stringappend_106_)) with
+ (shamt,(existT _ _stringappend_71_ _)) =>
+ match (string_drop _stringappend_70_
+ (build_ex _stringappend_71_)) with
| s_ => true
end
| None => false
@@ -24295,117 +22737,95 @@ Definition assembly_matches_prefix (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__98 : bool =>
(if (w__98) then
- match (shiftiop_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_90_,(existT _ _stringappend_91_ _)) =>
- returnm ((_stringappend_90_, build_ex _stringappend_91_)
- : (sop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((sop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__100 : (sop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_91_ _) :=
- w__100
- : (sop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_92_ :=
- string_drop _stringappend_0_ (build_ex _stringappend_91_) in
- match (spc_matches_prefix _stringappend_92_) with
- | Some (_stringappend_93_,(existT _ _stringappend_94_ _)) =>
- returnm ((_stringappend_93_, build_ex _stringappend_94_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__102 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_94_ _) :=
- w__102
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_95_ :=
- string_drop _stringappend_92_ (build_ex _stringappend_94_) in
- match (reg_name_matches_prefix _stringappend_95_) with
- | Some (_stringappend_96_,(existT _ _stringappend_97_ _)) =>
- returnm ((_stringappend_96_, build_ex _stringappend_97_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__104 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_97_ _) :=
- w__104
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_98_ :=
- string_drop _stringappend_95_ (build_ex _stringappend_97_) in
- sep_matches_prefix _stringappend_98_ >>= fun w__105 : option ((unit * {n : Z & ArithFact (n >=
+ (match (shiftiop_mnemonic_matches_prefix _stringappend_0_) with
+ | Some (op,(existT _ _stringappend_61_ _)) =>
+ returnm (op, build_ex _stringappend_61_)
+ | _ => exit tt : M ((sop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((sop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_61_ _) =>
+ let _stringappend_62_ :=
+ string_drop _stringappend_0_ (build_ex _stringappend_61_) in
+ (match (spc_matches_prefix _stringappend_62_) with
+ | Some (tt,(existT _ _stringappend_63_ _)) =>
+ returnm (tt, build_ex _stringappend_63_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_63_ _) =>
+ let _stringappend_64_ :=
+ string_drop _stringappend_62_ (build_ex _stringappend_63_) in
+ (match (reg_name_matches_prefix _stringappend_64_) with
+ | Some (rd,(existT _ _stringappend_65_ _)) =>
+ returnm (rd, build_ex _stringappend_65_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_65_ _) =>
+ let _stringappend_66_ :=
+ string_drop _stringappend_64_ (build_ex _stringappend_65_) in
+ sep_matches_prefix _stringappend_66_ >>= fun w__105 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__105 with
- | Some (_stringappend_99_,(existT _ _stringappend_100_ _)) =>
- returnm ((_stringappend_99_, build_ex _stringappend_100_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__107 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_100_ _) :=
- w__107
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_101_ :=
- string_drop _stringappend_98_ (build_ex _stringappend_100_) in
- match (reg_name_matches_prefix _stringappend_101_) with
- | Some (_stringappend_102_,(existT _ _stringappend_103_ _)) =>
- returnm ((_stringappend_102_, build_ex _stringappend_103_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__109 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_103_ _) :=
- w__109
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_104_ :=
- string_drop _stringappend_101_ (build_ex _stringappend_103_) in
- match (hex_bits_6_matches_prefix _stringappend_104_) with
- | Some (_stringappend_105_,(existT _ _stringappend_106_ _)) =>
- returnm ((_stringappend_105_, build_ex _stringappend_106_)
- : (mword 6 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 6 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__111 : (mword 6 * {n : Z & ArithFact (n >= 0)}) =>
- let '(shamt, existT _ _stringappend_106_ _) :=
- w__111
- : (mword 6 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_104_ (build_ex _stringappend_106_)) with
+ (match w__105 with
+ | Some (tt,(existT _ _stringappend_67_ _)) =>
+ returnm (tt, build_ex _stringappend_67_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_67_ _) =>
+ let _stringappend_68_ :=
+ string_drop _stringappend_66_ (build_ex _stringappend_67_) in
+ (match (reg_name_matches_prefix _stringappend_68_) with
+ | Some (rs1,(existT _ _stringappend_69_ _)) =>
+ returnm (rs1, build_ex _stringappend_69_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_69_ _) =>
+ let _stringappend_70_ :=
+ string_drop _stringappend_68_ (build_ex _stringappend_69_) in
+ (match (hex_bits_6_matches_prefix _stringappend_70_) with
+ | Some (shamt,(existT _ _stringappend_71_ _)) =>
+ returnm (shamt, build_ex _stringappend_71_)
+ | _ => exit tt : M ((mword 6 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 6 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(shamt, existT _ _stringappend_71_ _) =>
+ returnm ((match (string_drop _stringappend_70_ (build_ex _stringappend_71_)) with
| s_ =>
- Some (SHIFTIOP (shamt,rs1,rd,op),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ ((SHIFTIOP
+ (shamt, rs1, rd, op), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
match (rtype_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_108_,(existT _ _stringappend_109_ _)) =>
- let _stringappend_110_ :=
- string_drop _stringappend_0_ (build_ex _stringappend_109_) in
- match (spc_matches_prefix _stringappend_110_) with
- | Some (_stringappend_111_,(existT _ _stringappend_112_ _)) =>
- let _stringappend_113_ :=
- string_drop _stringappend_110_ (build_ex _stringappend_112_) in
- match (reg_name_matches_prefix _stringappend_113_) with
- | Some (_stringappend_114_,(existT _ _stringappend_115_ _)) =>
- let _stringappend_116_ :=
- string_drop _stringappend_113_ (build_ex _stringappend_115_) in
- sep_matches_prefix _stringappend_116_ >>= fun w__112 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op,(existT _ _stringappend_73_ _)) =>
+ let _stringappend_74_ :=
+ string_drop _stringappend_0_ (build_ex _stringappend_73_) in
+ match (spc_matches_prefix _stringappend_74_) with
+ | Some (tt,(existT _ _stringappend_75_ _)) =>
+ let _stringappend_76_ :=
+ string_drop _stringappend_74_ (build_ex _stringappend_75_) in
+ match (reg_name_matches_prefix _stringappend_76_) with
+ | Some (rd,(existT _ _stringappend_77_ _)) =>
+ let _stringappend_78_ :=
+ string_drop _stringappend_76_ (build_ex _stringappend_77_) in
+ sep_matches_prefix _stringappend_78_ >>= fun w__112 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__112 with
- | Some (_stringappend_117_,(existT _ _stringappend_118_ _)) =>
- let _stringappend_119_ :=
- string_drop _stringappend_116_ (build_ex _stringappend_118_) in
- match (reg_name_matches_prefix _stringappend_119_) with
- | Some (_stringappend_120_,(existT _ _stringappend_121_ _)) =>
- let _stringappend_122_ :=
- string_drop _stringappend_119_ (build_ex _stringappend_121_) in
- sep_matches_prefix _stringappend_122_ >>= fun w__113 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_79_ _)) =>
+ let _stringappend_80_ :=
+ string_drop _stringappend_78_ (build_ex _stringappend_79_) in
+ match (reg_name_matches_prefix _stringappend_80_) with
+ | Some (rs1,(existT _ _stringappend_81_ _)) =>
+ let _stringappend_82_ :=
+ string_drop _stringappend_80_ (build_ex _stringappend_81_) in
+ sep_matches_prefix _stringappend_82_ >>= fun w__113 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__113 with
- | Some
- (_stringappend_123_,(existT _ _stringappend_124_ _)) =>
- let _stringappend_125_ :=
- string_drop _stringappend_122_
- (build_ex
- _stringappend_124_) in
+ | Some (tt,(existT _ _stringappend_83_ _)) =>
+ let _stringappend_84_ :=
+ string_drop _stringappend_82_
+ (build_ex _stringappend_83_) in
if ((match (reg_name_matches_prefix
- _stringappend_125_) with
+ _stringappend_84_) with
| Some
- (_stringappend_126_,(existT _ _stringappend_127_ _)) =>
- match (string_drop _stringappend_125_
- (build_ex
- _stringappend_127_)) with
+ (rs2,(existT _ _stringappend_85_ _)) =>
+ match (string_drop _stringappend_84_
+ (build_ex _stringappend_85_)) with
| s_ => true
end
| None => false
@@ -24440,158 +22860,125 @@ Definition assembly_matches_prefix (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__118 : bool =>
(if (w__118) then
- match (rtype_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_108_,(existT _ _stringappend_109_ _)) =>
- returnm ((_stringappend_108_, build_ex _stringappend_109_)
- : (rop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((rop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__120 : (rop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_109_ _) :=
- w__120
- : (rop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_110_ :=
- string_drop _stringappend_0_ (build_ex _stringappend_109_) in
- match (spc_matches_prefix _stringappend_110_) with
- | Some (_stringappend_111_,(existT _ _stringappend_112_ _)) =>
- returnm ((_stringappend_111_, build_ex _stringappend_112_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__122 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_112_ _) :=
- w__122
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_113_ :=
- string_drop _stringappend_110_ (build_ex _stringappend_112_) in
- match (reg_name_matches_prefix _stringappend_113_) with
- | Some (_stringappend_114_,(existT _ _stringappend_115_ _)) =>
- returnm ((_stringappend_114_, build_ex _stringappend_115_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__124 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_115_ _) :=
- w__124
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_116_ :=
- string_drop _stringappend_113_ (build_ex _stringappend_115_) in
- sep_matches_prefix _stringappend_116_ >>= fun w__125 : option ((unit * {n : Z & ArithFact (n >=
+ (match (rtype_mnemonic_matches_prefix _stringappend_0_) with
+ | Some (op,(existT _ _stringappend_73_ _)) =>
+ returnm (op, build_ex _stringappend_73_)
+ | _ => exit tt : M ((rop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((rop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_73_ _) =>
+ let _stringappend_74_ :=
+ string_drop _stringappend_0_ (build_ex _stringappend_73_) in
+ (match (spc_matches_prefix _stringappend_74_) with
+ | Some (tt,(existT _ _stringappend_75_ _)) =>
+ returnm (tt, build_ex _stringappend_75_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_75_ _) =>
+ let _stringappend_76_ :=
+ string_drop _stringappend_74_ (build_ex _stringappend_75_) in
+ (match (reg_name_matches_prefix _stringappend_76_) with
+ | Some (rd,(existT _ _stringappend_77_ _)) =>
+ returnm (rd, build_ex _stringappend_77_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_77_ _) =>
+ let _stringappend_78_ :=
+ string_drop _stringappend_76_ (build_ex _stringappend_77_) in
+ sep_matches_prefix _stringappend_78_ >>= fun w__125 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__125 with
- | Some (_stringappend_117_,(existT _ _stringappend_118_ _)) =>
- returnm ((_stringappend_117_, build_ex _stringappend_118_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__127 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_118_ _) :=
- w__127
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_119_ :=
- string_drop _stringappend_116_ (build_ex _stringappend_118_) in
- match (reg_name_matches_prefix _stringappend_119_) with
- | Some (_stringappend_120_,(existT _ _stringappend_121_ _)) =>
- returnm ((_stringappend_120_, build_ex _stringappend_121_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__129 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_121_ _) :=
- w__129
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_122_ :=
- string_drop _stringappend_119_ (build_ex _stringappend_121_) in
- sep_matches_prefix _stringappend_122_ >>= fun w__130 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__125 with
+ | Some (tt,(existT _ _stringappend_79_ _)) =>
+ returnm (tt, build_ex _stringappend_79_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_79_ _) =>
+ let _stringappend_80_ :=
+ string_drop _stringappend_78_ (build_ex _stringappend_79_) in
+ (match (reg_name_matches_prefix _stringappend_80_) with
+ | Some (rs1,(existT _ _stringappend_81_ _)) =>
+ returnm (rs1, build_ex _stringappend_81_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_81_ _) =>
+ let _stringappend_82_ :=
+ string_drop _stringappend_80_ (build_ex _stringappend_81_) in
+ sep_matches_prefix _stringappend_82_ >>= fun w__130 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__130 with
- | Some (_stringappend_123_,(existT _ _stringappend_124_ _)) =>
- returnm ((_stringappend_123_, build_ex _stringappend_124_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__132 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_124_ _) :=
- w__132
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_125_ :=
- string_drop _stringappend_122_ (build_ex _stringappend_124_) in
- match (reg_name_matches_prefix _stringappend_125_) with
- | Some (_stringappend_126_,(existT _ _stringappend_127_ _)) =>
- returnm ((_stringappend_126_, build_ex _stringappend_127_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__134 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs2, existT _ _stringappend_127_ _) :=
- w__134
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_125_
- (build_ex
- _stringappend_127_)) with
+ (match w__130 with
+ | Some (tt,(existT _ _stringappend_83_ _)) =>
+ returnm (tt, build_ex _stringappend_83_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_83_ _) =>
+ let _stringappend_84_ :=
+ string_drop _stringappend_82_ (build_ex _stringappend_83_) in
+ (match (reg_name_matches_prefix _stringappend_84_) with
+ | Some (rs2,(existT _ _stringappend_85_ _)) =>
+ returnm (rs2, build_ex _stringappend_85_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_85_ _) =>
+ returnm ((match (string_drop _stringappend_84_ (build_ex _stringappend_85_)) with
| s_ =>
- Some (RTYPE (rs2,rs1,rd,op),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ ((RTYPE
+ (rs2, rs1, rd, op), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
and_boolM (returnm ((string_startswith _stringappend_0_ "l") : bool))
- (let _stringappend_129_ :=
- string_drop _stringappend_0_ (string_length "l") in
- match (size_mnemonic_matches_prefix _stringappend_129_) with
- | Some (_stringappend_130_,(existT _ _stringappend_131_ _)) =>
- let _stringappend_132_ :=
- string_drop _stringappend_129_ (build_ex _stringappend_131_) in
- match (maybe_u_matches_prefix _stringappend_132_) with
- | Some (_stringappend_133_,(existT _ _stringappend_134_ _)) =>
- let _stringappend_135_ :=
- string_drop _stringappend_132_ (build_ex _stringappend_134_) in
- match (maybe_aq_matches_prefix _stringappend_135_) with
- | Some (_stringappend_136_,(existT _ _stringappend_137_ _)) =>
- let _stringappend_138_ :=
- string_drop _stringappend_135_ (build_ex _stringappend_137_) in
- match (maybe_rl_matches_prefix _stringappend_138_) with
- | Some (_stringappend_139_,(existT _ _stringappend_140_ _)) =>
- let _stringappend_141_ :=
- string_drop _stringappend_138_ (build_ex _stringappend_140_) in
- match (spc_matches_prefix _stringappend_141_) with
- | Some (_stringappend_142_,(existT _ _stringappend_143_ _)) =>
- let _stringappend_144_ :=
- string_drop _stringappend_141_
- (build_ex
- _stringappend_143_) in
- match (reg_name_matches_prefix _stringappend_144_) with
- | Some (_stringappend_145_,(existT _ _stringappend_146_ _)) =>
- let _stringappend_147_ :=
- string_drop _stringappend_144_
- (build_ex
- _stringappend_146_) in
- sep_matches_prefix _stringappend_147_ >>= fun w__135 : option ((unit * {n : Z & ArithFact (n >=
+ (let _stringappend_87_ :=
+ string_drop _stringappend_0_ (build_ex (projT1 (string_length "l"))) in
+ match (size_mnemonic_matches_prefix _stringappend_87_) with
+ | Some (size,(existT _ _stringappend_88_ _)) =>
+ let _stringappend_89_ :=
+ string_drop _stringappend_87_ (build_ex _stringappend_88_) in
+ match (maybe_u_matches_prefix _stringappend_89_) with
+ | Some (is_unsigned,(existT _ _stringappend_90_ _)) =>
+ let _stringappend_91_ :=
+ string_drop _stringappend_89_ (build_ex _stringappend_90_) in
+ match (maybe_aq_matches_prefix _stringappend_91_) with
+ | Some (aq,(existT _ _stringappend_92_ _)) =>
+ let _stringappend_93_ :=
+ string_drop _stringappend_91_ (build_ex _stringappend_92_) in
+ match (maybe_rl_matches_prefix _stringappend_93_) with
+ | Some (rl,(existT _ _stringappend_94_ _)) =>
+ let _stringappend_95_ :=
+ string_drop _stringappend_93_ (build_ex _stringappend_94_) in
+ match (spc_matches_prefix _stringappend_95_) with
+ | Some (tt,(existT _ _stringappend_96_ _)) =>
+ let _stringappend_97_ :=
+ string_drop _stringappend_95_
+ (build_ex _stringappend_96_) in
+ match (reg_name_matches_prefix _stringappend_97_) with
+ | Some (rd,(existT _ _stringappend_98_ _)) =>
+ let _stringappend_99_ :=
+ string_drop _stringappend_97_
+ (build_ex _stringappend_98_) in
+ sep_matches_prefix _stringappend_99_ >>= fun w__135 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__135 with
- | Some
- (_stringappend_148_,(existT _ _stringappend_149_ _)) =>
- let _stringappend_150_ :=
- string_drop _stringappend_147_
- (build_ex
- _stringappend_149_) in
- match (reg_name_matches_prefix _stringappend_150_) with
- | Some
- (_stringappend_151_,(existT _ _stringappend_152_ _)) =>
- let _stringappend_153_ :=
- string_drop _stringappend_150_
- (build_ex
- _stringappend_152_) in
- sep_matches_prefix _stringappend_153_ >>= fun w__136 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_100_ _)) =>
+ let _stringappend_101_ :=
+ string_drop _stringappend_99_
+ (build_ex _stringappend_100_) in
+ match (reg_name_matches_prefix _stringappend_101_) with
+ | Some (rs1,(existT _ _stringappend_102_ _)) =>
+ let _stringappend_103_ :=
+ string_drop _stringappend_101_
+ (build_ex _stringappend_102_) in
+ sep_matches_prefix _stringappend_103_ >>= fun w__136 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__136 with
| Some
- (_stringappend_154_,(existT _ _stringappend_155_ _)) =>
- let _stringappend_156_ :=
- string_drop _stringappend_153_
- (build_ex
- _stringappend_155_) in
+ (tt,(existT _ _stringappend_104_ _)) =>
+ let _stringappend_105_ :=
+ string_drop _stringappend_103_
+ (build_ex _stringappend_104_) in
if ((match (hex_bits_12_matches_prefix
- _stringappend_156_) with
+ _stringappend_105_) with
| Some
- (_stringappend_157_,(existT _ _stringappend_158_ _)) =>
+ (imm,(existT _ _stringappend_106_ _)) =>
match (string_drop
- _stringappend_156_
- (build_ex
- _stringappend_158_)) with
+ _stringappend_105_
+ (build_ex _stringappend_106_)) with
| s_ => true
end
| None => false
@@ -24644,189 +23031,145 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__145 : bool =>
(if (w__145) then
- let _stringappend_129_ :=
- string_drop _stringappend_0_ (string_length "l") in
- match (size_mnemonic_matches_prefix _stringappend_129_) with
- | Some (_stringappend_130_,(existT _ _stringappend_131_ _)) =>
- returnm ((_stringappend_130_, build_ex _stringappend_131_)
- : (word_width * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((word_width * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__147 : (word_width * {n : Z & ArithFact (n >= 0)}) =>
- let '(size, existT _ _stringappend_131_ _) :=
- w__147
- : (word_width * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_132_ :=
- string_drop _stringappend_129_ (build_ex _stringappend_131_) in
- match (maybe_u_matches_prefix _stringappend_132_) with
- | Some (_stringappend_133_,(existT _ _stringappend_134_ _)) =>
- returnm ((_stringappend_133_, build_ex _stringappend_134_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__149 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(is_unsigned, existT _ _stringappend_134_ _) :=
- w__149
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_135_ :=
- string_drop _stringappend_132_ (build_ex _stringappend_134_) in
- match (maybe_aq_matches_prefix _stringappend_135_) with
- | Some (_stringappend_136_,(existT _ _stringappend_137_ _)) =>
- returnm ((_stringappend_136_, build_ex _stringappend_137_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__151 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(aq, existT _ _stringappend_137_ _) :=
- w__151
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_138_ :=
- string_drop _stringappend_135_ (build_ex _stringappend_137_) in
- match (maybe_rl_matches_prefix _stringappend_138_) with
- | Some (_stringappend_139_,(existT _ _stringappend_140_ _)) =>
- returnm ((_stringappend_139_, build_ex _stringappend_140_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__153 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(rl, existT _ _stringappend_140_ _) :=
- w__153
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_141_ :=
- string_drop _stringappend_138_ (build_ex _stringappend_140_) in
- match (spc_matches_prefix _stringappend_141_) with
- | Some (_stringappend_142_,(existT _ _stringappend_143_ _)) =>
- returnm ((_stringappend_142_, build_ex _stringappend_143_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__155 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_143_ _) :=
- w__155
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_144_ :=
- string_drop _stringappend_141_ (build_ex _stringappend_143_) in
- match (reg_name_matches_prefix _stringappend_144_) with
- | Some (_stringappend_145_,(existT _ _stringappend_146_ _)) =>
- returnm ((_stringappend_145_, build_ex _stringappend_146_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__157 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_146_ _) :=
- w__157
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_147_ :=
- string_drop _stringappend_144_ (build_ex _stringappend_146_) in
- sep_matches_prefix _stringappend_147_ >>= fun w__158 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_87_ :=
+ string_drop _stringappend_0_ (build_ex (projT1 (string_length "l"))) in
+ (match (size_mnemonic_matches_prefix _stringappend_87_) with
+ | Some (size,(existT _ _stringappend_88_ _)) =>
+ returnm (size, build_ex _stringappend_88_)
+ | _ => exit tt : M ((word_width * {n : Z & ArithFact (n >= 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >= 0)}))) >>= fun '(size, existT _ _stringappend_88_ _) =>
+ let _stringappend_89_ :=
+ string_drop _stringappend_87_ (build_ex _stringappend_88_) in
+ (match (maybe_u_matches_prefix _stringappend_89_) with
+ | Some (is_unsigned,(existT _ _stringappend_90_ _)) =>
+ returnm (is_unsigned, build_ex _stringappend_90_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(is_unsigned, existT _ _stringappend_90_ _) =>
+ let _stringappend_91_ :=
+ string_drop _stringappend_89_ (build_ex _stringappend_90_) in
+ (match (maybe_aq_matches_prefix _stringappend_91_) with
+ | Some (aq,(existT _ _stringappend_92_ _)) =>
+ returnm (aq, build_ex _stringappend_92_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(aq, existT _ _stringappend_92_ _) =>
+ let _stringappend_93_ :=
+ string_drop _stringappend_91_ (build_ex _stringappend_92_) in
+ (match (maybe_rl_matches_prefix _stringappend_93_) with
+ | Some (rl,(existT _ _stringappend_94_ _)) =>
+ returnm (rl, build_ex _stringappend_94_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rl, existT _ _stringappend_94_ _) =>
+ let _stringappend_95_ :=
+ string_drop _stringappend_93_ (build_ex _stringappend_94_) in
+ (match (spc_matches_prefix _stringappend_95_) with
+ | Some (tt,(existT _ _stringappend_96_ _)) =>
+ returnm (tt, build_ex _stringappend_96_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_96_ _) =>
+ let _stringappend_97_ :=
+ string_drop _stringappend_95_ (build_ex _stringappend_96_) in
+ (match (reg_name_matches_prefix _stringappend_97_) with
+ | Some (rd,(existT _ _stringappend_98_ _)) =>
+ returnm (rd, build_ex _stringappend_98_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_98_ _) =>
+ let _stringappend_99_ :=
+ string_drop _stringappend_97_ (build_ex _stringappend_98_) in
+ sep_matches_prefix _stringappend_99_ >>= fun w__158 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__158 with
- | Some (_stringappend_148_,(existT _ _stringappend_149_ _)) =>
- returnm ((_stringappend_148_, build_ex _stringappend_149_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__160 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_149_ _) :=
- w__160
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_150_ :=
- string_drop _stringappend_147_ (build_ex _stringappend_149_) in
- match (reg_name_matches_prefix _stringappend_150_) with
- | Some (_stringappend_151_,(existT _ _stringappend_152_ _)) =>
- returnm ((_stringappend_151_, build_ex _stringappend_152_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__162 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_152_ _) :=
- w__162
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_153_ :=
- string_drop _stringappend_150_ (build_ex _stringappend_152_) in
- sep_matches_prefix _stringappend_153_ >>= fun w__163 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__158 with
+ | Some (tt,(existT _ _stringappend_100_ _)) =>
+ returnm (tt, build_ex _stringappend_100_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_100_ _) =>
+ let _stringappend_101_ :=
+ string_drop _stringappend_99_ (build_ex _stringappend_100_) in
+ (match (reg_name_matches_prefix _stringappend_101_) with
+ | Some (rs1,(existT _ _stringappend_102_ _)) =>
+ returnm (rs1, build_ex _stringappend_102_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_102_ _) =>
+ let _stringappend_103_ :=
+ string_drop _stringappend_101_ (build_ex _stringappend_102_) in
+ sep_matches_prefix _stringappend_103_ >>= fun w__163 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__163 with
- | Some (_stringappend_154_,(existT _ _stringappend_155_ _)) =>
- returnm ((_stringappend_154_, build_ex _stringappend_155_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__165 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_155_ _) :=
- w__165
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_156_ :=
- string_drop _stringappend_153_ (build_ex _stringappend_155_) in
- match (hex_bits_12_matches_prefix _stringappend_156_) with
- | Some (_stringappend_157_,(existT _ _stringappend_158_ _)) =>
- returnm ((_stringappend_157_, build_ex _stringappend_158_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__167 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_158_ _) :=
- w__167
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_156_
- (build_ex
- _stringappend_158_)) with
+ (match w__163 with
+ | Some (tt,(existT _ _stringappend_104_ _)) =>
+ returnm (tt, build_ex _stringappend_104_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_104_ _) =>
+ let _stringappend_105_ :=
+ string_drop _stringappend_103_ (build_ex _stringappend_104_) in
+ (match (hex_bits_12_matches_prefix _stringappend_105_) with
+ | Some (imm,(existT _ _stringappend_106_ _)) =>
+ returnm (imm, build_ex _stringappend_106_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_106_ _) =>
+ returnm ((match (string_drop _stringappend_105_
+ (build_ex _stringappend_106_)) with
| s_ =>
- Some (LOAD (imm,rs1,rd,is_unsigned,size,aq,rl),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ ((LOAD
+ (imm, rs1, rd, is_unsigned, size, aq, rl), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
and_boolM (returnm ((string_startswith _stringappend_0_ "s") : bool))
- (let _stringappend_160_ :=
- string_drop _stringappend_0_ (string_length "s") in
- match (size_mnemonic_matches_prefix _stringappend_160_) with
- | Some (_stringappend_161_,(existT _ _stringappend_162_ _)) =>
- let _stringappend_163_ :=
- string_drop _stringappend_160_ (build_ex _stringappend_162_) in
- match (maybe_aq_matches_prefix _stringappend_163_) with
- | Some (_stringappend_164_,(existT _ _stringappend_165_ _)) =>
- let _stringappend_166_ :=
- string_drop _stringappend_163_ (build_ex _stringappend_165_) in
- match (maybe_rl_matches_prefix _stringappend_166_) with
- | Some (_stringappend_167_,(existT _ _stringappend_168_ _)) =>
- let _stringappend_169_ :=
- string_drop _stringappend_166_ (build_ex _stringappend_168_) in
- match (spc_matches_prefix _stringappend_169_) with
- | Some (_stringappend_170_,(existT _ _stringappend_171_ _)) =>
- let _stringappend_172_ :=
- string_drop _stringappend_169_
- (build_ex
- _stringappend_171_) in
- match (reg_name_matches_prefix _stringappend_172_) with
- | Some (_stringappend_173_,(existT _ _stringappend_174_ _)) =>
- let _stringappend_175_ :=
- string_drop _stringappend_172_
- (build_ex
- _stringappend_174_) in
- sep_matches_prefix _stringappend_175_ >>= fun w__168 : option ((unit * {n : Z & ArithFact (n >=
+ (let _stringappend_108_ :=
+ string_drop _stringappend_0_ (build_ex (projT1 (string_length "s"))) in
+ match (size_mnemonic_matches_prefix _stringappend_108_) with
+ | Some (size,(existT _ _stringappend_109_ _)) =>
+ let _stringappend_110_ :=
+ string_drop _stringappend_108_ (build_ex _stringappend_109_) in
+ match (maybe_aq_matches_prefix _stringappend_110_) with
+ | Some (aq,(existT _ _stringappend_111_ _)) =>
+ let _stringappend_112_ :=
+ string_drop _stringappend_110_ (build_ex _stringappend_111_) in
+ match (maybe_rl_matches_prefix _stringappend_112_) with
+ | Some (rl,(existT _ _stringappend_113_ _)) =>
+ let _stringappend_114_ :=
+ string_drop _stringappend_112_ (build_ex _stringappend_113_) in
+ match (spc_matches_prefix _stringappend_114_) with
+ | Some (tt,(existT _ _stringappend_115_ _)) =>
+ let _stringappend_116_ :=
+ string_drop _stringappend_114_
+ (build_ex _stringappend_115_) in
+ match (reg_name_matches_prefix _stringappend_116_) with
+ | Some (rd,(existT _ _stringappend_117_ _)) =>
+ let _stringappend_118_ :=
+ string_drop _stringappend_116_
+ (build_ex _stringappend_117_) in
+ sep_matches_prefix _stringappend_118_ >>= fun w__168 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__168 with
- | Some
- (_stringappend_176_,(existT _ _stringappend_177_ _)) =>
- let _stringappend_178_ :=
- string_drop _stringappend_175_
- (build_ex
- _stringappend_177_) in
- match (reg_name_matches_prefix _stringappend_178_) with
- | Some
- (_stringappend_179_,(existT _ _stringappend_180_ _)) =>
- let _stringappend_181_ :=
- string_drop _stringappend_178_
- (build_ex
- _stringappend_180_) in
- sep_matches_prefix _stringappend_181_ >>= fun w__169 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_119_ _)) =>
+ let _stringappend_120_ :=
+ string_drop _stringappend_118_
+ (build_ex _stringappend_119_) in
+ match (reg_name_matches_prefix _stringappend_120_) with
+ | Some (rs1,(existT _ _stringappend_121_ _)) =>
+ let _stringappend_122_ :=
+ string_drop _stringappend_120_
+ (build_ex _stringappend_121_) in
+ sep_matches_prefix _stringappend_122_ >>= fun w__169 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__169 with
| Some
- (_stringappend_182_,(existT _ _stringappend_183_ _)) =>
- let _stringappend_184_ :=
- string_drop _stringappend_181_
- (build_ex
- _stringappend_183_) in
+ (tt,(existT _ _stringappend_123_ _)) =>
+ let _stringappend_124_ :=
+ string_drop _stringappend_122_
+ (build_ex _stringappend_123_) in
if ((match (hex_bits_12_matches_prefix
- _stringappend_184_) with
+ _stringappend_124_) with
| Some
- (_stringappend_185_,(existT _ _stringappend_186_ _)) =>
+ (imm,(existT _ _stringappend_125_ _)) =>
match (string_drop
- _stringappend_184_
- (build_ex
- _stringappend_186_)) with
+ _stringappend_124_
+ (build_ex _stringappend_125_)) with
| s_ => true
end
| None => false
@@ -24874,162 +23217,127 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__177 : bool =>
(if (w__177) then
- let _stringappend_160_ :=
- string_drop _stringappend_0_ (string_length "s") in
- match (size_mnemonic_matches_prefix _stringappend_160_) with
- | Some (_stringappend_161_,(existT _ _stringappend_162_ _)) =>
- returnm ((_stringappend_161_, build_ex _stringappend_162_)
- : (word_width * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((word_width * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__179 : (word_width * {n : Z & ArithFact (n >= 0)}) =>
- let '(size, existT _ _stringappend_162_ _) :=
- w__179
- : (word_width * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_163_ :=
- string_drop _stringappend_160_ (build_ex _stringappend_162_) in
- match (maybe_aq_matches_prefix _stringappend_163_) with
- | Some (_stringappend_164_,(existT _ _stringappend_165_ _)) =>
- returnm ((_stringappend_164_, build_ex _stringappend_165_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__181 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(aq, existT _ _stringappend_165_ _) :=
- w__181
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_166_ :=
- string_drop _stringappend_163_ (build_ex _stringappend_165_) in
- match (maybe_rl_matches_prefix _stringappend_166_) with
- | Some (_stringappend_167_,(existT _ _stringappend_168_ _)) =>
- returnm ((_stringappend_167_, build_ex _stringappend_168_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__183 : (bool * {n : Z & ArithFact (n >= 0)}) =>
- let '(rl, existT _ _stringappend_168_ _) :=
- w__183
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_169_ :=
- string_drop _stringappend_166_ (build_ex _stringappend_168_) in
- match (spc_matches_prefix _stringappend_169_) with
- | Some (_stringappend_170_,(existT _ _stringappend_171_ _)) =>
- returnm ((_stringappend_170_, build_ex _stringappend_171_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__185 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_171_ _) :=
- w__185
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_172_ :=
- string_drop _stringappend_169_ (build_ex _stringappend_171_) in
- match (reg_name_matches_prefix _stringappend_172_) with
- | Some (_stringappend_173_,(existT _ _stringappend_174_ _)) =>
- returnm ((_stringappend_173_, build_ex _stringappend_174_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__187 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_174_ _) :=
- w__187
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_175_ :=
- string_drop _stringappend_172_ (build_ex _stringappend_174_) in
- sep_matches_prefix _stringappend_175_ >>= fun w__188 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_108_ :=
+ string_drop _stringappend_0_ (build_ex (projT1 (string_length "s"))) in
+ (match (size_mnemonic_matches_prefix _stringappend_108_) with
+ | Some (size,(existT _ _stringappend_109_ _)) =>
+ returnm (size, build_ex _stringappend_109_)
+ | _ => exit tt : M ((word_width * {n : Z & ArithFact (n >= 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >= 0)}))) >>= fun '(size, existT _ _stringappend_109_ _) =>
+ let _stringappend_110_ :=
+ string_drop _stringappend_108_ (build_ex _stringappend_109_) in
+ (match (maybe_aq_matches_prefix _stringappend_110_) with
+ | Some (aq,(existT _ _stringappend_111_ _)) =>
+ returnm (aq, build_ex _stringappend_111_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(aq, existT _ _stringappend_111_ _) =>
+ let _stringappend_112_ :=
+ string_drop _stringappend_110_ (build_ex _stringappend_111_) in
+ (match (maybe_rl_matches_prefix _stringappend_112_) with
+ | Some (rl,(existT _ _stringappend_113_ _)) =>
+ returnm (rl, build_ex _stringappend_113_)
+ | _ => exit tt : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rl, existT _ _stringappend_113_ _) =>
+ let _stringappend_114_ :=
+ string_drop _stringappend_112_ (build_ex _stringappend_113_) in
+ (match (spc_matches_prefix _stringappend_114_) with
+ | Some (tt,(existT _ _stringappend_115_ _)) =>
+ returnm (tt, build_ex _stringappend_115_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_115_ _) =>
+ let _stringappend_116_ :=
+ string_drop _stringappend_114_ (build_ex _stringappend_115_) in
+ (match (reg_name_matches_prefix _stringappend_116_) with
+ | Some (rd,(existT _ _stringappend_117_ _)) =>
+ returnm (rd, build_ex _stringappend_117_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_117_ _) =>
+ let _stringappend_118_ :=
+ string_drop _stringappend_116_ (build_ex _stringappend_117_) in
+ sep_matches_prefix _stringappend_118_ >>= fun w__188 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__188 with
- | Some (_stringappend_176_,(existT _ _stringappend_177_ _)) =>
- returnm ((_stringappend_176_, build_ex _stringappend_177_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__190 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_177_ _) :=
- w__190
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_178_ :=
- string_drop _stringappend_175_ (build_ex _stringappend_177_) in
- match (reg_name_matches_prefix _stringappend_178_) with
- | Some (_stringappend_179_,(existT _ _stringappend_180_ _)) =>
- returnm ((_stringappend_179_, build_ex _stringappend_180_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__192 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_180_ _) :=
- w__192
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_181_ :=
- string_drop _stringappend_178_ (build_ex _stringappend_180_) in
- sep_matches_prefix _stringappend_181_ >>= fun w__193 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__188 with
+ | Some (tt,(existT _ _stringappend_119_ _)) =>
+ returnm (tt, build_ex _stringappend_119_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_119_ _) =>
+ let _stringappend_120_ :=
+ string_drop _stringappend_118_ (build_ex _stringappend_119_) in
+ (match (reg_name_matches_prefix _stringappend_120_) with
+ | Some (rs1,(existT _ _stringappend_121_ _)) =>
+ returnm (rs1, build_ex _stringappend_121_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_121_ _) =>
+ let _stringappend_122_ :=
+ string_drop _stringappend_120_ (build_ex _stringappend_121_) in
+ sep_matches_prefix _stringappend_122_ >>= fun w__193 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__193 with
- | Some (_stringappend_182_,(existT _ _stringappend_183_ _)) =>
- returnm ((_stringappend_182_, build_ex _stringappend_183_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__195 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_183_ _) :=
- w__195
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_184_ :=
- string_drop _stringappend_181_ (build_ex _stringappend_183_) in
- match (hex_bits_12_matches_prefix _stringappend_184_) with
- | Some (_stringappend_185_,(existT _ _stringappend_186_ _)) =>
- returnm ((_stringappend_185_, build_ex _stringappend_186_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__197 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_186_ _) :=
- w__197
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_184_
- (build_ex
- _stringappend_186_)) with
+ (match w__193 with
+ | Some (tt,(existT _ _stringappend_123_ _)) =>
+ returnm (tt, build_ex _stringappend_123_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_123_ _) =>
+ let _stringappend_124_ :=
+ string_drop _stringappend_122_ (build_ex _stringappend_123_) in
+ (match (hex_bits_12_matches_prefix _stringappend_124_) with
+ | Some (imm,(existT _ _stringappend_125_ _)) =>
+ returnm (imm, build_ex _stringappend_125_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_125_ _) =>
+ returnm ((match (string_drop _stringappend_124_
+ (build_ex _stringappend_125_)) with
| s_ =>
- Some (STORE (imm,rs1,rd,size,aq,rl),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ ((STORE
+ (imm, rs1, rd, size, aq, rl), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
and_boolM
(returnm ((string_startswith _stringappend_0_ "addiw")
: bool))
- (let _stringappend_188_ :=
- string_drop _stringappend_0_ (string_length "addiw") in
- match (spc_matches_prefix _stringappend_188_) with
- | Some (_stringappend_189_,(existT _ _stringappend_190_ _)) =>
- let _stringappend_191_ :=
- string_drop _stringappend_188_ (build_ex _stringappend_190_) in
- match (reg_name_matches_prefix _stringappend_191_) with
- | Some (_stringappend_192_,(existT _ _stringappend_193_ _)) =>
- let _stringappend_194_ :=
- string_drop _stringappend_191_ (build_ex _stringappend_193_) in
- sep_matches_prefix _stringappend_194_ >>= fun w__198 : option ((unit * {n : Z & ArithFact (n >=
+ (let _stringappend_127_ :=
+ string_drop _stringappend_0_
+ (build_ex (projT1 (string_length "addiw"))) in
+ match (spc_matches_prefix _stringappend_127_) with
+ | Some (tt,(existT _ _stringappend_128_ _)) =>
+ let _stringappend_129_ :=
+ string_drop _stringappend_127_ (build_ex _stringappend_128_) in
+ match (reg_name_matches_prefix _stringappend_129_) with
+ | Some (rd,(existT _ _stringappend_130_ _)) =>
+ let _stringappend_131_ :=
+ string_drop _stringappend_129_ (build_ex _stringappend_130_) in
+ sep_matches_prefix _stringappend_131_ >>= fun w__198 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__198 with
- | Some (_stringappend_195_,(existT _ _stringappend_196_ _)) =>
- let _stringappend_197_ :=
- string_drop _stringappend_194_
- (build_ex
- _stringappend_196_) in
- match (reg_name_matches_prefix _stringappend_197_) with
- | Some (_stringappend_198_,(existT _ _stringappend_199_ _)) =>
- let _stringappend_200_ :=
- string_drop _stringappend_197_
- (build_ex
- _stringappend_199_) in
- sep_matches_prefix _stringappend_200_ >>= fun w__199 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_132_ _)) =>
+ let _stringappend_133_ :=
+ string_drop _stringappend_131_
+ (build_ex _stringappend_132_) in
+ match (reg_name_matches_prefix _stringappend_133_) with
+ | Some (rs1,(existT _ _stringappend_134_ _)) =>
+ let _stringappend_135_ :=
+ string_drop _stringappend_133_
+ (build_ex _stringappend_134_) in
+ sep_matches_prefix _stringappend_135_ >>= fun w__199 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__199 with
| Some
- (_stringappend_201_,(existT _ _stringappend_202_ _)) =>
- let _stringappend_203_ :=
- string_drop _stringappend_200_
- (build_ex
- _stringappend_202_) in
+ (tt,(existT _ _stringappend_136_ _)) =>
+ let _stringappend_137_ :=
+ string_drop _stringappend_135_
+ (build_ex _stringappend_136_) in
if ((match (hex_bits_12_matches_prefix
- _stringappend_203_) with
+ _stringappend_137_) with
| Some
- (_stringappend_204_,(existT _ _stringappend_205_ _)) =>
+ (imm,(existT _ _stringappend_138_ _)) =>
match (string_drop
- _stringappend_203_
- (build_ex
- _stringappend_205_)) with
+ _stringappend_137_
+ (build_ex _stringappend_138_)) with
| s_ => true
end
| None => false
@@ -25062,134 +23370,107 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__204 : bool =>
(if (w__204) then
- let _stringappend_188_ :=
- string_drop _stringappend_0_ (string_length "addiw") in
- match (spc_matches_prefix _stringappend_188_) with
- | Some (_stringappend_189_,(existT _ _stringappend_190_ _)) =>
- returnm ((_stringappend_189_, build_ex _stringappend_190_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__206 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_190_ _) :=
- w__206
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_191_ :=
- string_drop _stringappend_188_ (build_ex _stringappend_190_) in
- match (reg_name_matches_prefix _stringappend_191_) with
- | Some (_stringappend_192_,(existT _ _stringappend_193_ _)) =>
- returnm ((_stringappend_192_, build_ex _stringappend_193_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__208 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_193_ _) :=
- w__208
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_194_ :=
- string_drop _stringappend_191_ (build_ex _stringappend_193_) in
- sep_matches_prefix _stringappend_194_ >>= fun w__209 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_127_ :=
+ string_drop _stringappend_0_
+ (build_ex (projT1 (string_length "addiw"))) in
+ (match (spc_matches_prefix _stringappend_127_) with
+ | Some (tt,(existT _ _stringappend_128_ _)) =>
+ returnm (tt, build_ex _stringappend_128_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_128_ _) =>
+ let _stringappend_129_ :=
+ string_drop _stringappend_127_ (build_ex _stringappend_128_) in
+ (match (reg_name_matches_prefix _stringappend_129_) with
+ | Some (rd,(existT _ _stringappend_130_ _)) =>
+ returnm (rd, build_ex _stringappend_130_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_130_ _) =>
+ let _stringappend_131_ :=
+ string_drop _stringappend_129_ (build_ex _stringappend_130_) in
+ sep_matches_prefix _stringappend_131_ >>= fun w__209 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__209 with
- | Some (_stringappend_195_,(existT _ _stringappend_196_ _)) =>
- returnm ((_stringappend_195_, build_ex _stringappend_196_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__211 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_196_ _) :=
- w__211
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_197_ :=
- string_drop _stringappend_194_ (build_ex _stringappend_196_) in
- match (reg_name_matches_prefix _stringappend_197_) with
- | Some (_stringappend_198_,(existT _ _stringappend_199_ _)) =>
- returnm ((_stringappend_198_, build_ex _stringappend_199_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__213 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_199_ _) :=
- w__213
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_200_ :=
- string_drop _stringappend_197_ (build_ex _stringappend_199_) in
- sep_matches_prefix _stringappend_200_ >>= fun w__214 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__209 with
+ | Some (tt,(existT _ _stringappend_132_ _)) =>
+ returnm (tt, build_ex _stringappend_132_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_132_ _) =>
+ let _stringappend_133_ :=
+ string_drop _stringappend_131_ (build_ex _stringappend_132_) in
+ (match (reg_name_matches_prefix _stringappend_133_) with
+ | Some (rs1,(existT _ _stringappend_134_ _)) =>
+ returnm (rs1, build_ex _stringappend_134_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_134_ _) =>
+ let _stringappend_135_ :=
+ string_drop _stringappend_133_ (build_ex _stringappend_134_) in
+ sep_matches_prefix _stringappend_135_ >>= fun w__214 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__214 with
- | Some (_stringappend_201_,(existT _ _stringappend_202_ _)) =>
- returnm ((_stringappend_201_, build_ex _stringappend_202_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__216 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_202_ _) :=
- w__216
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_203_ :=
- string_drop _stringappend_200_ (build_ex _stringappend_202_) in
- match (hex_bits_12_matches_prefix _stringappend_203_) with
- | Some (_stringappend_204_,(existT _ _stringappend_205_ _)) =>
- returnm ((_stringappend_204_, build_ex _stringappend_205_)
- : (mword 12 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__218 : (mword 12 * {n : Z & ArithFact (n >= 0)}) =>
- let '(imm, existT _ _stringappend_205_ _) :=
- w__218
- : (mword 12 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_203_
- (build_ex
- _stringappend_205_)) with
+ (match w__214 with
+ | Some (tt,(existT _ _stringappend_136_ _)) =>
+ returnm (tt, build_ex _stringappend_136_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_136_ _) =>
+ let _stringappend_137_ :=
+ string_drop _stringappend_135_ (build_ex _stringappend_136_) in
+ (match (hex_bits_12_matches_prefix _stringappend_137_) with
+ | Some (imm,(existT _ _stringappend_138_ _)) =>
+ returnm (imm, build_ex _stringappend_138_)
+ | _ => exit tt : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(imm, existT _ _stringappend_138_ _) =>
+ returnm ((match (string_drop _stringappend_137_
+ (build_ex _stringappend_138_)) with
| s_ =>
- Some (ADDIW (imm,rs1,rd),
- sub_nat (string_length arg_) (string_length s_))
+ Some
+ ((ADDIW
+ (imm, rs1, rd), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
match (shiftw_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_207_,(existT _ _stringappend_208_ _)) =>
- let _stringappend_209_ :=
- string_drop _stringappend_0_ (build_ex _stringappend_208_) in
- match (spc_matches_prefix _stringappend_209_) with
- | Some (_stringappend_210_,(existT _ _stringappend_211_ _)) =>
- let _stringappend_212_ :=
- string_drop _stringappend_209_
- (build_ex
- _stringappend_211_) in
- match (reg_name_matches_prefix _stringappend_212_) with
- | Some (_stringappend_213_,(existT _ _stringappend_214_ _)) =>
- let _stringappend_215_ :=
- string_drop _stringappend_212_
- (build_ex
- _stringappend_214_) in
- sep_matches_prefix _stringappend_215_ >>= fun w__219 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op,(existT _ _stringappend_140_ _)) =>
+ let _stringappend_141_ :=
+ string_drop _stringappend_0_ (build_ex _stringappend_140_) in
+ match (spc_matches_prefix _stringappend_141_) with
+ | Some (tt,(existT _ _stringappend_142_ _)) =>
+ let _stringappend_143_ :=
+ string_drop _stringappend_141_
+ (build_ex _stringappend_142_) in
+ match (reg_name_matches_prefix _stringappend_143_) with
+ | Some (rd,(existT _ _stringappend_144_ _)) =>
+ let _stringappend_145_ :=
+ string_drop _stringappend_143_
+ (build_ex _stringappend_144_) in
+ sep_matches_prefix _stringappend_145_ >>= fun w__219 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__219 with
- | Some
- (_stringappend_216_,(existT _ _stringappend_217_ _)) =>
- let _stringappend_218_ :=
- string_drop _stringappend_215_
- (build_ex
- _stringappend_217_) in
- match (reg_name_matches_prefix _stringappend_218_) with
- | Some
- (_stringappend_219_,(existT _ _stringappend_220_ _)) =>
- let _stringappend_221_ :=
- string_drop _stringappend_218_
- (build_ex
- _stringappend_220_) in
- sep_matches_prefix _stringappend_221_ >>= fun w__220 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_146_ _)) =>
+ let _stringappend_147_ :=
+ string_drop _stringappend_145_
+ (build_ex _stringappend_146_) in
+ match (reg_name_matches_prefix _stringappend_147_) with
+ | Some (rs1,(existT _ _stringappend_148_ _)) =>
+ let _stringappend_149_ :=
+ string_drop _stringappend_147_
+ (build_ex _stringappend_148_) in
+ sep_matches_prefix _stringappend_149_ >>= fun w__220 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__220 with
| Some
- (_stringappend_222_,(existT _ _stringappend_223_ _)) =>
- let _stringappend_224_ :=
- string_drop _stringappend_221_
- (build_ex
- _stringappend_223_) in
+ (tt,(existT _ _stringappend_150_ _)) =>
+ let _stringappend_151_ :=
+ string_drop _stringappend_149_
+ (build_ex _stringappend_150_) in
if ((match (hex_bits_5_matches_prefix
- _stringappend_224_) with
+ _stringappend_151_) with
| Some
- (_stringappend_225_,(existT _ _stringappend_226_ _)) =>
+ (shamt,(existT _ _stringappend_152_ _)) =>
match (string_drop
- _stringappend_224_
- (build_ex
- _stringappend_226_)) with
+ _stringappend_151_
+ (build_ex _stringappend_152_)) with
| s_ => true
end
| None => false
@@ -25224,146 +23505,112 @@ Definition assembly_matches_prefix (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__225 : bool =>
(if (w__225) then
- match (shiftw_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_207_,(existT _ _stringappend_208_ _)) =>
- returnm ((_stringappend_207_, build_ex _stringappend_208_)
- : (sop * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((sop * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__227 : (sop * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_208_ _) :=
- w__227
- : (sop * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_209_ :=
- string_drop _stringappend_0_ (build_ex _stringappend_208_) in
- match (spc_matches_prefix _stringappend_209_) with
- | Some (_stringappend_210_,(existT _ _stringappend_211_ _)) =>
- returnm ((_stringappend_210_, build_ex _stringappend_211_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__229 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_211_ _) :=
- w__229
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_212_ :=
- string_drop _stringappend_209_ (build_ex _stringappend_211_) in
- match (reg_name_matches_prefix _stringappend_212_) with
- | Some (_stringappend_213_,(existT _ _stringappend_214_ _)) =>
- returnm ((_stringappend_213_, build_ex _stringappend_214_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__231 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_214_ _) :=
- w__231
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_215_ :=
- string_drop _stringappend_212_ (build_ex _stringappend_214_) in
- sep_matches_prefix _stringappend_215_ >>= fun w__232 : option ((unit * {n : Z & ArithFact (n >=
+ (match (shiftw_mnemonic_matches_prefix _stringappend_0_) with
+ | Some (op,(existT _ _stringappend_140_ _)) =>
+ returnm (op, build_ex _stringappend_140_)
+ | _ => exit tt : M ((sop * {n : Z & ArithFact (n >= 0)}))
+ end : M ((sop * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_140_ _) =>
+ let _stringappend_141_ :=
+ string_drop _stringappend_0_ (build_ex _stringappend_140_) in
+ (match (spc_matches_prefix _stringappend_141_) with
+ | Some (tt,(existT _ _stringappend_142_ _)) =>
+ returnm (tt, build_ex _stringappend_142_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_142_ _) =>
+ let _stringappend_143_ :=
+ string_drop _stringappend_141_ (build_ex _stringappend_142_) in
+ (match (reg_name_matches_prefix _stringappend_143_) with
+ | Some (rd,(existT _ _stringappend_144_ _)) =>
+ returnm (rd, build_ex _stringappend_144_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_144_ _) =>
+ let _stringappend_145_ :=
+ string_drop _stringappend_143_ (build_ex _stringappend_144_) in
+ sep_matches_prefix _stringappend_145_ >>= fun w__232 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__232 with
- | Some (_stringappend_216_,(existT _ _stringappend_217_ _)) =>
- returnm ((_stringappend_216_, build_ex _stringappend_217_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__234 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_217_ _) :=
- w__234
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_218_ :=
- string_drop _stringappend_215_ (build_ex _stringappend_217_) in
- match (reg_name_matches_prefix _stringappend_218_) with
- | Some (_stringappend_219_,(existT _ _stringappend_220_ _)) =>
- returnm ((_stringappend_219_, build_ex _stringappend_220_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__236 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_220_ _) :=
- w__236
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_221_ :=
- string_drop _stringappend_218_ (build_ex _stringappend_220_) in
- sep_matches_prefix _stringappend_221_ >>= fun w__237 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__232 with
+ | Some (tt,(existT _ _stringappend_146_ _)) =>
+ returnm (tt, build_ex _stringappend_146_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_146_ _) =>
+ let _stringappend_147_ :=
+ string_drop _stringappend_145_ (build_ex _stringappend_146_) in
+ (match (reg_name_matches_prefix _stringappend_147_) with
+ | Some (rs1,(existT _ _stringappend_148_ _)) =>
+ returnm (rs1, build_ex _stringappend_148_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_148_ _) =>
+ let _stringappend_149_ :=
+ string_drop _stringappend_147_ (build_ex _stringappend_148_) in
+ sep_matches_prefix _stringappend_149_ >>= fun w__237 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__237 with
- | Some (_stringappend_222_,(existT _ _stringappend_223_ _)) =>
- returnm ((_stringappend_222_, build_ex _stringappend_223_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__239 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_223_ _) :=
- w__239
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_224_ :=
- string_drop _stringappend_221_ (build_ex _stringappend_223_) in
- match (hex_bits_5_matches_prefix _stringappend_224_) with
- | Some (_stringappend_225_,(existT _ _stringappend_226_ _)) =>
- returnm ((_stringappend_225_, build_ex _stringappend_226_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__241 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(shamt, existT _ _stringappend_226_ _) :=
- w__241
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_224_
- (build_ex
- _stringappend_226_)) with
+ (match w__237 with
+ | Some (tt,(existT _ _stringappend_150_ _)) =>
+ returnm (tt, build_ex _stringappend_150_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_150_ _) =>
+ let _stringappend_151_ :=
+ string_drop _stringappend_149_ (build_ex _stringappend_150_) in
+ (match (hex_bits_5_matches_prefix _stringappend_151_) with
+ | Some (shamt,(existT _ _stringappend_152_ _)) =>
+ returnm (shamt, build_ex _stringappend_152_)
+ | _ => exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(shamt, existT _ _stringappend_152_ _) =>
+ returnm ((match (string_drop _stringappend_151_
+ (build_ex _stringappend_152_)) with
| s_ =>
- Some (SHIFTW (shamt,rs1,rd,op),
- sub_nat (string_length arg_)
- (string_length s_))
+ Some
+ ((SHIFTW
+ (shamt, rs1, rd, op), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
match (rtypew_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_228_,(existT _ _stringappend_229_ _)) =>
- let _stringappend_230_ :=
- string_drop _stringappend_0_ (build_ex _stringappend_229_) in
- match (spc_matches_prefix _stringappend_230_) with
- | Some (_stringappend_231_,(existT _ _stringappend_232_ _)) =>
- let _stringappend_233_ :=
- string_drop _stringappend_230_
- (build_ex
- _stringappend_232_) in
- match (reg_name_matches_prefix _stringappend_233_) with
- | Some
- (_stringappend_234_,(existT _ _stringappend_235_ _)) =>
- let _stringappend_236_ :=
- string_drop _stringappend_233_
- (build_ex
- _stringappend_235_) in
- sep_matches_prefix _stringappend_236_ >>= fun w__242 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op,(existT _ _stringappend_154_ _)) =>
+ let _stringappend_155_ :=
+ string_drop _stringappend_0_ (build_ex _stringappend_154_) in
+ match (spc_matches_prefix _stringappend_155_) with
+ | Some (tt,(existT _ _stringappend_156_ _)) =>
+ let _stringappend_157_ :=
+ string_drop _stringappend_155_
+ (build_ex _stringappend_156_) in
+ match (reg_name_matches_prefix _stringappend_157_) with
+ | Some (rd,(existT _ _stringappend_158_ _)) =>
+ let _stringappend_159_ :=
+ string_drop _stringappend_157_
+ (build_ex _stringappend_158_) in
+ sep_matches_prefix _stringappend_159_ >>= fun w__242 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__242 with
- | Some
- (_stringappend_237_,(existT _ _stringappend_238_ _)) =>
- let _stringappend_239_ :=
- string_drop _stringappend_236_
- (build_ex
- _stringappend_238_) in
- match (reg_name_matches_prefix _stringappend_239_) with
- | Some
- (_stringappend_240_,(existT _ _stringappend_241_ _)) =>
- let _stringappend_242_ :=
- string_drop _stringappend_239_
- (build_ex
- _stringappend_241_) in
- sep_matches_prefix _stringappend_242_ >>= fun w__243 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt,(existT _ _stringappend_160_ _)) =>
+ let _stringappend_161_ :=
+ string_drop _stringappend_159_
+ (build_ex _stringappend_160_) in
+ match (reg_name_matches_prefix _stringappend_161_) with
+ | Some (rs1,(existT _ _stringappend_162_ _)) =>
+ let _stringappend_163_ :=
+ string_drop _stringappend_161_
+ (build_ex _stringappend_162_) in
+ sep_matches_prefix _stringappend_163_ >>= fun w__243 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__243 with
| Some
- (_stringappend_243_,(existT _ _stringappend_244_ _)) =>
- let _stringappend_245_ :=
+ (tt,(existT _ _stringappend_164_ _)) =>
+ let _stringappend_165_ :=
string_drop
- _stringappend_242_
- (build_ex
- _stringappend_244_) in
+ _stringappend_163_
+ (build_ex _stringappend_164_) in
if ((match (reg_name_matches_prefix
- _stringappend_245_) with
+ _stringappend_165_) with
| Some
- (_stringappend_246_,(existT _ _stringappend_247_ _)) =>
+ (rs2,(existT _ _stringappend_166_ _)) =>
match (string_drop
- _stringappend_245_
- (build_ex
- _stringappend_247_)) with
+ _stringappend_165_
+ (build_ex _stringappend_166_)) with
| s_ => true
end
| None => false
@@ -25398,154 +23645,116 @@ Definition assembly_matches_prefix (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__248 : bool =>
(if (w__248) then
- match (rtypew_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_228_,(existT _ _stringappend_229_ _)) =>
- returnm ((_stringappend_228_, build_ex _stringappend_229_)
- : (ropw * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((ropw * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__250 : (ropw * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_229_ _) :=
- w__250
- : (ropw * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_230_ :=
- string_drop _stringappend_0_ (build_ex _stringappend_229_) in
- match (spc_matches_prefix _stringappend_230_) with
- | Some (_stringappend_231_,(existT _ _stringappend_232_ _)) =>
- returnm ((_stringappend_231_, build_ex _stringappend_232_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__252 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_232_ _) :=
- w__252
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_233_ :=
- string_drop _stringappend_230_
- (build_ex
- _stringappend_232_) in
- match (reg_name_matches_prefix _stringappend_233_) with
- | Some (_stringappend_234_,(existT _ _stringappend_235_ _)) =>
- returnm ((_stringappend_234_, build_ex _stringappend_235_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__254 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_235_ _) :=
- w__254
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_236_ :=
- string_drop _stringappend_233_
- (build_ex
- _stringappend_235_) in
- sep_matches_prefix _stringappend_236_ >>= fun w__255 : option ((unit * {n : Z & ArithFact (n >=
+ (match (rtypew_mnemonic_matches_prefix _stringappend_0_) with
+ | Some (op,(existT _ _stringappend_154_ _)) =>
+ returnm (op, build_ex _stringappend_154_)
+ | _ => exit tt : M ((ropw * {n : Z & ArithFact (n >= 0)}))
+ end : M ((ropw * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_154_ _) =>
+ let _stringappend_155_ :=
+ string_drop _stringappend_0_ (build_ex _stringappend_154_) in
+ (match (spc_matches_prefix _stringappend_155_) with
+ | Some (tt,(existT _ _stringappend_156_ _)) =>
+ returnm (tt, build_ex _stringappend_156_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_156_ _) =>
+ let _stringappend_157_ :=
+ string_drop _stringappend_155_
+ (build_ex _stringappend_156_) in
+ (match (reg_name_matches_prefix _stringappend_157_) with
+ | Some (rd,(existT _ _stringappend_158_ _)) =>
+ returnm (rd, build_ex _stringappend_158_)
+ | _ =>
+ exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_158_ _) =>
+ let _stringappend_159_ :=
+ string_drop _stringappend_157_
+ (build_ex _stringappend_158_) in
+ sep_matches_prefix _stringappend_159_ >>= fun w__255 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__255 with
- | Some (_stringappend_237_,(existT _ _stringappend_238_ _)) =>
- returnm ((_stringappend_237_, build_ex _stringappend_238_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__257 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_238_ _) :=
- w__257
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_239_ :=
- string_drop _stringappend_236_
- (build_ex
- _stringappend_238_) in
- match (reg_name_matches_prefix _stringappend_239_) with
- | Some (_stringappend_240_,(existT _ _stringappend_241_ _)) =>
- returnm ((_stringappend_240_, build_ex _stringappend_241_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__259 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_241_ _) :=
- w__259
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_242_ :=
- string_drop _stringappend_239_
- (build_ex
- _stringappend_241_) in
- sep_matches_prefix _stringappend_242_ >>= fun w__260 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__255 with
+ | Some (tt,(existT _ _stringappend_160_ _)) =>
+ returnm (tt, build_ex _stringappend_160_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_160_ _) =>
+ let _stringappend_161_ :=
+ string_drop _stringappend_159_
+ (build_ex _stringappend_160_) in
+ (match (reg_name_matches_prefix _stringappend_161_) with
+ | Some (rs1,(existT _ _stringappend_162_ _)) =>
+ returnm (rs1, build_ex _stringappend_162_)
+ | _ =>
+ exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_162_ _) =>
+ let _stringappend_163_ :=
+ string_drop _stringappend_161_
+ (build_ex _stringappend_162_) in
+ sep_matches_prefix _stringappend_163_ >>= fun w__260 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__260 with
- | Some (_stringappend_243_,(existT _ _stringappend_244_ _)) =>
- returnm ((_stringappend_243_, build_ex _stringappend_244_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__262 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_244_ _) :=
- w__262
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_245_ :=
- string_drop _stringappend_242_
- (build_ex
- _stringappend_244_) in
- match (reg_name_matches_prefix _stringappend_245_) with
- | Some (_stringappend_246_,(existT _ _stringappend_247_ _)) =>
- returnm ((_stringappend_246_, build_ex _stringappend_247_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__264 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs2, existT _ _stringappend_247_ _) :=
- w__264
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_245_
- (build_ex
- _stringappend_247_)) with
+ (match w__260 with
+ | Some (tt,(existT _ _stringappend_164_ _)) =>
+ returnm (tt, build_ex _stringappend_164_)
+ | _ => exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_164_ _) =>
+ let _stringappend_165_ :=
+ string_drop _stringappend_163_
+ (build_ex _stringappend_164_) in
+ (match (reg_name_matches_prefix _stringappend_165_) with
+ | Some (rs2,(existT _ _stringappend_166_ _)) =>
+ returnm (rs2, build_ex _stringappend_166_)
+ | _ =>
+ exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_166_ _) =>
+ returnm ((match (string_drop _stringappend_165_
+ (build_ex _stringappend_166_)) with
| s_ =>
- Some (RTYPEW (rs2,rs1,rd,op),
- sub_nat (string_length arg_)
- (string_length s_))
+ Some
+ ((RTYPEW
+ (rs2, rs1, rd, op), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
match (shiftiwop_mnemonic_matches_prefix _stringappend_0_) with
- | Some (_stringappend_249_,(existT _ _stringappend_250_ _)) =>
- let _stringappend_251_ :=
+ | Some (op,(existT _ _stringappend_168_ _)) =>
+ let _stringappend_169_ :=
string_drop _stringappend_0_
- (build_ex
- _stringappend_250_) in
- match (spc_matches_prefix _stringappend_251_) with
- | Some
- (_stringappend_252_,(existT _ _stringappend_253_ _)) =>
- let _stringappend_254_ :=
- string_drop _stringappend_251_
- (build_ex
- _stringappend_253_) in
- match (reg_name_matches_prefix _stringappend_254_) with
- | Some
- (_stringappend_255_,(existT _ _stringappend_256_ _)) =>
- let _stringappend_257_ :=
- string_drop _stringappend_254_
- (build_ex
- _stringappend_256_) in
- sep_matches_prefix _stringappend_257_ >>= fun w__265 : option ((unit * {n : Z & ArithFact (n >=
+ (build_ex _stringappend_168_) in
+ match (spc_matches_prefix _stringappend_169_) with
+ | Some (tt,(existT _ _stringappend_170_ _)) =>
+ let _stringappend_171_ :=
+ string_drop _stringappend_169_
+ (build_ex _stringappend_170_) in
+ match (reg_name_matches_prefix _stringappend_171_) with
+ | Some (rd,(existT _ _stringappend_172_ _)) =>
+ let _stringappend_173_ :=
+ string_drop _stringappend_171_
+ (build_ex _stringappend_172_) in
+ sep_matches_prefix _stringappend_173_ >>= fun w__265 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__265 with
| Some
- (_stringappend_258_,(existT _ _stringappend_259_ _)) =>
- let _stringappend_260_ :=
- string_drop _stringappend_257_
- (build_ex
- _stringappend_259_) in
+ (tt,(existT _ _stringappend_174_ _)) =>
+ let _stringappend_175_ :=
+ string_drop _stringappend_173_
+ (build_ex _stringappend_174_) in
if ((match (reg_name_matches_prefix
- _stringappend_260_) with
+ _stringappend_175_) with
| Some
- (_stringappend_261_,(existT _ _stringappend_262_ _)) =>
- let _stringappend_263_ :=
+ (rs1,(existT _ _stringappend_176_ _)) =>
+ let _stringappend_177_ :=
string_drop
- _stringappend_260_
- (build_ex
- _stringappend_262_) in
+ _stringappend_175_
+ (build_ex _stringappend_176_) in
if ((match (hex_bits_5_matches_prefix
- _stringappend_263_) with
+ _stringappend_177_) with
| Some
- (_stringappend_264_,(existT _ _stringappend_265_ _)) =>
+ (shamt,(existT _ _stringappend_178_ _)) =>
match (string_drop
- _stringappend_263_
- (build_ex
- _stringappend_265_)) with
+ _stringappend_177_
+ (build_ex _stringappend_178_)) with
| s_ => true
end
| None => false
@@ -25574,172 +23783,120 @@ Definition assembly_matches_prefix (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__268 : bool =>
(if (w__268) then
- match (shiftiwop_mnemonic_matches_prefix _stringappend_0_) with
- | Some
- (_stringappend_249_,(existT _ _stringappend_250_ _)) =>
- returnm ((_stringappend_249_,
- build_ex
- _stringappend_250_)
- : (sopw * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((sopw * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__270 : (sopw * {n : Z & ArithFact (n >= 0)}) =>
- let '(op, existT _ _stringappend_250_ _) :=
- w__270
- : (sopw * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_251_ :=
+ (match (shiftiwop_mnemonic_matches_prefix _stringappend_0_) with
+ | Some (op,(existT _ _stringappend_168_ _)) =>
+ returnm (op, build_ex _stringappend_168_)
+ | _ =>
+ exit tt : M ((sopw * {n : Z & ArithFact (n >= 0)}))
+ end : M ((sopw * {n : Z & ArithFact (n >= 0)}))) >>= fun '(op, existT _ _stringappend_168_ _) =>
+ let _stringappend_169_ :=
string_drop _stringappend_0_
- (build_ex
- _stringappend_250_) in
- match (spc_matches_prefix _stringappend_251_) with
- | Some
- (_stringappend_252_,(existT _ _stringappend_253_ _)) =>
- returnm ((_stringappend_252_,
- build_ex
- _stringappend_253_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__272 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_253_ _) :=
- w__272
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_254_ :=
- string_drop _stringappend_251_
- (build_ex
- _stringappend_253_) in
- match (reg_name_matches_prefix _stringappend_254_) with
- | Some
- (_stringappend_255_,(existT _ _stringappend_256_ _)) =>
- returnm ((_stringappend_255_,
- build_ex
- _stringappend_256_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__274 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rd, existT _ _stringappend_256_ _) :=
- w__274
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_257_ :=
- string_drop _stringappend_254_
- (build_ex
- _stringappend_256_) in
- sep_matches_prefix _stringappend_257_ >>= fun w__275 : option ((unit * {n : Z & ArithFact (n >=
+ (build_ex _stringappend_168_) in
+ (match (spc_matches_prefix _stringappend_169_) with
+ | Some (tt,(existT _ _stringappend_170_ _)) =>
+ returnm (tt, build_ex _stringappend_170_)
+ | _ =>
+ exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_170_ _) =>
+ let _stringappend_171_ :=
+ string_drop _stringappend_169_
+ (build_ex _stringappend_170_) in
+ (match (reg_name_matches_prefix _stringappend_171_) with
+ | Some (rd,(existT _ _stringappend_172_ _)) =>
+ returnm (rd, build_ex _stringappend_172_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_172_ _) =>
+ let _stringappend_173_ :=
+ string_drop _stringappend_171_
+ (build_ex _stringappend_172_) in
+ sep_matches_prefix _stringappend_173_ >>= fun w__275 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__275 with
- | Some
- (_stringappend_258_,(existT _ _stringappend_259_ _)) =>
- returnm ((_stringappend_258_,
- build_ex
- _stringappend_259_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__277 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_259_ _) :=
- w__277
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_260_ :=
- string_drop _stringappend_257_
- (build_ex
- _stringappend_259_) in
- match (reg_name_matches_prefix _stringappend_260_) with
- | Some
- (_stringappend_261_,(existT _ _stringappend_262_ _)) =>
- returnm ((_stringappend_261_,
- build_ex
- _stringappend_262_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__279 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(rs1, existT _ _stringappend_262_ _) :=
- w__279
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_263_ :=
- string_drop _stringappend_260_
- (build_ex
- _stringappend_262_) in
- match (hex_bits_5_matches_prefix _stringappend_263_) with
- | Some
- (_stringappend_264_,(existT _ _stringappend_265_ _)) =>
- returnm ((_stringappend_264_,
- build_ex
- _stringappend_265_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__281 : (mword 5 * {n : Z & ArithFact (n >= 0)}) =>
- let '(shamt, existT _ _stringappend_265_ _) :=
- w__281
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_263_
- (build_ex
- _stringappend_265_)) with
+ (match w__275 with
+ | Some (tt,(existT _ _stringappend_174_ _)) =>
+ returnm (tt, build_ex _stringappend_174_)
+ | _ =>
+ exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_174_ _) =>
+ let _stringappend_175_ :=
+ string_drop _stringappend_173_
+ (build_ex _stringappend_174_) in
+ (match (reg_name_matches_prefix _stringappend_175_) with
+ | Some (rs1,(existT _ _stringappend_176_ _)) =>
+ returnm (rs1, build_ex _stringappend_176_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_176_ _) =>
+ let _stringappend_177_ :=
+ string_drop _stringappend_175_
+ (build_ex _stringappend_176_) in
+ (match (hex_bits_5_matches_prefix _stringappend_177_) with
+ | Some (shamt,(existT _ _stringappend_178_ _)) =>
+ returnm (shamt, build_ex _stringappend_178_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(shamt, existT _ _stringappend_178_ _) =>
+ returnm ((match (string_drop _stringappend_177_
+ (build_ex _stringappend_178_)) with
| s_ =>
- Some (SHIFTIWOP (shamt,rs1,rd,op),
- sub_nat (string_length arg_)
- (string_length s_))
+ Some
+ ((SHIFTIWOP
+ (shamt, rs1, rd, op), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
match (mul_mnemonic_matches_prefix _stringappend_0_) with
| Some
- (_stringappend_267_,(existT _ _stringappend_268_ _)) =>
- let _stringappend_269_ :=
+ ((high, signed1, signed2),(existT _ _stringappend_180_ _)) =>
+ let _stringappend_181_ :=
string_drop _stringappend_0_
- (build_ex
- _stringappend_268_) in
- match (spc_matches_prefix _stringappend_269_) with
- | Some
- (_stringappend_270_,(existT _ _stringappend_271_ _)) =>
- let _stringappend_272_ :=
- string_drop _stringappend_269_
- (build_ex
- _stringappend_271_) in
- match (reg_name_matches_prefix _stringappend_272_) with
- | Some
- (_stringappend_273_,(existT _ _stringappend_274_ _)) =>
- let _stringappend_275_ :=
- string_drop _stringappend_272_
- (build_ex
- _stringappend_274_) in
- sep_matches_prefix _stringappend_275_ >>= fun w__282 : option ((unit * {n : Z & ArithFact (n >=
+ (build_ex _stringappend_180_) in
+ match (spc_matches_prefix _stringappend_181_) with
+ | Some (tt,(existT _ _stringappend_182_ _)) =>
+ let _stringappend_183_ :=
+ string_drop _stringappend_181_
+ (build_ex _stringappend_182_) in
+ match (reg_name_matches_prefix _stringappend_183_) with
+ | Some (rd,(existT _ _stringappend_184_ _)) =>
+ let _stringappend_185_ :=
+ string_drop _stringappend_183_
+ (build_ex _stringappend_184_) in
+ sep_matches_prefix _stringappend_185_ >>= fun w__282 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__282 with
- | Some
- (_stringappend_276_,(existT _ _stringappend_277_ _)) =>
- let _stringappend_278_ :=
- string_drop _stringappend_275_
- (build_ex
- _stringappend_277_) in
+ | Some (tt,(existT _ _stringappend_186_ _)) =>
+ let _stringappend_187_ :=
+ string_drop _stringappend_185_
+ (build_ex _stringappend_186_) in
match (reg_name_matches_prefix
- _stringappend_278_) with
- | Some
- (_stringappend_279_,(existT _ _stringappend_280_ _)) =>
- let _stringappend_281_ :=
- string_drop _stringappend_278_
- (build_ex
- _stringappend_280_) in
- sep_matches_prefix _stringappend_281_ >>= fun w__283 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_187_) with
+ | Some (rs1,(existT _ _stringappend_188_ _)) =>
+ let _stringappend_189_ :=
+ string_drop _stringappend_187_
+ (build_ex _stringappend_188_) in
+ sep_matches_prefix _stringappend_189_ >>= fun w__283 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__283 with
| Some
- (_stringappend_282_,(existT _ _stringappend_283_ _)) =>
- let _stringappend_284_ :=
+ (tt,(existT _ _stringappend_190_ _)) =>
+ let _stringappend_191_ :=
string_drop
- _stringappend_281_
- (build_ex
- _stringappend_283_) in
+ _stringappend_189_
+ (build_ex _stringappend_190_) in
if ((match (reg_name_matches_prefix
- _stringappend_284_) with
+ _stringappend_191_) with
| Some
- (_stringappend_285_,(existT _ _stringappend_286_ _)) =>
+ (rs2,(existT _ _stringappend_192_ _)) =>
match (string_drop
- _stringappend_284_
- (build_ex
- _stringappend_286_)) with
+ _stringappend_191_
+ (build_ex _stringappend_192_)) with
| s_ => true
end
| None => false
@@ -25774,207 +23931,146 @@ Definition assembly_matches_prefix (arg_ : string)
| None => returnm (false : bool)
end >>= fun w__288 : bool =>
(if (w__288) then
- match (mul_mnemonic_matches_prefix _stringappend_0_) with
- | Some
- (_stringappend_267_,(existT _ _stringappend_268_ _)) =>
- returnm ((_stringappend_267_,
- build_ex
- _stringappend_268_)
- : ((bool * bool * bool) * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M (((bool * bool * bool) * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__290 : ((bool * bool * bool) * {n : Z & ArithFact (n >=
- 0)}) =>
- let '((high, signed1, signed2), existT _ _stringappend_268_ _) :=
- w__290
- : ((bool * bool * bool) * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_269_ :=
+ (match (mul_mnemonic_matches_prefix _stringappend_0_) with
+ | Some
+ ((high, signed1, signed2),(existT _ _stringappend_180_ _)) =>
+ returnm ((high, signed1, signed2), build_ex _stringappend_180_)
+ | _ =>
+ exit tt
+ : M (((bool * bool * bool) * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M (((bool * bool * bool) * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '((high, signed1, signed2), existT _ _stringappend_180_ _) =>
+ let _stringappend_181_ :=
string_drop _stringappend_0_
- (build_ex
- _stringappend_268_) in
- match (spc_matches_prefix _stringappend_269_) with
- | Some
- (_stringappend_270_,(existT _ _stringappend_271_ _)) =>
- returnm ((_stringappend_270_,
- build_ex
- _stringappend_271_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__292 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_271_ _) :=
- w__292
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_272_ :=
- string_drop _stringappend_269_
- (build_ex
- _stringappend_271_) in
- match (reg_name_matches_prefix _stringappend_272_) with
- | Some
- (_stringappend_273_,(existT _ _stringappend_274_ _)) =>
- returnm ((_stringappend_273_,
- build_ex
- _stringappend_274_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__294 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_274_ _) :=
- w__294
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_275_ :=
- string_drop _stringappend_272_
- (build_ex
- _stringappend_274_) in
- sep_matches_prefix _stringappend_275_ >>= fun w__295 : option ((unit * {n : Z & ArithFact (n >=
+ (build_ex _stringappend_180_) in
+ (match (spc_matches_prefix _stringappend_181_) with
+ | Some (tt,(existT _ _stringappend_182_ _)) =>
+ returnm (tt, build_ex _stringappend_182_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_182_ _) =>
+ let _stringappend_183_ :=
+ string_drop _stringappend_181_
+ (build_ex _stringappend_182_) in
+ (match (reg_name_matches_prefix _stringappend_183_) with
+ | Some (rd,(existT _ _stringappend_184_ _)) =>
+ returnm (rd, build_ex _stringappend_184_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_184_ _) =>
+ let _stringappend_185_ :=
+ string_drop _stringappend_183_
+ (build_ex _stringappend_184_) in
+ sep_matches_prefix _stringappend_185_ >>= fun w__295 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__295 with
- | Some
- (_stringappend_276_,(existT _ _stringappend_277_ _)) =>
- returnm ((_stringappend_276_,
- build_ex
- _stringappend_277_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__297 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_277_ _) :=
- w__297
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_278_ :=
- string_drop _stringappend_275_
- (build_ex
- _stringappend_277_) in
- match (reg_name_matches_prefix _stringappend_278_) with
- | Some
- (_stringappend_279_,(existT _ _stringappend_280_ _)) =>
- returnm ((_stringappend_279_,
- build_ex
- _stringappend_280_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__299 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_280_ _) :=
- w__299
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_281_ :=
- string_drop _stringappend_278_
- (build_ex
- _stringappend_280_) in
- sep_matches_prefix _stringappend_281_ >>= fun w__300 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__295 with
+ | Some (tt,(existT _ _stringappend_186_ _)) =>
+ returnm (tt, build_ex _stringappend_186_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_186_ _) =>
+ let _stringappend_187_ :=
+ string_drop _stringappend_185_
+ (build_ex _stringappend_186_) in
+ (match (reg_name_matches_prefix _stringappend_187_) with
+ | Some (rs1,(existT _ _stringappend_188_ _)) =>
+ returnm (rs1, build_ex _stringappend_188_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_188_ _) =>
+ let _stringappend_189_ :=
+ string_drop _stringappend_187_
+ (build_ex _stringappend_188_) in
+ sep_matches_prefix _stringappend_189_ >>= fun w__300 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__300 with
- | Some
- (_stringappend_282_,(existT _ _stringappend_283_ _)) =>
- returnm ((_stringappend_282_,
- build_ex
- _stringappend_283_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__302 : (unit * {n : Z & ArithFact (n >= 0)}) =>
- let '(tt, existT _ _stringappend_283_ _) :=
- w__302
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_284_ :=
- string_drop _stringappend_281_
- (build_ex
- _stringappend_283_) in
- match (reg_name_matches_prefix _stringappend_284_) with
- | Some
- (_stringappend_285_,(existT _ _stringappend_286_ _)) =>
- returnm ((_stringappend_285_,
- build_ex
- _stringappend_286_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__304 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_286_ _) :=
- w__304
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_284_
- (build_ex
- _stringappend_286_)) with
+ (match w__300 with
+ | Some (tt,(existT _ _stringappend_190_ _)) =>
+ returnm (tt, build_ex _stringappend_190_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_190_ _) =>
+ let _stringappend_191_ :=
+ string_drop _stringappend_189_
+ (build_ex _stringappend_190_) in
+ (match (reg_name_matches_prefix _stringappend_191_) with
+ | Some (rs2,(existT _ _stringappend_192_ _)) =>
+ returnm (rs2, build_ex _stringappend_192_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_192_ _) =>
+ returnm ((match (string_drop _stringappend_191_
+ (build_ex _stringappend_192_)) with
| s_ =>
- Some (MUL (rs2,rs1,rd,high,signed1,signed2),
- sub_nat (string_length arg_)
- (string_length s_))
+ Some
+ ((MUL
+ (rs2, rs1, rd, high, signed1,
+ signed2), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
and_boolM
(returnm ((string_startswith _stringappend_0_ "div")
: bool))
- (let _stringappend_288_ :=
- string_drop _stringappend_0_ (string_length "div") in
- match (maybe_not_u_matches_prefix _stringappend_288_) with
- | Some
- (_stringappend_289_,(existT _ _stringappend_290_ _)) =>
- let _stringappend_291_ :=
- string_drop _stringappend_288_
- (build_ex
- _stringappend_290_) in
- match (spc_matches_prefix _stringappend_291_) with
- | Some
- (_stringappend_292_,(existT _ _stringappend_293_ _)) =>
- let _stringappend_294_ :=
- string_drop _stringappend_291_
- (build_ex
- _stringappend_293_) in
+ (let _stringappend_194_ :=
+ string_drop _stringappend_0_
+ (build_ex (projT1 (string_length "div"))) in
+ match (maybe_not_u_matches_prefix _stringappend_194_) with
+ | Some (s,(existT _ _stringappend_195_ _)) =>
+ let _stringappend_196_ :=
+ string_drop _stringappend_194_
+ (build_ex _stringappend_195_) in
+ match (spc_matches_prefix _stringappend_196_) with
+ | Some (tt,(existT _ _stringappend_197_ _)) =>
+ let _stringappend_198_ :=
+ string_drop _stringappend_196_
+ (build_ex _stringappend_197_) in
match (reg_name_matches_prefix
- _stringappend_294_) with
- | Some
- (_stringappend_295_,(existT _ _stringappend_296_ _)) =>
- let _stringappend_297_ :=
- string_drop _stringappend_294_
- (build_ex
- _stringappend_296_) in
- sep_matches_prefix _stringappend_297_ >>= fun w__305 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_198_) with
+ | Some (rd,(existT _ _stringappend_199_ _)) =>
+ let _stringappend_200_ :=
+ string_drop _stringappend_198_
+ (build_ex _stringappend_199_) in
+ sep_matches_prefix _stringappend_200_ >>= fun w__305 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__305 with
- | Some
- (_stringappend_298_,(existT _ _stringappend_299_ _)) =>
- let _stringappend_300_ :=
- string_drop _stringappend_297_
- (build_ex
- _stringappend_299_) in
+ | Some (tt,(existT _ _stringappend_201_ _)) =>
+ let _stringappend_202_ :=
+ string_drop _stringappend_200_
+ (build_ex _stringappend_201_) in
match (reg_name_matches_prefix
- _stringappend_300_) with
+ _stringappend_202_) with
| Some
- (_stringappend_301_,(existT _ _stringappend_302_ _)) =>
- let _stringappend_303_ :=
- string_drop _stringappend_300_
- (build_ex
- _stringappend_302_) in
- sep_matches_prefix _stringappend_303_ >>= fun w__306 : option ((unit * {n : Z & ArithFact (n >=
+ (rs1,(existT _ _stringappend_203_ _)) =>
+ let _stringappend_204_ :=
+ string_drop _stringappend_202_
+ (build_ex _stringappend_203_) in
+ sep_matches_prefix _stringappend_204_ >>= fun w__306 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__306 with
| Some
- (_stringappend_304_,(existT _ _stringappend_305_ _)) =>
- let _stringappend_306_ :=
+ (tt,(existT _ _stringappend_205_ _)) =>
+ let _stringappend_206_ :=
string_drop
- _stringappend_303_
- (build_ex
- _stringappend_305_) in
+ _stringappend_204_
+ (build_ex _stringappend_205_) in
if ((match (reg_name_matches_prefix
- _stringappend_306_) with
+ _stringappend_206_) with
| Some
- (_stringappend_307_,(existT _ _stringappend_308_ _)) =>
+ (rs2,(existT _ _stringappend_207_ _)) =>
match (string_drop
- _stringappend_306_
- (build_ex
- _stringappend_308_)) with
+ _stringappend_206_
+ (build_ex _stringappend_207_)) with
| s_ => true
end
| None => false
@@ -26012,148 +24108,91 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__312 : bool =>
(if (w__312) then
- let _stringappend_288_ :=
- string_drop _stringappend_0_ (string_length "div") in
- match (maybe_not_u_matches_prefix _stringappend_288_) with
- | Some
- (_stringappend_289_,(existT _ _stringappend_290_ _)) =>
- returnm ((_stringappend_289_,
- build_ex
- _stringappend_290_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__314 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_290_ _) :=
- w__314
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_291_ :=
- string_drop _stringappend_288_
- (build_ex
- _stringappend_290_) in
- match (spc_matches_prefix _stringappend_291_) with
- | Some
- (_stringappend_292_,(existT _ _stringappend_293_ _)) =>
- returnm ((_stringappend_292_,
- build_ex
- _stringappend_293_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__316 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_293_ _) :=
- w__316
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_294_ :=
- string_drop _stringappend_291_
- (build_ex
- _stringappend_293_) in
- match (reg_name_matches_prefix _stringappend_294_) with
- | Some
- (_stringappend_295_,(existT _ _stringappend_296_ _)) =>
- returnm ((_stringappend_295_,
- build_ex
- _stringappend_296_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__318 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_296_ _) :=
- w__318
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_297_ :=
- string_drop _stringappend_294_
- (build_ex
- _stringappend_296_) in
- sep_matches_prefix _stringappend_297_ >>= fun w__319 : option ((unit * {n : Z & ArithFact (n >=
+ let _stringappend_194_ :=
+ string_drop _stringappend_0_
+ (build_ex (projT1 (string_length "div"))) in
+ (match (maybe_not_u_matches_prefix
+ _stringappend_194_) with
+ | Some (s,(existT _ _stringappend_195_ _)) =>
+ returnm (s, build_ex _stringappend_195_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(s, existT _ _stringappend_195_ _) =>
+ let _stringappend_196_ :=
+ string_drop _stringappend_194_
+ (build_ex _stringappend_195_) in
+ (match (spc_matches_prefix _stringappend_196_) with
+ | Some (tt,(existT _ _stringappend_197_ _)) =>
+ returnm (tt, build_ex _stringappend_197_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_197_ _) =>
+ let _stringappend_198_ :=
+ string_drop _stringappend_196_
+ (build_ex _stringappend_197_) in
+ (match (reg_name_matches_prefix _stringappend_198_) with
+ | Some (rd,(existT _ _stringappend_199_ _)) =>
+ returnm (rd, build_ex _stringappend_199_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_199_ _) =>
+ let _stringappend_200_ :=
+ string_drop _stringappend_198_
+ (build_ex _stringappend_199_) in
+ sep_matches_prefix _stringappend_200_ >>= fun w__319 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__319 with
- | Some
- (_stringappend_298_,(existT _ _stringappend_299_ _)) =>
- returnm ((_stringappend_298_,
- build_ex
- _stringappend_299_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__321 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_299_ _) :=
- w__321
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_300_ :=
- string_drop _stringappend_297_
- (build_ex
- _stringappend_299_) in
- match (reg_name_matches_prefix _stringappend_300_) with
- | Some
- (_stringappend_301_,(existT _ _stringappend_302_ _)) =>
- returnm ((_stringappend_301_,
- build_ex
- _stringappend_302_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__323 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_302_ _) :=
- w__323
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_303_ :=
- string_drop _stringappend_300_
- (build_ex
- _stringappend_302_) in
- sep_matches_prefix _stringappend_303_ >>= fun w__324 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__319 with
+ | Some (tt,(existT _ _stringappend_201_ _)) =>
+ returnm (tt, build_ex _stringappend_201_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_201_ _) =>
+ let _stringappend_202_ :=
+ string_drop _stringappend_200_
+ (build_ex _stringappend_201_) in
+ (match (reg_name_matches_prefix _stringappend_202_) with
+ | Some (rs1,(existT _ _stringappend_203_ _)) =>
+ returnm (rs1, build_ex _stringappend_203_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_203_ _) =>
+ let _stringappend_204_ :=
+ string_drop _stringappend_202_
+ (build_ex _stringappend_203_) in
+ sep_matches_prefix _stringappend_204_ >>= fun w__324 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__324 with
- | Some
- (_stringappend_304_,(existT _ _stringappend_305_ _)) =>
- returnm ((_stringappend_304_,
- build_ex
- _stringappend_305_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__326 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_305_ _) :=
- w__326
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_306_ :=
- string_drop _stringappend_303_
- (build_ex
- _stringappend_305_) in
- match (reg_name_matches_prefix _stringappend_306_) with
- | Some
- (_stringappend_307_,(existT _ _stringappend_308_ _)) =>
- returnm ((_stringappend_307_,
- build_ex
- _stringappend_308_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__328 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_308_ _) :=
- w__328
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_306_
- (build_ex
- _stringappend_308_)) with
+ (match w__324 with
+ | Some (tt,(existT _ _stringappend_205_ _)) =>
+ returnm (tt, build_ex _stringappend_205_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_205_ _) =>
+ let _stringappend_206_ :=
+ string_drop _stringappend_204_
+ (build_ex _stringappend_205_) in
+ (match (reg_name_matches_prefix _stringappend_206_) with
+ | Some (rs2,(existT _ _stringappend_207_ _)) =>
+ returnm (rs2, build_ex _stringappend_207_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_207_ _) =>
+ returnm ((match (string_drop _stringappend_206_
+ (build_ex _stringappend_207_)) with
| s_ =>
- Some (DIV (rs2,rs1,rd,s),
- sub_nat (string_length arg_)
- (string_length s_))
+ Some
+ ((DIV
+ (rs2, rs1, rd, s), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
@@ -26161,68 +24200,58 @@ Definition assembly_matches_prefix (arg_ : string)
(returnm ((string_startswith _stringappend_0_
"rem")
: bool))
- (let _stringappend_310_ :=
+ (let _stringappend_209_ :=
string_drop _stringappend_0_
- (string_length "rem") in
+ (build_ex (projT1 (string_length "rem"))) in
match (maybe_not_u_matches_prefix
- _stringappend_310_) with
- | Some
- (_stringappend_311_,(existT _ _stringappend_312_ _)) =>
- let _stringappend_313_ :=
- string_drop _stringappend_310_
- (build_ex
- _stringappend_312_) in
- match (spc_matches_prefix _stringappend_313_) with
- | Some
- (_stringappend_314_,(existT _ _stringappend_315_ _)) =>
- let _stringappend_316_ :=
- string_drop _stringappend_313_
- (build_ex
- _stringappend_315_) in
+ _stringappend_209_) with
+ | Some (s,(existT _ _stringappend_210_ _)) =>
+ let _stringappend_211_ :=
+ string_drop _stringappend_209_
+ (build_ex _stringappend_210_) in
+ match (spc_matches_prefix _stringappend_211_) with
+ | Some (tt,(existT _ _stringappend_212_ _)) =>
+ let _stringappend_213_ :=
+ string_drop _stringappend_211_
+ (build_ex _stringappend_212_) in
match (reg_name_matches_prefix
- _stringappend_316_) with
- | Some
- (_stringappend_317_,(existT _ _stringappend_318_ _)) =>
- let _stringappend_319_ :=
- string_drop _stringappend_316_
- (build_ex
- _stringappend_318_) in
- sep_matches_prefix _stringappend_319_ >>= fun w__329 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_213_) with
+ | Some (rd,(existT _ _stringappend_214_ _)) =>
+ let _stringappend_215_ :=
+ string_drop _stringappend_213_
+ (build_ex _stringappend_214_) in
+ sep_matches_prefix _stringappend_215_ >>= fun w__329 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__329 with
| Some
- (_stringappend_320_,(existT _ _stringappend_321_ _)) =>
- let _stringappend_322_ :=
- string_drop _stringappend_319_
- (build_ex
- _stringappend_321_) in
+ (tt,(existT _ _stringappend_216_ _)) =>
+ let _stringappend_217_ :=
+ string_drop _stringappend_215_
+ (build_ex _stringappend_216_) in
match (reg_name_matches_prefix
- _stringappend_322_) with
+ _stringappend_217_) with
| Some
- (_stringappend_323_,(existT _ _stringappend_324_ _)) =>
- let _stringappend_325_ :=
- string_drop _stringappend_322_
- (build_ex
- _stringappend_324_) in
+ (rs1,(existT _ _stringappend_218_ _)) =>
+ let _stringappend_219_ :=
+ string_drop _stringappend_217_
+ (build_ex _stringappend_218_) in
sep_matches_prefix
- _stringappend_325_ >>= fun w__330 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_219_ >>= fun w__330 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__330 with
| Some
- (_stringappend_326_,(existT _ _stringappend_327_ _)) =>
- let _stringappend_328_ :=
+ (tt,(existT _ _stringappend_220_ _)) =>
+ let _stringappend_221_ :=
string_drop
- _stringappend_325_
- (build_ex
- _stringappend_327_) in
+ _stringappend_219_
+ (build_ex _stringappend_220_) in
if ((match (reg_name_matches_prefix
- _stringappend_328_) with
+ _stringappend_221_) with
| Some
- (_stringappend_329_,(existT _ _stringappend_330_ _)) =>
+ (rs2,(existT _ _stringappend_222_ _)) =>
match (string_drop
- _stringappend_328_
- (build_ex
- _stringappend_330_)) with
+ _stringappend_221_
+ (build_ex _stringappend_222_)) with
| s_ =>
true
end
@@ -26262,150 +24291,94 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__336 : bool =>
(if (w__336) then
- let _stringappend_310_ :=
+ let _stringappend_209_ :=
string_drop _stringappend_0_
- (string_length "rem") in
- match (maybe_not_u_matches_prefix
- _stringappend_310_) with
- | Some
- (_stringappend_311_,(existT _ _stringappend_312_ _)) =>
- returnm ((_stringappend_311_,
- build_ex
- _stringappend_312_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__338 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_312_ _) :=
- w__338
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_313_ :=
- string_drop _stringappend_310_
- (build_ex
- _stringappend_312_) in
- match (spc_matches_prefix _stringappend_313_) with
- | Some
- (_stringappend_314_,(existT _ _stringappend_315_ _)) =>
- returnm ((_stringappend_314_,
- build_ex
- _stringappend_315_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__340 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_315_ _) :=
- w__340
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_316_ :=
- string_drop _stringappend_313_
- (build_ex
- _stringappend_315_) in
- match (reg_name_matches_prefix _stringappend_316_) with
- | Some
- (_stringappend_317_,(existT _ _stringappend_318_ _)) =>
- returnm ((_stringappend_317_,
- build_ex
- _stringappend_318_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__342 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_318_ _) :=
- w__342
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_319_ :=
- string_drop _stringappend_316_
- (build_ex
- _stringappend_318_) in
- sep_matches_prefix _stringappend_319_ >>= fun w__343 : option ((unit * {n : Z & ArithFact (n >=
+ (build_ex (projT1 (string_length "rem"))) in
+ (match (maybe_not_u_matches_prefix
+ _stringappend_209_) with
+ | Some (s,(existT _ _stringappend_210_ _)) =>
+ returnm (s, build_ex _stringappend_210_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >= 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >= 0)}))) >>= fun '(s, existT _ _stringappend_210_ _) =>
+ let _stringappend_211_ :=
+ string_drop _stringappend_209_
+ (build_ex _stringappend_210_) in
+ (match (spc_matches_prefix _stringappend_211_) with
+ | Some (tt,(existT _ _stringappend_212_ _)) =>
+ returnm (tt, build_ex _stringappend_212_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_212_ _) =>
+ let _stringappend_213_ :=
+ string_drop _stringappend_211_
+ (build_ex _stringappend_212_) in
+ (match (reg_name_matches_prefix
+ _stringappend_213_) with
+ | Some (rd,(existT _ _stringappend_214_ _)) =>
+ returnm (rd, build_ex _stringappend_214_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rd, existT _ _stringappend_214_ _) =>
+ let _stringappend_215_ :=
+ string_drop _stringappend_213_
+ (build_ex _stringappend_214_) in
+ sep_matches_prefix _stringappend_215_ >>= fun w__343 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__343 with
- | Some
- (_stringappend_320_,(existT _ _stringappend_321_ _)) =>
- returnm ((_stringappend_320_,
- build_ex
- _stringappend_321_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__345 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_321_ _) :=
- w__345
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_322_ :=
- string_drop _stringappend_319_
- (build_ex
- _stringappend_321_) in
- match (reg_name_matches_prefix _stringappend_322_) with
- | Some
- (_stringappend_323_,(existT _ _stringappend_324_ _)) =>
- returnm ((_stringappend_323_,
- build_ex
- _stringappend_324_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__347 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_324_ _) :=
- w__347
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_325_ :=
- string_drop _stringappend_322_
- (build_ex
- _stringappend_324_) in
- sep_matches_prefix _stringappend_325_ >>= fun w__348 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__343 with
+ | Some (tt,(existT _ _stringappend_216_ _)) =>
+ returnm (tt, build_ex _stringappend_216_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_216_ _) =>
+ let _stringappend_217_ :=
+ string_drop _stringappend_215_
+ (build_ex _stringappend_216_) in
+ (match (reg_name_matches_prefix
+ _stringappend_217_) with
+ | Some (rs1,(existT _ _stringappend_218_ _)) =>
+ returnm (rs1, build_ex _stringappend_218_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs1, existT _ _stringappend_218_ _) =>
+ let _stringappend_219_ :=
+ string_drop _stringappend_217_
+ (build_ex _stringappend_218_) in
+ sep_matches_prefix _stringappend_219_ >>= fun w__348 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__348 with
- | Some
- (_stringappend_326_,(existT _ _stringappend_327_ _)) =>
- returnm ((_stringappend_326_,
- build_ex
- _stringappend_327_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__350 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_327_ _) :=
- w__350
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_328_ :=
- string_drop _stringappend_325_
- (build_ex
- _stringappend_327_) in
- match (reg_name_matches_prefix _stringappend_328_) with
- | Some
- (_stringappend_329_,(existT _ _stringappend_330_ _)) =>
- returnm ((_stringappend_329_,
- build_ex
- _stringappend_330_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__352 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_330_ _) :=
- w__352
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- returnm ((match (string_drop _stringappend_328_
- (build_ex
- _stringappend_330_)) with
+ (match w__348 with
+ | Some (tt,(existT _ _stringappend_220_ _)) =>
+ returnm (tt, build_ex _stringappend_220_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_220_ _) =>
+ let _stringappend_221_ :=
+ string_drop _stringappend_219_
+ (build_ex _stringappend_220_) in
+ (match (reg_name_matches_prefix
+ _stringappend_221_) with
+ | Some (rs2,(existT _ _stringappend_222_ _)) =>
+ returnm (rs2, build_ex _stringappend_222_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))) >>= fun '(rs2, existT _ _stringappend_222_ _) =>
+ returnm ((match (string_drop _stringappend_221_
+ (build_ex _stringappend_222_)) with
| s_ =>
- Some (REM (rs2,rs1,rd,s),
- sub_nat (string_length arg_)
- (string_length s_))
+ Some
+ ((REM
+ (rs2, rs1, rd, s), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
@@ -26413,60 +24386,52 @@ Definition assembly_matches_prefix (arg_ : string)
(returnm ((string_startswith _stringappend_0_
"mulw")
: bool))
- (let _stringappend_332_ :=
+ (let _stringappend_224_ :=
string_drop _stringappend_0_
- (string_length "mulw") in
- match (spc_matches_prefix _stringappend_332_) with
- | Some
- (_stringappend_333_,(existT _ _stringappend_334_ _)) =>
- let _stringappend_335_ :=
- string_drop _stringappend_332_
- (build_ex
- _stringappend_334_) in
+ (build_ex (projT1 (string_length "mulw"))) in
+ match (spc_matches_prefix _stringappend_224_) with
+ | Some (tt,(existT _ _stringappend_225_ _)) =>
+ let _stringappend_226_ :=
+ string_drop _stringappend_224_
+ (build_ex _stringappend_225_) in
match (reg_name_matches_prefix
- _stringappend_335_) with
- | Some
- (_stringappend_336_,(existT _ _stringappend_337_ _)) =>
- let _stringappend_338_ :=
- string_drop _stringappend_335_
- (build_ex
- _stringappend_337_) in
- sep_matches_prefix _stringappend_338_ >>= fun w__353 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_226_) with
+ | Some (rd,(existT _ _stringappend_227_ _)) =>
+ let _stringappend_228_ :=
+ string_drop _stringappend_226_
+ (build_ex _stringappend_227_) in
+ sep_matches_prefix _stringappend_228_ >>= fun w__353 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__353 with
| Some
- (_stringappend_339_,(existT _ _stringappend_340_ _)) =>
- let _stringappend_341_ :=
- string_drop _stringappend_338_
- (build_ex
- _stringappend_340_) in
+ (tt,(existT _ _stringappend_229_ _)) =>
+ let _stringappend_230_ :=
+ string_drop _stringappend_228_
+ (build_ex _stringappend_229_) in
match (reg_name_matches_prefix
- _stringappend_341_) with
+ _stringappend_230_) with
| Some
- (_stringappend_342_,(existT _ _stringappend_343_ _)) =>
- let _stringappend_344_ :=
- string_drop _stringappend_341_
- (build_ex
- _stringappend_343_) in
+ (rs1,(existT _ _stringappend_231_ _)) =>
+ let _stringappend_232_ :=
+ string_drop _stringappend_230_
+ (build_ex _stringappend_231_) in
sep_matches_prefix
- _stringappend_344_ >>= fun w__354 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_232_ >>= fun w__354 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__354 with
| Some
- (_stringappend_345_,(existT _ _stringappend_346_ _)) =>
- let _stringappend_347_ :=
+ (tt,(existT _ _stringappend_233_ _)) =>
+ let _stringappend_234_ :=
string_drop
- _stringappend_344_
- (build_ex
- _stringappend_346_) in
+ _stringappend_232_
+ (build_ex _stringappend_233_) in
if ((match (reg_name_matches_prefix
- _stringappend_347_) with
+ _stringappend_234_) with
| Some
- (_stringappend_348_,(existT _ _stringappend_349_ _)) =>
+ (rs2,(existT _ _stringappend_235_ _)) =>
match (string_drop
- _stringappend_347_
- (build_ex
- _stringappend_349_)) with
+ _stringappend_234_
+ (build_ex _stringappend_235_)) with
| s_ =>
true
end
@@ -26501,135 +24466,90 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__359 : bool =>
(if (w__359) then
- let _stringappend_332_ :=
+ let _stringappend_224_ :=
string_drop _stringappend_0_
- (string_length "mulw") in
- match (spc_matches_prefix _stringappend_332_) with
- | Some
- (_stringappend_333_,(existT _ _stringappend_334_ _)) =>
- returnm ((_stringappend_333_,
- build_ex
- _stringappend_334_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__361 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_334_ _) :=
- w__361
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_335_ :=
- string_drop _stringappend_332_
- (build_ex
- _stringappend_334_) in
- match (reg_name_matches_prefix
- _stringappend_335_) with
- | Some
- (_stringappend_336_,(existT _ _stringappend_337_ _)) =>
- returnm ((_stringappend_336_,
- build_ex
- _stringappend_337_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__363 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_337_ _) :=
- w__363
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_338_ :=
- string_drop _stringappend_335_
- (build_ex
- _stringappend_337_) in
- sep_matches_prefix _stringappend_338_ >>= fun w__364 : option ((unit * {n : Z & ArithFact (n >=
+ (build_ex (projT1 (string_length "mulw"))) in
+ (match (spc_matches_prefix _stringappend_224_) with
+ | Some (tt,(existT _ _stringappend_225_ _)) =>
+ returnm (tt, build_ex _stringappend_225_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_225_ _) =>
+ let _stringappend_226_ :=
+ string_drop _stringappend_224_
+ (build_ex _stringappend_225_) in
+ (match (reg_name_matches_prefix
+ _stringappend_226_) with
+ | Some (rd,(existT _ _stringappend_227_ _)) =>
+ returnm (rd, build_ex _stringappend_227_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_227_ _) =>
+ let _stringappend_228_ :=
+ string_drop _stringappend_226_
+ (build_ex _stringappend_227_) in
+ sep_matches_prefix _stringappend_228_ >>= fun w__364 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__364 with
- | Some
- (_stringappend_339_,(existT _ _stringappend_340_ _)) =>
- returnm ((_stringappend_339_,
- build_ex
- _stringappend_340_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__366 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_340_ _) :=
- w__366
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_341_ :=
- string_drop _stringappend_338_
- (build_ex
- _stringappend_340_) in
- match (reg_name_matches_prefix
- _stringappend_341_) with
- | Some
- (_stringappend_342_,(existT _ _stringappend_343_ _)) =>
- returnm ((_stringappend_342_,
- build_ex
- _stringappend_343_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__368 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_343_ _) :=
- w__368
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_344_ :=
- string_drop _stringappend_341_
- (build_ex
- _stringappend_343_) in
- sep_matches_prefix _stringappend_344_ >>= fun w__369 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__364 with
+ | Some (tt,(existT _ _stringappend_229_ _)) =>
+ returnm (tt, build_ex _stringappend_229_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_229_ _) =>
+ let _stringappend_230_ :=
+ string_drop _stringappend_228_
+ (build_ex _stringappend_229_) in
+ (match (reg_name_matches_prefix
+ _stringappend_230_) with
+ | Some (rs1,(existT _ _stringappend_231_ _)) =>
+ returnm (rs1, build_ex _stringappend_231_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_231_ _) =>
+ let _stringappend_232_ :=
+ string_drop _stringappend_230_
+ (build_ex _stringappend_231_) in
+ sep_matches_prefix _stringappend_232_ >>= fun w__369 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__369 with
- | Some
- (_stringappend_345_,(existT _ _stringappend_346_ _)) =>
- returnm ((_stringappend_345_,
- build_ex
- _stringappend_346_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__371 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_346_ _) :=
- w__371
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_347_ :=
- string_drop _stringappend_344_
- (build_ex
- _stringappend_346_) in
- match (reg_name_matches_prefix
- _stringappend_347_) with
- | Some
- (_stringappend_348_,(existT _ _stringappend_349_ _)) =>
- returnm ((_stringappend_348_,
- build_ex
- _stringappend_349_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__373 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_349_ _) :=
- w__373
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
+ (match w__369 with
+ | Some (tt,(existT _ _stringappend_233_ _)) =>
+ returnm (tt, build_ex _stringappend_233_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >= 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >= 0)}))) >>= fun '(tt, existT _ _stringappend_233_ _) =>
+ let _stringappend_234_ :=
+ string_drop _stringappend_232_
+ (build_ex _stringappend_233_) in
+ (match (reg_name_matches_prefix
+ _stringappend_234_) with
+ | Some (rs2,(existT _ _stringappend_235_ _)) =>
+ returnm (rs2, build_ex _stringappend_235_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_235_ _) =>
returnm ((match (string_drop
- _stringappend_347_
- (build_ex
- _stringappend_349_)) with
+ _stringappend_234_
+ (build_ex _stringappend_235_)) with
| s_ =>
- Some (MULW (rs2,rs1,rd),
- sub_nat
- (string_length arg_)
- (string_length s_))
+ Some
+ ((MULW
+ (rs2, rs1, rd), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >= 0)})))
else
@@ -26637,79 +24557,72 @@ Definition assembly_matches_prefix (arg_ : string)
(returnm ((string_startswith
_stringappend_0_ "div")
: bool))
- (let _stringappend_351_ :=
+ (let _stringappend_237_ :=
string_drop _stringappend_0_
- (string_length "div") in
+ (build_ex (projT1 (string_length "div"))) in
match (maybe_not_u_matches_prefix
- _stringappend_351_) with
- | Some
- (_stringappend_352_,(existT _ _stringappend_353_ _)) =>
- let _stringappend_354_ :=
- string_drop _stringappend_351_
- (build_ex
- _stringappend_353_) in
+ _stringappend_237_) with
+ | Some (s,(existT _ _stringappend_238_ _)) =>
+ let _stringappend_239_ :=
+ string_drop _stringappend_237_
+ (build_ex _stringappend_238_) in
and_boolM
(returnm ((string_startswith
- _stringappend_354_ "w")
+ _stringappend_239_ "w")
: bool))
- (let _stringappend_355_ :=
- string_drop _stringappend_354_
- (string_length "w") in
+ (let _stringappend_240_ :=
+ string_drop _stringappend_239_
+ (build_ex (projT1 (string_length
+ "w"))) in
match (spc_matches_prefix
- _stringappend_355_) with
+ _stringappend_240_) with
| Some
- (_stringappend_356_,(existT _ _stringappend_357_ _)) =>
- let _stringappend_358_ :=
- string_drop _stringappend_355_
- (build_ex
- _stringappend_357_) in
+ (tt,(existT _ _stringappend_241_ _)) =>
+ let _stringappend_242_ :=
+ string_drop _stringappend_240_
+ (build_ex _stringappend_241_) in
match (reg_name_matches_prefix
- _stringappend_358_) with
+ _stringappend_242_) with
| Some
- (_stringappend_359_,(existT _ _stringappend_360_ _)) =>
- let _stringappend_361_ :=
- string_drop _stringappend_358_
- (build_ex
- _stringappend_360_) in
+ (rd,(existT _ _stringappend_243_ _)) =>
+ let _stringappend_244_ :=
+ string_drop _stringappend_242_
+ (build_ex _stringappend_243_) in
sep_matches_prefix
- _stringappend_361_ >>= fun w__374 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_244_ >>= fun w__374 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__374 with
| Some
- (_stringappend_362_,(existT _ _stringappend_363_ _)) =>
- let _stringappend_364_ :=
+ (tt,(existT _ _stringappend_245_ _)) =>
+ let _stringappend_246_ :=
string_drop
- _stringappend_361_
- (build_ex
- _stringappend_363_) in
+ _stringappend_244_
+ (build_ex _stringappend_245_) in
match (reg_name_matches_prefix
- _stringappend_364_) with
+ _stringappend_246_) with
| Some
- (_stringappend_365_,(existT _ _stringappend_366_ _)) =>
- let _stringappend_367_ :=
+ (rs1,(existT _ _stringappend_247_ _)) =>
+ let _stringappend_248_ :=
string_drop
- _stringappend_364_
- (build_ex
- _stringappend_366_) in
+ _stringappend_246_
+ (build_ex _stringappend_247_) in
sep_matches_prefix
- _stringappend_367_ >>= fun w__375 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_248_ >>= fun w__375 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__375 with
| Some
- (_stringappend_368_,(existT _ _stringappend_369_ _)) =>
- let _stringappend_370_ :=
+ (tt,(existT _ _stringappend_249_ _)) =>
+ let _stringappend_250_ :=
string_drop
- _stringappend_367_
- (build_ex
- _stringappend_369_) in
+ _stringappend_248_
+ (build_ex _stringappend_249_) in
if ((match (reg_name_matches_prefix
- _stringappend_370_) with
+ _stringappend_250_) with
| Some
- (_stringappend_371_,(existT _ _stringappend_372_ _)) =>
+ (rs2,(existT _ _stringappend_251_ _)) =>
match (string_drop
- _stringappend_370_
- (build_ex
- _stringappend_372_)) with
+ _stringappend_250_
+ (build_ex _stringappend_251_)) with
| s_ =>
true
end
@@ -26757,162 +24670,119 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__382 : bool =>
(if (w__382) then
- let _stringappend_351_ :=
+ let _stringappend_237_ :=
string_drop _stringappend_0_
- (string_length "div") in
- match (maybe_not_u_matches_prefix
- _stringappend_351_) with
- | Some
- (_stringappend_352_,(existT _ _stringappend_353_ _)) =>
- returnm ((_stringappend_352_,
- build_ex
- _stringappend_353_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__384 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_353_ _) :=
- w__384
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_354_ :=
- string_drop _stringappend_351_
- (build_ex
- _stringappend_353_) in
- let _stringappend_355_ :=
- string_drop _stringappend_354_
- (string_length "w") in
- match (spc_matches_prefix
- _stringappend_355_) with
- | Some
- (_stringappend_356_,(existT _ _stringappend_357_ _)) =>
- returnm ((_stringappend_356_,
- build_ex
- _stringappend_357_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__386 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_357_ _) :=
- w__386
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_358_ :=
- string_drop _stringappend_355_
- (build_ex
- _stringappend_357_) in
- match (reg_name_matches_prefix
- _stringappend_358_) with
- | Some
- (_stringappend_359_,(existT _ _stringappend_360_ _)) =>
- returnm ((_stringappend_359_,
- build_ex
- _stringappend_360_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__388 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_360_ _) :=
- w__388
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_361_ :=
- string_drop _stringappend_358_
- (build_ex
- _stringappend_360_) in
- sep_matches_prefix _stringappend_361_ >>= fun w__389 : option ((unit * {n : Z & ArithFact (n >=
+ (build_ex (projT1 (string_length "div"))) in
+ (match (maybe_not_u_matches_prefix
+ _stringappend_237_) with
+ | Some (s,(existT _ _stringappend_238_ _)) =>
+ returnm (s, build_ex _stringappend_238_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(s, existT _ _stringappend_238_ _) =>
+ let _stringappend_239_ :=
+ string_drop _stringappend_237_
+ (build_ex _stringappend_238_) in
+ let _stringappend_240_ :=
+ string_drop _stringappend_239_
+ (build_ex (projT1 (string_length "w"))) in
+ (match (spc_matches_prefix
+ _stringappend_240_) with
+ | Some
+ (tt,(existT _ _stringappend_241_ _)) =>
+ returnm (tt, build_ex _stringappend_241_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_241_ _) =>
+ let _stringappend_242_ :=
+ string_drop _stringappend_240_
+ (build_ex _stringappend_241_) in
+ (match (reg_name_matches_prefix
+ _stringappend_242_) with
+ | Some
+ (rd,(existT _ _stringappend_243_ _)) =>
+ returnm (rd, build_ex _stringappend_243_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_243_ _) =>
+ let _stringappend_244_ :=
+ string_drop _stringappend_242_
+ (build_ex _stringappend_243_) in
+ sep_matches_prefix _stringappend_244_ >>= fun w__389 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__389 with
- | Some
- (_stringappend_362_,(existT _ _stringappend_363_ _)) =>
- returnm ((_stringappend_362_,
- build_ex
- _stringappend_363_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__391 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_363_ _) :=
- w__391
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_364_ :=
- string_drop _stringappend_361_
- (build_ex
- _stringappend_363_) in
- match (reg_name_matches_prefix
- _stringappend_364_) with
- | Some
- (_stringappend_365_,(existT _ _stringappend_366_ _)) =>
- returnm ((_stringappend_365_,
- build_ex
- _stringappend_366_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__393 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_366_ _) :=
- w__393
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_367_ :=
- string_drop _stringappend_364_
- (build_ex
- _stringappend_366_) in
- sep_matches_prefix _stringappend_367_ >>= fun w__394 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__389 with
+ | Some
+ (tt,(existT _ _stringappend_245_ _)) =>
+ returnm (tt, build_ex _stringappend_245_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_245_ _) =>
+ let _stringappend_246_ :=
+ string_drop _stringappend_244_
+ (build_ex _stringappend_245_) in
+ (match (reg_name_matches_prefix
+ _stringappend_246_) with
+ | Some
+ (rs1,(existT _ _stringappend_247_ _)) =>
+ returnm (rs1, build_ex _stringappend_247_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_247_ _) =>
+ let _stringappend_248_ :=
+ string_drop _stringappend_246_
+ (build_ex _stringappend_247_) in
+ sep_matches_prefix _stringappend_248_ >>= fun w__394 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__394 with
- | Some
- (_stringappend_368_,(existT _ _stringappend_369_ _)) =>
- returnm ((_stringappend_368_,
- build_ex
- _stringappend_369_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >= 0)}))
- end >>= fun w__396 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_369_ _) :=
- w__396
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_370_ :=
- string_drop _stringappend_367_
- (build_ex
- _stringappend_369_) in
- match (reg_name_matches_prefix
- _stringappend_370_) with
- | Some
- (_stringappend_371_,(existT _ _stringappend_372_ _)) =>
- returnm ((_stringappend_371_,
- build_ex
- _stringappend_372_)
- : (mword 5 * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__398 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_372_ _) :=
- w__398
- : (mword 5 * {n : Z & ArithFact (n >= 0)}) in
+ (match w__394 with
+ | Some
+ (tt,(existT _ _stringappend_249_ _)) =>
+ returnm (tt, build_ex _stringappend_249_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_249_ _) =>
+ let _stringappend_250_ :=
+ string_drop _stringappend_248_
+ (build_ex _stringappend_249_) in
+ (match (reg_name_matches_prefix
+ _stringappend_250_) with
+ | Some
+ (rs2,(existT _ _stringappend_251_ _)) =>
+ returnm (rs2, build_ex _stringappend_251_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_251_ _) =>
returnm ((match (string_drop
- _stringappend_370_
- (build_ex
- _stringappend_372_)) with
+ _stringappend_250_
+ (build_ex _stringappend_251_)) with
| s_ =>
- Some (DIVW (rs2,rs1,rd,s),
- sub_nat
- (string_length arg_)
- (string_length s_))
+ Some
+ ((DIVW
+ (rs2, rs1, rd, s), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -26921,80 +24791,75 @@ Definition assembly_matches_prefix (arg_ : string)
(returnm ((string_startswith
_stringappend_0_ "rem")
: bool))
- (let _stringappend_374_ :=
+ (let _stringappend_253_ :=
string_drop _stringappend_0_
- (string_length "rem") in
+ (build_ex (projT1 (string_length
+ "rem"))) in
match (maybe_not_u_matches_prefix
- _stringappend_374_) with
+ _stringappend_253_) with
| Some
- (_stringappend_375_,(existT _ _stringappend_376_ _)) =>
- let _stringappend_377_ :=
- string_drop _stringappend_374_
- (build_ex
- _stringappend_376_) in
+ (s,(existT _ _stringappend_254_ _)) =>
+ let _stringappend_255_ :=
+ string_drop _stringappend_253_
+ (build_ex _stringappend_254_) in
and_boolM
(returnm ((string_startswith
- _stringappend_377_ "w")
+ _stringappend_255_ "w")
: bool))
- (let _stringappend_378_ :=
- string_drop _stringappend_377_
- (string_length "w") in
+ (let _stringappend_256_ :=
+ string_drop _stringappend_255_
+ (build_ex (projT1 (string_length
+ "w"))) in
match (spc_matches_prefix
- _stringappend_378_) with
+ _stringappend_256_) with
| Some
- (_stringappend_379_,(existT _ _stringappend_380_ _)) =>
- let _stringappend_381_ :=
- string_drop _stringappend_378_
- (build_ex
- _stringappend_380_) in
+ (tt,(existT _ _stringappend_257_ _)) =>
+ let _stringappend_258_ :=
+ string_drop _stringappend_256_
+ (build_ex _stringappend_257_) in
match (reg_name_matches_prefix
- _stringappend_381_) with
+ _stringappend_258_) with
| Some
- (_stringappend_382_,(existT _ _stringappend_383_ _)) =>
- let _stringappend_384_ :=
+ (rd,(existT _ _stringappend_259_ _)) =>
+ let _stringappend_260_ :=
string_drop
- _stringappend_381_
- (build_ex
- _stringappend_383_) in
+ _stringappend_258_
+ (build_ex _stringappend_259_) in
sep_matches_prefix
- _stringappend_384_ >>= fun w__399 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_260_ >>= fun w__399 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__399 with
| Some
- (_stringappend_385_,(existT _ _stringappend_386_ _)) =>
- let _stringappend_387_ :=
+ (tt,(existT _ _stringappend_261_ _)) =>
+ let _stringappend_262_ :=
string_drop
- _stringappend_384_
- (build_ex
- _stringappend_386_) in
+ _stringappend_260_
+ (build_ex _stringappend_261_) in
match (reg_name_matches_prefix
- _stringappend_387_) with
+ _stringappend_262_) with
| Some
- (_stringappend_388_,(existT _ _stringappend_389_ _)) =>
- let _stringappend_390_ :=
+ (rs1,(existT _ _stringappend_263_ _)) =>
+ let _stringappend_264_ :=
string_drop
- _stringappend_387_
- (build_ex
- _stringappend_389_) in
+ _stringappend_262_
+ (build_ex _stringappend_263_) in
sep_matches_prefix
- _stringappend_390_ >>= fun w__400 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_264_ >>= fun w__400 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__400 with
| Some
- (_stringappend_391_,(existT _ _stringappend_392_ _)) =>
- let _stringappend_393_ :=
+ (tt,(existT _ _stringappend_265_ _)) =>
+ let _stringappend_266_ :=
string_drop
- _stringappend_390_
- (build_ex
- _stringappend_392_) in
+ _stringappend_264_
+ (build_ex _stringappend_265_) in
if ((match (reg_name_matches_prefix
- _stringappend_393_) with
+ _stringappend_266_) with
| Some
- (_stringappend_394_,(existT _ _stringappend_395_ _)) =>
+ (rs2,(existT _ _stringappend_267_ _)) =>
match (string_drop
- _stringappend_393_
- (build_ex
- _stringappend_395_)) with
+ _stringappend_266_
+ (build_ex _stringappend_267_)) with
| s_ =>
true
end
@@ -27045,172 +24910,121 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__407 : bool =>
(if (w__407) then
- let _stringappend_374_ :=
+ let _stringappend_253_ :=
string_drop _stringappend_0_
- (string_length "rem") in
- match (maybe_not_u_matches_prefix
- _stringappend_374_) with
- | Some
- (_stringappend_375_,(existT _ _stringappend_376_ _)) =>
- returnm ((_stringappend_375_,
- build_ex
- _stringappend_376_)
- : (bool * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__409 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_376_ _) :=
- w__409
- : (bool * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_377_ :=
- string_drop _stringappend_374_
- (build_ex
- _stringappend_376_) in
- let _stringappend_378_ :=
- string_drop _stringappend_377_
- (string_length "w") in
- match (spc_matches_prefix
- _stringappend_378_) with
- | Some
- (_stringappend_379_,(existT _ _stringappend_380_ _)) =>
- returnm ((_stringappend_379_,
- build_ex
- _stringappend_380_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__411 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_380_ _) :=
- w__411
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_381_ :=
- string_drop _stringappend_378_
- (build_ex
- _stringappend_380_) in
- match (reg_name_matches_prefix
- _stringappend_381_) with
- | Some
- (_stringappend_382_,(existT _ _stringappend_383_ _)) =>
- returnm ((_stringappend_382_,
- build_ex
- _stringappend_383_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__413 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_383_ _) :=
- w__413
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_384_ :=
- string_drop _stringappend_381_
- (build_ex
- _stringappend_383_) in
- sep_matches_prefix _stringappend_384_ >>= fun w__414 : option ((unit * {n : Z & ArithFact (n >=
+ (build_ex (projT1 (string_length
+ "rem"))) in
+ (match (maybe_not_u_matches_prefix
+ _stringappend_253_) with
+ | Some
+ (s,(existT _ _stringappend_254_ _)) =>
+ returnm (s, build_ex _stringappend_254_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(s, existT _ _stringappend_254_ _) =>
+ let _stringappend_255_ :=
+ string_drop _stringappend_253_
+ (build_ex _stringappend_254_) in
+ let _stringappend_256_ :=
+ string_drop _stringappend_255_
+ (build_ex (projT1 (string_length "w"))) in
+ (match (spc_matches_prefix
+ _stringappend_256_) with
+ | Some
+ (tt,(existT _ _stringappend_257_ _)) =>
+ returnm (tt, build_ex _stringappend_257_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_257_ _) =>
+ let _stringappend_258_ :=
+ string_drop _stringappend_256_
+ (build_ex _stringappend_257_) in
+ (match (reg_name_matches_prefix
+ _stringappend_258_) with
+ | Some
+ (rd,(existT _ _stringappend_259_ _)) =>
+ returnm (rd, build_ex _stringappend_259_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_259_ _) =>
+ let _stringappend_260_ :=
+ string_drop _stringappend_258_
+ (build_ex _stringappend_259_) in
+ sep_matches_prefix _stringappend_260_ >>= fun w__414 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__414 with
- | Some
- (_stringappend_385_,(existT _ _stringappend_386_ _)) =>
- returnm ((_stringappend_385_,
- build_ex
- _stringappend_386_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__416 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_386_ _) :=
- w__416
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_387_ :=
- string_drop _stringappend_384_
- (build_ex
- _stringappend_386_) in
- match (reg_name_matches_prefix
- _stringappend_387_) with
- | Some
- (_stringappend_388_,(existT _ _stringappend_389_ _)) =>
- returnm ((_stringappend_388_,
- build_ex
- _stringappend_389_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__418 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_389_ _) :=
- w__418
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_390_ :=
- string_drop _stringappend_387_
- (build_ex
- _stringappend_389_) in
- sep_matches_prefix _stringappend_390_ >>= fun w__419 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__414 with
+ | Some
+ (tt,(existT _ _stringappend_261_ _)) =>
+ returnm (tt, build_ex _stringappend_261_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_261_ _) =>
+ let _stringappend_262_ :=
+ string_drop _stringappend_260_
+ (build_ex _stringappend_261_) in
+ (match (reg_name_matches_prefix
+ _stringappend_262_) with
+ | Some
+ (rs1,(existT _ _stringappend_263_ _)) =>
+ returnm (rs1, build_ex _stringappend_263_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_263_ _) =>
+ let _stringappend_264_ :=
+ string_drop _stringappend_262_
+ (build_ex _stringappend_263_) in
+ sep_matches_prefix _stringappend_264_ >>= fun w__419 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__419 with
- | Some
- (_stringappend_391_,(existT _ _stringappend_392_ _)) =>
- returnm ((_stringappend_391_,
- build_ex
- _stringappend_392_)
- : (unit * {n : Z & ArithFact (n >= 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__421 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_392_ _) :=
- w__421
- : (unit * {n : Z & ArithFact (n >= 0)}) in
- let _stringappend_393_ :=
- string_drop _stringappend_390_
- (build_ex
- _stringappend_392_) in
- match (reg_name_matches_prefix
- _stringappend_393_) with
- | Some
- (_stringappend_394_,(existT _ _stringappend_395_ _)) =>
- returnm ((_stringappend_394_,
- build_ex
- _stringappend_395_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__423 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_395_ _) :=
- w__423
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
+ (match w__419 with
+ | Some
+ (tt,(existT _ _stringappend_265_ _)) =>
+ returnm (tt, build_ex _stringappend_265_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_265_ _) =>
+ let _stringappend_266_ :=
+ string_drop _stringappend_264_
+ (build_ex _stringappend_265_) in
+ (match (reg_name_matches_prefix
+ _stringappend_266_) with
+ | Some
+ (rs2,(existT _ _stringappend_267_ _)) =>
+ returnm (rs2, build_ex _stringappend_267_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_267_ _) =>
returnm ((match (string_drop
- _stringappend_393_
- (build_ex
- _stringappend_395_)) with
+ _stringappend_266_
+ (build_ex _stringappend_267_)) with
| s_ =>
- Some (REMW (rs2,rs1,rd,s),
- sub_nat
- (string_length arg_)
- (string_length s_))
+ Some
+ ((REMW
+ (rs2, rs1, rd, s), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -27219,48 +25033,45 @@ Definition assembly_matches_prefix (arg_ : string)
(returnm ((string_startswith
_stringappend_0_ "fence")
: bool))
- (let _stringappend_397_ :=
+ (let _stringappend_269_ :=
string_drop _stringappend_0_
- (string_length "fence") in
+ (build_ex (projT1 (string_length
+ "fence"))) in
match (spc_matches_prefix
- _stringappend_397_) with
+ _stringappend_269_) with
| Some
- (_stringappend_398_,(existT _ _stringappend_399_ _)) =>
- let _stringappend_400_ :=
- string_drop _stringappend_397_
- (build_ex
- _stringappend_399_) in
+ (tt,(existT _ _stringappend_270_ _)) =>
+ let _stringappend_271_ :=
+ string_drop _stringappend_269_
+ (build_ex _stringappend_270_) in
fence_bits_matches_prefix
- _stringappend_400_ >>= fun w__424 : option ((mword 4 * {n : Z & ArithFact (n >=
+ _stringappend_271_ >>= fun w__424 : option ((mword 4 * {n : Z & ArithFact (n >=
0)})) =>
match w__424 with
| Some
- (_stringappend_401_,(existT _ _stringappend_402_ _)) =>
- let _stringappend_403_ :=
- string_drop _stringappend_400_
- (build_ex
- _stringappend_402_) in
+ (pred,(existT _ _stringappend_272_ _)) =>
+ let _stringappend_273_ :=
+ string_drop _stringappend_271_
+ (build_ex _stringappend_272_) in
sep_matches_prefix
- _stringappend_403_ >>= fun w__425 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_273_ >>= fun w__425 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__425 with
| Some
- (_stringappend_404_,(existT _ _stringappend_405_ _)) =>
- let _stringappend_406_ :=
+ (tt,(existT _ _stringappend_274_ _)) =>
+ let _stringappend_275_ :=
string_drop
- _stringappend_403_
- (build_ex
- _stringappend_405_) in
+ _stringappend_273_
+ (build_ex _stringappend_274_) in
fence_bits_matches_prefix
- _stringappend_406_ >>= fun w__426 : option ((mword 4 * {n : Z & ArithFact (n >=
+ _stringappend_275_ >>= fun w__426 : option ((mword 4 * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__426 with
| Some
- (_stringappend_407_,(existT _ _stringappend_408_ _)) =>
+ (succ,(existT _ _stringappend_276_ _)) =>
match (string_drop
- _stringappend_406_
- (build_ex
- _stringappend_408_)) with
+ _stringappend_275_
+ (build_ex _stringappend_276_)) with
| s_ =>
true
end
@@ -27287,113 +25098,79 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__430 : bool =>
(if (w__430) then
- let _stringappend_397_ :=
+ let _stringappend_269_ :=
string_drop _stringappend_0_
- (string_length "fence") in
- match (spc_matches_prefix
- _stringappend_397_) with
- | Some
- (_stringappend_398_,(existT _ _stringappend_399_ _)) =>
- returnm ((_stringappend_398_,
- build_ex
- _stringappend_399_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__432 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_399_ _) :=
- w__432
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_400_ :=
- string_drop _stringappend_397_
- (build_ex
- _stringappend_399_) in
+ (build_ex (projT1 (string_length
+ "fence"))) in
+ (match (spc_matches_prefix
+ _stringappend_269_) with
+ | Some
+ (tt,(existT _ _stringappend_270_ _)) =>
+ returnm (tt, build_ex _stringappend_270_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_270_ _) =>
+ let _stringappend_271_ :=
+ string_drop _stringappend_269_
+ (build_ex _stringappend_270_) in
fence_bits_matches_prefix
- _stringappend_400_ >>= fun w__433 : option ((mword 4 * {n : Z & ArithFact (n >=
+ _stringappend_271_ >>= fun w__433 : option ((mword 4 * {n : Z & ArithFact (n >=
0)})) =>
- match w__433 with
- | Some
- (_stringappend_401_,(existT _ _stringappend_402_ _)) =>
- returnm ((_stringappend_401_,
- build_ex
- _stringappend_402_)
- : (mword 4 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 4 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__435 : (mword 4 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(pred, existT _ _stringappend_402_ _) :=
- w__435
- : (mword 4 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_403_ :=
- string_drop _stringappend_400_
- (build_ex
- _stringappend_402_) in
- sep_matches_prefix _stringappend_403_ >>= fun w__436 : option ((unit * {n : Z & ArithFact (n >=
+ (match w__433 with
+ | Some
+ (pred,(existT _ _stringappend_272_ _)) =>
+ returnm (pred, build_ex _stringappend_272_)
+ | _ =>
+ exit tt
+ : M ((mword 4 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 4 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(pred, existT _ _stringappend_272_ _) =>
+ let _stringappend_273_ :=
+ string_drop _stringappend_271_
+ (build_ex _stringappend_272_) in
+ sep_matches_prefix _stringappend_273_ >>= fun w__436 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__436 with
- | Some
- (_stringappend_404_,(existT _ _stringappend_405_ _)) =>
- returnm ((_stringappend_404_,
- build_ex
- _stringappend_405_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__438 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_405_ _) :=
- w__438
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_406_ :=
- string_drop _stringappend_403_
- (build_ex
- _stringappend_405_) in
+ (match w__436 with
+ | Some
+ (tt,(existT _ _stringappend_274_ _)) =>
+ returnm (tt, build_ex _stringappend_274_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_274_ _) =>
+ let _stringappend_275_ :=
+ string_drop _stringappend_273_
+ (build_ex _stringappend_274_) in
fence_bits_matches_prefix
- _stringappend_406_ >>= fun w__439 : option ((mword 4 * {n : Z & ArithFact (n >=
+ _stringappend_275_ >>= fun w__439 : option ((mword 4 * {n : Z & ArithFact (n >=
0)})) =>
- match w__439 with
- | Some
- (_stringappend_407_,(existT _ _stringappend_408_ _)) =>
- returnm ((_stringappend_407_,
- build_ex
- _stringappend_408_)
- : (mword 4 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 4 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__441 : (mword 4 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(succ, existT _ _stringappend_408_ _) :=
- w__441
- : (mword 4 * {n : Z & ArithFact (n >=
- 0)}) in
+ (match w__439 with
+ | Some
+ (succ,(existT _ _stringappend_276_ _)) =>
+ returnm (succ, build_ex _stringappend_276_)
+ | _ =>
+ exit tt
+ : M ((mword 4 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 4 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(succ, existT _ _stringappend_276_ _) =>
returnm ((match (string_drop
- _stringappend_406_
- (build_ex
- _stringappend_408_)) with
+ _stringappend_275_
+ (build_ex _stringappend_276_)) with
| s_ =>
- Some (FENCE (pred,succ),
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((FENCE
+ (pred, succ), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -27403,21 +25180,22 @@ Definition assembly_matches_prefix (arg_ : string)
"fence.i")
(match (string_drop
_stringappend_0_
- (string_length
- "fence.i")) with
+ (build_ex (projT1 (string_length
+ "fence.i")))) with
| s_ => true
end))) then
returnm ((match (string_drop
_stringappend_0_
- (string_length
- "fence.i")) with
+ (build_ex (projT1 (string_length
+ "fence.i")))) with
| s_ =>
- Some (FENCEI tt,
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((FENCEI
+ (tt), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -27426,21 +25204,22 @@ Definition assembly_matches_prefix (arg_ : string)
_stringappend_0_ "ecall")
(match (string_drop
_stringappend_0_
- (string_length
- "ecall")) with
+ (build_ex (projT1 (string_length
+ "ecall")))) with
| s_ => true
end))) then
returnm ((match (string_drop
_stringappend_0_
- (string_length
- "ecall")) with
+ (build_ex (projT1 (string_length
+ "ecall")))) with
| s_ =>
- Some (ECALL tt,
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((ECALL
+ (tt), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -27449,21 +25228,22 @@ Definition assembly_matches_prefix (arg_ : string)
_stringappend_0_ "mret")
(match (string_drop
_stringappend_0_
- (string_length
- "mret")) with
+ (build_ex (projT1 (string_length
+ "mret")))) with
| s_ => true
end))) then
returnm ((match (string_drop
_stringappend_0_
- (string_length
- "mret")) with
+ (build_ex (projT1 (string_length
+ "mret")))) with
| s_ =>
- Some (MRET tt,
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((MRET
+ (tt), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -27472,21 +25252,22 @@ Definition assembly_matches_prefix (arg_ : string)
_stringappend_0_ "sret")
(match (string_drop
_stringappend_0_
- (string_length
- "sret")) with
+ (build_ex (projT1 (string_length
+ "sret")))) with
| s_ => true
end))) then
returnm ((match (string_drop
_stringappend_0_
- (string_length
- "sret")) with
+ (build_ex (projT1 (string_length
+ "sret")))) with
| s_ =>
- Some (SRET tt,
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((SRET
+ (tt), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -27496,21 +25277,22 @@ Definition assembly_matches_prefix (arg_ : string)
"ebreak")
(match (string_drop
_stringappend_0_
- (string_length
- "ebreak")) with
+ (build_ex (projT1 (string_length
+ "ebreak")))) with
| s_ => true
end))) then
returnm ((match (string_drop
_stringappend_0_
- (string_length
- "ebreak")) with
+ (build_ex (projT1 (string_length
+ "ebreak")))) with
| s_ =>
- Some (EBREAK tt,
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((EBREAK
+ (tt), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -27519,21 +25301,22 @@ Definition assembly_matches_prefix (arg_ : string)
_stringappend_0_ "wfi")
(match (string_drop
_stringappend_0_
- (string_length
- "wfi")) with
+ (build_ex (projT1 (string_length
+ "wfi")))) with
| s_ => true
end))) then
returnm ((match (string_drop
_stringappend_0_
- (string_length
- "wfi")) with
+ (build_ex (projT1 (string_length
+ "wfi")))) with
| s_ =>
- Some (WFI tt,
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((WFI
+ (tt), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -27543,45 +25326,42 @@ Definition assembly_matches_prefix (arg_ : string)
_stringappend_0_
"sfence.vma")
: bool))
- (let _stringappend_416_ :=
+ (let _stringappend_284_ :=
string_drop _stringappend_0_
- (string_length "sfence.vma") in
+ (build_ex (projT1 (string_length
+ "sfence.vma"))) in
match (spc_matches_prefix
- _stringappend_416_) with
+ _stringappend_284_) with
| Some
- (_stringappend_417_,(existT _ _stringappend_418_ _)) =>
- let _stringappend_419_ :=
- string_drop _stringappend_416_
- (build_ex
- _stringappend_418_) in
+ (tt,(existT _ _stringappend_285_ _)) =>
+ let _stringappend_286_ :=
+ string_drop _stringappend_284_
+ (build_ex _stringappend_285_) in
match (reg_name_matches_prefix
- _stringappend_419_) with
+ _stringappend_286_) with
| Some
- (_stringappend_420_,(existT _ _stringappend_421_ _)) =>
- let _stringappend_422_ :=
+ (rs1,(existT _ _stringappend_287_ _)) =>
+ let _stringappend_288_ :=
string_drop
- _stringappend_419_
- (build_ex
- _stringappend_421_) in
+ _stringappend_286_
+ (build_ex _stringappend_287_) in
sep_matches_prefix
- _stringappend_422_ >>= fun w__442 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_288_ >>= fun w__442 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__442 with
| Some
- (_stringappend_423_,(existT _ _stringappend_424_ _)) =>
- let _stringappend_425_ :=
+ (tt,(existT _ _stringappend_289_ _)) =>
+ let _stringappend_290_ :=
string_drop
- _stringappend_422_
- (build_ex
- _stringappend_424_) in
+ _stringappend_288_
+ (build_ex _stringappend_289_) in
if ((match (reg_name_matches_prefix
- _stringappend_425_) with
+ _stringappend_290_) with
| Some
- (_stringappend_426_,(existT _ _stringappend_427_ _)) =>
+ (rs2,(existT _ _stringappend_291_ _)) =>
match (string_drop
- _stringappend_425_
- (build_ex
- _stringappend_427_)) with
+ _stringappend_290_
+ (build_ex _stringappend_291_)) with
| s_ =>
true
end
@@ -27609,110 +25389,76 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__445 : bool =>
(if (w__445) then
- let _stringappend_416_ :=
+ let _stringappend_284_ :=
string_drop _stringappend_0_
- (string_length "sfence.vma") in
- match (spc_matches_prefix
- _stringappend_416_) with
- | Some
- (_stringappend_417_,(existT _ _stringappend_418_ _)) =>
- returnm ((_stringappend_417_,
- build_ex
- _stringappend_418_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__447 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_418_ _) :=
- w__447
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_419_ :=
- string_drop _stringappend_416_
- (build_ex
- _stringappend_418_) in
- match (reg_name_matches_prefix
- _stringappend_419_) with
- | Some
- (_stringappend_420_,(existT _ _stringappend_421_ _)) =>
- returnm ((_stringappend_420_,
- build_ex
- _stringappend_421_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__449 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_421_ _) :=
- w__449
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_422_ :=
- string_drop _stringappend_419_
- (build_ex
- _stringappend_421_) in
+ (build_ex (projT1 (string_length
+ "sfence.vma"))) in
+ (match (spc_matches_prefix
+ _stringappend_284_) with
+ | Some
+ (tt,(existT _ _stringappend_285_ _)) =>
+ returnm (tt, build_ex _stringappend_285_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_285_ _) =>
+ let _stringappend_286_ :=
+ string_drop _stringappend_284_
+ (build_ex _stringappend_285_) in
+ (match (reg_name_matches_prefix
+ _stringappend_286_) with
+ | Some
+ (rs1,(existT _ _stringappend_287_ _)) =>
+ returnm (rs1, build_ex _stringappend_287_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_287_ _) =>
+ let _stringappend_288_ :=
+ string_drop _stringappend_286_
+ (build_ex _stringappend_287_) in
sep_matches_prefix
- _stringappend_422_ >>= fun w__450 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_288_ >>= fun w__450 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__450 with
- | Some
- (_stringappend_423_,(existT _ _stringappend_424_ _)) =>
- returnm ((_stringappend_423_,
- build_ex
- _stringappend_424_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__452 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_424_ _) :=
- w__452
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_425_ :=
- string_drop _stringappend_422_
- (build_ex
- _stringappend_424_) in
- match (reg_name_matches_prefix
- _stringappend_425_) with
- | Some
- (_stringappend_426_,(existT _ _stringappend_427_ _)) =>
- returnm ((_stringappend_426_,
- build_ex
- _stringappend_427_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__454 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_427_ _) :=
- w__454
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
+ (match w__450 with
+ | Some
+ (tt,(existT _ _stringappend_289_ _)) =>
+ returnm (tt, build_ex _stringappend_289_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_289_ _) =>
+ let _stringappend_290_ :=
+ string_drop _stringappend_288_
+ (build_ex _stringappend_289_) in
+ (match (reg_name_matches_prefix
+ _stringappend_290_) with
+ | Some
+ (rs2,(existT _ _stringappend_291_ _)) =>
+ returnm (rs2, build_ex _stringappend_291_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_291_ _) =>
returnm ((match (string_drop
- _stringappend_425_
- (build_ex
- _stringappend_427_)) with
+ _stringappend_290_
+ (build_ex _stringappend_291_)) with
| s_ =>
- Some (SFENCE_VMA (rs1,rs2),
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((SFENCE_VMA
+ (rs1, rs2), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -27722,74 +25468,68 @@ Definition assembly_matches_prefix (arg_ : string)
_stringappend_0_
"lr.")
: bool))
- (let _stringappend_429_ :=
+ (let _stringappend_293_ :=
string_drop _stringappend_0_
- (string_length "lr.") in
+ (build_ex (projT1 (string_length
+ "lr."))) in
match (maybe_aq_matches_prefix
- _stringappend_429_) with
+ _stringappend_293_) with
| Some
- (_stringappend_430_,(existT _ _stringappend_431_ _)) =>
- let _stringappend_432_ :=
+ (aq,(existT _ _stringappend_294_ _)) =>
+ let _stringappend_295_ :=
string_drop
- _stringappend_429_
- (build_ex
- _stringappend_431_) in
+ _stringappend_293_
+ (build_ex _stringappend_294_) in
match (maybe_rl_matches_prefix
- _stringappend_432_) with
+ _stringappend_295_) with
| Some
- (_stringappend_433_,(existT _ _stringappend_434_ _)) =>
- let _stringappend_435_ :=
+ (rl,(existT _ _stringappend_296_ _)) =>
+ let _stringappend_297_ :=
string_drop
- _stringappend_432_
- (build_ex
- _stringappend_434_) in
+ _stringappend_295_
+ (build_ex _stringappend_296_) in
match (size_mnemonic_matches_prefix
- _stringappend_435_) with
+ _stringappend_297_) with
| Some
- (_stringappend_436_,(existT _ _stringappend_437_ _)) =>
- let _stringappend_438_ :=
+ (size,(existT _ _stringappend_298_ _)) =>
+ let _stringappend_299_ :=
string_drop
- _stringappend_435_
- (build_ex
- _stringappend_437_) in
+ _stringappend_297_
+ (build_ex _stringappend_298_) in
match (spc_matches_prefix
- _stringappend_438_) with
+ _stringappend_299_) with
| Some
- (_stringappend_439_,(existT _ _stringappend_440_ _)) =>
- let _stringappend_441_ :=
+ (tt,(existT _ _stringappend_300_ _)) =>
+ let _stringappend_301_ :=
string_drop
- _stringappend_438_
- (build_ex
- _stringappend_440_) in
+ _stringappend_299_
+ (build_ex _stringappend_300_) in
match (reg_name_matches_prefix
- _stringappend_441_) with
+ _stringappend_301_) with
| Some
- (_stringappend_442_,(existT _ _stringappend_443_ _)) =>
- let _stringappend_444_ :=
+ (rd,(existT _ _stringappend_302_ _)) =>
+ let _stringappend_303_ :=
string_drop
- _stringappend_441_
- (build_ex
- _stringappend_443_) in
+ _stringappend_301_
+ (build_ex _stringappend_302_) in
sep_matches_prefix
- _stringappend_444_ >>= fun w__455 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_303_ >>= fun w__455 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if ((match w__455 with
| Some
- (_stringappend_445_,(existT _ _stringappend_446_ _)) =>
- let _stringappend_447_ :=
+ (tt,(existT _ _stringappend_304_ _)) =>
+ let _stringappend_305_ :=
string_drop
- _stringappend_444_
- (build_ex
- _stringappend_446_) in
+ _stringappend_303_
+ (build_ex _stringappend_304_) in
if
((match (reg_name_matches_prefix
- _stringappend_447_) with
+ _stringappend_305_) with
| Some
- (_stringappend_448_,(existT _ _stringappend_449_ _)) =>
+ (rs1,(existT _ _stringappend_306_ _)) =>
match (string_drop
- _stringappend_447_
- (build_ex
- _stringappend_449_)) with
+ _stringappend_305_
+ (build_ex _stringappend_306_)) with
| s_ =>
true
end
@@ -27847,185 +25587,125 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__461 : bool =>
(if (w__461) then
- let _stringappend_429_ :=
+ let _stringappend_293_ :=
string_drop _stringappend_0_
- (string_length "lr.") in
- match (maybe_aq_matches_prefix
- _stringappend_429_) with
- | Some
- (_stringappend_430_,(existT _ _stringappend_431_ _)) =>
- returnm ((_stringappend_430_,
- build_ex
- _stringappend_431_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__463 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(aq, existT _ _stringappend_431_ _) :=
- w__463
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_432_ :=
+ (build_ex (projT1 (string_length
+ "lr."))) in
+ (match (maybe_aq_matches_prefix
+ _stringappend_293_) with
+ | Some
+ (aq,(existT _ _stringappend_294_ _)) =>
+ returnm (aq, build_ex _stringappend_294_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(aq, existT _ _stringappend_294_ _) =>
+ let _stringappend_295_ :=
string_drop
- _stringappend_429_
- (build_ex
- _stringappend_431_) in
- match (maybe_rl_matches_prefix
- _stringappend_432_) with
- | Some
- (_stringappend_433_,(existT _ _stringappend_434_ _)) =>
- returnm ((_stringappend_433_,
- build_ex
- _stringappend_434_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__465 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rl, existT _ _stringappend_434_ _) :=
- w__465
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_435_ :=
+ _stringappend_293_
+ (build_ex _stringappend_294_) in
+ (match (maybe_rl_matches_prefix
+ _stringappend_295_) with
+ | Some
+ (rl,(existT _ _stringappend_296_ _)) =>
+ returnm (rl, build_ex _stringappend_296_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rl, existT _ _stringappend_296_ _) =>
+ let _stringappend_297_ :=
string_drop
- _stringappend_432_
- (build_ex
- _stringappend_434_) in
- match (size_mnemonic_matches_prefix
- _stringappend_435_) with
- | Some
- (_stringappend_436_,(existT _ _stringappend_437_ _)) =>
- returnm ((_stringappend_436_,
- build_ex
- _stringappend_437_)
- : (word_width * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((word_width * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__467 : (word_width * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(size, existT _ _stringappend_437_ _) :=
- w__467
- : (word_width * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_438_ :=
+ _stringappend_295_
+ (build_ex _stringappend_296_) in
+ (match (size_mnemonic_matches_prefix
+ _stringappend_297_) with
+ | Some
+ (size,(existT _ _stringappend_298_ _)) =>
+ returnm (size, build_ex _stringappend_298_)
+ | _ =>
+ exit tt
+ : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(size, existT _ _stringappend_298_ _) =>
+ let _stringappend_299_ :=
string_drop
- _stringappend_435_
- (build_ex
- _stringappend_437_) in
- match (spc_matches_prefix
- _stringappend_438_) with
- | Some
- (_stringappend_439_,(existT _ _stringappend_440_ _)) =>
- returnm ((_stringappend_439_,
- build_ex
- _stringappend_440_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__469 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_440_ _) :=
- w__469
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_441_ :=
+ _stringappend_297_
+ (build_ex _stringappend_298_) in
+ (match (spc_matches_prefix
+ _stringappend_299_) with
+ | Some
+ (tt,(existT _ _stringappend_300_ _)) =>
+ returnm (tt, build_ex _stringappend_300_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_300_ _) =>
+ let _stringappend_301_ :=
string_drop
- _stringappend_438_
- (build_ex
- _stringappend_440_) in
- match (reg_name_matches_prefix
- _stringappend_441_) with
- | Some
- (_stringappend_442_,(existT _ _stringappend_443_ _)) =>
- returnm ((_stringappend_442_,
- build_ex
- _stringappend_443_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__471 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_443_ _) :=
- w__471
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_444_ :=
+ _stringappend_299_
+ (build_ex _stringappend_300_) in
+ (match (reg_name_matches_prefix
+ _stringappend_301_) with
+ | Some
+ (rd,(existT _ _stringappend_302_ _)) =>
+ returnm (rd, build_ex _stringappend_302_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_302_ _) =>
+ let _stringappend_303_ :=
string_drop
- _stringappend_441_
- (build_ex
- _stringappend_443_) in
+ _stringappend_301_
+ (build_ex _stringappend_302_) in
sep_matches_prefix
- _stringappend_444_ >>= fun w__472 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_303_ >>= fun w__472 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__472 with
- | Some
- (_stringappend_445_,(existT _ _stringappend_446_ _)) =>
- returnm ((_stringappend_445_,
- build_ex
- _stringappend_446_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__474 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_446_ _) :=
- w__474
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_447_ :=
+ (match w__472 with
+ | Some
+ (tt,(existT _ _stringappend_304_ _)) =>
+ returnm (tt, build_ex _stringappend_304_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_304_ _) =>
+ let _stringappend_305_ :=
string_drop
- _stringappend_444_
- (build_ex
- _stringappend_446_) in
- match (reg_name_matches_prefix
- _stringappend_447_) with
- | Some
- (_stringappend_448_,(existT _ _stringappend_449_ _)) =>
- returnm ((_stringappend_448_,
- build_ex
- _stringappend_449_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__476 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_449_ _) :=
- w__476
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_303_
+ (build_ex _stringappend_304_) in
+ (match (reg_name_matches_prefix
+ _stringappend_305_) with
+ | Some
+ (rs1,(existT _ _stringappend_306_ _)) =>
+ returnm (rs1, build_ex _stringappend_306_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_306_ _) =>
returnm ((match (string_drop
- _stringappend_447_
- (build_ex
- _stringappend_449_)) with
+ _stringappend_305_
+ (build_ex _stringappend_306_)) with
| s_ =>
- Some (LOADRES (aq,rl,rs1,size,rd),
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((LOADRES
+ (aq, rl, rs1,
+ size, rd), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -28035,96 +25715,88 @@ Definition assembly_matches_prefix (arg_ : string)
_stringappend_0_
"sc.")
: bool))
- (let _stringappend_451_ :=
+ (let _stringappend_308_ :=
string_drop
_stringappend_0_
- (string_length "sc.") in
+ (build_ex (projT1 (string_length
+ "sc."))) in
match (maybe_aq_matches_prefix
- _stringappend_451_) with
+ _stringappend_308_) with
| Some
- (_stringappend_452_,(existT _ _stringappend_453_ _)) =>
- let _stringappend_454_ :=
+ (aq,(existT _ _stringappend_309_ _)) =>
+ let _stringappend_310_ :=
string_drop
- _stringappend_451_
- (build_ex
- _stringappend_453_) in
+ _stringappend_308_
+ (build_ex _stringappend_309_) in
match (maybe_rl_matches_prefix
- _stringappend_454_) with
+ _stringappend_310_) with
| Some
- (_stringappend_455_,(existT _ _stringappend_456_ _)) =>
- let _stringappend_457_ :=
+ (rl,(existT _ _stringappend_311_ _)) =>
+ let _stringappend_312_ :=
string_drop
- _stringappend_454_
- (build_ex
- _stringappend_456_) in
+ _stringappend_310_
+ (build_ex _stringappend_311_) in
match (size_mnemonic_matches_prefix
- _stringappend_457_) with
+ _stringappend_312_) with
| Some
- (_stringappend_458_,(existT _ _stringappend_459_ _)) =>
- let _stringappend_460_ :=
+ (size,(existT _ _stringappend_313_ _)) =>
+ let _stringappend_314_ :=
string_drop
- _stringappend_457_
- (build_ex
- _stringappend_459_) in
+ _stringappend_312_
+ (build_ex _stringappend_313_) in
match (spc_matches_prefix
- _stringappend_460_) with
+ _stringappend_314_) with
| Some
- (_stringappend_461_,(existT _ _stringappend_462_ _)) =>
- let _stringappend_463_ :=
+ (tt,(existT _ _stringappend_315_ _)) =>
+ let _stringappend_316_ :=
string_drop
- _stringappend_460_
- (build_ex
- _stringappend_462_) in
+ _stringappend_314_
+ (build_ex _stringappend_315_) in
match (reg_name_matches_prefix
- _stringappend_463_) with
+ _stringappend_316_) with
| Some
- (_stringappend_464_,(existT _ _stringappend_465_ _)) =>
- let _stringappend_466_ :=
+ (rd,(existT _ _stringappend_317_ _)) =>
+ let _stringappend_318_ :=
string_drop
- _stringappend_463_
- (build_ex
- _stringappend_465_) in
+ _stringappend_316_
+ (build_ex _stringappend_317_) in
sep_matches_prefix
- _stringappend_466_ >>= fun w__477 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_318_ >>= fun w__477 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__477 with
| Some
- (_stringappend_467_,(existT _ _stringappend_468_ _)) =>
- let _stringappend_469_ :=
+ (tt,(existT _ _stringappend_319_ _)) =>
+ let _stringappend_320_ :=
string_drop
- _stringappend_466_
- (build_ex
- _stringappend_468_) in
+ _stringappend_318_
+ (build_ex _stringappend_319_) in
match (reg_name_matches_prefix
- _stringappend_469_) with
+ _stringappend_320_) with
| Some
- (_stringappend_470_,(existT _ _stringappend_471_ _)) =>
- let _stringappend_472_ :=
+ (rs1,(existT _ _stringappend_321_ _)) =>
+ let _stringappend_322_ :=
string_drop
- _stringappend_469_
- (build_ex
- _stringappend_471_) in
+ _stringappend_320_
+ (build_ex _stringappend_321_) in
sep_matches_prefix
- _stringappend_472_ >>= fun w__478 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_322_ >>= fun w__478 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if
((match w__478 with
| Some
- (_stringappend_473_,(existT _ _stringappend_474_ _)) =>
- let _stringappend_475_ :=
+ (tt,(existT _ _stringappend_323_ _)) =>
+ let _stringappend_324_ :=
string_drop
- _stringappend_472_
- (build_ex
- _stringappend_474_) in
+ _stringappend_322_
+ (build_ex _stringappend_323_) in
if
((match (reg_name_matches_prefix
- _stringappend_475_) with
+ _stringappend_324_) with
| Some
- (_stringappend_476_,(existT _ _stringappend_477_ _)) =>
+ (rs2,(existT _ _stringappend_325_ _)) =>
match (string_drop
- _stringappend_475_
- (build_ex
- _stringappend_477_)) with
+ _stringappend_324_
+ (build_ex _stringappend_325_)) with
| s_ =>
true
end
@@ -28207,236 +25879,159 @@ Definition assembly_matches_prefix (arg_ : string)
else false)
: bool)) >>= fun w__486 : bool =>
(if (w__486) then
- let _stringappend_451_ :=
+ let _stringappend_308_ :=
string_drop
_stringappend_0_
- (string_length "sc.") in
- match (maybe_aq_matches_prefix
- _stringappend_451_) with
- | Some
- (_stringappend_452_,(existT _ _stringappend_453_ _)) =>
- returnm ((_stringappend_452_,
- build_ex
- _stringappend_453_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__488 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(aq, existT _ _stringappend_453_ _) :=
- w__488
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_454_ :=
+ (build_ex (projT1 (string_length
+ "sc."))) in
+ (match (maybe_aq_matches_prefix
+ _stringappend_308_) with
+ | Some
+ (aq,(existT _ _stringappend_309_ _)) =>
+ returnm (aq, build_ex _stringappend_309_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(aq, existT _ _stringappend_309_ _) =>
+ let _stringappend_310_ :=
string_drop
- _stringappend_451_
- (build_ex
- _stringappend_453_) in
- match (maybe_rl_matches_prefix
- _stringappend_454_) with
- | Some
- (_stringappend_455_,(existT _ _stringappend_456_ _)) =>
- returnm ((_stringappend_455_,
- build_ex
- _stringappend_456_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__490 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rl, existT _ _stringappend_456_ _) :=
- w__490
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_457_ :=
+ _stringappend_308_
+ (build_ex _stringappend_309_) in
+ (match (maybe_rl_matches_prefix
+ _stringappend_310_) with
+ | Some
+ (rl,(existT _ _stringappend_311_ _)) =>
+ returnm (rl, build_ex _stringappend_311_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rl, existT _ _stringappend_311_ _) =>
+ let _stringappend_312_ :=
string_drop
- _stringappend_454_
- (build_ex
- _stringappend_456_) in
- match (size_mnemonic_matches_prefix
- _stringappend_457_) with
- | Some
- (_stringappend_458_,(existT _ _stringappend_459_ _)) =>
- returnm ((_stringappend_458_,
- build_ex
- _stringappend_459_)
- : (word_width * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((word_width * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__492 : (word_width * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(size, existT _ _stringappend_459_ _) :=
- w__492
- : (word_width * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_460_ :=
+ _stringappend_310_
+ (build_ex _stringappend_311_) in
+ (match (size_mnemonic_matches_prefix
+ _stringappend_312_) with
+ | Some
+ (size,(existT _ _stringappend_313_ _)) =>
+ returnm (size, build_ex _stringappend_313_)
+ | _ =>
+ exit tt
+ : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(size, existT _ _stringappend_313_ _) =>
+ let _stringappend_314_ :=
string_drop
- _stringappend_457_
- (build_ex
- _stringappend_459_) in
- match (spc_matches_prefix
- _stringappend_460_) with
- | Some
- (_stringappend_461_,(existT _ _stringappend_462_ _)) =>
- returnm ((_stringappend_461_,
- build_ex
- _stringappend_462_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__494 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_462_ _) :=
- w__494
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_463_ :=
+ _stringappend_312_
+ (build_ex _stringappend_313_) in
+ (match (spc_matches_prefix
+ _stringappend_314_) with
+ | Some
+ (tt,(existT _ _stringappend_315_ _)) =>
+ returnm (tt, build_ex _stringappend_315_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_315_ _) =>
+ let _stringappend_316_ :=
string_drop
- _stringappend_460_
- (build_ex
- _stringappend_462_) in
- match (reg_name_matches_prefix
- _stringappend_463_) with
- | Some
- (_stringappend_464_,(existT _ _stringappend_465_ _)) =>
- returnm ((_stringappend_464_,
- build_ex
- _stringappend_465_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__496 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_465_ _) :=
- w__496
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_466_ :=
+ _stringappend_314_
+ (build_ex _stringappend_315_) in
+ (match (reg_name_matches_prefix
+ _stringappend_316_) with
+ | Some
+ (rd,(existT _ _stringappend_317_ _)) =>
+ returnm (rd, build_ex _stringappend_317_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_317_ _) =>
+ let _stringappend_318_ :=
string_drop
- _stringappend_463_
- (build_ex
- _stringappend_465_) in
+ _stringappend_316_
+ (build_ex _stringappend_317_) in
sep_matches_prefix
- _stringappend_466_ >>= fun w__497 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_318_ >>= fun w__497 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__497 with
- | Some
- (_stringappend_467_,(existT _ _stringappend_468_ _)) =>
- returnm ((_stringappend_467_,
- build_ex
- _stringappend_468_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__499 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_468_ _) :=
- w__499
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_469_ :=
+ (match w__497 with
+ | Some
+ (tt,(existT _ _stringappend_319_ _)) =>
+ returnm (tt, build_ex _stringappend_319_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_319_ _) =>
+ let _stringappend_320_ :=
string_drop
- _stringappend_466_
- (build_ex
- _stringappend_468_) in
- match (reg_name_matches_prefix
- _stringappend_469_) with
- | Some
- (_stringappend_470_,(existT _ _stringappend_471_ _)) =>
- returnm ((_stringappend_470_,
- build_ex
- _stringappend_471_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__501 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_471_ _) :=
- w__501
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_472_ :=
+ _stringappend_318_
+ (build_ex _stringappend_319_) in
+ (match (reg_name_matches_prefix
+ _stringappend_320_) with
+ | Some
+ (rs1,(existT _ _stringappend_321_ _)) =>
+ returnm (rs1, build_ex _stringappend_321_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_321_ _) =>
+ let _stringappend_322_ :=
string_drop
- _stringappend_469_
- (build_ex
- _stringappend_471_) in
+ _stringappend_320_
+ (build_ex _stringappend_321_) in
sep_matches_prefix
- _stringappend_472_ >>= fun w__502 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_322_ >>= fun w__502 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__502 with
- | Some
- (_stringappend_473_,(existT _ _stringappend_474_ _)) =>
- returnm ((_stringappend_473_,
- build_ex
- _stringappend_474_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__504 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_474_ _) :=
- w__504
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_475_ :=
+ (match w__502 with
+ | Some
+ (tt,(existT _ _stringappend_323_ _)) =>
+ returnm (tt, build_ex _stringappend_323_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_323_ _) =>
+ let _stringappend_324_ :=
string_drop
- _stringappend_472_
- (build_ex
- _stringappend_474_) in
- match (reg_name_matches_prefix
- _stringappend_475_) with
- | Some
- (_stringappend_476_,(existT _ _stringappend_477_ _)) =>
- returnm ((_stringappend_476_,
- build_ex
- _stringappend_477_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__506 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_477_ _) :=
- w__506
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_322_
+ (build_ex _stringappend_323_) in
+ (match (reg_name_matches_prefix
+ _stringappend_324_) with
+ | Some
+ (rs2,(existT _ _stringappend_325_ _)) =>
+ returnm (rs2, build_ex _stringappend_325_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_325_ _) =>
returnm ((match (string_drop
- _stringappend_475_
- (build_ex
- _stringappend_477_)) with
+ _stringappend_324_
+ (build_ex _stringappend_325_)) with
| s_ =>
- Some (STORECON (aq,rl,rs2,rs1,size,rd),
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((STORECON
+ (aq, rl,
+ rs2, rs1,
+ size, rd), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -28444,107 +26039,98 @@ Definition assembly_matches_prefix (arg_ : string)
match (amo_mnemonic_matches_prefix
_stringappend_0_) with
| Some
- (_stringappend_479_,(existT _ _stringappend_480_ _)) =>
- let _stringappend_481_ :=
+ (op,(existT _ _stringappend_327_ _)) =>
+ let _stringappend_328_ :=
string_drop
_stringappend_0_
- (build_ex
- _stringappend_480_) in
+ (build_ex _stringappend_327_) in
and_boolM
(returnm ((string_startswith
- _stringappend_481_
+ _stringappend_328_
".")
: bool))
- (let _stringappend_482_ :=
+ (let _stringappend_329_ :=
string_drop
- _stringappend_481_
- (string_length ".") in
+ _stringappend_328_
+ (build_ex (projT1 (string_length
+ "."))) in
match (size_mnemonic_matches_prefix
- _stringappend_482_) with
+ _stringappend_329_) with
| Some
- (_stringappend_483_,(existT _ _stringappend_484_ _)) =>
- let _stringappend_485_ :=
+ (width,(existT _ _stringappend_330_ _)) =>
+ let _stringappend_331_ :=
string_drop
- _stringappend_482_
- (build_ex
- _stringappend_484_) in
+ _stringappend_329_
+ (build_ex _stringappend_330_) in
match (maybe_aq_matches_prefix
- _stringappend_485_) with
+ _stringappend_331_) with
| Some
- (_stringappend_486_,(existT _ _stringappend_487_ _)) =>
- let _stringappend_488_ :=
+ (aq,(existT _ _stringappend_332_ _)) =>
+ let _stringappend_333_ :=
string_drop
- _stringappend_485_
- (build_ex
- _stringappend_487_) in
+ _stringappend_331_
+ (build_ex _stringappend_332_) in
match (maybe_rl_matches_prefix
- _stringappend_488_) with
+ _stringappend_333_) with
| Some
- (_stringappend_489_,(existT _ _stringappend_490_ _)) =>
- let _stringappend_491_ :=
+ (rl,(existT _ _stringappend_334_ _)) =>
+ let _stringappend_335_ :=
string_drop
- _stringappend_488_
- (build_ex
- _stringappend_490_) in
+ _stringappend_333_
+ (build_ex _stringappend_334_) in
match (spc_matches_prefix
- _stringappend_491_) with
+ _stringappend_335_) with
| Some
- (_stringappend_492_,(existT _ _stringappend_493_ _)) =>
- let _stringappend_494_ :=
+ (tt,(existT _ _stringappend_336_ _)) =>
+ let _stringappend_337_ :=
string_drop
- _stringappend_491_
- (build_ex
- _stringappend_493_) in
+ _stringappend_335_
+ (build_ex _stringappend_336_) in
match (reg_name_matches_prefix
- _stringappend_494_) with
+ _stringappend_337_) with
| Some
- (_stringappend_495_,(existT _ _stringappend_496_ _)) =>
- let _stringappend_497_ :=
+ (rd,(existT _ _stringappend_338_ _)) =>
+ let _stringappend_339_ :=
string_drop
- _stringappend_494_
- (build_ex
- _stringappend_496_) in
+ _stringappend_337_
+ (build_ex _stringappend_338_) in
sep_matches_prefix
- _stringappend_497_ >>= fun w__507 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_339_ >>= fun w__507 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__507 with
| Some
- (_stringappend_498_,(existT _ _stringappend_499_ _)) =>
- let _stringappend_500_ :=
+ (tt,(existT _ _stringappend_340_ _)) =>
+ let _stringappend_341_ :=
string_drop
- _stringappend_497_
- (build_ex
- _stringappend_499_) in
+ _stringappend_339_
+ (build_ex _stringappend_340_) in
match (reg_name_matches_prefix
- _stringappend_500_) with
+ _stringappend_341_) with
| Some
- (_stringappend_501_,(existT _ _stringappend_502_ _)) =>
- let _stringappend_503_ :=
+ (rs1,(existT _ _stringappend_342_ _)) =>
+ let _stringappend_343_ :=
string_drop
- _stringappend_500_
- (build_ex
- _stringappend_502_) in
+ _stringappend_341_
+ (build_ex _stringappend_342_) in
sep_matches_prefix
- _stringappend_503_ >>= fun w__508 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_343_ >>= fun w__508 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if
((match w__508 with
| Some
- (_stringappend_504_,(existT _ _stringappend_505_ _)) =>
- let _stringappend_506_ :=
+ (tt,(existT _ _stringappend_344_ _)) =>
+ let _stringappend_345_ :=
string_drop
- _stringappend_503_
- (build_ex
- _stringappend_505_) in
+ _stringappend_343_
+ (build_ex _stringappend_344_) in
if
((match (reg_name_matches_prefix
- _stringappend_506_) with
+ _stringappend_345_) with
| Some
- (_stringappend_507_,(existT _ _stringappend_508_ _)) =>
+ (rs2,(existT _ _stringappend_346_ _)) =>
match (string_drop
- _stringappend_506_
- (build_ex
- _stringappend_508_)) with
+ _stringappend_345_
+ (build_ex _stringappend_346_)) with
| s_ =>
true
end
@@ -28643,260 +26229,177 @@ Definition assembly_matches_prefix (arg_ : string)
returnm (false : bool)
end >>= fun w__517 : bool =>
(if (w__517) then
- match (amo_mnemonic_matches_prefix
- _stringappend_0_) with
- | Some
- (_stringappend_479_,(existT _ _stringappend_480_ _)) =>
- returnm ((_stringappend_479_,
- build_ex
- _stringappend_480_)
- : (amoop * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((amoop * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__519 : (amoop * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(op, existT _ _stringappend_480_ _) :=
- w__519
- : (amoop * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_481_ :=
+ (match (amo_mnemonic_matches_prefix
+ _stringappend_0_) with
+ | Some
+ (op,(existT _ _stringappend_327_ _)) =>
+ returnm (op, build_ex _stringappend_327_)
+ | _ =>
+ exit tt
+ : M ((amoop * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((amoop * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(op, existT _ _stringappend_327_ _) =>
+ let _stringappend_328_ :=
string_drop
_stringappend_0_
- (build_ex
- _stringappend_480_) in
- let _stringappend_482_ :=
+ (build_ex _stringappend_327_) in
+ let _stringappend_329_ :=
string_drop
- _stringappend_481_
- (string_length ".") in
- match (size_mnemonic_matches_prefix
- _stringappend_482_) with
- | Some
- (_stringappend_483_,(existT _ _stringappend_484_ _)) =>
- returnm ((_stringappend_483_,
- build_ex
- _stringappend_484_)
- : (word_width * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((word_width * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__521 : (word_width * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(width, existT _ _stringappend_484_ _) :=
- w__521
- : (word_width * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_485_ :=
+ _stringappend_328_
+ (build_ex (projT1 (string_length
+ "."))) in
+ (match (size_mnemonic_matches_prefix
+ _stringappend_329_) with
+ | Some
+ (width,(existT _ _stringappend_330_ _)) =>
+ returnm (width, build_ex _stringappend_330_)
+ | _ =>
+ exit tt
+ : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((word_width * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(width, existT _ _stringappend_330_ _) =>
+ let _stringappend_331_ :=
string_drop
- _stringappend_482_
- (build_ex
- _stringappend_484_) in
- match (maybe_aq_matches_prefix
- _stringappend_485_) with
- | Some
- (_stringappend_486_,(existT _ _stringappend_487_ _)) =>
- returnm ((_stringappend_486_,
- build_ex
- _stringappend_487_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__523 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(aq, existT _ _stringappend_487_ _) :=
- w__523
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_488_ :=
+ _stringappend_329_
+ (build_ex _stringappend_330_) in
+ (match (maybe_aq_matches_prefix
+ _stringappend_331_) with
+ | Some
+ (aq,(existT _ _stringappend_332_ _)) =>
+ returnm (aq, build_ex _stringappend_332_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(aq, existT _ _stringappend_332_ _) =>
+ let _stringappend_333_ :=
string_drop
- _stringappend_485_
- (build_ex
- _stringappend_487_) in
- match (maybe_rl_matches_prefix
- _stringappend_488_) with
- | Some
- (_stringappend_489_,(existT _ _stringappend_490_ _)) =>
- returnm ((_stringappend_489_,
- build_ex
- _stringappend_490_)
- : (bool * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((bool * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__525 : (bool * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rl, existT _ _stringappend_490_ _) :=
- w__525
- : (bool * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_491_ :=
+ _stringappend_331_
+ (build_ex _stringappend_332_) in
+ (match (maybe_rl_matches_prefix
+ _stringappend_333_) with
+ | Some
+ (rl,(existT _ _stringappend_334_ _)) =>
+ returnm (rl, build_ex _stringappend_334_)
+ | _ =>
+ exit tt
+ : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((bool * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rl, existT _ _stringappend_334_ _) =>
+ let _stringappend_335_ :=
string_drop
- _stringappend_488_
- (build_ex
- _stringappend_490_) in
- match (spc_matches_prefix
- _stringappend_491_) with
- | Some
- (_stringappend_492_,(existT _ _stringappend_493_ _)) =>
- returnm ((_stringappend_492_,
- build_ex
- _stringappend_493_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__527 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_493_ _) :=
- w__527
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_494_ :=
+ _stringappend_333_
+ (build_ex _stringappend_334_) in
+ (match (spc_matches_prefix
+ _stringappend_335_) with
+ | Some
+ (tt,(existT _ _stringappend_336_ _)) =>
+ returnm (tt, build_ex _stringappend_336_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_336_ _) =>
+ let _stringappend_337_ :=
string_drop
- _stringappend_491_
- (build_ex
- _stringappend_493_) in
- match (reg_name_matches_prefix
- _stringappend_494_) with
- | Some
- (_stringappend_495_,(existT _ _stringappend_496_ _)) =>
- returnm ((_stringappend_495_,
- build_ex
- _stringappend_496_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__529 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_496_ _) :=
- w__529
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_497_ :=
+ _stringappend_335_
+ (build_ex _stringappend_336_) in
+ (match (reg_name_matches_prefix
+ _stringappend_337_) with
+ | Some
+ (rd,(existT _ _stringappend_338_ _)) =>
+ returnm (rd, build_ex _stringappend_338_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_338_ _) =>
+ let _stringappend_339_ :=
string_drop
- _stringappend_494_
- (build_ex
- _stringappend_496_) in
+ _stringappend_337_
+ (build_ex _stringappend_338_) in
sep_matches_prefix
- _stringappend_497_ >>= fun w__530 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_339_ >>= fun w__530 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__530 with
- | Some
- (_stringappend_498_,(existT _ _stringappend_499_ _)) =>
- returnm ((_stringappend_498_,
- build_ex
- _stringappend_499_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__532 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_499_ _) :=
- w__532
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_500_ :=
+ (match w__530 with
+ | Some
+ (tt,(existT _ _stringappend_340_ _)) =>
+ returnm (tt, build_ex _stringappend_340_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_340_ _) =>
+ let _stringappend_341_ :=
string_drop
- _stringappend_497_
- (build_ex
- _stringappend_499_) in
- match (reg_name_matches_prefix
- _stringappend_500_) with
- | Some
- (_stringappend_501_,(existT _ _stringappend_502_ _)) =>
- returnm ((_stringappend_501_,
- build_ex
- _stringappend_502_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__534 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_502_ _) :=
- w__534
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_503_ :=
+ _stringappend_339_
+ (build_ex _stringappend_340_) in
+ (match (reg_name_matches_prefix
+ _stringappend_341_) with
+ | Some
+ (rs1,(existT _ _stringappend_342_ _)) =>
+ returnm (rs1, build_ex _stringappend_342_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_342_ _) =>
+ let _stringappend_343_ :=
string_drop
- _stringappend_500_
- (build_ex
- _stringappend_502_) in
+ _stringappend_341_
+ (build_ex _stringappend_342_) in
sep_matches_prefix
- _stringappend_503_ >>= fun w__535 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_343_ >>= fun w__535 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__535 with
- | Some
- (_stringappend_504_,(existT _ _stringappend_505_ _)) =>
- returnm ((_stringappend_504_,
- build_ex
- _stringappend_505_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__537 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_505_ _) :=
- w__537
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_506_ :=
+ (match w__535 with
+ | Some
+ (tt,(existT _ _stringappend_344_ _)) =>
+ returnm (tt, build_ex _stringappend_344_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_344_ _) =>
+ let _stringappend_345_ :=
string_drop
- _stringappend_503_
- (build_ex
- _stringappend_505_) in
- match (reg_name_matches_prefix
- _stringappend_506_) with
- | Some
- (_stringappend_507_,(existT _ _stringappend_508_ _)) =>
- returnm ((_stringappend_507_,
- build_ex
- _stringappend_508_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__539 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs2, existT _ _stringappend_508_ _) :=
- w__539
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_343_
+ (build_ex _stringappend_344_) in
+ (match (reg_name_matches_prefix
+ _stringappend_345_) with
+ | Some
+ (rs2,(existT _ _stringappend_346_ _)) =>
+ returnm (rs2, build_ex _stringappend_346_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs2, existT _ _stringappend_346_ _) =>
returnm ((match (string_drop
- _stringappend_506_
- (build_ex
- _stringappend_508_)) with
+ _stringappend_345_
+ (build_ex _stringappend_346_)) with
| s_ =>
- Some (AMO (op,aq,rl,rs2,rs1,width,rd),
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((AMO
+ (op,
+ aq, rl,
+ rs2,
+ rs1,
+ width,
+ rd), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -28904,81 +26407,74 @@ Definition assembly_matches_prefix (arg_ : string)
match (csr_mnemonic_matches_prefix
_stringappend_0_) with
| Some
- (_stringappend_510_,(existT _ _stringappend_511_ _)) =>
- let _stringappend_512_ :=
+ (op,(existT _ _stringappend_348_ _)) =>
+ let _stringappend_349_ :=
string_drop
_stringappend_0_
- (build_ex
- _stringappend_511_) in
+ (build_ex _stringappend_348_) in
and_boolM
(returnm ((string_startswith
- _stringappend_512_
+ _stringappend_349_
"i")
: bool))
- (let _stringappend_513_ :=
+ (let _stringappend_350_ :=
string_drop
- _stringappend_512_
- (string_length
- "i") in
+ _stringappend_349_
+ (build_ex (projT1 (string_length
+ "i"))) in
match (spc_matches_prefix
- _stringappend_513_) with
+ _stringappend_350_) with
| Some
- (_stringappend_514_,(existT _ _stringappend_515_ _)) =>
- let _stringappend_516_ :=
+ (tt,(existT _ _stringappend_351_ _)) =>
+ let _stringappend_352_ :=
string_drop
- _stringappend_513_
- (build_ex
- _stringappend_515_) in
+ _stringappend_350_
+ (build_ex _stringappend_351_) in
match (reg_name_matches_prefix
- _stringappend_516_) with
+ _stringappend_352_) with
| Some
- (_stringappend_517_,(existT _ _stringappend_518_ _)) =>
- let _stringappend_519_ :=
+ (rd,(existT _ _stringappend_353_ _)) =>
+ let _stringappend_354_ :=
string_drop
- _stringappend_516_
- (build_ex
- _stringappend_518_) in
+ _stringappend_352_
+ (build_ex _stringappend_353_) in
sep_matches_prefix
- _stringappend_519_ >>= fun w__540 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_354_ >>= fun w__540 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__540 with
| Some
- (_stringappend_520_,(existT _ _stringappend_521_ _)) =>
- let _stringappend_522_ :=
+ (tt,(existT _ _stringappend_355_ _)) =>
+ let _stringappend_356_ :=
string_drop
- _stringappend_519_
- (build_ex
- _stringappend_521_) in
+ _stringappend_354_
+ (build_ex _stringappend_355_) in
match (hex_bits_5_matches_prefix
- _stringappend_522_) with
+ _stringappend_356_) with
| Some
- (_stringappend_523_,(existT _ _stringappend_524_ _)) =>
- let _stringappend_525_ :=
+ (rs1,(existT _ _stringappend_357_ _)) =>
+ let _stringappend_358_ :=
string_drop
- _stringappend_522_
- (build_ex
- _stringappend_524_) in
+ _stringappend_356_
+ (build_ex _stringappend_357_) in
sep_matches_prefix
- _stringappend_525_ >>= fun w__541 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_358_ >>= fun w__541 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if
((match w__541 with
| Some
- (_stringappend_526_,(existT _ _stringappend_527_ _)) =>
- let _stringappend_528_ :=
+ (tt,(existT _ _stringappend_359_ _)) =>
+ let _stringappend_360_ :=
string_drop
- _stringappend_525_
- (build_ex
- _stringappend_527_) in
+ _stringappend_358_
+ (build_ex _stringappend_359_) in
if
((match (csr_name_map_matches_prefix
- _stringappend_528_) with
+ _stringappend_360_) with
| Some
- (_stringappend_529_,(existT _ _stringappend_530_ _)) =>
+ (csr,(existT _ _stringappend_361_ _)) =>
match (string_drop
- _stringappend_528_
- (build_ex
- _stringappend_530_)) with
+ _stringappend_360_
+ (build_ex _stringappend_361_)) with
| s_ =>
true
end
@@ -29047,188 +26543,131 @@ Definition assembly_matches_prefix (arg_ : string)
: bool)
end >>= fun w__547 : bool =>
(if (w__547) then
- match (csr_mnemonic_matches_prefix
- _stringappend_0_) with
- | Some
- (_stringappend_510_,(existT _ _stringappend_511_ _)) =>
- returnm ((_stringappend_510_,
- build_ex
- _stringappend_511_)
- : (csrop * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((csrop * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__549 : (csrop * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(op, existT _ _stringappend_511_ _) :=
- w__549
- : (csrop * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_512_ :=
+ (match (csr_mnemonic_matches_prefix
+ _stringappend_0_) with
+ | Some
+ (op,(existT _ _stringappend_348_ _)) =>
+ returnm (op, build_ex _stringappend_348_)
+ | _ =>
+ exit tt
+ : M ((csrop * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((csrop * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(op, existT _ _stringappend_348_ _) =>
+ let _stringappend_349_ :=
string_drop
_stringappend_0_
- (build_ex
- _stringappend_511_) in
- let _stringappend_513_ :=
+ (build_ex _stringappend_348_) in
+ let _stringappend_350_ :=
string_drop
- _stringappend_512_
- (string_length "i") in
- match (spc_matches_prefix
- _stringappend_513_) with
- | Some
- (_stringappend_514_,(existT _ _stringappend_515_ _)) =>
- returnm ((_stringappend_514_,
- build_ex
- _stringappend_515_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__551 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_515_ _) :=
- w__551
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_516_ :=
+ _stringappend_349_
+ (build_ex (projT1 (string_length
+ "i"))) in
+ (match (spc_matches_prefix
+ _stringappend_350_) with
+ | Some
+ (tt,(existT _ _stringappend_351_ _)) =>
+ returnm (tt, build_ex _stringappend_351_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_351_ _) =>
+ let _stringappend_352_ :=
string_drop
- _stringappend_513_
- (build_ex
- _stringappend_515_) in
- match (reg_name_matches_prefix
- _stringappend_516_) with
- | Some
- (_stringappend_517_,(existT _ _stringappend_518_ _)) =>
- returnm ((_stringappend_517_,
- build_ex
- _stringappend_518_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__553 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_518_ _) :=
- w__553
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_519_ :=
+ _stringappend_350_
+ (build_ex _stringappend_351_) in
+ (match (reg_name_matches_prefix
+ _stringappend_352_) with
+ | Some
+ (rd,(existT _ _stringappend_353_ _)) =>
+ returnm (rd, build_ex _stringappend_353_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_353_ _) =>
+ let _stringappend_354_ :=
string_drop
- _stringappend_516_
- (build_ex
- _stringappend_518_) in
+ _stringappend_352_
+ (build_ex _stringappend_353_) in
sep_matches_prefix
- _stringappend_519_ >>= fun w__554 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_354_ >>= fun w__554 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__554 with
- | Some
- (_stringappend_520_,(existT _ _stringappend_521_ _)) =>
- returnm ((_stringappend_520_,
- build_ex
- _stringappend_521_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__556 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_521_ _) :=
- w__556
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_522_ :=
+ (match w__554 with
+ | Some
+ (tt,(existT _ _stringappend_355_ _)) =>
+ returnm (tt, build_ex _stringappend_355_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_355_ _) =>
+ let _stringappend_356_ :=
string_drop
- _stringappend_519_
- (build_ex
- _stringappend_521_) in
- match (hex_bits_5_matches_prefix
- _stringappend_522_) with
- | Some
- (_stringappend_523_,(existT _ _stringappend_524_ _)) =>
- returnm ((_stringappend_523_,
- build_ex
- _stringappend_524_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__558 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_524_ _) :=
- w__558
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_525_ :=
+ _stringappend_354_
+ (build_ex _stringappend_355_) in
+ (match (hex_bits_5_matches_prefix
+ _stringappend_356_) with
+ | Some
+ (rs1,(existT _ _stringappend_357_ _)) =>
+ returnm (rs1, build_ex _stringappend_357_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_357_ _) =>
+ let _stringappend_358_ :=
string_drop
- _stringappend_522_
- (build_ex
- _stringappend_524_) in
+ _stringappend_356_
+ (build_ex _stringappend_357_) in
sep_matches_prefix
- _stringappend_525_ >>= fun w__559 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_358_ >>= fun w__559 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__559 with
- | Some
- (_stringappend_526_,(existT _ _stringappend_527_ _)) =>
- returnm ((_stringappend_526_,
- build_ex
- _stringappend_527_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__561 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_527_ _) :=
- w__561
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_528_ :=
+ (match w__559 with
+ | Some
+ (tt,(existT _ _stringappend_359_ _)) =>
+ returnm (tt, build_ex _stringappend_359_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_359_ _) =>
+ let _stringappend_360_ :=
string_drop
- _stringappend_525_
- (build_ex
- _stringappend_527_) in
- match (csr_name_map_matches_prefix
- _stringappend_528_) with
- | Some
- (_stringappend_529_,(existT _ _stringappend_530_ _)) =>
- returnm ((_stringappend_529_,
- build_ex
- _stringappend_530_)
- : (mword 12 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 12 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__563 : (mword 12 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(csr, existT _ _stringappend_530_ _) :=
- w__563
- : (mword 12 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_358_
+ (build_ex _stringappend_359_) in
+ (match (csr_name_map_matches_prefix
+ _stringappend_360_) with
+ | Some
+ (csr,(existT _ _stringappend_361_ _)) =>
+ returnm (csr, build_ex _stringappend_361_)
+ | _ =>
+ exit tt
+ : M ((mword 12 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(csr, existT _ _stringappend_361_ _) =>
returnm ((match (string_drop
- _stringappend_528_
- (build_ex
- _stringappend_530_)) with
+ _stringappend_360_
+ (build_ex _stringappend_361_)) with
| s_ =>
- Some (CSR (csr,rs1,rd,true,op),
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((CSR
+ (csr,
+ rs1,
+ rd,
+ true,
+ op), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -29236,71 +26675,64 @@ Definition assembly_matches_prefix (arg_ : string)
match (csr_mnemonic_matches_prefix
_stringappend_0_) with
| Some
- (_stringappend_532_,(existT _ _stringappend_533_ _)) =>
- let _stringappend_534_ :=
+ (op,(existT _ _stringappend_363_ _)) =>
+ let _stringappend_364_ :=
string_drop
_stringappend_0_
- (build_ex
- _stringappend_533_) in
+ (build_ex _stringappend_363_) in
match (spc_matches_prefix
- _stringappend_534_) with
+ _stringappend_364_) with
| Some
- (_stringappend_535_,(existT _ _stringappend_536_ _)) =>
- let _stringappend_537_ :=
+ (tt,(existT _ _stringappend_365_ _)) =>
+ let _stringappend_366_ :=
string_drop
- _stringappend_534_
- (build_ex
- _stringappend_536_) in
+ _stringappend_364_
+ (build_ex _stringappend_365_) in
match (reg_name_matches_prefix
- _stringappend_537_) with
+ _stringappend_366_) with
| Some
- (_stringappend_538_,(existT _ _stringappend_539_ _)) =>
- let _stringappend_540_ :=
+ (rd,(existT _ _stringappend_367_ _)) =>
+ let _stringappend_368_ :=
string_drop
- _stringappend_537_
- (build_ex
- _stringappend_539_) in
+ _stringappend_366_
+ (build_ex _stringappend_367_) in
sep_matches_prefix
- _stringappend_540_ >>= fun w__564 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_368_ >>= fun w__564 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
match w__564 with
| Some
- (_stringappend_541_,(existT _ _stringappend_542_ _)) =>
- let _stringappend_543_ :=
+ (tt,(existT _ _stringappend_369_ _)) =>
+ let _stringappend_370_ :=
string_drop
- _stringappend_540_
- (build_ex
- _stringappend_542_) in
+ _stringappend_368_
+ (build_ex _stringappend_369_) in
match (reg_name_matches_prefix
- _stringappend_543_) with
+ _stringappend_370_) with
| Some
- (_stringappend_544_,(existT _ _stringappend_545_ _)) =>
- let _stringappend_546_ :=
+ (rs1,(existT _ _stringappend_371_ _)) =>
+ let _stringappend_372_ :=
string_drop
- _stringappend_543_
- (build_ex
- _stringappend_545_) in
+ _stringappend_370_
+ (build_ex _stringappend_371_) in
sep_matches_prefix
- _stringappend_546_ >>= fun w__565 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_372_ >>= fun w__565 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
returnm ((if
((match w__565 with
| Some
- (_stringappend_547_,(existT _ _stringappend_548_ _)) =>
- let _stringappend_549_ :=
+ (tt,(existT _ _stringappend_373_ _)) =>
+ let _stringappend_374_ :=
string_drop
- _stringappend_546_
- (build_ex
- _stringappend_548_) in
+ _stringappend_372_
+ (build_ex _stringappend_373_) in
if
((match (csr_name_map_matches_prefix
- _stringappend_549_) with
+ _stringappend_374_) with
| Some
- (_stringappend_550_,(existT _ _stringappend_551_ _)) =>
+ (csr,(existT _ _stringappend_375_ _)) =>
match (string_drop
- _stringappend_549_
- (build_ex
- _stringappend_551_)) with
+ _stringappend_374_
+ (build_ex _stringappend_375_)) with
| s_ =>
true
end
@@ -29365,184 +26797,126 @@ Definition assembly_matches_prefix (arg_ : string)
: bool)
end >>= fun w__570 : bool =>
(if (w__570) then
- match (csr_mnemonic_matches_prefix
- _stringappend_0_) with
- | Some
- (_stringappend_532_,(existT _ _stringappend_533_ _)) =>
- returnm ((_stringappend_532_,
- build_ex
- _stringappend_533_)
- : (csrop * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((csrop * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__572 : (csrop * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(op, existT _ _stringappend_533_ _) :=
- w__572
- : (csrop * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_534_ :=
+ (match (csr_mnemonic_matches_prefix
+ _stringappend_0_) with
+ | Some
+ (op,(existT _ _stringappend_363_ _)) =>
+ returnm (op, build_ex _stringappend_363_)
+ | _ =>
+ exit tt
+ : M ((csrop * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((csrop * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(op, existT _ _stringappend_363_ _) =>
+ let _stringappend_364_ :=
string_drop
_stringappend_0_
- (build_ex
- _stringappend_533_) in
- match (spc_matches_prefix
- _stringappend_534_) with
- | Some
- (_stringappend_535_,(existT _ _stringappend_536_ _)) =>
- returnm ((_stringappend_535_,
- build_ex
- _stringappend_536_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__574 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_536_ _) :=
- w__574
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_537_ :=
+ (build_ex _stringappend_363_) in
+ (match (spc_matches_prefix
+ _stringappend_364_) with
+ | Some
+ (tt,(existT _ _stringappend_365_ _)) =>
+ returnm (tt, build_ex _stringappend_365_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_365_ _) =>
+ let _stringappend_366_ :=
string_drop
- _stringappend_534_
- (build_ex
- _stringappend_536_) in
- match (reg_name_matches_prefix
- _stringappend_537_) with
- | Some
- (_stringappend_538_,(existT _ _stringappend_539_ _)) =>
- returnm ((_stringappend_538_,
- build_ex
- _stringappend_539_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__576 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rd, existT _ _stringappend_539_ _) :=
- w__576
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_540_ :=
+ _stringappend_364_
+ (build_ex _stringappend_365_) in
+ (match (reg_name_matches_prefix
+ _stringappend_366_) with
+ | Some
+ (rd,(existT _ _stringappend_367_ _)) =>
+ returnm (rd, build_ex _stringappend_367_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rd, existT _ _stringappend_367_ _) =>
+ let _stringappend_368_ :=
string_drop
- _stringappend_537_
- (build_ex
- _stringappend_539_) in
+ _stringappend_366_
+ (build_ex _stringappend_367_) in
sep_matches_prefix
- _stringappend_540_ >>= fun w__577 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_368_ >>= fun w__577 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__577 with
- | Some
- (_stringappend_541_,(existT _ _stringappend_542_ _)) =>
- returnm ((_stringappend_541_,
- build_ex
- _stringappend_542_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__579 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_542_ _) :=
- w__579
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_543_ :=
+ (match w__577 with
+ | Some
+ (tt,(existT _ _stringappend_369_ _)) =>
+ returnm (tt, build_ex _stringappend_369_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_369_ _) =>
+ let _stringappend_370_ :=
string_drop
- _stringappend_540_
- (build_ex
- _stringappend_542_) in
- match (reg_name_matches_prefix
- _stringappend_543_) with
- | Some
- (_stringappend_544_,(existT _ _stringappend_545_ _)) =>
- returnm ((_stringappend_544_,
- build_ex
- _stringappend_545_)
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 5 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__581 : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(rs1, existT _ _stringappend_545_ _) :=
- w__581
- : (mword 5 * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_546_ :=
+ _stringappend_368_
+ (build_ex _stringappend_369_) in
+ (match (reg_name_matches_prefix
+ _stringappend_370_) with
+ | Some
+ (rs1,(existT _ _stringappend_371_ _)) =>
+ returnm (rs1, build_ex _stringappend_371_)
+ | _ =>
+ exit tt
+ : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 5 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(rs1, existT _ _stringappend_371_ _) =>
+ let _stringappend_372_ :=
string_drop
- _stringappend_543_
- (build_ex
- _stringappend_545_) in
+ _stringappend_370_
+ (build_ex _stringappend_371_) in
sep_matches_prefix
- _stringappend_546_ >>= fun w__582 : option ((unit * {n : Z & ArithFact (n >=
+ _stringappend_372_ >>= fun w__582 : option ((unit * {n : Z & ArithFact (n >=
0)})) =>
- match w__582 with
- | Some
- (_stringappend_547_,(existT _ _stringappend_548_ _)) =>
- returnm ((_stringappend_547_,
- build_ex
- _stringappend_548_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__584 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_548_ _) :=
- w__584
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_549_ :=
+ (match w__582 with
+ | Some
+ (tt,(existT _ _stringappend_373_ _)) =>
+ returnm (tt, build_ex _stringappend_373_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_373_ _) =>
+ let _stringappend_374_ :=
string_drop
- _stringappend_546_
- (build_ex
- _stringappend_548_) in
- match (csr_name_map_matches_prefix
- _stringappend_549_) with
- | Some
- (_stringappend_550_,(existT _ _stringappend_551_ _)) =>
- returnm ((_stringappend_550_,
- build_ex
- _stringappend_551_)
- : (mword 12 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 12 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__586 : (mword 12 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(csr, existT _ _stringappend_551_ _) :=
- w__586
- : (mword 12 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_372_
+ (build_ex _stringappend_373_) in
+ (match (csr_name_map_matches_prefix
+ _stringappend_374_) with
+ | Some
+ (csr,(existT _ _stringappend_375_ _)) =>
+ returnm (csr, build_ex _stringappend_375_)
+ | _ =>
+ exit tt
+ : M ((mword 12 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 12 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(csr, existT _ _stringappend_375_ _) =>
returnm ((match (string_drop
- _stringappend_549_
- (build_ex
- _stringappend_551_)) with
+ _stringappend_374_
+ (build_ex _stringappend_375_)) with
| s_ =>
- Some (CSR (csr,rs1,rd,false,op),
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((CSR
+ (csr,
+ rs1,
+ rd,
+ false,
+ op), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -29550,29 +26924,27 @@ Definition assembly_matches_prefix (arg_ : string)
(string_startswith
_stringappend_0_
"illegal")
- (let _stringappend_553_ :=
+ (let _stringappend_377_ :=
string_drop
_stringappend_0_
- (string_length
- "illegal") in
+ (build_ex (projT1 (string_length
+ "illegal"))) in
if ((match (spc_matches_prefix
- _stringappend_553_) with
+ _stringappend_377_) with
| Some
- (_stringappend_554_,(existT _ _stringappend_555_ _)) =>
- let _stringappend_556_ :=
+ (tt,(existT _ _stringappend_378_ _)) =>
+ let _stringappend_379_ :=
string_drop
- _stringappend_553_
- (build_ex
- _stringappend_555_) in
+ _stringappend_377_
+ (build_ex _stringappend_378_) in
if
((match (hex_bits_32_matches_prefix
- _stringappend_556_) with
+ _stringappend_379_) with
| Some
- (_stringappend_557_,(existT _ _stringappend_558_ _)) =>
+ (s,(existT _ _stringappend_380_ _)) =>
match (string_drop
- _stringappend_556_
- (build_ex
- _stringappend_558_)) with
+ _stringappend_379_
+ (build_ex _stringappend_380_)) with
| s_ =>
true
end
@@ -29591,65 +26963,48 @@ Definition assembly_matches_prefix (arg_ : string)
else
false)))
then
- let _stringappend_553_ :=
+ let _stringappend_377_ :=
string_drop
_stringappend_0_
- (string_length
- "illegal") in
- match (spc_matches_prefix
- _stringappend_553_) with
- | Some
- (_stringappend_554_,(existT _ _stringappend_555_ _)) =>
- returnm ((_stringappend_554_,
- build_ex
- _stringappend_555_)
- : (unit * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((unit * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__588 : (unit * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(tt, existT _ _stringappend_555_ _) :=
- w__588
- : (unit * {n : Z & ArithFact (n >=
- 0)}) in
- let _stringappend_556_ :=
+ (build_ex (projT1 (string_length
+ "illegal"))) in
+ (match (spc_matches_prefix
+ _stringappend_377_) with
+ | Some
+ (tt,(existT _ _stringappend_378_ _)) =>
+ returnm (tt, build_ex _stringappend_378_)
+ | _ =>
+ exit tt
+ : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((unit * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(tt, existT _ _stringappend_378_ _) =>
+ let _stringappend_379_ :=
string_drop
- _stringappend_553_
- (build_ex
- _stringappend_555_) in
- match (hex_bits_32_matches_prefix
- _stringappend_556_) with
- | Some
- (_stringappend_557_,(existT _ _stringappend_558_ _)) =>
- returnm ((_stringappend_557_,
- build_ex
- _stringappend_558_)
- : (mword 32 * {n : Z & ArithFact (n >=
- 0)}))
- | _ =>
- exit tt
- : M ((mword 32 * {n : Z & ArithFact (n >=
- 0)}))
- end >>= fun w__590 : (mword 32 * {n : Z & ArithFact (n >=
- 0)}) =>
- let '(s, existT _ _stringappend_558_ _) :=
- w__590
- : (mword 32 * {n : Z & ArithFact (n >=
- 0)}) in
+ _stringappend_377_
+ (build_ex _stringappend_378_) in
+ (match (hex_bits_32_matches_prefix
+ _stringappend_379_) with
+ | Some
+ (s,(existT _ _stringappend_380_ _)) =>
+ returnm (s, build_ex _stringappend_380_)
+ | _ =>
+ exit tt
+ : M ((mword 32 * {n : Z & ArithFact (n >=
+ 0)}))
+ end : M ((mword 32 * {n : Z & ArithFact (n >=
+ 0)}))) >>= fun '(s, existT _ _stringappend_380_ _) =>
returnm ((match (string_drop
- _stringappend_556_
- (build_ex
- _stringappend_558_)) with
+ _stringappend_379_
+ (build_ex _stringappend_380_)) with
| s_ =>
- Some (ILLEGAL s,
- sub_nat
- (string_length
- arg_)
- (string_length
- s_))
+ Some
+ ((ILLEGAL
+ (s), build_ex (projT1 (sub_nat
+ (build_ex (projT1 (string_length
+ arg_)))
+ (build_ex (projT1 (string_length
+ s_)))))))
end)
: option ((ast * {n : Z & ArithFact (n >=
0)})))
@@ -29698,11 +27053,12 @@ Definition encdec_forwards (arg_ : ast)
(match arg_ with
| UTYPE (imm,rd,op) =>
returnm ((concat_vec (imm : mword 20) (concat_vec (rd : mword 5) (encdec_uop_forwards op)))
- : mword 32)
+ : mword (20 + 12))
| RISCV_JAL (v__172,rd) =>
let imm_19 : bits 1 := subrange_vec_dec v__172 20 20 in
- let imm_7_0 : bits 8 := subrange_vec_dec v__172 19 12 in
let imm_8 : bits 1 := subrange_vec_dec v__172 11 11 in
+ let imm_7_0 : bits 8 := subrange_vec_dec v__172 19 12 in
+ let imm_19 : bits 1 := subrange_vec_dec v__172 20 20 in
let imm_18_13 : bits 6 := subrange_vec_dec v__172 10 5 in
let imm_12_9 : bits 4 := subrange_vec_dec v__172 4 1 in
returnm ((concat_vec (imm_19 : bits 1)
@@ -29718,12 +27074,13 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (rs1 : mword 5)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (rd : mword 5) (vec_of_bits [B1;B1;B0;B0;B1;B1;B1] : mword 7)))))
- : mword 32)
- | BTYPE (v__173,rs2,rs1,op) =>
- let imm7_6 : bits 1 := subrange_vec_dec v__173 12 12 in
- let imm5_0 : bits 1 := subrange_vec_dec v__173 11 11 in
- let imm7_5_0 : bits 6 := subrange_vec_dec v__173 10 5 in
- let imm5_4_1 : bits 4 := subrange_vec_dec v__173 4 1 in
+ : mword (12 + 20))
+ | BTYPE (v__174,rs2,rs1,op) =>
+ let imm7_6 : bits 1 := subrange_vec_dec v__174 12 12 in
+ let imm7_6 : bits 1 := subrange_vec_dec v__174 12 12 in
+ let imm7_5_0 : bits 6 := subrange_vec_dec v__174 10 5 in
+ let imm5_4_1 : bits 4 := subrange_vec_dec v__174 4 1 in
+ let imm5_0 : bits 1 := subrange_vec_dec v__174 11 11 in
returnm ((concat_vec (imm7_6 : bits 1)
(concat_vec (imm7_5_0 : bits 6)
(concat_vec (rs2 : mword 5)
@@ -29738,7 +27095,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (rs1 : mword 5)
(concat_vec (encdec_iop_forwards op)
(concat_vec (rd : mword 5) (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7)))))
- : mword 32)
+ : mword (12 + 20))
| SHIFTIOP (shamt,rs1,rd,RISCV_SLLI) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6)
(concat_vec (shamt : mword 6)
@@ -29746,7 +27103,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (6 + 26))
| SHIFTIOP (shamt,rs1,rd,RISCV_SRLI) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6)
(concat_vec (shamt : mword 6)
@@ -29754,7 +27111,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (6 + 26))
| SHIFTIOP (shamt,rs1,rd,RISCV_SRAI) =>
returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0] : mword 6)
(concat_vec (shamt : mword 6)
@@ -29762,7 +27119,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (6 + 26))
| RTYPE (rs2,rs1,rd,RISCV_ADD) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29770,7 +27127,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPE (rs2,rs1,rd,RISCV_SUB) =>
returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29778,7 +27135,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPE (rs2,rs1,rd,RISCV_SLL) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29786,7 +27143,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPE (rs2,rs1,rd,RISCV_SLT) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29794,7 +27151,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B1;B0] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPE (rs2,rs1,rd,RISCV_SLTU) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29802,7 +27159,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B1;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPE (rs2,rs1,rd,RISCV_XOR) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29810,7 +27167,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPE (rs2,rs1,rd,RISCV_SRL) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29818,7 +27175,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPE (rs2,rs1,rd,RISCV_SRA) =>
returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29826,7 +27183,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPE (rs2,rs1,rd,RISCV_OR) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29834,7 +27191,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B1;B1;B0] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPE (rs2,rs1,rd,RISCV_AND) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29842,7 +27199,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B1;B1;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| LOAD (imm,rs1,rd,is_unsigned,size,false,false) =>
returnm ((concat_vec (imm : mword 12)
(concat_vec (rs1 : mword 5)
@@ -29850,10 +27207,11 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (size_bits_forwards size)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B0;B0;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
- | STORE (v__174,rs2,rs1,size,false,false) =>
- let imm7 : bits 7 := subrange_vec_dec v__174 11 5 in
- let imm5 : bits 5 := subrange_vec_dec v__174 4 0 in
+ : mword (12 + 20))
+ | STORE (v__176,rs2,rs1,size,false,false) =>
+ let imm7 : bits 7 := subrange_vec_dec v__176 11 5 in
+ let imm7 : bits 7 := subrange_vec_dec v__176 11 5 in
+ let imm5 : bits 5 := subrange_vec_dec v__176 4 0 in
returnm ((concat_vec (imm7 : bits 7)
(concat_vec (rs2 : mword 5)
(concat_vec (rs1 : mword 5)
@@ -29867,7 +27225,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (rs1 : mword 5)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (rd : mword 5) (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7)))))
- : mword 32)
+ : mword (12 + 20))
| SHIFTW (shamt,rs1,rd,RISCV_SLLI) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (shamt : mword 5)
@@ -29875,7 +27233,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| SHIFTW (shamt,rs1,rd,RISCV_SRLI) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (shamt : mword 5)
@@ -29883,7 +27241,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| SHIFTW (shamt,rs1,rd,RISCV_SRAI) =>
returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (shamt : mword 5)
@@ -29891,7 +27249,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPEW (rs2,rs1,rd,RISCV_ADDW) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29899,7 +27257,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPEW (rs2,rs1,rd,RISCV_SUBW) =>
returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29907,7 +27265,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPEW (rs2,rs1,rd,RISCV_SLLW) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29915,7 +27273,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPEW (rs2,rs1,rd,RISCV_SRLW) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29923,7 +27281,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| RTYPEW (rs2,rs1,rd,RISCV_SRAW) =>
returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29931,7 +27289,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| SHIFTIWOP (shamt,rs1,rd,RISCV_SLLIW) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (shamt : mword 5)
@@ -29939,7 +27297,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| SHIFTIWOP (shamt,rs1,rd,RISCV_SRLIW) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (shamt : mword 5)
@@ -29947,7 +27305,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| SHIFTIWOP (shamt,rs1,rd,RISCV_SRAIW) =>
returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)
(concat_vec (shamt : mword 5)
@@ -29955,16 +27313,16 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| MUL (rs2,rs1,rd,high,signed1,signed2) =>
- encdec_mul_op_forwards high signed1 signed2 >>= fun w__0 : bits 3 =>
+ encdec_mul_op_forwards high signed1 signed2 >>= fun w__0 : mword 3 =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)
(concat_vec (rs2 : mword 5)
(concat_vec (rs1 : mword 5)
(concat_vec (w__0 : bits 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| DIV (rs2,rs1,rd,s) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29973,7 +27331,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (bool_not_bits_forwards s)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))))))
- : mword 32)
+ : mword (7 + 25))
| REM (rs2,rs1,rd,s) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29982,7 +27340,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (bool_not_bits_forwards s)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))))))
- : mword 32)
+ : mword (7 + 25))
| MULW (rs2,rs1,rd) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29990,7 +27348,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| DIVW (rs2,rs1,rd,s) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -29999,7 +27357,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (bool_not_bits_forwards s)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7)))))))
- : mword 32)
+ : mword (7 + 25))
| REMW (rs2,rs1,rd,s) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -30008,7 +27366,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (bool_not_bits_forwards s)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7)))))))
- : mword 32)
+ : mword (7 + 25))
| FENCE (pred,succ) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0] : mword 4)
(concat_vec (pred : mword 4)
@@ -30017,21 +27375,21 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
(vec_of_bits [B0;B0;B0;B1;B1;B1;B1] : mword 7)))))))
- : mword 32)
+ : mword (4 + 28))
| FENCEI (tt) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)
(concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
(concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
(concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
(vec_of_bits [B0;B0;B0;B1;B1;B1;B1] : mword 7)))))
- : mword 32)
+ : mword (12 + 20))
| ECALL (tt) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)
(concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
(vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7)))))
- : mword 32)
+ : mword (12 + 20))
| MRET (tt) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B1;B1;B0;B0;B0] : mword 7)
(concat_vec (vec_of_bits [B0;B0;B0;B1;B0] : mword 5)
@@ -30039,7 +27397,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
(vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| SRET (tt) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B1;B0;B0;B0] : mword 7)
(concat_vec (vec_of_bits [B0;B0;B0;B1;B0] : mword 5)
@@ -30047,21 +27405,21 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
(vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| EBREAK (tt) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12)
(concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
(vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7)))))
- : mword 32)
+ : mword (12 + 20))
| WFI (tt) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12)
(concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
(vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7)))))
- : mword 32)
+ : mword (12 + 20))
| SFENCE_VMA (rs1,rs2) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B1;B0;B0;B1] : mword 7)
(concat_vec (rs2 : mword 5)
@@ -30069,7 +27427,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
(concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
(vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (7 + 25))
| LOADRES (aq,rl,rs1,size,rd) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B1;B0] : mword 5)
(concat_vec (bool_bits_forwards aq)
@@ -30080,7 +27438,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (size_bits_forwards size)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))))))))
- : mword 32)
+ : mword (5 + 27))
| STORECON (aq,rl,rs2,rs1,size,rd) =>
returnm ((concat_vec (vec_of_bits [B0;B0;B0;B1;B1] : mword 5)
(concat_vec (bool_bits_forwards aq)
@@ -30091,7 +27449,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (size_bits_forwards size)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))))))))
- : mword 32)
+ : mword (5 + 27))
| AMO (op,aq,rl,rs2,rs1,size,rd) =>
returnm ((concat_vec (encdec_amoop_forwards op)
(concat_vec (bool_bits_forwards aq)
@@ -30102,7 +27460,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (size_bits_forwards size)
(concat_vec (rd : mword 5)
(vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))))))))
- : mword 32)
+ : mword (5 + 27))
| CSR (csr,rs1,rd,is_imm,op) =>
returnm ((concat_vec (csr : mword 12)
(concat_vec (rs1 : mword 5)
@@ -30110,7 +27468,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (encdec_csrop_forwards op)
(concat_vec (rd : mword 5)
(vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword 32)
+ : mword (12 + 20))
| STOP_FETCHING (tt) =>
returnm ((concat_vec
(vec_of_bits [B1;B1;B1;B1;B1;B0;B1;B0;B1;B1;B0;B1;B1;B1;B1;B0] : mword 16)
@@ -30119,7 +27477,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0] : mword 2)
(concat_vec (vec_of_bits [B0;B1;B0] : mword 3)
(vec_of_bits [B1;B1] : mword 2))))))
- : mword 32)
+ : mword (16 + 16))
| THREAD_START (tt) =>
returnm ((concat_vec
(vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B1;B1;B0;B1;B1;B1;B1;B0] : mword 16)
@@ -30128,7 +27486,7 @@ Definition encdec_forwards (arg_ : ast)
(concat_vec (vec_of_bits [B0;B0] : mword 2)
(concat_vec (vec_of_bits [B0;B1;B0] : mword 3)
(vec_of_bits [B1;B1] : mword 2))))))
- : mword 32)
+ : mword (16 + 16))
| ILLEGAL (s) => returnm (s : mword 32)
| _ => exit tt : M (mword 32)
end)
@@ -30136,782 +27494,747 @@ Definition encdec_forwards (arg_ : ast)
Definition encdec_backwards (arg_ : mword 32)
: ast :=
- let v__175 := arg_ in
- if ((let _mappingpatterns_23_ := subrange_vec_dec v__175 6 0 in
- encdec_uop_backwards_matches _mappingpatterns_23_)) then
- let imm : mword 20 := subrange_vec_dec v__175 31 12 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- let _mappingpatterns_23_ := subrange_vec_dec v__175 6 0 in
+ let v__177 := arg_ in
+ if ((let _mappingpatterns_23_ := subrange_vec_dec v__177 6 0 in
+ andb (encdec_uop_backwards_matches _mappingpatterns_23_)
+ (if ((encdec_uop_backwards_matches _mappingpatterns_23_)) then
+ let op := encdec_uop_backwards _mappingpatterns_23_ in
+ true
+ else false))) then
+ let imm : mword 20 := subrange_vec_dec v__177 31 12 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let imm : mword 20 := subrange_vec_dec v__177 31 12 in
+ let _mappingpatterns_23_ := subrange_vec_dec v__177 6 0 in
let op := encdec_uop_backwards _mappingpatterns_23_ in
- UTYPE (imm,rd,op)
- else if ((let p0_ := subrange_vec_dec v__175 6 0 in
- eq_vec p0_ (vec_of_bits [B1;B1;B0;B1;B1;B1;B1] : mword 7))) then
- let imm_19 : bits 1 := subrange_vec_dec v__175 31 31 in
- let imm_18_13 : bits 6 := subrange_vec_dec v__175 30 25 in
- let imm_12_9 : bits 4 := subrange_vec_dec v__175 24 21 in
- let imm_8 : bits 1 := subrange_vec_dec v__175 20 20 in
- let imm_7_0 : bits 8 := subrange_vec_dec v__175 19 12 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RISCV_JAL (concat_vec (imm_19 : bits 1)
- (concat_vec (imm_7_0 : bits 8)
- (concat_vec (imm_8 : bits 1)
- (concat_vec (imm_18_13 : bits 6)
- (concat_vec (imm_12_9 : bits 4) (vec_of_bits [B0] : mword 1))))),rd)
- else if ((let p0_ := subrange_vec_dec v__175 14 12 in
- let p1_ := subrange_vec_dec v__175 6 0 in
- andb (eq_vec p1_ (vec_of_bits [B1;B1;B0;B0;B1;B1;B1] : mword 7))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0] : mword 3)))) then
- let imm : mword 12 := subrange_vec_dec v__175 31 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RISCV_JALR (imm,rs1,rd)
- else if ((let _mappingpatterns_24_ := subrange_vec_dec v__175 14 12 in
- let p0_ := subrange_vec_dec v__175 6 0 in
- andb (encdec_bop_backwards_matches _mappingpatterns_24_)
- (eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B0;B1;B1] : mword 7)))) then
- let imm7_6 : bits 1 := subrange_vec_dec v__175 31 31 in
- let imm7_5_0 : bits 6 := subrange_vec_dec v__175 30 25 in
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let _mappingpatterns_24_ := subrange_vec_dec v__175 14 12 in
- let imm5_4_1 : bits 4 := subrange_vec_dec v__175 11 8 in
- let imm5_0 : bits 1 := subrange_vec_dec v__175 7 7 in
+ UTYPE
+ (imm, rd, op)
+ else if ((eq_vec (subrange_vec_dec v__177 6 0) (vec_of_bits [B1;B1;B0;B1;B1;B1;B1] : mword 7)))
+ then
+ let imm_19 : bits 1 := subrange_vec_dec v__177 31 31 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let imm_8 : bits 1 := subrange_vec_dec v__177 20 20 in
+ let imm_7_0 : bits 8 := subrange_vec_dec v__177 19 12 in
+ let imm_19 : bits 1 := subrange_vec_dec v__177 31 31 in
+ let imm_18_13 : bits 6 := subrange_vec_dec v__177 30 25 in
+ let imm_12_9 : bits 4 := subrange_vec_dec v__177 24 21 in
+ RISCV_JAL
+ (concat_vec (imm_19 : bits 1)
+ (concat_vec (imm_7_0 : bits 8)
+ (concat_vec (imm_8 : bits 1)
+ (concat_vec (imm_18_13 : bits 6)
+ (concat_vec (imm_12_9 : bits 4) (vec_of_bits [B0] : mword 1))))), rd)
+ else if ((andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0) (vec_of_bits [B1;B1;B0;B0;B1;B1;B1] : mword 7))))
+ then
+ let imm : mword 12 := subrange_vec_dec v__177 31 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let imm : mword 12 := subrange_vec_dec v__177 31 20 in
+ RISCV_JALR
+ (imm, rs1, rd)
+ else if ((andb
+ (let _mappingpatterns_24_ := subrange_vec_dec v__177 14 12 in
+ andb (encdec_bop_backwards_matches _mappingpatterns_24_)
+ (if ((encdec_bop_backwards_matches _mappingpatterns_24_)) then
+ let op := encdec_bop_backwards _mappingpatterns_24_ in
+ true
+ else false))
+ (eq_vec (subrange_vec_dec v__177 6 0) (vec_of_bits [B1;B1;B0;B0;B0;B1;B1] : mword 7))))
+ then
+ let imm7_6 : bits 1 := subrange_vec_dec v__177 31 31 in
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let imm7_6 : bits 1 := subrange_vec_dec v__177 31 31 in
+ let imm7_5_0 : bits 6 := subrange_vec_dec v__177 30 25 in
+ let imm5_4_1 : bits 4 := subrange_vec_dec v__177 11 8 in
+ let imm5_0 : bits 1 := subrange_vec_dec v__177 7 7 in
+ let _mappingpatterns_24_ := subrange_vec_dec v__177 14 12 in
let op := encdec_bop_backwards _mappingpatterns_24_ in
- BTYPE (concat_vec (imm7_6 : bits 1)
- (concat_vec (imm5_0 : bits 1)
- (concat_vec (imm7_5_0 : bits 6)
- (concat_vec (imm5_4_1 : bits 4) (vec_of_bits [B0] : mword 1)))),rs2,rs1,op)
- else if ((let _mappingpatterns_25_ := subrange_vec_dec v__175 14 12 in
- let p0_ := subrange_vec_dec v__175 6 0 in
- andb (encdec_iop_backwards_matches _mappingpatterns_25_)
- (eq_vec p0_ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7)))) then
- let imm : mword 12 := subrange_vec_dec v__175 31 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let _mappingpatterns_25_ := subrange_vec_dec v__175 14 12 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
+ BTYPE
+ (concat_vec (imm7_6 : bits 1)
+ (concat_vec (imm5_0 : bits 1)
+ (concat_vec (imm7_5_0 : bits 6)
+ (concat_vec (imm5_4_1 : bits 4) (vec_of_bits [B0] : mword 1)))), rs2, rs1, op)
+ else if ((andb
+ (let _mappingpatterns_25_ := subrange_vec_dec v__177 14 12 in
+ andb (encdec_iop_backwards_matches _mappingpatterns_25_)
+ (if ((encdec_iop_backwards_matches _mappingpatterns_25_)) then
+ let op := encdec_iop_backwards _mappingpatterns_25_ in
+ true
+ else false))
+ (eq_vec (subrange_vec_dec v__177 6 0) (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))
+ then
+ let imm : mword 12 := subrange_vec_dec v__177 31 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let imm : mword 12 := subrange_vec_dec v__177 31 20 in
+ let _mappingpatterns_25_ := subrange_vec_dec v__177 14 12 in
let op := encdec_iop_backwards _mappingpatterns_25_ in
- ITYPE (imm,rs1,rd,op)
- else if ((let p0_ := subrange_vec_dec v__175 31 26 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6)))) then
- let shamt : mword 6 := subrange_vec_dec v__175 25 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- SHIFTIOP (shamt,rs1,rd,RISCV_SLLI)
- else if ((let p0_ := subrange_vec_dec v__175 31 26 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6)))) then
- let shamt : mword 6 := subrange_vec_dec v__175 25 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- SHIFTIOP (shamt,rs1,rd,RISCV_SRLI)
- else if ((let p0_ := subrange_vec_dec v__175 31 26 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0] : mword 6)))) then
- let shamt : mword 6 := subrange_vec_dec v__175 25 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- SHIFTIOP (shamt,rs1,rd,RISCV_SRAI)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPE (rs2,rs1,rd,RISCV_ADD)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPE (rs2,rs1,rd,RISCV_SUB)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPE (rs2,rs1,rd,RISCV_SLL)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B1;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPE (rs2,rs1,rd,RISCV_SLT)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B1;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPE (rs2,rs1,rd,RISCV_SLTU)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPE (rs2,rs1,rd,RISCV_XOR)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPE (rs2,rs1,rd,RISCV_SRL)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPE (rs2,rs1,rd,RISCV_SRA)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B1;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPE (rs2,rs1,rd,RISCV_OR)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B1;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPE (rs2,rs1,rd,RISCV_AND)
- else if ((let _mappingpatterns_26_ := subrange_vec_dec v__175 14 14 in
- let _mappingpatterns_27_ := subrange_vec_dec v__175 13 12 in
- let p0_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (size_bits_backwards_matches _mappingpatterns_27_)
- (bool_bits_backwards_matches _mappingpatterns_26_))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B1;B1] : mword 7)))) then
- let imm : mword 12 := subrange_vec_dec v__175 31 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let _mappingpatterns_26_ := subrange_vec_dec v__175 14 14 in
- let _mappingpatterns_27_ := subrange_vec_dec v__175 13 12 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
+ ITYPE
+ (imm, rs1, rd, op)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 26) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))) then
+ let shamt : mword 6 := subrange_vec_dec v__177 25 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ SHIFTIOP
+ (shamt, rs1, rd, RISCV_SLLI)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 26) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))) then
+ let shamt : mword 6 := subrange_vec_dec v__177 25 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ SHIFTIOP
+ (shamt, rs1, rd, RISCV_SRLI)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 26) (vec_of_bits [B0;B1;B0;B0;B0;B0] : mword 6))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))) then
+ let shamt : mword 6 := subrange_vec_dec v__177 25 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ SHIFTIOP
+ (shamt, rs1, rd, RISCV_SRAI)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPE
+ (rs2, rs1, rd, RISCV_ADD)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPE
+ (rs2, rs1, rd, RISCV_SUB)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPE
+ (rs2, rs1, rd, RISCV_SLL)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B1;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPE
+ (rs2, rs1, rd, RISCV_SLT)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B1;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPE
+ (rs2, rs1, rd, RISCV_SLTU)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B1;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPE
+ (rs2, rs1, rd, RISCV_XOR)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPE
+ (rs2, rs1, rd, RISCV_SRL)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPE
+ (rs2, rs1, rd, RISCV_SRA)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B1;B1;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPE
+ (rs2, rs1, rd, RISCV_OR)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B1;B1;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPE
+ (rs2, rs1, rd, RISCV_AND)
+ else if ((andb
+ (let _mappingpatterns_27_ := subrange_vec_dec v__177 13 12 in
+ let _mappingpatterns_26_ := subrange_vec_dec v__177 14 14 in
+ andb (size_bits_backwards_matches _mappingpatterns_27_)
+ (if ((size_bits_backwards_matches _mappingpatterns_27_)) then
+ let size := size_bits_backwards _mappingpatterns_27_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_26_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_26_)) then
+ let is_unsigned := bool_bits_backwards _mappingpatterns_26_ in
+ true
+ else false)
+ else false))
+ (eq_vec (subrange_vec_dec v__177 6 0) (vec_of_bits [B0;B0;B0;B0;B0;B1;B1] : mword 7))))
+ then
+ let imm : mword 12 := subrange_vec_dec v__177 31 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let imm : mword 12 := subrange_vec_dec v__177 31 20 in
+ let _mappingpatterns_27_ := subrange_vec_dec v__177 13 12 in
+ let _mappingpatterns_26_ := subrange_vec_dec v__177 14 14 in
let size := size_bits_backwards _mappingpatterns_27_ in
let is_unsigned := bool_bits_backwards _mappingpatterns_26_ in
- LOAD (imm,rs1,rd,is_unsigned,size,false,false)
- else if ((let p0_ := subrange_vec_dec v__175 14 14 in
- let _mappingpatterns_28_ := subrange_vec_dec v__175 13 12 in
- let p1_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (size_bits_backwards_matches _mappingpatterns_28_)
- (eq_vec p1_ (vec_of_bits [B0;B1;B0;B0;B0;B1;B1] : mword 7)))
- (eq_vec p0_ (vec_of_bits [B0] : mword 1)))) then
- let imm7 : bits 7 := subrange_vec_dec v__175 31 25 in
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let _mappingpatterns_28_ := subrange_vec_dec v__175 13 12 in
- let imm5 : bits 5 := subrange_vec_dec v__175 11 7 in
+ LOAD
+ (imm, rs1, rd, is_unsigned, size, false, false)
+ else if ((andb
+ (let _mappingpatterns_28_ := subrange_vec_dec v__177 13 12 in
+ andb (size_bits_backwards_matches _mappingpatterns_28_)
+ (if ((size_bits_backwards_matches _mappingpatterns_28_)) then
+ let size := size_bits_backwards _mappingpatterns_28_ in
+ true
+ else false))
+ (andb (eq_vec (subrange_vec_dec v__177 14 14) (vec_of_bits [B0] : mword 1))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B0;B0;B0;B1;B1] : mword 7))))) then
+ let imm7 : bits 7 := subrange_vec_dec v__177 31 25 in
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let imm7 : bits 7 := subrange_vec_dec v__177 31 25 in
+ let imm5 : bits 5 := subrange_vec_dec v__177 11 7 in
+ let _mappingpatterns_28_ := subrange_vec_dec v__177 13 12 in
let size := size_bits_backwards _mappingpatterns_28_ in
- STORE (concat_vec (imm7 : bits 7) (imm5 : bits 5),rs2,rs1,size,false,false)
- else if ((let p0_ := subrange_vec_dec v__175 14 12 in
- let p1_ := subrange_vec_dec v__175 6 0 in
- andb (eq_vec p1_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0] : mword 3)))) then
- let imm : mword 12 := subrange_vec_dec v__175 31 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- ADDIW (imm,rs1,rd)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let shamt : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- SHIFTW (shamt,rs1,rd,RISCV_SLLI)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let shamt : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- SHIFTW (shamt,rs1,rd,RISCV_SRLI)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)))) then
- let shamt : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- SHIFTW (shamt,rs1,rd,RISCV_SRAI)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPEW (rs2,rs1,rd,RISCV_ADDW)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPEW (rs2,rs1,rd,RISCV_SUBW)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPEW (rs2,rs1,rd,RISCV_SLLW)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPEW (rs2,rs1,rd,RISCV_SRLW)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- RTYPEW (rs2,rs1,rd,RISCV_SRAW)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let shamt : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- SHIFTIWOP (shamt,rs1,rd,RISCV_SLLIW)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
- let shamt : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- SHIFTIWOP (shamt,rs1,rd,RISCV_SRLIW)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)))) then
- let shamt : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- SHIFTIWOP (shamt,rs1,rd,RISCV_SRAIW)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let _mappingpatterns_29_ : bits 3 := subrange_vec_dec v__175 14 12 in
- let p1_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (encdec_mul_op_backwards_matches _mappingpatterns_29_)
- (eq_vec p1_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let _mappingpatterns_29_ : bits 3 := subrange_vec_dec v__175 14 12 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
+ STORE
+ (concat_vec (imm7 : bits 7) (imm5 : bits 5), rs2, rs1, size, false, false)
+ else if ((andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0) (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))
+ then
+ let imm : mword 12 := subrange_vec_dec v__177 31 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let imm : mword 12 := subrange_vec_dec v__177 31 20 in
+ ADDIW
+ (imm, rs1, rd)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))) then
+ let shamt : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ SHIFTW
+ (shamt, rs1, rd, RISCV_SLLI)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))) then
+ let shamt : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ SHIFTW
+ (shamt, rs1, rd, RISCV_SRLI)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))) then
+ let shamt : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ SHIFTW
+ (shamt, rs1, rd, RISCV_SRAI)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPEW
+ (rs2, rs1, rd, RISCV_ADDW)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPEW
+ (rs2, rs1, rd, RISCV_SUBW)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPEW
+ (rs2, rs1, rd, RISCV_SLLW)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPEW
+ (rs2, rs1, rd, RISCV_SRLW)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ RTYPEW
+ (rs2, rs1, rd, RISCV_SRAW)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))) then
+ let shamt : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ SHIFTIWOP
+ (shamt, rs1, rd, RISCV_SLLIW)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))) then
+ let shamt : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ SHIFTIWOP
+ (shamt, rs1, rd, RISCV_SRLIW)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))) then
+ let shamt : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ SHIFTIWOP
+ (shamt, rs1, rd, RISCV_SRAIW)
+ else if ((andb
+ (let _mappingpatterns_29_ : bits 3 := subrange_vec_dec v__177 14 12 in
+ andb (encdec_mul_op_backwards_matches _mappingpatterns_29_)
+ (if ((encdec_mul_op_backwards_matches _mappingpatterns_29_)) then
+ let '(high, signed1, signed2) := encdec_mul_op_backwards _mappingpatterns_29_ in
+ true
+ else false))
+ (andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let _mappingpatterns_29_ : bits 3 := subrange_vec_dec v__177 14 12 in
let '(high, signed1, signed2) := encdec_mul_op_backwards _mappingpatterns_29_ in
- MUL (rs2,rs1,rd,high,signed1,signed2)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 13 in
- let _mappingpatterns_30_ := subrange_vec_dec v__175 12 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb (bool_not_bits_backwards_matches _mappingpatterns_30_)
- (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))
- (eq_vec p1_ (vec_of_bits [B1;B0] : mword 2)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let _mappingpatterns_30_ := subrange_vec_dec v__175 12 12 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
+ MUL
+ (rs2, rs1, rd, high, signed1, signed2)
+ else if ((andb
+ (let _mappingpatterns_30_ := subrange_vec_dec v__177 12 12 in
+ andb (bool_not_bits_backwards_matches _mappingpatterns_30_)
+ (if ((bool_not_bits_backwards_matches _mappingpatterns_30_)) then
+ let s := bool_not_bits_backwards _mappingpatterns_30_ in
+ true
+ else false))
+ (andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 13) (vec_of_bits [B1;B0] : mword 2))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let _mappingpatterns_30_ := subrange_vec_dec v__177 12 12 in
let s := bool_not_bits_backwards _mappingpatterns_30_ in
- DIV (rs2,rs1,rd,s)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 13 in
- let _mappingpatterns_31_ := subrange_vec_dec v__175 12 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb (bool_not_bits_backwards_matches _mappingpatterns_31_)
- (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))
- (eq_vec p1_ (vec_of_bits [B1;B1] : mword 2)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let _mappingpatterns_31_ := subrange_vec_dec v__175 12 12 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
+ DIV
+ (rs2, rs1, rd, s)
+ else if ((andb
+ (let _mappingpatterns_31_ := subrange_vec_dec v__177 12 12 in
+ andb (bool_not_bits_backwards_matches _mappingpatterns_31_)
+ (if ((bool_not_bits_backwards_matches _mappingpatterns_31_)) then
+ let s := bool_not_bits_backwards _mappingpatterns_31_ in
+ true
+ else false))
+ (andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 13) (vec_of_bits [B1;B1] : mword 2))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let _mappingpatterns_31_ := subrange_vec_dec v__177 12 12 in
let s := bool_not_bits_backwards _mappingpatterns_31_ in
- REM (rs2,rs1,rd,s)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
- MULW (rs2,rs1,rd)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 13 in
- let _mappingpatterns_32_ := subrange_vec_dec v__175 12 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb (bool_not_bits_backwards_matches _mappingpatterns_32_)
- (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7)))
- (eq_vec p1_ (vec_of_bits [B1;B0] : mword 2)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let _mappingpatterns_32_ := subrange_vec_dec v__175 12 12 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
+ REM
+ (rs2, rs1, rd, s)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ MULW
+ (rs2, rs1, rd)
+ else if ((andb
+ (let _mappingpatterns_32_ := subrange_vec_dec v__177 12 12 in
+ andb (bool_not_bits_backwards_matches _mappingpatterns_32_)
+ (if ((bool_not_bits_backwards_matches _mappingpatterns_32_)) then
+ let s := bool_not_bits_backwards _mappingpatterns_32_ in
+ true
+ else false))
+ (andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 13) (vec_of_bits [B1;B0] : mword 2))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7)))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let _mappingpatterns_32_ := subrange_vec_dec v__177 12 12 in
let s := bool_not_bits_backwards _mappingpatterns_32_ in
- DIVW (rs2,rs1,rd,s)
- else if ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 13 in
- let _mappingpatterns_33_ := subrange_vec_dec v__175 12 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb (bool_not_bits_backwards_matches _mappingpatterns_33_)
- (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7)))
- (eq_vec p1_ (vec_of_bits [B1;B1] : mword 2)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let _mappingpatterns_33_ := subrange_vec_dec v__175 12 12 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
+ DIVW
+ (rs2, rs1, rd, s)
+ else if ((andb
+ (let _mappingpatterns_33_ := subrange_vec_dec v__177 12 12 in
+ andb (bool_not_bits_backwards_matches _mappingpatterns_33_)
+ (if ((bool_not_bits_backwards_matches _mappingpatterns_33_)) then
+ let s := bool_not_bits_backwards _mappingpatterns_33_ in
+ true
+ else false))
+ (andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__177 14 13) (vec_of_bits [B1;B1] : mword 2))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7)))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let _mappingpatterns_33_ := subrange_vec_dec v__177 12 12 in
let s := bool_not_bits_backwards _mappingpatterns_33_ in
- REMW (rs2,rs1,rd,s)
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__175 31 28 in
- let p1_ := subrange_vec_dec v__175 19 15 in
- let p2_ := subrange_vec_dec v__175 14 12 in
- let p3_ := subrange_vec_dec v__175 11 7 in
- let p4_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb
- (andb (eq_vec p4_ (vec_of_bits [B0;B0;B0;B1;B1;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p3_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p2_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0] : mword 4)))) then
- let pred : mword 4 := subrange_vec_dec v__175 27 24 in
- let succ : mword 4 := subrange_vec_dec v__175 23 20 in
- FENCE (pred,succ)
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__175 31 20 in
- let p1_ := subrange_vec_dec v__175 19 15 in
- let p2_ := subrange_vec_dec v__175 14 12 in
- let p3_ := subrange_vec_dec v__175 11 7 in
- let p4_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb
- (andb (eq_vec p4_ (vec_of_bits [B0;B0;B0;B1;B1;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p3_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p2_ (vec_of_bits [B0;B0;B1] : mword 3)))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)))) then
- FENCEI tt
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__175 31 20 in
- let p1_ := subrange_vec_dec v__175 19 15 in
- let p2_ := subrange_vec_dec v__175 14 12 in
- let p3_ := subrange_vec_dec v__175 11 7 in
- let p4_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb
- (andb (eq_vec p4_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p3_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p2_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)))) then
- ECALL tt
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 24 20 in
- let p2_ := subrange_vec_dec v__175 19 15 in
- let p3_ := subrange_vec_dec v__175 14 12 in
- let p4_ := subrange_vec_dec v__175 11 7 in
- let p5_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb
- (andb
- (andb (eq_vec p5_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p4_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p3_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (Z.eqb
- (projT1 ((regbits_to_regno p2_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0] : mword 7)))) then
- MRET tt
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 24 20 in
- let p2_ := subrange_vec_dec v__175 19 15 in
- let p3_ := subrange_vec_dec v__175 14 12 in
- let p4_ := subrange_vec_dec v__175 11 7 in
- let p5_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb
- (andb
- (andb (eq_vec p5_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p4_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p3_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (Z.eqb
- (projT1 ((regbits_to_regno p2_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0] : mword 7)))) then
- SRET tt
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__175 31 20 in
- let p1_ := subrange_vec_dec v__175 19 15 in
- let p2_ := subrange_vec_dec v__175 14 12 in
- let p3_ := subrange_vec_dec v__175 11 7 in
- let p4_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb
- (andb (eq_vec p4_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p3_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p2_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12)))) then
- EBREAK tt
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__175 31 20 in
- let p1_ := subrange_vec_dec v__175 19 15 in
- let p2_ := subrange_vec_dec v__175 14 12 in
- let p3_ := subrange_vec_dec v__175 11 7 in
- let p4_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb
- (andb (eq_vec p4_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p3_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p2_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12)))) then
- WFI tt
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__175 31 25 in
- let p1_ := subrange_vec_dec v__175 14 12 in
- let p2_ := subrange_vec_dec v__175 11 7 in
- let p3_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb (eq_vec p3_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p2_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B1] : mword 7)))) then
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- SFENCE_VMA (rs1,rs2)
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__175 31 27 in
- let _mappingpatterns_34_ := subrange_vec_dec v__175 26 26 in
- let _mappingpatterns_35_ := subrange_vec_dec v__175 25 25 in
- let p1_ := subrange_vec_dec v__175 24 20 in
- let p2_ := subrange_vec_dec v__175 14 14 in
- let _mappingpatterns_36_ := subrange_vec_dec v__175 13 12 in
- let p3_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb
- (andb
- (andb
- (andb (size_bits_backwards_matches _mappingpatterns_36_)
- (bool_bits_backwards_matches _mappingpatterns_35_))
- (bool_bits_backwards_matches _mappingpatterns_34_))
- (eq_vec p3_ (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))
- (eq_vec p2_ (vec_of_bits [B0] : mword 1)))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (Z.eqb (projT1 ((regbits_to_regno p0_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))) then
- let _mappingpatterns_34_ := subrange_vec_dec v__175 26 26 in
- let _mappingpatterns_35_ := subrange_vec_dec v__175 25 25 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let _mappingpatterns_36_ := subrange_vec_dec v__175 13 12 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
+ REMW
+ (rs2, rs1, rd, s)
+ else if ((andb (eq_vec (subrange_vec_dec v__177 31 28) (vec_of_bits [B0;B0;B0;B0] : mword 4))
+ (eq_vec (subrange_vec_dec v__177 19 0)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1;B1;B1]
+ : mword 20)))) then
+ let succ : mword 4 := subrange_vec_dec v__177 23 20 in
+ let pred : mword 4 := subrange_vec_dec v__177 27 24 in
+ FENCE
+ (pred, succ)
+ else if ((eq_vec 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]
+ : mword 32))) then
+ FENCEI
+ (tt)
+ else if ((eq_vec 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]
+ : mword 32))) then
+ ECALL
+ (tt)
+ else if ((eq_vec 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]
+ : mword 32))) then
+ MRET
+ (tt)
+ else if ((eq_vec 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]
+ : mword 32))) then
+ SRET
+ (tt)
+ else if ((eq_vec 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]
+ : mword 32))) then
+ EBREAK
+ (tt)
+ else if ((eq_vec 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]
+ : mword 32))) then
+ WFI
+ (tt)
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__177 31 25)
+ (vec_of_bits [B0;B0;B0;B1;B0;B0;B1] : mword 7))
+ (eq_vec (subrange_vec_dec v__177 14 0)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1] : mword 15)))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ SFENCE_VMA
+ (rs1, rs2)
+ else if sumbool_of_bool ((andb
+ (let _mappingpatterns_36_ := subrange_vec_dec v__177 13 12 in
+ let _mappingpatterns_35_ := subrange_vec_dec v__177 25 25 in
+ let _mappingpatterns_34_ := subrange_vec_dec v__177 26 26 in
+ andb (size_bits_backwards_matches _mappingpatterns_36_)
+ (if ((size_bits_backwards_matches _mappingpatterns_36_)) then
+ let size := size_bits_backwards _mappingpatterns_36_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_35_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_35_)) then
+ let rl := bool_bits_backwards _mappingpatterns_35_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_34_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_34_))
+ then
+ let aq := bool_bits_backwards _mappingpatterns_34_ in
+ true
+ else false)
+ else false)
+ else false))
+ (andb
+ (Z.eqb
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (subrange_vec_dec v__177 31 27))))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B1;B0]
+ : mword 5))))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))
+ (andb
+ (Z.eqb
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (subrange_vec_dec v__177 24 20))))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B0]
+ : mword 5))))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))
+ (andb
+ (eq_vec (subrange_vec_dec v__177 14 14)
+ (vec_of_bits [B0] : mword 1))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))))))
+ then
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let _mappingpatterns_36_ := subrange_vec_dec v__177 13 12 in
+ let _mappingpatterns_35_ := subrange_vec_dec v__177 25 25 in
+ let _mappingpatterns_34_ := subrange_vec_dec v__177 26 26 in
let size := size_bits_backwards _mappingpatterns_36_ in
let rl := bool_bits_backwards _mappingpatterns_35_ in
let aq := bool_bits_backwards _mappingpatterns_34_ in
- LOADRES (aq,rl,rs1,size,rd)
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__175 31 27 in
- let _mappingpatterns_37_ := subrange_vec_dec v__175 26 26 in
- let _mappingpatterns_38_ := subrange_vec_dec v__175 25 25 in
- let p1_ := subrange_vec_dec v__175 14 14 in
- let _mappingpatterns_39_ := subrange_vec_dec v__175 13 12 in
- let p2_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb
- (andb
- (andb (size_bits_backwards_matches _mappingpatterns_39_)
- (bool_bits_backwards_matches _mappingpatterns_38_))
- (bool_bits_backwards_matches _mappingpatterns_37_))
- (eq_vec p2_ (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))
- (eq_vec p1_ (vec_of_bits [B0] : mword 1)))
- (Z.eqb (projT1 ((regbits_to_regno p0_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B1] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))) then
- let _mappingpatterns_37_ := subrange_vec_dec v__175 26 26 in
- let _mappingpatterns_38_ := subrange_vec_dec v__175 25 25 in
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let _mappingpatterns_39_ := subrange_vec_dec v__175 13 12 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
+ LOADRES
+ (aq, rl, rs1, size, rd)
+ else if sumbool_of_bool ((andb
+ (let _mappingpatterns_39_ := subrange_vec_dec v__177 13 12 in
+ let _mappingpatterns_38_ := subrange_vec_dec v__177 25 25 in
+ let _mappingpatterns_37_ := subrange_vec_dec v__177 26 26 in
+ andb (size_bits_backwards_matches _mappingpatterns_39_)
+ (if ((size_bits_backwards_matches _mappingpatterns_39_)) then
+ let size := size_bits_backwards _mappingpatterns_39_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_38_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_38_)) then
+ let rl := bool_bits_backwards _mappingpatterns_38_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_37_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_37_))
+ then
+ let aq := bool_bits_backwards _mappingpatterns_37_ in
+ true
+ else false)
+ else false)
+ else false))
+ (andb
+ (Z.eqb
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (subrange_vec_dec v__177 31 27))))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B1;B1]
+ : mword 5))))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))
+ (andb
+ (eq_vec (subrange_vec_dec v__177 14 14)
+ (vec_of_bits [B0] : mword 1))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let _mappingpatterns_39_ := subrange_vec_dec v__177 13 12 in
+ let _mappingpatterns_38_ := subrange_vec_dec v__177 25 25 in
+ let _mappingpatterns_37_ := subrange_vec_dec v__177 26 26 in
let size := size_bits_backwards _mappingpatterns_39_ in
let rl := bool_bits_backwards _mappingpatterns_38_ in
let aq := bool_bits_backwards _mappingpatterns_37_ in
- STORECON (aq,rl,rs2,rs1,size,rd)
- else if ((let _mappingpatterns_40_ := subrange_vec_dec v__175 31 27 in
- let _mappingpatterns_41_ := subrange_vec_dec v__175 26 26 in
- let _mappingpatterns_42_ := subrange_vec_dec v__175 25 25 in
- let p0_ := subrange_vec_dec v__175 14 14 in
- let _mappingpatterns_43_ := subrange_vec_dec v__175 13 12 in
- let p1_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb
- (andb
- (andb
- (andb (size_bits_backwards_matches _mappingpatterns_43_)
- (bool_bits_backwards_matches _mappingpatterns_42_))
- (bool_bits_backwards_matches _mappingpatterns_41_))
- (encdec_amoop_backwards_matches _mappingpatterns_40_))
- (eq_vec p1_ (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))
- (eq_vec p0_ (vec_of_bits [B0] : mword 1)))) then
- let _mappingpatterns_40_ := subrange_vec_dec v__175 31 27 in
- let _mappingpatterns_41_ := subrange_vec_dec v__175 26 26 in
- let _mappingpatterns_42_ := subrange_vec_dec v__175 25 25 in
- let rs2 : mword 5 := subrange_vec_dec v__175 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let _mappingpatterns_43_ := subrange_vec_dec v__175 13 12 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
+ STORECON
+ (aq, rl, rs2, rs1, size, rd)
+ else if ((andb
+ (let _mappingpatterns_40_ := subrange_vec_dec v__177 31 27 in
+ let _mappingpatterns_43_ := subrange_vec_dec v__177 13 12 in
+ let _mappingpatterns_42_ := subrange_vec_dec v__177 25 25 in
+ let _mappingpatterns_41_ := subrange_vec_dec v__177 26 26 in
+ let _mappingpatterns_40_ := subrange_vec_dec v__177 31 27 in
+ andb (size_bits_backwards_matches _mappingpatterns_43_)
+ (if ((size_bits_backwards_matches _mappingpatterns_43_)) then
+ let size := size_bits_backwards _mappingpatterns_43_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_42_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_42_)) then
+ let rl := bool_bits_backwards _mappingpatterns_42_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_41_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_41_)) then
+ let aq := bool_bits_backwards _mappingpatterns_41_ in
+ andb (encdec_amoop_backwards_matches _mappingpatterns_40_)
+ (if ((encdec_amoop_backwards_matches _mappingpatterns_40_)) then
+ let op := encdec_amoop_backwards _mappingpatterns_40_ in
+ true
+ else false)
+ else false)
+ else false)
+ else false))
+ (andb (eq_vec (subrange_vec_dec v__177 14 14) (vec_of_bits [B0] : mword 1))
+ (eq_vec (subrange_vec_dec v__177 6 0)
+ (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7))))) then
+ let _mappingpatterns_40_ := subrange_vec_dec v__177 31 27 in
+ let rs2 : mword 5 := subrange_vec_dec v__177 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let _mappingpatterns_43_ := subrange_vec_dec v__177 13 12 in
+ let _mappingpatterns_42_ := subrange_vec_dec v__177 25 25 in
+ let _mappingpatterns_41_ := subrange_vec_dec v__177 26 26 in
+ let _mappingpatterns_40_ := subrange_vec_dec v__177 31 27 in
let size := size_bits_backwards _mappingpatterns_43_ in
let rl := bool_bits_backwards _mappingpatterns_42_ in
let aq := bool_bits_backwards _mappingpatterns_41_ in
let op := encdec_amoop_backwards _mappingpatterns_40_ in
- AMO (op,aq,rl,rs2,rs1,size,rd)
- else if ((let _mappingpatterns_44_ := subrange_vec_dec v__175 14 14 in
- let _mappingpatterns_45_ := subrange_vec_dec v__175 13 12 in
- let p0_ := subrange_vec_dec v__175 6 0 in
- andb
- (andb (encdec_csrop_backwards_matches _mappingpatterns_45_)
- (bool_bits_backwards_matches _mappingpatterns_44_))
- (eq_vec p0_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7)))) then
- let csr : mword 12 := subrange_vec_dec v__175 31 20 in
- let rs1 : mword 5 := subrange_vec_dec v__175 19 15 in
- let _mappingpatterns_44_ := subrange_vec_dec v__175 14 14 in
- let _mappingpatterns_45_ := subrange_vec_dec v__175 13 12 in
- let rd : mword 5 := subrange_vec_dec v__175 11 7 in
+ AMO
+ (op, aq, rl, rs2, rs1, size, rd)
+ else if ((andb
+ (let _mappingpatterns_45_ := subrange_vec_dec v__177 13 12 in
+ let _mappingpatterns_44_ := subrange_vec_dec v__177 14 14 in
+ andb (encdec_csrop_backwards_matches _mappingpatterns_45_)
+ (if ((encdec_csrop_backwards_matches _mappingpatterns_45_)) then
+ let op := encdec_csrop_backwards _mappingpatterns_45_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_44_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_44_)) then
+ let is_imm := bool_bits_backwards _mappingpatterns_44_ in
+ true
+ else false)
+ else false))
+ (eq_vec (subrange_vec_dec v__177 6 0) (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))))
+ then
+ let csr : mword 12 := subrange_vec_dec v__177 31 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__177 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__177 11 7 in
+ let csr : mword 12 := subrange_vec_dec v__177 31 20 in
+ let _mappingpatterns_45_ := subrange_vec_dec v__177 13 12 in
+ let _mappingpatterns_44_ := subrange_vec_dec v__177 14 14 in
let op := encdec_csrop_backwards _mappingpatterns_45_ in
let is_imm := bool_bits_backwards _mappingpatterns_44_ in
- CSR (csr,rs1,rd,is_imm,op)
- else if ((let p0_ := subrange_vec_dec v__175 31 16 in
- let p1_ := subrange_vec_dec v__175 15 8 in
- let p2_ := subrange_vec_dec v__175 7 7 in
- let p3_ := subrange_vec_dec v__175 6 5 in
- let p4_ := subrange_vec_dec v__175 4 2 in
- let p5_ := subrange_vec_dec v__175 1 0 in
- andb
- (andb
- (andb
- (andb
- (andb (eq_vec p5_ (vec_of_bits [B1;B1] : mword 2))
- (eq_vec p4_ (vec_of_bits [B0;B1;B0] : mword 3)))
- (eq_vec p3_ (vec_of_bits [B0;B0] : mword 2)))
- (eq_vec p2_ (vec_of_bits [B0] : mword 1)))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword 8)))
- (eq_vec p0_ (vec_of_bits [B1;B1;B1;B1;B1;B0;B1;B0;B1;B1;B0;B1;B1;B1;B1;B0] : mword 16))))
- then
- STOP_FETCHING tt
- else if ((let p0_ := subrange_vec_dec v__175 31 16 in
- let p1_ := subrange_vec_dec v__175 15 8 in
- let p2_ := subrange_vec_dec v__175 7 7 in
- let p3_ := subrange_vec_dec v__175 6 5 in
- let p4_ := subrange_vec_dec v__175 4 2 in
- let p5_ := subrange_vec_dec v__175 1 0 in
- andb
- (andb
- (andb
- (andb
- (andb (eq_vec p5_ (vec_of_bits [B1;B1] : mword 2))
- (eq_vec p4_ (vec_of_bits [B0;B1;B0] : mword 3)))
- (eq_vec p3_ (vec_of_bits [B0;B0] : mword 2)))
- (eq_vec p2_ (vec_of_bits [B0] : mword 1)))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword 8)))
- (eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B1;B1;B0;B1;B1;B1;B1;B0] : mword 16))))
- then
- THREAD_START tt
- else ILLEGAL v__175.
+ CSR
+ (csr, rs1, rd, is_imm, op)
+ else if ((eq_vec 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]
+ : mword 32))) then
+ STOP_FETCHING
+ (tt)
+ else if ((eq_vec 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]
+ : mword 32))) then
+ THREAD_START
+ (tt)
+ else ILLEGAL (v__177).
Definition encdec_forwards_matches (arg_ : ast)
: bool :=
match arg_ with
| UTYPE (imm,rd,op) => true
- | RISCV_JAL (v__227,rd) =>
- if ((let p0_ := subrange_vec_dec v__227 0 0 in
- eq_vec p0_ (vec_of_bits [B0] : mword 1))) then
- true
+ | RISCV_JAL (v__391,rd) =>
+ if ((eq_vec (subrange_vec_dec v__391 0 0) (vec_of_bits [B0] : mword 1))) then true
else
- let g__17 := RISCV_JAL (v__227,rd) in
+ let g__13 := RISCV_JAL (v__391, rd) in
false
| RISCV_JALR (imm,rs1,rd) => true
- | BTYPE (v__228,rs2,rs1,op) =>
- if ((let p0_ := subrange_vec_dec v__228 0 0 in
- eq_vec p0_ (vec_of_bits [B0] : mword 1))) then
- true
+ | BTYPE (v__393,rs2,rs1,op) =>
+ if ((eq_vec (subrange_vec_dec v__393 0 0) (vec_of_bits [B0] : mword 1))) then true
else
- let g__17 := BTYPE (v__228,rs2,rs1,op) in
+ let g__13 := BTYPE (v__393, rs2, rs1, op) in
false
| ITYPE (imm,rs1,rd,op) => true
| SHIFTIOP (shamt,rs1,rd,RISCV_SLLI) => true
@@ -30928,7 +28251,7 @@ Definition encdec_forwards_matches (arg_ : ast)
| RTYPE (rs2,rs1,rd,RISCV_OR) => true
| RTYPE (rs2,rs1,rd,RISCV_AND) => true
| LOAD (imm,rs1,rd,is_unsigned,size,false,false) => true
- | STORE (v__229,rs2,rs1,size,false,false) => true
+ | STORE (v__395,rs2,rs1,size,false,false) => true
| ADDIW (imm,rs1,rd) => true
| SHIFTW (shamt,rs1,rd,RISCV_SLLI) => true
| SHIFTW (shamt,rs1,rd,RISCV_SRLI) => true
@@ -30962,631 +28285,539 @@ Definition encdec_forwards_matches (arg_ : ast)
| STOP_FETCHING (tt) => true
| THREAD_START (tt) => true
| ILLEGAL (s) => true
- | g__17 => false
+ | g__13 => false
end.
Definition encdec_backwards_matches (arg_ : mword 32)
: bool :=
- let v__230 := arg_ in
- if ((let _mappingpatterns_0_ := subrange_vec_dec v__230 6 0 in
- encdec_uop_backwards_matches _mappingpatterns_0_)) then
- let _mappingpatterns_0_ := subrange_vec_dec v__230 6 0 in
+ let v__396 := arg_ in
+ if ((let _mappingpatterns_0_ := subrange_vec_dec v__396 6 0 in
+ andb (encdec_uop_backwards_matches _mappingpatterns_0_)
+ (if ((encdec_uop_backwards_matches _mappingpatterns_0_)) then
+ let op := encdec_uop_backwards _mappingpatterns_0_ in
+ true
+ else false))) then
+ let _mappingpatterns_0_ := subrange_vec_dec v__396 6 0 in
let op := encdec_uop_backwards _mappingpatterns_0_ in
true
- else if ((let p0_ := subrange_vec_dec v__230 6 0 in
- eq_vec p0_ (vec_of_bits [B1;B1;B0;B1;B1;B1;B1] : mword 7))) then
+ else if ((eq_vec (subrange_vec_dec v__396 6 0) (vec_of_bits [B1;B1;B0;B1;B1;B1;B1] : mword 7)))
+ then
true
- else if ((let p0_ := subrange_vec_dec v__230 14 12 in
- let p1_ := subrange_vec_dec v__230 6 0 in
- andb (eq_vec p1_ (vec_of_bits [B1;B1;B0;B0;B1;B1;B1] : mword 7))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0] : mword 3)))) then
+ else if ((andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0) (vec_of_bits [B1;B1;B0;B0;B1;B1;B1] : mword 7))))
+ then
true
- else if ((let _mappingpatterns_1_ := subrange_vec_dec v__230 14 12 in
- let p0_ := subrange_vec_dec v__230 6 0 in
- andb (encdec_bop_backwards_matches _mappingpatterns_1_)
- (eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B0;B1;B1] : mword 7)))) then
- let _mappingpatterns_1_ := subrange_vec_dec v__230 14 12 in
+ else if ((andb
+ (let _mappingpatterns_1_ := subrange_vec_dec v__396 14 12 in
+ andb (encdec_bop_backwards_matches _mappingpatterns_1_)
+ (if ((encdec_bop_backwards_matches _mappingpatterns_1_)) then
+ let op := encdec_bop_backwards _mappingpatterns_1_ in
+ true
+ else false))
+ (eq_vec (subrange_vec_dec v__396 6 0) (vec_of_bits [B1;B1;B0;B0;B0;B1;B1] : mword 7))))
+ then
+ let _mappingpatterns_1_ := subrange_vec_dec v__396 14 12 in
let op := encdec_bop_backwards _mappingpatterns_1_ in
true
- else if ((let _mappingpatterns_2_ := subrange_vec_dec v__230 14 12 in
- let p0_ := subrange_vec_dec v__230 6 0 in
- andb (encdec_iop_backwards_matches _mappingpatterns_2_)
- (eq_vec p0_ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7)))) then
- let _mappingpatterns_2_ := subrange_vec_dec v__230 14 12 in
+ else if ((andb
+ (let _mappingpatterns_2_ := subrange_vec_dec v__396 14 12 in
+ andb (encdec_iop_backwards_matches _mappingpatterns_2_)
+ (if ((encdec_iop_backwards_matches _mappingpatterns_2_)) then
+ let op := encdec_iop_backwards _mappingpatterns_2_ in
+ true
+ else false))
+ (eq_vec (subrange_vec_dec v__396 6 0) (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))
+ then
+ let _mappingpatterns_2_ := subrange_vec_dec v__396 14 12 in
let op := encdec_iop_backwards _mappingpatterns_2_ in
true
- else if ((let p0_ := subrange_vec_dec v__230 31 26 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 26) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 26 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 26) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 26 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0] : mword 6)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 26) (vec_of_bits [B0;B1;B0;B0;B0;B0] : mword 6))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B1;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B1;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B1;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B1;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B1;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B1;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B1;B1;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B1;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B1;B1;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
true
- else if ((let _mappingpatterns_3_ := subrange_vec_dec v__230 14 14 in
- let _mappingpatterns_4_ := subrange_vec_dec v__230 13 12 in
- let p0_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (size_bits_backwards_matches _mappingpatterns_4_)
- (bool_bits_backwards_matches _mappingpatterns_3_))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B1;B1] : mword 7)))) then
- let _mappingpatterns_3_ := subrange_vec_dec v__230 14 14 in
- let _mappingpatterns_4_ := subrange_vec_dec v__230 13 12 in
+ else if ((andb
+ (let _mappingpatterns_4_ := subrange_vec_dec v__396 13 12 in
+ let _mappingpatterns_3_ := subrange_vec_dec v__396 14 14 in
+ andb (size_bits_backwards_matches _mappingpatterns_4_)
+ (if ((size_bits_backwards_matches _mappingpatterns_4_)) then
+ let size := size_bits_backwards _mappingpatterns_4_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_3_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_3_)) then
+ let is_unsigned := bool_bits_backwards _mappingpatterns_3_ in
+ true
+ else false)
+ else false))
+ (eq_vec (subrange_vec_dec v__396 6 0) (vec_of_bits [B0;B0;B0;B0;B0;B1;B1] : mword 7))))
+ then
+ let _mappingpatterns_4_ := subrange_vec_dec v__396 13 12 in
+ let _mappingpatterns_3_ := subrange_vec_dec v__396 14 14 in
let size := size_bits_backwards _mappingpatterns_4_ in
let is_unsigned := bool_bits_backwards _mappingpatterns_3_ in
true
- else if ((let p0_ := subrange_vec_dec v__230 14 14 in
- let _mappingpatterns_5_ := subrange_vec_dec v__230 13 12 in
- let p1_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (size_bits_backwards_matches _mappingpatterns_5_)
- (eq_vec p1_ (vec_of_bits [B0;B1;B0;B0;B0;B1;B1] : mword 7)))
- (eq_vec p0_ (vec_of_bits [B0] : mword 1)))) then
- let _mappingpatterns_5_ := subrange_vec_dec v__230 13 12 in
+ else if ((andb
+ (let _mappingpatterns_5_ := subrange_vec_dec v__396 13 12 in
+ andb (size_bits_backwards_matches _mappingpatterns_5_)
+ (if ((size_bits_backwards_matches _mappingpatterns_5_)) then
+ let size := size_bits_backwards _mappingpatterns_5_ in
+ true
+ else false))
+ (andb (eq_vec (subrange_vec_dec v__396 14 14) (vec_of_bits [B0] : mword 1))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B0;B0;B0;B1;B1] : mword 7))))) then
+ let _mappingpatterns_5_ := subrange_vec_dec v__396 13 12 in
let size := size_bits_backwards _mappingpatterns_5_ in
true
- else if ((let p0_ := subrange_vec_dec v__230 14 12 in
- let p1_ := subrange_vec_dec v__230 6 0 in
- andb (eq_vec p1_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0] : mword 3)))) then
+ else if ((andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0) (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))
+ then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B1;B0;B1] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B1;B0;B1] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let _mappingpatterns_6_ : bits 3 := subrange_vec_dec v__230 14 12 in
- let p1_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (encdec_mul_op_backwards_matches _mappingpatterns_6_)
- (eq_vec p1_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)))) then
- let _mappingpatterns_6_ : bits 3 := subrange_vec_dec v__230 14 12 in
+ else if ((andb
+ (let _mappingpatterns_6_ : bits 3 := subrange_vec_dec v__396 14 12 in
+ andb (encdec_mul_op_backwards_matches _mappingpatterns_6_)
+ (if ((encdec_mul_op_backwards_matches _mappingpatterns_6_)) then
+ let '(high, signed1, signed2) := encdec_mul_op_backwards _mappingpatterns_6_ in
+ true
+ else false))
+ (andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))) then
+ let _mappingpatterns_6_ : bits 3 := subrange_vec_dec v__396 14 12 in
let '(high, signed1, signed2) := encdec_mul_op_backwards _mappingpatterns_6_ in
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 13 in
- let _mappingpatterns_7_ := subrange_vec_dec v__230 12 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb (bool_not_bits_backwards_matches _mappingpatterns_7_)
- (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))
- (eq_vec p1_ (vec_of_bits [B1;B0] : mword 2)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)))) then
- let _mappingpatterns_7_ := subrange_vec_dec v__230 12 12 in
+ else if ((andb
+ (let _mappingpatterns_7_ := subrange_vec_dec v__396 12 12 in
+ andb (bool_not_bits_backwards_matches _mappingpatterns_7_)
+ (if ((bool_not_bits_backwards_matches _mappingpatterns_7_)) then
+ let s := bool_not_bits_backwards _mappingpatterns_7_ in
+ true
+ else false))
+ (andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 13) (vec_of_bits [B1;B0] : mword 2))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))))) then
+ let _mappingpatterns_7_ := subrange_vec_dec v__396 12 12 in
let s := bool_not_bits_backwards _mappingpatterns_7_ in
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 13 in
- let _mappingpatterns_8_ := subrange_vec_dec v__230 12 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb (bool_not_bits_backwards_matches _mappingpatterns_8_)
- (eq_vec p2_ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))
- (eq_vec p1_ (vec_of_bits [B1;B1] : mword 2)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)))) then
- let _mappingpatterns_8_ := subrange_vec_dec v__230 12 12 in
+ else if ((andb
+ (let _mappingpatterns_8_ := subrange_vec_dec v__396 12 12 in
+ andb (bool_not_bits_backwards_matches _mappingpatterns_8_)
+ (if ((bool_not_bits_backwards_matches _mappingpatterns_8_)) then
+ let s := bool_not_bits_backwards _mappingpatterns_8_ in
+ true
+ else false))
+ (andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 13) (vec_of_bits [B1;B1] : mword 2))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))))) then
+ let _mappingpatterns_8_ := subrange_vec_dec v__396 12 12 in
let s := bool_not_bits_backwards _mappingpatterns_8_ in
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 12) (vec_of_bits [B0;B0;B0] : mword 3))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 13 in
- let _mappingpatterns_9_ := subrange_vec_dec v__230 12 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb (bool_not_bits_backwards_matches _mappingpatterns_9_)
- (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7)))
- (eq_vec p1_ (vec_of_bits [B1;B0] : mword 2)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)))) then
- let _mappingpatterns_9_ := subrange_vec_dec v__230 12 12 in
+ else if ((andb
+ (let _mappingpatterns_9_ := subrange_vec_dec v__396 12 12 in
+ andb (bool_not_bits_backwards_matches _mappingpatterns_9_)
+ (if ((bool_not_bits_backwards_matches _mappingpatterns_9_)) then
+ let s := bool_not_bits_backwards _mappingpatterns_9_ in
+ true
+ else false))
+ (andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 13) (vec_of_bits [B1;B0] : mword 2))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7)))))) then
+ let _mappingpatterns_9_ := subrange_vec_dec v__396 12 12 in
let s := bool_not_bits_backwards _mappingpatterns_9_ in
true
- else if ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 13 in
- let _mappingpatterns_10_ := subrange_vec_dec v__230 12 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb (bool_not_bits_backwards_matches _mappingpatterns_10_)
- (eq_vec p2_ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7)))
- (eq_vec p1_ (vec_of_bits [B1;B1] : mword 2)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)))) then
- let _mappingpatterns_10_ := subrange_vec_dec v__230 12 12 in
+ else if ((andb
+ (let _mappingpatterns_10_ := subrange_vec_dec v__396 12 12 in
+ andb (bool_not_bits_backwards_matches _mappingpatterns_10_)
+ (if ((bool_not_bits_backwards_matches _mappingpatterns_10_)) then
+ let s := bool_not_bits_backwards _mappingpatterns_10_ in
+ true
+ else false))
+ (andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7))
+ (andb (eq_vec (subrange_vec_dec v__396 14 13) (vec_of_bits [B1;B1] : mword 2))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7)))))) then
+ let _mappingpatterns_10_ := subrange_vec_dec v__396 12 12 in
let s := bool_not_bits_backwards _mappingpatterns_10_ in
true
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__230 31 28 in
- let p1_ := subrange_vec_dec v__230 19 15 in
- let p2_ := subrange_vec_dec v__230 14 12 in
- let p3_ := subrange_vec_dec v__230 11 7 in
- let p4_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb
- (andb (eq_vec p4_ (vec_of_bits [B0;B0;B0;B1;B1;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p3_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p2_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0] : mword 4)))) then
+ else if ((andb (eq_vec (subrange_vec_dec v__396 31 28) (vec_of_bits [B0;B0;B0;B0] : mword 4))
+ (eq_vec (subrange_vec_dec v__396 19 0)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1;B1;B1]
+ : mword 20)))) then
true
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__230 31 20 in
- let p1_ := subrange_vec_dec v__230 19 15 in
- let p2_ := subrange_vec_dec v__230 14 12 in
- let p3_ := subrange_vec_dec v__230 11 7 in
- let p4_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb
- (andb (eq_vec p4_ (vec_of_bits [B0;B0;B0;B1;B1;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p3_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p2_ (vec_of_bits [B0;B0;B1] : mword 3)))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)))) then
+ else if ((eq_vec 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]
+ : mword 32))) then
true
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__230 31 20 in
- let p1_ := subrange_vec_dec v__230 19 15 in
- let p2_ := subrange_vec_dec v__230 14 12 in
- let p3_ := subrange_vec_dec v__230 11 7 in
- let p4_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb
- (andb (eq_vec p4_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p3_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p2_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)))) then
+ else if ((eq_vec 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]
+ : mword 32))) then
true
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 24 20 in
- let p2_ := subrange_vec_dec v__230 19 15 in
- let p3_ := subrange_vec_dec v__230 14 12 in
- let p4_ := subrange_vec_dec v__230 11 7 in
- let p5_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb
- (andb
- (andb (eq_vec p5_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p4_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p3_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (Z.eqb
- (projT1 ((regbits_to_regno p2_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B1;B1;B0;B0;B0] : mword 7)))) then
+ else if ((eq_vec 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]
+ : mword 32))) then
true
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 24 20 in
- let p2_ := subrange_vec_dec v__230 19 15 in
- let p3_ := subrange_vec_dec v__230 14 12 in
- let p4_ := subrange_vec_dec v__230 11 7 in
- let p5_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb
- (andb
- (andb (eq_vec p5_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p4_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p3_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (Z.eqb
- (projT1 ((regbits_to_regno p2_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0] : mword 7)))) then
+ else if ((eq_vec 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]
+ : mword 32))) then
true
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__230 31 20 in
- let p1_ := subrange_vec_dec v__230 19 15 in
- let p2_ := subrange_vec_dec v__230 14 12 in
- let p3_ := subrange_vec_dec v__230 11 7 in
- let p4_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb
- (andb (eq_vec p4_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p3_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p2_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12)))) then
+ else if ((eq_vec 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]
+ : mword 32))) then
true
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__230 31 20 in
- let p1_ := subrange_vec_dec v__230 19 15 in
- let p2_ := subrange_vec_dec v__230 14 12 in
- let p3_ := subrange_vec_dec v__230 11 7 in
- let p4_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb
- (andb (eq_vec p4_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p3_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p2_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12)))) then
+ else if ((eq_vec 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]
+ : mword 32))) then
true
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__230 31 25 in
- let p1_ := subrange_vec_dec v__230 14 12 in
- let p2_ := subrange_vec_dec v__230 11 7 in
- let p3_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb (eq_vec p3_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))
- (Z.eqb
- (projT1 ((regbits_to_regno p2_)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0] : mword 3)))
- (eq_vec p0_ (vec_of_bits [B0;B0;B0;B1;B0;B0;B1] : mword 7)))) then
+ else if ((andb
+ (eq_vec (subrange_vec_dec v__396 31 25)
+ (vec_of_bits [B0;B0;B0;B1;B0;B0;B1] : mword 7))
+ (eq_vec (subrange_vec_dec v__396 14 0)
+ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1] : mword 15)))) then
true
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__230 31 27 in
- let _mappingpatterns_11_ := subrange_vec_dec v__230 26 26 in
- let _mappingpatterns_12_ := subrange_vec_dec v__230 25 25 in
- let p1_ := subrange_vec_dec v__230 24 20 in
- let p2_ := subrange_vec_dec v__230 14 14 in
- let _mappingpatterns_13_ := subrange_vec_dec v__230 13 12 in
- let p3_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb
- (andb
- (andb
- (andb (size_bits_backwards_matches _mappingpatterns_13_)
- (bool_bits_backwards_matches _mappingpatterns_12_))
- (bool_bits_backwards_matches _mappingpatterns_11_))
- (eq_vec p3_ (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))
- (eq_vec p2_ (vec_of_bits [B0] : mword 1)))
- (Z.eqb (projT1 ((regbits_to_regno p1_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))
- (Z.eqb (projT1 ((regbits_to_regno p0_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B0] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))) then
- let _mappingpatterns_11_ := subrange_vec_dec v__230 26 26 in
- let _mappingpatterns_12_ := subrange_vec_dec v__230 25 25 in
- let _mappingpatterns_13_ := subrange_vec_dec v__230 13 12 in
+ else if sumbool_of_bool ((andb
+ (let _mappingpatterns_13_ := subrange_vec_dec v__396 13 12 in
+ let _mappingpatterns_12_ := subrange_vec_dec v__396 25 25 in
+ let _mappingpatterns_11_ := subrange_vec_dec v__396 26 26 in
+ andb (size_bits_backwards_matches _mappingpatterns_13_)
+ (if ((size_bits_backwards_matches _mappingpatterns_13_)) then
+ let size := size_bits_backwards _mappingpatterns_13_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_12_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_12_)) then
+ let rl := bool_bits_backwards _mappingpatterns_12_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_11_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_11_))
+ then
+ let aq := bool_bits_backwards _mappingpatterns_11_ in
+ true
+ else false)
+ else false)
+ else false))
+ (andb
+ (Z.eqb
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (subrange_vec_dec v__396 31 27))))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B1;B0]
+ : mword 5))))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))
+ (andb
+ (Z.eqb
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (subrange_vec_dec v__396 24 20))))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B0;B0]
+ : mword 5))))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))
+ (andb
+ (eq_vec (subrange_vec_dec v__396 14 14)
+ (vec_of_bits [B0] : mword 1))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))))))
+ then
+ let _mappingpatterns_13_ := subrange_vec_dec v__396 13 12 in
+ let _mappingpatterns_12_ := subrange_vec_dec v__396 25 25 in
+ let _mappingpatterns_11_ := subrange_vec_dec v__396 26 26 in
let size := size_bits_backwards _mappingpatterns_13_ in
let rl := bool_bits_backwards _mappingpatterns_12_ in
let aq := bool_bits_backwards _mappingpatterns_11_ in
true
- else if sumbool_of_bool ((let p0_ := subrange_vec_dec v__230 31 27 in
- let _mappingpatterns_14_ := subrange_vec_dec v__230 26 26 in
- let _mappingpatterns_15_ := subrange_vec_dec v__230 25 25 in
- let p1_ := subrange_vec_dec v__230 14 14 in
- let _mappingpatterns_16_ := subrange_vec_dec v__230 13 12 in
- let p2_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb
- (andb
- (andb (size_bits_backwards_matches _mappingpatterns_16_)
- (bool_bits_backwards_matches _mappingpatterns_15_))
- (bool_bits_backwards_matches _mappingpatterns_14_))
- (eq_vec p2_ (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))
- (eq_vec p1_ (vec_of_bits [B0] : mword 1)))
- (Z.eqb (projT1 ((regbits_to_regno p0_) : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
- (projT1 ((regbits_to_regno (vec_of_bits [B0;B0;B0;B1;B1] : mword 5))
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))))) then
- let _mappingpatterns_14_ := subrange_vec_dec v__230 26 26 in
- let _mappingpatterns_15_ := subrange_vec_dec v__230 25 25 in
- let _mappingpatterns_16_ := subrange_vec_dec v__230 13 12 in
+ else if sumbool_of_bool ((andb
+ (let _mappingpatterns_16_ := subrange_vec_dec v__396 13 12 in
+ let _mappingpatterns_15_ := subrange_vec_dec v__396 25 25 in
+ let _mappingpatterns_14_ := subrange_vec_dec v__396 26 26 in
+ andb (size_bits_backwards_matches _mappingpatterns_16_)
+ (if ((size_bits_backwards_matches _mappingpatterns_16_)) then
+ let size := size_bits_backwards _mappingpatterns_16_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_15_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_15_)) then
+ let rl := bool_bits_backwards _mappingpatterns_15_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_14_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_14_))
+ then
+ let aq := bool_bits_backwards _mappingpatterns_14_ in
+ true
+ else false)
+ else false)
+ else false))
+ (andb
+ (Z.eqb
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (subrange_vec_dec v__396 31 27))))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)}))
+ (projT1 ((build_ex (projT1 (regbits_to_regno
+ (vec_of_bits [B0;B0;B0;B1;B1]
+ : mword 5))))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))
+ (andb
+ (eq_vec (subrange_vec_dec v__396 14 14)
+ (vec_of_bits [B0] : mword 1))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))))) then
+ let _mappingpatterns_16_ := subrange_vec_dec v__396 13 12 in
+ let _mappingpatterns_15_ := subrange_vec_dec v__396 25 25 in
+ let _mappingpatterns_14_ := subrange_vec_dec v__396 26 26 in
let size := size_bits_backwards _mappingpatterns_16_ in
let rl := bool_bits_backwards _mappingpatterns_15_ in
let aq := bool_bits_backwards _mappingpatterns_14_ in
true
- else if ((let _mappingpatterns_17_ := subrange_vec_dec v__230 31 27 in
- let _mappingpatterns_18_ := subrange_vec_dec v__230 26 26 in
- let _mappingpatterns_19_ := subrange_vec_dec v__230 25 25 in
- let p0_ := subrange_vec_dec v__230 14 14 in
- let _mappingpatterns_20_ := subrange_vec_dec v__230 13 12 in
- let p1_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb
- (andb
- (andb
- (andb (size_bits_backwards_matches _mappingpatterns_20_)
- (bool_bits_backwards_matches _mappingpatterns_19_))
- (bool_bits_backwards_matches _mappingpatterns_18_))
- (encdec_amoop_backwards_matches _mappingpatterns_17_))
- (eq_vec p1_ (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))
- (eq_vec p0_ (vec_of_bits [B0] : mword 1)))) then
- let _mappingpatterns_17_ := subrange_vec_dec v__230 31 27 in
- let _mappingpatterns_18_ := subrange_vec_dec v__230 26 26 in
- let _mappingpatterns_19_ := subrange_vec_dec v__230 25 25 in
- let _mappingpatterns_20_ := subrange_vec_dec v__230 13 12 in
+ else if ((andb
+ (let _mappingpatterns_17_ := subrange_vec_dec v__396 31 27 in
+ let _mappingpatterns_20_ := subrange_vec_dec v__396 13 12 in
+ let _mappingpatterns_19_ := subrange_vec_dec v__396 25 25 in
+ let _mappingpatterns_18_ := subrange_vec_dec v__396 26 26 in
+ let _mappingpatterns_17_ := subrange_vec_dec v__396 31 27 in
+ andb (size_bits_backwards_matches _mappingpatterns_20_)
+ (if ((size_bits_backwards_matches _mappingpatterns_20_)) then
+ let size := size_bits_backwards _mappingpatterns_20_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_19_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_19_)) then
+ let rl := bool_bits_backwards _mappingpatterns_19_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_18_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_18_)) then
+ let aq := bool_bits_backwards _mappingpatterns_18_ in
+ andb (encdec_amoop_backwards_matches _mappingpatterns_17_)
+ (if ((encdec_amoop_backwards_matches _mappingpatterns_17_)) then
+ let op := encdec_amoop_backwards _mappingpatterns_17_ in
+ true
+ else false)
+ else false)
+ else false)
+ else false))
+ (andb (eq_vec (subrange_vec_dec v__396 14 14) (vec_of_bits [B0] : mword 1))
+ (eq_vec (subrange_vec_dec v__396 6 0)
+ (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7))))) then
+ let _mappingpatterns_17_ := subrange_vec_dec v__396 31 27 in
+ let _mappingpatterns_20_ := subrange_vec_dec v__396 13 12 in
+ let _mappingpatterns_19_ := subrange_vec_dec v__396 25 25 in
+ let _mappingpatterns_18_ := subrange_vec_dec v__396 26 26 in
+ let _mappingpatterns_17_ := subrange_vec_dec v__396 31 27 in
let size := size_bits_backwards _mappingpatterns_20_ in
let rl := bool_bits_backwards _mappingpatterns_19_ in
let aq := bool_bits_backwards _mappingpatterns_18_ in
let op := encdec_amoop_backwards _mappingpatterns_17_ in
true
- else if ((let _mappingpatterns_21_ := subrange_vec_dec v__230 14 14 in
- let _mappingpatterns_22_ := subrange_vec_dec v__230 13 12 in
- let p0_ := subrange_vec_dec v__230 6 0 in
- andb
- (andb (encdec_csrop_backwards_matches _mappingpatterns_22_)
- (bool_bits_backwards_matches _mappingpatterns_21_))
- (eq_vec p0_ (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7)))) then
- let _mappingpatterns_21_ := subrange_vec_dec v__230 14 14 in
- let _mappingpatterns_22_ := subrange_vec_dec v__230 13 12 in
+ else if ((andb
+ (let _mappingpatterns_22_ := subrange_vec_dec v__396 13 12 in
+ let _mappingpatterns_21_ := subrange_vec_dec v__396 14 14 in
+ andb (encdec_csrop_backwards_matches _mappingpatterns_22_)
+ (if ((encdec_csrop_backwards_matches _mappingpatterns_22_)) then
+ let op := encdec_csrop_backwards _mappingpatterns_22_ in
+ andb (bool_bits_backwards_matches _mappingpatterns_21_)
+ (if ((bool_bits_backwards_matches _mappingpatterns_21_)) then
+ let is_imm := bool_bits_backwards _mappingpatterns_21_ in
+ true
+ else false)
+ else false))
+ (eq_vec (subrange_vec_dec v__396 6 0) (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))))
+ then
+ let _mappingpatterns_22_ := subrange_vec_dec v__396 13 12 in
+ let _mappingpatterns_21_ := subrange_vec_dec v__396 14 14 in
let op := encdec_csrop_backwards _mappingpatterns_22_ in
let is_imm := bool_bits_backwards _mappingpatterns_21_ in
true
- else if ((let p0_ := subrange_vec_dec v__230 31 16 in
- let p1_ := subrange_vec_dec v__230 15 8 in
- let p2_ := subrange_vec_dec v__230 7 7 in
- let p3_ := subrange_vec_dec v__230 6 5 in
- let p4_ := subrange_vec_dec v__230 4 2 in
- let p5_ := subrange_vec_dec v__230 1 0 in
- andb
- (andb
- (andb
- (andb
- (andb (eq_vec p5_ (vec_of_bits [B1;B1] : mword 2))
- (eq_vec p4_ (vec_of_bits [B0;B1;B0] : mword 3)))
- (eq_vec p3_ (vec_of_bits [B0;B0] : mword 2)))
- (eq_vec p2_ (vec_of_bits [B0] : mword 1)))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword 8)))
- (eq_vec p0_ (vec_of_bits [B1;B1;B1;B1;B1;B0;B1;B0;B1;B1;B0;B1;B1;B1;B1;B0] : mword 16))))
- then
+ else if ((eq_vec 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]
+ : mword 32))) then
true
- else if ((let p0_ := subrange_vec_dec v__230 31 16 in
- let p1_ := subrange_vec_dec v__230 15 8 in
- let p2_ := subrange_vec_dec v__230 7 7 in
- let p3_ := subrange_vec_dec v__230 6 5 in
- let p4_ := subrange_vec_dec v__230 4 2 in
- let p5_ := subrange_vec_dec v__230 1 0 in
- andb
- (andb
- (andb
- (andb
- (andb (eq_vec p5_ (vec_of_bits [B1;B1] : mword 2))
- (eq_vec p4_ (vec_of_bits [B0;B1;B0] : mword 3)))
- (eq_vec p3_ (vec_of_bits [B0;B0] : mword 2)))
- (eq_vec p2_ (vec_of_bits [B0] : mword 1)))
- (eq_vec p1_ (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword 8)))
- (eq_vec p0_ (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B1;B1;B0;B1;B1;B1;B1;B0] : mword 16))))
- then
+ else if ((eq_vec 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]
+ : mword 32))) then
true
else true.
Definition print_insn (merge_var : ast)
: M (string) :=
match merge_var with
- | NOP (g__13) => returnm ("nop" : string)
+ | NOP (tt) => returnm ("nop" : string)
| C_ADDI4SPN (rdc,nzimm) =>
returnm ((String.append "c.addi4spn "
(String.append ((reg_name_abi (creg2reg_bits rdc)) : string)
@@ -31735,10 +28966,10 @@ Definition print_insn (merge_var : ast)
(String.append ((reg_name_abi rsd) : string)
(String.append ", " ((reg_name_abi rs2) : string))))
: string)
- | STOP_FETCHING (g__14) => returnm ("stop_fetching" : string)
- | THREAD_START (g__15) => returnm ("thread_start" : string)
+ | STOP_FETCHING (tt) => returnm ("stop_fetching" : string)
+ | THREAD_START (tt) => returnm ("thread_start" : string)
| ILLEGAL (s) => returnm ((String.append "illegal " (string_of_bits s)) : string)
- | C_ILLEGAL (g__16) => returnm ("c.illegal" : string)
+ | C_ILLEGAL (tt) => returnm ("c.illegal" : string)
| insn => (assembly_forwards insn) : M (string)
end.
@@ -31762,36 +28993,39 @@ Definition fetch '(tt : unit)
: bool)) (haveRVC tt >>= fun w__4 : bool => returnm ((negb w__4) : bool)))
: M (bool)) >>= fun w__6 : bool =>
(if (w__6) then
- (read_reg PC_ref : M (mword 64)) >>= fun w__7 : mword 64 =>
- returnm ((F_Error (E_Fetch_Addr_Align,w__7))
- : FetchResult)
+ (read_reg PC_ref : M (mword 64)) >>= fun w__7 : xlenbits =>
+ returnm ((F_Error
+ (E_Fetch_Addr_Align, w__7))
+ : FetchResult )
else
- (read_reg PC_ref : M (mword 64)) >>= fun w__8 : mword 64 =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__8 : xlenbits =>
translateAddr w__8 Execute Instruction >>= fun w__9 : TR_Result =>
(match w__9 with
| TR_Failure (e) =>
- (read_reg PC_ref : M (mword 64)) >>= fun w__10 : mword 64 =>
- returnm ((F_Error (e,w__10))
- : FetchResult)
+ (read_reg PC_ref : M (mword 64)) >>= fun w__10 : xlenbits =>
+ returnm ((F_Error
+ (e, w__10))
+ : FetchResult )
| TR_Address (ppclo) =>
- checked_mem_read Instruction ppclo 2 >>= fun w__11 : MemoryOpResult (mword 16) =>
+ checked_mem_read Instruction ppclo 2 >>= fun w__11 : MemoryOpResult (mword (8 * 2)) =>
(match w__11 with
| MemException (e) =>
- (read_reg PC_ref : M (mword 64)) >>= fun w__12 : mword 64 =>
- returnm ((F_Error (E_Fetch_Access_Fault,w__12))
- : FetchResult)
+ (read_reg PC_ref : M (mword 64)) >>= fun w__12 : xlenbits =>
+ returnm ((F_Error
+ (E_Fetch_Access_Fault, w__12))
+ : FetchResult )
| MemValue (ilo) =>
- (if ((isRVC ilo)) then returnm ((F_RVC ilo) : FetchResult)
+ (if ((isRVC ilo)) then returnm ((F_RVC (ilo)) : FetchResult )
else
- (read_reg PC_ref : M (mword 64)) >>= fun w__13 : mword 64 =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__13 : xlenbits =>
let PChi : xlenbits := add_vec_int w__13 2 in
translateAddr PChi Execute Instruction >>= fun w__14 : TR_Result =>
(match w__14 with
- | TR_Failure (e) => returnm ((F_Error (e,PChi)) : FetchResult)
+ | TR_Failure (e) => returnm ((F_Error (e, PChi)) : FetchResult )
| TR_Address (ppchi) =>
- checked_mem_read Instruction ppchi 2 >>= fun w__15 : MemoryOpResult (mword 16) =>
+ checked_mem_read Instruction ppchi 2 >>= fun w__15 : MemoryOpResult (mword (8 * 2)) =>
returnm ((match w__15 with
- | MemException (e) => F_Error (E_Fetch_Access_Fault,PChi)
+ | MemException (e) => F_Error (E_Fetch_Access_Fault, PChi)
| MemValue (ihi) => F_Base (concat_vec ihi ilo)
end)
: FetchResult)
@@ -31816,12 +29050,11 @@ Definition step (step_no : Z)
let '_ :=
(print_bits "Handling interrupt: " ((interruptType_to_bits intr) : mword 4))
: unit in
- handle_interrupt intr priv >> returnm ((false, false) : (bool * bool))
+ handle_interrupt intr priv >> returnm (false, false)
| None =>
fetch tt >>= fun w__5 : FetchResult =>
(match w__5 with
- | F_Error (e,addr) =>
- handle_mem_exception addr e >> returnm ((false, false) : (bool * bool))
+ | F_Error (e,addr) => handle_mem_exception addr e >> returnm (false, false)
| F_RVC (h) =>
(match (decodeCompressed h) with
| None =>
@@ -31838,7 +29071,7 @@ Definition step (step_no : Z)
(String.append " ("
(String.append (string_of_bits h) ") <no-decode>")))))))))
: unit in
- handle_decode_exception (EXTZ 64 h) >> returnm ((false, true) : (bool * bool))
+ handle_decode_exception (EXTZ 64 h) >> returnm (false, true)
| Some (ast) =>
read_reg cur_privilege_ref >>= fun w__8 : Privilege =>
(read_reg PC_ref : M (mword 64)) >>= fun w__9 : xlenbits =>
@@ -31855,9 +29088,9 @@ Definition step (step_no : Z)
(String.append (string_of_bits h)
(String.append ") " (w__10 : string)))))))))))
: unit in
- (read_reg PC_ref : M (mword 64)) >>= fun w__11 : mword 64 =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__11 : xlenbits =>
write_reg nextPC_ref (add_vec_int w__11 2) >>
- execute ast >>= fun w__12 : bool => returnm ((w__12, true) : (bool * bool))
+ execute ast >>= fun w__12 : bool => returnm (w__12, true)
end)
: M ((bool * bool))
| F_Base (w) =>
@@ -31876,7 +29109,7 @@ Definition step (step_no : Z)
(String.append " ("
(String.append (string_of_bits w) ") <no-decode>")))))))))
: unit in
- handle_decode_exception (EXTZ 64 w) >> returnm ((false, true) : (bool * bool))
+ handle_decode_exception (EXTZ 64 w) >> returnm (false, true)
| Some (ast) =>
read_reg cur_privilege_ref >>= fun w__16 : Privilege =>
(read_reg PC_ref : M (mword 64)) >>= fun w__17 : xlenbits =>
@@ -31893,9 +29126,9 @@ Definition step (step_no : Z)
(String.append (string_of_bits w)
(String.append ") " (w__18 : string)))))))))))
: unit in
- (read_reg PC_ref : M (mword 64)) >>= fun w__19 : mword 64 =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__19 : xlenbits =>
write_reg nextPC_ref (add_vec_int w__19 4) >>
- execute ast >>= fun w__20 : bool => returnm ((w__20, true) : (bool * bool))
+ execute ast >>= fun w__20 : bool => returnm (w__20, true)
end)
: M ((bool * bool))
end)
@@ -31920,23 +29153,31 @@ Definition loop '(tt : unit)
else returnm (tt : unit)) >>
let step_no : Z :=
if (stepped) then
- projT1 (add_range ((ex_int step_no) : {n : Z & ArithFact (True)}) (build_ex 1))
+ projT1 (add_range
+ (build_ex (projT1 ((build_ex (projT1 (ex_int step_no)))
+ : {n : Z & ArithFact (True)}))) (build_ex 1))
else step_no in
read_reg htif_done_ref >>= fun w__2 : bool =>
(if (w__2) then
(read_reg htif_exit_code_ref : M (mword 64)) >>= fun w__3 : xlenbits =>
- let '(existT _ exit_val _) := uint w__3 in
+ let 'exit_val := projT1 (uint w__3) in
returnm (let '_ :=
(if sumbool_of_bool ((Z.eqb exit_val 0)) then print_endline "SUCCESS"
else print_int "FAILURE: " exit_val)
: unit in
i)
else
- let i := (projT1 (add_range ((ex_int i) : {n : Z & ArithFact (True)}) (build_ex 1))) : Z in
- (if sumbool_of_bool ((Z.eqb (projT1 ((ex_int i) : {n : Z & ArithFact (True)}))
- insns_per_tick)) then
+ let i :=
+ projT1 ((build_ex (projT1 (add_range
+ (build_ex (projT1 ((build_ex (projT1 (ex_int i)))
+ : {n : Z & ArithFact (True)}))) (build_ex 1))))
+ : {arg0_ex46277_n : Z & {rangevar : Z & ArithFact ((arg0_ex46277_n + 1) <= rangevar /\
+ rangevar <= (arg0_ex46277_n + 1))}}) in
+ (if sumbool_of_bool ((Z.eqb
+ (projT1 ((build_ex (projT1 (ex_int projT1i)))
+ : {n : Z & ArithFact (True)})) insns_per_tick)) then
tick_clock tt >> tick_platform tt >> returnm 0
- else returnm i)
+ else returnm projT1i)
: M (Z)) >>= fun i : Z =>
returnm (i, step_no))) >>= fun '(i, step_no) =>
returnm (tt
@@ -31960,20 +29201,20 @@ Definition read_kind_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 11)}
Definition num_of_read_kind (arg_ : read_kind)
: {e : Z & ArithFact (0 <= e /\ e <= 11)} :=
- match arg_ with
- | Read_plain => build_ex 0
- | Read_reserve => build_ex 1
- | Read_acquire => build_ex 2
- | Read_exclusive => build_ex 3
- | Read_exclusive_acquire => build_ex 4
- | Read_stream => build_ex 5
- | Read_RISCV_acquire => build_ex 6
- | Read_RISCV_strong_acquire => build_ex 7
- | Read_RISCV_reserved => build_ex 8
- | Read_RISCV_reserved_acquire => build_ex 9
- | Read_RISCV_reserved_strong_acquire => build_ex 10
- | Read_X86_locked => build_ex 11
- end.
+ build_ex(match arg_ with
+ | Read_plain => 0
+ | Read_reserve => 1
+ | Read_acquire => 2
+ | Read_exclusive => 3
+ | Read_exclusive_acquire => 4
+ | Read_stream => 5
+ | Read_RISCV_acquire => 6
+ | Read_RISCV_strong_acquire => 7
+ | Read_RISCV_reserved => 8
+ | Read_RISCV_reserved_acquire => 9
+ | Read_RISCV_reserved_strong_acquire => 10
+ | Read_X86_locked => 11
+ end).
Definition write_kind_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 10)}
: write_kind :=
@@ -31992,19 +29233,19 @@ Definition write_kind_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 10)}
Definition num_of_write_kind (arg_ : write_kind)
: {e : Z & ArithFact (0 <= e /\ e <= 10)} :=
- match arg_ with
- | Write_plain => build_ex 0
- | Write_conditional => build_ex 1
- | Write_release => build_ex 2
- | Write_exclusive => build_ex 3
- | Write_exclusive_release => build_ex 4
- | Write_RISCV_release => build_ex 5
- | Write_RISCV_strong_release => build_ex 6
- | Write_RISCV_conditional => build_ex 7
- | Write_RISCV_conditional_release => build_ex 8
- | Write_RISCV_conditional_strong_release => build_ex 9
- | Write_X86_locked => build_ex 10
- end.
+ build_ex(match arg_ with
+ | Write_plain => 0
+ | Write_conditional => 1
+ | Write_release => 2
+ | Write_exclusive => 3
+ | Write_exclusive_release => 4
+ | Write_RISCV_release => 5
+ | Write_RISCV_strong_release => 6
+ | Write_RISCV_conditional => 7
+ | Write_RISCV_conditional_release => 8
+ | Write_RISCV_conditional_strong_release => 9
+ | Write_X86_locked => 10
+ end).
Definition barrier_kind_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 22)}
: barrier_kind :=
@@ -32032,35 +29273,35 @@ Definition barrier_kind_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 22)}
else if sumbool_of_bool ((Z.eqb p0_ 20)) then Barrier_RISCV_w_r
else if sumbool_of_bool ((Z.eqb p0_ 21)) then Barrier_RISCV_i
else Barrier_x86_MFENCE.
-(*
+
Definition num_of_barrier_kind (arg_ : barrier_kind)
: {e : Z & ArithFact (0 <= e /\ e <= 22)} :=
- match arg_ with
- | Barrier_Sync => build_ex 0
- | Barrier_LwSync => build_ex 1
- | Barrier_Eieio => build_ex 2
- | Barrier_Isync => build_ex 3
- | Barrier_DMB => build_ex 4
- | Barrier_DMB_ST => build_ex 5
- | Barrier_DMB_LD => build_ex 6
- | Barrier_DSB => build_ex 7
- | Barrier_DSB_ST => build_ex 8
- | Barrier_DSB_LD => build_ex 9
- | Barrier_ISB => build_ex 10
- | Barrier_MIPS_SYNC => build_ex 11
- | Barrier_RISCV_rw_rw => build_ex 12
- | Barrier_RISCV_r_rw => build_ex 13
- | Barrier_RISCV_r_r => build_ex 14
- | Barrier_RISCV_rw_w => build_ex 15
- | Barrier_RISCV_w_w => build_ex 16
- | Barrier_RISCV_w_rw => build_ex 17
- | Barrier_RISCV_rw_r => build_ex 18
- | Barrier_RISCV_r_w => build_ex 19
- | Barrier_RISCV_w_r => build_ex 20
- | Barrier_RISCV_i => build_ex 21
- | Barrier_x86_MFENCE => build_ex 22
- end.
-*)
+ build_ex(match arg_ with
+ | Barrier_Sync => 0
+ | Barrier_LwSync => 1
+ | Barrier_Eieio => 2
+ | Barrier_Isync => 3
+ | Barrier_DMB => 4
+ | Barrier_DMB_ST => 5
+ | Barrier_DMB_LD => 6
+ | Barrier_DSB => 7
+ | Barrier_DSB_ST => 8
+ | Barrier_DSB_LD => 9
+ | Barrier_ISB => 10
+ | Barrier_MIPS_SYNC => 11
+ | Barrier_RISCV_rw_rw => 12
+ | Barrier_RISCV_r_rw => 13
+ | Barrier_RISCV_r_r => 14
+ | Barrier_RISCV_rw_w => 15
+ | Barrier_RISCV_w_w => 16
+ | Barrier_RISCV_w_rw => 17
+ | Barrier_RISCV_rw_r => 18
+ | Barrier_RISCV_r_w => 19
+ | Barrier_RISCV_w_r => 20
+ | Barrier_RISCV_i => 21
+ | Barrier_x86_MFENCE => 22
+ end).
+
Definition trans_kind_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
: trans_kind :=
let p0_ := arg_ in
@@ -32070,230 +29311,251 @@ Definition trans_kind_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
Definition num_of_trans_kind (arg_ : trans_kind)
: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
- match arg_ with
- | Transaction_start => build_ex 0
- | Transaction_commit => build_ex 1
- | Transaction_abort => build_ex 2
- end.
+ build_ex(match arg_ with
+ | Transaction_start => 0
+ | Transaction_commit => 1
+ | Transaction_abort => 2
+ end).
-Let GPRstr : vec string 32 :=
+Definition GPRstr : vec string 32 :=
vec_of_list_len ["x31";"x30";"x29";"x28";"x27";"x26";"x25";"x24";"x23";"x22";"x21";"x20";"x19";"x18";"x17";"x16";"x15";"x14";"x13";"x12";"x11";
"x10";"x9";"x8";"x7";"x6";"x5";"x4";"x3";"x2";"x1";"x0"].
-Let CIA_fp := RFull "CIA".
-Let NIA_fp := RFull "NIA".
-(*
+Hint Unfold GPRstr : sail.
+Definition CIA_fp := RFull ("CIA").
+Hint Unfold CIA_fp : sail.
+Definition NIA_fp := RFull ("NIA").
+Hint Unfold NIA_fp : sail.
Definition initial_analysis (instr : ast)
: M ((list regfp * list regfp * list regfp * list niafp * diafp * instruction_kind)) :=
let iR := [] : regfps in
let oR := [] : regfps in
let aR := [] : regfps in
- let ik := (IK_simple tt) : instruction_kind in
- let Nias := [NIAFP_successor tt] : niafps in
- let Dia := (DIAFP_none tt) : diafp in
+ let ik := (IK_simple (tt)) : instruction_kind in
+ let Nias := [NIAFP_successor (tt)] : niafps in
+ let Dia := (DIAFP_none (tt)) : diafp in
match instr with
- | EBREAK (_) => returnm (Nias, aR, iR, ik, oR)
+ | EBREAK (tt) => returnm (Nias, aR, iR, ik, oR)
| UTYPE (imm,rd,op) =>
let oR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
oR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
returnm (Nias, aR, iR, ik, oR)
| RISCV_JAL (imm,rd) =>
let oR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
oR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
let offset : bits 64 := EXTS 64 imm in
- (read_reg PC_ref : M (mword 64)) >>= fun w__0 : mword 64 =>
+ (read_reg PC_ref : M (mword 64)) >>= fun w__0 : xlenbits =>
let Nias : niafps := [NIAFP_concrete_address (add_vec w__0 offset)] in
- let ik : instruction_kind := IK_branch tt in
+ let ik : instruction_kind := IK_branch (tt) in
returnm (Nias, aR, iR, ik, oR)
| RISCV_JALR (imm,rs,rd) =>
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let oR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
oR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
let offset : bits 64 := EXTS 64 imm in
- let Nias : niafps := [NIAFP_indirect_address tt] in
- let ik : instruction_kind := IK_branch tt in
+ let Nias : niafps := [NIAFP_indirect_address (tt)] in
+ let ik : instruction_kind := IK_branch (tt) in
returnm (Nias, aR, iR, ik, oR)
| BTYPE (imm,rs2,rs1,op) =>
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs2)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs2)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs1)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
- let ik := (IK_branch tt) : instruction_kind in
+ let ik := (IK_branch (tt)) : instruction_kind in
let offset : bits 64 := EXTS 64 imm in
- (read_reg PC_ref : M (mword 64)) >>= fun w__1 : mword 64 =>
- let Nias : niafps := [NIAFP_concrete_address (add_vec w__1 offset);NIAFP_successor tt] in
+ (read_reg PC_ref : M (mword 64)) >>= fun w__1 : xlenbits =>
+ let Nias : niafps := [NIAFP_concrete_address (add_vec w__1 offset);NIAFP_successor (tt)] in
returnm (Nias, aR, iR, ik, oR)
| ITYPE (imm,rs,rd,op) =>
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let oR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
oR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
returnm (Nias, aR, iR, ik, oR)
| SHIFTIOP (imm,rs,rd,op) =>
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let oR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
oR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
returnm (Nias, aR, iR, ik, oR)
| RTYPE (rs2,rs1,rd,op) =>
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs2)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs2)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs1)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let oR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
oR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
returnm (Nias, aR, iR, ik, oR)
| CSR (csr,rs1,rd,is_imm,op) =>
let isWrite : bool :=
match op with
| CSRRW => true
- | _ => if (is_imm) then neq_int (projT1 (uint rs1)) 0 else neq_int (projT1 (uint rs1)) 0
+ | _ =>
+ if (is_imm) then neq_range (build_ex (projT1 (uint rs1))) (build_ex 0)
+ else neq_range (build_ex (projT1 (uint rs1))) (build_ex 0)
end in
let iR : regfps := (RFull (csr_name csr)) :: iR in
let iR : regfps :=
if ((negb is_imm)) then
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs1)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR
else iR in
let oR : regfps := if (isWrite) then (RFull (csr_name csr)) :: oR else oR in
let oR : regfps :=
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
returnm (Nias, aR, iR, ik, oR)
| LOAD (imm,rs,rd,unsign,width,aq,rl) =>
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let oR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
oR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
let aR := iR : regfps in
match (aq, rl) with
- | (false, false) => returnm ((IK_mem_read Read_plain) : instruction_kind)
- | (true, false) => returnm ((IK_mem_read Read_RISCV_acquire) : instruction_kind)
- | (true, true) => returnm ((IK_mem_read Read_RISCV_strong_acquire) : instruction_kind)
+ | (false, false) => returnm ((IK_mem_read (Read_plain)) : instruction_kind )
+ | (true, false) => returnm ((IK_mem_read (Read_RISCV_acquire)) : instruction_kind )
+ | (true, true) => returnm ((IK_mem_read (Read_RISCV_strong_acquire)) : instruction_kind )
| _ =>
(internal_error "LOAD type not implemented in initial_analysis") : M (instruction_kind)
end >>= fun w__3 : instruction_kind =>
@@ -32302,38 +29564,41 @@ Definition initial_analysis (instr : ast)
| STORE (imm,rs2,rs1,width,aq,rl) =>
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs2)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs2)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs1)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let aR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
aR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs1)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
aR in
match (aq, rl) with
- | (false, false) => returnm ((IK_mem_write Write_plain) : instruction_kind)
- | (false, true) => returnm ((IK_mem_write Write_RISCV_release) : instruction_kind)
- | (true, true) => returnm ((IK_mem_write Write_RISCV_strong_release) : instruction_kind)
+ | (false, false) => returnm ((IK_mem_write (Write_plain)) : instruction_kind )
+ | (false, true) => returnm ((IK_mem_write (Write_RISCV_release)) : instruction_kind )
+ | (true, true) => returnm ((IK_mem_write (Write_RISCV_strong_release)) : instruction_kind )
| _ =>
(internal_error "STORE type not implemented in initial_analysis") : M (instruction_kind)
end >>= fun w__5 : instruction_kind =>
@@ -32342,122 +29607,139 @@ Definition initial_analysis (instr : ast)
| ADDIW (imm,rs,rd) =>
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let oR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
oR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
returnm (Nias, aR, iR, ik, oR)
| SHIFTW (imm,rs,rd,op) =>
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let oR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
oR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
returnm (Nias, aR, iR, ik, oR)
| RTYPEW (rs2,rs1,rd,op) =>
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs2)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs2)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs1)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let oR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
oR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
returnm (Nias, aR, iR, ik, oR)
| FENCE (pred,succ) =>
match (pred, succ) with
- | (v__282, v__283) =>
- (if ((andb (eq_vec (subrange_vec_dec v__282 1 0) (vec_of_bits [B1;B1] : mword 2))
- (eq_vec (subrange_vec_dec v__283 1 0) (vec_of_bits [B1;B1] : mword 2)))) then
- returnm ((IK_barrier Barrier_RISCV_rw_rw)
- : instruction_kind)
- else if ((andb (eq_vec (subrange_vec_dec v__282 1 0) (vec_of_bits [B1;B0] : mword 2))
- (eq_vec (subrange_vec_dec v__283 1 0) (vec_of_bits [B1;B1] : mword 2)))) then
- returnm ((IK_barrier Barrier_RISCV_r_rw)
- : instruction_kind)
- else if ((andb (eq_vec (subrange_vec_dec v__282 1 0) (vec_of_bits [B1;B0] : mword 2))
- (eq_vec (subrange_vec_dec v__283 1 0) (vec_of_bits [B1;B0] : mword 2)))) then
- returnm ((IK_barrier Barrier_RISCV_r_r)
- : instruction_kind)
- else if ((andb (eq_vec (subrange_vec_dec v__282 1 0) (vec_of_bits [B1;B1] : mword 2))
- (eq_vec (subrange_vec_dec v__283 1 0) (vec_of_bits [B0;B1] : mword 2)))) then
- returnm ((IK_barrier Barrier_RISCV_rw_w)
- : instruction_kind)
- else if ((andb (eq_vec (subrange_vec_dec v__282 1 0) (vec_of_bits [B0;B1] : mword 2))
- (eq_vec (subrange_vec_dec v__283 1 0) (vec_of_bits [B0;B1] : mword 2)))) then
- returnm ((IK_barrier Barrier_RISCV_w_w)
- : instruction_kind)
- else if ((andb (eq_vec (subrange_vec_dec v__282 1 0) (vec_of_bits [B0;B1] : mword 2))
- (eq_vec (subrange_vec_dec v__283 1 0) (vec_of_bits [B1;B1] : mword 2)))) then
- returnm ((IK_barrier Barrier_RISCV_w_rw)
- : instruction_kind)
- else if ((andb (eq_vec (subrange_vec_dec v__282 1 0) (vec_of_bits [B1;B1] : mword 2))
- (eq_vec (subrange_vec_dec v__283 1 0) (vec_of_bits [B1;B0] : mword 2)))) then
- returnm ((IK_barrier Barrier_RISCV_rw_r)
- : instruction_kind)
- else if ((andb (eq_vec (subrange_vec_dec v__282 1 0) (vec_of_bits [B1;B0] : mword 2))
- (eq_vec (subrange_vec_dec v__283 1 0) (vec_of_bits [B0;B1] : mword 2)))) then
- returnm ((IK_barrier Barrier_RISCV_r_w)
- : instruction_kind)
- else if ((andb (eq_vec (subrange_vec_dec v__282 1 0) (vec_of_bits [B0;B1] : mword 2))
- (eq_vec (subrange_vec_dec v__283 1 0) (vec_of_bits [B1;B0] : mword 2)))) then
- returnm ((IK_barrier Barrier_RISCV_w_r)
- : instruction_kind)
- else if ((andb (eq_vec (subrange_vec_dec v__282 1 0) (vec_of_bits [B0;B0] : mword 2))
- (eq_vec (subrange_vec_dec v__283 1 0) (vec_of_bits [B0;B0] : mword 2)))) then
- returnm ((IK_simple tt)
- : instruction_kind)
+ | (v__610, v__611) =>
+ (if ((andb (eq_vec (subrange_vec_dec v__610 1 0) (vec_of_bits [B1;B1] : mword 2))
+ (eq_vec (subrange_vec_dec v__611 1 0) (vec_of_bits [B1;B1] : mword 2)))) then
+ returnm ((IK_barrier
+ (Barrier_RISCV_rw_rw))
+ : instruction_kind )
+ else if ((andb (eq_vec (subrange_vec_dec v__610 1 0) (vec_of_bits [B1;B0] : mword 2))
+ (eq_vec (subrange_vec_dec v__611 1 0) (vec_of_bits [B1;B1] : mword 2)))) then
+ returnm ((IK_barrier
+ (Barrier_RISCV_r_rw))
+ : instruction_kind )
+ else if ((andb (eq_vec (subrange_vec_dec v__610 1 0) (vec_of_bits [B1;B0] : mword 2))
+ (eq_vec (subrange_vec_dec v__611 1 0) (vec_of_bits [B1;B0] : mword 2)))) then
+ returnm ((IK_barrier
+ (Barrier_RISCV_r_r))
+ : instruction_kind )
+ else if ((andb (eq_vec (subrange_vec_dec v__610 1 0) (vec_of_bits [B1;B1] : mword 2))
+ (eq_vec (subrange_vec_dec v__611 1 0) (vec_of_bits [B0;B1] : mword 2)))) then
+ returnm ((IK_barrier
+ (Barrier_RISCV_rw_w))
+ : instruction_kind )
+ else if ((andb (eq_vec (subrange_vec_dec v__610 1 0) (vec_of_bits [B0;B1] : mword 2))
+ (eq_vec (subrange_vec_dec v__611 1 0) (vec_of_bits [B0;B1] : mword 2)))) then
+ returnm ((IK_barrier
+ (Barrier_RISCV_w_w))
+ : instruction_kind )
+ else if ((andb (eq_vec (subrange_vec_dec v__610 1 0) (vec_of_bits [B0;B1] : mword 2))
+ (eq_vec (subrange_vec_dec v__611 1 0) (vec_of_bits [B1;B1] : mword 2)))) then
+ returnm ((IK_barrier
+ (Barrier_RISCV_w_rw))
+ : instruction_kind )
+ else if ((andb (eq_vec (subrange_vec_dec v__610 1 0) (vec_of_bits [B1;B1] : mword 2))
+ (eq_vec (subrange_vec_dec v__611 1 0) (vec_of_bits [B1;B0] : mword 2)))) then
+ returnm ((IK_barrier
+ (Barrier_RISCV_rw_r))
+ : instruction_kind )
+ else if ((andb (eq_vec (subrange_vec_dec v__610 1 0) (vec_of_bits [B1;B0] : mword 2))
+ (eq_vec (subrange_vec_dec v__611 1 0) (vec_of_bits [B0;B1] : mword 2)))) then
+ returnm ((IK_barrier
+ (Barrier_RISCV_r_w))
+ : instruction_kind )
+ else if ((andb (eq_vec (subrange_vec_dec v__610 1 0) (vec_of_bits [B0;B1] : mword 2))
+ (eq_vec (subrange_vec_dec v__611 1 0) (vec_of_bits [B1;B0] : mword 2)))) then
+ returnm ((IK_barrier
+ (Barrier_RISCV_w_r))
+ : instruction_kind )
+ else if ((andb (eq_vec (subrange_vec_dec v__610 1 0) (vec_of_bits [B0;B0] : mword 2))
+ (eq_vec (subrange_vec_dec v__611 1 0) (vec_of_bits [B0;B0] : mword 2)))) then
+ returnm ((IK_simple
+ (tt))
+ : instruction_kind )
else
(internal_error "barrier type not implemented in initial_analysis")
: M (instruction_kind))
@@ -32465,36 +29747,38 @@ Definition initial_analysis (instr : ast)
end >>= fun w__17 : instruction_kind =>
let ik : instruction_kind := w__17 in
returnm (Nias, aR, iR, ik, oR)
- | FENCEI (_) =>
- let ik : instruction_kind := IK_simple tt in
+ | FENCEI (tt) =>
+ let ik : instruction_kind := IK_simple (tt) in
returnm (Nias, aR, iR, ik, oR)
| LOADRES (aq,rl,rs1,width,rd) =>
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs1)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let oR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
oR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
let aR := iR : regfps in
match (aq, rl) with
- | (false, false) => returnm ((IK_mem_read Read_RISCV_reserved) : instruction_kind)
- | (true, false) => returnm ((IK_mem_read Read_RISCV_reserved_acquire) : instruction_kind)
+ | (false, false) => returnm ((IK_mem_read (Read_RISCV_reserved)) : instruction_kind )
+ | (true, false) => returnm ((IK_mem_read (Read_RISCV_reserved_acquire)) : instruction_kind )
| (true, true) =>
- returnm ((IK_mem_read Read_RISCV_reserved_strong_acquire) : instruction_kind)
+ returnm ((IK_mem_read (Read_RISCV_reserved_strong_acquire)) : instruction_kind )
| (false, true) =>
(internal_error "LOADRES type not implemented in initial_analysis") : M (instruction_kind)
end >>= fun w__19 : instruction_kind =>
@@ -32503,50 +29787,54 @@ Definition initial_analysis (instr : ast)
| STORECON (aq,rl,rs2,rs1,width,rd) =>
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs2)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs2)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs1)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let aR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
aR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs1)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
aR in
let oR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
oR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
match (aq, rl) with
- | (false, false) => returnm ((IK_mem_write Write_RISCV_conditional) : instruction_kind)
+ | (false, false) => returnm ((IK_mem_write (Write_RISCV_conditional)) : instruction_kind )
| (false, true) =>
- returnm ((IK_mem_write Write_RISCV_conditional_release) : instruction_kind)
+ returnm ((IK_mem_write (Write_RISCV_conditional_release)) : instruction_kind )
| (true, true) =>
- returnm ((IK_mem_write Write_RISCV_conditional_strong_release) : instruction_kind)
+ returnm ((IK_mem_write (Write_RISCV_conditional_strong_release)) : instruction_kind )
| (true, false) =>
(internal_error "STORECON type not implemented in initial_analysis")
: M (instruction_kind)
@@ -32556,58 +29844,61 @@ Definition initial_analysis (instr : ast)
| AMO (op,aq,rl,rs2,rs1,width,rd) =>
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs2)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs2)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs2)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let iR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
iR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs1)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
iR in
let aR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rs1)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
aR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rs1)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rs1)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
aR in
let oR : regfps :=
if sumbool_of_bool ((Z.eqb
- (projT1 ((regbits_to_regno rd)
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
: {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})) 0)) then
oR
else
- (RFull (vec_access_dec GPRstr
- (projT1 ((regbits_to_regno rd)
- : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
+ (RFull
+ (vec_access_dec GPRstr
+ (projT1 ((build_ex (projT1 (regbits_to_regno rd)))
+ : {n : Z & ArithFact (0 <= n /\ (n + 1) <= 32)})))) ::
oR in
let ik : instruction_kind :=
match (aq, rl) with
- | (false, false) => IK_mem_rmw (Read_RISCV_reserved,Write_RISCV_conditional)
- | (false, true) => IK_mem_rmw (Read_RISCV_reserved,Write_RISCV_conditional_release)
- | (true, false) => IK_mem_rmw (Read_RISCV_reserved_acquire,Write_RISCV_conditional)
- | (true, true) => IK_mem_rmw (Read_RISCV_reserved_acquire,Write_RISCV_conditional_release)
+ | (false, false) => IK_mem_rmw (Read_RISCV_reserved, Write_RISCV_conditional)
+ | (false, true) => IK_mem_rmw (Read_RISCV_reserved, Write_RISCV_conditional_release)
+ | (true, false) => IK_mem_rmw (Read_RISCV_reserved_acquire, Write_RISCV_conditional)
+ | (true, true) => IK_mem_rmw (Read_RISCV_reserved_acquire, Write_RISCV_conditional_release)
end in
returnm (Nias, aR, iR, ik, oR)
| _ => returnm (Nias, aR, iR, ik, oR)
end >>= fun '(Nias, aR, iR, ik, oR) =>
- returnm ((iR, oR, aR, Nias, Dia, ik)
- : (list regfp * list regfp * list regfp * list niafp * diafp * instruction_kind)).
-*)
-Let initial_regstate : regstate :=
+ returnm (iR, oR, aR, Nias, Dia, ik).
+
+Definition initial_regstate : regstate :=
{| tlb39 := None;
htif_exit_code :=
(vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
@@ -33161,5 +30452,6 @@ Let initial_regstate : regstate :=
B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
B0]
: mword 64) |}.
+Hint Unfold initial_regstate : sail.
End Content.
diff --git a/snapshots/coq-riscv/sail/riscv/riscv_extras.v b/snapshots/coq-riscv/sail/riscv/riscv_extras.v
index 3f1fe7e0..820e3f3a 100644
--- a/snapshots/coq-riscv/sail/riscv/riscv_extras.v
+++ b/snapshots/coq-riscv/sail/riscv/riscv_extras.v
@@ -9,60 +9,32 @@ Import List.ListNotations.
Axiom real : Type.
+Definition MEM_fence_rw_rw {rv e} (_:unit) : monad rv unit e := barrier Barrier_RISCV_rw_rw.
+Definition MEM_fence_r_rw {rv e} (_:unit) : monad rv unit e := barrier Barrier_RISCV_r_rw.
+Definition MEM_fence_r_r {rv e} (_:unit) : monad rv unit e := barrier Barrier_RISCV_r_r.
+Definition MEM_fence_rw_w {rv e} (_:unit) : monad rv unit e := barrier Barrier_RISCV_rw_w.
+Definition MEM_fence_w_w {rv e} (_:unit) : monad rv unit e := barrier Barrier_RISCV_w_w.
+Definition MEM_fence_w_rw {rv e} (_:unit) : monad rv unit e := barrier Barrier_RISCV_w_rw.
+Definition MEM_fence_rw_r {rv e} (_:unit) : monad rv unit e := barrier Barrier_RISCV_rw_r.
+Definition MEM_fence_r_w {rv e} (_:unit) : monad rv unit e := barrier Barrier_RISCV_r_w.
+Definition MEM_fence_w_r {rv e} (_:unit) : monad rv unit e := barrier Barrier_RISCV_w_r.
+Definition MEM_fence_i {rv e} (_:unit) : monad rv unit e := barrier Barrier_RISCV_i.
(*
-val MEMr : forall 'regval 'a 'b 'e. Bitvector 'a, Bitvector 'b => 'a -> integer -> monad 'regval 'b 'e
-val MEMr_reserve : forall 'regval 'a 'b 'e. Bitvector 'a, Bitvector 'b => 'a -> integer -> monad 'regval 'b 'e
-val MEMr_tag : forall 'regval 'a 'b 'e. Bitvector 'a, Bitvector 'b => 'a -> integer -> monad 'regval (bool * 'b) 'e
-val MEMr_tag_reserve : forall 'regval 'a 'b 'e. Bitvector 'a, Bitvector 'b => 'a -> integer -> monad 'regval (bool * 'b) 'e
+val MEMea : forall 'rv 'a 'e. Size 'a => bitvector 'a -> integer -> monad 'rv unit 'e
+val MEMea_release : forall 'rv 'a 'e. Size 'a => bitvector 'a -> integer -> monad 'rv unit 'e
+val MEMea_strong_release : forall 'rv 'a 'e. Size 'a => bitvector 'a -> integer -> monad 'rv unit 'e
+val MEMea_conditional : forall 'rv 'a 'e. Size 'a => bitvector 'a -> integer -> monad 'rv unit 'e
+val MEMea_conditional_release : forall 'rv 'a 'e. Size 'a => bitvector 'a -> integer -> monad 'rv unit 'e
+val MEMea_conditional_strong_release : forall 'rv 'a 'e. Size 'a => bitvector 'a -> integer -> monad 'rv unit 'e
*)
-Definition MEMr {regval a b e} `{ArithFact (b >= 0)} (addr : mword a) size : monad regval (mword b) e := read_mem Read_plain addr size.
-Definition MEMr_reserve {regval a b e} `{ArithFact (b >= 0)} (addr : mword a) size : monad regval (mword b) e := read_mem Read_reserve addr size.
+Definition MEMea {rv a e} (addr : mword a) size : monad rv unit e := write_mem_ea Write_plain addr size.
+Definition MEMea_release {rv a e} (addr : mword a) size : monad rv unit e := write_mem_ea Write_RISCV_release addr size.
+Definition MEMea_strong_release {rv a e} (addr : mword a) size : monad rv unit e := write_mem_ea Write_RISCV_strong_release addr size.
+Definition MEMea_conditional {rv a e} (addr : mword a) size : monad rv unit e := write_mem_ea Write_RISCV_conditional addr size.
+Definition MEMea_conditional_release {rv a e} (addr : mword a) size : monad rv unit e := write_mem_ea Write_RISCV_conditional_release addr size.
+Definition MEMea_conditional_strong_release {rv a e} (addr : mword a) size : monad rv unit e
+ := write_mem_ea Write_RISCV_conditional_strong_release addr size.
-(*val read_tag_bool : forall 'regval 'a 'e. Bitvector 'a => 'a -> monad 'regval bool 'e*)
-Definition read_tag_bool {regval a e} (addr : mword a) : monad regval bool e :=
- read_tag addr >>= fun t =>
- maybe_fail "read_tag_bool" (bool_of_bitU t).
-
-(*val write_tag_bool : forall 'regval 'a 'e. Bitvector 'a => 'a -> bool -> monad 'regval unit 'e*)
-Definition write_tag_bool {regval a e} (addr : mword a) t : monad regval unit e :=
- write_tag addr (bitU_of_bool t) >>= fun _ => returnm tt.
-
-Definition MEMr_tag {regval a b e} `{ArithFact (b >= 0)} (addr : mword a) size : monad regval (bool * mword b) e :=
- read_mem Read_plain addr size >>= fun v =>
- read_tag_bool addr >>= fun t =>
- returnm (t, v).
-
-Definition MEMr_tag_reserve {regval a b e} `{ArithFact (b >= 0)} (addr : mword a) size : monad regval (bool * mword b) e :=
- read_mem Read_plain addr size >>= fun v =>
- read_tag_bool addr >>= fun t =>
- returnm (t, v).
-
-(*
-val MEMea : forall 'regval 'a 'e. Bitvector 'a => 'a -> integer -> monad 'regval unit 'e
-val MEMea_conditional : forall 'regval 'a 'e. Bitvector 'a => 'a -> integer -> monad 'regval unit 'e
-val MEMea_tag : forall 'regval 'a 'e. Bitvector 'a => 'a -> integer -> monad 'regval unit 'e
-val MEMea_tag_conditional : forall 'regval 'a 'e. Bitvector 'a => 'a -> integer -> monad 'regval unit 'e
-*)
-Definition MEMea {regval a e} (addr : mword a) size : monad regval unit e := write_mem_ea Write_plain addr size.
-Definition MEMea_conditional {regval a e} (addr : mword a) size : monad regval unit e := write_mem_ea Write_conditional addr size.
-
-Definition MEMea_tag {regval a e} (addr : mword a) size : monad regval unit e := write_mem_ea Write_plain addr size.
-Definition MEMea_tag_conditional {regval a e} (addr : mword a) size : monad regval unit e := write_mem_ea Write_conditional addr size.
-
-(*
-val MEMval : forall 'regval 'a 'b 'e. Bitvector 'a, Bitvector 'b => 'a -> integer -> 'b -> monad 'regval unit 'e
-val MEMval_conditional : forall 'regval 'a 'b 'e. Bitvector 'a, Bitvector 'b => 'a -> integer -> 'b -> monad 'regval bool 'e
-val MEMval_tag : forall 'regval 'a 'b 'e. Bitvector 'a, Bitvector 'b => 'a -> integer -> bool -> 'b -> monad 'regval unit 'e
-val MEMval_tag_conditional : forall 'regval 'a 'b 'e. Bitvector 'a, Bitvector 'b => 'a -> integer -> bool -> 'b -> monad 'regval bool 'e
-*)
-Definition MEMval {regval a b e} (_ : mword a) (size : Z) (v : mword b) : monad regval unit e := write_mem_val v >>= fun _ => returnm tt.
-Definition MEMval_conditional {regval a b e} (_ : mword a) (size : Z) (v : mword b) : monad regval bool e := write_mem_val v >>= fun b => returnm (if b then true else false).
-Definition MEMval_tag {regval a b e} (addr : mword a) (size : Z) t (v : mword b) : monad regval unit e := write_mem_val v >>= fun _ => write_tag_bool addr t >>= fun _ => returnm tt.
-Definition MEMval_tag_conditional {regval a b e} (addr : mword a) (size : Z) t (v : mword b) : monad regval bool e := write_mem_val v >>= fun b => write_tag_bool addr t >>= fun _ => returnm (if b then true else false).
-
-(*val MEM_sync : forall 'regval 'e. unit -> monad 'regval unit 'e*)
-
-Definition MEM_sync {regval e} (_:unit) : monad regval unit e := barrier Barrier_MIPS_SYNC.
(* Some wrappers copied from aarch64_extras *)
(* TODO: Harmonise into a common library *)
diff --git a/snapshots/coq-riscv/sail/riscv/riscv_types.v b/snapshots/coq-riscv/sail/riscv/riscv_types.v
index 2bcebd62..6d3817c6 100644
--- a/snapshots/coq-riscv/sail/riscv/riscv_types.v
+++ b/snapshots/coq-riscv/sail/riscv/riscv_types.v
@@ -614,83 +614,84 @@ Notation "{[ r 'with' 'PC' := e ]}" := ({| PC := e; tlb39 := tlb39 r; htif_exit_
Definition Counteren_of_regval (merge_var : register_value)
: option Counteren :=
- match merge_var with | Regval_Counteren (v) => Some v | g__12 => None end.
+ match merge_var with | Regval_Counteren (v) => Some (v) | g__12 => None end.
-Definition regval_of_Counteren (v : Counteren) : register_value := Regval_Counteren v.
+Definition regval_of_Counteren (v : Counteren) : register_value := Regval_Counteren (v).
Definition Mcause_of_regval (merge_var : register_value)
: option Mcause :=
- match merge_var with | Regval_Mcause (v) => Some v | g__11 => None end.
+ match merge_var with | Regval_Mcause (v) => Some (v) | g__11 => None end.
-Definition regval_of_Mcause (v : Mcause) : register_value := Regval_Mcause v.
+Definition regval_of_Mcause (v : Mcause) : register_value := Regval_Mcause (v).
Definition Medeleg_of_regval (merge_var : register_value)
: option Medeleg :=
- match merge_var with | Regval_Medeleg (v) => Some v | g__10 => None end.
+ match merge_var with | Regval_Medeleg (v) => Some (v) | g__10 => None end.
-Definition regval_of_Medeleg (v : Medeleg) : register_value := Regval_Medeleg v.
+Definition regval_of_Medeleg (v : Medeleg) : register_value := Regval_Medeleg (v).
Definition Minterrupts_of_regval (merge_var : register_value)
: option Minterrupts :=
- match merge_var with | Regval_Minterrupts (v) => Some v | g__9 => None end.
+ match merge_var with | Regval_Minterrupts (v) => Some (v) | g__9 => None end.
-Definition regval_of_Minterrupts (v : Minterrupts) : register_value := Regval_Minterrupts v.
+Definition regval_of_Minterrupts (v : Minterrupts) : register_value := Regval_Minterrupts (v).
Definition Misa_of_regval (merge_var : register_value)
: option Misa :=
- match merge_var with | Regval_Misa (v) => Some v | g__8 => None end.
+ match merge_var with | Regval_Misa (v) => Some (v) | g__8 => None end.
-Definition regval_of_Misa (v : Misa) : register_value := Regval_Misa v.
+Definition regval_of_Misa (v : Misa) : register_value := Regval_Misa (v).
Definition Mstatus_of_regval (merge_var : register_value)
: option Mstatus :=
- match merge_var with | Regval_Mstatus (v) => Some v | g__7 => None end.
+ match merge_var with | Regval_Mstatus (v) => Some (v) | g__7 => None end.
-Definition regval_of_Mstatus (v : Mstatus) : register_value := Regval_Mstatus v.
+Definition regval_of_Mstatus (v : Mstatus) : register_value := Regval_Mstatus (v).
Definition Mtvec_of_regval (merge_var : register_value)
: option Mtvec :=
- match merge_var with | Regval_Mtvec (v) => Some v | g__6 => None end.
+ match merge_var with | Regval_Mtvec (v) => Some (v) | g__6 => None end.
-Definition regval_of_Mtvec (v : Mtvec) : register_value := Regval_Mtvec v.
+Definition regval_of_Mtvec (v : Mtvec) : register_value := Regval_Mtvec (v).
Definition Privilege_of_regval (merge_var : register_value)
: option Privilege :=
- match merge_var with | Regval_Privilege (v) => Some v | g__5 => None end.
+ match merge_var with | Regval_Privilege (v) => Some (v) | g__5 => None end.
-Definition regval_of_Privilege (v : Privilege) : register_value := Regval_Privilege v.
+Definition regval_of_Privilege (v : Privilege) : register_value := Regval_Privilege (v).
Definition Sedeleg_of_regval (merge_var : register_value)
: option Sedeleg :=
- match merge_var with | Regval_Sedeleg (v) => Some v | g__4 => None end.
+ match merge_var with | Regval_Sedeleg (v) => Some (v) | g__4 => None end.
-Definition regval_of_Sedeleg (v : Sedeleg) : register_value := Regval_Sedeleg v.
+Definition regval_of_Sedeleg (v : Sedeleg) : register_value := Regval_Sedeleg (v).
Definition Sinterrupts_of_regval (merge_var : register_value)
: option Sinterrupts :=
- match merge_var with | Regval_Sinterrupts (v) => Some v | g__3 => None end.
+ match merge_var with | Regval_Sinterrupts (v) => Some (v) | g__3 => None end.
-Definition regval_of_Sinterrupts (v : Sinterrupts) : register_value := Regval_Sinterrupts v.
+Definition regval_of_Sinterrupts (v : Sinterrupts) : register_value := Regval_Sinterrupts (v).
Definition TLB39_Entry_of_regval (merge_var : register_value)
: option TLB39_Entry :=
- match merge_var with | Regval_TLB39_Entry (v) => Some v | g__2 => None end.
+ match merge_var with | Regval_TLB39_Entry (v) => Some (v) | g__2 => None end.
-Definition regval_of_TLB39_Entry (v : TLB39_Entry) : register_value := Regval_TLB39_Entry v.
+Definition regval_of_TLB39_Entry (v : TLB39_Entry) : register_value := Regval_TLB39_Entry (v).
Definition bool_of_regval (merge_var : register_value)
: option bool :=
- match merge_var with | Regval_bool (v) => Some v | g__1 => None end.
+ match merge_var with | Regval_bool (v) => Some (v) | g__1 => None end.
-Definition regval_of_bool (v : bool) : register_value := Regval_bool v.
+Definition regval_of_bool (v : bool) : register_value := Regval_bool (v).
Definition vector_64_dec_bit_of_regval (merge_var : register_value)
: option (mword 64) :=
- match merge_var with | Regval_vector_64_dec_bit (v) => Some v | g__0 => None end.
+ match merge_var with | Regval_vector_64_dec_bit (v) => Some (v) | g__0 => None end.
Definition regval_of_vector_64_dec_bit (v : mword 64)
: register_value :=
- Regval_vector_64_dec_bit v.
+ Regval_vector_64_dec_bit
+ (v).