aboutsummaryrefslogtreecommitdiff
path: root/prover_snapshots/coq/lib
diff options
context:
space:
mode:
Diffstat (limited to 'prover_snapshots/coq/lib')
-rw-r--r--prover_snapshots/coq/lib/sail/Hoare.v810
-rw-r--r--prover_snapshots/coq/lib/sail/Makefile26
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_impl_base.v1103
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_instr_kinds.v332
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_operators.v232
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_operators_bitlists.v182
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_operators_mwords.v544
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_prompt.v229
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_prompt_monad.v367
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_real.v36
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_state.v167
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_state_lemmas.v819
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_state_lifting.v61
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_state_monad.v323
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_state_monad_lemmas.v542
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_string.v194
-rw-r--r--prover_snapshots/coq/lib/sail/Sail2_values.v2490
-rw-r--r--prover_snapshots/coq/lib/sail/_CoqProject2
18 files changed, 8459 insertions, 0 deletions
diff --git a/prover_snapshots/coq/lib/sail/Hoare.v b/prover_snapshots/coq/lib/sail/Hoare.v
new file mode 100644
index 0000000..d23ff32
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Hoare.v
@@ -0,0 +1,810 @@
+Require Import String ZArith.
+Require Import Sail2_state_monad Sail2_prompt Sail2_state Sail2_state_monad_lemmas.
+Require Import Sail2_state_lemmas.
+
+(*adhoc_overloading
+ Monad_Syntax.bind State_monad.bindS*)
+
+(*section \<open>Hoare logic for the state, exception and nondeterminism monad\<close>
+
+subsection \<open>Hoare triples\<close>
+*)
+Definition predS regs := sequential_state regs -> Prop.
+
+Definition PrePost {Regs A E} (P : predS Regs) (f : monadS Regs A E) (Q : result A E -> predS Regs) : Prop :=
+ (*"\<lbrace>_\<rbrace> _ \<lbrace>_\<rbrace>"*)
+ forall s, P s -> (forall r s', List.In (r, s') (f s) -> Q r s').
+
+Notation "{{ P }} m {{ Q }}" := (PrePost P m Q).
+
+(*
+lemma PrePostI:
+ assumes "\<And>s r s'. P s \<Longrightarrow> (r, s') \<in> f s \<Longrightarrow> Q r s'"
+ shows "PrePost P f Q"
+ using assms unfolding PrePost_def by auto
+
+lemma PrePost_elim:
+ assumes "PrePost P f Q" and "P s" and "(r, s') \<in> f s"
+ obtains "Q r s'"
+ using assms by (fastforce simp: PrePost_def)
+*)
+Lemma PrePost_consequence Regs X E (A P : predS Regs) (f : monadS Regs X E) (B Q : result X E -> predS Regs) :
+ PrePost A f B ->
+ (forall s, P s -> A s) ->
+ (forall v s, B v s -> Q v s) ->
+ PrePost P f Q.
+intros Triple PA BQ.
+intros s Pre r s' IN.
+specialize (Triple s).
+auto.
+Qed.
+
+Lemma PrePost_strengthen_pre Regs X E (A B : predS Regs) (f : monadS Regs X E) (C : result X E -> predS Regs) :
+ PrePost A f C ->
+ (forall s, B s -> A s) ->
+ PrePost B f C.
+eauto using PrePost_consequence.
+Qed.
+
+Lemma PrePost_weaken_post Regs X E (A : predS Regs) (f : monadS Regs X E) (B C : result X E -> predS Regs) :
+ PrePost A f B ->
+ (forall v s, B v s -> C v s) ->
+ PrePost A f C.
+eauto using PrePost_consequence.
+Qed.
+
+Lemma PrePost_True_post (*[PrePost_atomI, intro, simp]:*) Regs A E (P : predS Regs) (m : monadS Regs A E) :
+ PrePost P m (fun _ _ => True).
+unfold PrePost. auto.
+Qed.
+
+Lemma PrePost_any Regs A E (m : monadS Regs A E) (Q : result A E -> predS Regs) :
+ PrePost (fun s => forall r s', List.In (r, s') (m s) -> Q r s') m Q.
+unfold PrePost. auto.
+Qed.
+
+Lemma PrePost_returnS (*[intro, PrePost_atomI]:*) Regs A E (P : result A E -> predS Regs) (x : A) :
+ PrePost (P (Value x)) (returnS x) P.
+unfold PrePost, returnS.
+intros s p r s' IN.
+simpl in IN.
+destruct IN as [[=] | []].
+subst; auto.
+Qed.
+
+Lemma PrePost_bindS (*[intro, PrePost_compositeI]:*) Regs A B E (m : monadS Regs A E) (f : A -> monadS Regs B E) (P : predS Regs) (Q : result B E -> predS Regs) (R : A -> predS Regs) :
+ (forall s a s', List.In (Value a, s') (m s) -> PrePost (R a) (f a) Q) ->
+ (PrePost P m (fun r => match r with Value a => R a | Ex e => Q (Ex e) end)) ->
+ PrePost P (bindS m f) Q.
+intros F M s Pre r s' IN.
+destruct (bindS_cases IN) as [(a & a' & s'' & [= ->] & IN' & IN'') | [(e & [= ->] & IN') | (e & a & s'' & [= ->] & IN' & IN'')]].
+* eapply F. apply IN'. specialize (M s Pre (Value a') s'' IN'). apply M. assumption.
+* specialize (M _ Pre _ _ IN'). apply M.
+* specialize (M _ Pre _ _ IN'). simpl in M. eapply F; eauto.
+Qed.
+
+Lemma PrePost_bindS_ignore Regs A B E (m : monadS Regs A E) (f : monadS Regs B E) (P : predS Regs) (Q : result B E -> predS Regs) (R : predS Regs) :
+ PrePost R f Q ->
+ PrePost P m (fun r => match r with Value a => R | Ex e => Q (Ex e) end) ->
+ PrePost P (bindS m (fun _ => f)) Q.
+intros F M.
+eapply PrePost_bindS; eauto.
+* intros. apply F.
+* apply M.
+Qed.
+
+Lemma PrePost_bindS_unit Regs B E (m : monadS Regs unit E) (f : unit -> monadS Regs B E) P Q R :
+ PrePost R (f tt) Q ->
+ PrePost P m (fun r => match r with Value a => R | Ex e => Q (Ex e) end) ->
+ PrePost P (bindS m f) Q.
+intros F M.
+eapply PrePost_bindS with (R := fun _ => R).
+* intros. destruct a. apply F.
+* apply M.
+Qed.
+
+Lemma PrePost_readS (*[intro, PrePost_atomI]:*) Regs A E (P : result A E -> predS Regs) f :
+ PrePost (fun s => P (Value (f s)) s) (readS f) P.
+unfold PrePost, readS, returnS.
+intros s Pre r s' [H | []].
+inversion H; subst.
+assumption.
+Qed.
+
+Lemma PrePost_updateS (*[intro, PrePost_atomI]:*) Regs E (P : result unit E -> predS Regs) f :
+ PrePost (fun s => P (Value tt) (f s)) (updateS f) P.
+unfold PrePost, readS, returnS.
+intros s Pre r s' [H | []].
+inversion H; subst.
+assumption.
+Qed.
+
+Lemma PrePost_if Regs A E b (f g : monadS Regs A E) P Q :
+ (b = true -> PrePost P f Q) ->
+ (b = false -> PrePost P g Q) ->
+ PrePost P (if b then f else g) Q.
+intros T F.
+destruct b; auto.
+Qed.
+
+Lemma PrePost_if_branch (*[PrePost_compositeI]:*) Regs A E b (f g : monadS Regs A E) Pf Pg Q :
+ (b = true -> PrePost Pf f Q) ->
+ (b = false -> PrePost Pg g Q) ->
+ PrePost (if b then Pf else Pg) (if b then f else g) Q.
+destruct b; auto.
+Qed.
+
+Lemma PrePost_if_then Regs A E b (f g : monadS Regs A E) P Q :
+ b = true ->
+ PrePost P f Q ->
+ PrePost P (if b then f else g) Q.
+intros; subst; auto.
+Qed.
+
+Lemma PrePost_if_else Regs A E b (f g : monadS Regs A E) P Q :
+ b = false ->
+ PrePost P g Q ->
+ PrePost P (if b then f else g) Q.
+intros; subst; auto.
+Qed.
+
+Lemma PrePost_prod_cases (*[PrePost_compositeI]:*) Regs A B E (f : A -> B -> monadS Regs A E) P Q x :
+ PrePost P (f (fst x) (snd x)) Q ->
+ PrePost P (match x with (a, b) => f a b end) Q.
+destruct x; auto.
+Qed.
+
+Lemma PrePost_option_cases (*[PrePost_compositeI]:*) Regs A B E x (s : A -> monadS Regs B E) n PS PN Q :
+ (forall a, PrePost (PS a) (s a) Q) ->
+ PrePost PN n Q ->
+ PrePost (match x with Some a => PS a | None => PN end) (match x with Some a => s a | None => n end) Q.
+destruct x; auto.
+Qed.
+
+Lemma PrePost_let (*[intro, PrePost_compositeI]:*) Regs A B E y (m : A -> monadS Regs B E) P Q :
+ PrePost P (m y) Q ->
+ PrePost P (let x := y in m x) Q.
+auto.
+Qed.
+
+Lemma PrePost_and_boolS (*[PrePost_compositeI]:*) Regs E (l r : monadS Regs bool E) P Q R :
+ PrePost R r Q ->
+ PrePost P l (fun r => match r with Value true => R | _ => Q r end) ->
+ PrePost P (and_boolS l r) Q.
+intros Hr Hl.
+unfold and_boolS.
+eapply PrePost_bindS.
+2: { instantiate (1 := fun a => if a then R else Q (Value false)).
+ eapply PrePost_weaken_post.
+ apply Hl.
+ intros [[|] | ] s H; auto. }
+* intros. destruct a; eauto.
+ apply PrePost_returnS.
+Qed.
+
+Lemma PrePost_or_boolS (*[PrePost_compositeI]:*) Regs E (l r : monadS Regs bool E) P Q R :
+ PrePost R r Q ->
+ PrePost P l (fun r => match r with Value false => R | _ => Q r end) ->
+ PrePost P (or_boolS l r) Q.
+intros Hr Hl.
+unfold or_boolS.
+eapply PrePost_bindS.
+* intros.
+ instantiate (1 := fun a => if a then Q (Value true) else R).
+ destruct a; eauto.
+ apply PrePost_returnS.
+* eapply PrePost_weaken_post.
+ apply Hl.
+ intros [[|] | ] s H; auto.
+Qed.
+
+Lemma PrePost_failS (*[intro, PrePost_atomI]:*) Regs A E msg (Q : result A E -> predS Regs) :
+ PrePost (Q (Ex (Failure msg))) (failS msg) Q.
+intros s Pre r s' [[= <- <-] | []].
+assumption.
+Qed.
+
+Lemma PrePost_assert_expS (*[intro, PrePost_atomI]:*) Regs E (c : bool) m (P : result unit E -> predS Regs) :
+ PrePost (if c then P (Value tt) else P (Ex (Failure m))) (assert_expS c m) P.
+destruct c; simpl.
+* apply PrePost_returnS.
+* apply PrePost_failS.
+Qed.
+
+Lemma PrePost_chooseS (*[intro, PrePost_atomI]:*) Regs A E xs (Q : result A E -> predS Regs) :
+ PrePost (fun s => forall x, List.In x xs -> Q (Value x) s) (chooseS xs) Q.
+unfold PrePost, chooseS.
+intros s IN r s' IN'.
+apply List.in_map_iff in IN'.
+destruct IN' as (x & [= <- <-] & IN').
+auto.
+Qed.
+
+Lemma case_result_combine (*[simp]:*) A E X r (Q : result A E -> X) :
+ (match r with Value a => Q (Value a) | Ex e => Q (Ex e) end) = Q r.
+destruct r; auto.
+Qed.
+
+Lemma PrePost_foreachS_Nil (*[intro, simp, PrePost_atomI]:*) Regs A Vars E vars body (Q : result Vars E -> predS Regs) :
+ PrePost (Q (Value vars)) (foreachS (A := A) nil vars body) Q.
+simpl. apply PrePost_returnS.
+Qed.
+
+Lemma PrePost_foreachS_Cons Regs A Vars E (x : A) xs vars body (Q : result Vars E -> predS Regs) :
+ (forall s vars' s', List.In (Value vars', s') (body x vars s) -> PrePost (Q (Value vars')) (foreachS xs vars' body) Q) ->
+ PrePost (Q (Value vars)) (body x vars) Q ->
+ PrePost (Q (Value vars)) (foreachS (x :: xs) vars body) Q.
+intros XS X.
+simpl.
+eapply PrePost_bindS.
+* apply XS.
+* apply PrePost_weaken_post with (B := Q).
+ assumption.
+ intros; rewrite case_result_combine.
+ assumption.
+Qed.
+
+Lemma PrePost_foreachS_invariant Regs A Vars E (xs : list A) vars body (Q : result Vars E -> predS Regs) :
+ (forall x vars, List.In x xs -> PrePost (Q (Value vars)) (body x vars) Q) ->
+ PrePost (Q (Value vars)) (foreachS xs vars body) Q.
+revert vars.
+induction xs.
+* intros. apply PrePost_foreachS_Nil.
+* intros. apply PrePost_foreachS_Cons.
+ + auto with datatypes.
+ + apply H. auto with datatypes.
+Qed.
+
+(*subsection \<open>Hoare quadruples\<close>
+
+text \<open>It is often convenient to treat the exception case separately. For this purpose, we use
+a Hoare logic similar to the one used in [1]. It features not only Hoare triples, but also quadruples
+with two postconditions: one for the case where the computation succeeds, and one for the case where
+there is an exception.
+
+[1] D. Cock, G. Klein, and T. Sewell, ‘Secure Microkernels, State Monads and Scalable Refinement’,
+in Theorem Proving in Higher Order Logics, 2008, pp. 167–182.\<close>
+*)
+Definition PrePostE {Regs A Ety} (P : predS Regs) (f : monadS Regs A Ety) (Q : A -> predS Regs) (E : ex Ety -> predS Regs) : Prop :=
+(* ("\<lbrace>_\<rbrace> _ \<lbrace>_ \<bar> _\<rbrace>")*)
+ PrePost P f (fun v => match v with Value a => Q a | Ex e => E e end).
+
+Notation "{{ P }} m {{ Q | X }}" := (PrePostE P m Q X).
+
+(*lemmas PrePost_defs = PrePost_def PrePostE_def*)
+
+Lemma PrePostE_I (*[case_names Val Err]:*) Regs A Ety (P : predS Regs) f (Q : A -> predS Regs) (E : ex Ety -> predS Regs) :
+ (forall s a s', P s -> List.In (Value a, s') (f s) -> Q a s') ->
+ (forall s e s', P s -> List.In (Ex e, s') (f s) -> E e s') ->
+ PrePostE P f Q E.
+intros. unfold PrePostE.
+unfold PrePost.
+intros s Pre [a | e] s' IN; eauto.
+Qed.
+
+Lemma PrePostE_PrePost Regs A Ety P m (Q : A -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePost P m (fun v => match v with Value a => Q a | Ex e => E e end) ->
+ PrePostE P m Q E.
+auto.
+Qed.
+
+Lemma PrePostE_elim Regs A Ety P f r s s' (Q : A -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE P f Q E ->
+ P s ->
+ List.In (r, s') (f s) ->
+ (exists v, r = Value v /\ Q v s') \/
+ (exists e, r = Ex e /\ E e s').
+intros PP Pre IN.
+specialize (PP _ Pre _ _ IN).
+destruct r; eauto.
+Qed.
+
+Lemma PrePostE_consequence Regs Aty Ety (P : predS Regs) f A B C (Q : Aty -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE A f B C ->
+ (forall s, P s -> A s) ->
+ (forall v s, B v s -> Q v s) ->
+ (forall e s, C e s -> E e s) ->
+ PrePostE P f Q E.
+intros PP PA BQ CE.
+intros s Pre [a | e] s' IN.
+* apply BQ. specialize (PP _ (PA _ Pre) _ _ IN).
+ apply PP.
+* apply CE. specialize (PP _ (PA _ Pre) _ _ IN).
+ apply PP.
+Qed.
+
+Lemma PrePostE_strengthen_pre Regs Aty Ety (P : predS Regs) f R (Q : Aty -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE R f Q E ->
+ (forall s, P s -> R s) ->
+ PrePostE P f Q E.
+intros PP PR.
+eapply PrePostE_consequence; eauto.
+Qed.
+
+Lemma PrePostE_weaken_post Regs Aty Ety (A : predS Regs) f (B C : Aty -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE A f B E ->
+ (forall v s, B v s -> C v s) ->
+ PrePostE A f C E.
+intros PP BC.
+eauto using PrePostE_consequence.
+Qed.
+
+Lemma PrePostE_weaken_Epost Regs Aty Ety (A : predS Regs) f (B : Aty -> predS Regs) (E F : ex Ety -> predS Regs) :
+ PrePostE A f B E ->
+ (forall v s, E v s -> F v s) ->
+ PrePostE A f B F.
+intros PP EF.
+eauto using PrePostE_consequence.
+Qed.
+(*named_theorems PrePostE_compositeI
+named_theorems PrePostE_atomI*)
+
+Lemma PrePostE_conj_conds Regs Aty Ety (P1 P2 : predS Regs) m (Q1 Q2 : Aty -> predS Regs) (E1 E2 : ex Ety -> predS Regs) :
+ PrePostE P1 m Q1 E1 ->
+ PrePostE P2 m Q2 E2 ->
+ PrePostE (fun s => P1 s /\ P2 s) m (fun r s => Q1 r s /\ Q2 r s) (fun e s => E1 e s /\ E2 e s).
+intros H1 H2.
+apply PrePostE_I.
+* intros s a s' [p1 p2] IN.
+ specialize (H1 _ p1 _ _ IN).
+ specialize (H2 _ p2 _ _ IN).
+ simpl in *.
+ auto.
+* intros s a s' [p1 p2] IN.
+ specialize (H1 _ p1 _ _ IN).
+ specialize (H2 _ p2 _ _ IN).
+ simpl in *.
+ auto.
+Qed.
+
+(*lemmas PrePostE_conj_conds_consequence = PrePostE_conj_conds[THEN PrePostE_consequence]*)
+
+Lemma PrePostE_post_mp Regs Aty Ety (P : predS Regs) m (Q Q' : Aty -> predS Regs) (E: ex Ety -> predS Regs) :
+ PrePostE P m Q' E ->
+ PrePostE P m (fun r s => Q' r s -> Q r s) E ->
+ PrePostE P m Q E.
+intros H1 H2.
+eapply PrePostE_conj_conds in H1. 2: apply H2.
+eapply PrePostE_consequence. apply H1. all: simpl; intuition.
+Qed.
+
+Lemma PrePostE_cong Regs Aty Ety (P1 P2 : predS Regs) m1 m2 (Q1 Q2 : Aty -> predS Regs) (E1 E2 : ex Ety -> predS Regs) :
+ (forall s, P1 s <-> P2 s) ->
+ (forall s, P1 s -> m1 s = m2 s) ->
+ (forall r s, Q1 r s <-> Q2 r s) ->
+ (forall e s, E1 e s <-> E2 e s) ->
+ PrePostE P1 m1 Q1 E1 <-> PrePostE P2 m2 Q2 E2.
+intros P12 m12 Q12 E12.
+unfold PrePostE, PrePost.
+split.
+* intros. apply P12 in H0. rewrite <- m12 in H1; auto. specialize (H _ H0 _ _ H1).
+ destruct r; [ apply Q12 | apply E12]; auto.
+* intros. rewrite m12 in H1; auto. apply P12 in H0. specialize (H _ H0 _ _ H1).
+ destruct r; [ apply Q12 | apply E12]; auto.
+Qed.
+
+Lemma PrePostE_True_post (*[PrePostE_atomI, intro, simp]:*) Regs A E P (m : monadS Regs A E) :
+ PrePostE P m (fun _ _ => True) (fun _ _ => True).
+intros s Pre [a | e]; auto.
+Qed.
+
+Lemma PrePostE_any Regs A Ety m (Q : result A Ety -> predS Regs) E :
+ PrePostE (Ety := Ety) (fun s => forall r s', List.In (r, s') (m s) -> match r with Value a => Q a s' | Ex e => E e s' end) m Q E.
+apply PrePostE_I.
+intros. apply (H (Value a)); auto.
+intros. apply (H (Ex e)); auto.
+Qed.
+
+Lemma PrePostE_returnS (*[PrePostE_atomI, intro, simp]:*) Regs A E P (x : A) (Q : ex E -> predS Regs) :
+ PrePostE (P x) (returnS x) P Q.
+unfold PrePostE, PrePost.
+intros s Pre r s' [[= <- <-] | []].
+assumption.
+Qed.
+
+Lemma PrePostE_bindS (*[intro, PrePostE_compositeI]:*) Regs A B Ety P m (f : A -> monadS Regs B Ety) Q R E :
+ (forall s a s', List.In (Value a, s') (m s) -> PrePostE (R a) (f a) Q E) ->
+ PrePostE P m R E ->
+ PrePostE P (bindS m f) Q E.
+intros.
+unfold PrePostE in *.
+eauto using PrePost_bindS.
+Qed.
+
+Lemma PrePostE_bindS_ignore Regs A B Ety (P : predS Regs) (m : monadS Regs A Ety) (f : monadS Regs B Ety) R Q E :
+ PrePostE R f Q E ->
+ PrePostE P m (fun _ => R) E ->
+ PrePostE P (bindS m (fun _ => f)) Q E.
+apply PrePost_bindS_ignore.
+Qed.
+
+Lemma PrePostE_bindS_unit Regs A Ety (P : predS Regs) (m : monadS Regs unit Ety) (f : unit -> monadS Regs A Ety) Q R E :
+ PrePostE R (f tt) Q E ->
+ PrePostE P m (fun _ => R) E ->
+ PrePostE P (bindS m f) Q E.
+apply PrePost_bindS_unit.
+Qed.
+
+Lemma PrePostE_readS (*[PrePostE_atomI, intro]:*) Regs A Ety (P : predS Regs) f (Q : result A Ety -> predS Regs) E :
+ PrePostE (Ety := Ety) (fun s => Q (f s) s) (readS f) Q E.
+unfold PrePostE, PrePost, readS.
+intros s Pre [a | e] s' [[= <- <-] | []].
+assumption.
+Qed.
+
+Lemma PrePostE_updateS (*[PrePostE_atomI, intro]:*) Regs Ety f (Q : unit -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE (fun s => Q tt (f s)) (updateS f) Q E.
+intros s Pre [a | e] s' [[= <- <-] | []].
+assumption.
+Qed.
+
+Lemma PrePostE_if_branch (*[PrePostE_compositeI]:*) Regs A Ety (b : bool) (f g : monadS Regs A Ety) Pf Pg Q E :
+ (b = true -> PrePostE Pf f Q E) ->
+ (b = false -> PrePostE Pg g Q E) ->
+ PrePostE (if b then Pf else Pg) (if b then f else g) Q E.
+destruct b; auto.
+Qed.
+
+Lemma PrePostE_if Regs A Ety (b : bool) (f g : monadS Regs A Ety) P Q E :
+ (b = true -> PrePostE P f Q E) ->
+ (b = false -> PrePostE P g Q E) ->
+ PrePostE P (if b then f else g) Q E.
+destruct b; auto.
+Qed.
+
+Lemma PrePostE_if_then Regs A Ety (b : bool) (f g : monadS Regs A Ety) P Q E :
+ b = true ->
+ PrePostE P f Q E ->
+ PrePostE P (if b then f else g) Q E.
+intros; subst; auto.
+Qed.
+
+Lemma PrePostE_if_else Regs A Ety (b : bool) (f g : monadS Regs A Ety) P Q E :
+ b = false ->
+ PrePostE P g Q E ->
+ PrePostE P (if b then f else g) Q E.
+intros; subst; auto.
+Qed.
+
+Lemma PrePostE_prod_cases (*[PrePostE_compositeI]:*) Regs A B C Ety x (f : A -> B -> monadS Regs C Ety) P Q E :
+ PrePostE P (f (fst x) (snd x)) Q E ->
+ PrePostE P (match x with (a, b) => f a b end) Q E.
+destruct x; auto.
+Qed.
+
+Lemma PrePostE_option_cases (*[PrePostE_compositeI]:*) Regs A B Ety x (s : option A -> monadS Regs B Ety) n PS PN Q E :
+ (forall a, PrePostE (PS a) (s a) Q E) ->
+ PrePostE PN n Q E ->
+ PrePostE (match x with Some a => PS a | None => PN end) (match x with Some a => s a | None => n end) Q E.
+apply PrePost_option_cases.
+Qed.
+
+Lemma PrePostE_sum_cases (*[PrePostE_compositeI]:*) Regs A B C Ety x (l : A -> monadS Regs C Ety) (r : B -> monadS Regs C Ety) Pl Pr Q E :
+ (forall a, PrePostE (Pl a) (l a) Q E) ->
+ (forall b, PrePostE (Pr b) (r b) Q E) ->
+ PrePostE (match x with inl a => Pl a | inr b => Pr b end) (match x with inl a => l a | inr b => r b end) Q E.
+intros; destruct x; auto.
+Qed.
+
+Lemma PrePostE_let (*[PrePostE_compositeI]:*) Regs A B Ety y (m : A -> monadS Regs B Ety) P Q E :
+ PrePostE P (m y) Q E ->
+ PrePostE P (let x := y in m x) Q E.
+auto.
+Qed.
+
+Lemma PrePostE_and_boolS (*[PrePostE_compositeI]:*) Regs Ety (l r : monadS Regs bool Ety) P Q R E :
+ PrePostE R r Q E ->
+ PrePostE P l (fun r => if r then R else Q false) E ->
+ PrePostE P (and_boolS l r) Q E.
+intros Hr Hl.
+unfold and_boolS.
+eapply PrePostE_bindS.
+* intros.
+ instantiate (1 := fun a => if a then R else Q false).
+ destruct a; eauto.
+ apply PrePostE_returnS.
+* assumption.
+Qed.
+
+Lemma PrePostE_or_boolS (*[PrePostE_compositeI]:*) Regs Ety (l r : monadS Regs bool Ety) P Q R E :
+ PrePostE R r Q E ->
+ PrePostE P l (fun r => if r then Q true else R) E ->
+ PrePostE P (or_boolS l r) Q E.
+intros Hr Hl.
+unfold or_boolS.
+eapply PrePostE_bindS.
+* intros.
+ instantiate (1 := fun a => if a then Q true else R).
+ destruct a; eauto.
+ apply PrePostE_returnS.
+* assumption.
+Qed.
+
+Lemma PrePostE_failS (*[PrePostE_atomI, intro]:*) Regs A Ety msg (Q : A -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE (E (Failure msg)) (failS msg) Q E.
+unfold PrePostE, PrePost, failS.
+intros s Pre r s' [[= <- <-] | []].
+assumption.
+Qed.
+
+Lemma PrePostE_assert_expS (*[PrePostE_atomI, intro]:*) Regs Ety (c : bool) m P (Q : ex Ety -> predS Regs) :
+ PrePostE (if c then P tt else Q (Failure m)) (assert_expS c m) P Q.
+unfold assert_expS.
+destruct c; auto using PrePostE_returnS, PrePostE_failS.
+Qed.
+
+Lemma PrePostE_maybe_failS (*[PrePostE_atomI]:*) Regs A Ety msg v (Q : A -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE (fun s => match v with Some v => Q v s | None => E (Failure msg) s end) (maybe_failS msg v) Q E.
+unfold maybe_failS.
+destruct v; auto using PrePostE_returnS, PrePostE_failS.
+Qed.
+
+Lemma PrePostE_exitS (*[PrePostE_atomI, intro]:*) Regs A Ety msg (Q : A -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE (E (Failure "exit")) (exitS msg) Q E.
+unfold exitS.
+apply PrePostE_failS.
+Qed.
+
+Lemma PrePostE_chooseS (*[intro, PrePostE_atomI]:*) Regs A Ety (xs : list A) (Q : A -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE (fun s => forall x, List.In x xs -> Q x s) (chooseS xs) Q E.
+unfold chooseS.
+intros s IN r s' IN'.
+apply List.in_map_iff in IN'.
+destruct IN' as (x & [= <- <-] & IN').
+auto.
+Qed.
+
+Lemma PrePostE_throwS (*[PrePostE_atomI]:*) Regs A Ety e (Q : A -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE (E (Throw e)) (throwS e) Q E.
+unfold throwS.
+intros s Pre r s' [[= <- <-] | []].
+assumption.
+Qed.
+
+Lemma PrePostE_try_catchS (*[PrePostE_compositeI]:*) Regs A E1 E2 m h P (Ph : E1 -> predS Regs) (Q : A -> predS Regs) (E : ex E2 -> predS Regs) :
+ (forall s e s', List.In (Ex (Throw e), s') (m s) -> PrePostE (Ph e) (h e) Q E) ->
+ PrePostE P m Q (fun ex => match ex with Throw e => Ph e | Failure msg => E (Failure msg) end) ->
+ PrePostE P (try_catchS m h) Q E.
+intros.
+intros s Pre r s' IN.
+destruct (try_catchS_cases IN) as [(a' & [= ->] & IN') | [(msg & [= ->] & IN') | (e & s'' & IN1 & IN2)]].
+* specialize (H0 _ Pre _ _ IN'). apply H0.
+* specialize (H0 _ Pre _ _ IN'). apply H0.
+* specialize (H _ _ _ IN1). specialize (H0 _ Pre _ _ IN1). simpl in *.
+ specialize (H _ H0 _ _ IN2). apply H.
+Qed.
+
+Lemma PrePostE_catch_early_returnS (*[PrePostE_compositeI]:*) Regs A Ety m P (Q : A -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE P m Q (fun ex => match ex with Throw (inl a) => Q a | Throw (inr e) => E (Throw e) | Failure msg => E (Failure msg) end) ->
+ PrePostE P (catch_early_returnS m) Q E.
+unfold catch_early_returnS.
+intro H.
+apply PrePostE_try_catchS with (Ph := fun e => match e with inl a => Q a | inr e => E (Throw e) end).
+* intros. destruct e.
+ + apply PrePostE_returnS.
+ + apply PrePostE_throwS.
+* apply H.
+Qed.
+
+Lemma PrePostE_early_returnS (*[PrePostE_atomI]:*) Regs A E1 E2 r (Q : A -> predS Regs) (E : ex (E1 + E2) -> predS Regs) :
+ PrePostE (E (Throw (inl r))) (early_returnS r) Q E.
+unfold early_returnS.
+apply PrePostE_throwS.
+Qed.
+
+Lemma PrePostE_liftRS (*[PrePostE_compositeI]:*) Regs A E1 E2 m P (Q : A -> predS Regs) (E : ex (E1 + E2) -> predS Regs) :
+ PrePostE P m Q (fun ex => match ex with Throw e => E (Throw (inr e)) | Failure msg => E (Failure msg) end) ->
+ PrePostE P (liftRS m) Q E.
+unfold liftRS.
+apply PrePostE_try_catchS.
+auto using PrePostE_throwS.
+Qed.
+
+Lemma PrePostE_foreachS_Cons Regs A Vars Ety (x : A) xs vars body (Q : Vars -> predS Regs) (E : ex Ety -> predS Regs) :
+ (forall s vars' s', List.In (Value vars', s') (body x vars s) -> PrePostE (Q vars') (foreachS xs vars' body) Q E) ->
+ PrePostE (Q vars) (body x vars) Q E ->
+ PrePostE (Q vars) (foreachS (x :: xs) vars body) Q E.
+intros.
+simpl.
+apply PrePostE_bindS with (R := Q); auto.
+Qed.
+
+Lemma PrePostE_foreachS_invariant Regs A Vars Ety (xs : list A) vars body (Q : Vars -> predS Regs) (E : ex Ety -> predS Regs) :
+ (forall x vars, List.In x xs -> PrePostE (Q vars) (body x vars) Q E) ->
+ PrePostE (Q vars) (foreachS xs vars body) Q E.
+unfold PrePostE.
+intros H.
+apply PrePost_foreachS_invariant with (Q := fun v => match v with Value a => Q a | Ex e => E e end).
+auto.
+Qed.
+
+
+Lemma PrePostE_use_pre Regs A Ety m (P : predS Regs) (Q : A -> predS Regs) (E : ex Ety -> predS Regs) :
+ (forall s, P s -> PrePostE P m Q E) ->
+ PrePostE P m Q E.
+unfold PrePostE, PrePost.
+intros H s p r s' IN.
+eapply H; eauto.
+Qed.
+
+Local Open Scope Z.
+Local Opaque _limit_reduces.
+Ltac gen_reduces :=
+ match goal with |- context[@_limit_reduces ?a ?b ?c] => generalize (@_limit_reduces a b c) end.
+
+
+Lemma PrePostE_untilST Regs Vars Ety vars measure cond (body : Vars -> monadS Regs Vars Ety) Inv Inv' (Q : Vars -> predS Regs) E :
+ (forall vars, PrePostE (Inv' Q vars) (cond vars) (fun c s' => Inv Q vars s' /\ (c = true -> Q vars s')) E) ->
+ (forall vars, PrePostE (Inv Q vars) (body vars) (fun vars' s' => Inv' Q vars' s' /\ measure vars' < measure vars) E) ->
+ (forall vars s, Inv Q vars s -> measure vars >= 0) ->
+ PrePostE (Inv Q vars) (untilST vars measure cond body) Q E.
+
+intros Hcond Hbody Hmeasure.
+unfold untilST.
+apply PrePostE_use_pre. intros s0 Pre0.
+assert (measure vars >= 0) as Hlimit_0 by eauto. clear s0 Pre0.
+remember (measure vars) as limit eqn: Heqlimit in Hlimit_0 |- *.
+assert (measure vars <= limit) as Hlimit by omega. clear Heqlimit.
+generalize (Sail2_prompt.Zwf_guarded limit).
+revert vars Hlimit.
+apply Wf_Z.natlike_ind with (x := limit).
+* intros vars Hmeasure_limit [acc]. simpl.
+ eapply PrePostE_bindS; [ | apply Hbody ].
+ intros s vars' s' IN.
+ eapply PrePostE_bindS with (R := (fun c s' => (Inv Q vars' s' /\ (c = true -> Q vars' s')) /\ measure vars' < measure vars)).
+ 2: {
+ apply PrePostE_weaken_Epost with (E := (fun e s' => E e s' /\ measure vars' < measure vars)). 2: tauto.
+ eapply PrePostE_conj_conds.
+ apply Hcond.
+ apply PrePostE_I; tauto.
+ }
+ intros.
+ destruct a.
+ - eapply PrePostE_strengthen_pre; try apply PrePostE_returnS.
+ intros ? [[? ?] ?]; auto.
+ - apply PrePostE_I;
+ intros ? ? ? [[Pre ?] ?] ?; exfalso;
+ specialize (Hmeasure _ _ Pre); omega.
+* intros limit' Hlimit' IH vars Hmeasure_limit [acc].
+ simpl.
+ destruct (Z_ge_dec _ _); try omega.
+ eapply PrePostE_bindS; [ | apply Hbody].
+ intros s vars' s' IN.
+ eapply PrePostE_bindS with (R := (fun c s' => (Inv Q vars' s' /\ (c = true -> Q vars' s')) /\ measure vars' < measure vars)).
+ 2: {
+ apply PrePostE_weaken_Epost with (E := (fun e s' => E e s' /\ measure vars' < measure vars)). 2: tauto.
+ eapply PrePostE_conj_conds.
+ apply Hcond.
+ apply PrePostE_I; tauto.
+ }
+ intros.
+ destruct a.
+ - eapply PrePostE_strengthen_pre; try apply PrePostE_returnS.
+ intros ? [[? ?] ?]; auto.
+ - gen_reduces.
+ replace (Z.succ limit' - 1) with limit'; [ | omega].
+ intro acc'.
+ apply PrePostE_use_pre. intros sx [[Pre _] Hreduces].
+ apply Hmeasure in Pre.
+ eapply PrePostE_strengthen_pre; [apply IH | ].
+ + omega.
+ + tauto.
+* omega.
+Qed.
+
+
+Lemma PrePostE_untilST_pure_cond Regs Vars Ety vars measure cond (body : Vars -> monadS Regs Vars Ety) Inv (Q : Vars -> predS Regs) E :
+ (forall vars, PrePostE (Inv Q vars) (body vars) (fun vars' s' => Inv Q vars' s' /\ measure vars' < measure vars /\ (cond vars' = true -> Q vars' s')) E) ->
+ (forall vars s, Inv Q vars s -> measure vars >= 0) ->
+ (PrePostE (Inv Q vars) (untilST vars measure (fun vars => returnS (cond vars)) body) Q E).
+intros Hbody Hmeasure.
+apply PrePostE_untilST with (Inv' := fun Q vars s => Inv Q vars s /\ (cond vars = true -> Q vars s)).
+* intro.
+ apply PrePostE_returnS with (P := fun c s' => Inv Q vars0 s' /\ (c = true -> Q vars0 s')).
+* intro.
+ eapply PrePost_weaken_post; [ apply Hbody | ].
+ simpl. intros [a |e]; eauto. tauto.
+* apply Hmeasure.
+Qed.
+
+Local Close Scope Z.
+
+(*
+lemma PrePostE_liftState_untilM:
+ assumes dom: (forall s, Inv Q vars s -> untilM_dom (vars, cond, body))
+ and cond: (forall vars, PrePostE (Inv' Q vars) (liftState r (cond vars)) (fun c s' => Inv Q vars s' /\ (c \<longrightarrow> Q vars s')) E)
+ and body: (forall vars, PrePostE (Inv Q vars) (liftState r (body vars)) (Inv' Q) E)
+ shows "PrePostE (Inv Q vars) (liftState r (untilM vars cond body)) Q E"
+proof -
+ have domS: "untilS_dom (vars, liftState r \<circ> cond, liftState r \<circ> body, s)" if "Inv Q vars s" for s
+ using dom that by (intro untilM_dom_untilS_dom)
+ then have "PrePostE (Inv Q vars) (untilS vars (liftState r \<circ> cond) (liftState r \<circ> body)) Q E"
+ using cond body by (auto intro: PrePostE_untilS simp: comp_def)
+ moreover have "liftState r (untilM vars cond body) s = untilS vars (liftState r \<circ> cond) (liftState r \<circ> body) s"
+ if "Inv Q vars s" for s
+ unfolding liftState_untilM[OF domS[OF that] dom[OF that]] ..
+ ultimately show ?thesis by (auto cong: PrePostE_cong)
+qed
+
+lemma PrePostE_liftState_untilM_pure_cond:
+ assumes dom: (forall s, Inv Q vars s -> untilM_dom (vars, return \<circ> cond, body)"
+ and body: (forall vars, PrePostE (Inv Q vars) (liftState r (body vars)) (fun vars' s' => Inv Q vars' s' /\ (cond vars' \<longrightarrow> Q vars' s')) E"
+ shows "PrePostE (Inv Q vars) (liftState r (untilM vars (return \<circ> cond) body)) Q E"
+ using assms by (intro PrePostE_liftState_untilM) (auto simp: comp_def liftState_simp)
+*)
+Lemma PrePostE_choose_boolS_any (*[PrePostE_atomI]:*) Regs Ety unit_val (Q : bool -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE (fun s => forall b, Q b s) (choose_boolS unit_val) Q E.
+unfold choose_boolS, seqS.
+eapply PrePostE_strengthen_pre.
+apply PrePostE_chooseS.
+simpl. intros. destruct x; auto.
+Qed.
+
+Lemma PrePostE_bool_of_bitU_nondetS_any Regs Ety b (Q : bool -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE (fun s => forall b, Q b s) (bool_of_bitU_nondetS b) Q E.
+unfold bool_of_bitU_nondetS, undefined_boolS.
+destruct b.
+* intros s Pre r s' [[= <- <-] | []]. auto.
+* intros s Pre r s' [[= <- <-] | []]. auto.
+* apply PrePostE_choose_boolS_any.
+Qed.
+(*
+Lemma PrePostE_bools_of_bits_nondetS_any:
+ PrePostE (fun s => forall bs, Q bs s) (bools_of_bits_nondetS bs) Q E.
+ unfolding bools_of_bits_nondetS_def
+ by (rule PrePostE_weaken_post[where B = "fun _ s => forall bs, Q bs s"], rule PrePostE_strengthen_pre,
+ (rule PrePostE_foreachS_invariant[OF PrePostE_strengthen_pre] PrePostE_bindS PrePostE_returnS
+ PrePostE_bool_of_bitU_nondetS_any)+)
+ auto
+*)
+Lemma PrePostE_choose_boolsS_any Regs Ety n (Q : list bool -> predS Regs) (E : ex Ety -> predS Regs) :
+ PrePostE (fun s => forall bs, Q bs s) (choose_boolsS n) Q E.
+unfold choose_boolsS, genlistS.
+apply PrePostE_weaken_post with (B := fun _ s => forall bs, Q bs s).
+* apply PrePostE_foreachS_invariant with (Q := fun _ s => forall bs, Q bs s).
+ intros. apply PrePostE_bindS with (R := fun _ s => forall bs, Q bs s).
+ + intros. apply PrePostE_returnS with (P := fun _ s => forall bs, Q bs s).
+ + eapply PrePostE_strengthen_pre.
+ apply PrePostE_choose_boolS_any.
+ intuition.
+* intuition.
+Qed.
+
+Lemma nth_error_exists {A} {l : list A} {n} :
+ n < Datatypes.length l -> exists x, List.In x l /\ List.nth_error l n = Some x.
+revert n. induction l.
+* simpl. intros. apply PeanoNat.Nat.nlt_0_r in H. destruct H.
+* intros. destruct n.
+ + exists a. auto with datatypes.
+ + simpl in H. apply Lt.lt_S_n in H.
+ destruct (IHl n H) as [x H1].
+ intuition eauto with datatypes.
+Qed.
+
+Lemma nth_error_modulo {A} {xs : list A} n :
+ xs <> nil ->
+ exists x, List.In x xs /\ List.nth_error xs (PeanoNat.Nat.modulo n (Datatypes.length xs)) = Some x.
+intro notnil.
+assert (Datatypes.length xs <> 0) by (rewrite List.length_zero_iff_nil; auto).
+assert (PeanoNat.Nat.modulo n (Datatypes.length xs) < Datatypes.length xs) by auto using PeanoNat.Nat.mod_upper_bound.
+destruct (nth_error_exists H0) as [x [H1 H2]].
+exists x.
+auto.
+Qed.
+
+Lemma PrePostE_internal_pick Regs A Ety (xs : list A) (Q : A -> predS Regs) (E : ex Ety -> predS Regs) :
+ xs <> nil ->
+ PrePostE (fun s => forall x, List.In x xs -> Q x s) (internal_pickS xs) Q E.
+unfold internal_pickS.
+intro notnil.
+eapply PrePostE_bindS with (R := fun _ s => forall x, List.In x xs -> Q x s).
+* intros.
+ destruct (nth_error_modulo (Sail2_values.nat_of_bools a) notnil) as (x & IN & nth).
+ rewrite nth.
+ eapply PrePostE_strengthen_pre.
+ apply PrePostE_returnS.
+ intuition.
+* eapply PrePostE_strengthen_pre.
+ apply PrePostE_choose_boolsS_any.
+ intuition.
+Qed.
diff --git a/prover_snapshots/coq/lib/sail/Makefile b/prover_snapshots/coq/lib/sail/Makefile
new file mode 100644
index 0000000..fa453d9
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Makefile
@@ -0,0 +1,26 @@
+BBV_DIR?=../../../bbv
+
+CORESRC=Sail2_prompt_monad.v Sail2_prompt.v Sail2_impl_base.v Sail2_instr_kinds.v Sail2_operators_bitlists.v Sail2_operators_mwords.v Sail2_operators.v Sail2_values.v Sail2_state_monad.v Sail2_state.v Sail2_state_lifting.v Sail2_string.v Sail2_real.v
+PROOFSRC=Sail2_state_monad_lemmas.v Sail2_state_lemmas.v Hoare.v
+SRC=$(CORESRC) $(PROOFSRC)
+
+COQ_LIBS = -R . Sail -R "$(BBV_DIR)/theories" bbv
+
+TARGETS=$(SRC:.v=.vo)
+
+.PHONY: all clean *.ide
+
+all: $(TARGETS)
+clean:
+ rm -f -- $(TARGETS) $(TARGETS:.vo=.glob) $(TARGETS:%.vo=.%.aux) deps
+
+%.vo: %.v
+ coqc $(COQ_LIBS) $<
+
+%.ide: %.v
+ coqide $(COQ_LIBS) $<
+
+deps: $(SRC)
+ coqdep $(COQ_LIBS) $(SRC) > deps
+
+-include deps
diff --git a/prover_snapshots/coq/lib/sail/Sail2_impl_base.v b/prover_snapshots/coq/lib/sail/Sail2_impl_base.v
new file mode 100644
index 0000000..464c290
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_impl_base.v
@@ -0,0 +1,1103 @@
+(*========================================================================*)
+(* Sail *)
+(* *)
+(* Copyright (c) 2013-2017 *)
+(* Kathyrn Gray *)
+(* Shaked Flur *)
+(* Stephen Kell *)
+(* Gabriel Kerneis *)
+(* Robert Norton-Wright *)
+(* Christopher Pulte *)
+(* Peter Sewell *)
+(* Alasdair Armstrong *)
+(* Brian Campbell *)
+(* Thomas Bauereiss *)
+(* Anthony Fox *)
+(* Jon French *)
+(* Dominic Mulligan *)
+(* Stephen Kell *)
+(* Mark Wassell *)
+(* *)
+(* All rights reserved. *)
+(* *)
+(* This software was developed by the University of Cambridge Computer *)
+(* Laboratory as part of the Rigorous Engineering of Mainstream Systems *)
+(* (REMS) project, funded by EPSRC grant EP/K008528/1. *)
+(* *)
+(* Redistribution and use in source and binary forms, with or without *)
+(* modification, are permitted provided that the following conditions *)
+(* are met: *)
+(* 1. Redistributions of source code must retain the above copyright *)
+(* notice, this list of conditions and the following disclaimer. *)
+(* 2. Redistributions in binary form must reproduce the above copyright *)
+(* notice, this list of conditions and the following disclaimer in *)
+(* the documentation and/or other materials provided with the *)
+(* distribution. *)
+(* *)
+(* THIS SOFTWARE IS PROVIDED BY THE AUTHOR AND CONTRIBUTORS ``AS IS'' *)
+(* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED *)
+(* TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A *)
+(* PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE AUTHOR OR *)
+(* CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, *)
+(* SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT *)
+(* LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF *)
+(* USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND *)
+(* ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, *)
+(* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT *)
+(* OF THE USE OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF *)
+(* SUCH DAMAGE. *)
+(*========================================================================*)
+
+Require Import Sail2_instr_kinds.
+
+(*
+class ( EnumerationType 'a )
+ val toNat : 'a -> nat
+end
+
+
+val enumeration_typeCompare : forall 'a. EnumerationType 'a => 'a -> 'a -> ordering
+let ~{ocaml} enumeration_typeCompare e1 e2 =
+ compare (toNat e1) (toNat e2)
+let inline {ocaml} enumeration_typeCompare = defaultCompare
+
+
+default_instance forall 'a. EnumerationType 'a => (Ord 'a)
+ let compare = enumeration_typeCompare
+ let (<) r1 r2 = (enumeration_typeCompare r1 r2) = LT
+ let (<=) r1 r2 = (enumeration_typeCompare r1 r2) <> GT
+ let (>) r1 r2 = (enumeration_typeCompare r1 r2) = GT
+ let (>=) r1 r2 = (enumeration_typeCompare r1 r2) <> LT
+end
+
+
+
+(* maybe isn't a member of type Ord - this should be in the Lem standard library*)
+instance forall 'a. Ord 'a => (Ord (maybe 'a))
+ let compare = maybeCompare compare
+ let (<) r1 r2 = (maybeCompare compare r1 r2) = LT
+ let (<=) r1 r2 = (maybeCompare compare r1 r2) <> GT
+ let (>) r1 r2 = (maybeCompare compare r1 r2) = GT
+ let (>=) r1 r2 = (maybeCompare compare r1 r2) <> LT
+end
+
+type word8 = nat (* bounded at a byte, for when lem supports it*)
+
+type end_flag =
+ | E_big_endian
+ | E_little_endian
+
+type bit =
+ | Bitc_zero
+ | Bitc_one
+
+type bit_lifted =
+ | Bitl_zero
+ | Bitl_one
+ | Bitl_undef (* used for modelling h/w arch unspecified bits *)
+ | Bitl_unknown (* used for interpreter analysis exhaustive execution *)
+
+type direction =
+ | D_increasing
+ | D_decreasing
+
+let dir_of_bool is_inc = if is_inc then D_increasing else D_decreasing
+let bool_of_dir = function
+ | D_increasing -> true
+ | D_decreasing -> false
+ end
+
+(* at some point this should probably not mention bit_lifted anymore *)
+type register_value = <|
+ rv_bits: list bit_lifted (* MSB first, smallest index number *);
+ rv_dir: direction;
+ rv_start: nat ;
+ rv_start_internal: nat;
+ (*when dir is increasing, rv_start = rv_start_internal.
+ Otherwise, tells interpreter how to reconstruct a proper decreasing value*)
+ |>
+
+type byte_lifted = Byte_lifted of list bit_lifted (* of length 8 *) (*MSB first everywhere*)
+
+type instruction_field_value = list bit
+
+type byte = Byte of list bit (* of length 8 *) (*MSB first everywhere*)
+
+type address_lifted = Address_lifted of list byte_lifted (* of length 8 for 64bit machines*) * maybe integer
+(* for both values of end_flag, MSBy first *)
+
+type memory_byte = byte_lifted (* of length 8 *) (*MSB first everywhere*)
+
+type memory_value = list memory_byte
+(* the list is of length >=1 *)
+(* the head of the list is the byte stored at the lowest address;
+when calling a Sail function with a wmv effect, the least significant 8
+bits of the bit vector passed to the function will be interpreted as
+the lowest address byte; similarly, when calling a Sail function with
+rmem effect, the lowest address byte will be placed in the least
+significant 8 bits of the bit vector returned by the function; this
+behaviour is consistent with little-endian. *)
+
+
+(* not sure which of these is more handy yet *)
+type address = Address of list byte (* of length 8 *) * integer
+(* type address = Address of integer *)
+
+type opcode = Opcode of list byte (* of length 4 *)
+
+(** typeclass instantiations *)
+
+instance (EnumerationType bit)
+ let toNat = function
+ | Bitc_zero -> 0
+ | Bitc_one -> 1
+ end
+end
+
+instance (EnumerationType bit_lifted)
+ let toNat = function
+ | Bitl_zero -> 0
+ | Bitl_one -> 1
+ | Bitl_undef -> 2
+ | Bitl_unknown -> 3
+ end
+end
+
+let ~{ocaml} byte_liftedCompare (Byte_lifted b1) (Byte_lifted b2) = compare b1 b2
+let inline {ocaml} byte_liftedCompare = defaultCompare
+
+let ~{ocaml} byte_liftedLess b1 b2 = byte_liftedCompare b1 b2 = LT
+let ~{ocaml} byte_liftedLessEq b1 b2 = byte_liftedCompare b1 b2 <> GT
+let ~{ocaml} byte_liftedGreater b1 b2 = byte_liftedCompare b1 b2 = GT
+let ~{ocaml} byte_liftedGreaterEq b1 b2 = byte_liftedCompare b1 b2 <> LT
+
+let inline {ocaml} byte_liftedLess = defaultLess
+let inline {ocaml} byte_liftedLessEq = defaultLessEq
+let inline {ocaml} byte_liftedGreater = defaultGreater
+let inline {ocaml} byte_liftedGreaterEq = defaultGreaterEq
+
+instance (Ord byte_lifted)
+ let compare = byte_liftedCompare
+ let (<) = byte_liftedLess
+ let (<=) = byte_liftedLessEq
+ let (>) = byte_liftedGreater
+ let (>=) = byte_liftedGreaterEq
+end
+
+let ~{ocaml} byteCompare (Byte b1) (Byte b2) = compare b1 b2
+let inline {ocaml} byteCompare = defaultCompare
+
+let ~{ocaml} byteLess b1 b2 = byteCompare b1 b2 = LT
+let ~{ocaml} byteLessEq b1 b2 = byteCompare b1 b2 <> GT
+let ~{ocaml} byteGreater b1 b2 = byteCompare b1 b2 = GT
+let ~{ocaml} byteGreaterEq b1 b2 = byteCompare b1 b2 <> LT
+
+let inline {ocaml} byteLess = defaultLess
+let inline {ocaml} byteLessEq = defaultLessEq
+let inline {ocaml} byteGreater = defaultGreater
+let inline {ocaml} byteGreaterEq = defaultGreaterEq
+
+instance (Ord byte)
+ let compare = byteCompare
+ let (<) = byteLess
+ let (<=) = byteLessEq
+ let (>) = byteGreater
+ let (>=) = byteGreaterEq
+end
+
+
+
+
+
+let ~{ocaml} opcodeCompare (Opcode o1) (Opcode o2) =
+ compare o1 o2
+let {ocaml} opcodeCompare = defaultCompare
+
+let ~{ocaml} opcodeLess b1 b2 = opcodeCompare b1 b2 = LT
+let ~{ocaml} opcodeLessEq b1 b2 = opcodeCompare b1 b2 <> GT
+let ~{ocaml} opcodeGreater b1 b2 = opcodeCompare b1 b2 = GT
+let ~{ocaml} opcodeGreaterEq b1 b2 = opcodeCompare b1 b2 <> LT
+
+let inline {ocaml} opcodeLess = defaultLess
+let inline {ocaml} opcodeLessEq = defaultLessEq
+let inline {ocaml} opcodeGreater = defaultGreater
+let inline {ocaml} opcodeGreaterEq = defaultGreaterEq
+
+instance (Ord opcode)
+ let compare = opcodeCompare
+ let (<) = opcodeLess
+ let (<=) = opcodeLessEq
+ let (>) = opcodeGreater
+ let (>=) = opcodeGreaterEq
+end
+
+let addressCompare (Address b1 i1) (Address b2 i2) = compare i1 i2
+(* this cannot be defaultCompare for OCaml because addresses contain big ints *)
+
+let addressLess b1 b2 = addressCompare b1 b2 = LT
+let addressLessEq b1 b2 = addressCompare b1 b2 <> GT
+let addressGreater b1 b2 = addressCompare b1 b2 = GT
+let addressGreaterEq b1 b2 = addressCompare b1 b2 <> LT
+
+instance (SetType address)
+ let setElemCompare = addressCompare
+end
+
+instance (Ord address)
+ let compare = addressCompare
+ let (<) = addressLess
+ let (<=) = addressLessEq
+ let (>) = addressGreater
+ let (>=) = addressGreaterEq
+end
+
+let {coq; ocaml} addressEqual a1 a2 = (addressCompare a1 a2) = EQ
+let inline {hol; isabelle} addressEqual = unsafe_structural_equality
+
+let {coq; ocaml} addressInequal a1 a2 = not (addressEqual a1 a2)
+let inline {hol; isabelle} addressInequal = unsafe_structural_inequality
+
+instance (Eq address)
+ let (=) = addressEqual
+ let (<>) = addressInequal
+end
+
+let ~{ocaml} directionCompare d1 d2 =
+ match (d1, d2) with
+ | (D_decreasing, D_increasing) -> GT
+ | (D_increasing, D_decreasing) -> LT
+ | _ -> EQ
+ end
+let inline {ocaml} directionCompare = defaultCompare
+
+let ~{ocaml} directionLess b1 b2 = directionCompare b1 b2 = LT
+let ~{ocaml} directionLessEq b1 b2 = directionCompare b1 b2 <> GT
+let ~{ocaml} directionGreater b1 b2 = directionCompare b1 b2 = GT
+let ~{ocaml} directionGreaterEq b1 b2 = directionCompare b1 b2 <> LT
+
+let inline {ocaml} directionLess = defaultLess
+let inline {ocaml} directionLessEq = defaultLessEq
+let inline {ocaml} directionGreater = defaultGreater
+let inline {ocaml} directionGreaterEq = defaultGreaterEq
+
+instance (Ord direction)
+ let compare = directionCompare
+ let (<) = directionLess
+ let (<=) = directionLessEq
+ let (>) = directionGreater
+ let (>=) = directionGreaterEq
+end
+
+instance (Show direction)
+ let show = function D_increasing -> "D_increasing" | D_decreasing -> "D_decreasing" end
+end
+
+let ~{ocaml} register_valueCompare rv1 rv2 =
+ compare (rv1.rv_bits, rv1.rv_dir, rv1.rv_start, rv1.rv_start_internal)
+ (rv2.rv_bits, rv2.rv_dir, rv2.rv_start, rv2.rv_start_internal)
+let inline {ocaml} register_valueCompare = defaultCompare
+
+let ~{ocaml} register_valueLess b1 b2 = register_valueCompare b1 b2 = LT
+let ~{ocaml} register_valueLessEq b1 b2 = register_valueCompare b1 b2 <> GT
+let ~{ocaml} register_valueGreater b1 b2 = register_valueCompare b1 b2 = GT
+let ~{ocaml} register_valueGreaterEq b1 b2 = register_valueCompare b1 b2 <> LT
+
+let inline {ocaml} register_valueLess = defaultLess
+let inline {ocaml} register_valueLessEq = defaultLessEq
+let inline {ocaml} register_valueGreater = defaultGreater
+let inline {ocaml} register_valueGreaterEq = defaultGreaterEq
+
+instance (Ord register_value)
+ let compare = register_valueCompare
+ let (<) = register_valueLess
+ let (<=) = register_valueLessEq
+ let (>) = register_valueGreater
+ let (>=) = register_valueGreaterEq
+end
+
+let address_liftedCompare (Address_lifted b1 i1) (Address_lifted b2 i2) =
+ compare (i1,b1) (i2,b2)
+(* this cannot be defaultCompare for OCaml because address_lifteds contain big
+ ints *)
+
+let address_liftedLess b1 b2 = address_liftedCompare b1 b2 = LT
+let address_liftedLessEq b1 b2 = address_liftedCompare b1 b2 <> GT
+let address_liftedGreater b1 b2 = address_liftedCompare b1 b2 = GT
+let address_liftedGreaterEq b1 b2 = address_liftedCompare b1 b2 <> LT
+
+instance (Ord address_lifted)
+ let compare = address_liftedCompare
+ let (<) = address_liftedLess
+ let (<=) = address_liftedLessEq
+ let (>) = address_liftedGreater
+ let (>=) = address_liftedGreaterEq
+end
+
+(* Registers *)
+type slice = (nat * nat)
+
+type reg_name =
+ (* do we really need this here if ppcmem already has this information by itself? *)
+| Reg of string * nat * nat * direction
+(*Name of the register, accessing the entire register, the start and size of this register, and its direction *)
+
+| Reg_slice of string * nat * direction * slice
+(* Name of the register, accessing from the bit indexed by the first
+to the bit indexed by the second integer of the slice, inclusive. For
+machineDef* the first is a smaller number or equal to the second, adjusted
+to reflect the correct span direction in the interpreter side. *)
+
+| Reg_field of string * nat * direction * string * slice
+(*Name of the register, start and direction, and name of the field of the register
+accessed. The slice specifies where this field is in the register*)
+
+| Reg_f_slice of string * nat * direction * string * slice * slice
+(* The first four components are as in Reg_field; the final slice
+specifies a part of the field, indexed w.r.t. the register as a whole *)
+
+let register_base_name : reg_name -> string = function
+ | Reg s _ _ _ -> s
+ | Reg_slice s _ _ _ -> s
+ | Reg_field s _ _ _ _ -> s
+ | Reg_f_slice s _ _ _ _ _ -> s
+ end
+
+let slice_of_reg_name : reg_name -> slice = function
+ | Reg _ start width D_increasing -> (start, start + width -1)
+ | Reg _ start width D_decreasing -> (start - width - 1, start)
+ | Reg_slice _ _ _ sl -> sl
+ | Reg_field _ _ _ _ sl -> sl
+ | Reg_f_slice _ _ _ _ _ sl -> sl
+ end
+
+let width_of_reg_name (r: reg_name) : nat =
+ let width_of_slice (i, j) = (* j - i + 1 in *)
+
+ (integerFromNat j) - (integerFromNat i) + 1
+ $> abs $> natFromInteger
+ in
+ match r with
+ | Reg _ _ width _ -> width
+ | Reg_slice _ _ _ sl -> width_of_slice sl
+ | Reg_field _ _ _ _ sl -> width_of_slice sl
+ | Reg_f_slice _ _ _ _ _ sl -> width_of_slice sl
+ end
+
+let reg_name_non_empty_intersection (r: reg_name) (r': reg_name) : bool =
+ register_base_name r = register_base_name r' &&
+ let (i1, i2) = slice_of_reg_name r in
+ let (i1', i2') = slice_of_reg_name r' in
+ i1' <= i2 && i2' >= i1
+
+let reg_nameCompare r1 r2 =
+ compare (register_base_name r1,slice_of_reg_name r1)
+ (register_base_name r2,slice_of_reg_name r2)
+
+let reg_nameLess b1 b2 = reg_nameCompare b1 b2 = LT
+let reg_nameLessEq b1 b2 = reg_nameCompare b1 b2 <> GT
+let reg_nameGreater b1 b2 = reg_nameCompare b1 b2 = GT
+let reg_nameGreaterEq b1 b2 = reg_nameCompare b1 b2 <> LT
+
+instance (Ord reg_name)
+ let compare = reg_nameCompare
+ let (<) = reg_nameLess
+ let (<=) = reg_nameLessEq
+ let (>) = reg_nameGreater
+ let (>=) = reg_nameGreaterEq
+end
+
+let {coq;ocaml} reg_nameEqual a1 a2 = (reg_nameCompare a1 a2) = EQ
+let {hol;isabelle} reg_nameEqual = unsafe_structural_equality
+let {coq;ocaml} reg_nameInequal a1 a2 = not (reg_nameEqual a1 a2)
+let {hol;isabelle} reg_nameInequal = unsafe_structural_inequality
+
+instance (Eq reg_name)
+ let (=) = reg_nameEqual
+ let (<>) = reg_nameInequal
+end
+
+instance (SetType reg_name)
+ let setElemCompare = reg_nameCompare
+end
+
+let direction_of_reg_name r = match r with
+ | Reg _ _ _ d -> d
+ | Reg_slice _ _ d _ -> d
+ | Reg_field _ _ d _ _ -> d
+ | Reg_f_slice _ _ d _ _ _ -> d
+ end
+
+let start_of_reg_name r = match r with
+ | Reg _ start _ _ -> start
+ | Reg_slice _ start _ _ -> start
+ | Reg_field _ start _ _ _ -> start
+ | Reg_f_slice _ start _ _ _ _ -> start
+end
+
+(* Data structures for building up instructions *)
+
+(* read_kind, write_kind, barrier_kind, trans_kind and instruction_kind have
+ been moved to sail_instr_kinds.lem. This removes the dependency of the
+ shallow embedding on the rest of sail_impl_base.lem, and helps avoid name
+ clashes between the different monad types. *)
+
+type event =
+ | E_read_mem of read_kind * address_lifted * nat * maybe (list reg_name)
+ | E_read_memt of read_kind * address_lifted * nat * maybe (list reg_name)
+ | E_write_mem of write_kind * address_lifted * nat * maybe (list reg_name) * memory_value * maybe (list reg_name)
+ | E_write_ea of write_kind * address_lifted * nat * maybe (list reg_name)
+ | E_excl_res
+ | E_write_memv of maybe address_lifted * memory_value * maybe (list reg_name)
+ | E_write_memvt of maybe address_lifted * (bit_lifted * memory_value) * maybe (list reg_name)
+ | E_barrier of barrier_kind
+ | E_footprint
+ | E_read_reg of reg_name
+ | E_write_reg of reg_name * register_value
+ | E_escape
+ | E_error of string
+
+
+let eventCompare e1 e2 =
+ match (e1,e2) with
+ | (E_read_mem rk1 v1 i1 tr1, E_read_mem rk2 v2 i2 tr2) ->
+ compare (rk1, (v1,i1,tr1)) (rk2,(v2, i2, tr2))
+ | (E_read_memt rk1 v1 i1 tr1, E_read_memt rk2 v2 i2 tr2) ->
+ compare (rk1, (v1,i1,tr1)) (rk2,(v2, i2, tr2))
+ | (E_write_mem wk1 v1 i1 tr1 v1' tr1', E_write_mem wk2 v2 i2 tr2 v2' tr2') ->
+ compare ((wk1,v1,i1),(tr1,v1',tr1')) ((wk2,v2,i2),(tr2,v2',tr2'))
+ | (E_write_ea wk1 a1 i1 tr1, E_write_ea wk2 a2 i2 tr2) ->
+ compare (wk1, (a1, i1, tr1)) (wk2, (a2, i2, tr2))
+ | (E_excl_res, E_excl_res) -> EQ
+ | (E_write_memv _ mv1 tr1, E_write_memv _ mv2 tr2) -> compare (mv1,tr1) (mv2,tr2)
+ | (E_write_memvt _ mv1 tr1, E_write_memvt _ mv2 tr2) -> compare (mv1,tr1) (mv2,tr2)
+ | (E_barrier bk1, E_barrier bk2) -> compare bk1 bk2
+ | (E_read_reg r1, E_read_reg r2) -> compare r1 r2
+ | (E_write_reg r1 v1, E_write_reg r2 v2) -> compare (r1,v1) (r2,v2)
+ | (E_error s1, E_error s2) -> compare s1 s2
+ | (E_escape,E_escape) -> EQ
+ | (E_read_mem _ _ _ _, _) -> LT
+ | (E_write_mem _ _ _ _ _ _, _) -> LT
+ | (E_write_ea _ _ _ _, _) -> LT
+ | (E_excl_res, _) -> LT
+ | (E_write_memv _ _ _, _) -> LT
+ | (E_barrier _, _) -> LT
+ | (E_read_reg _, _) -> LT
+ | (E_write_reg _ _, _) -> LT
+ | _ -> GT
+ end
+
+let eventLess b1 b2 = eventCompare b1 b2 = LT
+let eventLessEq b1 b2 = eventCompare b1 b2 <> GT
+let eventGreater b1 b2 = eventCompare b1 b2 = GT
+let eventGreaterEq b1 b2 = eventCompare b1 b2 <> LT
+
+instance (Ord event)
+ let compare = eventCompare
+ let (<) = eventLess
+ let (<=) = eventLessEq
+ let (>) = eventGreater
+ let (>=) = eventGreaterEq
+end
+
+instance (SetType event)
+ let setElemCompare = compare
+end
+
+
+(* the address_lifted types should go away here and be replaced by address *)
+type with_aux 'o = 'o * maybe ((unit -> (string * string)) * ((list (reg_name * register_value)) -> list event))
+type outcome 'a 'e =
+ (* Request to read memory, value is location to read, integer is size to read,
+ followed by registers that were used in computing that size *)
+ | Read_mem of (read_kind * address_lifted * nat) * (memory_value -> with_aux (outcome 'a 'e))
+ (* Tell the system a write is imminent, at address lifted, of size nat *)
+ | Write_ea of (write_kind * address_lifted * nat) * (with_aux (outcome 'a 'e))
+ (* Request the result of store-exclusive *)
+ | Excl_res of (bool -> with_aux (outcome 'a 'e))
+ (* Request to write memory at last signalled address. Memory value should be 8
+ times the size given in ea signal *)
+ | Write_memv of memory_value * (bool -> with_aux (outcome 'a 'e))
+ (* Request a memory barrier *)
+ | Barrier of barrier_kind * with_aux (outcome 'a 'e)
+ (* Tell the system to dynamically recalculate dependency footprint *)
+ | Footprint of with_aux (outcome 'a 'e)
+ (* Request to read register, will track dependency when mode.track_values *)
+ | Read_reg of reg_name * (register_value -> with_aux (outcome 'a 'e))
+ (* Request to write register *)
+ | Write_reg of (reg_name * register_value) * with_aux (outcome 'a 'e)
+ | Escape of maybe string
+ (*Result of a failed assert with possible error message to report*)
+ | Fail of maybe string
+ (* Exception of type 'e *)
+ | Exception of 'e
+ | Internal of (maybe string * maybe (unit -> string)) * with_aux (outcome 'a 'e)
+ | Done of 'a
+ | Error of string
+
+type outcome_s 'a 'e = with_aux (outcome 'a 'e)
+(* first string : output of instruction_stack_to_string
+ second string: output of local_variables_to_string *)
+
+(** operations and coercions on basic values *)
+
+val word8_to_bitls : word8 -> list bit_lifted
+val bitls_to_word8 : list bit_lifted -> word8
+
+val integer_of_word8_list : list word8 -> integer
+val word8_list_of_integer : integer -> integer -> list word8
+
+val concretizable_bitl : bit_lifted -> bool
+val concretizable_bytl : byte_lifted -> bool
+val concretizable_bytls : list byte_lifted -> bool
+
+let concretizable_bitl = function
+ | Bitl_zero -> true
+ | Bitl_one -> true
+ | Bitl_undef -> false
+ | Bitl_unknown -> false
+end
+
+let concretizable_bytl (Byte_lifted bs) = List.all concretizable_bitl bs
+let concretizable_bytls = List.all concretizable_bytl
+
+(* constructing values *)
+
+val build_register_value : list bit_lifted -> direction -> nat -> nat -> register_value
+let build_register_value bs dir width start_index =
+ <| rv_bits = bs;
+ rv_dir = dir; (* D_increasing for Power, D_decreasing for ARM *)
+ rv_start_internal = start_index;
+ rv_start = if dir = D_increasing
+ then start_index
+ else (start_index+1) - width; (* Smaller index, as in Power, for external interaction *)
+ |>
+
+val register_value : bit_lifted -> direction -> nat -> nat -> register_value
+let register_value b dir width start_index =
+ build_register_value (List.replicate width b) dir width start_index
+
+val register_value_zeros : direction -> nat -> nat -> register_value
+let register_value_zeros dir width start_index =
+ register_value Bitl_zero dir width start_index
+
+val register_value_ones : direction -> nat -> nat -> register_value
+let register_value_ones dir width start_index =
+ register_value Bitl_one dir width start_index
+
+val register_value_for_reg : reg_name -> list bit_lifted -> register_value
+let register_value_for_reg r bs : register_value =
+ let () = ensure (width_of_reg_name r = List.length bs)
+ ("register_value_for_reg (\"" ^ show (register_base_name r) ^ "\") length mismatch: "
+ ^ show (width_of_reg_name r) ^ " vs " ^ show (List.length bs))
+ in
+ let (j1, j2) = slice_of_reg_name r in
+ let d = direction_of_reg_name r in
+ <| rv_bits = bs;
+ rv_dir = d;
+ rv_start_internal = if d = D_increasing then j1 else (start_of_reg_name r) - j1;
+ rv_start = j1;
+ |>
+
+val byte_lifted_undef : byte_lifted
+let byte_lifted_undef = Byte_lifted (List.replicate 8 Bitl_undef)
+
+val byte_lifted_unknown : byte_lifted
+let byte_lifted_unknown = Byte_lifted (List.replicate 8 Bitl_unknown)
+
+val memory_value_unknown : nat (*the number of bytes*) -> memory_value
+let memory_value_unknown (width:nat) : memory_value =
+ List.replicate width byte_lifted_unknown
+
+val memory_value_undef : nat (*the number of bytes*) -> memory_value
+let memory_value_undef (width:nat) : memory_value =
+ List.replicate width byte_lifted_undef
+
+val match_endianness : forall 'a. end_flag -> list 'a -> list 'a
+let match_endianness endian l =
+ match endian with
+ | E_little_endian -> List.reverse l
+ | E_big_endian -> l
+ end
+
+(* lengths *)
+
+val memory_value_length : memory_value -> nat
+let memory_value_length (mv:memory_value) = List.length mv
+
+
+(* aux fns *)
+
+val maybe_all : forall 'a. list (maybe 'a) -> maybe (list 'a)
+let rec maybe_all' xs acc =
+ match xs with
+ | [] -> Just (List.reverse acc)
+ | Nothing :: _ -> Nothing
+ | (Just y)::xs' -> maybe_all' xs' (y::acc)
+ end
+let maybe_all xs = maybe_all' xs []
+
+(** coercions *)
+
+(* bits and bytes *)
+
+let bit_to_bool = function (* TODO: rename bool_of_bit *)
+ | Bitc_zero -> false
+ | Bitc_one -> true
+end
+
+
+val bit_lifted_of_bit : bit -> bit_lifted
+let bit_lifted_of_bit b =
+ match b with
+ | Bitc_zero -> Bitl_zero
+ | Bitc_one -> Bitl_one
+ end
+
+val bit_of_bit_lifted : bit_lifted -> maybe bit
+let bit_of_bit_lifted bl =
+ match bl with
+ | Bitl_zero -> Just Bitc_zero
+ | Bitl_one -> Just Bitc_one
+ | Bitl_undef -> Nothing
+ | Bitl_unknown -> Nothing
+ end
+
+
+val byte_lifted_of_byte : byte -> byte_lifted
+let byte_lifted_of_byte (Byte bs) : byte_lifted = Byte_lifted (List.map bit_lifted_of_bit bs)
+
+val byte_of_byte_lifted : byte_lifted -> maybe byte
+let byte_of_byte_lifted bl =
+ match bl with
+ | Byte_lifted bls ->
+ match maybe_all (List.map bit_of_bit_lifted bls) with
+ | Nothing -> Nothing
+ | Just bs -> Just (Byte bs)
+ end
+ end
+
+
+val bytes_of_bits : list bit -> list byte (*assumes (length bits) mod 8 = 0*)
+let rec bytes_of_bits bits = match bits with
+ | [] -> []
+ | b0::b1::b2::b3::b4::b5::b6::b7::bits ->
+ (Byte [b0;b1;b2;b3;b4;b5;b6;b7])::(bytes_of_bits bits)
+ | _ -> failwith "bytes_of_bits not given bits divisible by 8"
+end
+
+val byte_lifteds_of_bit_lifteds : list bit_lifted -> list byte_lifted (*assumes (length bits) mod 8 = 0*)
+let rec byte_lifteds_of_bit_lifteds bits = match bits with
+ | [] -> []
+ | b0::b1::b2::b3::b4::b5::b6::b7::bits ->
+ (Byte_lifted [b0;b1;b2;b3;b4;b5;b6;b7])::(byte_lifteds_of_bit_lifteds bits)
+ | _ -> failwith "byte_lifteds of bit_lifteds not given bits divisible by 8"
+end
+
+
+val byte_of_memory_byte : memory_byte -> maybe byte
+let byte_of_memory_byte = byte_of_byte_lifted
+
+val memory_byte_of_byte : byte -> memory_byte
+let memory_byte_of_byte = byte_lifted_of_byte
+
+
+(* to and from nat *)
+
+(* this natFromBoolList could move to the Lem word.lem library *)
+val natFromBoolList : list bool -> nat
+let rec natFromBoolListAux (acc : nat) (bl : list bool) =
+ match bl with
+ | [] -> acc
+ | (true :: bl') -> natFromBoolListAux ((acc * 2) + 1) bl'
+ | (false :: bl') -> natFromBoolListAux (acc * 2) bl'
+ end
+let natFromBoolList bl =
+ natFromBoolListAux 0 (List.reverse bl)
+
+
+val nat_of_bit_list : list bit -> nat
+let nat_of_bit_list b =
+ natFromBoolList (List.reverse (List.map bit_to_bool b))
+ (* natFromBoolList takes a list with LSB first, for consistency with rest of Lem word library, so we reverse it. twice. *)
+
+
+(* to and from integer *)
+
+val integer_of_bit_list : list bit -> integer
+let integer_of_bit_list b =
+ integerFromBoolList (false,(List.reverse (List.map bit_to_bool b)))
+ (* integerFromBoolList takes a list with LSB first, so we reverse it *)
+
+val bit_list_of_integer : nat -> integer -> list bit
+let bit_list_of_integer len b =
+ List.map (fun b -> if b then Bitc_one else Bitc_zero)
+ (reverse (boolListFrombitSeq len (bitSeqFromInteger Nothing b)))
+
+val integer_of_byte_list : list byte -> integer
+let integer_of_byte_list bytes = integer_of_bit_list (List.concatMap (fun (Byte bs) -> bs) bytes)
+
+val byte_list_of_integer : nat -> integer -> list byte
+let byte_list_of_integer (len:nat) (a:integer):list byte =
+ let bits = bit_list_of_integer (len * 8) a in bytes_of_bits bits
+
+
+val integer_of_address : address -> integer
+let integer_of_address (a:address):integer =
+ match a with
+ | Address bs i -> i
+ end
+
+val address_of_integer : integer -> address
+let address_of_integer (i:integer):address =
+ Address (byte_list_of_integer 8 i) i
+
+(* to and from signed-integer *)
+
+val signed_integer_of_bit_list : list bit -> integer
+let signed_integer_of_bit_list b =
+ match b with
+ | [] -> failwith "empty bit list"
+ | Bitc_zero :: b' ->
+ integerFromBoolList (false,(List.reverse (List.map bit_to_bool b)))
+ | Bitc_one :: b' ->
+ let b'_val = integerFromBoolList (false,(List.reverse (List.map bit_to_bool b'))) in
+ (* integerFromBoolList takes a list with LSB first, so we reverse it *)
+ let msb_val = integerPow 2 ((List.length b) - 1) in
+ b'_val - msb_val
+ end
+
+
+(* regarding a list of int as a list of bytes in memory, MSB lowest-address first, convert to an integer *)
+val integer_address_of_int_list : list int -> integer
+let rec integerFromIntListAux (acc: integer) (is: list int) =
+ match is with
+ | [] -> acc
+ | (i :: is') -> integerFromIntListAux ((acc * 256) + integerFromInt i) is'
+ end
+let integer_address_of_int_list (is: list int) =
+ integerFromIntListAux 0 is
+
+val address_of_byte_list : list byte -> address
+let address_of_byte_list bs =
+ if List.length bs <> 8 then failwith "address_of_byte_list given list not of length 8" else
+ Address bs (integer_of_byte_list bs)
+
+let address_of_byte_lifted_list bls =
+ match maybe_all (List.map byte_of_byte_lifted bls) with
+ | Nothing -> Nothing
+ | Just bs -> Just (address_of_byte_list bs)
+ end
+
+(* operations on addresses *)
+
+val add_address_nat : address -> nat -> address
+let add_address_nat (a:address) (i:nat) : address =
+ address_of_integer ((integer_of_address a) + (integerFromNat i))
+
+val clear_low_order_bits_of_address : address -> address
+let clear_low_order_bits_of_address a =
+ match a with
+ | Address [b0;b1;b2;b3;b4;b5;b6;b7] i ->
+ match b7 with
+ | Byte [bt0;bt1;bt2;bt3;bt4;bt5;bt6;bt7] ->
+ let b7' = Byte [bt0;bt1;bt2;bt3;bt4;bt5;Bitc_zero;Bitc_zero] in
+ let bytes = [b0;b1;b2;b3;b4;b5;b6;b7'] in
+ Address bytes (integer_of_byte_list bytes)
+ | _ -> failwith "Byte does not contain 8 bits"
+ end
+ | _ -> failwith "Address does not contain 8 bytes"
+ end
+
+
+
+val byte_list_of_memory_value : end_flag -> memory_value -> maybe (list byte)
+let byte_list_of_memory_value endian mv =
+ match_endianness endian mv
+ $> List.map byte_of_memory_byte
+ $> maybe_all
+
+
+val integer_of_memory_value : end_flag -> memory_value -> maybe integer
+let integer_of_memory_value endian (mv:memory_value):maybe integer =
+ match byte_list_of_memory_value endian mv with
+ | Just bs -> Just (integer_of_byte_list bs)
+ | Nothing -> Nothing
+ end
+
+val memory_value_of_integer : end_flag -> nat -> integer -> memory_value
+let memory_value_of_integer endian (len:nat) (i:integer):memory_value =
+ List.map byte_lifted_of_byte (byte_list_of_integer len i)
+ $> match_endianness endian
+
+
+val integer_of_register_value : register_value -> maybe integer
+let integer_of_register_value (rv:register_value):maybe integer =
+ match maybe_all (List.map bit_of_bit_lifted rv.rv_bits) with
+ | Nothing -> Nothing
+ | Just bs -> Just (integer_of_bit_list bs)
+ end
+
+(* NOTE: register_value_for_reg_of_integer might be easier to use *)
+val register_value_of_integer : nat -> nat -> direction -> integer -> register_value
+let register_value_of_integer (len:nat) (start:nat) (dir:direction) (i:integer):register_value =
+ let bs = bit_list_of_integer len i in
+ build_register_value (List.map bit_lifted_of_bit bs) dir len start
+
+val register_value_for_reg_of_integer : reg_name -> integer -> register_value
+let register_value_for_reg_of_integer (r: reg_name) (i:integer) : register_value =
+ register_value_of_integer (width_of_reg_name r) (start_of_reg_name r) (direction_of_reg_name r) i
+
+(* *)
+
+val opcode_of_bytes : byte -> byte -> byte -> byte -> opcode
+let opcode_of_bytes b0 b1 b2 b3 : opcode = Opcode [b0;b1;b2;b3]
+
+val register_value_of_address : address -> direction -> register_value
+let register_value_of_address (Address bytes _) dir : register_value =
+ let bits = List.concatMap (fun (Byte bs) -> List.map bit_lifted_of_bit bs) bytes in
+ <| rv_bits = bits;
+ rv_dir = dir;
+ rv_start = 0;
+ rv_start_internal = if dir = D_increasing then 0 else (List.length bits) - 1
+ |>
+
+val register_value_of_memory_value : memory_value -> direction -> register_value
+let register_value_of_memory_value bytes dir : register_value =
+ let bitls = List.concatMap (fun (Byte_lifted bs) -> bs) bytes in
+ <| rv_bits = bitls;
+ rv_dir = dir;
+ rv_start = 0;
+ rv_start_internal = if dir = D_increasing then 0 else (List.length bitls) - 1
+ |>
+
+val memory_value_of_register_value: register_value -> memory_value
+let memory_value_of_register_value r =
+ (byte_lifteds_of_bit_lifteds r.rv_bits)
+
+val address_lifted_of_register_value : register_value -> maybe address_lifted
+(* returning Nothing iff the register value is not 64 bits wide, but
+allowing Bitl_undef and Bitl_unknown *)
+let address_lifted_of_register_value (rv:register_value) : maybe address_lifted =
+ if List.length rv.rv_bits <> 64 then Nothing
+ else
+ Just (Address_lifted (byte_lifteds_of_bit_lifteds rv.rv_bits)
+ (if List.all concretizable_bitl rv.rv_bits
+ then match (maybe_all (List.map bit_of_bit_lifted rv.rv_bits)) with
+ | (Just(bits)) -> Just (integer_of_bit_list bits)
+ | Nothing -> Nothing end
+ else Nothing))
+
+val address_of_address_lifted : address_lifted -> maybe address
+(* returning Nothing iff the address contains any Bitl_undef or Bitl_unknown *)
+let address_of_address_lifted (al:address_lifted): maybe address =
+ match al with
+ | Address_lifted bls (Just i)->
+ match maybe_all ((List.map byte_of_byte_lifted) bls) with
+ | Nothing -> Nothing
+ | Just bs -> Just (Address bs i)
+ end
+ | _ -> Nothing
+end
+
+val address_of_register_value : register_value -> maybe address
+(* returning Nothing iff the register value is not 64 bits wide, or contains Bitl_undef or Bitl_unknown *)
+let address_of_register_value (rv:register_value) : maybe address =
+ match address_lifted_of_register_value rv with
+ | Nothing -> Nothing
+ | Just al ->
+ match address_of_address_lifted al with
+ | Nothing -> Nothing
+ | Just a -> Just a
+ end
+ end
+
+let address_of_memory_value (endian: end_flag) (mv:memory_value) : maybe address =
+ match byte_list_of_memory_value endian mv with
+ | Nothing -> Nothing
+ | Just bs ->
+ if List.length bs <> 8 then Nothing else
+ Just (address_of_byte_list bs)
+ end
+
+val byte_of_int : int -> byte
+let byte_of_int (i:int) : byte =
+ Byte (bit_list_of_integer 8 (integerFromInt i))
+
+val memory_byte_of_int : int -> memory_byte
+let memory_byte_of_int (i:int) : memory_byte =
+ memory_byte_of_byte (byte_of_int i)
+
+(*
+val int_of_memory_byte : int -> maybe memory_byte
+let int_of_memory_byte (mb:memory_byte) : int =
+ failwith "TODO"
+*)
+
+
+
+val memory_value_of_address_lifted : end_flag -> address_lifted -> memory_value
+let memory_value_of_address_lifted endian (Address_lifted bs _ :address_lifted) =
+ match_endianness endian bs
+
+val byte_list_of_address : address -> list byte
+let byte_list_of_address (Address bs _) : list byte = bs
+
+val memory_value_of_address : end_flag -> address -> memory_value
+let memory_value_of_address endian (Address bs _) =
+ match_endianness endian bs
+ $> List.map byte_lifted_of_byte
+
+val byte_list_of_opcode : opcode -> list byte
+let byte_list_of_opcode (Opcode bs) : list byte = bs
+
+(** ****************************************** *)
+(** show type class instantiations *)
+(** ****************************************** *)
+
+(* matching printing_functions.ml *)
+val stringFromReg_name : reg_name -> string
+let stringFromReg_name r =
+ let norm_sl start dir (first,second) = (first,second)
+ (* match dir with
+ | D_increasing -> (first,second)
+ | D_decreasing -> (start - first, start - second)
+ end *)
+ in
+ match r with
+ | Reg s start size dir -> s
+ | Reg_slice s start dir sl ->
+ let (first,second) = norm_sl start dir sl in
+ s ^ "[" ^ show first ^ (if (first = second) then "" else ".." ^ (show second)) ^ "]"
+ | Reg_field s start dir f sl ->
+ let (first,second) = norm_sl start dir sl in
+ s ^ "." ^ f ^ " (" ^ (show start) ^ ", " ^ (show dir) ^ ", " ^ (show first) ^ ", " ^ (show second) ^ ")"
+ | Reg_f_slice s start dir f (first1,second1) (first,second) ->
+ let (first,second) =
+ match dir with
+ | D_increasing -> (first,second)
+ | D_decreasing -> (start - first, start - second)
+ end in
+ s ^ "." ^ f ^ "]" ^ show first ^ (if (first = second) then "" else ".." ^ (show second)) ^ "]"
+ end
+
+instance (Show reg_name)
+ let show = stringFromReg_name
+end
+
+
+(* hex pp of integers, adapting the Lem string_extra.lem code *)
+val stringFromNaturalHexHelper : natural -> list char -> list char
+let rec stringFromNaturalHexHelper n acc =
+ if n = 0 then
+ acc
+ else
+ stringFromNaturalHexHelper (n / 16) (String_extra.chr (natFromNatural (let nd = n mod 16 in if nd <=9 then nd + 48 else nd - 10 + 97)) :: acc)
+
+val stringFromNaturalHex : natural -> string
+let (*~{ocaml;hol}*) stringFromNaturalHex n =
+ if n = 0 then "0" else toString (stringFromNaturalHexHelper n [])
+
+val stringFromIntegerHex : integer -> string
+let (*~{ocaml}*) stringFromIntegerHex i =
+ if i < 0 then
+ "-" ^ stringFromNaturalHex (naturalFromInteger i)
+ else
+ stringFromNaturalHex (naturalFromInteger i)
+
+
+let stringFromAddress (Address bs i) =
+ let i' = integer_of_byte_list bs in
+ if i=i' then
+(*TODO: ideally this should be made to match the src/pp.ml pp_address; the following very roughly matches what's used in the ppcmem UI, enough to make exceptions readable *)
+ if i < 65535 then
+ show i
+ else
+ stringFromIntegerHex i
+ else
+ "stringFromAddress bytes and integer mismatch"
+
+instance (Show address)
+ let show = stringFromAddress
+end
+
+let stringFromByte_lifted bl =
+ match byte_of_byte_lifted bl with
+ | Nothing -> "u?"
+ | Just (Byte bits) ->
+ let i = integer_of_bit_list bits in
+ show i
+ end
+
+instance (Show byte_lifted)
+ let show = stringFromByte_lifted
+end
+
+(* possible next instruction address options *)
+type nia =
+ | NIA_successor
+ | NIA_concrete_address of address
+ | NIA_indirect_address
+
+let niaCompare n1 n2 = match (n1,n2) with
+ | (NIA_successor, NIA_successor) -> EQ
+ | (NIA_successor, _) -> LT
+ | (_, NIA_successor) -> GT
+ | (NIA_concrete_address a1, NIA_concrete_address a2) -> compare a1 a2
+ | (NIA_concrete_address _, _) -> LT
+ | (_, NIA_concrete_address _) -> GT
+ | (NIA_indirect_address, NIA_indirect_address) -> EQ
+ (* | (NIA_indirect_address, _) -> LT
+ | (_, NIA_indirect_address) -> GT *)
+ end
+
+instance (Ord nia)
+ let compare = niaCompare
+ let (<) n1 n2 = (niaCompare n1 n2) = LT
+ let (<=) n1 n2 = (niaCompare n1 n2) <> GT
+ let (>) n1 n2 = (niaCompare n1 n2) = GT
+ let (>=) n1 n2 = (niaCompare n1 n2) <> LT
+end
+
+let stringFromNia = function
+ | NIA_successor -> "NIA_successor"
+ | NIA_concrete_address a -> "NIA_concrete_address " ^ show a
+ | NIA_indirect_address -> "NIA_indirect_address"
+end
+
+instance (Show nia)
+ let show = stringFromNia
+end
+
+type dia =
+ | DIA_none
+ | DIA_concrete_address of address
+ | DIA_register of reg_name
+
+let diaCompare d1 d2 = match (d1, d2) with
+ | (DIA_none, DIA_none) -> EQ
+ | (DIA_none, _) -> LT
+ | (DIA_concrete_address a1, DIA_none) -> GT
+ | (DIA_concrete_address a1, DIA_concrete_address a2) -> compare a1 a2
+ | (DIA_concrete_address a1, _) -> LT
+ | (DIA_register r1, DIA_register r2) -> compare r1 r2
+ | (DIA_register _, _) -> GT
+end
+
+instance (Ord dia)
+ let compare = diaCompare
+ let (<) n1 n2 = (diaCompare n1 n2) = LT
+ let (<=) n1 n2 = (diaCompare n1 n2) <> GT
+ let (>) n1 n2 = (diaCompare n1 n2) = GT
+ let (>=) n1 n2 = (diaCompare n1 n2) <> LT
+end
+
+let stringFromDia = function
+ | DIA_none -> "DIA_none"
+ | DIA_concrete_address a -> "DIA_concrete_address " ^ show a
+ | DIA_register r -> "DIA_delayed_register " ^ show r
+end
+
+instance (Show dia)
+ let show = stringFromDia
+end
+*)
diff --git a/prover_snapshots/coq/lib/sail/Sail2_instr_kinds.v b/prover_snapshots/coq/lib/sail/Sail2_instr_kinds.v
new file mode 100644
index 0000000..d03d5e6
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_instr_kinds.v
@@ -0,0 +1,332 @@
+(*========================================================================*)
+(* Sail *)
+(* *)
+(* Copyright (c) 2013-2017 *)
+(* Kathyrn Gray *)
+(* Shaked Flur *)
+(* Stephen Kell *)
+(* Gabriel Kerneis *)
+(* Robert Norton-Wright *)
+(* Christopher Pulte *)
+(* Peter Sewell *)
+(* Alasdair Armstrong *)
+(* Brian Campbell *)
+(* Thomas Bauereiss *)
+(* Anthony Fox *)
+(* Jon French *)
+(* Dominic Mulligan *)
+(* Stephen Kell *)
+(* Mark Wassell *)
+(* *)
+(* All rights reserved. *)
+(* *)
+(* This software was developed by the University of Cambridge Computer *)
+(* Laboratory as part of the Rigorous Engineering of Mainstream Systems *)
+(* (REMS) project, funded by EPSRC grant EP/K008528/1. *)
+(* *)
+(* Redistribution and use in source and binary forms, with or without *)
+(* modification, are permitted provided that the following conditions *)
+(* are met: *)
+(* 1. Redistributions of source code must retain the above copyright *)
+(* notice, this list of conditions and the following disclaimer. *)
+(* 2. Redistributions in binary form must reproduce the above copyright *)
+(* notice, this list of conditions and the following disclaimer in *)
+(* the documentation and/or other materials provided with the *)
+(* distribution. *)
+(* *)
+(* THIS SOFTWARE IS PROVIDED BY THE AUTHOR AND CONTRIBUTORS ``AS IS'' *)
+(* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED *)
+(* TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A *)
+(* PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE AUTHOR OR *)
+(* CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, *)
+(* SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT *)
+(* LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF *)
+(* USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND *)
+(* ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, *)
+(* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT *)
+(* OF THE USE OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF *)
+(* SUCH DAMAGE. *)
+(*========================================================================*)
+
+Require Import DecidableClass.
+
+Class EnumerationType (A : Type) := {
+ toNat : A -> nat
+}.
+
+(*
+val enumeration_typeCompare : forall 'a. EnumerationType 'a => 'a -> 'a -> ordering
+let ~{ocaml} enumeration_typeCompare e1 e2 :=
+ compare (toNat e1) (toNat e2)
+let inline {ocaml} enumeration_typeCompare := defaultCompare
+
+
+default_instance forall 'a. EnumerationType 'a => (Ord 'a)
+ let compare := enumeration_typeCompare
+ let (<) r1 r2 := (enumeration_typeCompare r1 r2) = LT
+ let (<=) r1 r2 := (enumeration_typeCompare r1 r2) <> GT
+ let (>) r1 r2 := (enumeration_typeCompare r1 r2) = GT
+ let (>=) r1 r2 := (enumeration_typeCompare r1 r2) <> LT
+end
+*)
+
+(* Data structures for building up instructions *)
+
+(* careful: changes in the read/write/barrier kinds have to be
+ reflected in deep_shallow_convert *)
+Inductive read_kind :=
+ (* common reads *)
+ | Read_plain
+ (* Power reads *)
+ | Read_reserve
+ (* AArch64 reads *)
+ | Read_acquire | Read_exclusive | Read_exclusive_acquire | Read_stream
+ (* RISC-V reads *)
+ | Read_RISCV_acquire | Read_RISCV_strong_acquire
+ | Read_RISCV_reserved | Read_RISCV_reserved_acquire
+ | Read_RISCV_reserved_strong_acquire
+ (* x86 reads *)
+ | Read_X86_locked (* the read part of a lock'd instruction (rmw) *)
+.
+Scheme Equality for read_kind.
+(*
+instance (Show read_kind)
+ let show := function
+ | Read_plain -> "Read_plain"
+ | Read_reserve -> "Read_reserve"
+ | Read_acquire -> "Read_acquire"
+ | Read_exclusive -> "Read_exclusive"
+ | Read_exclusive_acquire -> "Read_exclusive_acquire"
+ | Read_stream -> "Read_stream"
+ | Read_RISCV_acquire -> "Read_RISCV_acquire"
+ | Read_RISCV_strong_acquire -> "Read_RISCV_strong_acquire"
+ | Read_RISCV_reserved -> "Read_RISCV_reserved"
+ | Read_RISCV_reserved_acquire -> "Read_RISCV_reserved_acquire"
+ | Read_RISCV_reserved_strong_acquire -> "Read_RISCV_reserved_strong_acquire"
+ | Read_X86_locked -> "Read_X86_locked"
+ end
+end
+*)
+Inductive write_kind :=
+ (* common writes *)
+ | Write_plain
+ (* Power writes *)
+ | Write_conditional
+ (* AArch64 writes *)
+ | Write_release | Write_exclusive | Write_exclusive_release
+ (* RISC-V *)
+ | Write_RISCV_release | Write_RISCV_strong_release
+ | Write_RISCV_conditional | Write_RISCV_conditional_release
+ | Write_RISCV_conditional_strong_release
+ (* x86 writes *)
+ | Write_X86_locked (* the write part of a lock'd instruction (rmw) *)
+.
+Scheme Equality for write_kind.
+(*
+instance (Show write_kind)
+ let show := function
+ | Write_plain -> "Write_plain"
+ | Write_conditional -> "Write_conditional"
+ | Write_release -> "Write_release"
+ | Write_exclusive -> "Write_exclusive"
+ | Write_exclusive_release -> "Write_exclusive_release"
+ | Write_RISCV_release -> "Write_RISCV_release"
+ | Write_RISCV_strong_release -> "Write_RISCV_strong_release"
+ | Write_RISCV_conditional -> "Write_RISCV_conditional"
+ | Write_RISCV_conditional_release -> "Write_RISCV_conditional_release"
+ | Write_RISCV_conditional_strong_release -> "Write_RISCV_conditional_strong_release"
+ | Write_X86_locked -> "Write_X86_locked"
+ end
+end
+*)
+
+Inductive a64_barrier_domain :=
+ A64_FullShare
+ | A64_InnerShare
+ | A64_OuterShare
+ | A64_NonShare.
+
+Inductive a64_barrier_type :=
+ A64_barrier_all
+ | A64_barrier_LD
+ | A64_barrier_ST.
+
+Inductive barrier_kind :=
+ (* Power barriers *)
+ | Barrier_Sync : unit -> barrier_kind
+ | Barrier_LwSync : unit -> barrier_kind
+ | Barrier_Eieio : unit -> barrier_kind
+ | Barrier_Isync : unit -> barrier_kind
+ (* AArch64 barriers *)
+ | Barrier_DMB : a64_barrier_domain -> a64_barrier_type -> barrier_kind
+ | Barrier_DSB : a64_barrier_domain -> a64_barrier_type -> barrier_kind
+ | Barrier_ISB : unit -> barrier_kind
+ (* | Barrier_TM_COMMIT*)
+ (* MIPS barriers *)
+ | Barrier_MIPS_SYNC : unit -> barrier_kind
+ (* RISC-V barriers *)
+ | Barrier_RISCV_rw_rw : unit -> barrier_kind
+ | Barrier_RISCV_r_rw : unit -> barrier_kind
+ | Barrier_RISCV_r_r : unit -> barrier_kind
+ | Barrier_RISCV_rw_w : unit -> barrier_kind
+ | Barrier_RISCV_w_w : unit -> barrier_kind
+ | Barrier_RISCV_w_rw : unit -> barrier_kind
+ | Barrier_RISCV_rw_r : unit -> barrier_kind
+ | Barrier_RISCV_r_w : unit -> barrier_kind
+ | Barrier_RISCV_w_r : unit -> barrier_kind
+ | Barrier_RISCV_tso : unit -> barrier_kind
+ | Barrier_RISCV_i : unit -> barrier_kind
+ (* X86 *)
+ | Barrier_x86_MFENCE : unit -> barrier_kind.
+Scheme Equality for barrier_kind.
+
+(*
+instance (Show barrier_kind)
+ let show := function
+ | Barrier_Sync -> "Barrier_Sync"
+ | Barrier_LwSync -> "Barrier_LwSync"
+ | Barrier_Eieio -> "Barrier_Eieio"
+ | Barrier_Isync -> "Barrier_Isync"
+ | Barrier_DMB -> "Barrier_DMB"
+ | Barrier_DMB_ST -> "Barrier_DMB_ST"
+ | Barrier_DMB_LD -> "Barrier_DMB_LD"
+ | Barrier_DSB -> "Barrier_DSB"
+ | Barrier_DSB_ST -> "Barrier_DSB_ST"
+ | Barrier_DSB_LD -> "Barrier_DSB_LD"
+ | Barrier_ISB -> "Barrier_ISB"
+ | Barrier_TM_COMMIT -> "Barrier_TM_COMMIT"
+ | Barrier_MIPS_SYNC -> "Barrier_MIPS_SYNC"
+ | Barrier_RISCV_rw_rw -> "Barrier_RISCV_rw_rw"
+ | Barrier_RISCV_r_rw -> "Barrier_RISCV_r_rw"
+ | Barrier_RISCV_r_r -> "Barrier_RISCV_r_r"
+ | Barrier_RISCV_rw_w -> "Barrier_RISCV_rw_w"
+ | Barrier_RISCV_w_w -> "Barrier_RISCV_w_w"
+ | Barrier_RISCV_w_rw -> "Barrier_RISCV_w_rw"
+ | Barrier_RISCV_rw_r -> "Barrier_RISCV_rw_r"
+ | Barrier_RISCV_r_w -> "Barrier_RISCV_r_w"
+ | Barrier_RISCV_w_r -> "Barrier_RISCV_w_r"
+ | Barrier_RISCV_tso -> "Barrier_RISCV_tso"
+ | Barrier_RISCV_i -> "Barrier_RISCV_i"
+ | Barrier_x86_MFENCE -> "Barrier_x86_MFENCE"
+ end
+end*)
+
+Inductive trans_kind :=
+ (* AArch64 *)
+ | Transaction_start | Transaction_commit | Transaction_abort.
+Scheme Equality for trans_kind.
+(*
+instance (Show trans_kind)
+ let show := function
+ | Transaction_start -> "Transaction_start"
+ | Transaction_commit -> "Transaction_commit"
+ | Transaction_abort -> "Transaction_abort"
+ end
+end*)
+
+Inductive instruction_kind :=
+ | IK_barrier : barrier_kind -> 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 : 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 : unit -> instruction_kind.
+
+(*
+instance (Show instruction_kind)
+ let show := function
+ | IK_barrier barrier_kind -> "IK_barrier " ^ (show barrier_kind)
+ | IK_mem_read read_kind -> "IK_mem_read " ^ (show read_kind)
+ | IK_mem_write write_kind -> "IK_mem_write " ^ (show write_kind)
+ | IK_mem_rmw (r, w) -> "IK_mem_rmw " ^ (show r) ^ " " ^ (show w)
+ | IK_branch -> "IK_branch"
+ | IK_trans trans_kind -> "IK_trans " ^ (show trans_kind)
+ | IK_simple -> "IK_simple"
+ end
+end
+*)
+
+Definition read_is_exclusive r :=
+match r with
+ | Read_plain => false
+ | Read_reserve => true
+ | Read_acquire => false
+ | Read_exclusive => true
+ | Read_exclusive_acquire => true
+ | Read_stream => false
+ | Read_RISCV_acquire => false
+ | Read_RISCV_strong_acquire => false
+ | Read_RISCV_reserved => true
+ | Read_RISCV_reserved_acquire => true
+ | Read_RISCV_reserved_strong_acquire => true
+ | Read_X86_locked => true
+end.
+
+
+(*
+instance (EnumerationType read_kind)
+ let toNat := function
+ | 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
+end
+
+instance (EnumerationType write_kind)
+ let toNat := function
+ | 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
+end
+
+instance (EnumerationType barrier_kind)
+ let toNat := function
+ | 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_TM_COMMIT -> 11
+ | Barrier_MIPS_SYNC -> 12
+ | Barrier_RISCV_rw_rw -> 13
+ | Barrier_RISCV_r_rw -> 14
+ | Barrier_RISCV_r_r -> 15
+ | Barrier_RISCV_rw_w -> 16
+ | Barrier_RISCV_w_w -> 17
+ | Barrier_RISCV_w_rw -> 18
+ | Barrier_RISCV_rw_r -> 19
+ | Barrier_RISCV_r_w -> 20
+ | Barrier_RISCV_w_r -> 21
+ | Barrier_RISCV_tso -> 22
+ | Barrier_RISCV_i -> 23
+ | Barrier_x86_MFENCE -> 24
+ end
+end
+*)
diff --git a/prover_snapshots/coq/lib/sail/Sail2_operators.v b/prover_snapshots/coq/lib/sail/Sail2_operators.v
new file mode 100644
index 0000000..ab02c4a
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_operators.v
@@ -0,0 +1,232 @@
+Require Import Sail2_values.
+Require List.
+Import List.ListNotations.
+
+(*** Bit vector operations *)
+
+Section Bitvectors.
+Context {a b c} `{Bitvector a} `{Bitvector b} `{Bitvector c}.
+
+(*val concat_bv : forall 'a 'b 'c. Bitvector 'a, Bitvector 'b, Bitvector 'c => 'a -> 'b -> 'c*)
+Definition concat_bv (l : a) (r : b) : list bitU := bits_of l ++ bits_of r.
+
+(*val cons_bv : forall 'a 'b 'c. Bitvector 'a, Bitvector 'b => bitU -> 'a -> 'b*)
+Definition cons_bv b' (v : a) : list bitU := b' :: bits_of v.
+
+Definition cast_unit_bv b : list bitU := [b].
+Definition bv_of_bit len b : list bitU := extz_bits len [b].
+
+(*Definition most_significant v := match bits_of v with
+ | cons b _ => b
+ | _ => failwith "most_significant applied to empty vector"
+ end.
+
+Definition get_max_representable_in sign (n : integer) : integer :=
+ if (n = 64) then match sign with | true -> max_64 | false -> max_64u end
+ else if (n=32) then match sign with | true -> max_32 | false -> max_32u end
+ else if (n=8) then max_8
+ else if (n=5) then max_5
+ else match sign with | true -> integerPow 2 ((natFromInteger n) -1)
+ | false -> integerPow 2 (natFromInteger n)
+ end
+
+Definition get_min_representable_in _ (n : integer) : integer :=
+ if n = 64 then min_64
+ else if n = 32 then min_32
+ else if n = 8 then min_8
+ else if n = 5 then min_5
+ else 0 - (integerPow 2 (natFromInteger n))
+
+val arith_op_bv_int : forall 'a 'b. Bitvector 'a =>
+ (integer -> integer -> integer) -> bool -> 'a -> integer -> 'a*)
+Definition arith_op_bv_int {a} `{Bitvector a} (op : Z -> Z -> Z) (sign : bool) (l : a) (r : Z) : a :=
+ let r' := of_int (length l) r in
+ arith_op_bv op sign l r'.
+
+(*val arith_op_int_bv : forall 'a 'b. Bitvector 'a =>
+ (integer -> integer -> integer) -> bool -> integer -> 'a -> 'a*)
+Definition arith_op_int_bv {a} `{Bitvector a} (op : Z -> Z -> Z) (sign : bool) (l : Z) (r : a) : a :=
+ let l' := of_int (length r) l in
+ arith_op_bv op sign l' r.
+(*
+Definition add_bv_int := arith_op_bv_int Zplus false 1.
+Definition sadd_bv_int := arith_op_bv_int Zplus true 1.
+Definition sub_bv_int := arith_op_bv_int Zminus false 1.
+Definition mult_bv_int := arith_op_bv_int Zmult false 2.
+Definition smult_bv_int := arith_op_bv_int Zmult true 2.
+
+(*val arith_op_int_bv : forall 'a 'b. Bitvector 'a, Bitvector 'b =>
+ (integer -> integer -> integer) -> bool -> integer -> integer -> 'a -> 'b
+Definition arith_op_int_bv op sign size l r :=
+ let r' = int_of_bv sign r in
+ let n = op l r' in
+ of_int (size * length r) n
+
+Definition add_int_bv = arith_op_int_bv integerAdd false 1
+Definition sadd_int_bv = arith_op_int_bv integerAdd true 1
+Definition sub_int_bv = arith_op_int_bv integerMinus false 1
+Definition mult_int_bv = arith_op_int_bv integerMult false 2
+Definition smult_int_bv = arith_op_int_bv integerMult true 2
+
+Definition arith_op_bv_bit op sign (size : integer) l r :=
+ let l' = int_of_bv sign l in
+ let n = op l' (match r with | B1 -> (1 : integer) | _ -> 0 end) in
+ of_int (size * length l) n
+
+Definition add_bv_bit := arith_op_bv_bit integerAdd false 1
+Definition sadd_bv_bit := arith_op_bv_bit integerAdd true 1
+Definition sub_bv_bit := arith_op_bv_bit integerMinus true 1
+
+val arith_op_overflow_bv : forall 'a 'b. Bitvector 'a, Bitvector 'b =>
+ (integer -> integer -> integer) -> bool -> integer -> 'a -> 'a -> ('b * bitU * bitU)
+Definition arith_op_overflow_bv op sign size l r :=
+ let len := length l in
+ let act_size := len * size in
+ let (l_sign,r_sign) := (int_of_bv sign l,int_of_bv sign r) in
+ let (l_unsign,r_unsign) := (int_of_bv false l,int_of_bv false r) in
+ let n := op l_sign r_sign in
+ let n_unsign := op l_unsign r_unsign in
+ let correct_size := of_int act_size n in
+ let one_more_size_u := bits_of_int (act_size + 1) n_unsign in
+ let overflow :=
+ if n <= get_max_representable_in sign len &&
+ n >= get_min_representable_in sign len
+ then B0 else B1 in
+ let c_out := most_significant one_more_size_u in
+ (correct_size,overflow,c_out)
+
+Definition add_overflow_bv := arith_op_overflow_bv integerAdd false 1
+Definition add_overflow_bv_signed := arith_op_overflow_bv integerAdd true 1
+Definition sub_overflow_bv := arith_op_overflow_bv integerMinus false 1
+Definition sub_overflow_bv_signed := arith_op_overflow_bv integerMinus true 1
+Definition mult_overflow_bv := arith_op_overflow_bv integerMult false 2
+Definition mult_overflow_bv_signed := arith_op_overflow_bv integerMult true 2
+
+val arith_op_overflow_bv_bit : forall 'a 'b. Bitvector 'a, Bitvector 'b =>
+ (integer -> integer -> integer) -> bool -> integer -> 'a -> bitU -> ('b * bitU * bitU)
+Definition arith_op_overflow_bv_bit op sign size l r_bit :=
+ let act_size := length l * size in
+ let l' := int_of_bv sign l in
+ let l_u := int_of_bv false l in
+ let (n,nu,changed) := match r_bit with
+ | B1 -> (op l' 1, op l_u 1, true)
+ | B0 -> (l',l_u,false)
+ | BU -> failwith "arith_op_overflow_bv_bit applied to undefined bit"
+ end in
+ let correct_size := of_int act_size n in
+ let one_larger := bits_of_int (act_size + 1) nu in
+ let overflow :=
+ if changed
+ then
+ if n <= get_max_representable_in sign act_size && n >= get_min_representable_in sign act_size
+ then B0 else B1
+ else B0 in
+ (correct_size,overflow,most_significant one_larger)
+
+Definition add_overflow_bv_bit := arith_op_overflow_bv_bit integerAdd false 1
+Definition add_overflow_bv_bit_signed := arith_op_overflow_bv_bit integerAdd true 1
+Definition sub_overflow_bv_bit := arith_op_overflow_bv_bit integerMinus false 1
+Definition sub_overflow_bv_bit_signed := arith_op_overflow_bv_bit integerMinus true 1
+
+type shift := LL_shift | RR_shift | RR_shift_arith | LL_rot | RR_rot
+
+val shift_op_bv : forall 'a. Bitvector 'a => shift -> 'a -> integer -> 'a
+Definition shift_op_bv op v n :=
+ match op with
+ | LL_shift ->
+ of_bits (get_bits true v n (length v - 1) ++ repeat [B0] n)
+ | RR_shift ->
+ of_bits (repeat [B0] n ++ get_bits true v 0 (length v - n - 1))
+ | RR_shift_arith ->
+ of_bits (repeat [most_significant v] n ++ get_bits true v 0 (length v - n - 1))
+ | LL_rot ->
+ of_bits (get_bits true v n (length v - 1) ++ get_bits true v 0 (n - 1))
+ | RR_rot ->
+ of_bits (get_bits false v 0 (n - 1) ++ get_bits false v n (length v - 1))
+ end
+
+Definition shiftl_bv := shift_op_bv LL_shift (*"<<"*)
+Definition shiftr_bv := shift_op_bv RR_shift (*">>"*)
+Definition arith_shiftr_bv := shift_op_bv RR_shift_arith
+Definition rotl_bv := shift_op_bv LL_rot (*"<<<"*)
+Definition rotr_bv := shift_op_bv LL_rot (*">>>"*)
+
+Definition shiftl_mword w n := Machine_word.shiftLeft w (natFromInteger n)
+Definition shiftr_mword w n := Machine_word.shiftRight w (natFromInteger n)
+Definition rotl_mword w n := Machine_word.rotateLeft (natFromInteger n) w
+Definition rotr_mword w n := Machine_word.rotateRight (natFromInteger n) w
+
+Definition rec arith_op_no0 (op : integer -> integer -> integer) l r :=
+ if r = 0
+ then Nothing
+ else Just (op l r)
+
+val arith_op_bv_no0 : forall 'a 'b. Bitvector 'a, Bitvector 'b =>
+ (integer -> integer -> integer) -> bool -> integer -> 'a -> 'a -> 'b
+Definition arith_op_bv_no0 op sign size l r :=
+ let act_size := length l * size in
+ let (l',r') := (int_of_bv sign l,int_of_bv sign r) in
+ let n := arith_op_no0 op l' r' in
+ let (representable,n') :=
+ match n with
+ | Just n' ->
+ (n' <= get_max_representable_in sign act_size &&
+ n' >= get_min_representable_in sign act_size, n')
+ | _ -> (false,0)
+ end in
+ if representable then (of_int act_size n') else (of_bits (repeat [BU] act_size))
+
+Definition mod_bv := arith_op_bv_no0 hardware_mod false 1
+Definition quot_bv := arith_op_bv_no0 hardware_quot false 1
+Definition quot_bv_signed := arith_op_bv_no0 hardware_quot true 1
+
+Definition mod_mword := Machine_word.modulo
+Definition quot_mword := Machine_word.unsignedDivide
+Definition quot_mword_signed := Machine_word.signedDivide
+
+Definition arith_op_bv_int_no0 op sign size l r :=
+ arith_op_bv_no0 op sign size l (of_int (length l) r)
+
+Definition quot_bv_int := arith_op_bv_int_no0 hardware_quot false 1
+Definition mod_bv_int := arith_op_bv_int_no0 hardware_mod false 1
+*)
+Definition replicate_bits_bv {a b} `{Bitvector a} `{Bitvector b} (v : a) count : b := of_bits (repeat (bits_of v) count).
+Import List.
+Import ListNotations.
+Definition duplicate_bit_bv {a} `{Bitvector a} bit len : a := replicate_bits_bv [bit] len.
+
+(*val eq_bv : forall 'a. Bitvector 'a => 'a -> 'a -> bool*)
+Definition eq_bv {A} `{Bitvector A} (l : A) r := (unsigned l =? unsigned r).
+
+(*val neq_bv : forall 'a. Bitvector 'a => 'a -> 'a -> bool*)
+Definition neq_bv (l : a) (r :a) : bool := (negb (unsigned l =? unsigned r)).
+(*
+val ucmp_bv : forall 'a. Bitvector 'a => (integer -> integer -> bool) -> 'a -> 'a -> bool
+Definition ucmp_bv cmp l r := cmp (unsigned l) (unsigned r)
+
+val scmp_bv : forall 'a. Bitvector 'a => (integer -> integer -> bool) -> 'a -> 'a -> bool
+Definition scmp_bv cmp l r := cmp (signed l) (signed r)
+
+Definition ult_bv := ucmp_bv (<)
+Definition slt_bv := scmp_bv (<)
+Definition ugt_bv := ucmp_bv (>)
+Definition sgt_bv := scmp_bv (>)
+Definition ulteq_bv := ucmp_bv (<=)
+Definition slteq_bv := scmp_bv (<=)
+Definition ugteq_bv := ucmp_bv (>=)
+Definition sgteq_bv := scmp_bv (>=)
+*)
+
+(*val get_slice_int_bv : forall 'a. Bitvector 'a => integer -> integer -> integer -> 'a*)*)
+Definition get_slice_int_bv {a} `{Bitvector a} len n lo : a :=
+ let hi := lo + len - 1 in
+ let bs := bools_of_int (hi + 1) n in
+ of_bools (subrange_list false bs hi lo).
+
+(*val set_slice_int_bv : forall 'a. Bitvector 'a => integer -> integer -> integer -> 'a -> integer
+Definition set_slice_int_bv {a} `{Bitvector a} len n lo (v : a) :=
+ let hi := lo + len - 1 in
+ let bs := bits_of_int (hi + 1) n in
+ maybe_failwith (signed_of_bits (update_subrange_list false bs hi lo (bits_of v))).*)
+
+End Bitvectors.
diff --git a/prover_snapshots/coq/lib/sail/Sail2_operators_bitlists.v b/prover_snapshots/coq/lib/sail/Sail2_operators_bitlists.v
new file mode 100644
index 0000000..dbd8215
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_operators_bitlists.v
@@ -0,0 +1,182 @@
+Require Import Sail2_values.
+Require Import Sail2_operators.
+
+(*
+
+(* Specialisation of operators to bit lists *)
+
+val access_vec_inc : list bitU -> integer -> bitU
+let access_vec_inc = access_bv_inc
+
+val access_vec_dec : list bitU -> integer -> bitU
+let access_vec_dec = access_bv_dec
+
+val update_vec_inc : list bitU -> integer -> bitU -> list bitU
+let update_vec_inc = update_bv_inc
+
+val update_vec_dec : list bitU -> integer -> bitU -> list bitU
+let update_vec_dec = update_bv_dec
+
+val subrange_vec_inc : list bitU -> integer -> integer -> list bitU
+let subrange_vec_inc = subrange_bv_inc
+
+val subrange_vec_dec : list bitU -> integer -> integer -> list bitU
+let subrange_vec_dec = subrange_bv_dec
+
+val update_subrange_vec_inc : list bitU -> integer -> integer -> list bitU -> list bitU
+let update_subrange_vec_inc = update_subrange_bv_inc
+
+val update_subrange_vec_dec : list bitU -> integer -> integer -> list bitU -> list bitU
+let update_subrange_vec_dec = update_subrange_bv_dec
+
+val extz_vec : integer -> list bitU -> list bitU
+let extz_vec = extz_bv
+
+val exts_vec : integer -> list bitU -> list bitU
+let exts_vec = exts_bv
+
+val concat_vec : list bitU -> list bitU -> list bitU
+let concat_vec = concat_bv
+
+val cons_vec : bitU -> list bitU -> list bitU
+let cons_vec = cons_bv
+
+val bool_of_vec : mword ty1 -> bitU
+let bool_of_vec = bool_of_bv
+
+val cast_unit_vec : bitU -> mword ty1
+let cast_unit_vec = cast_unit_bv
+
+val vec_of_bit : integer -> bitU -> list bitU
+let vec_of_bit = bv_of_bit
+
+val msb : list bitU -> bitU
+let msb = most_significant
+
+val int_of_vec : bool -> list bitU -> integer
+let int_of_vec = int_of_bv
+
+val string_of_vec : list bitU -> string
+let string_of_vec = string_of_bv
+
+val and_vec : list bitU -> list bitU -> list bitU
+val or_vec : list bitU -> list bitU -> list bitU
+val xor_vec : list bitU -> list bitU -> list bitU
+val not_vec : list bitU -> list bitU
+let and_vec = and_bv
+let or_vec = or_bv
+let xor_vec = xor_bv
+let not_vec = not_bv
+
+val add_vec : list bitU -> list bitU -> list bitU
+val sadd_vec : list bitU -> list bitU -> list bitU
+val sub_vec : list bitU -> list bitU -> list bitU
+val mult_vec : list bitU -> list bitU -> list bitU
+val smult_vec : list bitU -> list bitU -> list bitU
+let add_vec = add_bv
+let sadd_vec = sadd_bv
+let sub_vec = sub_bv
+let mult_vec = mult_bv
+let smult_vec = smult_bv
+
+val add_vec_int : list bitU -> integer -> list bitU
+val sadd_vec_int : list bitU -> integer -> list bitU
+val sub_vec_int : list bitU -> integer -> list bitU
+val mult_vec_int : list bitU -> integer -> list bitU
+val smult_vec_int : list bitU -> integer -> list bitU
+let add_vec_int = add_bv_int
+let sadd_vec_int = sadd_bv_int
+let sub_vec_int = sub_bv_int
+let mult_vec_int = mult_bv_int
+let smult_vec_int = smult_bv_int
+
+val add_int_vec : integer -> list bitU -> list bitU
+val sadd_int_vec : integer -> list bitU -> list bitU
+val sub_int_vec : integer -> list bitU -> list bitU
+val mult_int_vec : integer -> list bitU -> list bitU
+val smult_int_vec : integer -> list bitU -> list bitU
+let add_int_vec = add_int_bv
+let sadd_int_vec = sadd_int_bv
+let sub_int_vec = sub_int_bv
+let mult_int_vec = mult_int_bv
+let smult_int_vec = smult_int_bv
+
+val add_vec_bit : list bitU -> bitU -> list bitU
+val sadd_vec_bit : list bitU -> bitU -> list bitU
+val sub_vec_bit : list bitU -> bitU -> list bitU
+let add_vec_bit = add_bv_bit
+let sadd_vec_bit = sadd_bv_bit
+let sub_vec_bit = sub_bv_bit
+
+val add_overflow_vec : list bitU -> list bitU -> (list bitU * bitU * bitU)
+val add_overflow_vec_signed : list bitU -> list bitU -> (list bitU * bitU * bitU)
+val sub_overflow_vec : list bitU -> list bitU -> (list bitU * bitU * bitU)
+val sub_overflow_vec_signed : list bitU -> list bitU -> (list bitU * bitU * bitU)
+val mult_overflow_vec : list bitU -> list bitU -> (list bitU * bitU * bitU)
+val mult_overflow_vec_signed : list bitU -> list bitU -> (list bitU * bitU * bitU)
+let add_overflow_vec = add_overflow_bv
+let add_overflow_vec_signed = add_overflow_bv_signed
+let sub_overflow_vec = sub_overflow_bv
+let sub_overflow_vec_signed = sub_overflow_bv_signed
+let mult_overflow_vec = mult_overflow_bv
+let mult_overflow_vec_signed = mult_overflow_bv_signed
+
+val add_overflow_vec_bit : list bitU -> bitU -> (list bitU * bitU * bitU)
+val add_overflow_vec_bit_signed : list bitU -> bitU -> (list bitU * bitU * bitU)
+val sub_overflow_vec_bit : list bitU -> bitU -> (list bitU * bitU * bitU)
+val sub_overflow_vec_bit_signed : list bitU -> bitU -> (list bitU * bitU * bitU)
+let add_overflow_vec_bit = add_overflow_bv_bit
+let add_overflow_vec_bit_signed = add_overflow_bv_bit_signed
+let sub_overflow_vec_bit = sub_overflow_bv_bit
+let sub_overflow_vec_bit_signed = sub_overflow_bv_bit_signed
+
+val shiftl : list bitU -> integer -> list bitU
+val shiftr : list bitU -> integer -> list bitU
+val arith_shiftr : list bitU -> integer -> list bitU
+val rotl : list bitU -> integer -> list bitU
+val rotr : list bitU -> integer -> list bitU
+let shiftl = shiftl_bv
+let shiftr = shiftr_bv
+let arith_shiftr = arith_shiftr_bv
+let rotl = rotl_bv
+let rotr = rotr_bv
+
+val mod_vec : list bitU -> list bitU -> list bitU
+val quot_vec : list bitU -> list bitU -> list bitU
+val quot_vec_signed : list bitU -> list bitU -> list bitU
+let mod_vec = mod_bv
+let quot_vec = quot_bv
+let quot_vec_signed = quot_bv_signed
+
+val mod_vec_int : list bitU -> integer -> list bitU
+val quot_vec_int : list bitU -> integer -> list bitU
+let mod_vec_int = mod_bv_int
+let quot_vec_int = quot_bv_int
+
+val replicate_bits : list bitU -> integer -> list bitU
+let replicate_bits = replicate_bits_bv
+
+val duplicate : bitU -> integer -> list bitU
+let duplicate = duplicate_bit_bv
+
+val eq_vec : list bitU -> list bitU -> bool
+val neq_vec : list bitU -> list bitU -> bool
+val ult_vec : list bitU -> list bitU -> bool
+val slt_vec : list bitU -> list bitU -> bool
+val ugt_vec : list bitU -> list bitU -> bool
+val sgt_vec : list bitU -> list bitU -> bool
+val ulteq_vec : list bitU -> list bitU -> bool
+val slteq_vec : list bitU -> list bitU -> bool
+val ugteq_vec : list bitU -> list bitU -> bool
+val sgteq_vec : list bitU -> list bitU -> bool
+let eq_vec = eq_bv
+let neq_vec = neq_bv
+let ult_vec = ult_bv
+let slt_vec = slt_bv
+let ugt_vec = ugt_bv
+let sgt_vec = sgt_bv
+let ulteq_vec = ulteq_bv
+let slteq_vec = slteq_bv
+let ugteq_vec = ugteq_bv
+let sgteq_vec = sgteq_bv
+*)
diff --git a/prover_snapshots/coq/lib/sail/Sail2_operators_mwords.v b/prover_snapshots/coq/lib/sail/Sail2_operators_mwords.v
new file mode 100644
index 0000000..697bc4a
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_operators_mwords.v
@@ -0,0 +1,544 @@
+Require Import Sail2_values.
+Require Import Sail2_operators.
+Require Import Sail2_prompt_monad.
+Require Import Sail2_prompt.
+Require Import bbv.Word.
+Require bbv.BinNotation.
+Require Import Arith.
+Require Import ZArith.
+Require Import Omega.
+Require Import Eqdep_dec.
+
+Fixpoint cast_positive (T : positive -> Type) (p q : positive) : T p -> p = q -> T q.
+refine (
+match p, q with
+| xH, xH => fun x _ => x
+| xO p', xO q' => fun x e => cast_positive (fun x => T (xO x)) p' q' x _
+| xI p', xI q' => fun x e => cast_positive (fun x => T (xI x)) p' q' x _
+| _, _ => _
+end); congruence.
+Defined.
+
+Definition cast_T {T : Z -> Type} {m n} : forall (x : T m) (eq : m = n), T n.
+refine (match m,n with
+| Z0, Z0 => fun x _ => x
+| Zneg p1, Zneg p2 => fun x e => cast_positive (fun p => T (Zneg p)) p1 p2 x _
+| Zpos p1, Zpos p2 => fun x e => cast_positive (fun p => T (Zpos p)) p1 p2 x _
+| _,_ => _
+end); congruence.
+Defined.
+
+Lemma cast_positive_refl : forall p T x (e : p = p),
+ cast_positive T p p x e = x.
+induction p.
+* intros. simpl. rewrite IHp; auto.
+* intros. simpl. rewrite IHp; auto.
+* reflexivity.
+Qed.
+
+Lemma cast_T_refl {T : Z -> Type} {m} {H:m = m} (x : T m) : cast_T x H = x.
+destruct m.
+* reflexivity.
+* simpl. rewrite cast_positive_refl. reflexivity.
+* simpl. rewrite cast_positive_refl. reflexivity.
+Qed.
+
+Definition autocast {T : Z -> Type} {m n} (x : T m) `{H:ArithFact (m = n)} : T n :=
+ cast_T x (use_ArithFact H).
+
+Definition autocast_m {rv e m n} (x : monad rv (mword m) e) `{H:ArithFact (m = n)} : monad rv (mword n) e :=
+ x >>= fun x => returnm (cast_T x (use_ArithFact H)).
+
+Definition cast_word {m n} (x : Word.word m) (eq : m = n) : Word.word n :=
+ DepEqNat.nat_cast _ eq x.
+
+Lemma cast_word_refl {m} {H:m = m} (x : word m) : cast_word x H = x.
+rewrite (UIP_refl_nat _ H).
+apply nat_cast_same.
+Qed.
+
+Definition mword_of_nat {m} : Word.word m -> mword (Z.of_nat m).
+refine (match m return word m -> mword (Z.of_nat m) with
+| O => fun x => x
+| S m' => fun x => nat_cast _ _ x
+end).
+rewrite SuccNat2Pos.id_succ.
+reflexivity.
+Defined.
+
+Definition cast_to_mword {m n} (x : Word.word m) : Z.of_nat m = n -> mword n.
+refine (match n return Z.of_nat m = n -> mword n with
+| Z0 => fun _ => WO
+| Zpos p => fun eq => cast_T (mword_of_nat x) eq
+| Zneg p => _
+end).
+intro eq.
+exfalso. destruct m; simpl in *; congruence.
+Defined.
+
+(*
+(* Specialisation of operators to machine words *)
+
+val access_vec_inc : forall 'a. Size 'a => mword 'a -> integer -> bitU*)
+Definition access_vec_inc {a} : mword a -> Z -> bitU := access_mword_inc.
+
+(*val access_vec_dec : forall 'a. Size 'a => mword 'a -> integer -> bitU*)
+Definition access_vec_dec {a} : mword a -> Z -> bitU := access_mword_dec.
+
+(*val update_vec_inc : forall 'a. Size 'a => mword 'a -> integer -> bitU -> mword 'a*)
+(* TODO: probably ought to use a monadic version instead, but using bad default for
+ type compatibility just now *)
+Definition update_vec_inc {a} (w : mword a) i b : mword a :=
+ opt_def w (update_mword_inc w i b).
+
+(*val update_vec_dec : forall 'a. Size 'a => mword 'a -> integer -> bitU -> mword 'a*)
+Definition update_vec_dec {a} (w : mword a) i b : mword a := opt_def w (update_mword_dec w i b).
+
+Lemma subrange_lemma0 {n m o} `{ArithFact (0 <= o)} `{ArithFact (o <= m < n)} : (Z.to_nat o <= Z.to_nat m < Z.to_nat n)%nat.
+intros.
+unwrap_ArithFacts.
+split.
++ apply Z2Nat.inj_le; omega.
++ apply Z2Nat.inj_lt; omega.
+Qed.
+Lemma subrange_lemma1 {n m o} : (o <= m < n -> n = m + 1 + (n - (m + 1)))%nat.
+intros. omega.
+Qed.
+Lemma subrange_lemma2 {n m o} : (o <= m < n -> m+1 = o+(m-o+1))%nat.
+omega.
+Qed.
+Lemma subrange_lemma3 {n m o} `{ArithFact (0 <= o)} `{ArithFact (o <= m < n)} :
+ Z.of_nat (Z.to_nat m - Z.to_nat o + 1)%nat = m - o + 1.
+unwrap_ArithFacts.
+rewrite Nat2Z.inj_add.
+rewrite Nat2Z.inj_sub.
+repeat rewrite Z2Nat.id; try omega.
+reflexivity.
+apply Z2Nat.inj_le; omega.
+Qed.
+
+Definition subrange_vec_dec {n} (v : mword n) m o `{ArithFact (0 <= o)} `{ArithFact (o <= m < n)} : mword (m - o + 1) :=
+ let n := Z.to_nat n in
+ let m := Z.to_nat m in
+ let o := Z.to_nat o in
+ let prf : (o <= m < n)%nat := subrange_lemma0 in
+ let w := get_word v in
+ cast_to_mword (split2 o (m-o+1)
+ (cast_word (split1 (m+1) (n-(m+1)) (cast_word w (subrange_lemma1 prf)))
+ (subrange_lemma2 prf))) subrange_lemma3.
+
+Definition subrange_vec_inc {n} (v : mword n) m o `{ArithFact (0 <= m)} `{ArithFact (m <= o < n)} : mword (o - m + 1) := autocast (subrange_vec_dec v (n-1-m) (n-1-o)).
+
+(* TODO: get rid of bogus default *)
+Parameter dummy_vector : forall {n} `{ArithFact (n >= 0)}, mword n.
+
+(*val update_subrange_vec_inc : forall 'a 'b. Size 'a, Size 'b => mword 'a -> integer -> integer -> mword 'b -> mword 'a*)
+Definition update_subrange_vec_inc {a b} (v : mword a) i j (w : mword b) : mword a :=
+ opt_def dummy_vector (of_bits (update_subrange_bv_inc v i j w)).
+
+(*val update_subrange_vec_dec : forall 'a 'b. Size 'a, Size 'b => mword 'a -> integer -> integer -> mword 'b -> mword 'a*)
+Definition update_subrange_vec_dec {a b} (v : mword a) i j (w : mword b) : mword a :=
+ opt_def dummy_vector (of_bits (update_subrange_bv_dec v i j w)).
+
+Lemma mword_nonneg {a} : mword a -> a >= 0.
+destruct a;
+auto using Z.le_ge, Zle_0_pos with zarith.
+destruct 1.
+Qed.
+
+(*val extz_vec : forall 'a 'b. Size 'a, Size 'b => integer -> mword 'a -> mword 'b*)
+Definition extz_vec {a b} `{ArithFact (b >= a)} (n : Z) (v : mword a) : mword b.
+refine (cast_to_mword (Word.zext (get_word v) (Z.to_nat (b - a))) _).
+unwrap_ArithFacts.
+assert (a >= 0). { apply mword_nonneg. assumption. }
+rewrite <- Z2Nat.inj_add; try omega.
+rewrite Zplus_minus.
+apply Z2Nat.id.
+auto with zarith.
+Defined.
+
+(*val exts_vec : forall 'a 'b. Size 'a, Size 'b => integer -> mword 'a -> mword 'b*)
+Definition exts_vec {a b} `{ArithFact (b >= a)} (n : Z) (v : mword a) : mword b.
+refine (cast_to_mword (Word.sext (get_word v) (Z.to_nat (b - a))) _).
+unwrap_ArithFacts.
+assert (a >= 0). { apply mword_nonneg. assumption. }
+rewrite <- Z2Nat.inj_add; try omega.
+rewrite Zplus_minus.
+apply Z2Nat.id.
+auto with zarith.
+Defined.
+
+Definition zero_extend {a} (v : mword a) (n : Z) `{ArithFact (n >= a)} : mword n := extz_vec n v.
+
+Definition sign_extend {a} (v : mword a) (n : Z) `{ArithFact (n >= a)} : mword n := exts_vec n v.
+
+Definition zeros (n : Z) `{ArithFact (n >= 0)} : mword n.
+refine (cast_to_mword (Word.wzero (Z.to_nat n)) _).
+unwrap_ArithFacts.
+apply Z2Nat.id.
+auto with zarith.
+Defined.
+
+Lemma truncate_eq {m n} : m >= 0 -> m <= n -> (Z.to_nat n = Z.to_nat m + (Z.to_nat n - Z.to_nat m))%nat.
+intros.
+assert ((Z.to_nat m <= Z.to_nat n)%nat).
+{ apply Z2Nat.inj_le; omega. }
+omega.
+Qed.
+Lemma truncateLSB_eq {m n} : m >= 0 -> m <= n -> (Z.to_nat n = (Z.to_nat n - Z.to_nat m) + Z.to_nat m)%nat.
+intros.
+assert ((Z.to_nat m <= Z.to_nat n)%nat).
+{ apply Z2Nat.inj_le; omega. }
+omega.
+Qed.
+
+Definition vector_truncate {n} (v : mword n) (m : Z) `{ArithFact (m >= 0)} `{ArithFact (m <= n)} : mword m :=
+ cast_to_mword (Word.split1 _ _ (cast_word (get_word v) (ltac:(unwrap_ArithFacts; apply truncate_eq; auto) : Z.to_nat n = Z.to_nat m + (Z.to_nat n - Z.to_nat m))%nat)) (ltac:(unwrap_ArithFacts; apply Z2Nat.id; omega) : Z.of_nat (Z.to_nat m) = m).
+
+Definition vector_truncateLSB {n} (v : mword n) (m : Z) `{ArithFact (m >= 0)} `{ArithFact (m <= n)} : mword m :=
+ cast_to_mword (Word.split2 _ _ (cast_word (get_word v) (ltac:(unwrap_ArithFacts; apply truncateLSB_eq; auto) : Z.to_nat n = (Z.to_nat n - Z.to_nat m) + Z.to_nat m)%nat)) (ltac:(unwrap_ArithFacts; apply Z2Nat.id; omega) : Z.of_nat (Z.to_nat m) = m).
+
+Lemma concat_eq {a b} : a >= 0 -> b >= 0 -> Z.of_nat (Z.to_nat b + Z.to_nat a)%nat = a + b.
+intros.
+rewrite Nat2Z.inj_add.
+rewrite Z2Nat.id; auto with zarith.
+rewrite Z2Nat.id; auto with zarith.
+Qed.
+
+
+(*val concat_vec : forall 'a 'b 'c. Size 'a, Size 'b, Size 'c => mword 'a -> mword 'b -> mword 'c*)
+Definition concat_vec {a b} (v : mword a) (w : mword b) : mword (a + b) :=
+ cast_to_mword (Word.combine (get_word w) (get_word v)) (ltac:(solve [auto using concat_eq, mword_nonneg with zarith]) : Z.of_nat (Z.to_nat b + Z.to_nat a)%nat = a + b).
+
+(*val cons_vec : forall 'a 'b 'c. Size 'a, Size 'b => bitU -> mword 'a -> mword 'b*)
+(*Definition cons_vec {a b} : bitU -> mword a -> mword b := cons_bv.*)
+
+(*val bool_of_vec : mword ty1 -> bitU
+Definition bool_of_vec := bool_of_bv
+
+val cast_unit_vec : bitU -> mword ty1
+Definition cast_unit_vec := cast_unit_bv
+
+val vec_of_bit : forall 'a. Size 'a => integer -> bitU -> mword 'a
+Definition vec_of_bit := bv_of_bit*)
+
+Require Import bbv.NatLib.
+
+Lemma Npow2_pow {n} : (2 ^ (N.of_nat n) = Npow2 n)%N.
+induction n.
+* reflexivity.
+* rewrite Nnat.Nat2N.inj_succ.
+ rewrite N.pow_succ_r'.
+ rewrite IHn.
+ rewrite Npow2_S.
+ rewrite Word.Nmul_two.
+ reflexivity.
+Qed.
+
+Program Definition uint {a} (x : mword a) : {z : Z & ArithFact (0 <= z /\ z <= 2 ^ a - 1)} :=
+ existT _ (Z.of_N (Word.wordToN (get_word x))) _.
+Next Obligation.
+constructor.
+constructor.
+* apply N2Z.is_nonneg.
+* assert (2 ^ a - 1 = Z.of_N (2 ^ (Z.to_N a) - 1)). {
+ rewrite N2Z.inj_sub.
+ * rewrite N2Z.inj_pow.
+ rewrite Z2N.id; auto.
+ destruct a; auto with zarith. destruct x.
+ * apply N.le_trans with (m := (2^0)%N); auto using N.le_refl.
+ apply N.pow_le_mono_r.
+ inversion 1.
+ apply N.le_0_l.
+ }
+ rewrite H.
+ apply N2Z.inj_le.
+ rewrite N.sub_1_r.
+ apply N.lt_le_pred.
+ rewrite <- Z_nat_N.
+ rewrite Npow2_pow.
+ apply Word.wordToN_bound.
+Defined.
+
+Lemma Zpow_pow2 {n} : 2 ^ Z.of_nat n = Z.of_nat (pow2 n).
+induction n.
+* reflexivity.
+* rewrite pow2_S_z.
+ rewrite Nat2Z.inj_succ.
+ rewrite Z.pow_succ_r; auto with zarith.
+Qed.
+
+Program Definition sint {a} `{ArithFact (a > 0)} (x : mword a) : {z : Z & ArithFact (-(2^(a-1)) <= z /\ z <= 2 ^ (a-1) - 1)} :=
+ existT _ (Word.wordToZ (get_word x)) _.
+Next Obligation.
+destruct H.
+destruct a; try inversion fact.
+constructor.
+generalize (get_word x).
+rewrite <- positive_nat_Z.
+destruct (Pos2Nat.is_succ p) as [n eq].
+rewrite eq.
+rewrite Nat2Z.id.
+intro w.
+destruct (Word.wordToZ_size' w) as [LO HI].
+replace 1 with (Z.of_nat 1); auto.
+rewrite <- Nat2Z.inj_sub; auto with arith.
+simpl.
+rewrite <- minus_n_O.
+rewrite Zpow_pow2.
+rewrite Z.sub_1_r.
+rewrite <- Z.lt_le_pred.
+auto.
+Defined.
+
+Definition sint0 {a} `{ArithFact (a >= 0)} (x : mword a) : Z :=
+ if sumbool_of_bool (Z.eqb a 0) then 0 else projT1 (sint x).
+
+Lemma length_list_pos : forall {A} {l:list A}, length_list l >= 0.
+unfold length_list.
+auto with zarith.
+Qed.
+Hint Resolve length_list_pos : sail.
+
+Definition vec_of_bits (l:list bitU) : mword (length_list l) := opt_def dummy_vector (of_bits l).
+(*
+
+val msb : forall 'a. Size 'a => mword 'a -> bitU
+Definition msb := most_significant
+
+val int_of_vec : forall 'a. Size 'a => bool -> mword 'a -> integer
+Definition int_of_vec := int_of_bv
+
+val string_of_vec : forall 'a. Size 'a => mword 'a -> string*)
+Definition string_of_bits {n} (w : mword n) : string := string_of_bv w.
+Definition with_word' {n} (P : Type -> Type) : (forall n, Word.word n -> P (Word.word n)) -> mword n -> P (mword n) := fun f w => @with_word n _ (f (Z.to_nat n)) w.
+Definition word_binop {n} (f : forall n, Word.word n -> Word.word n -> Word.word n) : mword n -> mword n -> mword n := with_word' (fun x => x -> x) f.
+Definition word_unop {n} (f : forall n, Word.word n -> Word.word n) : mword n -> mword n := with_word' (fun x => x) f.
+
+
+(*
+val and_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> mword 'a
+val or_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> mword 'a
+val xor_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> mword 'a
+val not_vec : forall 'a. Size 'a => mword 'a -> mword 'a*)
+Definition and_vec {n} : mword n -> mword n -> mword n := word_binop Word.wand.
+Definition or_vec {n} : mword n -> mword n -> mword n := word_binop Word.wor.
+Definition xor_vec {n} : mword n -> mword n -> mword n := word_binop Word.wxor.
+Definition not_vec {n} : mword n -> mword n := word_unop Word.wnot.
+
+(*val add_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> mword 'a
+val sadd_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> mword 'a
+val sub_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> mword 'a
+val mult_vec : forall 'a 'b. Size 'a, Size 'b => mword 'a -> mword 'a -> mword 'b
+val smult_vec : forall 'a 'b. Size 'a, Size 'b => mword 'a -> mword 'a -> mword 'b*)
+Definition add_vec {n} : mword n -> mword n -> mword n := word_binop Word.wplus.
+(*Definition sadd_vec {n} : mword n -> mword n -> mword n := sadd_bv w.*)
+Definition sub_vec {n} : mword n -> mword n -> mword n := word_binop Word.wminus.
+Definition mult_vec {n m} `{ArithFact (m >= n)} (l : mword n) (r : mword n) : mword m :=
+ word_binop Word.wmult (zero_extend l _) (zero_extend r _).
+Definition mults_vec {n m} `{ArithFact (m >= n)} (l : mword n) (r : mword n) : mword m :=
+ word_binop Word.wmult (sign_extend l _) (sign_extend r _).
+
+(*val add_vec_int : forall 'a. Size 'a => mword 'a -> integer -> mword 'a
+val sadd_vec_int : forall 'a. Size 'a => mword 'a -> integer -> mword 'a
+val sub_vec_int : forall 'a. Size 'a => mword 'a -> integer -> mword 'a
+val mult_vec_int : forall 'a 'b. Size 'a, Size 'b => mword 'a -> integer -> mword 'b
+val smult_vec_int : forall 'a 'b. Size 'a, Size 'b => mword 'a -> integer -> mword 'b*)
+Definition add_vec_int {a} (l : mword a) (r : Z) : mword a := arith_op_bv_int Z.add false l r.
+Definition sadd_vec_int {a} (l : mword a) (r : Z) : mword a := arith_op_bv_int Z.add true l r.
+Definition sub_vec_int {a} (l : mword a) (r : Z) : mword a := arith_op_bv_int Z.sub false l r.
+(*Definition mult_vec_int {a b} : mword a -> Z -> mword b := mult_bv_int.
+Definition smult_vec_int {a b} : mword a -> Z -> mword b := smult_bv_int.*)
+
+(*val add_int_vec : forall 'a. Size 'a => integer -> mword 'a -> mword 'a
+val sadd_int_vec : forall 'a. Size 'a => integer -> mword 'a -> mword 'a
+val sub_int_vec : forall 'a. Size 'a => integer -> mword 'a -> mword 'a
+val mult_int_vec : forall 'a 'b. Size 'a, Size 'b => integer -> mword 'a -> mword 'b
+val smult_int_vec : forall 'a 'b. Size 'a, Size 'b => integer -> mword 'a -> mword 'b
+Definition add_int_vec := add_int_bv
+Definition sadd_int_vec := sadd_int_bv
+Definition sub_int_vec := sub_int_bv
+Definition mult_int_vec := mult_int_bv
+Definition smult_int_vec := smult_int_bv
+
+val add_vec_bit : forall 'a. Size 'a => mword 'a -> bitU -> mword 'a
+val sadd_vec_bit : forall 'a. Size 'a => mword 'a -> bitU -> mword 'a
+val sub_vec_bit : forall 'a. Size 'a => mword 'a -> bitU -> mword 'a
+Definition add_vec_bit := add_bv_bit
+Definition sadd_vec_bit := sadd_bv_bit
+Definition sub_vec_bit := sub_bv_bit
+
+val add_overflow_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> (mword 'a * bitU * bitU)
+val add_overflow_vec_signed : forall 'a. Size 'a => mword 'a -> mword 'a -> (mword 'a * bitU * bitU)
+val sub_overflow_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> (mword 'a * bitU * bitU)
+val sub_overflow_vec_signed : forall 'a. Size 'a => mword 'a -> mword 'a -> (mword 'a * bitU * bitU)
+val mult_overflow_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> (mword 'a * bitU * bitU)
+val mult_overflow_vec_signed : forall 'a. Size 'a => mword 'a -> mword 'a -> (mword 'a * bitU * bitU)
+Definition add_overflow_vec := add_overflow_bv
+Definition add_overflow_vec_signed := add_overflow_bv_signed
+Definition sub_overflow_vec := sub_overflow_bv
+Definition sub_overflow_vec_signed := sub_overflow_bv_signed
+Definition mult_overflow_vec := mult_overflow_bv
+Definition mult_overflow_vec_signed := mult_overflow_bv_signed
+
+val add_overflow_vec_bit : forall 'a. Size 'a => mword 'a -> bitU -> (mword 'a * bitU * bitU)
+val add_overflow_vec_bit_signed : forall 'a. Size 'a => mword 'a -> bitU -> (mword 'a * bitU * bitU)
+val sub_overflow_vec_bit : forall 'a. Size 'a => mword 'a -> bitU -> (mword 'a * bitU * bitU)
+val sub_overflow_vec_bit_signed : forall 'a. Size 'a => mword 'a -> bitU -> (mword 'a * bitU * bitU)
+Definition add_overflow_vec_bit := add_overflow_bv_bit
+Definition add_overflow_vec_bit_signed := add_overflow_bv_bit_signed
+Definition sub_overflow_vec_bit := sub_overflow_bv_bit
+Definition sub_overflow_vec_bit_signed := sub_overflow_bv_bit_signed
+
+val shiftl : forall 'a. Size 'a => mword 'a -> integer -> mword 'a
+val shiftr : forall 'a. Size 'a => mword 'a -> integer -> mword 'a
+val arith_shiftr : forall 'a. Size 'a => mword 'a -> integer -> mword 'a
+val rotl : forall 'a. Size 'a => mword 'a -> integer -> mword 'a
+val rotr : forall 'a. Size 'a => mword 'a -> integer -> mword 'a*)
+(* TODO: check/redefine behaviour on out-of-range n *)
+Definition shiftl {a} (v : mword a) n : mword a := with_word (P := id) (fun w => Word.wlshift' w (Z.to_nat n)) v.
+Definition shiftr {a} (v : mword a) n : mword a := with_word (P := id) (fun w => Word.wrshift' w (Z.to_nat n)) v.
+Definition arith_shiftr {a} (v : mword a) n : mword a := with_word (P := id) (fun w => Word.wrshifta' w (Z.to_nat n)) v.
+(*
+Definition rotl := rotl_bv
+Definition rotr := rotr_bv
+
+val mod_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> mword 'a
+val quot_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> mword 'a
+val quot_vec_signed : forall 'a. Size 'a => mword 'a -> mword 'a -> mword 'a
+Definition mod_vec := mod_bv
+Definition quot_vec := quot_bv
+Definition quot_vec_signed := quot_bv_signed
+
+val mod_vec_int : forall 'a. Size 'a => mword 'a -> integer -> mword 'a
+val quot_vec_int : forall 'a. Size 'a => mword 'a -> integer -> mword 'a
+Definition mod_vec_int := mod_bv_int
+Definition quot_vec_int := quot_bv_int
+
+val replicate_bits : forall 'a 'b. Size 'a, Size 'b => mword 'a -> integer -> mword 'b*)
+Fixpoint replicate_bits_aux {a} (w : Word.word a) (n : nat) : Word.word (n * a) :=
+match n with
+| O => Word.WO
+| S m => Word.combine w (replicate_bits_aux w m)
+end.
+Lemma replicate_ok {n a} `{ArithFact (n >= 0)} `{ArithFact (a >= 0)} :
+ Z.of_nat (Z.to_nat n * Z.to_nat a) = a * n.
+destruct H. destruct H0.
+rewrite <- Z2Nat.id; auto with zarith.
+rewrite Z2Nat.inj_mul; auto with zarith.
+rewrite Nat.mul_comm. reflexivity.
+Qed.
+Definition replicate_bits {a} (w : mword a) (n : Z) `{ArithFact (n >= 0)} : mword (a * n) :=
+ cast_to_mword (replicate_bits_aux (get_word w) (Z.to_nat n)) replicate_ok.
+
+(*val duplicate : forall 'a. Size 'a => bitU -> integer -> mword 'a
+Definition duplicate := duplicate_bit_bv
+
+val eq_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> bool
+val neq_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> bool
+val ult_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> bool
+val slt_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> bool
+val ugt_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> bool
+val sgt_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> bool
+val ulteq_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> bool
+val slteq_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> bool
+val ugteq_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> bool
+val sgteq_vec : forall 'a. Size 'a => mword 'a -> mword 'a -> bool*)
+Definition eq_vec {n} (x : mword n) (y : mword n) : bool := Word.weqb (get_word x) (get_word y).
+Definition neq_vec {n} (x : mword n) (y : mword n) : bool := negb (eq_vec x y).
+(*Definition ult_vec := ult_bv.
+Definition slt_vec := slt_bv.
+Definition ugt_vec := ugt_bv.
+Definition sgt_vec := sgt_bv.
+Definition ulteq_vec := ulteq_bv.
+Definition slteq_vec := slteq_bv.
+Definition ugteq_vec := ugteq_bv.
+Definition sgteq_vec := sgteq_bv.
+
+*)
+
+Definition eq_vec_dec {n} : forall (x y : mword n), {x = y} + {x <> y}.
+refine (match n with
+| Z0 => _
+| Zpos m => _
+| Zneg m => _
+end).
+* simpl. apply Word.weq.
+* simpl. apply Word.weq.
+* simpl. destruct x.
+Defined.
+
+Instance Decidable_eq_mword {n} : forall (x y : mword n), Decidable (x = y) :=
+ Decidable_eq_from_dec eq_vec_dec.
+
+Program Fixpoint reverse_endianness_word {n} (bits : word n) : word n :=
+ match n with
+ | S (S (S (S (S (S (S (S m))))))) =>
+ combine
+ (reverse_endianness_word (split2 8 m bits))
+ (split1 8 m bits)
+ | _ => bits
+ end.
+Next Obligation.
+omega.
+Qed.
+
+Definition reverse_endianness {n} (bits : mword n) := with_word (P := id) reverse_endianness_word bits.
+
+Definition get_slice_int {a} `{ArithFact (a >= 0)} : Z -> Z -> Z -> mword a := get_slice_int_bv.
+
+Definition set_slice n m (v : mword n) x (w : mword m) : mword n :=
+ update_subrange_vec_dec v (x + m - 1) x w.
+
+Definition set_slice_int len n lo (v : mword len) : Z :=
+ let hi := lo + len - 1 in
+ (* We don't currently have a constraint on lo in the sail prelude, so let's
+ avoid one here. *)
+ if sumbool_of_bool (Z.gtb hi 0) then
+ let bs : mword (hi + 1) := mword_of_int n in
+ (int_of_mword true (update_subrange_vec_dec bs hi lo v))
+ else n.
+
+(* Variant of bitvector slicing for the ARM model with few constraints *)
+Definition slice {m} (v : mword m) lo len `{ArithFact (0 <= len)} : mword len :=
+ if sumbool_of_bool (orb (len =? 0) (lo <? 0))
+ then zeros len
+ else
+ if sumbool_of_bool (lo + len - 1 >=? m)
+ then if sumbool_of_bool (lo <? m)
+ then zero_extend (subrange_vec_dec v (m - 1) lo) len
+ else zeros len
+ else autocast (subrange_vec_dec v (lo + len - 1) lo).
+
+(*
+Lemma slice_is_ok m (v : mword m) lo len
+ (H1 : 0 <= lo) (H2 : 0 < len) (H3: lo + len < m) :
+ slice v lo len = autocast (subrange_vec_dec v (lo + len - 1) lo).
+unfold slice.
+destruct (sumbool_of_bool _).
+* exfalso.
+ unbool_comparisons.
+ omega.
+* destruct (sumbool_of_bool _).
+ + exfalso.
+ unbool_comparisons.
+ omega.
+ + f_equal.
+ f_equal.
+*)
+
+Import ListNotations.
+Definition count_leading_zeros {N : Z} (x : mword N) `{ArithFact (N >= 1)}
+: {n : Z & ArithFact (0 <= n /\ n <= N)} :=
+ let r : {n : Z & ArithFact (0 <= n /\ n <= N)} := build_ex N in
+ foreach_Z_up 0 (N - 1) 1 r
+ (fun i _ r =>
+ (if ((eq_vec (vec_of_bits [access_vec_dec x i] : mword 1) (vec_of_bits [B1] : mword 1)))
+ then build_ex
+ (Z.sub (Z.sub (length_mword x) i) 1)
+ : {n : Z & ArithFact (0 <= n /\ n <= N)}
+ else r))
+ .
+
+Definition prerr_bits {a} (s : string) (bs : mword a) : unit := tt.
+Definition print_bits {a} (s : string) (bs : mword a) : unit := tt.
diff --git a/prover_snapshots/coq/lib/sail/Sail2_prompt.v b/prover_snapshots/coq/lib/sail/Sail2_prompt.v
new file mode 100644
index 0000000..79bf87e
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_prompt.v
@@ -0,0 +1,229 @@
+(*Require Import Sail_impl_base*)
+Require Import Sail2_values.
+Require Import Sail2_prompt_monad.
+Require Export ZArith.Zwf.
+Require Import List.
+Import ListNotations.
+(*
+
+val iter_aux : forall 'rv 'a 'e. integer -> (integer -> 'a -> monad 'rv unit 'e) -> list 'a -> monad 'rv unit 'e
+let rec iter_aux i f xs = match xs with
+ | x :: xs -> f i x >> iter_aux (i + 1) f xs
+ | [] -> return ()
+ end
+
+declare {isabelle} termination_argument iter_aux = automatic
+
+val iteri : forall 'rv 'a 'e. (integer -> 'a -> monad 'rv unit 'e) -> list 'a -> monad 'rv unit 'e
+let iteri f xs = iter_aux 0 f xs
+
+val iter : forall 'rv 'a 'e. ('a -> monad 'rv unit 'e) -> list 'a -> monad 'rv unit 'e
+let iter f xs = iteri (fun _ x -> f x) xs
+
+val foreachM : forall 'a 'rv 'vars 'e.
+ list 'a -> 'vars -> ('a -> 'vars -> monad 'rv 'vars 'e) -> monad 'rv 'vars 'e*)
+Fixpoint foreachM {a rv Vars e} (l : list a) (vars : Vars) (body : a -> Vars -> monad rv Vars e) : monad rv Vars e :=
+match l with
+| [] => returnm vars
+| (x :: xs) =>
+ body x vars >>= fun vars =>
+ foreachM xs vars body
+end.
+
+Fixpoint foreach_ZM_up' {rv e Vars} from to step off n `{ArithFact (0 < step)} `{ArithFact (0 <= off)} (vars : Vars) (body : forall (z : Z) `(ArithFact (from <= z <= to)), Vars -> monad rv Vars e) {struct n} : monad rv Vars e :=
+ if sumbool_of_bool (from + off <=? to) then
+ match n with
+ | O => returnm vars
+ | S n => body (from + off) _ vars >>= fun vars => foreach_ZM_up' from to step (off + step) n vars body
+ end
+ else returnm vars.
+
+Fixpoint foreach_ZM_down' {rv e Vars} from to step off n `{ArithFact (0 < step)} `{ArithFact (off <= 0)} (vars : Vars) (body : forall (z : Z) `(ArithFact (to <= z <= from)), Vars -> monad rv Vars e) {struct n} : monad rv Vars e :=
+ if sumbool_of_bool (to <=? from + off) then
+ match n with
+ | O => returnm vars
+ | S n => body (from + off) _ vars >>= fun vars => foreach_ZM_down' from to step (off - step) n vars body
+ end
+ else returnm vars.
+
+Definition foreach_ZM_up {rv e Vars} from to step vars body `{ArithFact (0 < step)} :=
+ foreach_ZM_up' (rv := rv) (e := e) (Vars := Vars) from to step 0 (S (Z.abs_nat (from - to))) vars body.
+Definition foreach_ZM_down {rv e Vars} from to step vars body `{ArithFact (0 < step)} :=
+ foreach_ZM_down' (rv := rv) (e := e) (Vars := Vars) from to step 0 (S (Z.abs_nat (from - to))) vars body.
+
+(*declare {isabelle} termination_argument foreachM = automatic*)
+
+Definition genlistM {A RV E} (f : nat -> monad RV A E) (n : nat) : monad RV (list A) E :=
+ let indices := List.seq 0 n in
+ foreachM indices [] (fun n xs => (f n >>= (fun x => returnm (xs ++ [x])))).
+
+(*val and_boolM : forall 'rv 'e. monad 'rv bool 'e -> monad 'rv bool 'e -> monad 'rv bool 'e*)
+Definition and_boolM {rv E} (l : monad rv bool E) (r : monad rv bool E) : monad rv bool E :=
+ l >>= (fun l => if l then r else returnm false).
+
+Definition and_boolMP {rv E} {P Q R:bool->Prop} (x : monad rv {b:bool & ArithFact (P b)} E) (y : monad rv {b:bool & ArithFact (Q b)} E)
+ `{H:ArithFact (forall l r, P l -> (l = true -> Q r) -> R (andb l r))}
+ : monad rv {b:bool & ArithFact (R b)} E.
+refine (
+ x >>= fun '(existT _ x (Build_ArithFact _ p)) => (if x return P x -> _ then
+ fun p => y >>= fun '(existT _ y _) => returnm (existT _ y _)
+ else fun p => returnm (existT _ false _)) p
+).
+* constructor. destruct H. destruct a0. change y with (andb true y). auto.
+* constructor. destruct H. change false with (andb false false). apply fact.
+ assumption.
+ congruence.
+Defined.
+
+(*val or_boolM : forall 'rv 'e. monad 'rv bool 'e -> monad 'rv bool 'e -> monad 'rv bool 'e*)
+Definition or_boolM {rv E} (l : monad rv bool E) (r : monad rv bool E) : monad rv bool E :=
+ l >>= (fun l => if l then returnm true else r).
+Definition or_boolMP {rv E} {P Q R:bool -> Prop} (l : monad rv {b : bool & ArithFact (P b)} E) (r : monad rv {b : bool & ArithFact (Q b)} E)
+ `{ArithFact (forall l r, P l -> (l = false -> Q r) -> R (orb l r))}
+ : monad rv {b : bool & ArithFact (R b)} E.
+refine (
+ l >>= fun '(existT _ l (Build_ArithFact _ p)) =>
+ (if l return P l -> _ then fun p => returnm (existT _ true _)
+ else fun p => r >>= fun '(existT _ r _) => returnm (existT _ r _)) p
+).
+* constructor. destruct H. change true with (orb true true). apply fact. assumption. congruence.
+* constructor. destruct H. destruct a0. change r with (orb false r). auto.
+Defined.
+
+Definition build_trivial_ex {rv E} {T:Type} (x:monad rv T E) : monad rv {x : T & ArithFact True} E :=
+ x >>= fun x => returnm (existT _ x (Build_ArithFact _ I)).
+
+(*val bool_of_bitU_fail : forall 'rv 'e. bitU -> monad 'rv bool 'e*)
+Definition bool_of_bitU_fail {rv E} (b : bitU) : monad rv bool E :=
+match b with
+ | B0 => returnm false
+ | B1 => returnm true
+ | BU => Fail "bool_of_bitU"
+end.
+
+(*val bool_of_bitU_oracle : forall 'rv 'e. bitU -> monad 'rv bool 'e*)
+Definition bool_of_bitU_oracle {rv E} (b : bitU) : monad rv bool E :=
+match b with
+ | B0 => returnm false
+ | B1 => returnm true
+ | BU => undefined_bool tt
+end.
+
+(* For termination of recursive functions. We don't name assertions, so use
+ the type class mechanism to find it. *)
+Definition _limit_reduces {_limit} (_acc:Acc (Zwf 0) _limit) `{ArithFact (_limit >= 0)} : Acc (Zwf 0) (_limit - 1).
+refine (Acc_inv _acc _).
+destruct H.
+red.
+omega.
+Defined.
+
+(* A version of well-foundedness of measures with a guard to ensure that
+ definitions can be reduced without inspecting proofs, based on a coq-club
+ thread featuring Barras, Gonthier and Gregoire, see
+ https://sympa.inria.fr/sympa/arc/coq-club/2007-07/msg00014.html *)
+
+Fixpoint pos_guard_wf {A:Type} {R:A -> A -> Prop} (p:positive) : well_founded R -> well_founded R :=
+ match p with
+ | xH => fun wfR x => Acc_intro x (fun y _ => wfR y)
+ | xO p' => fun wfR x => let F := pos_guard_wf p' in Acc_intro x (fun y _ => F (F
+wfR) y)
+ | xI p' => fun wfR x => let F := pos_guard_wf p' in Acc_intro x (fun y _ => F (F
+wfR) y)
+ end.
+
+Definition Zwf_guarded (z:Z) : Acc (Zwf 0) z :=
+ Acc_intro _ (fun y H => match z with
+ | Zpos p => pos_guard_wf p (Zwf_well_founded _) _
+ | Zneg p => pos_guard_wf p (Zwf_well_founded _) _
+ | Z0 => Zwf_well_founded _ _
+ end).
+
+(*val whileM : forall 'rv 'vars 'e. 'vars -> ('vars -> monad 'rv bool 'e) ->
+ ('vars -> monad 'rv 'vars 'e) -> monad 'rv 'vars 'e*)
+Fixpoint whileMT' {RV Vars E} limit (vars : Vars) (cond : Vars -> monad RV bool E) (body : Vars -> monad RV Vars E) (acc : Acc (Zwf 0) limit) : monad RV Vars E :=
+ if Z_ge_dec limit 0 then
+ cond vars >>= fun cond_val =>
+ if cond_val then
+ body vars >>= fun vars => whileMT' (limit - 1) vars cond body (_limit_reduces acc)
+ else returnm vars
+ else Fail "Termination limit reached".
+
+Definition whileMT {RV Vars E} (vars : Vars) (measure : Vars -> Z) (cond : Vars -> monad RV bool E) (body : Vars -> monad RV Vars E) : monad RV Vars E :=
+ let limit := measure vars in
+ whileMT' limit vars cond body (Zwf_guarded limit).
+
+(*val untilM : forall 'rv 'vars 'e. 'vars -> ('vars -> monad 'rv bool 'e) ->
+ ('vars -> monad 'rv 'vars 'e) -> monad 'rv 'vars 'e*)
+Fixpoint untilMT' {RV Vars E} limit (vars : Vars) (cond : Vars -> monad RV bool E) (body : Vars -> monad RV Vars E) (acc : Acc (Zwf 0) limit) : monad RV Vars E :=
+ if Z_ge_dec limit 0 then
+ body vars >>= fun vars =>
+ cond vars >>= fun cond_val =>
+ if cond_val then returnm vars else untilMT' (limit - 1) vars cond body (_limit_reduces acc)
+ else Fail "Termination limit reached".
+
+Definition untilMT {RV Vars E} (vars : Vars) (measure : Vars -> Z) (cond : Vars -> monad RV bool E) (body : Vars -> monad RV Vars E) : monad RV Vars E :=
+ let limit := measure vars in
+ untilMT' limit vars cond body (Zwf_guarded limit).
+
+(*let write_two_regs r1 r2 vec =
+ let is_inc =
+ let is_inc_r1 = is_inc_of_reg r1 in
+ let is_inc_r2 = is_inc_of_reg r2 in
+ let () = ensure (is_inc_r1 = is_inc_r2)
+ "write_two_regs called with vectors of different direction" in
+ is_inc_r1 in
+
+ let (size_r1 : integer) = size_of_reg r1 in
+ let (start_vec : integer) = get_start vec in
+ let size_vec = length vec in
+ let r1_v =
+ if is_inc
+ then slice vec start_vec (size_r1 - start_vec - 1)
+ else slice vec start_vec (start_vec - size_r1 - 1) in
+ let r2_v =
+ if is_inc
+ then slice vec (size_r1 - start_vec) (size_vec - start_vec)
+ else slice vec (start_vec - size_r1) (start_vec - size_vec) in
+ write_reg r1 r1_v >> write_reg r2 r2_v*)
+
+Definition choose_bools {RV E} (descr : string) (n : nat) : monad RV (list bool) E :=
+ genlistM (fun _ => choose_bool descr) n.
+
+Definition choose {RV A E} (descr : string) (xs : list A) : monad RV A E :=
+ (* Use sufficiently many nondeterministically chosen bits and convert into an
+ index into the list *)
+ choose_bools descr (List.length xs) >>= fun bs =>
+ let idx := ((nat_of_bools bs) mod List.length xs)%nat in
+ match List.nth_error xs idx with
+ | Some x => returnm x
+ | None => Fail ("choose " ++ descr)
+ end.
+
+Definition internal_pick {rv a e} (xs : list a) : monad rv a e :=
+ choose "internal_pick" xs.
+
+Fixpoint undefined_word_nat {rv e} n : monad rv (Word.word n) e :=
+ match n with
+ | O => returnm Word.WO
+ | S m =>
+ choose_bool "undefined_word_nat" >>= fun b =>
+ undefined_word_nat m >>= fun t =>
+ returnm (Word.WS b t)
+ end.
+
+Definition undefined_bitvector {rv e} n `{ArithFact (n >= 0)} : monad rv (mword n) e :=
+ undefined_word_nat (Z.to_nat n) >>= fun w =>
+ returnm (word_to_mword w).
+
+(* If we need to build an existential after a monadic operation, assume that
+ we can do it entirely from the type. *)
+
+Definition build_ex_m {rv e} {T:Type} (x:monad rv T e) {P:T -> Prop} `{H:forall x, ArithFact (P x)} : monad rv {x : T & ArithFact (P x)} e :=
+ x >>= fun y => returnm (existT _ y (H y)).
+
+Definition projT1_m {rv e} {T:Type} {P:T -> Prop} (x: monad rv {x : T & P x} e) : monad rv T e :=
+ x >>= fun y => returnm (projT1 y).
+
+Definition derive_m {rv e} {T:Type} {P Q:T -> Prop} (x : monad rv {x : T & P x} e) `{forall x, ArithFact (P x) -> ArithFact (Q x)} : monad rv {x : T & (ArithFact (Q x))} e :=
+ x >>= fun y => returnm (build_ex (projT1 y)).
diff --git a/prover_snapshots/coq/lib/sail/Sail2_prompt_monad.v b/prover_snapshots/coq/lib/sail/Sail2_prompt_monad.v
new file mode 100644
index 0000000..b26a2ff
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_prompt_monad.v
@@ -0,0 +1,367 @@
+Require Import String.
+(*Require Import Sail_impl_base*)
+Require Import Sail2_instr_kinds.
+Require Import Sail2_values.
+Require bbv.Word.
+Import ListNotations.
+
+Definition register_name := string.
+Definition address := list bitU.
+
+Inductive monad regval a e :=
+ | Done : a -> monad regval a e
+ (* Read a number of bytes from memory, returned in little endian order,
+ with or without a tag. The first nat specifies the address, the second
+ the number of bytes. *)
+ | Read_mem : read_kind -> nat -> nat -> (list memory_byte -> monad regval a e) -> monad regval a e
+ | Read_memt : read_kind -> nat -> nat -> ((list memory_byte * bitU) -> monad regval a e) -> monad regval a e
+ (* Tell the system a write is imminent, at the given address and with the
+ given size. *)
+ | Write_ea : write_kind -> nat -> nat -> monad regval a e -> monad regval a e
+ (* Request the result : store-exclusive *)
+ | Excl_res : (bool -> monad regval a e) -> monad regval a e
+ (* Request to write a memory value of the given size at the given address,
+ with or without a tag. *)
+ | Write_mem : write_kind -> nat -> nat -> list memory_byte -> (bool -> monad regval a e) -> monad regval a e
+ | Write_memt : write_kind -> nat -> nat -> list memory_byte -> bitU -> (bool -> monad regval a e) -> monad regval a e
+ (* Tell the system to dynamically recalculate dependency footprint *)
+ | Footprint : monad regval a e -> monad regval a e
+ (* Request a memory barrier *)
+ | Barrier : barrier_kind -> monad regval a e -> monad regval a e
+ (* Request to read register, will track dependency when mode.track_values *)
+ | Read_reg : register_name -> (regval -> monad regval a e) -> monad regval a e
+ (* Request to write register *)
+ | Write_reg : register_name -> regval -> monad regval a e -> monad regval a e
+ (* Request to choose a Boolean, e.g. to resolve an undefined bit. The string
+ argument may be used to provide information to the system about what the
+ Boolean is going to be used for. *)
+ | Choose : string -> (bool -> monad regval a e) -> monad regval a e
+ (* Print debugging or tracing information *)
+ | Print : string -> monad regval a e -> monad regval a e
+ (*Result of a failed assert with possible error message to report*)
+ | Fail : string -> monad regval a e
+ (* Exception of type e *)
+ | Exception : e -> monad regval a e.
+
+Arguments Done [_ _ _].
+Arguments Read_mem [_ _ _].
+Arguments Read_memt [_ _ _].
+Arguments Write_ea [_ _ _].
+Arguments Excl_res [_ _ _].
+Arguments Write_mem [_ _ _].
+Arguments Write_memt [_ _ _].
+Arguments Footprint [_ _ _].
+Arguments Barrier [_ _ _].
+Arguments Read_reg [_ _ _].
+Arguments Write_reg [_ _ _].
+Arguments Choose [_ _ _].
+Arguments Print [_ _ _].
+Arguments Fail [_ _ _].
+Arguments Exception [_ _ _].
+
+Inductive event {regval} :=
+ | E_read_mem : read_kind -> nat -> nat -> list memory_byte -> event
+ | E_read_memt : read_kind -> nat -> nat -> (list memory_byte * bitU) -> event
+ | E_write_mem : write_kind -> nat -> nat -> list memory_byte -> bool -> event
+ | E_write_memt : write_kind -> nat -> nat -> list memory_byte -> bitU -> bool -> event
+ | E_write_ea : write_kind -> nat -> nat -> event
+ | E_excl_res : bool -> event
+ | E_barrier : barrier_kind -> event
+ | E_footprint : event
+ | E_read_reg : register_name -> regval -> event
+ | E_write_reg : register_name -> regval -> event
+ | E_choose : string -> bool -> event
+ | E_print : string -> event.
+Arguments event : clear implicits.
+
+Definition trace regval := list (event regval).
+
+(*val return : forall rv a e. a -> monad rv a e*)
+Definition returnm {rv A E} (a : A) : monad rv A E := Done a.
+
+(*val bind : forall rv a b e. monad rv a e -> (a -> monad rv b e) -> monad rv b e*)
+Fixpoint bind {rv A B E} (m : monad rv A E) (f : A -> monad rv B E) := match m with
+ | Done a => f a
+ | Read_mem rk a sz k => Read_mem rk a sz (fun v => bind (k v) f)
+ | Read_memt rk a sz k => Read_memt rk a sz (fun v => bind (k v) f)
+ | Write_mem wk a sz v k => Write_mem wk a sz v (fun v => bind (k v) f)
+ | Write_memt wk a sz v t k => Write_memt wk a sz v t (fun v => bind (k v) f)
+ | Read_reg descr k => Read_reg descr (fun v => bind (k v) f)
+ | Excl_res k => Excl_res (fun v => bind (k v) f)
+ | Choose descr k => Choose descr (fun v => bind (k v) f)
+ | Write_ea wk a sz k => Write_ea wk a sz (bind k f)
+ | Footprint k => Footprint (bind k f)
+ | Barrier bk k => Barrier bk (bind k f)
+ | Write_reg r v k => Write_reg r v (bind k f)
+ | Print msg k => Print msg (bind k f)
+ | Fail descr => Fail descr
+ | Exception e => Exception e
+end.
+
+Notation "m >>= f" := (bind m f) (at level 50, left associativity).
+(*val (>>) : forall rv b e. monad rv unit e -> monad rv b e -> monad rv b e*)
+Definition bind0 {rv A E} (m : monad rv unit E) (n : monad rv A E) :=
+ m >>= fun (_ : unit) => n.
+Notation "m >> n" := (bind0 m n) (at level 50, left associativity).
+
+(*val exit : forall rv a e. unit -> monad rv a e*)
+Definition exit {rv A E} (_ : unit) : monad rv A E := Fail "exit".
+
+(*val choose_bool : forall 'rv 'e. string -> monad 'rv bool 'e*)
+Definition choose_bool {rv E} descr : monad rv bool E := Choose descr returnm.
+
+(*val undefined_bool : forall 'rv 'e. unit -> monad 'rv bool 'e*)
+Definition undefined_bool {rv e} (_:unit) : monad rv bool e := choose_bool "undefined_bool".
+
+Definition undefined_unit {rv e} (_:unit) : monad rv unit e := returnm tt.
+
+(*val assert_exp : forall rv e. bool -> string -> monad rv unit e*)
+Definition assert_exp {rv E} (exp :bool) msg : monad rv unit E :=
+ if exp then Done tt else Fail msg.
+
+Definition assert_exp' {rv E} (exp :bool) msg : monad rv (exp = true) E :=
+ if exp return monad rv (exp = true) E then Done eq_refl else Fail msg.
+Definition bindH {rv A P E} (m : monad rv P E) (n : monad rv A E) :=
+ m >>= fun (H : P) => n.
+Notation "m >>> n" := (bindH m n) (at level 50, left associativity).
+
+(*val throw : forall rv a e. e -> monad rv a e*)
+Definition throw {rv A E} e : monad rv A E := Exception e.
+
+(*val try_catch : forall rv a e1 e2. monad rv a e1 -> (e1 -> monad rv a e2) -> monad rv a e2*)
+Fixpoint try_catch {rv A E1 E2} (m : monad rv A E1) (h : E1 -> monad rv A E2) := match m with
+ | Done a => Done a
+ | Read_mem rk a sz k => Read_mem rk a sz (fun v => try_catch (k v) h)
+ | Read_memt rk a sz k => Read_memt rk a sz (fun v => try_catch (k v) h)
+ | Write_mem wk a sz v k => Write_mem wk a sz v (fun v => try_catch (k v) h)
+ | Write_memt wk a sz v t k => Write_memt wk a sz v t (fun v => try_catch (k v) h)
+ | Read_reg descr k => Read_reg descr (fun v => try_catch (k v) h)
+ | Excl_res k => Excl_res (fun v => try_catch (k v) h)
+ | Choose descr k => Choose descr (fun v => try_catch (k v) h)
+ | Write_ea wk a sz k => Write_ea wk a sz (try_catch k h)
+ | Footprint k => Footprint (try_catch k h)
+ | Barrier bk k => Barrier bk (try_catch k h)
+ | Write_reg r v k => Write_reg r v (try_catch k h)
+ | Print msg k => Print msg (try_catch k h)
+ | Fail descr => Fail descr
+ | Exception e => h e
+end.
+
+(* For early return, we abuse exceptions by throwing and catching
+ the return value. The exception type is "either r e", where "inr e"
+ represents a proper exception and "inl r" an early return : value "r". *)
+Definition monadR rv a r e := monad rv a (sum r e).
+
+(*val early_return : forall rv a r e. r -> monadR rv a r e*)
+Definition early_return {rv A R E} (r : R) : monadR rv A R E := throw (inl r).
+
+(*val catch_early_return : forall rv a e. monadR rv a a e -> monad rv a e*)
+Definition catch_early_return {rv A E} (m : monadR rv A A E) :=
+ try_catch m
+ (fun r => match r with
+ | inl a => returnm a
+ | inr e => throw e
+ end).
+
+(* Lift to monad with early return by wrapping exceptions *)
+(*val liftR : forall rv a r e. monad rv a e -> monadR rv a r e*)
+Definition liftR {rv A R E} (m : monad rv A E) : monadR rv A R E :=
+ try_catch m (fun e => throw (inr e)).
+
+(* Catch exceptions in the presence : early returns *)
+(*val try_catchR : forall rv a r e1 e2. monadR rv a r e1 -> (e1 -> monadR rv a r e2) -> monadR rv a r e2*)
+Definition try_catchR {rv A R E1 E2} (m : monadR rv A R E1) (h : E1 -> monadR rv A R E2) :=
+ try_catch m
+ (fun r => match r with
+ | inl r => throw (inl r)
+ | inr e => h e
+ end).
+
+(*val maybe_fail : forall 'rv 'a 'e. string -> maybe 'a -> monad 'rv 'a 'e*)
+Definition maybe_fail {rv A E} msg (x : option A) : monad rv A E :=
+match x with
+ | Some a => returnm a
+ | None => Fail msg
+end.
+
+(*val read_memt_bytes : forall 'rv 'a 'b 'e. Bitvector 'a, Bitvector 'b => read_kind -> 'a -> integer -> monad 'rv (list memory_byte * bitU) 'e*)
+Definition read_memt_bytes {rv A E} rk (addr : mword A) sz : monad rv (list memory_byte * bitU) E :=
+ Read_memt rk (Word.wordToNat (get_word addr)) (Z.to_nat sz) returnm.
+
+(*val read_memt : forall 'rv 'a 'b 'e. Bitvector 'a, Bitvector 'b => read_kind -> 'a -> integer -> monad 'rv ('b * bitU) 'e*)
+Definition read_memt {rv A B E} `{ArithFact (B >= 0)} rk (addr : mword A) sz : monad rv (mword B * bitU) E :=
+ bind
+ (read_memt_bytes rk addr sz)
+ (fun '(bytes, tag) =>
+ match of_bits (bits_of_mem_bytes bytes) with
+ | Some v => returnm (v, tag)
+ | None => Fail "bits_of_mem_bytes"
+ end).
+
+(*val read_mem_bytes : forall 'rv 'a 'b 'e. Bitvector 'a, Bitvector 'b => read_kind -> 'a -> integer -> monad 'rv (list memory_byte) 'e*)
+Definition read_mem_bytes {rv A E} rk (addr : mword A) sz : monad rv (list memory_byte) E :=
+ Read_mem rk (Word.wordToNat (get_word addr)) (Z.to_nat sz) returnm.
+
+(*val read_mem : forall 'rv 'a 'b 'e. Bitvector 'a, Bitvector 'b => read_kind -> 'a -> integer -> monad 'rv 'b 'e*)
+Definition read_mem {rv A B E} `{ArithFact (B >= 0)} rk (addrsz : Z) (addr : mword A) sz : monad rv (mword B) E :=
+ bind
+ (read_mem_bytes rk addr sz)
+ (fun bytes =>
+ maybe_fail "bits_of_mem_bytes" (of_bits (bits_of_mem_bytes bytes))).
+
+(*val excl_result : forall rv e. unit -> monad rv bool e*)
+Definition excl_result {rv e} (_:unit) : monad rv bool e :=
+ let k successful := (returnm successful) in
+ Excl_res k.
+
+Definition write_mem_ea {rv a E} wk (addrsz : Z) (addr: mword a) sz : monad rv unit E :=
+ Write_ea wk (Word.wordToNat (get_word addr)) (Z.to_nat sz) (Done tt).
+
+(*val write_mem : forall 'rv 'a 'b 'e. Bitvector 'a, Bitvector 'b =>
+ write_kind -> integer -> 'a -> integer -> 'b -> monad 'rv bool 'e*)
+Definition write_mem {rv a b E} wk (addrsz : Z) (addr : mword a) sz (v : mword b) : monad rv bool E :=
+ match (mem_bytes_of_bits v, Word.wordToNat (get_word addr)) with
+ | (Some v, addr) =>
+ Write_mem wk addr (Z.to_nat sz) v returnm
+ | _ => Fail "write_mem"
+ end.
+
+(*val write_memt : forall 'rv 'a 'b 'e. Bitvector 'a, Bitvector 'b =>
+ write_kind -> 'a -> integer -> 'b -> bitU -> monad 'rv bool 'e*)
+Definition write_memt {rv a b E} wk (addr : mword a) sz (v : mword b) tag : monad rv bool E :=
+ match (mem_bytes_of_bits v, Word.wordToNat (get_word addr)) with
+ | (Some v, addr) =>
+ Write_memt wk addr (Z.to_nat sz) v tag returnm
+ | _ => Fail "write_mem"
+ end.
+
+Definition read_reg {s rv a e} (reg : register_ref s rv a) : monad rv a e :=
+ let k v :=
+ match reg.(of_regval) v with
+ | Some v => Done v
+ | None => Fail "read_reg: unrecognised value"
+ end
+ in
+ Read_reg reg.(name) k.
+
+(* TODO
+val read_reg_range : forall 's 'r 'rv 'a 'e. Bitvector 'a => register_ref 's 'rv 'r -> integer -> integer -> monad 'rv 'a 'e
+let read_reg_range reg i j =
+ read_reg_aux of_bits (external_reg_slice reg (nat_of_int i,nat_of_int j))
+
+let read_reg_bit reg i =
+ read_reg_aux (fun v -> v) (external_reg_slice reg (nat_of_int i,nat_of_int i)) >>= fun v ->
+ return (extract_only_element v)
+
+let read_reg_field reg regfield =
+ read_reg_aux (external_reg_field_whole reg regfield)
+
+let read_reg_bitfield reg regfield =
+ read_reg_aux (external_reg_field_whole reg regfield) >>= fun v ->
+ return (extract_only_element v)*)
+
+Definition reg_deref {s rv a e} := @read_reg s rv a e.
+
+(*Parameter write_reg : forall {s rv a e}, register_ref s rv a -> a -> monad rv unit e.*)
+Definition write_reg {s rv a e} (reg : register_ref s rv a) (v : a) : monad rv unit e :=
+ Write_reg reg.(name) (reg.(regval_of) v) (Done tt).
+
+(* TODO
+let write_reg reg v =
+ write_reg_aux (external_reg_whole reg) v
+let write_reg_range reg i j v =
+ write_reg_aux (external_reg_slice reg (nat_of_int i,nat_of_int j)) v
+let write_reg_pos reg i v =
+ let iN = nat_of_int i in
+ write_reg_aux (external_reg_slice reg (iN,iN)) [v]
+let write_reg_bit = write_reg_pos
+let write_reg_field reg regfield v =
+ write_reg_aux (external_reg_field_whole reg regfield.field_name) v
+let write_reg_field_bit reg regfield bit =
+ write_reg_aux (external_reg_field_whole reg regfield.field_name)
+ (Vector [bit] 0 (is_inc_of_reg reg))
+let write_reg_field_range reg regfield i j v =
+ write_reg_aux (external_reg_field_slice reg regfield.field_name (nat_of_int i,nat_of_int j)) v
+let write_reg_field_pos reg regfield i v =
+ write_reg_field_range reg regfield i i [v]
+let write_reg_field_bit = write_reg_field_pos*)
+
+(*val barrier : forall rv e. barrier_kind -> monad rv unit e*)
+Definition barrier {rv e} bk : monad rv unit e := Barrier bk (Done tt).
+
+(*val footprint : forall rv e. unit -> monad rv unit e*)
+Definition footprint {rv e} (_ : unit) : monad rv unit e := Footprint (Done tt).
+
+(* Event traces *)
+
+Local Open Scope bool_scope.
+
+(*val emitEvent : forall 'regval 'a 'e. Eq 'regval => monad 'regval 'a 'e -> event 'regval -> maybe (monad 'regval 'a 'e)*)
+Definition emitEvent {Regval A E} `{forall (x y : Regval), Decidable (x = y)} (m : monad Regval A E) (e : event Regval) : option (monad Regval A E) :=
+ match (e, m) with
+ | (E_read_mem rk a sz v, Read_mem rk' a' sz' k) =>
+ if read_kind_beq rk' rk && Nat.eqb a' a && Nat.eqb sz' sz then Some (k v) else None
+ | (E_read_memt rk a sz vt, Read_memt rk' a' sz' k) =>
+ if read_kind_beq rk' rk && Nat.eqb a' a && Nat.eqb sz' sz then Some (k vt) else None
+ | (E_write_mem wk a sz v r, Write_mem wk' a' sz' v' k) =>
+ if write_kind_beq wk' wk && Nat.eqb a' a && Nat.eqb sz' sz && generic_eq v' v then Some (k r) else None
+ | (E_write_memt wk a sz v tag r, Write_memt wk' a' sz' v' tag' k) =>
+ if write_kind_beq wk' wk && Nat.eqb a' a && Nat.eqb sz' sz && generic_eq v' v && generic_eq tag' tag then Some (k r) else None
+ | (E_read_reg r v, Read_reg r' k) =>
+ if generic_eq r' r then Some (k v) else None
+ | (E_write_reg r v, Write_reg r' v' k) =>
+ if generic_eq r' r && generic_eq v' v then Some k else None
+ | (E_write_ea wk a sz, Write_ea wk' a' sz' k) =>
+ if write_kind_beq wk' wk && Nat.eqb a' a && Nat.eqb sz' sz then Some k else None
+ | (E_barrier bk, Barrier bk' k) =>
+ if barrier_kind_beq bk' bk then Some k else None
+ | (E_print m, Print m' k) =>
+ if generic_eq m' m then Some k else None
+ | (E_excl_res v, Excl_res k) => Some (k v)
+ | (E_choose descr v, Choose descr' k) => if generic_eq descr' descr then Some (k v) else None
+ | (E_footprint, Footprint k) => Some k
+ | _ => None
+end.
+
+Definition option_bind {A B : Type} (a : option A) (f : A -> option B) : option B :=
+match a with
+| Some x => f x
+| None => None
+end.
+
+(*val runTrace : forall 'regval 'a 'e. Eq 'regval => trace 'regval -> monad 'regval 'a 'e -> maybe (monad 'regval 'a 'e)*)
+Fixpoint runTrace {Regval A E} `{forall (x y : Regval), Decidable (x = y)} (t : trace Regval) (m : monad Regval A E) : option (monad Regval A E) :=
+match t with
+ | [] => Some m
+ | e :: t' => option_bind (emitEvent m e) (runTrace t')
+end.
+
+(*val final : forall 'regval 'a 'e. monad 'regval 'a 'e -> bool*)
+Definition final {Regval A E} (m : monad Regval A E) : bool :=
+match m with
+ | Done _ => true
+ | Fail _ => true
+ | Exception _ => true
+ | _ => false
+end.
+
+(*val hasTrace : forall 'regval 'a 'e. Eq 'regval => trace 'regval -> monad 'regval 'a 'e -> bool*)
+Definition hasTrace {Regval A E} `{forall (x y : Regval), Decidable (x = y)} (t : trace Regval) (m : monad Regval A E) : bool :=
+match runTrace t m with
+ | Some m => final m
+ | None => false
+end.
+
+(*val hasException : forall 'regval 'a 'e. Eq 'regval => trace 'regval -> monad 'regval 'a 'e -> bool*)
+Definition hasException {Regval A E} `{forall (x y : Regval), Decidable (x = y)} (t : trace Regval) (m : monad Regval A E) :=
+match runTrace t m with
+ | Some (Exception _) => true
+ | _ => false
+end.
+
+(*val hasFailure : forall 'regval 'a 'e. Eq 'regval => trace 'regval -> monad 'regval 'a 'e -> bool*)
+Definition hasFailure {Regval A E} `{forall (x y : Regval), Decidable (x = y)} (t : trace Regval) (m : monad Regval A E) :=
+match runTrace t m with
+ | Some (Fail _) => true
+ | _ => false
+end.
diff --git a/prover_snapshots/coq/lib/sail/Sail2_real.v b/prover_snapshots/coq/lib/sail/Sail2_real.v
new file mode 100644
index 0000000..494e36d
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_real.v
@@ -0,0 +1,36 @@
+Require Export Rbase.
+Require Import Reals.
+Require Export ROrderedType.
+Require Import Sail2_values.
+
+(* "Decidable" in a classical sense... *)
+Instance Decidable_eq_real : forall (x y : R), Decidable (x = y) :=
+ Decidable_eq_from_dec Req_dec.
+
+Definition realFromFrac (num denom : Z) : R := Rdiv (IZR num) (IZR denom).
+
+Definition neg_real := Ropp.
+Definition mult_real := Rmult.
+Definition sub_real := Rminus.
+Definition add_real := Rplus.
+Definition div_real := Rdiv.
+Definition sqrt_real := sqrt.
+Definition abs_real := Rabs.
+
+(* Use flocq definitions, but without making the whole library a dependency. *)
+Definition round_down (x : R) := (up x - 1)%Z.
+Definition round_up (x : R) := (- round_down (- x))%Z.
+
+Definition to_real := IZR.
+
+Definition eq_real := Reqb.
+Definition gteq_real (x y : R) : bool := if Rge_dec x y then true else false.
+Definition lteq_real (x y : R) : bool := if Rle_dec x y then true else false.
+Definition gt_real (x y : R) : bool := if Rgt_dec x y then true else false.
+Definition lt_real (x y : R) : bool := if Rlt_dec x y then true else false.
+
+(* Export select definitions from outside of Rbase *)
+Definition pow_real := powerRZ.
+
+Definition print_real (_ : string) (_ : R) : unit := tt.
+Definition prerr_real (_ : string) (_ : R) : unit := tt.
diff --git a/prover_snapshots/coq/lib/sail/Sail2_state.v b/prover_snapshots/coq/lib/sail/Sail2_state.v
new file mode 100644
index 0000000..dc635cb
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_state.v
@@ -0,0 +1,167 @@
+(*Require Import Sail_impl_base*)
+Require Import Sail2_values.
+Require Import Sail2_prompt_monad.
+Require Import Sail2_prompt.
+Require Import Sail2_state_monad.
+Import ListNotations.
+
+(*val iterS_aux : forall 'rv 'a 'e. integer -> (integer -> 'a -> monadS 'rv unit 'e) -> list 'a -> monadS 'rv unit 'e*)
+Fixpoint iterS_aux {RV A E} i (f : Z -> A -> monadS RV unit E) (xs : list A) :=
+ match xs with
+ | x :: xs => f i x >>$ iterS_aux (i + 1) f xs
+ | [] => returnS tt
+ end.
+
+(*val iteriS : forall 'rv 'a 'e. (integer -> 'a -> monadS 'rv unit 'e) -> list 'a -> monadS 'rv unit 'e*)
+Definition iteriS {RV A E} (f : Z -> A -> monadS RV unit E) (xs : list A) : monadS RV unit E :=
+ iterS_aux 0 f xs.
+
+(*val iterS : forall 'rv 'a 'e. ('a -> monadS 'rv unit 'e) -> list 'a -> monadS 'rv unit 'e*)
+Definition iterS {RV A E} (f : A -> monadS RV unit E) (xs : list A) : monadS RV unit E :=
+ iteriS (fun _ x => f x) xs.
+
+(*val foreachS : forall 'a 'rv 'vars 'e.
+ list 'a -> 'vars -> ('a -> 'vars -> monadS 'rv 'vars 'e) -> monadS 'rv 'vars 'e*)
+Fixpoint foreachS {A RV Vars E} (xs : list A) (vars : Vars) (body : A -> Vars -> monadS RV Vars E) : monadS RV Vars E :=
+ match xs with
+ | [] => returnS vars
+ | x :: xs =>
+ body x vars >>$= fun vars =>
+ foreachS xs vars body
+end.
+
+(*val genlistS : forall 'a 'rv 'e. (nat -> monadS 'rv 'a 'e) -> nat -> monadS 'rv (list 'a) 'e*)
+Definition genlistS {A RV E} (f : nat -> monadS RV A E) n : monadS RV (list A) E :=
+ let indices := List.seq 0 n in
+ foreachS indices [] (fun n xs => (f n >>$= (fun x => returnS (xs ++ [x])))).
+
+(*val and_boolS : forall 'rv 'e. monadS 'rv bool 'e -> monadS 'rv bool 'e -> monadS 'rv bool 'e*)
+Definition and_boolS {RV E} (l r : monadS RV bool E) : monadS RV bool E :=
+ l >>$= (fun l => if l then r else returnS false).
+
+(*val or_boolS : forall 'rv 'e. monadS 'rv bool 'e -> monadS 'rv bool 'e -> monadS 'rv bool 'e*)
+Definition or_boolS {RV E} (l r : monadS RV bool E) : monadS RV bool E :=
+ l >>$= (fun l => if l then returnS true else r).
+
+Definition and_boolSP {rv E} {P Q R:bool->Prop} (x : monadS rv {b:bool & ArithFact (P b)} E) (y : monadS rv {b:bool & ArithFact (Q b)} E)
+ `{H:ArithFact (forall l r, P l -> (l = true -> Q r) -> R (andb l r))}
+ : monadS rv {b:bool & ArithFact (R b)} E.
+refine (
+ x >>$= fun '(existT _ x (Build_ArithFact _ p)) => (if x return P x -> _ then
+ fun p => y >>$= fun '(existT _ y _) => returnS (existT _ y _)
+ else fun p => returnS (existT _ false _)) p
+).
+* constructor. destruct H. destruct a0. change y with (andb true y). auto.
+* constructor. destruct H. change false with (andb false false). apply fact.
+ assumption.
+ congruence.
+Defined.
+Definition or_boolSP {rv E} {P Q R:bool -> Prop} (l : monadS rv {b : bool & ArithFact (P b)} E) (r : monadS rv {b : bool & ArithFact (Q b)} E)
+ `{ArithFact (forall l r, P l -> (l = false -> Q r) -> R (orb l r))}
+ : monadS rv {b : bool & ArithFact (R b)} E.
+refine (
+ l >>$= fun '(existT _ l (Build_ArithFact _ p)) =>
+ (if l return P l -> _ then fun p => returnS (existT _ true _)
+ else fun p => r >>$= fun '(existT _ r _) => returnS (existT _ r _)) p
+).
+* constructor. destruct H. change true with (orb true true). apply fact. assumption. congruence.
+* constructor. destruct H. destruct a0. change r with (orb false r). auto.
+Defined.
+
+(*val bool_of_bitU_fail : forall 'rv 'e. bitU -> monadS 'rv bool 'e*)
+Definition bool_of_bitU_fail {RV E} (b : bitU) : monadS RV bool E :=
+match b with
+ | B0 => returnS false
+ | B1 => returnS true
+ | BU => failS "bool_of_bitU"
+end.
+
+(*val bool_of_bitU_nondetS : forall 'rv 'e. bitU -> monadS 'rv bool 'e*)
+Definition bool_of_bitU_nondetS {RV E} (b : bitU) : monadS RV bool E :=
+match b with
+ | B0 => returnS false
+ | B1 => returnS true
+ | BU => undefined_boolS tt
+end.
+
+(*val bools_of_bits_nondetS : forall 'rv 'e. list bitU -> monadS 'rv (list bool) 'e*)
+Definition bools_of_bits_nondetS {RV E} bits : monadS RV (list bool) E :=
+ foreachS bits []
+ (fun b bools =>
+ bool_of_bitU_nondetS b >>$= (fun b =>
+ returnS (bools ++ [b]))).
+
+(*val of_bits_nondetS : forall 'rv 'a 'e. Bitvector 'a => list bitU -> monadS 'rv 'a 'e*)
+Definition of_bits_nondetS {RV A E} bits `{ArithFact (A >= 0)} : monadS RV (mword A) E :=
+ bools_of_bits_nondetS bits >>$= (fun bs =>
+ returnS (of_bools bs)).
+
+(*val of_bits_failS : forall 'rv 'a 'e. Bitvector 'a => list bitU -> monadS 'rv 'a 'e*)
+Definition of_bits_failS {RV A E} bits `{ArithFact (A >= 0)} : monadS RV (mword A) E :=
+ maybe_failS "of_bits" (of_bits bits).
+
+(*val mword_nondetS : forall 'rv 'a 'e. Size 'a => unit -> monadS 'rv (mword 'a) 'e
+let mword_nondetS () =
+ bools_of_bits_nondetS (repeat [BU] (integerFromNat size)) >>$= (fun bs ->
+ returnS (wordFromBitlist bs))
+
+
+val whileS : forall 'rv 'vars 'e. 'vars -> ('vars -> monadS 'rv bool 'e) ->
+ ('vars -> monadS 'rv 'vars 'e) -> monadS 'rv 'vars 'e
+let rec whileS vars cond body s =
+ (cond vars >>$= (fun cond_val s' ->
+ if cond_val then
+ (body vars >>$= (fun vars s'' -> whileS vars cond body s'')) s'
+ else returnS vars s')) s
+
+val untilS : forall 'rv 'vars 'e. 'vars -> ('vars -> monadS 'rv bool 'e) ->
+ ('vars -> monadS 'rv 'vars 'e) -> monadS 'rv 'vars 'e
+let rec untilS vars cond body s =
+ (body vars >>$= (fun vars s' ->
+ (cond vars >>$= (fun cond_val s'' ->
+ if cond_val then returnS vars s'' else untilS vars cond body s'')) s')) s
+*)
+
+Fixpoint whileST' {RV Vars E} limit (vars : Vars) (cond : Vars -> monadS RV bool E) (body : Vars -> monadS RV Vars E) (acc : Acc (Zwf 0) limit) : monadS RV Vars E :=
+ if Z_ge_dec limit 0 then
+ cond vars >>$= fun cond_val =>
+ if cond_val then
+ body vars >>$= fun vars => whileST' (limit - 1) vars cond body (_limit_reduces acc)
+ else returnS vars
+ else failS "Termination limit reached".
+
+Definition whileST {RV Vars E} (vars : Vars) measure (cond : Vars -> monadS RV bool E) (body : Vars -> monadS RV Vars E) : monadS RV Vars E :=
+ let limit := measure vars in
+ whileST' limit vars cond body (Zwf_guarded limit).
+
+(*val untilM : forall 'rv 'vars 'e. 'vars -> ('vars -> monad 'rv bool 'e) ->
+ ('vars -> monad 'rv 'vars 'e) -> monad 'rv 'vars 'e*)
+Fixpoint untilST' {RV Vars E} limit (vars : Vars) (cond : Vars -> monadS RV bool E) (body : Vars -> monadS RV Vars E) (acc : Acc (Zwf 0) limit) : monadS RV Vars E :=
+ if Z_ge_dec limit 0 then
+ body vars >>$= fun vars =>
+ cond vars >>$= fun cond_val =>
+ if cond_val then returnS vars else untilST' (limit - 1) vars cond body (_limit_reduces acc)
+ else failS "Termination limit reached".
+
+Definition untilST {RV Vars E} (vars : Vars) measure (cond : Vars -> monadS RV bool E) (body : Vars -> monadS RV Vars E) : monadS RV Vars E :=
+ let limit := measure vars in
+ untilST' limit vars cond body (Zwf_guarded limit).
+
+
+(*val choose_boolsS : forall 'rv 'e. nat -> monadS 'rv (list bool) 'e*)
+Definition choose_boolsS {RV E} n : monadS RV (list bool) E :=
+ genlistS (fun _ => choose_boolS tt) n.
+
+(* TODO: Replace by chooseS and prove equivalence to prompt monad version *)
+(*val internal_pickS : forall 'rv 'a 'e. list 'a -> monadS 'rv 'a 'e*)
+Definition internal_pickS {RV A E} (xs : list A) : monadS RV A E :=
+ (* Use sufficiently many nondeterministically chosen bits and convert into an
+ index into the list *)
+ choose_boolsS (List.length xs) >>$= fun bs =>
+ let idx := ((nat_of_bools bs) mod List.length xs)%nat in
+ match List.nth_error xs idx with
+ | Some x => returnS x
+ | None => failS "choose internal_pick"
+ end.
+
+
diff --git a/prover_snapshots/coq/lib/sail/Sail2_state_lemmas.v b/prover_snapshots/coq/lib/sail/Sail2_state_lemmas.v
new file mode 100644
index 0000000..c07016d
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_state_lemmas.v
@@ -0,0 +1,819 @@
+Require Import Sail2_values Sail2_prompt_monad Sail2_prompt Sail2_state_monad Sail2_state Sail2_state Sail2_state_lifting.
+Require Import Sail2_state_monad_lemmas.
+
+Local Open Scope equiv_scope.
+
+(* Monad lifting *)
+
+Lemma liftState_bind Regval Regs A B E {r : Sail2_values.register_accessors Regs Regval} {m : monad Regval A E} {f : A -> monad Regval B E} :
+ liftState r (bind m f) === bindS (liftState r m) (fun x => liftState r (f x)).
+induction m; simpl; autorewrite with state; auto using bindS_cong.
+Qed.
+Hint Rewrite liftState_bind : liftState.
+
+(* TODO: I want a general tactic for this, but abstracting the hint db out
+ appears to break.
+ This does beta reduction when no rules apply to try and allow more rules to apply
+ (e.g., the application of f to x in the above lemma may introduce a beta redex). *)
+Ltac rewrite_liftState := rewrite_strat topdown (choice (progress try hints liftState) progress eval cbn beta).
+
+Lemma liftState_return Regval Regs A E {r : Sail2_values.register_accessors Regs Regval} {a :A} :
+ liftState (E:=E) r (returnm a) = returnS a.
+reflexivity.
+Qed.
+Hint Rewrite liftState_return : liftState.
+
+(*
+Lemma Value_liftState_Run:
+ List.In (Value a, s') (liftState r m s)
+ exists t, Run m t a.
+ by (use assms in \<open>induction r m arbitrary: s s' rule: liftState.induct\<close>;
+ simp add: failS_def throwS_def returnS_def del: read_regvalS.simps;
+ blast elim: Value_bindS_elim)
+
+lemmas liftState_if_distrib[liftState_simp] = if_distrib[where f = "liftState ra" for ra]
+*)
+Lemma liftState_if_distrib Regs Regval A E {r x y} {c : bool} :
+ @liftState Regs Regval A E r (if c then x else y) = if c then liftState r x else liftState r y.
+destruct c; reflexivity.
+Qed.
+Lemma liftState_if_distrib_sumbool {Regs Regval A E P Q r x y} {c : sumbool P Q} :
+ @liftState Regs Regval A E r (if c then x else y) = if c then liftState r x else liftState r y.
+destruct c; reflexivity.
+Qed.
+
+Lemma Value_bindS_iff {Regs A B E} {f : A -> monadS Regs B E} {b m s s''} :
+ List.In (Value b, s'') (bindS m f s) <-> (exists a s', List.In (Value a, s') (m s) /\ List.In (Value b, s'') (f a s')).
+split.
+* intro H.
+ apply bindS_cases in H.
+ destruct H as [(? & ? & ? & [= <-] & ? & ?) | [(? & [= <-] & ?) | (? & ? & ? & [= <-] & ? & ?)]];
+ eauto.
+* intros (? & ? & ? & ?).
+ eauto with bindS_intros.
+Qed.
+
+Lemma Ex_bindS_iff {Regs A B E} {f : A -> monadS Regs B E} {m e s s''} :
+ List.In (Ex e, s'') (bindS m f s) <-> List.In (Ex e, s'') (m s) \/ (exists a s', List.In (Value a, s') (m s) /\ List.In (Ex e, s'') (f a s')).
+split.
+* intro H.
+ apply bindS_cases in H.
+ destruct H as [(? & ? & ? & [= <-] & ? & ?) | [(? & [= <-] & ?) | (? & ? & ? & [= <-] & ? & ?)]];
+ eauto.
+* intros [H | (? & ? & H1 & H2)];
+ eauto with bindS_intros.
+Qed.
+
+Lemma liftState_throw Regs Regval A E {r} {e : E} :
+ @liftState Regval Regs A E r (throw e) = throwS e.
+reflexivity.
+Qed.
+Lemma liftState_assert Regs Regval E {r c msg} :
+ @liftState Regval Regs _ E r (assert_exp c msg) = assert_expS c msg.
+destruct c; reflexivity.
+Qed.
+Lemma liftState_exit Regs Regval A E r :
+ @liftState Regval Regs A E r (exit tt) = exitS tt.
+reflexivity.
+Qed.
+Lemma liftState_exclResult Regs Regval E r :
+ @liftState Regs Regval _ E r (excl_result tt) = excl_resultS tt.
+reflexivity.
+Qed.
+Lemma liftState_barrier Regs Regval E r bk :
+ @liftState Regs Regval _ E r (barrier bk) = returnS tt.
+reflexivity.
+Qed.
+Lemma liftState_footprint Regs Regval E r :
+ @liftState Regs Regval _ E r (footprint tt) = returnS tt.
+reflexivity.
+Qed.
+Lemma liftState_choose_bool Regs Regval E r descr :
+ @liftState Regs Regval _ E r (choose_bool descr) = choose_boolS tt.
+reflexivity.
+Qed.
+(*declare undefined_boolS_def[simp]*)
+Lemma liftState_undefined Regs Regval E r :
+ @liftState Regs Regval _ E r (undefined_bool tt) = undefined_boolS tt.
+reflexivity.
+Qed.
+Lemma liftState_maybe_fail Regs Regval A E r msg x :
+ @liftState Regs Regval A E r (maybe_fail msg x) = maybe_failS msg x.
+destruct x; reflexivity.
+Qed.
+Lemma liftState_and_boolM Regs Regval E r x y :
+ @liftState Regs Regval _ E r (and_boolM x y) === and_boolS (liftState r x) (liftState r y).
+unfold and_boolM, and_boolS.
+rewrite liftState_bind.
+apply bindS_cong; auto.
+intros. rewrite liftState_if_distrib.
+reflexivity.
+Qed.
+Lemma liftState_and_boolMP Regs Regval E P Q R r x y H :
+ @liftState Regs Regval _ E r (@and_boolMP _ _ P Q R x y H) === and_boolSP (liftState r x) (liftState r y).
+unfold and_boolMP, and_boolSP.
+rewrite liftState_bind.
+apply bindS_cong; auto.
+intros [[|] [A]].
+* rewrite liftState_bind;
+ simpl;
+ apply bindS_cong; auto;
+ intros [a' A'];
+ rewrite liftState_return;
+ reflexivity.
+* rewrite liftState_return.
+ reflexivity.
+Qed.
+
+Lemma liftState_or_boolM Regs Regval E r x y :
+ @liftState Regs Regval _ E r (or_boolM x y) === or_boolS (liftState r x) (liftState r y).
+unfold or_boolM, or_boolS.
+rewrite liftState_bind.
+apply bindS_cong; auto.
+intros. rewrite liftState_if_distrib.
+reflexivity.
+Qed.
+Lemma liftState_or_boolMP Regs Regval E P Q R r x y H :
+ @liftState Regs Regval _ E r (@or_boolMP _ _ P Q R x y H) === or_boolSP (liftState r x) (liftState r y).
+unfold or_boolMP, or_boolSP.
+rewrite liftState_bind.
+simpl.
+apply bindS_cong; auto.
+intros [[|] [A]].
+* rewrite liftState_return.
+ reflexivity.
+* rewrite liftState_bind;
+ simpl;
+ apply bindS_cong; auto;
+ intros [a' A'];
+ rewrite liftState_return;
+ reflexivity.
+Qed.
+Hint Rewrite liftState_throw liftState_assert liftState_exit liftState_exclResult
+ liftState_barrier liftState_footprint liftState_choose_bool
+ liftState_undefined liftState_maybe_fail
+ liftState_and_boolM liftState_and_boolMP
+ liftState_or_boolM liftState_or_boolMP
+ : liftState.
+
+Lemma liftState_try_catch Regs Regval A E1 E2 r m h :
+ @liftState Regs Regval A E2 r (try_catch (E1 := E1) m h) === try_catchS (liftState r m) (fun e => liftState r (h e)).
+induction m; intros; simpl; autorewrite with state;
+solve
+[ auto
+| erewrite try_catchS_bindS_no_throw; intros;
+ only 2,3: (autorewrite with ignore_throw; reflexivity);
+ apply bindS_cong; auto
+].
+Qed.
+Hint Rewrite liftState_try_catch : liftState.
+
+Lemma liftState_early_return Regs Regval A R E r x :
+ liftState (Regs := Regs) r (@early_return Regval A R E x) = early_returnS x.
+reflexivity.
+Qed.
+Hint Rewrite liftState_early_return : liftState.
+
+Lemma liftState_catch_early_return (*[liftState_simp]:*) Regs Regval A E r m :
+ liftState (Regs := Regs) r (@catch_early_return Regval A E m) === catch_early_returnS (liftState r m).
+unfold catch_early_return, catch_early_returnS.
+rewrite_liftState.
+apply try_catchS_cong; auto.
+intros [a | e] s'; auto.
+Qed.
+Hint Rewrite liftState_catch_early_return : liftState.
+
+Lemma liftState_liftR Regs Regval A R E r m :
+ liftState (Regs := Regs) r (@liftR Regval A R E m) === liftRS (liftState r m).
+unfold liftR, liftRS.
+rewrite_liftState.
+reflexivity.
+Qed.
+Hint Rewrite liftState_liftR : liftState.
+
+Lemma liftState_try_catchR Regs Regval A R E1 E2 r m h :
+ liftState (Regs := Regs) r (@try_catchR Regval A R E1 E2 m h) === try_catchRS (liftState r m) (fun x => liftState r (h x)).
+unfold try_catchR, try_catchRS. rewrite_liftState.
+apply try_catchS_cong; auto.
+intros [r' | e] s'; auto.
+Qed.
+Hint Rewrite liftState_try_catchR : liftState.
+(*
+Lemma liftState_bool_of_bitU_nondet Regs Regval :
+ "liftState r (bool_of_bitU_nondet b) = bool_of_bitU_nondetS b"
+ by (cases b; auto simp: bool_of_bitU_nondet_def bool_of_bitU_nondetS_def liftState_simp)
+Hint Rewrite liftState_bool_of_bitU_nondet : liftState.
+*)
+Lemma liftState_read_memt Regs Regval A B E H rk a sz r :
+ liftState (Regs := Regs) r (@read_memt Regval A B E H rk a sz) === read_memtS rk a sz.
+unfold read_memt, read_memt_bytes, read_memtS, maybe_failS. simpl.
+apply bindS_cong; auto.
+intros [byte bit].
+destruct (option_map _); auto.
+Qed.
+Hint Rewrite liftState_read_memt : liftState.
+
+Lemma liftState_read_mem Regs Regval A B E H rk asz a sz r :
+ liftState (Regs := Regs) r (@read_mem Regval A B E H rk asz a sz) === read_memS rk a sz.
+unfold read_mem, read_memS, read_memtS. simpl.
+unfold read_mem_bytesS, read_memt_bytesS.
+repeat rewrite bindS_assoc.
+apply bindS_cong; auto.
+intros [ bytes | ]; auto. simpl.
+apply bindS_cong; auto.
+intros [byte bit].
+rewrite bindS_returnS_left. rewrite_liftState.
+destruct (option_map _); auto.
+Qed.
+Hint Rewrite liftState_read_mem : liftState.
+
+Lemma liftState_write_mem_ea Regs Regval A E rk asz a sz r :
+ liftState (Regs := Regs) r (@write_mem_ea Regval A E rk asz a sz) = returnS tt.
+reflexivity.
+Qed.
+Hint Rewrite liftState_write_mem_ea : liftState.
+
+Lemma liftState_write_memt Regs Regval A B E wk addr sz v t r :
+ liftState (Regs := Regs) r (@write_memt Regval A B E wk addr sz v t) = write_memtS wk addr sz v t.
+unfold write_memt, write_memtS.
+destruct (Sail2_values.mem_bytes_of_bits v); auto.
+Qed.
+Hint Rewrite liftState_write_memt : liftState.
+
+Lemma liftState_write_mem Regs Regval A B E wk addrsize addr sz v r :
+ liftState (Regs := Regs) r (@write_mem Regval A B E wk addrsize addr sz v) = write_memS wk addr sz v.
+unfold write_mem, write_memS, write_memtS.
+destruct (Sail2_values.mem_bytes_of_bits v); simpl; auto.
+Qed.
+Hint Rewrite liftState_write_mem : liftState.
+
+Lemma bindS_rw_left Regs A B E m1 m2 (f : A -> monadS Regs B E) s :
+ m1 s = m2 s ->
+ bindS m1 f s = bindS m2 f s.
+intro H. unfold bindS. rewrite H. reflexivity.
+Qed.
+
+Lemma liftState_read_reg_readS Regs Regval A E reg get_regval' set_regval' :
+ (forall s, map_bind reg.(of_regval) (get_regval' reg.(name) s) = Some (reg.(read_from) s)) ->
+ liftState (Regs := Regs) (get_regval', set_regval') (@read_reg _ Regval A E reg) === readS (fun x => reg.(read_from) (ss_regstate x)).
+intros.
+unfold read_reg. simpl. unfold readS. intro s.
+erewrite bindS_rw_left. 2: {
+ apply bindS_returnS_left.
+}
+specialize (H (ss_regstate s)).
+destruct (get_regval' _ _) as [v | ]; only 2: discriminate H.
+rewrite bindS_returnS_left.
+simpl in *.
+rewrite H.
+reflexivity.
+Qed.
+
+Lemma liftState_write_reg_updateS Regs Regval A E get_regval' set_regval' reg (v : A) :
+ (forall s, set_regval' (name reg) (regval_of reg v) s = Some (write_to reg v s)) ->
+ liftState (Regs := Regs) (Regval := Regval) (E := E) (get_regval', set_regval') (write_reg reg v) === updateS (fun s => {| ss_regstate := (write_to reg v s.(ss_regstate)); ss_memstate := s.(ss_memstate); ss_tagstate := s.(ss_tagstate) |}).
+intros. intro s.
+unfold write_reg. simpl. unfold readS, seqS.
+erewrite bindS_rw_left. 2: {
+ apply bindS_returnS_left.
+}
+specialize (H (ss_regstate s)).
+destruct (set_regval' _ _) as [v' | ]; only 2: discriminate H.
+injection H as H1.
+unfold updateS.
+rewrite <- H1.
+reflexivity.
+Qed.
+(*
+Lemma liftState_iter_aux Regs Regval A E :
+ liftState r (iter_aux i f xs) = iterS_aux i (fun i x => liftState r (f i x)) xs.
+ by (induction i "\<lambda>i x. liftState r (f i x)" xs rule: iterS_aux.induct)
+ (auto simp: liftState_simp cong: bindS_cong)
+Hint Rewrite liftState_iter_aux : liftState.
+
+lemma liftState_iteri[liftState_simp]:
+ "liftState r (iteri f xs) = iteriS (\<lambda>i x. liftState r (f i x)) xs"
+ by (auto simp: iteri_def iteriS_def liftState_simp)
+
+lemma liftState_iter[liftState_simp]:
+ "liftState r (iter f xs) = iterS (liftState r \<circ> f) xs"
+ by (auto simp: iter_def iterS_def liftState_simp)
+*)
+Lemma liftState_foreachM Regs Regval A Vars E (xs : list A) (vars : Vars) (body : A -> Vars -> monad Regval Vars E) r :
+ liftState (Regs := Regs) r (foreachM xs vars body) === foreachS xs vars (fun x vars => liftState r (body x vars)).
+revert vars.
+induction xs as [ | h t].
+* reflexivity.
+* intros vars. simpl.
+ rewrite_liftState.
+ apply bindS_cong; auto.
+Qed.
+Hint Rewrite liftState_foreachM : liftState.
+
+Lemma foreachS_cong {A RV Vars E} xs vars f f' :
+ (forall a vars, f a vars === f' a vars) ->
+ @foreachS A RV Vars E xs vars f === foreachS xs vars f'.
+intro H.
+revert vars.
+induction xs.
+* reflexivity.
+* intros. simpl.
+ rewrite H.
+ apply bindS_cong; auto.
+Qed.
+
+Add Parametric Morphism {Regs A Vars E : Type} : (@foreachS A Regs Vars E)
+ with signature eq ==> eq ==> equiv ==> equiv as foreachS_morphism.
+apply foreachS_cong.
+Qed.
+
+(*Tactic Notation "sail_rewrite" ident(hintdb) := rewrite_strat topdown (choice (hints hintdb) progress eval cbn beta).
+Ltac sail_rewrite hintdb := rewrite_strat topdown (choice (hints hintdb) progress eval cbn beta).*)
+
+Lemma liftState_genlistM Regs Regval A E r f n :
+ liftState (Regs := Regs) r (@genlistM A Regval E f n) === genlistS (fun x => liftState r (f x)) n.
+unfold genlistM, genlistS.
+rewrite_liftState.
+reflexivity.
+Qed.
+Hint Rewrite liftState_genlistM : liftState.
+
+Add Parametric Morphism {A RV E : Type} : (@genlistS A RV E)
+ with signature equiv ==> eq ==> equiv as genlistS_morphism.
+intros f g EQ n.
+unfold genlistS.
+apply foreachS_cong.
+intros m vars.
+rewrite EQ.
+reflexivity.
+Qed.
+
+Lemma liftState_choose_bools Regs Regval E descr n r :
+ liftState (Regs := Regs) r (@choose_bools Regval E descr n) === choose_boolsS n.
+unfold choose_bools, choose_boolsS.
+rewrite_liftState.
+reflexivity.
+Qed.
+Hint Rewrite liftState_choose_bools : liftState.
+
+(*
+Lemma liftState_bools_of_bits_nondet[liftState_simp]:
+ "liftState r (bools_of_bits_nondet bs) = bools_of_bits_nondetS bs"
+ unfolding bools_of_bits_nondet_def bools_of_bits_nondetS_def
+ by (auto simp: liftState_simp comp_def)
+Hint Rewrite liftState_choose_bools : liftState.
+*)
+
+Lemma liftState_internal_pick Regs Regval A E r (xs : list A) :
+ liftState (Regs := Regs) (Regval := Regval) (E := E) r (internal_pick xs) === internal_pickS xs.
+unfold internal_pick, internal_pickS.
+unfold choose.
+rewrite_liftState.
+apply bindS_cong; auto.
+intros.
+destruct (nth_error _ _); auto.
+Qed.
+Hint Rewrite liftState_internal_pick : liftState.
+
+Lemma liftRS_returnS (*[simp]:*) A R Regs E x :
+ @liftRS A R Regs E (returnS x) = returnS x.
+reflexivity.
+Qed.
+
+Lemma concat_singleton A (xs : list A) :
+ concat (xs::nil) = xs.
+simpl.
+rewrite app_nil_r.
+reflexivity.
+Qed.
+
+Lemma liftRS_bindS Regs A B R E (m : monadS Regs A E) (f : A -> monadS Regs B E) :
+ @liftRS B R Regs E (bindS m f) === bindS (liftRS m) (fun x => liftRS (f x)).
+intro s.
+unfold liftRS, try_catchS, bindS, throwS, returnS.
+induction (m s) as [ | [[a | [msg | e]] t]].
+* reflexivity.
+* simpl. rewrite flat_map_app. rewrite IHl. reflexivity.
+* simpl. rewrite IHl. reflexivity.
+* simpl. rewrite IHl. reflexivity.
+Qed.
+
+Lemma liftRS_assert_expS_True (*[simp]:*) Regs R E msg :
+ @liftRS _ R Regs E (assert_expS true msg) = returnS tt.
+reflexivity.
+Qed.
+
+(*
+lemma untilM_domI:
+ fixes V :: "'vars \<Rightarrow> nat"
+ assumes "Inv vars"
+ and "\<And>vars t vars' t'. \<lbrakk>Inv vars; Run (body vars) t vars'; Run (cond vars') t' False\<rbrakk> \<Longrightarrow> V vars' < V vars \<and> Inv vars'"
+ shows "untilM_dom (vars, cond, body)"
+ using assms
+ by (induction vars rule: measure_induct_rule[where f = V])
+ (auto intro: untilM.domintros)
+
+lemma untilM_dom_untilS_dom:
+ assumes "untilM_dom (vars, cond, body)"
+ shows "untilS_dom (vars, liftState r \<circ> cond, liftState r \<circ> body, s)"
+ using assms
+ by (induction vars cond body arbitrary: s rule: untilM.pinduct)
+ (rule untilS.domintros, auto elim!: Value_liftState_Run)
+
+lemma measure2_induct:
+ fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> nat"
+ assumes "\<And>x1 y1. (\<And>x2 y2. f x2 y2 < f x1 y1 \<Longrightarrow> P x2 y2) \<Longrightarrow> P x1 y1"
+ shows "P x y"
+proof -
+ have "P (fst x) (snd x)" for x
+ by (induction x rule: measure_induct_rule[where f = "\<lambda>x. f (fst x) (snd x)"]) (auto intro: assms)
+ then show ?thesis by auto
+qed
+
+lemma untilS_domI:
+ fixes V :: "'vars \<Rightarrow> 'regs sequential_state \<Rightarrow> nat"
+ assumes "Inv vars s"
+ and "\<And>vars s vars' s' s''.
+ \<lbrakk>Inv vars s; (Value vars', s') \<in> body vars s; (Value False, s'') \<in> cond vars' s'\<rbrakk>
+ \<Longrightarrow> V vars' s'' < V vars s \<and> Inv vars' s''"
+ shows "untilS_dom (vars, cond, body, s)"
+ using assms
+ by (induction vars s rule: measure2_induct[where f = V])
+ (auto intro: untilS.domintros)
+
+lemma whileS_dom_step:
+ assumes "whileS_dom (vars, cond, body, s)"
+ and "(Value True, s') \<in> cond vars s"
+ and "(Value vars', s'') \<in> body vars s'"
+ shows "whileS_dom (vars', cond, body, s'')"
+ by (use assms in \<open>induction vars cond body s arbitrary: vars' s' s'' rule: whileS.pinduct\<close>)
+ (auto intro: whileS.domintros)
+
+lemma whileM_dom_step:
+ assumes "whileM_dom (vars, cond, body)"
+ and "Run (cond vars) t True"
+ and "Run (body vars) t' vars'"
+ shows "whileM_dom (vars', cond, body)"
+ by (use assms in \<open>induction vars cond body arbitrary: vars' t t' rule: whileM.pinduct\<close>)
+ (auto intro: whileM.domintros)
+
+lemma whileM_dom_ex_step:
+ assumes "whileM_dom (vars, cond, body)"
+ and "\<exists>t. Run (cond vars) t True"
+ and "\<exists>t'. Run (body vars) t' vars'"
+ shows "whileM_dom (vars', cond, body)"
+ using assms by (blast intro: whileM_dom_step)
+
+lemmas whileS_pinduct = whileS.pinduct[case_names Step]
+
+lemma liftState_whileM:
+ assumes "whileS_dom (vars, liftState r \<circ> cond, liftState r \<circ> body, s)"
+ and "whileM_dom (vars, cond, body)"
+ shows "liftState r (whileM vars cond body) s = whileS vars (liftState r \<circ> cond) (liftState r \<circ> body) s"
+proof (use assms in \<open>induction vars "liftState r \<circ> cond" "liftState r \<circ> body" s rule: whileS.pinduct\<close>)
+ case Step: (1 vars s)
+ note domS = Step(1) and IH = Step(2) and domM = Step(3)
+ show ?case unfolding whileS.psimps[OF domS] whileM.psimps[OF domM] liftState_bind
+ proof (intro bindS_ext_cong, goal_cases cond while)
+ case (while a s')
+ have "bindS (liftState r (body vars)) (liftState r \<circ> (\<lambda>vars. whileM vars cond body)) s' =
+ bindS (liftState r (body vars)) (\<lambda>vars. whileS vars (liftState r \<circ> cond) (liftState r \<circ> body)) s'"
+ if "a"
+ proof (intro bindS_ext_cong, goal_cases body while')
+ case (while' vars' s'')
+ have "whileM_dom (vars', cond, body)" proof (rule whileM_dom_ex_step[OF domM])
+ show "\<exists>t. Run (cond vars) t True" using while that by (auto elim: Value_liftState_Run)
+ show "\<exists>t'. Run (body vars) t' vars'" using while' that by (auto elim: Value_liftState_Run)
+ qed
+ then show ?case using while while' that IH by auto
+ qed auto
+ then show ?case by (auto simp: liftState_simp)
+ qed auto
+qed
+*)
+
+Local Opaque _limit_reduces.
+Ltac gen_reduces :=
+ match goal with |- context[@_limit_reduces ?a ?b ?c] => generalize (@_limit_reduces a b c) end.
+
+(* TODO: rewrite_liftState is performing really badly here. We could add liftState_if_distrib
+ to the hint db, but then it starts failing in a way that causes the whole rewriting to fail. *)
+
+Lemma liftState_whileM RV Vars E r measure vars cond (body : Vars -> monad RV Vars E) :
+ liftState (Regs := RV) r (whileMT vars measure cond body) === whileST vars measure (fun vars => liftState r (cond vars)) (fun vars => liftState r (body vars)).
+unfold whileMT, whileST.
+generalize (measure vars) as limit. intro.
+revert vars.
+destruct (Z.le_decidable 0 limit).
+* generalize (Zwf_guarded limit) as acc.
+ apply Wf_Z.natlike_ind with (x := limit).
+ + intros [acc] *; simpl.
+ match goal with |- context [Build_ArithFact _ ?prf] => generalize prf; intros ?Proof end.
+ rewrite_liftState.
+ setoid_rewrite liftState_if_distrib.
+ apply bindS_cong; auto.
+ destruct a; rewrite_liftState; auto.
+ apply bindS_cong; auto.
+ intros. destruct (_limit_reduces _). simpl.
+ reflexivity.
+ + clear limit H.
+ intros limit H IH [acc] vars s. simpl.
+ destruct (Z_ge_dec _ _); try omega.
+ autorewrite with liftState.
+ apply bindS_ext_cong; auto.
+ intros. rewrite liftState_if_distrib.
+ destruct a; autorewrite with liftState; auto.
+ apply bindS_ext_cong; auto.
+ intros.
+ gen_reduces.
+ replace (Z.succ limit - 1) with limit; try omega. intro acc'.
+ apply IH.
+ + assumption.
+* intros. simpl.
+ destruct (Z_ge_dec _ _); try omega.
+ reflexivity.
+Qed.
+
+(*
+lemma untilM_dom_step:
+ assumes "untilM_dom (vars, cond, body)"
+ and "Run (body vars) t vars'"
+ and "Run (cond vars') t' False"
+ shows "untilM_dom (vars', cond, body)"
+ by (use assms in \<open>induction vars cond body arbitrary: vars' t t' rule: untilM.pinduct\<close>)
+ (auto intro: untilM.domintros)
+
+lemma untilM_dom_ex_step:
+ assumes "untilM_dom (vars, cond, body)"
+ and "\<exists>t. Run (body vars) t vars'"
+ and "\<exists>t'. Run (cond vars') t' False"
+ shows "untilM_dom (vars', cond, body)"
+ using assms by (blast intro: untilM_dom_step)
+
+lemma liftState_untilM:
+ assumes "untilS_dom (vars, liftState r \<circ> cond, liftState r \<circ> body, s)"
+ and "untilM_dom (vars, cond, body)"
+ shows "liftState r (untilM vars cond body) s = untilS vars (liftState r \<circ> cond) (liftState r \<circ> body) s"
+proof (use assms in \<open>induction vars "liftState r \<circ> cond" "liftState r \<circ> body" s rule: untilS.pinduct\<close>)
+ case Step: (1 vars s)
+ note domS = Step(1) and IH = Step(2) and domM = Step(3)
+ show ?case unfolding untilS.psimps[OF domS] untilM.psimps[OF domM] liftState_bind
+ proof (intro bindS_ext_cong, goal_cases body k)
+ case (k vars' s')
+ show ?case unfolding comp_def liftState_bind
+ proof (intro bindS_ext_cong, goal_cases cond until)
+ case (until a s'')
+ have "untilM_dom (vars', cond, body)" if "\<not>a"
+ proof (rule untilM_dom_ex_step[OF domM])
+ show "\<exists>t. Run (body vars) t vars'" using k by (auto elim: Value_liftState_Run)
+ show "\<exists>t'. Run (cond vars') t' False" using until that by (auto elim: Value_liftState_Run)
+ qed
+ then show ?case using k until IH by (auto simp: comp_def liftState_simp)
+ qed auto
+ qed auto
+qed*)
+
+Lemma liftState_untilM RV Vars E r measure vars cond (body : Vars -> monad RV Vars E) :
+ liftState (Regs := RV) r (untilMT vars measure cond body) === untilST vars measure (fun vars => liftState r (cond vars)) (fun vars => liftState r (body vars)).
+unfold untilMT, untilST.
+generalize (measure vars) as limit. intro.
+revert vars.
+destruct (Z.le_decidable 0 limit).
+* generalize (Zwf_guarded limit) as acc.
+ apply Wf_Z.natlike_ind with (x := limit).
+ + intros [acc] * s; simpl.
+(* TODO rewrite_liftState.*)
+autorewrite with liftState.
+ apply bindS_ext_cong; auto.
+ intros. autorewrite with liftState.
+ apply bindS_ext_cong; auto.
+ intros. rewrite liftState_if_distrib.
+ destruct a0; auto.
+ destruct (_limit_reduces _). simpl.
+ reflexivity.
+ + clear limit H.
+ intros limit H IH [acc] vars s. simpl.
+ destruct (Z_ge_dec _ _); try omega.
+ autorewrite with liftState.
+ apply bindS_ext_cong; auto.
+ intros. autorewrite with liftState; auto.
+ apply bindS_ext_cong; auto.
+ intros. rewrite liftState_if_distrib.
+ destruct a0; autorewrite with liftState; auto.
+ gen_reduces.
+ replace (Z.succ limit - 1) with limit; try omega. intro acc'.
+ apply IH.
+ + assumption.
+* intros. simpl.
+ destruct (Z_ge_dec _ _); try omega.
+ reflexivity.
+Qed.
+
+(*
+
+text \<open>Simplification rules for monadic Boolean connectives\<close>
+
+lemma if_return_return[simp]: "(if a then return True else return False) = return a" by auto
+
+lemma and_boolM_simps[simp]:
+ "and_boolM (return b) (return c) = return (b \<and> c)"
+ "and_boolM x (return True) = x"
+ "and_boolM x (return False) = x \<bind> (\<lambda>_. return False)"
+ "\<And>x y z. and_boolM (x \<bind> y) z = (x \<bind> (\<lambda>r. and_boolM (y r) z))"
+ by (auto simp: and_boolM_def)
+
+lemma and_boolM_return_if:
+ "and_boolM (return b) y = (if b then y else return False)"
+ by (auto simp: and_boolM_def)
+
+lemma and_boolM_return_return_and[simp]: "and_boolM (return l) (return r) = return (l \<and> r)"
+ by (auto simp: and_boolM_def)
+
+lemmas and_boolM_if_distrib[simp] = if_distrib[where f = "\<lambda>x. and_boolM x y" for y]
+
+lemma or_boolM_simps[simp]:
+ "or_boolM (return b) (return c) = return (b \<or> c)"
+ "or_boolM x (return True) = x \<bind> (\<lambda>_. return True)"
+ "or_boolM x (return False) = x"
+ "\<And>x y z. or_boolM (x \<bind> y) z = (x \<bind> (\<lambda>r. or_boolM (y r) z))"
+ by (auto simp: or_boolM_def)
+
+lemma or_boolM_return_if:
+ "or_boolM (return b) y = (if b then return True else y)"
+ by (auto simp: or_boolM_def)
+
+lemma or_boolM_return_return_or[simp]: "or_boolM (return l) (return r) = return (l \<or> r)"
+ by (auto simp: or_boolM_def)
+
+lemmas or_boolM_if_distrib[simp] = if_distrib[where f = "\<lambda>x. or_boolM x y" for y]
+
+lemma if_returnS_returnS[simp]: "(if a then returnS True else returnS False) = returnS a" by auto
+
+lemma and_boolS_simps[simp]:
+ "and_boolS (returnS b) (returnS c) = returnS (b \<and> c)"
+ "and_boolS x (returnS True) = x"
+ "and_boolS x (returnS False) = bindS x (\<lambda>_. returnS False)"
+ "\<And>x y z. and_boolS (bindS x y) z = (bindS x (\<lambda>r. and_boolS (y r) z))"
+ by (auto simp: and_boolS_def)
+
+lemma and_boolS_returnS_if:
+ "and_boolS (returnS b) y = (if b then y else returnS False)"
+ by (auto simp: and_boolS_def)
+
+lemmas and_boolS_if_distrib[simp] = if_distrib[where f = "\<lambda>x. and_boolS x y" for y]
+
+lemma and_boolS_returnS_True[simp]: "and_boolS (returnS True) c = c"
+ by (auto simp: and_boolS_def)
+
+lemma or_boolS_simps[simp]:
+ "or_boolS (returnS b) (returnS c) = returnS (b \<or> c)"
+ "or_boolS (returnS False) m = m"
+ "or_boolS x (returnS True) = bindS x (\<lambda>_. returnS True)"
+ "or_boolS x (returnS False) = x"
+ "\<And>x y z. or_boolS (bindS x y) z = (bindS x (\<lambda>r. or_boolS (y r) z))"
+ by (auto simp: or_boolS_def)
+
+lemma or_boolS_returnS_if:
+ "or_boolS (returnS b) y = (if b then returnS True else y)"
+ by (auto simp: or_boolS_def)
+
+lemmas or_boolS_if_distrib[simp] = if_distrib[where f = "\<lambda>x. or_boolS x y" for y]
+
+lemma Run_or_boolM_E:
+ assumes "Run (or_boolM l r) t a"
+ obtains "Run l t True" and "a"
+ | tl tr where "Run l tl False" and "Run r tr a" and "t = tl @ tr"
+ using assms by (auto simp: or_boolM_def elim!: Run_bindE Run_ifE Run_returnE)
+
+lemma Run_and_boolM_E:
+ assumes "Run (and_boolM l r) t a"
+ obtains "Run l t False" and "\<not>a"
+ | tl tr where "Run l tl True" and "Run r tr a" and "t = tl @ tr"
+ using assms by (auto simp: and_boolM_def elim!: Run_bindE Run_ifE Run_returnE)
+
+lemma maybe_failS_Some[simp]: "maybe_failS msg (Some v) = returnS v"
+ by (auto simp: maybe_failS_def)
+
+text \<open>Event traces\<close>
+
+lemma Some_eq_bind_conv: "Some x = Option.bind f g \<longleftrightarrow> (\<exists>y. f = Some y \<and> g y = Some x)"
+ unfolding bind_eq_Some_conv[symmetric] by auto
+
+lemma if_then_Some_eq_Some_iff: "((if b then Some x else None) = Some y) \<longleftrightarrow> (b \<and> y = x)"
+ by auto
+
+lemma Some_eq_if_then_Some_iff: "(Some y = (if b then Some x else None)) \<longleftrightarrow> (b \<and> y = x)"
+ by auto
+
+lemma emitEventS_update_cases:
+ assumes "emitEventS ra e s = Some s'"
+ obtains
+ (Write_mem) wk addr sz v tag r
+ where "e = E_write_memt wk addr sz v tag r \<or> (e = E_write_mem wk addr sz v r \<and> tag = B0)"
+ and "s' = put_mem_bytes addr sz v tag s"
+ | (Write_reg) r v rs'
+ where "e = E_write_reg r v" and "(snd ra) r v (regstate s) = Some rs'"
+ and "s' = s\<lparr>regstate := rs'\<rparr>"
+ | (Read) "s' = s"
+ using assms
+ by (elim emitEventS.elims)
+ (auto simp: Some_eq_bind_conv bind_eq_Some_conv if_then_Some_eq_Some_iff Some_eq_if_then_Some_iff)
+
+lemma runTraceS_singleton[simp]: "runTraceS ra [e] s = emitEventS ra e s"
+ by (cases "emitEventS ra e s"; auto)
+
+lemma runTraceS_ConsE:
+ assumes "runTraceS ra (e # t) s = Some s'"
+ obtains s'' where "emitEventS ra e s = Some s''" and "runTraceS ra t s'' = Some s'"
+ using assms by (auto simp: bind_eq_Some_conv)
+
+lemma runTraceS_ConsI:
+ assumes "emitEventS ra e s = Some s'" and "runTraceS ra t s' = Some s''"
+ shows "runTraceS ra (e # t) s = Some s''"
+ using assms by auto
+
+lemma runTraceS_Cons_tl:
+ assumes "emitEventS ra e s = Some s'"
+ shows "runTraceS ra (e # t) s = runTraceS ra t s'"
+ using assms by (elim emitEventS.elims) (auto simp: Some_eq_bind_conv bind_eq_Some_conv)
+
+lemma runTraceS_appendE:
+ assumes "runTraceS ra (t @ t') s = Some s'"
+ obtains s'' where "runTraceS ra t s = Some s''" and "runTraceS ra t' s'' = Some s'"
+proof -
+ have "\<exists>s''. runTraceS ra t s = Some s'' \<and> runTraceS ra t' s'' = Some s'"
+ proof (use assms in \<open>induction t arbitrary: s\<close>)
+ case (Cons e t)
+ from Cons.prems
+ obtain s_e where "emitEventS ra e s = Some s_e" and "runTraceS ra (t @ t') s_e = Some s'"
+ by (auto elim: runTraceS_ConsE simp: bind_eq_Some_conv)
+ with Cons.IH[of s_e] show ?case by (auto intro: runTraceS_ConsI)
+ qed auto
+ then show ?thesis using that by blast
+qed
+
+lemma runTraceS_nth_split:
+ assumes "runTraceS ra t s = Some s'" and n: "n < length t"
+ obtains s1 s2 where "runTraceS ra (take n t) s = Some s1"
+ and "emitEventS ra (t ! n) s1 = Some s2"
+ and "runTraceS ra (drop (Suc n) t) s2 = Some s'"
+proof -
+ have "runTraceS ra (take n t @ t ! n # drop (Suc n) t) s = Some s'"
+ using assms
+ by (auto simp: id_take_nth_drop[OF n, symmetric])
+ then show thesis by (blast elim: runTraceS_appendE runTraceS_ConsE intro: that)
+qed
+
+text \<open>Memory accesses\<close>
+
+lemma get_mem_bytes_put_mem_bytes_same_addr:
+ assumes "length v = sz"
+ shows "get_mem_bytes addr sz (put_mem_bytes addr sz v tag s) = Some (v, if sz > 0 then tag else B1)"
+proof (unfold assms[symmetric], induction v rule: rev_induct)
+ case Nil
+ then show ?case by (auto simp: get_mem_bytes_def)
+next
+ case (snoc x xs)
+ then show ?case
+ by (cases tag)
+ (auto simp: get_mem_bytes_def put_mem_bytes_def Let_def and_bit_eq_iff foldl_and_bit_eq_iff
+ cong: option.case_cong split: if_splits option.splits)
+qed
+
+lemma memstate_put_mem_bytes:
+ assumes "length v = sz"
+ shows "memstate (put_mem_bytes addr sz v tag s) addr' =
+ (if addr' \<in> {addr..<addr+sz} then Some (v ! (addr' - addr)) else memstate s addr')"
+ unfolding assms[symmetric]
+ by (induction v rule: rev_induct) (auto simp: put_mem_bytes_def nth_Cons nth_append Let_def)
+
+lemma tagstate_put_mem_bytes:
+ assumes "length v = sz"
+ shows "tagstate (put_mem_bytes addr sz v tag s) addr' =
+ (if addr' \<in> {addr..<addr+sz} then Some tag else tagstate s addr')"
+ unfolding assms[symmetric]
+ by (induction v rule: rev_induct) (auto simp: put_mem_bytes_def nth_Cons nth_append Let_def)
+
+lemma get_mem_bytes_cong:
+ assumes "\<forall>addr'. addr \<le> addr' \<and> addr' < addr + sz \<longrightarrow>
+ (memstate s' addr' = memstate s addr' \<and> tagstate s' addr' = tagstate s addr')"
+ shows "get_mem_bytes addr sz s' = get_mem_bytes addr sz s"
+proof (use assms in \<open>induction sz\<close>)
+ case 0
+ then show ?case by (auto simp: get_mem_bytes_def)
+next
+ case (Suc sz)
+ then show ?case
+ by (auto simp: get_mem_bytes_def Let_def
+ intro!: map_option_cong map_cong foldl_cong
+ arg_cong[where f = just_list] arg_cong2[where f = and_bit])
+qed
+
+lemma get_mem_bytes_tagged_tagstate:
+ assumes "get_mem_bytes addr sz s = Some (v, B1)"
+ shows "\<forall>addr' \<in> {addr..<addr + sz}. tagstate s addr' = Some B1"
+ using assms
+ by (auto simp: get_mem_bytes_def foldl_and_bit_eq_iff Let_def split: option.splits)
+
+end
+*) \ No newline at end of file
diff --git a/prover_snapshots/coq/lib/sail/Sail2_state_lifting.v b/prover_snapshots/coq/lib/sail/Sail2_state_lifting.v
new file mode 100644
index 0000000..1544c3c
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_state_lifting.v
@@ -0,0 +1,61 @@
+Require Import Sail2_values.
+Require Import Sail2_prompt_monad.
+Require Import Sail2_prompt.
+Require Import Sail2_state_monad.
+Import ListNotations.
+
+(* Lifting from prompt monad to state monad *)
+(*val liftState : forall 'regval 'regs 'a 'e. register_accessors 'regs 'regval -> monad 'regval 'a 'e -> monadS 'regs 'a 'e*)
+Fixpoint liftState {Regval Regs A E} (ra : register_accessors Regs Regval) (m : monad Regval A E) : monadS Regs A E :=
+ match m with
+ | (Done a) => returnS a
+ | (Read_mem rk a sz k) => bindS (read_mem_bytesS rk a sz) (fun v => liftState ra (k v))
+ | (Read_memt rk a sz k) => bindS (read_memt_bytesS rk a sz) (fun v => liftState ra (k v))
+ | (Write_mem wk a sz v k) => bindS (write_mem_bytesS wk a sz v) (fun v => liftState ra (k v))
+ | (Write_memt wk a sz v t k) => bindS (write_memt_bytesS wk a sz v t) (fun v => liftState ra (k v))
+ | (Read_reg r k) => bindS (read_regvalS ra r) (fun v => liftState ra (k v))
+ | (Excl_res k) => bindS (excl_resultS tt) (fun v => liftState ra (k v))
+ | (Choose _ k) => bindS (choose_boolS tt) (fun v => liftState ra (k v))
+ | (Write_reg r v k) => seqS (write_regvalS ra r v) (liftState ra k)
+ | (Write_ea _ _ _ k) => liftState ra k
+ | (Footprint k) => liftState ra k
+ | (Barrier _ k) => liftState ra k
+ | (Print _ k) => liftState ra k (* TODO *)
+ | (Fail descr) => failS descr
+ | (Exception e) => throwS e
+end.
+
+Local Open Scope bool_scope.
+
+(*val emitEventS : forall 'regval 'regs 'a 'e. Eq 'regval => register_accessors 'regs 'regval -> event 'regval -> sequential_state 'regs -> maybe (sequential_state 'regs)*)
+Definition emitEventS {Regval Regs} `{forall (x y : Regval), Decidable (x = y)} (ra : register_accessors Regs Regval) (e : event Regval) (s : sequential_state Regs) : option (sequential_state Regs) :=
+match e with
+ | E_read_mem _ addr sz v =>
+ option_bind (get_mem_bytes addr sz s) (fun '(v', _) =>
+ if generic_eq v' v then Some s else None)
+ | E_read_memt _ addr sz (v, tag) =>
+ option_bind (get_mem_bytes addr sz s) (fun '(v', tag') =>
+ if generic_eq v' v && generic_eq tag' tag then Some s else None)
+ | E_write_mem _ addr sz v success =>
+ if success then Some (put_mem_bytes addr sz v B0 s) else None
+ | E_write_memt _ addr sz v tag success =>
+ if success then Some (put_mem_bytes addr sz v tag s) else None
+ | E_read_reg r v =>
+ let (read_reg, _) := ra in
+ option_bind (read_reg r s.(ss_regstate)) (fun v' =>
+ if generic_eq v' v then Some s else None)
+ | E_write_reg r v =>
+ let (_, write_reg) := ra in
+ option_bind (write_reg r v s.(ss_regstate)) (fun rs' =>
+ Some {| ss_regstate := rs'; ss_memstate := s.(ss_memstate); ss_tagstate := s.(ss_tagstate) |})
+ | _ => Some s
+end.
+
+Local Close Scope bool_scope.
+
+(*val runTraceS : forall 'regval 'regs 'a 'e. Eq 'regval => register_accessors 'regs 'regval -> trace 'regval -> sequential_state 'regs -> maybe (sequential_state 'regs)*)
+Fixpoint runTraceS {Regval Regs} `{forall (x y : Regval), Decidable (x = y)} (ra : register_accessors Regs Regval) (t : trace Regval) (s : sequential_state Regs) : option (sequential_state Regs) :=
+match t with
+ | [] => Some s
+ | e :: t' => option_bind (emitEventS ra e s) (runTraceS ra t')
+end.
diff --git a/prover_snapshots/coq/lib/sail/Sail2_state_monad.v b/prover_snapshots/coq/lib/sail/Sail2_state_monad.v
new file mode 100644
index 0000000..552fa68
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_state_monad.v
@@ -0,0 +1,323 @@
+Require Import Sail2_instr_kinds.
+Require Import Sail2_values.
+Require FMapList.
+Require Import OrderedType.
+Require OrderedTypeEx.
+Require Import List.
+Require bbv.Word.
+Import ListNotations.
+
+(* TODO: revisit choice of FMapList *)
+Module NatMap := FMapList.Make(OrderedTypeEx.Nat_as_OT).
+
+Definition Memstate : Type := NatMap.t memory_byte.
+Definition Tagstate : Type := NatMap.t bitU.
+(* type regstate = map string (vector bitU) *)
+
+(* We deviate from the Lem library and prefix the fields with ss_ to avoid
+ name clashes. *)
+Record sequential_state {Regs} :=
+ { ss_regstate : Regs;
+ ss_memstate : Memstate;
+ ss_tagstate : Tagstate }.
+Arguments sequential_state : clear implicits.
+
+(*val init_state : forall 'regs. 'regs -> sequential_state 'regs*)
+Definition init_state {Regs} regs : sequential_state Regs :=
+ {| ss_regstate := regs;
+ ss_memstate := NatMap.empty _;
+ ss_tagstate := NatMap.empty _ |}.
+
+Inductive ex E :=
+ | Failure : string -> ex E
+ | Throw : E -> ex E.
+Arguments Failure {E} _.
+Arguments Throw {E} _.
+
+Inductive result A E :=
+ | Value : A -> result A E
+ | Ex : ex E -> result A E.
+Arguments Value {A} {E} _.
+Arguments Ex {A} {E} _.
+
+(* State, nondeterminism and exception monad with result value type 'a
+ and exception type 'e. *)
+(* TODO: the list was originally a set, can we reasonably go back to a set? *)
+Definition monadS Regs a e : Type :=
+ sequential_state Regs -> list (result a e * sequential_state Regs).
+
+(*val returnS : forall 'regs 'a 'e. 'a -> monadS 'regs 'a 'e*)
+Definition returnS {Regs A E} (a:A) : monadS Regs A E := fun s => [(Value a,s)].
+
+(*val bindS : forall 'regs 'a 'b 'e. monadS 'regs 'a 'e -> ('a -> monadS 'regs 'b 'e) -> monadS 'regs 'b 'e*)
+Definition bindS {Regs A B E} (m : monadS Regs A E) (f : A -> monadS Regs B E) : monadS Regs B E :=
+ fun (s : sequential_state Regs) =>
+ List.flat_map (fun v => match v with
+ | (Value a, s') => f a s'
+ | (Ex e, s') => [(Ex e, s')]
+ end) (m s).
+
+(*val seqS: forall 'regs 'b 'e. monadS 'regs unit 'e -> monadS 'regs 'b 'e -> monadS 'regs 'b 'e*)
+Definition seqS {Regs B E} (m : monadS Regs unit E) (n : monadS Regs B E) : monadS Regs B E :=
+ bindS m (fun (_ : unit) => n).
+(*
+let inline (>>$=) = bindS
+let inline (>>$) = seqS
+*)
+Notation "m >>$= f" := (bindS m f) (at level 50, left associativity).
+Notation "m >>$ n" := (seqS m n) (at level 50, left associativity).
+
+(*val chooseS : forall 'regs 'a 'e. SetType 'a => list 'a -> monadS 'regs 'a 'e*)
+Definition chooseS {Regs A E} (xs : list A) : monadS Regs A E :=
+ fun s => (List.map (fun x => (Value x, s)) xs).
+
+(*val readS : forall 'regs 'a 'e. (sequential_state 'regs -> 'a) -> monadS 'regs 'a 'e*)
+Definition readS {Regs A E} (f : sequential_state Regs -> A) : monadS Regs A E :=
+ (fun s => returnS (f s) s).
+
+(*val updateS : forall 'regs 'e. (sequential_state 'regs -> sequential_state 'regs) -> monadS 'regs unit 'e*)
+Definition updateS {Regs E} (f : sequential_state Regs -> sequential_state Regs) : monadS Regs unit E :=
+ (fun s => returnS tt (f s)).
+
+(*val failS : forall 'regs 'a 'e. string -> monadS 'regs 'a 'e*)
+Definition failS {Regs A E} msg : monadS Regs A E :=
+ fun s => [(Ex (Failure msg), s)].
+
+(*val choose_boolS : forall 'regval 'regs 'a 'e. unit -> monadS 'regs bool 'e*)
+Definition choose_boolS {Regs E} (_:unit) : monadS Regs bool E :=
+ chooseS [false; true].
+Definition undefined_boolS {Regs E} := @choose_boolS Regs E.
+
+(*val exitS : forall 'regs 'e 'a. unit -> monadS 'regs 'a 'e*)
+Definition exitS {Regs A E} (_:unit) : monadS Regs A E := failS "exit".
+
+(*val throwS : forall 'regs 'a 'e. 'e -> monadS 'regs 'a 'e*)
+Definition throwS {Regs A E} (e : E) :monadS Regs A E :=
+ fun s => [(Ex (Throw e), s)].
+
+(*val try_catchS : forall 'regs 'a 'e1 'e2. monadS 'regs 'a 'e1 -> ('e1 -> monadS 'regs 'a 'e2) -> monadS 'regs 'a 'e2*)
+Definition try_catchS {Regs A E1 E2} (m : monadS Regs A E1) (h : E1 -> monadS Regs A E2) : monadS Regs A E2 :=
+fun s =>
+ List.flat_map (fun v => match v with
+ | (Value a, s') => returnS a s'
+ | (Ex (Throw e), s') => h e s'
+ | (Ex (Failure msg), s') => [(Ex (Failure msg), s')]
+ end) (m s).
+
+(*val assert_expS : forall 'regs 'e. bool -> string -> monadS 'regs unit 'e*)
+Definition assert_expS {Regs E} (exp : bool) (msg : string) : monadS Regs unit E :=
+ if exp then returnS tt else failS msg.
+
+(* For early return, we abuse exceptions by throwing and catching
+ the return value. The exception type is "either 'r 'e", where "Right e"
+ represents a proper exception and "Left r" an early return of value "r". *)
+Definition monadRS Regs A R E := monadS Regs A (sum R E).
+
+(*val early_returnS : forall 'regs 'a 'r 'e. 'r -> monadRS 'regs 'a 'r 'e*)
+Definition early_returnS {Regs A R E} (r : R) : monadRS Regs A R E := throwS (inl r).
+
+(*val catch_early_returnS : forall 'regs 'a 'e. monadRS 'regs 'a 'a 'e -> monadS 'regs 'a 'e*)
+Definition catch_early_returnS {Regs A E} (m : monadRS Regs A A E) : monadS Regs A E :=
+ try_catchS m
+ (fun v => match v with
+ | inl a => returnS a
+ | inr e => throwS e
+ end).
+
+(* Lift to monad with early return by wrapping exceptions *)
+(*val liftRS : forall 'a 'r 'regs 'e. monadS 'regs 'a 'e -> monadRS 'regs 'a 'r 'e*)
+Definition liftRS {A R Regs E} (m : monadS Regs A E) : monadRS Regs A R E :=
+ try_catchS m (fun e => throwS (inr e)).
+
+(* Catch exceptions in the presence of early returns *)
+(*val try_catchRS : forall 'regs 'a 'r 'e1 'e2. monadRS 'regs 'a 'r 'e1 -> ('e1 -> monadRS 'regs 'a 'r 'e2) -> monadRS 'regs 'a 'r 'e2*)
+Definition try_catchRS {Regs A R E1 E2} (m : monadRS Regs A R E1) (h : E1 -> monadRS Regs A R E2) : monadRS Regs A R E2 :=
+ try_catchS m
+ (fun v => match v with
+ | inl r => throwS (inl r)
+ | inr e => h e
+ end).
+
+(*val maybe_failS : forall 'regs 'a 'e. string -> maybe 'a -> monadS 'regs 'a 'e*)
+Definition maybe_failS {Regs A E} msg (v : option A) : monadS Regs A E :=
+match v with
+ | Some a => returnS a
+ | None => failS msg
+end.
+
+(*val read_tagS : forall 'regs 'a 'e. Bitvector 'a => 'a -> monadS 'regs bitU 'e*)
+Definition read_tagS {Regs A E} (addr : mword A) : monadS Regs bitU E :=
+ let addr := Word.wordToNat (get_word addr) in
+ readS (fun s => opt_def B0 (NatMap.find addr s.(ss_tagstate))).
+
+Fixpoint genlist_acc {A:Type} (f : nat -> A) n acc : list A :=
+ match n with
+ | O => acc
+ | S n' => genlist_acc f n' (f n' :: acc)
+ end.
+Definition genlist {A} f n := @genlist_acc A f n [].
+
+
+(* Read bytes from memory and return in little endian order *)
+(*val get_mem_bytes : forall 'regs. nat -> nat -> sequential_state 'regs -> maybe (list memory_byte * bitU)*)
+Definition get_mem_bytes {Regs} addr sz (s : sequential_state Regs) : option (list memory_byte * bitU) :=
+ let addrs := genlist (fun n => addr + n)%nat sz in
+ let read_byte s addr := NatMap.find addr s.(ss_memstate) in
+ let read_tag s addr := opt_def B0 (NatMap.find addr s.(ss_tagstate)) in
+ option_map
+ (fun mem_val => (mem_val, List.fold_left and_bit (List.map (read_tag s) addrs) B1))
+ (just_list (List.map (read_byte s) addrs)).
+
+(*val read_memt_bytesS : forall 'regs 'e. read_kind -> nat -> nat -> monadS 'regs (list memory_byte * bitU) 'e*)
+Definition read_memt_bytesS {Regs E} (_ : read_kind) addr sz : monadS Regs (list memory_byte * bitU) E :=
+ readS (get_mem_bytes addr sz) >>$=
+ maybe_failS "read_memS".
+
+(*val read_mem_bytesS : forall 'regs 'e. read_kind -> nat -> nat -> monadS 'regs (list memory_byte) 'e*)
+Definition read_mem_bytesS {Regs E} (rk : read_kind) addr sz : monadS Regs (list memory_byte) E :=
+ read_memt_bytesS rk addr sz >>$= (fun '(bytes, _) =>
+ returnS bytes).
+
+(*val read_memtS : forall 'regs 'e 'a 'b. Bitvector 'a, Bitvector 'b => read_kind -> 'a -> integer -> monadS 'regs ('b * bitU) 'e*)
+Definition read_memtS {Regs E A B} (rk : read_kind) (a : mword A) sz `{ArithFact (B >= 0)} : monadS Regs (mword B * bitU) E :=
+ let a := Word.wordToNat (get_word a) in
+ read_memt_bytesS rk a (Z.to_nat sz) >>$= (fun '(bytes, tag) =>
+ maybe_failS "bits_of_mem_bytes" (of_bits (bits_of_mem_bytes bytes)) >>$= (fun mem_val =>
+ returnS (mem_val, tag))).
+
+(*val read_memS : forall 'regs 'e 'a 'b. Bitvector 'a, Bitvector 'b => read_kind -> 'a -> integer -> monadS 'regs 'b 'e*)
+Definition read_memS {Regs E A B} rk (a : mword A) sz `{ArithFact (B >= 0)} : monadS Regs (mword B) E :=
+ read_memtS rk a sz >>$= (fun '(bytes, _) =>
+ returnS bytes).
+
+(*val excl_resultS : forall 'regs 'e. unit -> monadS 'regs bool 'e*)
+Definition excl_resultS {Regs E} : unit -> monadS Regs bool E :=
+ (* TODO: This used to be more deterministic, checking a flag in the state
+ whether an exclusive load has occurred before. However, this does not
+ seem very precise; it might be safer to overapproximate the possible
+ behaviours by always making a nondeterministic choice. *)
+ @undefined_boolS Regs E.
+
+(* Write little-endian list of bytes to given address *)
+(*val put_mem_bytes : forall 'regs. nat -> nat -> list memory_byte -> bitU -> sequential_state 'regs -> sequential_state 'regs*)
+Definition put_mem_bytes {Regs} addr sz (v : list memory_byte) (tag : bitU) (s : sequential_state Regs) : sequential_state Regs :=
+ let addrs := genlist (fun n => addr + n)%nat sz in
+ let a_v := List.combine addrs v in
+ let write_byte mem '(addr, v) := NatMap.add addr v mem in
+ let write_tag mem addr := NatMap.add addr tag mem in
+ {| ss_regstate := s.(ss_regstate);
+ ss_memstate := List.fold_left write_byte a_v s.(ss_memstate);
+ ss_tagstate := List.fold_left write_tag addrs s.(ss_tagstate) |}.
+
+(*val write_memt_bytesS : forall 'regs 'e. write_kind -> nat -> nat -> list memory_byte -> bitU -> monadS 'regs bool 'e*)
+Definition write_memt_bytesS {Regs E} (_ : write_kind) addr sz (v : list memory_byte) (t : bitU) : monadS Regs bool E :=
+ updateS (put_mem_bytes addr sz v t) >>$
+ returnS true.
+
+(*val write_mem_bytesS : forall 'regs 'e. write_kind -> nat -> nat -> list memory_byte -> monadS 'regs bool 'e*)
+Definition write_mem_bytesS {Regs E} wk addr sz (v : list memory_byte) : monadS Regs bool E :=
+ write_memt_bytesS wk addr sz v B0.
+
+(*val write_memtS : forall 'regs 'e 'a 'b. Bitvector 'a, Bitvector 'b =>
+ write_kind -> 'a -> integer -> 'b -> bitU -> monadS 'regs bool 'e*)
+Definition write_memtS {Regs E A B} wk (addr : mword A) sz (v : mword B) (t : bitU) : monadS Regs bool E :=
+ match (Word.wordToNat (get_word addr), mem_bytes_of_bits v) with
+ | (addr, Some v) => write_memt_bytesS wk addr (Z.to_nat sz) v t
+ | _ => failS "write_mem"
+ end.
+
+(*val write_memS : forall 'regs 'e 'a 'b. Bitvector 'a, Bitvector 'b =>
+ write_kind -> 'a -> integer -> 'b -> monadS 'regs bool 'e*)
+Definition write_memS {Regs E A B} wk (addr : mword A) sz (v : mword B) : monadS Regs bool E :=
+ write_memtS wk addr sz v B0.
+
+(*val read_regS : forall 'regs 'rv 'a 'e. register_ref 'regs 'rv 'a -> monadS 'regs 'a 'e*)
+Definition read_regS {Regs RV A E} (reg : register_ref Regs RV A) : monadS Regs A E :=
+ readS (fun s => reg.(read_from) s.(ss_regstate)).
+
+(* TODO
+let read_reg_range reg i j state =
+ let v = slice (get_reg state (name_of_reg reg)) i j in
+ [(Value (vec_to_bvec v),state)]
+let read_reg_bit reg i state =
+ let v = access (get_reg state (name_of_reg reg)) i in
+ [(Value v,state)]
+let read_reg_field reg regfield =
+ let (i,j) = register_field_indices reg regfield in
+ read_reg_range reg i j
+let read_reg_bitfield reg regfield =
+ let (i,_) = register_field_indices reg regfield in
+ read_reg_bit reg i *)
+
+(*val read_regvalS : forall 'regs 'rv 'e.
+ register_accessors 'regs 'rv -> string -> monadS 'regs 'rv 'e*)
+Definition read_regvalS {Regs RV E} (acc : register_accessors Regs RV) reg : monadS Regs RV E :=
+ let '(read, _) := acc in
+ readS (fun s => read reg s.(ss_regstate)) >>$= (fun v => match v with
+ | Some v => returnS v
+ | None => failS ("read_regvalS " ++ reg)
+ end).
+
+(*val write_regvalS : forall 'regs 'rv 'e.
+ register_accessors 'regs 'rv -> string -> 'rv -> monadS 'regs unit 'e*)
+Definition write_regvalS {Regs RV E} (acc : register_accessors Regs RV) reg (v : RV) : monadS Regs unit E :=
+ let '(_, write) := acc in
+ readS (fun s => write reg v s.(ss_regstate)) >>$= (fun x => match x with
+ | Some rs' => updateS (fun s => {| ss_regstate := rs'; ss_memstate := s.(ss_memstate); ss_tagstate := s.(ss_tagstate) |})
+ | None => failS ("write_regvalS " ++ reg)
+ end).
+
+(*val write_regS : forall 'regs 'rv 'a 'e. register_ref 'regs 'rv 'a -> 'a -> monadS 'regs unit 'e*)
+Definition write_regS {Regs RV A E} (reg : register_ref Regs RV A) (v:A) : monadS Regs unit E :=
+ updateS (fun s => {| ss_regstate := reg.(write_to) v s.(ss_regstate); ss_memstate := s.(ss_memstate); ss_tagstate := s.(ss_tagstate) |}).
+
+(* TODO
+val update_reg : forall 'regs 'rv 'a 'b 'e. register_ref 'regs 'rv 'a -> ('a -> 'b -> 'a) -> 'b -> monadS 'regs unit 'e
+let update_reg reg f v state =
+ let current_value = get_reg state reg in
+ let new_value = f current_value v in
+ [(Value (), set_reg state reg new_value)]
+
+let write_reg_field reg regfield = update_reg reg regfield.set_field
+
+val update_reg_range : forall 'regs 'rv 'a 'b. Bitvector 'a, Bitvector 'b => register_ref 'regs 'rv 'a -> integer -> integer -> 'a -> 'b -> 'a
+let update_reg_range reg i j reg_val new_val = set_bits (reg.is_inc) reg_val i j (bits_of new_val)
+let write_reg_range reg i j = update_reg reg (update_reg_range reg i j)
+
+let update_reg_pos reg i reg_val x = update_list reg.is_inc reg_val i x
+let write_reg_pos reg i = update_reg reg (update_reg_pos reg i)
+
+let update_reg_bit reg i reg_val bit = set_bit (reg.is_inc) reg_val i (to_bitU bit)
+let write_reg_bit reg i = update_reg reg (update_reg_bit reg i)
+
+let update_reg_field_range regfield i j reg_val new_val =
+ let current_field_value = regfield.get_field reg_val in
+ let new_field_value = set_bits (regfield.field_is_inc) current_field_value i j (bits_of new_val) in
+ regfield.set_field reg_val new_field_value
+let write_reg_field_range reg regfield i j = update_reg reg (update_reg_field_range regfield i j)
+
+let update_reg_field_pos regfield i reg_val x =
+ let current_field_value = regfield.get_field reg_val in
+ let new_field_value = update_list regfield.field_is_inc current_field_value i x in
+ regfield.set_field reg_val new_field_value
+let write_reg_field_pos reg regfield i = update_reg reg (update_reg_field_pos regfield i)
+
+let update_reg_field_bit regfield i reg_val bit =
+ let current_field_value = regfield.get_field reg_val in
+ let new_field_value = set_bit (regfield.field_is_inc) current_field_value i (to_bitU bit) in
+ regfield.set_field reg_val new_field_value
+let write_reg_field_bit reg regfield i = update_reg reg (update_reg_field_bit regfield i)*)
+
+(* TODO Add Show typeclass for value and exception type *)
+(*val show_result : forall 'a 'e. result 'a 'e -> string*)
+Definition show_result {A E} (x : result A E) : string := match x with
+ | Value _ => "Value ()"
+ | Ex (Failure msg) => "Failure " ++ msg
+ | Ex (Throw _) => "Throw"
+end.
+
+(*val prerr_results : forall 'a 'e 's. SetType 's => set (result 'a 'e * 's) -> unit*)
+Definition prerr_results {A E S} (rs : list (result A E * S)) : unit := tt.
+(* let _ = Set.map (fun (r, _) -> let _ = prerr_endline (show_result r) in ()) rs in
+ ()*)
+
diff --git a/prover_snapshots/coq/lib/sail/Sail2_state_monad_lemmas.v b/prover_snapshots/coq/lib/sail/Sail2_state_monad_lemmas.v
new file mode 100644
index 0000000..99fef32
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_state_monad_lemmas.v
@@ -0,0 +1,542 @@
+Require Import Sail2_state_monad.
+(*Require Import Sail2_values_lemmas.*)
+Require Export Setoid.
+Require Export Morphisms Equivalence.
+
+(* Ensure that pointwise equality on states is the preferred notion of
+ equivalence for the state monad. *)
+Local Open Scope equiv_scope.
+Instance monadS_equivalence {Regs A E} :
+ Equivalence (pointwise_relation (sequential_state Regs) (@eq (list (result A E * sequential_state Regs)))) | 9.
+split; apply _.
+Qed.
+
+Global Instance refl_eq_subrelation {A : Type} {R : A -> A -> Prop} `{Reflexive A R} : subrelation eq R.
+intros x y EQ. subst. reflexivity.
+Qed.
+
+Hint Extern 4 (_ === _) => reflexivity.
+Hint Extern 4 (_ === _) => symmetry.
+
+Lemma bindS_ext_cong (*[fundef_cong]:*) {Regs A B E}
+ {m1 m2 : monadS Regs A E} {f1 f2 : A -> monadS Regs B E} s :
+ m1 s = m2 s ->
+ (forall a s', List.In (Value a, s') (m2 s) -> f1 a s' = f2 a s') ->
+ bindS m1 f1 s = bindS m2 f2 s.
+intros.
+unfold bindS.
+rewrite H.
+rewrite !List.flat_map_concat_map.
+f_equal.
+apply List.map_ext_in.
+intros [[a|a] s'] H_in; auto.
+Qed.
+
+(* Weaker than the Isabelle version, but avoids talking about individual states *)
+Lemma bindS_cong (*[fundef_cong]:*) Regs A B E m1 m2 (f1 f2 : A -> monadS Regs B E) :
+ m1 === m2 ->
+ (forall a, f1 a === f2 a) ->
+ bindS m1 f1 === bindS m2 f2.
+intros M F s.
+apply bindS_ext_cong; intros; auto.
+apply F.
+Qed.
+
+Add Parametric Morphism {Regs A B E : Type} : (@bindS Regs A B E)
+ with signature equiv ==> equiv ==> equiv as bindS_morphism.
+auto using bindS_cong.
+Qed.
+
+Lemma bindS_returnS_left Regs A B E {x : A} {f : A -> monadS Regs B E} :
+ bindS (returnS x) f === f x.
+intro s.
+unfold returnS, bindS.
+simpl.
+auto using List.app_nil_r.
+Qed.
+Hint Rewrite bindS_returnS_left : state.
+
+Lemma bindS_returnS_right Regs A E {m : monadS Regs A E} :
+ bindS m returnS === m.
+intro s.
+unfold returnS, bindS.
+induction (m s) as [|[[a|a] s'] t]; auto;
+simpl;
+rewrite IHt;
+reflexivity.
+Qed.
+Hint Rewrite bindS_returnS_right : state.
+
+Lemma bindS_readS {Regs A E} {f} {m : A -> monadS Regs A E} {s} :
+ bindS (readS f) m s = m (f s) s.
+unfold readS, bindS.
+simpl.
+rewrite List.app_nil_r.
+reflexivity.
+Qed.
+
+Lemma bindS_updateS {Regs A E} {f : sequential_state Regs -> sequential_state Regs} {m : unit -> monadS Regs A E} {s} :
+ bindS (updateS f) m s = m tt (f s).
+unfold updateS, bindS.
+simpl.
+auto using List.app_nil_r.
+Qed.
+
+Lemma bindS_assertS_true Regs A E msg {f : unit -> monadS Regs A E} :
+ bindS (assert_expS true msg) f === f tt.
+intro s.
+unfold assert_expS, bindS.
+simpl.
+auto using List.app_nil_r.
+Qed.
+Hint Rewrite bindS_assertS_true : state.
+
+Lemma bindS_chooseS_returnS (*[simp]:*) Regs A B E {xs : list A} {f : A -> B} :
+ bindS (Regs := Regs) (E := E) (chooseS xs) (fun x => returnS (f x)) === chooseS (List.map f xs).
+intro s.
+unfold chooseS, bindS, returnS.
+induction xs; auto.
+simpl. rewrite IHxs.
+reflexivity.
+Qed.
+Hint Rewrite bindS_chooseS_returnS : state.
+
+Lemma result_cases : forall (A E : Type) (P : result A E -> Prop),
+ (forall a, P (Value a)) ->
+ (forall e, P (Ex (Throw e))) ->
+ (forall msg, P (Ex (Failure msg))) ->
+ forall r, P r.
+intros.
+destruct r; auto.
+destruct e; auto.
+Qed.
+
+Lemma result_state_cases {A E S} {P : result A E * S -> Prop} :
+ (forall a s, P (Value a, s)) ->
+ (forall e s, P (Ex (Throw e), s)) ->
+ (forall msg s, P (Ex (Failure msg), s)) ->
+ forall rs, P rs.
+intros.
+destruct rs as [[a|[e|msg]] s]; auto.
+Qed.
+
+(* TODO: needs sets, not lists
+Lemma monadS_ext_eqI {Regs A E} {m m' : monadS Regs A E} s :
+ (forall a s', List.In (Value a, s') (m s) <-> List.In (Value a, s') (m' s)) ->
+ (forall e s', List.In (Ex (Throw e), s') (m s) <-> List.In (Ex (Throw e), s') (m' s)) ->
+ (forall msg s', List.In (Ex (Failure msg), s') (m s) <-> List.In (Ex (Failure msg), s') (m' s)) ->
+ m s = m' s.
+proof (intro set_eqI)
+ fix x
+ show "x \<in> m s \<longleftrightarrow> x \<in> m' s" using assms by (cases x rule: result_state_cases) auto
+qed
+
+lemma monadS_eqI:
+ fixes m m' :: "('regs, 'a, 'e) monadS"
+ assumes "\<And>s a s'. (Value a, s') \<in> m s \<longleftrightarrow> (Value a, s') \<in> m' s"
+ and "\<And>s e s'. (Ex (Throw e), s') \<in> m s \<longleftrightarrow> (Ex (Throw e), s') \<in> m' s"
+ and "\<And>s msg s'. (Ex (Failure msg), s') \<in> m s \<longleftrightarrow> (Ex (Failure msg), s') \<in> m' s"
+ shows "m = m'"
+ using assms by (intro ext monadS_ext_eqI)
+*)
+
+Lemma bindS_cases {Regs A B E} {m} {f : A -> monadS Regs B E} {r s s'} :
+ List.In (r, s') (bindS m f s) ->
+ (exists a a' s'', r = Value a /\ List.In (Value a', s'') (m s) /\ List.In (Value a, s') (f a' s'')) \/
+ (exists e, r = Ex e /\ List.In (Ex e, s') (m s)) \/
+ (exists e a s'', r = Ex e /\ List.In (Value a, s'') (m s) /\ List.In (Ex e, s') (f a s'')).
+unfold bindS.
+intro IN.
+apply List.in_flat_map in IN.
+destruct IN as [[r' s''] [INr' INr]].
+destruct r' as [a'|e'].
+* destruct r as [a|e].
+ + left. eauto 10.
+ + right; right. eauto 10.
+* right; left. simpl in INr. destruct INr as [|[]]. inversion H. subst. eauto 10.
+Qed.
+
+Lemma bindS_intro_Value {Regs A B E} {m} {f : A -> monadS Regs B E} {s a s' a' s''} :
+ List.In (Value a', s'') (m s) -> List.In (Value a, s') (f a' s'') -> List.In (Value a, s') (bindS m f s).
+intros; unfold bindS.
+apply List.in_flat_map.
+eauto.
+Qed.
+Lemma bindS_intro_Ex_left {Regs A B E} {m} {f : A -> monadS Regs B E} {s e s'} :
+ List.In (Ex e, s') (m s) -> List.In (Ex e, s') (bindS m f s).
+intros; unfold bindS.
+apply List.in_flat_map.
+exists (Ex e, s').
+auto with datatypes.
+Qed.
+Lemma bindS_intro_Ex_right {Regs A B E} {m} {f : A -> monadS Regs B E} {s e s' a s''} :
+ List.In (Ex e, s') (f a s'') -> List.In (Value a, s'') (m s) -> List.In (Ex e, s') (bindS m f s).
+intros; unfold bindS.
+apply List.in_flat_map.
+eauto.
+Qed.
+Hint Resolve bindS_intro_Value bindS_intro_Ex_left bindS_intro_Ex_right : bindS_intros.
+
+Lemma bindS_assoc Regs A B C E {m} {f : A -> monadS Regs B E} {g : B -> monadS Regs C E} :
+ bindS (bindS m f) g === bindS m (fun x => bindS (f x) g).
+intro s.
+unfold bindS.
+induction (m s) as [ | [[a | e] t]].
+* reflexivity.
+* simpl. rewrite <- IHl.
+ rewrite !List.flat_map_concat_map.
+ rewrite List.map_app.
+ rewrite List.concat_app.
+ reflexivity.
+* simpl. rewrite IHl. reflexivity.
+Qed.
+Hint Rewrite bindS_assoc : state.
+
+Lemma bindS_failS Regs A B E {msg} {f : A -> monadS Regs B E} :
+ bindS (failS msg) f = failS msg.
+reflexivity.
+Qed.
+Hint Rewrite bindS_failS : state.
+
+Lemma bindS_throwS Regs A B E {e} {f : A -> monadS Regs B E} :
+ bindS (throwS e) f = throwS e.
+reflexivity.
+Qed.
+Hint Rewrite bindS_throwS : state.
+
+(*declare seqS_def[simp]*)
+Lemma seqS_def Regs A E m (m' : monadS Regs A E) :
+ m >>$ m' = m >>$= (fun _ => m').
+reflexivity.
+Qed.
+Hint Rewrite seqS_def : state.
+
+Lemma Value_bindS_elim {Regs A B E} {a m} {f : A -> monadS Regs B E} {s s'} :
+ List.In (Value a, s') (bindS m f s) ->
+ exists s'' a', List.In (Value a', s'') (m s) /\ List.In (Value a, s') (f a' s'').
+intro H.
+apply bindS_cases in H.
+destruct H as [(a0 & a' & s'' & [= <-] & [*]) | [(e & [= ] & _) | (_ & _ & _ & [= ] & _)]].
+eauto.
+Qed.
+
+Lemma Ex_bindS_elim {Regs A B E} {e m s s'} {f : A -> monadS Regs B E} :
+ List.In (Ex e, s') (bindS m f s) ->
+ List.In (Ex e, s') (m s) \/
+ exists s'' a', List.In (Value a', s'') (m s) /\ List.In (Ex e, s') (f a' s'').
+intro H.
+apply bindS_cases in H.
+destruct H as [(? & ? & ? & [= ] & _) | [(? & [= <-] & X) | (? & ? & ? & [= <-] & X)]];
+eauto.
+Qed.
+
+Lemma try_catchS_returnS Regs A E1 E2 {a} {h : E1 -> monadS Regs A E2}:
+ try_catchS (returnS a) h = returnS a.
+reflexivity.
+Qed.
+Hint Rewrite try_catchS_returnS : state.
+Lemma try_catchS_failS Regs A E1 E2 {msg} {h : E1 -> monadS Regs A E2}:
+ try_catchS (failS msg) h = failS msg.
+reflexivity.
+Qed.
+Hint Rewrite try_catchS_failS : state.
+Lemma try_catchS_throwS Regs A E1 E2 {e} {h : E1 -> monadS Regs A E2}:
+ try_catchS (throwS e) h === h e.
+intro s.
+unfold try_catchS, throwS.
+simpl.
+auto using List.app_nil_r.
+Qed.
+Hint Rewrite try_catchS_throwS : state.
+
+Lemma try_catchS_cong (*[cong]:*) {Regs A E1 E2 m1 m2} {h1 h2 : E1 -> monadS Regs A E2} :
+ m1 === m2 ->
+ (forall e, h1 e === h2 e) ->
+ try_catchS m1 h1 === try_catchS m2 h2.
+intros H1 H2 s.
+unfold try_catchS.
+rewrite H1.
+rewrite !List.flat_map_concat_map.
+f_equal.
+apply List.map_ext_in.
+intros [[a|[e|msg]] s'] H_in; auto. apply H2.
+Qed.
+
+Add Parametric Morphism {Regs A E1 E2 : Type} : (@try_catchS Regs A E1 E2)
+ with signature equiv ==> equiv ==> equiv as try_catchS_morphism.
+intros. auto using try_catchS_cong.
+Qed.
+
+Add Parametric Morphism {Regs A E : Type} : (@catch_early_returnS Regs A E)
+ with signature equiv ==> equiv as catch_early_returnS_morphism.
+intros.
+unfold catch_early_returnS.
+rewrite H.
+reflexivity.
+Qed.
+
+Lemma try_catchS_cases {Regs A E1 E2 m} {h : E1 -> monadS Regs A E2} {r s s'} :
+ List.In (r, s') (try_catchS m h s) ->
+ (exists a, r = Value a /\ List.In (Value a, s') (m s)) \/
+ (exists msg, r = Ex (Failure msg) /\ List.In (Ex (Failure msg), s') (m s)) \/
+ (exists e s'', List.In (Ex (Throw e), s'') (m s) /\ List.In (r, s') (h e s'')).
+unfold try_catchS.
+intro IN.
+apply List.in_flat_map in IN.
+destruct IN as [[r' s''] [INr' INr]].
+destruct r' as [a'|[e'|msg]].
+* left. simpl in INr. destruct INr as [[= <- <-] | []]. eauto 10.
+* simpl in INr. destruct INr as [[= <- <-] | []]. eauto 10.
+* eauto 10.
+Qed.
+
+Lemma try_catchS_intros {Regs A E1 E2} {m} {h : E1 -> monadS Regs A E2} :
+ (forall s a s', List.In (Value a, s') (m s) -> List.In (Value a, s') (try_catchS m h s)) /\
+ (forall s msg s', List.In (Ex (Failure msg), s') (m s) -> List.In (Ex (Failure msg), s') (try_catchS m h s)) /\
+ (forall s e s'' r s', List.In (Ex (Throw e), s'') (m s) -> List.In (r, s') (h e s'') -> List.In (r, s') (try_catchS m h s)).
+repeat split; unfold try_catchS; intros;
+apply List.in_flat_map.
+* eexists; split; [ apply H | ]. simpl. auto.
+* eexists; split; [ apply H | ]. simpl. auto.
+* eexists; split; [ apply H | ]. simpl. auto.
+Qed.
+
+Lemma no_Ex_basic_builtins (*[simp]:*) {Regs E} {s s' : sequential_state Regs} {e : ex E} :
+ (forall A (a:A), ~ List.In (Ex e, s') (returnS a s)) /\
+ (forall A (f : _ -> A), ~ List.In (Ex e, s') (readS f s)) /\
+ (forall f, ~ List.In (Ex e, s') (updateS f s)) /\
+ (forall A (xs : list A), ~ List.In (Ex e, s') (chooseS xs s)).
+repeat split; intros;
+unfold returnS, readS, updateS, chooseS; simpl;
+try intuition congruence.
+* intro H.
+ apply List.in_map_iff in H.
+ destruct H as [x [X _]].
+ congruence.
+Qed.
+
+Import List.ListNotations.
+Definition ignore_throw_aux {A E1 E2 S} (rs : result A E1 * S) : list (result A E2 * S) :=
+match rs with
+| (Value a, s') => [(Value a, s')]
+| (Ex (Throw e), s') => []
+| (Ex (Failure msg), s') => [(Ex (Failure msg), s')]
+end.
+Definition ignore_throw {A E1 E2 S} (m : S -> list (result A E1 * S)) s : list (result A E2 * S) :=
+ List.flat_map ignore_throw_aux (m s).
+
+Lemma ignore_throw_cong {Regs A E1 E2} {m1 m2 : monadS Regs A E1} :
+ m1 === m2 ->
+ ignore_throw (E2 := E2) m1 === ignore_throw m2.
+intros H s.
+unfold ignore_throw.
+rewrite H.
+reflexivity.
+Qed.
+
+Lemma ignore_throw_aux_member_simps (*[simp]:*) {A E1 E2 S} {s' : S} {ms} :
+ (forall a:A, List.In (Value a, s') (ignore_throw_aux (E1 := E1) (E2 := E2) ms) <-> ms = (Value a, s')) /\
+ (forall e, ~ List.In (Ex (E := E2) (Throw e), s') (ignore_throw_aux ms)) /\
+ (forall msg, List.In (Ex (E := E2) (Failure msg), s') (ignore_throw_aux ms) <-> ms = (Ex (Failure msg), s')).
+destruct ms as [[a' | [e' | msg']] s]; simpl;
+intuition congruence.
+Qed.
+
+Lemma ignore_throw_member_simps (*[simp]:*) {A E1 E2 S} {s s' : S} {m} :
+ (forall {a:A}, List.In (Value (E := E2) a, s') (ignore_throw m s) <-> List.In (Value (E := E1) a, s') (m s)) /\
+ (forall {a:A}, List.In (Value (E := E2) a, s') (ignore_throw m s) <-> List.In (Value a, s') (m s)) /\
+ (forall e, ~ List.In (Ex (E := E2) (Throw e), s') (ignore_throw m s)) /\
+ (forall {msg}, List.In (Ex (E := E2) (Failure msg), s') (ignore_throw m s) <-> List.In (Ex (Failure msg), s') (m s)).
+unfold ignore_throw.
+repeat apply conj; intros; try apply conj;
+rewrite ?List.in_flat_map;
+solve
+[ intros [x [H1 H2]]; apply ignore_throw_aux_member_simps in H2; congruence
+| intro H; eexists; split; [ apply H | apply ignore_throw_aux_member_simps; reflexivity] ].
+Qed.
+
+Lemma ignore_throw_cases {A E S} {m : S -> list (result A E * S)} {r s s'} :
+ ignore_throw m s = m s ->
+ List.In (r, s') (m s) ->
+ (exists a, r = Value a) \/
+ (exists msg, r = Ex (Failure msg)).
+destruct r as [a | [e | msg]]; eauto.
+* intros H1 H2. rewrite <- H1 in H2.
+ apply ignore_throw_member_simps in H2.
+ destruct H2.
+Qed.
+
+(* *** *)
+Lemma flat_map_app {A B} {f : A -> list B} {l1 l2} :
+ List.flat_map f (l1 ++ l2) = (List.flat_map f l1 ++ List.flat_map f l2)%list.
+rewrite !List.flat_map_concat_map.
+rewrite List.map_app, List.concat_app.
+reflexivity.
+Qed.
+
+Lemma ignore_throw_bindS (*[simp]:*) Regs A B E E2 {m} {f : A -> monadS Regs B E} :
+ ignore_throw (E2 := E2) (bindS m f) === bindS (ignore_throw m) (fun s => ignore_throw (f s)).
+intro s.
+unfold bindS, ignore_throw.
+induction (m s) as [ | [[a | [e | msg]] t]].
+* reflexivity.
+* simpl. rewrite <- IHl. rewrite flat_map_app. reflexivity.
+* simpl. rewrite <- IHl. reflexivity.
+* simpl. apply IHl.
+Qed.
+Hint Rewrite ignore_throw_bindS : ignore_throw.
+
+Lemma try_catchS_bindS_no_throw {Regs A B E1 E2} {m1 : monadS Regs A E1} {m2 : monadS Regs A E2} {f : A -> monadS Regs B _} {h} :
+ ignore_throw m1 === m1 ->
+ ignore_throw m1 === m2 ->
+ try_catchS (bindS m1 f) h === bindS m2 (fun a => try_catchS (f a) h).
+intros Ignore1 Ignore2.
+transitivity ((ignore_throw m1 >>$= (fun a => try_catchS (f a) h))).
+* intro s.
+ unfold bindS, try_catchS, ignore_throw.
+ specialize (Ignore1 s). revert Ignore1. unfold ignore_throw.
+ induction (m1 s) as [ | [[a | [e | msg]] t]]; auto.
+ + intro Ig. simpl. rewrite flat_map_app. rewrite IHl. auto. injection Ig. auto.
+ + intro Ig. simpl. rewrite IHl. reflexivity. injection Ig. auto.
+ + intro Ig. exfalso. clear -Ig.
+ assert (List.In (Ex (Throw msg), t) (List.flat_map ignore_throw_aux l)).
+ simpl in Ig. rewrite Ig. simpl. auto.
+ apply List.in_flat_map in H.
+ destruct H as [x [H1 H2]].
+ apply ignore_throw_aux_member_simps in H2.
+ assumption.
+* apply bindS_cong; auto.
+Qed.
+
+Lemma concat_map_singleton {A B} {f : A -> B} {a : list A} :
+ List.concat (List.map (fun x => [f x]%list) a) = List.map f a.
+induction a; simpl; try rewrite IHa; auto with datatypes.
+Qed.
+
+(*lemma no_throw_basic_builtins[simp]:*)
+Lemma no_throw_basic_builtins_1 Regs A E E2 {a : A} :
+ ignore_throw (E1 := E2) (returnS a) = @returnS Regs A E a.
+reflexivity. Qed.
+Lemma no_throw_basic_builtins_2 Regs A E E2 {f : sequential_state Regs -> A} :
+ ignore_throw (E1 := E) (E2 := E2) (readS f) = readS f.
+reflexivity. Qed.
+Lemma no_throw_basic_builtins_3 Regs E E2 {f : sequential_state Regs -> sequential_state Regs} :
+ ignore_throw (E1 := E) (E2 := E2) (updateS f) = updateS f.
+reflexivity. Qed.
+Lemma no_throw_basic_builtins_4 Regs A E1 E2 {xs : list A} :
+ ignore_throw (E1 := E1) (chooseS xs) === @chooseS Regs A E2 xs.
+intro s.
+unfold ignore_throw, chooseS.
+rewrite List.flat_map_concat_map, List.map_map. simpl.
+rewrite concat_map_singleton.
+reflexivity.
+Qed.
+Lemma no_throw_basic_builtins_5 Regs E1 E2 :
+ ignore_throw (E1 := E1) (choose_boolS tt) = @choose_boolS Regs E2 tt.
+reflexivity. Qed.
+Lemma no_throw_basic_builtins_6 Regs A E1 E2 msg :
+ ignore_throw (E1 := E1) (failS msg) = @failS Regs A E2 msg.
+reflexivity. Qed.
+Lemma no_throw_basic_builtins_7 Regs A E1 E2 msg x :
+ ignore_throw (E1 := E1) (maybe_failS msg x) = @maybe_failS Regs A E2 msg x.
+destruct x; reflexivity. Qed.
+
+Hint Rewrite no_throw_basic_builtins_1 no_throw_basic_builtins_2
+ no_throw_basic_builtins_3 no_throw_basic_builtins_4
+ no_throw_basic_builtins_5 no_throw_basic_builtins_6
+ no_throw_basic_builtins_7 : ignore_throw.
+
+Lemma ignore_throw_option_case_distrib_1 Regs B C E1 E2 (c : sequential_state Regs -> option B) s (n : monadS Regs C E1) (f : B -> monadS Regs C E1) :
+ ignore_throw (E2 := E2) (match c s with None => n | Some b => f b end) s =
+ match c s with None => ignore_throw n s | Some b => ignore_throw (f b) s end.
+destruct (c s); auto.
+Qed.
+Lemma ignore_throw_option_case_distrib_2 Regs B C E1 E2 (c : option B) (n : monadS Regs C E1) (f : B -> monadS Regs C E1) :
+ ignore_throw (E2 := E2) (match c with None => n | Some b => f b end) =
+ match c with None => ignore_throw n | Some b => ignore_throw (f b) end.
+destruct c; auto.
+Qed.
+
+Lemma ignore_throw_let_distrib Regs A B E1 E2 (y : A) (f : A -> monadS Regs B E1) :
+ ignore_throw (E2 := E2) (let x := y in f x) = (let x := y in ignore_throw (f x)).
+reflexivity.
+Qed.
+
+Lemma no_throw_mem_builtins_1 Regs E1 E2 rk a sz :
+ ignore_throw (E2 := E2) (@read_memt_bytesS Regs E1 rk a sz) === read_memt_bytesS rk a sz.
+unfold read_memt_bytesS. autorewrite with ignore_throw.
+apply bindS_cong; auto. intros. autorewrite with ignore_throw. reflexivity.
+Qed.
+Hint Rewrite no_throw_mem_builtins_1 : ignore_throw.
+Lemma no_throw_mem_builtins_2 Regs E1 E2 rk a sz :
+ ignore_throw (E2 := E2) (@read_mem_bytesS Regs E1 rk a sz) === read_mem_bytesS rk a sz.
+unfold read_mem_bytesS. autorewrite with ignore_throw.
+apply bindS_cong; intros; autorewrite with ignore_throw; auto.
+destruct a0; reflexivity.
+Qed.
+Hint Rewrite no_throw_mem_builtins_2 : ignore_throw.
+Lemma no_throw_mem_builtins_3 Regs A E1 E2 a :
+ ignore_throw (E2 := E2) (@read_tagS Regs A E1 a) === read_tagS a.
+reflexivity. Qed.
+Hint Rewrite no_throw_mem_builtins_3 : ignore_throw.
+Lemma no_throw_mem_builtins_4 Regs A V E1 E2 rk a sz H :
+ ignore_throw (E2 := E2) (@read_memtS Regs E1 A V rk a sz H) === read_memtS rk a sz.
+unfold read_memtS. autorewrite with ignore_throw.
+apply bindS_cong; intros; autorewrite with ignore_throw.
+reflexivity. destruct a0; simpl. autorewrite with ignore_throw.
+reflexivity.
+Qed.
+Hint Rewrite no_throw_mem_builtins_4 : ignore_throw.
+Lemma no_throw_mem_builtins_5 Regs A V E1 E2 rk a sz H :
+ ignore_throw (E2 := E2) (@read_memS Regs E1 A V rk a sz H) === read_memS rk a sz.
+unfold read_memS. autorewrite with ignore_throw.
+apply bindS_cong; intros; autorewrite with ignore_throw; auto.
+destruct a0; auto.
+Qed.
+Hint Rewrite no_throw_mem_builtins_5 : ignore_throw.
+Lemma no_throw_mem_builtins_6 Regs E1 E2 wk addr sz v t :
+ ignore_throw (E2 := E2) (@write_memt_bytesS Regs E1 wk addr sz v t) === write_memt_bytesS wk addr sz v t.
+unfold write_memt_bytesS. unfold seqS. autorewrite with ignore_throw.
+reflexivity.
+Qed.
+Hint Rewrite no_throw_mem_builtins_6 : ignore_throw.
+Lemma no_throw_mem_builtins_7 Regs E1 E2 wk addr sz v :
+ ignore_throw (E2 := E2) (@write_mem_bytesS Regs E1 wk addr sz v) === write_mem_bytesS wk addr sz v.
+unfold write_mem_bytesS. autorewrite with ignore_throw. reflexivity.
+Qed.
+Hint Rewrite no_throw_mem_builtins_7 : ignore_throw.
+Lemma no_throw_mem_builtins_8 Regs E1 E2 A B wk addr sz v t :
+ ignore_throw (E2 := E2) (@write_memtS Regs E1 A B wk addr sz v t) === write_memtS wk addr sz v t.
+unfold write_memtS. rewrite ignore_throw_option_case_distrib_2.
+destruct (Sail2_values.mem_bytes_of_bits v); autorewrite with ignore_throw; auto.
+Qed.
+Hint Rewrite no_throw_mem_builtins_8 : ignore_throw.
+Lemma no_throw_mem_builtins_9 Regs E1 E2 A B wk addr sz v :
+ ignore_throw (E2 := E2) (@write_memS Regs E1 A B wk addr sz v) === write_memS wk addr sz v.
+unfold write_memS. autorewrite with ignore_throw; auto.
+Qed.
+Hint Rewrite no_throw_mem_builtins_9 : ignore_throw.
+Lemma no_throw_mem_builtins_10 Regs E1 E2 :
+ ignore_throw (E2 := E2) (@excl_resultS Regs E1 tt) === excl_resultS tt.
+reflexivity. Qed.
+Hint Rewrite no_throw_mem_builtins_10 : ignore_throw.
+Lemma no_throw_mem_builtins_11 Regs E1 E2 :
+ ignore_throw (E2 := E2) (@undefined_boolS Regs E1 tt) === undefined_boolS tt.
+reflexivity. Qed.
+Hint Rewrite no_throw_mem_builtins_11 : ignore_throw.
+
+Lemma no_throw_read_regvalS Regs RV E1 E2 r reg_name :
+ ignore_throw (E2 := E2) (@read_regvalS Regs RV E1 r reg_name) === read_regvalS r reg_name.
+destruct r; simpl. autorewrite with ignore_throw.
+apply bindS_cong; intros; auto. rewrite ignore_throw_option_case_distrib_2.
+autorewrite with ignore_throw. reflexivity.
+Qed.
+Hint Rewrite no_throw_read_regvalS : ignore_throw.
+
+Lemma no_throw_write_regvalS Regs RV E1 E2 r reg_name v :
+ ignore_throw (E2 := E2) (@write_regvalS Regs RV E1 r reg_name v) === write_regvalS r reg_name v.
+destruct r; simpl. autorewrite with ignore_throw.
+apply bindS_cong; intros; auto. rewrite ignore_throw_option_case_distrib_2.
+autorewrite with ignore_throw. reflexivity.
+Qed.
+Hint Rewrite no_throw_write_regvalS : ignore_throw.
diff --git a/prover_snapshots/coq/lib/sail/Sail2_string.v b/prover_snapshots/coq/lib/sail/Sail2_string.v
new file mode 100644
index 0000000..a0a2393
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_string.v
@@ -0,0 +1,194 @@
+Require Import Sail2_values.
+Require Import Coq.Strings.Ascii.
+
+Definition string_sub (s : string) (start : Z) (len : Z) : string :=
+ String.substring (Z.to_nat start) (Z.to_nat len) s.
+
+Definition string_startswith s expected :=
+ let prefix := String.substring 0 (String.length expected) s in
+ generic_eq prefix expected.
+
+Definition string_drop s (n : Z) `{ArithFact (n >= 0)} :=
+ let n := Z.to_nat n in
+ String.substring n (String.length s - n) s.
+
+Definition string_take s (n : Z) `{ArithFact (n >= 0)} :=
+ let n := Z.to_nat n in
+ String.substring 0 n s.
+
+Definition string_length s : {n : Z & ArithFact (n >= 0)} :=
+ build_ex (Z.of_nat (String.length s)).
+
+Definition string_append := String.append.
+
+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
+ | EmptyString => 0
+ | String " " t => S (n_leading_spaces t)
+ | _ => 0
+ end.
+
+Definition opt_spc_matches_prefix s : option (unit * {n : Z & ArithFact (n >= 0)}) :=
+ Some (tt, build_ex (Z.of_nat (n_leading_spaces s))).
+
+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.
+
+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_1_matches_prefix s := hex_bits_n_matches_prefix 1 s.
+Definition hex_bits_2_matches_prefix s := hex_bits_n_matches_prefix 2 s.
+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_17_matches_prefix s := hex_bits_n_matches_prefix 17 s.
+Definition hex_bits_18_matches_prefix s := hex_bits_n_matches_prefix 18 s.
+Definition hex_bits_19_matches_prefix s := hex_bits_n_matches_prefix 19 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_22_matches_prefix s := hex_bits_n_matches_prefix 22 s.
+Definition hex_bits_23_matches_prefix s := hex_bits_n_matches_prefix 23 s.
+Definition hex_bits_24_matches_prefix s := hex_bits_n_matches_prefix 24 s.
+Definition hex_bits_25_matches_prefix s := hex_bits_n_matches_prefix 25 s.
+Definition hex_bits_26_matches_prefix s := hex_bits_n_matches_prefix 26 s.
+Definition hex_bits_27_matches_prefix s := hex_bits_n_matches_prefix 27 s.
+Definition hex_bits_28_matches_prefix s := hex_bits_n_matches_prefix 28 s.
+Definition hex_bits_29_matches_prefix s := hex_bits_n_matches_prefix 29 s.
+Definition hex_bits_30_matches_prefix s := hex_bits_n_matches_prefix 30 s.
+Definition hex_bits_31_matches_prefix s := hex_bits_n_matches_prefix 31 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.
+Definition hex_bits_48_matches_prefix s := hex_bits_n_matches_prefix 48 s.
+Definition hex_bits_64_matches_prefix s := hex_bits_n_matches_prefix 64 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.
+
+Definition decimal_string_of_bv {a} `{Bitvector a} (bv : a) : string :=
+ match unsigned bv with
+ | None => "?"
+ | Some i => string_of_int i
+ end.
+
+Definition decimal_string_of_bits {n} (bv : mword n) : string := decimal_string_of_bv bv.
+
+
+(* Some aliases for compatibility. *)
+Definition dec_str := string_of_int.
+Definition concat_str := String.append.
diff --git a/prover_snapshots/coq/lib/sail/Sail2_values.v b/prover_snapshots/coq/lib/sail/Sail2_values.v
new file mode 100644
index 0000000..208f5c8
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/Sail2_values.v
@@ -0,0 +1,2490 @@
+(* Version of sail_values.lem that uses Lems machine words library *)
+
+(*Require Import Sail_impl_base*)
+Require Export ZArith.
+Require Import Ascii.
+Require Export String.
+Require Import bbv.Word.
+Require Export List.
+Require Export Sumbool.
+Require Export DecidableClass.
+Require Import Eqdep_dec.
+Require Export Zeuclid.
+Require Import Psatz.
+Import ListNotations.
+
+Open Scope Z.
+
+Module Z_eq_dec.
+Definition U := Z.
+Definition eq_dec := Z.eq_dec.
+End Z_eq_dec.
+Module ZEqdep := DecidableEqDep (Z_eq_dec).
+
+
+(* Constraint solving basics. A HintDb which unfolding hints and lemmata
+ can be added to, and a typeclass to wrap constraint arguments in to
+ trigger automatic solving. *)
+Create HintDb sail.
+Class ArithFact (P : Prop) := { fact : P }.
+Lemma use_ArithFact {P} `(ArithFact P) : P.
+apply fact.
+Defined.
+
+(* Allow setoid rewriting through ArithFact *)
+Require Import Coq.Classes.Morphisms.
+Require Import Coq.Program.Basics.
+Require Import Coq.Program.Tactics.
+Section Morphism.
+Local Obligation Tactic := try solve [simpl_relation | firstorder auto].
+
+Global Program Instance ArithFact_iff_morphism :
+ Proper (iff ==> iff) ArithFact.
+End Morphism.
+
+
+Definition build_ex {T:Type} (n:T) {P:T -> Prop} `{H:ArithFact (P n)} : {x : T & ArithFact (P x)} :=
+ existT _ n H.
+
+
+Definition generic_eq {T:Type} (x y:T) `{Decidable (x = y)} := Decidable_witness.
+Definition generic_neq {T:Type} (x y:T) `{Decidable (x = y)} := negb Decidable_witness.
+Lemma generic_eq_true {T} {x y:T} `{Decidable (x = y)} : generic_eq x y = true -> x = y.
+apply Decidable_spec.
+Qed.
+Lemma generic_eq_false {T} {x y:T} `{Decidable (x = y)} : generic_eq x y = false -> x <> y.
+unfold generic_eq.
+intros H1 H2.
+rewrite <- Decidable_spec in H2.
+congruence.
+Qed.
+Lemma generic_neq_true {T} {x y:T} `{Decidable (x = y)} : generic_neq x y = true -> x <> y.
+unfold generic_neq.
+intros H1 H2.
+rewrite <- Decidable_spec in H2.
+destruct Decidable_witness; simpl in *;
+congruence.
+Qed.
+Lemma generic_neq_false {T} {x y:T} `{Decidable (x = y)} : generic_neq x y = false -> x = y.
+unfold generic_neq.
+intro H1.
+rewrite <- Decidable_spec.
+destruct Decidable_witness; simpl in *;
+congruence.
+Qed.
+Instance Decidable_eq_from_dec {T:Type} (eqdec: forall x y : T, {x = y} + {x <> y}) :
+ forall (x y : T), Decidable (eq x y) := {
+ Decidable_witness := proj1_sig (bool_of_sumbool (eqdec x y))
+}.
+destruct (eqdec x y); simpl; split; congruence.
+Defined.
+
+Instance Decidable_eq_string : forall (x y : string), Decidable (x = y) :=
+ Decidable_eq_from_dec String.string_dec.
+
+Instance Decidable_eq_pair {A B : Type} `(DA : forall x y : A, Decidable (x = y), DB : forall x y : B, Decidable (x = y)) : forall x y : A*B, Decidable (x = y) :=
+{ Decidable_witness := andb (@Decidable_witness _ (DA (fst x) (fst y)))
+(@Decidable_witness _ (DB (snd x) (snd y))) }.
+destruct x as [x1 x2].
+destruct y as [y1 y2].
+simpl.
+destruct (DA x1 y1) as [b1 H1];
+destruct (DB x2 y2) as [b2 H2];
+simpl.
+split.
+* intro H.
+ apply Bool.andb_true_iff in H.
+ destruct H as [H1b H2b].
+ apply H1 in H1b.
+ apply H2 in H2b.
+ congruence.
+* intro. inversion H.
+ subst.
+ apply Bool.andb_true_iff.
+ tauto.
+Qed.
+
+Definition generic_dec {T:Type} (x y:T) `{Decidable (x = y)} : {x = y} + {x <> y}.
+refine ((if Decidable_witness as b return (b = true <-> x = y -> _) then fun H' => _ else fun H' => _) Decidable_spec).
+* left. tauto.
+* right. intuition.
+Defined.
+
+Instance Decidable_eq_list {A : Type} `(D : forall x y : A, Decidable (x = y)) : forall (x y : list A), Decidable (x = y) :=
+ Decidable_eq_from_dec (list_eq_dec (fun x y => generic_dec x y)).
+
+(* Used by generated code that builds Decidable equality instances for records. *)
+Ltac cmp_record_field x y :=
+ let H := fresh "H" in
+ case (generic_dec x y);
+ intro H; [ |
+ refine (Build_Decidable _ false _);
+ split; [congruence | intros Z; destruct H; injection Z; auto]
+ ].
+
+
+
+(* Project away range constraints in comparisons *)
+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.
+
+(*val pow : Z -> Z -> Z*)
+Definition pow m n := m ^ n.
+
+Program Definition pow2 n : {z : Z & ArithFact (2 ^ n <= z <= 2 ^ n)} := existT _ (pow 2 n) _.
+Next Obligation.
+constructor.
+unfold pow.
+auto using Z.le_refl.
+Qed.
+
+Lemma ZEuclid_div_pos : forall x y, y > 0 -> x >= 0 -> ZEuclid.div x y >= 0.
+intros.
+unfold ZEuclid.div.
+change 0 with (0 * 0).
+apply Zmult_ge_compat; auto with zarith.
+* apply Z.le_ge. apply Z.sgn_nonneg. apply Z.ge_le. auto with zarith.
+* apply Z_div_ge0; auto. apply Z.lt_gt. apply Z.abs_pos. auto with zarith.
+Qed.
+
+Lemma ZEuclid_pos_div : forall x y, y > 0 -> ZEuclid.div x y >= 0 -> x >= 0.
+intros x y GT.
+ specialize (ZEuclid.div_mod x y);
+ specialize (ZEuclid.mod_always_pos x y);
+ generalize (ZEuclid.modulo x y);
+ generalize (ZEuclid.div x y);
+ intros.
+nia.
+Qed.
+
+Lemma ZEuclid_div_ge : forall x y, y > 0 -> x >= 0 -> x - ZEuclid.div x y >= 0.
+intros.
+unfold ZEuclid.div.
+rewrite Z.sgn_pos; auto with zarith.
+rewrite Z.mul_1_l.
+apply Z.le_ge.
+apply Zle_minus_le_0.
+apply Z.div_le_upper_bound.
+* apply Z.abs_pos. auto with zarith.
+* rewrite Z.mul_comm.
+ assert (0 < Z.abs y). {
+ apply Z.abs_pos.
+ omega.
+ }
+ revert H1.
+ generalize (Z.abs y). intros. nia.
+Qed.
+
+Lemma ZEuclid_div_mod0 : forall x y, y <> 0 ->
+ ZEuclid.modulo x y = 0 ->
+ y * ZEuclid.div x y = x.
+intros x y H1 H2.
+rewrite Zplus_0_r_reverse at 1.
+rewrite <- H2.
+symmetry.
+apply ZEuclid.div_mod.
+assumption.
+Qed.
+
+Hint Resolve ZEuclid_div_pos ZEuclid_pos_div ZEuclid_div_ge ZEuclid_div_mod0 : sail.
+
+
+(*
+Definition inline lt := (<)
+Definition inline gt := (>)
+Definition inline lteq := (<=)
+Definition inline gteq := (>=)
+
+val eq : forall a. Eq a => a -> a -> bool
+Definition inline eq l r := (l = r)
+
+val neq : forall a. Eq a => a -> a -> bool*)
+Definition neq l r := (negb (l =? r)). (* Z only *)
+
+(*let add_int l r := integerAdd l r
+Definition add_signed l r := integerAdd l r
+Definition sub_int l r := integerMinus l r
+Definition mult_int l r := integerMult l r
+Definition div_int l r := integerDiv l r
+Definition div_nat l r := natDiv l r
+Definition power_int_nat l r := integerPow l r
+Definition power_int_int l r := integerPow l (Z.to_nat r)
+Definition negate_int i := integerNegate i
+Definition min_int l r := integerMin l r
+Definition max_int l r := integerMax l r
+
+Definition add_real l r := realAdd l r
+Definition sub_real l r := realMinus l r
+Definition mult_real l r := realMult l r
+Definition div_real l r := realDiv l r
+Definition negate_real r := realNegate r
+Definition abs_real r := realAbs r
+Definition power_real b e := realPowInteger b e*)
+
+Definition print_endline (_ : string) : unit := tt.
+Definition prerr_endline (_ : string) : unit := tt.
+Definition prerr (_ : string) : unit := tt.
+Definition print_int (_ : string) (_ : Z) : unit := tt.
+Definition prerr_int (_ : string) (_ : Z) : unit := tt.
+Definition putchar (_ : Z) : unit := tt.
+
+Definition shl_int := Z.shiftl.
+Definition shr_int := Z.shiftr.
+
+(*
+Definition or_bool l r := (l || r)
+Definition and_bool l r := (l && r)
+Definition xor_bool l r := xor l r
+*)
+Definition append_list {A:Type} (l : list A) r := l ++ r.
+Definition length_list {A:Type} (xs : list A) := Z.of_nat (List.length xs).
+Definition take_list {A:Type} n (xs : list A) := firstn (Z.to_nat n) xs.
+Definition drop_list {A:Type} n (xs : list A) := skipn (Z.to_nat n) xs.
+(*
+val repeat : forall a. list a -> Z -> list a*)
+Fixpoint repeat' {a} (xs : list a) n :=
+ match n with
+ | O => []
+ | S n => xs ++ repeat' xs n
+ end.
+Lemma repeat'_length {a} {xs : list a} {n : nat} : List.length (repeat' xs n) = (n * List.length xs)%nat.
+induction n.
+* reflexivity.
+* simpl.
+ rewrite app_length.
+ auto with arith.
+Qed.
+Definition repeat {a} (xs : list a) (n : Z) :=
+ if n <=? 0 then []
+ else repeat' xs (Z.to_nat n).
+Lemma repeat_length {a} {xs : list a} {n : Z} (H : n >= 0) : length_list (repeat xs n) = n * length_list xs.
+unfold length_list, repeat.
+destruct n.
++ reflexivity.
++ simpl (List.length _).
+ rewrite repeat'_length.
+ rewrite Nat2Z.inj_mul.
+ rewrite positive_nat_Z.
+ reflexivity.
++ exfalso.
+ auto with zarith.
+Qed.
+
+(*declare {isabelle} termination_argument repeat = automatic
+
+Definition duplicate_to_list bit length := repeat [bit] length
+
+Fixpoint replace bs (n : Z) b' := match bs with
+ | [] => []
+ | b :: bs =>
+ if n = 0 then b' :: bs
+ else b :: replace bs (n - 1) b'
+ end
+declare {isabelle} termination_argument replace = automatic
+
+Definition upper n := n
+
+(* Modulus operation corresponding to quot below -- result
+ has sign of dividend. *)
+Definition hardware_mod (a: Z) (b:Z) : Z :=
+ let m := (abs a) mod (abs b) in
+ if a < 0 then ~m else m
+
+(* There are different possible answers for integer divide regarding
+rounding behaviour on negative operands. Positive operands always
+round down so derive the one we want (trucation towards zero) from
+that *)
+Definition hardware_quot (a:Z) (b:Z) : Z :=
+ let q := (abs a) / (abs b) in
+ if ((a<0) = (b<0)) then
+ q (* same sign -- result positive *)
+ else
+ ~q (* different sign -- result negative *)
+
+Definition max_64u := (integerPow 2 64) - 1
+Definition max_64 := (integerPow 2 63) - 1
+Definition min_64 := 0 - (integerPow 2 63)
+Definition max_32u := (4294967295 : Z)
+Definition max_32 := (2147483647 : Z)
+Definition min_32 := (0 - 2147483648 : Z)
+Definition max_8 := (127 : Z)
+Definition min_8 := (0 - 128 : Z)
+Definition max_5 := (31 : Z)
+Definition min_5 := (0 - 32 : Z)
+*)
+
+(* just_list takes a list of maybes and returns Some xs if all elements have
+ a value, and None if one of the elements is None. *)
+(*val just_list : forall a. list (option a) -> option (list a)*)
+Fixpoint just_list {A} (l : list (option A)) := match l with
+ | [] => Some []
+ | (x :: xs) =>
+ match (x, just_list xs) with
+ | (Some x, Some xs) => Some (x :: xs)
+ | (_, _) => None
+ end
+ end.
+(*declare {isabelle} termination_argument just_list = automatic
+
+lemma just_list_spec:
+ ((forall xs. (just_list xs = None) <-> List.elem None xs) &&
+ (forall xs es. (just_list xs = Some es) <-> (xs = List.map Some es)))*)
+
+Lemma just_list_length {A} : forall (l : list (option A)) (l' : list A),
+ Some l' = just_list l -> List.length l = List.length l'.
+induction l.
+* intros.
+ simpl in H.
+ inversion H.
+ reflexivity.
+* intros.
+ destruct a; simplify_eq H.
+ simpl in *.
+ destruct (just_list l); simplify_eq H.
+ intros.
+ subst.
+ simpl.
+ f_equal.
+ apply IHl.
+ reflexivity.
+Qed.
+
+Lemma just_list_length_Z {A} : forall (l : list (option A)) l', Some l' = just_list l -> length_list l = length_list l'.
+unfold length_list.
+intros.
+f_equal.
+auto using just_list_length.
+Qed.
+
+Fixpoint member_Z_list (x : Z) (l : list Z) : bool :=
+match l with
+| [] => false
+| h::t => if x =? h then true else member_Z_list x t
+end.
+
+Lemma member_Z_list_In {x l} : member_Z_list x l = true <-> In x l.
+induction l.
+* simpl. split. congruence. tauto.
+* simpl. destruct (x =? a) eqn:H.
+ + rewrite Z.eqb_eq in H. subst. tauto.
+ + rewrite Z.eqb_neq in H. split.
+ - intro Heq. right. apply IHl. assumption.
+ - intros [bad | good]. congruence. apply IHl. assumption.
+Qed.
+
+(*** Bits *)
+Inductive bitU := B0 | B1 | BU.
+
+Scheme Equality for bitU.
+Definition eq_bit := bitU_beq.
+Instance Decidable_eq_bit : forall (x y : bitU), Decidable (x = y) :=
+ Decidable_eq_from_dec bitU_eq_dec.
+
+Definition showBitU b :=
+match b with
+ | B0 => "O"
+ | B1 => "I"
+ | BU => "U"
+end%string.
+
+Definition bitU_char b :=
+match b with
+| B0 => "0"
+| B1 => "1"
+| BU => "?"
+end%char.
+
+(*instance (Show bitU)
+ let show := showBitU
+end*)
+
+Class BitU (a : Type) : Type := {
+ to_bitU : a -> bitU;
+ of_bitU : bitU -> a
+}.
+
+Instance bitU_BitU : (BitU bitU) := {
+ to_bitU b := b;
+ of_bitU b := b
+}.
+
+Definition bool_of_bitU bu := match bu with
+ | B0 => Some false
+ | B1 => Some true
+ | BU => None
+ end.
+
+Definition bitU_of_bool (b : bool) := if b then B1 else B0.
+
+(*Instance bool_BitU : (BitU bool) := {
+ to_bitU := bitU_of_bool;
+ of_bitU := bool_of_bitU
+}.*)
+
+Definition cast_bit_bool := bool_of_bitU.
+(*
+Definition bit_lifted_of_bitU bu := match bu with
+ | B0 => Bitl_zero
+ | B1 => Bitl_one
+ | BU => Bitl_undef
+ end.
+
+Definition bitU_of_bit := function
+ | Bitc_zero => B0
+ | Bitc_one => B1
+ end.
+
+Definition bit_of_bitU := function
+ | B0 => Bitc_zero
+ | B1 => Bitc_one
+ | BU => failwith "bit_of_bitU: BU"
+ end.
+
+Definition bitU_of_bit_lifted := function
+ | Bitl_zero => B0
+ | Bitl_one => B1
+ | Bitl_undef => BU
+ | Bitl_unknown => failwith "bitU_of_bit_lifted Bitl_unknown"
+ end.
+*)
+Definition not_bit b :=
+match b with
+ | B1 => B0
+ | B0 => B1
+ | BU => BU
+ end.
+
+(*val is_one : Z -> bitU*)
+Definition is_one (i : Z) :=
+ if i =? 1 then B1 else B0.
+
+Definition binop_bit op x y :=
+ match (x, y) with
+ | (BU,_) => BU (*Do we want to do this or to respect | of I and & of B0 rules?*)
+ | (_,BU) => BU (*Do we want to do this or to respect | of I and & of B0 rules?*)
+(* | (x,y) => bitU_of_bool (op (bool_of_bitU x) (bool_of_bitU y))*)
+ | (B0,B0) => bitU_of_bool (op false false)
+ | (B0,B1) => bitU_of_bool (op false true)
+ | (B1,B0) => bitU_of_bool (op true false)
+ | (B1,B1) => bitU_of_bool (op true true)
+ end.
+
+(*val and_bit : bitU -> bitU -> bitU*)
+Definition and_bit := binop_bit andb.
+
+(*val or_bit : bitU -> bitU -> bitU*)
+Definition or_bit := binop_bit orb.
+
+(*val xor_bit : bitU -> bitU -> bitU*)
+Definition xor_bit := binop_bit xorb.
+
+(*val (&.) : bitU -> bitU -> bitU
+Definition inline (&.) x y := and_bit x y
+
+val (|.) : bitU -> bitU -> bitU
+Definition inline (|.) x y := or_bit x y
+
+val (+.) : bitU -> bitU -> bitU
+Definition inline (+.) x y := xor_bit x y
+*)
+
+(*** Bool lists ***)
+
+(*val bools_of_nat_aux : integer -> natural -> list bool -> list bool*)
+Fixpoint bools_of_nat_aux len (x : nat) (acc : list bool) : list bool :=
+ match len with
+ | O => acc
+ | S len' => bools_of_nat_aux len' (x / 2) ((if x mod 2 =? 1 then true else false) :: acc)
+ end %nat.
+ (*else (if x mod 2 = 1 then true else false) :: bools_of_nat_aux (x / 2)*)
+(*declare {isabelle} termination_argument bools_of_nat_aux = automatic*)
+Definition bools_of_nat len n := bools_of_nat_aux (Z.to_nat len) n [] (*List.reverse (bools_of_nat_aux n)*).
+
+(*val nat_of_bools_aux : natural -> list bool -> natural*)
+Fixpoint nat_of_bools_aux (acc : nat) (bs : list bool) : nat :=
+ match bs with
+ | [] => acc
+ | true :: bs => nat_of_bools_aux ((2 * acc) + 1) bs
+ | false :: bs => nat_of_bools_aux (2 * acc) bs
+end.
+(*declare {isabelle; hol} termination_argument nat_of_bools_aux = automatic*)
+Definition nat_of_bools bs := nat_of_bools_aux 0 bs.
+
+(*val unsigned_of_bools : list bool -> integer*)
+Definition unsigned_of_bools bs := Z.of_nat (nat_of_bools bs).
+
+(*val signed_of_bools : list bool -> integer*)
+Definition signed_of_bools bs :=
+ match bs with
+ | true :: _ => 0 - (1 + (unsigned_of_bools (List.map negb bs)))
+ | false :: _ => unsigned_of_bools bs
+ | [] => 0 (* Treat empty list as all zeros *)
+ end.
+
+(*val int_of_bools : bool -> list bool -> integer*)
+Definition int_of_bools (sign : bool) bs := if sign then signed_of_bools bs else unsigned_of_bools bs.
+
+(*val pad_list : forall 'a. 'a -> list 'a -> integer -> list 'a*)
+Fixpoint pad_list_nat {a} (x : a) (xs : list a) n :=
+ match n with
+ | O => xs
+ | S n' => pad_list_nat x (x :: xs) n'
+ end.
+(*declare {isabelle} termination_argument pad_list = automatic*)
+Definition pad_list {a} x xs n := @pad_list_nat a x xs (Z.to_nat n).
+
+Definition ext_list {a} pad len (xs : list a) :=
+ let longer := len - (Z.of_nat (List.length xs)) in
+ if longer <? 0 then skipn (Z.abs_nat (longer)) xs
+ else pad_list pad xs longer.
+
+(*let extz_bools len bs = ext_list false len bs*)
+Definition exts_bools len bs :=
+ match bs with
+ | true :: _ => ext_list true len bs
+ | _ => ext_list false len bs
+ end.
+
+Fixpoint add_one_bool_ignore_overflow_aux bits := match bits with
+ | [] => []
+ | false :: bits => true :: bits
+ | true :: bits => false :: add_one_bool_ignore_overflow_aux bits
+end.
+(*declare {isabelle; hol} termination_argument add_one_bool_ignore_overflow_aux = automatic*)
+
+Definition add_one_bool_ignore_overflow bits :=
+ List.rev (add_one_bool_ignore_overflow_aux (List.rev bits)).
+
+(* Ported from Lem, bad for large n.
+Definition bools_of_int len n :=
+ let bs_abs := bools_of_nat len (Z.abs_nat n) in
+ if n >=? 0 then bs_abs
+ else add_one_bool_ignore_overflow (List.map negb bs_abs).
+*)
+Fixpoint bitlistFromWord_rev {n} w :=
+match w with
+| WO => []
+| WS b w => b :: bitlistFromWord_rev w
+end.
+Definition bitlistFromWord {n} w :=
+ List.rev (@bitlistFromWord_rev n w).
+
+Definition bools_of_int len n :=
+ let w := Word.ZToWord (Z.to_nat len) n in
+ bitlistFromWord w.
+
+(*** Bit lists ***)
+
+(*val bits_of_nat_aux : natural -> list bitU*)
+Fixpoint bits_of_nat_aux n x :=
+ match n,x with
+ | O,_ => []
+ | _,O => []
+ | S n, S _ => (if x mod 2 =? 1 then B1 else B0) :: bits_of_nat_aux n (x / 2)
+ end%nat.
+(**declare {isabelle} termination_argument bits_of_nat_aux = automatic*)
+Definition bits_of_nat n := List.rev (bits_of_nat_aux n n).
+
+(*val nat_of_bits_aux : natural -> list bitU -> natural*)
+Fixpoint nat_of_bits_aux acc bs := match bs with
+ | [] => Some acc
+ | B1 :: bs => nat_of_bits_aux ((2 * acc) + 1) bs
+ | B0 :: bs => nat_of_bits_aux (2 * acc) bs
+ | BU :: bs => None
+end%nat.
+(*declare {isabelle} termination_argument nat_of_bits_aux = automatic*)
+Definition nat_of_bits bits := nat_of_bits_aux 0 bits.
+
+Definition not_bits := List.map not_bit.
+
+Definition binop_bits op bsl bsr :=
+ List.fold_right (fun '(bl, br) acc => binop_bit op bl br :: acc) [] (List.combine bsl bsr).
+(*
+Definition and_bits := binop_bits (&&)
+Definition or_bits := binop_bits (||)
+Definition xor_bits := binop_bits xor
+
+val unsigned_of_bits : list bitU -> Z*)
+Definition unsigned_of_bits bits :=
+match just_list (List.map bool_of_bitU bits) with
+| Some bs => Some (unsigned_of_bools bs)
+| None => None
+end.
+
+(*val signed_of_bits : list bitU -> Z*)
+Definition signed_of_bits bits :=
+ match just_list (List.map bool_of_bitU bits) with
+ | Some bs => Some (signed_of_bools bs)
+ | None => None
+ end.
+
+(*val int_of_bits : bool -> list bitU -> maybe integer*)
+Definition int_of_bits (sign : bool) bs :=
+ if sign then signed_of_bits bs else unsigned_of_bits bs.
+
+(*val pad_bitlist : bitU -> list bitU -> Z -> list bitU*)
+Fixpoint pad_bitlist_nat (b : bitU) bits n :=
+match n with
+| O => bits
+| S n' => pad_bitlist_nat b (b :: bits) n'
+end.
+Definition pad_bitlist b bits n := pad_bitlist_nat b bits (Z.to_nat n). (* Negative n will come out as 0 *)
+(* if n <= 0 then bits else pad_bitlist b (b :: bits) (n - 1).
+declare {isabelle} termination_argument pad_bitlist = automatic*)
+
+Definition ext_bits pad len bits :=
+ let longer := len - (Z.of_nat (List.length bits)) in
+ if longer <? 0 then skipn (Z.abs_nat longer) bits
+ else pad_bitlist pad bits longer.
+
+Definition extz_bits len bits := ext_bits B0 len bits.
+Parameter undefined_list_bitU : list bitU.
+Definition exts_bits len bits :=
+ match bits with
+ | BU :: _ => undefined_list_bitU (*failwith "exts_bits: undefined bit"*)
+ | B1 :: _ => ext_bits B1 len bits
+ | _ => ext_bits B0 len bits
+ end.
+
+Fixpoint add_one_bit_ignore_overflow_aux bits := match bits with
+ | [] => []
+ | B0 :: bits => B1 :: bits
+ | B1 :: bits => B0 :: add_one_bit_ignore_overflow_aux bits
+ | BU :: _ => undefined_list_bitU (*failwith "add_one_bit_ignore_overflow: undefined bit"*)
+end.
+(*declare {isabelle} termination_argument add_one_bit_ignore_overflow_aux = automatic*)
+
+Definition add_one_bit_ignore_overflow bits :=
+ rev (add_one_bit_ignore_overflow_aux (rev bits)).
+
+Definition bitlist_of_int n :=
+ let bits_abs := B0 :: bits_of_nat (Z.abs_nat n) in
+ if n >=? 0 then bits_abs
+ else add_one_bit_ignore_overflow (not_bits bits_abs).
+
+Definition bits_of_int len n := exts_bits len (bitlist_of_int n).
+
+(*val arith_op_bits :
+ (integer -> integer -> integer) -> bool -> list bitU -> list bitU -> list bitU*)
+Definition arith_op_bits (op : Z -> Z -> Z) (sign : bool) l r :=
+ match (int_of_bits sign l, int_of_bits sign r) with
+ | (Some li, Some ri) => bits_of_int (length_list l) (op li ri)
+ | (_, _) => repeat [BU] (length_list l)
+ end.
+
+
+Definition char_of_nibble x :=
+ match x with
+ | (B0, B0, B0, B0) => Some "0"%char
+ | (B0, B0, B0, B1) => Some "1"%char
+ | (B0, B0, B1, B0) => Some "2"%char
+ | (B0, B0, B1, B1) => Some "3"%char
+ | (B0, B1, B0, B0) => Some "4"%char
+ | (B0, B1, B0, B1) => Some "5"%char
+ | (B0, B1, B1, B0) => Some "6"%char
+ | (B0, B1, B1, B1) => Some "7"%char
+ | (B1, B0, B0, B0) => Some "8"%char
+ | (B1, B0, B0, B1) => Some "9"%char
+ | (B1, B0, B1, B0) => Some "A"%char
+ | (B1, B0, B1, B1) => Some "B"%char
+ | (B1, B1, B0, B0) => Some "C"%char
+ | (B1, B1, B0, B1) => Some "D"%char
+ | (B1, B1, B1, B0) => Some "E"%char
+ | (B1, B1, B1, B1) => Some "F"%char
+ | _ => None
+ end.
+
+Fixpoint hexstring_of_bits bs := match bs with
+ | b1 :: b2 :: b3 :: b4 :: bs =>
+ let n := char_of_nibble (b1, b2, b3, b4) in
+ let s := hexstring_of_bits bs in
+ match (n, s) with
+ | (Some n, Some s) => Some (String n s)
+ | _ => None
+ end
+ | [] => Some EmptyString
+ | _ => None
+ end%string.
+
+Fixpoint binstring_of_bits bs := match bs with
+ | b :: bs => String (bitU_char b) (binstring_of_bits bs)
+ | [] => EmptyString
+ end.
+
+Definition show_bitlist bs :=
+ match hexstring_of_bits bs with
+ | Some s => String "0" (String "x" s)
+ | None => String "0" (String "b" (binstring_of_bits bs))
+ end.
+
+(*** List operations *)
+(*
+Definition inline (^^) := append_list
+
+val subrange_list_inc : forall a. list a -> Z -> Z -> list a*)
+Definition subrange_list_inc {A} (xs : list A) i j :=
+ let toJ := firstn (Z.to_nat j + 1) xs in
+ let fromItoJ := skipn (Z.to_nat i) toJ in
+ fromItoJ.
+
+(*val subrange_list_dec : forall a. list a -> Z -> Z -> list a*)
+Definition subrange_list_dec {A} (xs : list A) i j :=
+ let top := (length_list xs) - 1 in
+ subrange_list_inc xs (top - i) (top - j).
+
+(*val subrange_list : forall a. bool -> list a -> Z -> Z -> list a*)
+Definition subrange_list {A} (is_inc : bool) (xs : list A) i j :=
+ if is_inc then subrange_list_inc xs i j else subrange_list_dec xs i j.
+
+Definition splitAt {A} n (l : list A) := (firstn n l, skipn n l).
+
+(*val update_subrange_list_inc : forall a. list a -> Z -> Z -> list a -> list a*)
+Definition update_subrange_list_inc {A} (xs : list A) i j xs' :=
+ let (toJ,suffix) := splitAt (Z.to_nat j + 1) xs in
+ let (prefix,_fromItoJ) := splitAt (Z.to_nat i) toJ in
+ prefix ++ xs' ++ suffix.
+
+(*val update_subrange_list_dec : forall a. list a -> Z -> Z -> list a -> list a*)
+Definition update_subrange_list_dec {A} (xs : list A) i j xs' :=
+ let top := (length_list xs) - 1 in
+ update_subrange_list_inc xs (top - i) (top - j) xs'.
+
+(*val update_subrange_list : forall a. bool -> list a -> Z -> Z -> list a -> list a*)
+Definition update_subrange_list {A} (is_inc : bool) (xs : list A) i j xs' :=
+ if is_inc then update_subrange_list_inc xs i j xs' else update_subrange_list_dec xs i j xs'.
+
+Open Scope nat.
+Fixpoint nth_in_range {A} (n:nat) (l:list A) : n < length l -> A.
+refine
+ (match n, l with
+ | O, h::_ => fun _ => h
+ | S m, _::t => fun H => nth_in_range A m t _
+ | _,_ => fun H => _
+ end).
+exfalso. inversion H.
+exfalso. inversion H.
+simpl in H. omega.
+Defined.
+
+Lemma nth_in_range_is_nth : forall A n (l : list A) d (H : n < length l),
+ nth_in_range n l H = nth n l d.
+intros until d. revert n.
+induction l; intros n H.
+* inversion H.
+* destruct n.
+ + reflexivity.
+ + apply IHl.
+Qed.
+
+Lemma nth_Z_nat {A} {n} {xs : list A} :
+ (0 <= n)%Z -> (n < length_list xs)%Z -> Z.to_nat n < length xs.
+unfold length_list.
+intros nonneg bounded.
+rewrite Z2Nat.inj_lt in bounded; auto using Zle_0_nat.
+rewrite Nat2Z.id in bounded.
+assumption.
+Qed.
+
+(*
+Lemma nth_top_aux {A} {n} {xs : list A} : Z.to_nat n < length xs -> let top := ((length_list xs) - 1)%Z in Z.to_nat (top - n)%Z < length xs.
+unfold length_list.
+generalize (length xs).
+intro n0.
+rewrite <- (Nat2Z.id n0).
+intro H.
+apply Z2Nat.inj_lt.
+* omega.
+*)
+
+Close Scope nat.
+
+(*val access_list_inc : forall a. list a -> Z -> a*)
+Definition access_list_inc {A} (xs : list A) n `{ArithFact (0 <= n)} `{ArithFact (n < length_list xs)} := nth_in_range (Z.to_nat n) xs (nth_Z_nat (use_ArithFact _) (use_ArithFact _)).
+
+(*val access_list_dec : forall a. list a -> Z -> a*)
+Definition access_list_dec {A} (xs : list A) n `{ArithFact (0 <= n)} `{ArithFact (n < length_list xs)} : A.
+refine (
+ let top := (length_list xs) - 1 in
+ @access_list_inc A xs (top - n) _ _).
+constructor. apply use_ArithFact in H. apply use_ArithFact in H0. omega.
+constructor. apply use_ArithFact in H. apply use_ArithFact in H0. omega.
+Defined.
+
+(*val access_list : forall a. bool -> list a -> Z -> a*)
+Definition access_list {A} (is_inc : bool) (xs : list A) n `{ArithFact (0 <= n)} `{ArithFact (n < length_list xs)} :=
+ if is_inc then access_list_inc xs n else access_list_dec xs n.
+
+Definition access_list_opt_inc {A} (xs : list A) n := nth_error xs (Z.to_nat n).
+
+(*val access_list_dec : forall a. list a -> Z -> a*)
+Definition access_list_opt_dec {A} (xs : list A) n :=
+ let top := (length_list xs) - 1 in
+ access_list_opt_inc xs (top - n).
+
+(*val access_list : forall a. bool -> list a -> Z -> a*)
+Definition access_list_opt {A} (is_inc : bool) (xs : list A) n :=
+ if is_inc then access_list_opt_inc xs n else access_list_opt_dec xs n.
+
+Definition list_update {A} (xs : list A) n x := firstn n xs ++ x :: skipn (S n) xs.
+
+(*val update_list_inc : forall a. list a -> Z -> a -> list a*)
+Definition update_list_inc {A} (xs : list A) n x := list_update xs (Z.to_nat n) x.
+
+(*val update_list_dec : forall a. list a -> Z -> a -> list a*)
+Definition update_list_dec {A} (xs : list A) n x :=
+ let top := (length_list xs) - 1 in
+ update_list_inc xs (top - n) x.
+
+(*val update_list : forall a. bool -> list a -> Z -> a -> list a*)
+Definition update_list {A} (is_inc : bool) (xs : list A) n x :=
+ if is_inc then update_list_inc xs n x else update_list_dec xs n x.
+
+(*Definition extract_only_element := function
+ | [] => failwith "extract_only_element called for empty list"
+ | [e] => e
+ | _ => failwith "extract_only_element called for list with more elements"
+end*)
+
+(*** Machine words *)
+
+Definition mword (n : Z) :=
+ match n with
+ | Zneg _ => False
+ | Z0 => word 0
+ | Zpos p => word (Pos.to_nat p)
+ end.
+
+Definition get_word {n} : mword n -> word (Z.to_nat n) :=
+ match n with
+ | Zneg _ => fun x => match x with end
+ | Z0 => fun x => x
+ | Zpos p => fun x => x
+ end.
+
+Definition with_word {n} {P : Type -> Type} : (word (Z.to_nat n) -> P (word (Z.to_nat n))) -> mword n -> P (mword n) :=
+match n with
+| Zneg _ => fun f w => match w with end
+| Z0 => fun f w => f w
+| Zpos _ => fun f w => f w
+end.
+
+Program Definition to_word {n} : n >= 0 -> word (Z.to_nat n) -> mword n :=
+ match n with
+ | Zneg _ => fun H _ => _
+ | Z0 => fun _ w => w
+ | Zpos _ => fun _ w => w
+ end.
+
+Definition word_to_mword {n} (w : word (Z.to_nat n)) `{H:ArithFact (n >= 0)} : mword n :=
+ to_word (match H with Build_ArithFact _ H' => H' end) w.
+
+(*val length_mword : forall a. mword a -> Z*)
+Definition length_mword {n} (w : mword n) := n.
+
+(*val slice_mword_dec : forall a b. mword a -> Z -> Z -> mword b*)
+(*Definition slice_mword_dec w i j := word_extract (Z.to_nat i) (Z.to_nat j) w.
+
+val slice_mword_inc : forall a b. mword a -> Z -> Z -> mword b
+Definition slice_mword_inc w i j :=
+ let top := (length_mword w) - 1 in
+ slice_mword_dec w (top - i) (top - j)
+
+val slice_mword : forall a b. bool -> mword a -> Z -> Z -> mword b
+Definition slice_mword is_inc w i j := if is_inc then slice_mword_inc w i j else slice_mword_dec w i j
+
+val update_slice_mword_dec : forall a b. mword a -> Z -> Z -> mword b -> mword a
+Definition update_slice_mword_dec w i j w' := word_update w (Z.to_nat i) (Z.to_nat j) w'
+
+val update_slice_mword_inc : forall a b. mword a -> Z -> Z -> mword b -> mword a
+Definition update_slice_mword_inc w i j w' :=
+ let top := (length_mword w) - 1 in
+ update_slice_mword_dec w (top - i) (top - j) w'
+
+val update_slice_mword : forall a b. bool -> mword a -> Z -> Z -> mword b -> mword a
+Definition update_slice_mword is_inc w i j w' :=
+ if is_inc then update_slice_mword_inc w i j w' else update_slice_mword_dec w i j w'
+
+val access_mword_dec : forall a. mword a -> Z -> bitU*)
+Parameter undefined_bit : bool.
+Definition getBit {n} :=
+match n with
+| O => fun (w : word O) i => undefined_bit
+| S n => fun (w : word (S n)) i => wlsb (wrshift' w i)
+end.
+
+Definition access_mword_dec {m} (w : mword m) n := bitU_of_bool (getBit (get_word w) (Z.to_nat n)).
+
+(*val access_mword_inc : forall a. mword a -> Z -> bitU*)
+Definition access_mword_inc {m} (w : mword m) n :=
+ let top := (length_mword w) - 1 in
+ access_mword_dec w (top - n).
+
+(*Parameter access_mword : forall {a}, bool -> mword a -> Z -> bitU.*)
+Definition access_mword {a} (is_inc : bool) (w : mword a) n :=
+ if is_inc then access_mword_inc w n else access_mword_dec w n.
+
+Definition setBit {n} :=
+match n with
+| O => fun (w : word O) i b => w
+| S n => fun (w : word (S n)) i (b : bool) =>
+ let bit : word (S n) := wlshift' (natToWord _ 1) i in
+ let mask : word (S n) := wnot bit in
+ let masked := wand mask w in
+ if b then masked else wor masked bit
+end.
+
+(*val update_mword_bool_dec : forall 'a. mword 'a -> integer -> bool -> mword 'a*)
+Definition update_mword_bool_dec {a} (w : mword a) n b : mword a :=
+ with_word (P := id) (fun w => setBit w (Z.to_nat n) b) w.
+Definition update_mword_dec {a} (w : mword a) n b :=
+ match bool_of_bitU b with
+ | Some bl => Some (update_mword_bool_dec w n bl)
+ | None => None
+ end.
+
+(*val update_mword_inc : forall a. mword a -> Z -> bitU -> mword a*)
+Definition update_mword_inc {a} (w : mword a) n b :=
+ let top := (length_mword w) - 1 in
+ update_mword_dec w (top - n) b.
+
+(*Parameter update_mword : forall {a}, bool -> mword a -> Z -> bitU -> mword a.*)
+Definition update_mword {a} (is_inc : bool) (w : mword a) n b :=
+ if is_inc then update_mword_inc w n b else update_mword_dec w n b.
+
+(*val int_of_mword : forall 'a. bool -> mword 'a -> integer*)
+Definition int_of_mword {a} `{ArithFact (a >= 0)} (sign : bool) (w : mword a) :=
+ if sign then wordToZ (get_word w) else Z.of_N (wordToN (get_word w)).
+
+
+(*val mword_of_int : forall a. Size a => Z -> Z -> mword a
+Definition mword_of_int len n :=
+ let w := wordFromInteger n in
+ if (length_mword w = len) then w else failwith "unexpected word length"
+*)
+Program Definition mword_of_int {len} `{H:ArithFact (len >= 0)} n : mword len :=
+match len with
+| Zneg _ => _
+| Z0 => ZToWord 0 n
+| Zpos p => ZToWord (Pos.to_nat p) n
+end.
+Next Obligation.
+destruct H.
+auto.
+Defined.
+(*
+(* Translating between a type level number (itself n) and an integer *)
+
+Definition size_itself_int x := Z.of_nat (size_itself x)
+
+(* NB: the corresponding sail type is forall n. atom(n) -> itself(n),
+ the actual integer is ignored. *)
+
+val make_the_value : forall n. Z -> itself n
+Definition inline make_the_value x := the_value
+*)
+
+Fixpoint wordFromBitlist_rev l : word (length l) :=
+match l with
+| [] => WO
+| b::t => WS b (wordFromBitlist_rev t)
+end.
+Definition wordFromBitlist l : word (length l) :=
+ nat_cast _ (List.rev_length l) (wordFromBitlist_rev (List.rev l)).
+
+Local Open Scope nat.
+
+Fixpoint nat_diff {T : nat -> Type} n m {struct n} :
+forall
+ (lt : forall p, T n -> T (n + p))
+ (eq : T m -> T m)
+ (gt : forall p, T (m + p) -> T m), T n -> T m :=
+(match n, m return (forall p, T n -> T (n + p)) -> (T m -> T m) -> (forall p, T (m + p) -> T m) -> T n -> T m with
+| O, O => fun lt eq gt => eq
+| S n', O => fun lt eq gt => gt _
+| O, S m' => fun lt eq gt => lt _
+| S n', S m' => @nat_diff (fun x => T (S x)) n' m'
+end).
+
+Definition fit_bbv_word {n m} : word n -> word m :=
+nat_diff n m
+ (fun p w => nat_cast _ (Nat.add_comm _ _) (extz w p))
+ (fun w => w)
+ (fun p w => split2 _ _ (nat_cast _ (Nat.add_comm _ _) w)).
+
+Local Close Scope nat.
+
+(*** Bitvectors *)
+
+Class Bitvector (a:Type) : Type := {
+ bits_of : a -> list bitU;
+ of_bits : list bitU -> option a;
+ of_bools : list bool -> a;
+ (* The first parameter specifies the desired length of the bitvector *)
+ of_int : Z -> Z -> a;
+ length : a -> Z;
+ unsigned : a -> option Z;
+ signed : a -> option Z;
+ arith_op_bv : (Z -> Z -> Z) -> bool -> a -> a -> a
+}.
+
+Instance bitlist_Bitvector {a : Type} `{BitU a} : (Bitvector (list a)) := {
+ bits_of v := List.map to_bitU v;
+ of_bits v := Some (List.map of_bitU v);
+ of_bools v := List.map of_bitU (List.map bitU_of_bool v);
+ of_int len n := List.map of_bitU (bits_of_int len n);
+ length := length_list;
+ unsigned v := unsigned_of_bits (List.map to_bitU v);
+ signed v := signed_of_bits (List.map to_bitU v);
+ arith_op_bv op sign l r := List.map of_bitU (arith_op_bits op sign (List.map to_bitU l) (List.map to_bitU r))
+}.
+
+Class ReasonableSize (a : Z) : Prop := {
+ isPositive : a >= 0
+}.
+
+(* Omega doesn't know about In, but can handle disjunctions. *)
+Ltac unfold_In :=
+repeat match goal with
+| H:context [member_Z_list _ _ = true] |- _ => rewrite member_Z_list_In in H
+| H:context [In ?x (?y :: ?t)] |- _ => change (In x (y :: t)) with (y = x \/ In x t) in H
+| H:context [In ?x []] |- _ => change (In x []) with False in H
+| |- context [member_Z_list _ _ = true] => rewrite member_Z_list_In
+| |- context [In ?x (?y :: ?t)] => change (In x (y :: t)) with (y = x \/ In x t)
+| |- context [In ?x []] => change (In x []) with False
+end.
+
+(* Definitions in the context that involve proof for other constraints can
+ break some of the constraint solving tactics, so prune definition bodies
+ down to integer types. *)
+Ltac not_Z_bool ty := match ty with Z => fail 1 | bool => fail 1 | _ => idtac end.
+Ltac clear_non_Z_bool_defns :=
+ repeat match goal with H := _ : ?X |- _ => not_Z_bool X; clearbody H end.
+Ltac clear_irrelevant_defns :=
+repeat match goal with X := _ |- _ =>
+ match goal with |- context[X] => idtac end ||
+ match goal with _ : context[X] |- _ => idtac end || clear X
+end.
+
+Lemma ArithFact_mword (a : Z) (w : mword a) : ArithFact (a >= 0).
+constructor.
+destruct a.
+auto with zarith.
+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']; clear H end.
+Ltac unbool_comparisons :=
+ repeat match goal with
+ | H:context [Z.geb _ _] |- _ => rewrite Z.geb_leb in H
+ | H:context [Z.gtb _ _] |- _ => rewrite Z.gtb_ltb in H
+ | H:context [Z.leb _ _ = true] |- _ => rewrite Z.leb_le in H
+ | H:context [Z.ltb _ _ = true] |- _ => rewrite Z.ltb_lt in H
+ | H:context [Z.eqb _ _ = true] |- _ => rewrite Z.eqb_eq in H
+ | H:context [Z.leb _ _ = false] |- _ => rewrite Z.leb_gt in H
+ | H:context [Z.ltb _ _ = false] |- _ => rewrite Z.ltb_ge in H
+ | H:context [Z.eqb _ _ = false] |- _ => rewrite Z.eqb_neq in H
+ | H:context [orb _ _ = true] |- _ => rewrite Bool.orb_true_iff in H
+ | H:context [orb _ _ = false] |- _ => rewrite Bool.orb_false_iff in H
+ | H:context [andb _ _ = true] |- _ => rewrite Bool.andb_true_iff in H
+ | H:context [andb _ _ = false] |- _ => rewrite Bool.andb_false_iff in H
+ | H:context [negb _ = true] |- _ => rewrite Bool.negb_true_iff in H
+ | H:context [negb _ = false] |- _ => rewrite Bool.negb_false_iff in H
+ | H:context [generic_eq _ _ = true] |- _ => apply generic_eq_true in H
+ | H:context [generic_eq _ _ = false] |- _ => apply generic_eq_false in H
+ | H:context [generic_neq _ _ = true] |- _ => apply generic_neq_true in H
+ | H:context [generic_neq _ _ = false] |- _ => apply generic_neq_false in H
+ | H:context [_ <> true] |- _ => rewrite Bool.not_true_iff_false in H
+ | H:context [_ <> false] |- _ => rewrite Bool.not_false_iff_true in H
+ end.
+Ltac unbool_comparisons_goal :=
+ repeat match goal with
+ | |- context [Z.geb _ _] => setoid_rewrite Z.geb_leb
+ | |- context [Z.gtb _ _] => setoid_rewrite Z.gtb_ltb
+ | |- context [Z.leb _ _ = true] => setoid_rewrite Z.leb_le
+ | |- context [Z.ltb _ _ = true] => setoid_rewrite Z.ltb_lt
+ | |- context [Z.eqb _ _ = true] => setoid_rewrite Z.eqb_eq
+ | |- context [Z.leb _ _ = false] => setoid_rewrite Z.leb_gt
+ | |- context [Z.ltb _ _ = false] => setoid_rewrite Z.ltb_ge
+ | |- context [Z.eqb _ _ = false] => setoid_rewrite Z.eqb_neq
+ | |- context [orb _ _ = true] => setoid_rewrite Bool.orb_true_iff
+ | |- context [orb _ _ = false] => setoid_rewrite Bool.orb_false_iff
+ | |- context [andb _ _ = true] => setoid_rewrite Bool.andb_true_iff
+ | |- context [andb _ _ = false] => setoid_rewrite Bool.andb_false_iff
+ | |- context [negb _ = true] => setoid_rewrite Bool.negb_true_iff
+ | |- context [negb _ = false] => setoid_rewrite Bool.negb_false_iff
+ | |- context [generic_eq _ _ = true] => apply generic_eq_true
+ | |- context [generic_eq _ _ = false] => apply generic_eq_false
+ | |- context [generic_neq _ _ = true] => apply generic_neq_true
+ | |- context [generic_neq _ _ = false] => apply generic_neq_false
+ | |- context [_ <> true] => setoid_rewrite Bool.not_true_iff_false
+ | |- context [_ <> false] => setoid_rewrite Bool.not_false_iff_true
+ end.
+
+(* Split up dependent pairs to get at proofs of properties *)
+Ltac extract_properties :=
+ (* Properties of local definitions *)
+ repeat match goal with H := 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 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 :=
+ repeat match goal with |- context [length_list ?X] =>
+ let r := (eval cbn in (length_list X)) in
+ change (length_list X) with r
+ end.
+(* TODO: can we restrict this to concrete terms? *)
+Ltac reduce_pow :=
+ repeat match goal with H:context [Z.pow ?X ?Y] |- _ =>
+ let r := (eval cbn in (Z.pow X Y)) in
+ change (Z.pow X Y) with r in H
+ end;
+ repeat match goal with |- context [Z.pow ?X ?Y] =>
+ let r := (eval cbn in (Z.pow X Y)) in
+ change (Z.pow X Y) with r
+ end.
+Ltac dump_context :=
+ repeat match goal with
+ | 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.
+Lemma True_left {P:Prop} : (True /\ P) <-> P.
+tauto.
+Qed.
+Lemma True_right {P:Prop} : (P /\ True) <-> P.
+tauto.
+Qed.
+
+(* Turn exists into metavariables like eexists, except put in dummy values when
+ the variable is unused. This is used so that we can use eauto with a low
+ search bound that doesn't include the exists. (Not terribly happy with
+ how this works...) *)
+Ltac drop_Z_exists :=
+repeat
+ match goal with |- @ex Z ?p =>
+ let a := eval hnf in (p 0) in
+ let b := eval hnf in (p 1) in
+ match a with b => exists 0 | _ => eexists end
+ end.
+(*
+ match goal with |- @ex Z (fun x => @?p x) =>
+ let xx := fresh "x" in
+ evar (xx : Z);
+ let a := eval hnf in (p xx) in
+ match a with context [xx] => eexists | _ => exists 0 end;
+ instantiate (xx := 0);
+ clear xx
+ end.
+*)
+(* For boolean solving we just use plain metavariables *)
+Ltac drop_bool_exists :=
+repeat match goal with |- @ex bool _ => eexists end.
+
+(* The linear solver doesn't like existentials. *)
+Ltac destruct_exists :=
+ repeat match goal with H:@ex Z _ |- _ => destruct H end;
+ repeat match goal with H:@ex bool _ |- _ => destruct H end.
+
+(* The ASL to Sail translator sometimes puts constraints of the form
+ p | not(q) into function signatures, then the body case splits on q.
+ The filter_disjunctions tactic simplifies hypotheses by obtaining p. *)
+
+Lemma truefalse : true = false <-> False.
+intuition.
+Qed.
+Lemma falsetrue : false = true <-> False.
+intuition.
+Qed.
+Lemma or_False_l P : False \/ P <-> P.
+intuition.
+Qed.
+Lemma or_False_r P : P \/ False <-> P.
+intuition.
+Qed.
+
+Ltac filter_disjunctions :=
+ repeat match goal with
+ | H1:?P \/ ?t1 = ?t2, H2: ?t3 = ?t4 |- _ =>
+ (* I used to use non-linear matching above, but Coq is happy to match up
+ to conversion, including more unfolding than we normally do. *)
+ constr_eq t1 t3; constr_eq t2 t4; clear H1
+ | H1:context [?P \/ ?t = true], H2: ?t = false |- _ => is_var t; rewrite H2 in H1
+ | H1:context [?P \/ ?t = false], H2: ?t = true |- _ => is_var t; rewrite H2 in H1
+ | H1:context [?t = true \/ ?P], H2: ?t = false |- _ => is_var t; rewrite H2 in H1
+ | H1:context [?t = false \/ ?P], H2: ?t = true |- _ => is_var t; rewrite H2 in H1
+ end;
+ rewrite ?truefalse, ?falsetrue, ?or_False_l, ?or_False_r in *;
+ (* We may have uncovered more conjunctions *)
+ repeat match goal with H:and _ _ |- _ => destruct H end.
+
+(* Turn x := if _ then ... into x = ... \/ x = ... *)
+
+Ltac Z_if_to_or :=
+ repeat match goal with x := ?t : Z |- _ =>
+ let rec build_goal t :=
+ match t with
+ | if _ then ?y else ?z =>
+ let Hy := build_goal y in
+ let Hz := build_goal z in
+ constr:(Hy \/ Hz)
+ | ?y => constr:(x = y)
+ end
+ in
+ let rec split_hyp t :=
+ match t with
+ | if ?b then ?y else ?z =>
+ destruct b in x; [split_hyp y| split_hyp z]
+ | _ => idtac
+ end
+ in
+ let g := build_goal t in
+ assert g by (clear -x; split_hyp t; auto);
+ clearbody x
+ end.
+
+(* Once we've done everything else, get rid of irrelevant bool and Z bindings
+ to help the brute force solver *)
+Ltac clear_irrelevant_bindings :=
+ repeat
+ match goal with
+ | b : bool |- _ =>
+ lazymatch goal with
+ | _ : context [b] |- _ => fail
+ | |- context [b] => fail
+ | _ => clear b
+ end
+ | x : Z |- _ =>
+ lazymatch goal with
+ | _ : context [x] |- _ => fail
+ | |- context [x] => fail
+ | _ => clear x
+ end
+ | H:?x |- _ =>
+ let s := type of x in
+ lazymatch s with
+ | Prop =>
+ match x with
+ | context [?v] => is_var v; fail 1
+ | _ => clear H
+ end
+ | _ => fail
+ end
+ end.
+
+(* Currently, the ASL to Sail translation produces some constraints of the form
+ P \/ x = true, P \/ x = false, which are simplified by the tactic below. In
+ future the translation is likely to be cleverer, and this won't be
+ necessary. *)
+(* TODO: remove duplication with filter_disjunctions *)
+Lemma remove_unnecessary_casesplit {P:Prop} {x} :
+ P \/ x = true -> P \/ x = false -> P.
+ intuition congruence.
+Qed.
+Lemma remove_eq_false_true {P:Prop} {x} :
+ x = true -> P \/ x = false -> P.
+intros H1 [H|H]; congruence.
+Qed.
+Lemma remove_eq_true_false {P:Prop} {x} :
+ x = false -> P \/ x = true -> P.
+intros H1 [H|H]; congruence.
+Qed.
+Ltac remove_unnecessary_casesplit :=
+repeat match goal with
+| H1 : ?P \/ ?v = true, H2 : ?v = true |- _ => clear H1
+| H1 : ?P \/ ?v = true, H2 : ?v = false |- _ => apply (remove_eq_true_false H2) in H1
+| H1 : ?P \/ ?v = false, H2 : ?v = false |- _ => clear H1
+| H1 : ?P \/ ?v = false, H2 : ?v = true |- _ => apply (remove_eq_false_true H2) in H1
+| H1 : ?P \/ ?v1 = true, H2 : ?P \/ ?v2 = false |- _ =>
+ constr_eq v1 v2;
+ is_var v1;
+ apply (remove_unnecessary_casesplit H1) in H2;
+ clear H1
+ (* There are worse cases where the hypotheses are different, so we actually
+ do the casesplit *)
+| H1 : _ \/ ?v = true, H2 : _ \/ ?v = false |- _ =>
+ is_var v;
+ destruct v;
+ [ clear H1; destruct H2; [ | congruence ]
+ | clear H2; destruct H1; [ | congruence ]
+ ]
+end;
+(* We may have uncovered more conjunctions *)
+repeat match goal with H:and _ _ |- _ => destruct H end.
+
+Ltac generalize_embedded_proofs :=
+ repeat match goal with H:context [?X] |- _ =>
+ match type of X with ArithFact _ =>
+ generalize dependent X
+ end
+ end;
+ intros.
+
+Lemma iff_equal_l {T:Type} {P:Prop} {x:T} : (x = x <-> P) -> P.
+tauto.
+Qed.
+Lemma iff_equal_r {T:Type} {P:Prop} {x:T} : (P <-> x = x) -> P.
+tauto.
+Qed.
+
+Lemma iff_known_l {P Q : Prop} : P -> P <-> Q -> Q.
+tauto.
+Qed.
+Lemma iff_known_r {P Q : Prop} : P -> Q <-> P -> Q.
+tauto.
+Qed.
+
+Ltac clean_up_props :=
+ repeat match goal with
+ (* I did try phrasing these as rewrites, but Coq was oddly reluctant to use them *)
+ | H:?x = ?x <-> _ |- _ => apply iff_equal_l in H
+ | H:_ <-> ?x = ?x |- _ => apply iff_equal_r in H
+ | H:context[true = false] |- _ => rewrite truefalse in H
+ | H:context[false = true] |- _ => rewrite falsetrue in H
+ | H1:?P <-> False, H2:context[?Q] |- _ => constr_eq P Q; rewrite -> H1 in H2
+ | H1:False <-> ?P, H2:context[?Q] |- _ => constr_eq P Q; rewrite <- H1 in H2
+ | H1:?P, H2:?Q <-> ?R |- _ => constr_eq P Q; apply (iff_known_l H1) in H2
+ | H1:?P, H2:?R <-> ?Q |- _ => constr_eq P Q; apply (iff_known_r H1) in H2
+ | H:context[_ \/ False] |- _ => rewrite or_False_r in H
+ | H:context[False \/ _] |- _ => rewrite or_False_l in H
+ (* omega doesn't cope well with extra "True"s in the goal.
+ Check that they actually appear because setoid_rewrite can fill in evars. *)
+ | |- context[True /\ _] => setoid_rewrite True_left
+ | |- context[_ /\ True] => setoid_rewrite True_right
+ end;
+ remove_unnecessary_casesplit.
+
+Ltac prepare_for_solver :=
+(*dump_context;*)
+ generalize_embedded_proofs;
+ clear_irrelevant_defns;
+ clear_non_Z_bool_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;
+ destruct_exists;
+ unbool_comparisons;
+ unbool_comparisons_goal;
+ repeat match goal with H:and _ _ |- _ => destruct H end;
+ remove_unnecessary_casesplit;
+ unfold_In; (* after unbool_comparisons to deal with && and || *)
+ reduce_list_lengths;
+ reduce_pow;
+ filter_disjunctions;
+ Z_if_to_or;
+ clear_irrelevant_bindings;
+ subst;
+ clean_up_props.
+
+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 bruteforce_bool_exists :=
+match goal with
+| |- exists _ : bool,_ => solve [ exists true; bruteforce_bool_exists
+ | exists false; bruteforce_bool_exists ]
+| _ => tauto
+end.
+
+Lemma or_iff_cong : forall A B C D, A <-> B -> C <-> D -> A \/ C <-> B \/ D.
+intros.
+tauto.
+Qed.
+
+Lemma and_iff_cong : forall A B C D, A <-> B -> C <-> D -> A /\ C <-> B /\ D.
+intros.
+tauto.
+Qed.
+
+Ltac solve_euclid :=
+repeat match goal with
+| |- context [ZEuclid.modulo ?x ?y] =>
+ specialize (ZEuclid.div_mod x y);
+ specialize (ZEuclid.mod_always_pos x y);
+ generalize (ZEuclid.modulo x y);
+ generalize (ZEuclid.div x y);
+ intros
+| |- context [ZEuclid.div ?x ?y] =>
+ specialize (ZEuclid.div_mod x y);
+ specialize (ZEuclid.mod_always_pos x y);
+ generalize (ZEuclid.modulo x y);
+ generalize (ZEuclid.div x y);
+ intros
+end;
+nia.
+
+(* A more ambitious brute force existential solver. *)
+
+Ltac guess_ex_solver :=
+ match goal with
+ | |- @ex bool ?t =>
+ match t with
+ | context [@eq bool ?b _] =>
+ solve [ exists b; guess_ex_solver
+ | exists (negb b); rewrite ?Bool.negb_true_iff, ?Bool.negb_false_iff;
+ guess_ex_solver ]
+ end
+(* | b : bool |- @ex bool _ => exists b; guess_ex_solver
+ | b : bool |- @ex bool _ =>
+ exists (negb b); rewrite ?Bool.negb_true_iff, ?Bool.negb_false_iff;
+ guess_ex_solver*)
+ | |- @ex bool _ => exists true; guess_ex_solver
+ | |- @ex bool _ => exists false; guess_ex_solver
+ | x : Z |- @ex Z _ => exists x; guess_ex_solver
+ | _ => solve [tauto | eauto 3 with zarith sail | omega | intuition]
+ end.
+
+(* A straightforward solver for simple problems like
+
+ exists ..., _ = true \/ _ = false /\ _ = true <-> _ = true \/ _ = true
+*)
+
+Ltac form_iff_true :=
+repeat match goal with
+| |- ?l <-> _ = true =>
+ let rec aux t :=
+ match t with
+ | _ = true \/ _ = true => rewrite <- Bool.orb_true_iff
+ | _ = true /\ _ = true => rewrite <- Bool.andb_true_iff
+ | _ = false => rewrite <- Bool.negb_true_iff
+ | ?l \/ ?r => aux l || aux r
+ | ?l /\ ?r => aux l || aux r
+ end
+ in aux l
+ end.
+Ltac simple_split_iff :=
+ repeat
+ match goal with
+ | |- _ /\ _ <-> _ /\ _ => apply and_iff_cong
+ | |- _ \/ _ <-> _ \/ _ => apply or_iff_cong
+ end.
+Ltac simple_ex_iff :=
+ match goal with
+ | |- @ex _ _ => eexists; simple_ex_iff
+ | |- _ <-> _ =>
+ simple_split_iff;
+ form_iff_true;
+ solve [apply iff_refl | eassumption]
+ end.
+
+(* Another attempt at similar goals, this time allowing for conjuncts to move
+ around, and filling in integer existentials and redundant boolean ones.
+ TODO: generalise / combine with simple_ex_iff. *)
+
+Ltac ex_iff_construct_bool_witness :=
+let rec search x y :=
+ lazymatch y with
+ | x => constr:(true)
+ | ?y1 /\ ?y2 =>
+ let b1 := search x y1 in
+ let b2 := search x y2 in
+ constr:(orb b1 b2)
+ | _ => constr:(false)
+ end
+in
+let rec make_clause x :=
+ lazymatch x with
+ | ?l = true => l
+ | ?l = false => constr:(negb l)
+ | @eq Z ?l ?n => constr:(Z.eqb l n)
+ | ?p \/ ?q =>
+ let p' := make_clause p in
+ let q' := make_clause q in
+ constr:(orb p' q')
+ | _ => fail
+ end in
+let add_clause x xs :=
+ let l := make_clause x in
+ match xs with
+ | true => l
+ | _ => constr:(andb l xs)
+ end
+in
+let rec construct_ex l r x :=
+ lazymatch l with
+ | ?l1 /\ ?l2 =>
+ let y := construct_ex l1 r x in
+ construct_ex l2 r y
+ | _ =>
+ let present := search l r in
+ lazymatch eval compute in present with true => x | _ => add_clause l x end
+ end
+in
+let witness := match goal with
+| |- ?l <-> ?r => construct_ex l r constr:(true)
+end in
+instantiate (1 := witness).
+
+Ltac ex_iff_fill_in_ints :=
+ let rec search l r y :=
+ match y with
+ | l = r => idtac
+ | ?v = r => is_evar v; unify v l
+ | ?y1 /\ ?y2 => first [search l r y1 | search l r y2]
+ | _ => fail
+ end
+ in
+ match goal with
+ | |- ?l <-> ?r =>
+ let rec traverse l :=
+ lazymatch l with
+ | ?l1 /\ ?l2 =>
+ traverse l1; traverse l2
+ | @eq Z ?x ?y => search x y r
+ | _ => idtac
+ end
+ in traverse l
+ end.
+
+Ltac ex_iff_fill_in_bools :=
+ let rec traverse t :=
+ lazymatch t with
+ | ?v = ?t => try (is_evar v; unify v t)
+ | ?p /\ ?q => traverse p; traverse q
+ | _ => idtac
+ end
+ in match goal with
+ | |- _ <-> ?r => traverse r
+ end.
+
+Ltac conjuncts_iff_solve :=
+ ex_iff_fill_in_ints;
+ ex_iff_construct_bool_witness;
+ ex_iff_fill_in_bools;
+ unbool_comparisons_goal;
+ clear;
+ intuition.
+
+Ltac ex_iff_solve :=
+ match goal with
+ | |- @ex _ _ => eexists; ex_iff_solve
+ (* Range constraints are attached to the right *)
+ | |- _ /\ _ => split; [ex_iff_solve | omega]
+ | |- _ <-> _ => conjuncts_iff_solve
+ end.
+
+
+Lemma iff_false_left {P Q R : Prop} : (false = true) <-> Q -> (false = true) /\ P <-> Q /\ R.
+intuition.
+Qed.
+
+(* Very simple proofs for trivial arithmetic. Preferable to running omega/lia because
+ they can get bogged down if they see large definitions; should also guarantee small
+ proof terms. *)
+Lemma Z_compare_lt_eq : Lt = Eq -> False. congruence. Qed.
+Lemma Z_compare_lt_gt : Lt = Gt -> False. congruence. Qed.
+Lemma Z_compare_eq_lt : Eq = Lt -> False. congruence. Qed.
+Lemma Z_compare_eq_gt : Eq = Gt -> False. congruence. Qed.
+Lemma Z_compare_gt_lt : Gt = Lt -> False. congruence. Qed.
+Lemma Z_compare_gt_eq : Gt = Eq -> False. congruence. Qed.
+Ltac z_comparisons :=
+ (* Don't try terms with variables - reduction may be expensive *)
+ match goal with |- context[?x] => is_var x; fail 1 | |- _ => idtac end;
+ solve [
+ exact eq_refl
+ | exact Z_compare_lt_eq
+ | exact Z_compare_lt_gt
+ | exact Z_compare_eq_lt
+ | exact Z_compare_eq_gt
+ | exact Z_compare_gt_lt
+ | exact Z_compare_gt_eq
+ ].
+
+(* Try to get the linear arithmetic solver to do booleans. *)
+
+Lemma b2z_true x : x = true <-> Z.b2z x = 1.
+destruct x; compute; split; congruence.
+Qed.
+
+Lemma b2z_false x : x = false <-> Z.b2z x = 0.
+destruct x; compute; split; congruence.
+Qed.
+
+Lemma b2z_tf x : 0 <= Z.b2z x <= 1.
+destruct x; simpl; omega.
+Qed.
+
+Ltac solve_bool_with_Z :=
+ subst;
+ rewrite ?truefalse, ?falsetrue, ?or_False_l, ?or_False_r in *;
+ (* I did try phrasing these as rewrites, but Coq was oddly reluctant to use them *)
+ repeat match goal with
+ | H:?x = ?x <-> _ |- _ => apply iff_equal_l in H
+ | H:_ <-> ?x = ?x |- _ => apply iff_equal_r in H
+ end;
+ repeat match goal with
+ | H:context [?v = true] |- _ => is_var v; rewrite (b2z_true v) in *
+ | |- context [?v = true] => is_var v; rewrite (b2z_true v) in *
+ | H:context [?v = false] |- _ => is_var v; rewrite (b2z_false v) in *
+ | |- context [?v = false] => is_var v; rewrite (b2z_false v) in *
+ end;
+ repeat match goal with
+ | _:context [Z.b2z ?v] |- _ => generalize (b2z_tf v); generalize dependent (Z.b2z v)
+ | |- context [Z.b2z ?v] => generalize (b2z_tf v); generalize dependent (Z.b2z v)
+ end;
+ intros;
+ lia.
+
+
+(* Redefine this to add extra solver tactics *)
+Ltac sail_extra_tactic := fail.
+
+Ltac main_solver :=
+ solve
+ [ apply ArithFact_mword; assumption
+ | z_comparisons
+ | omega with Z
+ (* Try sail hints before dropping the existential *)
+ | subst; eauto 3 with zarith sail
+ (* The datatypes hints give us some list handling, esp In *)
+ | subst; drop_Z_exists;
+ repeat match goal with |- and _ _ => split end;
+ eauto 3 with datatypes zarith sail
+ | subst; match goal with |- context [ZEuclid.div] => solve_euclid
+ | |- context [ZEuclid.modulo] => solve_euclid
+ end
+ | match goal with |- context [Z.mul] => nia end
+ (* If we have a disjunction from a set constraint on a variable we can often
+ solve a goal by trying them (admittedly this is quite heavy handed...) *)
+ | subst; drop_Z_exists;
+ let aux x :=
+ is_var x;
+ intuition (subst;auto with datatypes)
+ in
+ match goal with
+ | _:(@eq Z _ ?x) \/ (@eq Z _ ?x) \/ _ |- context[?x] => aux x
+ | _:(@eq Z ?x _) \/ (@eq Z ?x _) \/ _ |- context[?x] => aux x
+ | _:(@eq Z _ ?x) \/ (@eq Z _ ?x) \/ _, _:@eq Z ?y (ZEuclid.div ?x _) |- context[?y] => is_var x; aux y
+ | _:(@eq Z ?x _) \/ (@eq Z ?x _) \/ _, _:@eq Z ?y (ZEuclid.div ?x _) |- context[?y] => is_var x; aux y
+ end
+ (* Booleans - and_boolMP *)
+ | solve_bool_with_Z
+ | simple_ex_iff
+ | ex_iff_solve
+ | drop_bool_exists; solve [eauto using iff_refl, or_iff_cong, and_iff_cong | intuition]
+ | match goal with |- (forall l r:bool, _ -> _ -> exists _ : bool, _) =>
+ let r := fresh "r" in
+ let H1 := fresh "H" in
+ let H2 := fresh "H" in
+ intros [|] r H1 H2;
+ let t2 := type of H2 in
+ match t2 with
+ | ?b = ?b -> _ =>
+ destruct (H2 eq_refl);
+ repeat match goal with H:@ex _ _ |- _ => destruct H end;
+ simple_ex_iff
+ | ?b = _ -> _ =>
+ repeat match goal with H:@ex _ _ |- _ => destruct H end;
+ clear H2;
+ repeat match goal with
+ | |- @ex bool _ => exists b
+ | |- @ex Z _ => exists 0
+ end;
+ intuition
+ end
+ end
+ | match goal with |- (forall l r:bool, _ -> _ -> @ex _ _) =>
+ let H1 := fresh "H" in
+ let H2 := fresh "H" in
+ intros [|] [|] H1 H2;
+ repeat match goal with H:?X = ?X -> _ |- _ => specialize (H eq_refl) end;
+ repeat match goal with H:@ex _ _ |- _ => destruct H end;
+ guess_ex_solver
+ end
+ | match goal with |- @ex _ _ => guess_ex_solver end
+(* While firstorder was quite effective at dealing with existentially quantified
+ goals from boolean expressions, it attempts lazy normalization of terms,
+ which blows up on integer comparisons with large constants.
+ | match goal with |- context [@eq bool _ _] =>
+ (* Don't use auto for the fallback to keep runtime down *)
+ firstorder fail
+ end*)
+ | sail_extra_tactic
+ | idtac "Unable to solve constraint"; dump_context; fail
+ ].
+
+(* Omega can get upset by local definitions that are projections from value/proof pairs.
+ Complex goals can use prepare_for_solver to extract facts; this tactic can be used
+ for simpler proofs without using prepare_for_solver. *)
+Ltac simple_omega :=
+ repeat match goal with
+ H := projT1 _ |- _ => clearbody H
+ end; omega.
+
+Ltac solve_unknown :=
+ match goal with |- (ArithFact (?x ?y)) =>
+ is_evar x;
+ idtac "Warning: unknown constraint";
+ let t := type of y in
+ unify x (fun (_ : t) => True);
+ exact (Build_ArithFact _ I)
+ 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 solve_unknown;
+match goal with |- ArithFact (?x <= ?x <= ?x) => try (exact trivial_range) | _ => idtac end;
+try fill_in_evar_eq;
+try match goal with |- context [projT1 ?X] => apply (ArithFact_self_proof X) end;
+(* Trying reflexivity will fill in more complex metavariable examples than
+ fill_in_evar_eq above, e.g., 8 * n = 8 * ?Goal3 *)
+try (constructor; reflexivity);
+try (constructor; repeat match goal with |- and _ _ => split end; z_comparisons);
+try (constructor; simple_omega);
+prepare_for_solver;
+(*dump_context;*)
+constructor;
+repeat match goal with |- and _ _ => split end;
+main_solver.
+
+(* Add an indirection so that you can redefine run_solver to fail to get
+ slow running constraints into proof mode. *)
+Ltac run_solver := solve_arithfact.
+Hint Extern 0 (ArithFact _) => run_solver : typeclass_instances.
+
+Hint Unfold length_mword : sail.
+
+Lemma unit_comparison_lemma : true = true <-> True.
+intuition.
+Qed.
+Hint Resolve unit_comparison_lemma : sail.
+
+Definition neq_atom (x : Z) (y : Z) : bool := negb (Z.eqb x y).
+Hint Unfold neq_atom : sail.
+
+Lemma ReasonableSize_witness (a : Z) (w : mword a) : ReasonableSize a.
+constructor.
+destruct a.
+auto with zarith.
+auto using Z.le_ge, Zle_0_pos.
+destruct w.
+Qed.
+
+Hint Extern 0 (ReasonableSize ?A) => (unwrap_ArithFacts; solve [apply ReasonableSize_witness; assumption | constructor; omega]) : typeclass_instances.
+
+Definition to_range (x : Z) : {y : Z & ArithFact (x <= y <= x)} := build_ex x.
+
+
+
+Instance mword_Bitvector {a : Z} `{ArithFact (a >= 0)} : (Bitvector (mword a)) := {
+ bits_of v := List.map bitU_of_bool (bitlistFromWord (get_word v));
+ of_bits v := option_map (fun bl => to_word isPositive (fit_bbv_word (wordFromBitlist bl))) (just_list (List.map bool_of_bitU v));
+ of_bools v := to_word isPositive (fit_bbv_word (wordFromBitlist v));
+ of_int len z := mword_of_int z; (* cheat a little *)
+ length v := a;
+ unsigned v := Some (Z.of_N (wordToN (get_word v)));
+ signed v := Some (wordToZ (get_word v));
+ arith_op_bv op sign l r := mword_of_int (op (int_of_mword sign l) (int_of_mword sign r))
+}.
+
+Section Bitvector_defs.
+Context {a b} `{Bitvector a} `{Bitvector b}.
+
+Definition opt_def {a} (def:a) (v:option a) :=
+match v with
+| Some x => x
+| None => def
+end.
+
+(* The Lem version is partial, but lets go with BU here to avoid constraints for now *)
+Definition access_bv_inc (v : a) n := opt_def BU (access_list_opt_inc (bits_of v) n).
+Definition access_bv_dec (v : a) n := opt_def BU (access_list_opt_dec (bits_of v) n).
+
+Definition update_bv_inc (v : a) n b := update_list true (bits_of v) n b.
+Definition update_bv_dec (v : a) n b := update_list false (bits_of v) n b.
+
+Definition subrange_bv_inc (v : a) i j := subrange_list true (bits_of v) i j.
+Definition subrange_bv_dec (v : a) i j := subrange_list true (bits_of v) i j.
+
+Definition update_subrange_bv_inc (v : a) i j (v' : b) := update_subrange_list true (bits_of v) i j (bits_of v').
+Definition update_subrange_bv_dec (v : a) i j (v' : b) := update_subrange_list false (bits_of v) i j (bits_of v').
+
+(*val extz_bv : forall a b. Bitvector a, Bitvector b => Z -> a -> b*)
+Definition extz_bv n (v : a) : option b := of_bits (extz_bits n (bits_of v)).
+
+(*val exts_bv : forall a b. Bitvector a, Bitvector b => Z -> a -> b*)
+Definition exts_bv n (v : a) : option b := of_bits (exts_bits n (bits_of v)).
+
+(*val string_of_bv : forall a. Bitvector a => a -> string *)
+Definition string_of_bv v := show_bitlist (bits_of v).
+
+End Bitvector_defs.
+
+(*** Bytes and addresses *)
+
+Definition memory_byte := list bitU.
+
+(*val byte_chunks : forall a. list a -> option (list (list a))*)
+Fixpoint byte_chunks {a} (bs : list a) := match bs with
+ | [] => Some []
+ | a::b::c::d::e::f::g::h::rest =>
+ match byte_chunks rest with
+ | None => None
+ | Some rest => Some ([a;b;c;d;e;f;g;h] :: rest)
+ end
+ | _ => None
+end.
+(*declare {isabelle} termination_argument byte_chunks = automatic*)
+
+Section BytesBits.
+Context {a} `{Bitvector a}.
+
+(*val bytes_of_bits : forall a. Bitvector a => a -> option (list memory_byte)*)
+Definition bytes_of_bits (bs : a) := byte_chunks (bits_of bs).
+
+(*val bits_of_bytes : forall a. Bitvector a => list memory_byte -> a*)
+Definition bits_of_bytes (bs : list memory_byte) : list bitU := List.concat (List.map bits_of bs).
+
+Definition mem_bytes_of_bits (bs : a) := option_map (@rev (list bitU)) (bytes_of_bits bs).
+Definition bits_of_mem_bytes (bs : list memory_byte) := bits_of_bytes (List.rev bs).
+
+End BytesBits.
+
+(*val bitv_of_byte_lifteds : list Sail_impl_base.byte_lifted -> list bitU
+Definition bitv_of_byte_lifteds v :=
+ foldl (fun x (Byte_lifted y) => x ++ (List.map bitU_of_bit_lifted y)) [] v
+
+val bitv_of_bytes : list Sail_impl_base.byte -> list bitU
+Definition bitv_of_bytes v :=
+ foldl (fun x (Byte y) => x ++ (List.map bitU_of_bit y)) [] v
+
+val byte_lifteds_of_bitv : list bitU -> list byte_lifted
+Definition byte_lifteds_of_bitv bits :=
+ let bits := List.map bit_lifted_of_bitU bits in
+ byte_lifteds_of_bit_lifteds bits
+
+val bytes_of_bitv : list bitU -> list byte
+Definition bytes_of_bitv bits :=
+ let bits := List.map bit_of_bitU bits in
+ bytes_of_bits bits
+
+val bit_lifteds_of_bitUs : list bitU -> list bit_lifted
+Definition bit_lifteds_of_bitUs bits := List.map bit_lifted_of_bitU bits
+
+val bit_lifteds_of_bitv : list bitU -> list bit_lifted
+Definition bit_lifteds_of_bitv v := bit_lifteds_of_bitUs v
+
+
+val address_lifted_of_bitv : list bitU -> address_lifted
+Definition address_lifted_of_bitv v :=
+ let byte_lifteds := byte_lifteds_of_bitv v in
+ let maybe_address_integer :=
+ match (maybe_all (List.map byte_of_byte_lifted byte_lifteds)) with
+ | Some bs => Some (integer_of_byte_list bs)
+ | _ => None
+ end in
+ Address_lifted byte_lifteds maybe_address_integer
+
+val bitv_of_address_lifted : address_lifted -> list bitU
+Definition bitv_of_address_lifted (Address_lifted bs _) := bitv_of_byte_lifteds bs
+
+val address_of_bitv : list bitU -> address
+Definition address_of_bitv v :=
+ let bytes := bytes_of_bitv v in
+ address_of_byte_list bytes*)
+
+Fixpoint reverse_endianness_list (bits : list bitU) :=
+ match bits with
+ | _ :: _ :: _ :: _ :: _ :: _ :: _ :: _ :: t =>
+ reverse_endianness_list t ++ firstn 8 bits
+ | _ => bits
+ end.
+
+(*** Registers *)
+
+Definition register_field := string.
+Definition register_field_index : Type := string * (Z * Z). (* name, start and end *)
+
+Inductive register :=
+ | Register : string * (* name *)
+ Z * (* length *)
+ Z * (* start index *)
+ bool * (* is increasing *)
+ list register_field_index
+ -> register
+ | UndefinedRegister : Z -> register (* length *)
+ | RegisterPair : register * register -> register.
+
+Record register_ref regstate regval a :=
+ { name : string;
+ (*is_inc : bool;*)
+ read_from : regstate -> a;
+ write_to : a -> regstate -> regstate;
+ of_regval : regval -> option a;
+ regval_of : a -> regval }.
+Notation "{[ r 'with' 'name' := e ]}" := ({| name := e; read_from := read_from r; write_to := write_to r; of_regval := of_regval r; regval_of := regval_of r |}).
+Notation "{[ r 'with' 'read_from' := e ]}" := ({| read_from := e; name := name r; write_to := write_to r; of_regval := of_regval r; regval_of := regval_of r |}).
+Notation "{[ r 'with' 'write_to' := e ]}" := ({| write_to := e; name := name r; read_from := read_from r; of_regval := of_regval r; regval_of := regval_of r |}).
+Notation "{[ r 'with' 'of_regval' := e ]}" := ({| of_regval := e; name := name r; read_from := read_from r; write_to := write_to r; regval_of := regval_of r |}).
+Notation "{[ r 'with' 'regval_of' := e ]}" := ({| regval_of := e; name := name r; read_from := read_from r; write_to := write_to r; of_regval := of_regval r |}).
+Arguments name [_ _ _].
+Arguments read_from [_ _ _].
+Arguments write_to [_ _ _].
+Arguments of_regval [_ _ _].
+Arguments regval_of [_ _ _].
+
+(* Register accessors: pair of functions for reading and writing register values *)
+Definition register_accessors regstate regval : Type :=
+ ((string -> regstate -> option regval) *
+ (string -> regval -> regstate -> option regstate)).
+
+Record field_ref regtype a :=
+ { field_name : string;
+ field_start : Z;
+ field_is_inc : bool;
+ get_field : regtype -> a;
+ set_field : regtype -> a -> regtype }.
+Arguments field_name [_ _].
+Arguments field_start [_ _].
+Arguments field_is_inc [_ _].
+Arguments get_field [_ _].
+Arguments set_field [_ _].
+
+(*
+(*let name_of_reg := function
+ | Register name _ _ _ _ => name
+ | UndefinedRegister _ => failwith "name_of_reg UndefinedRegister"
+ | RegisterPair _ _ => failwith "name_of_reg RegisterPair"
+end
+
+Definition size_of_reg := function
+ | Register _ size _ _ _ => size
+ | UndefinedRegister size => size
+ | RegisterPair _ _ => failwith "size_of_reg RegisterPair"
+end
+
+Definition start_of_reg := function
+ | Register _ _ start _ _ => start
+ | UndefinedRegister _ => failwith "start_of_reg UndefinedRegister"
+ | RegisterPair _ _ => failwith "start_of_reg RegisterPair"
+end
+
+Definition is_inc_of_reg := function
+ | Register _ _ _ is_inc _ => is_inc
+ | UndefinedRegister _ => failwith "is_inc_of_reg UndefinedRegister"
+ | RegisterPair _ _ => failwith "in_inc_of_reg RegisterPair"
+end
+
+Definition dir_of_reg := function
+ | Register _ _ _ is_inc _ => dir_of_bool is_inc
+ | UndefinedRegister _ => failwith "dir_of_reg UndefinedRegister"
+ | RegisterPair _ _ => failwith "dir_of_reg RegisterPair"
+end
+
+Definition size_of_reg_nat reg := Z.to_nat (size_of_reg reg)
+Definition start_of_reg_nat reg := Z.to_nat (start_of_reg reg)
+
+val register_field_indices_aux : register -> register_field -> option (Z * Z)
+Fixpoint register_field_indices_aux register rfield :=
+ match register with
+ | Register _ _ _ _ rfields => List.lookup rfield rfields
+ | RegisterPair r1 r2 =>
+ let m_indices := register_field_indices_aux r1 rfield in
+ if isSome m_indices then m_indices else register_field_indices_aux r2 rfield
+ | UndefinedRegister _ => None
+ end
+
+val register_field_indices : register -> register_field -> Z * Z
+Definition register_field_indices register rfield :=
+ match register_field_indices_aux register rfield with
+ | Some indices => indices
+ | None => failwith "Invalid register/register-field combination"
+ end
+
+Definition register_field_indices_nat reg regfield=
+ let (i,j) := register_field_indices reg regfield in
+ (Z.to_nat i,Z.to_nat j)*)
+
+(*let rec external_reg_value reg_name v :=
+ let (internal_start, external_start, direction) :=
+ match reg_name with
+ | Reg _ start size dir =>
+ (start, (if dir = D_increasing then start else (start - (size +1))), dir)
+ | Reg_slice _ reg_start dir (slice_start, _) =>
+ ((if dir = D_increasing then slice_start else (reg_start - slice_start)),
+ slice_start, dir)
+ | Reg_field _ reg_start dir _ (slice_start, _) =>
+ ((if dir = D_increasing then slice_start else (reg_start - slice_start)),
+ slice_start, dir)
+ | Reg_f_slice _ reg_start dir _ _ (slice_start, _) =>
+ ((if dir = D_increasing then slice_start else (reg_start - slice_start)),
+ slice_start, dir)
+ end in
+ let bits := bit_lifteds_of_bitv v in
+ <| rv_bits := bits;
+ rv_dir := direction;
+ rv_start := external_start;
+ rv_start_internal := internal_start |>
+
+val internal_reg_value : register_value -> list bitU
+Definition internal_reg_value v :=
+ List.map bitU_of_bit_lifted v.rv_bits
+ (*(Z.of_nat v.rv_start_internal)
+ (v.rv_dir = D_increasing)*)
+
+
+Definition external_slice (d:direction) (start:nat) ((i,j):(nat*nat)) :=
+ match d with
+ (*This is the case the thread/concurrecny model expects, so no change needed*)
+ | D_increasing => (i,j)
+ | D_decreasing => let slice_i = start - i in
+ let slice_j = (i - j) + slice_i in
+ (slice_i,slice_j)
+ end *)
+
+(* TODO
+Definition external_reg_whole r :=
+ Reg (r.name) (Z.to_nat r.start) (Z.to_nat r.size) (dir_of_bool r.is_inc)
+
+Definition external_reg_slice r (i,j) :=
+ let start := Z.to_nat r.start in
+ let dir := dir_of_bool r.is_inc in
+ Reg_slice (r.name) start dir (external_slice dir start (i,j))
+
+Definition external_reg_field_whole reg rfield :=
+ let (m,n) := register_field_indices_nat reg rfield in
+ let start := start_of_reg_nat reg in
+ let dir := dir_of_reg reg in
+ Reg_field (name_of_reg reg) start dir rfield (external_slice dir start (m,n))
+
+Definition external_reg_field_slice reg rfield (i,j) :=
+ let (m,n) := register_field_indices_nat reg rfield in
+ let start := start_of_reg_nat reg in
+ let dir := dir_of_reg reg in
+ Reg_f_slice (name_of_reg reg) start dir rfield
+ (external_slice dir start (m,n))
+ (external_slice dir start (i,j))*)
+
+(*val external_mem_value : list bitU -> memory_value
+Definition external_mem_value v :=
+ byte_lifteds_of_bitv v $> List.reverse
+
+val internal_mem_value : memory_value -> list bitU
+Definition internal_mem_value bytes :=
+ List.reverse bytes $> bitv_of_byte_lifteds*)
+
+
+val foreach : forall a vars.
+ (list a) -> vars -> (a -> vars -> vars) -> vars*)
+Fixpoint foreach {a Vars} (l : list a) (vars : Vars) (body : a -> Vars -> Vars) : Vars :=
+match l with
+| [] => vars
+| (x :: xs) => foreach xs (body x vars) body
+end.
+
+(*declare {isabelle} termination_argument foreach = automatic
+
+val index_list : Z -> Z -> Z -> list Z*)
+Fixpoint index_list' from to step n :=
+ if orb (andb (step >? 0) (from <=? to)) (andb (step <? 0) (to <=? from)) then
+ match n with
+ | O => []
+ | S n => from :: index_list' (from + step) to step n
+ end
+ else [].
+
+Definition index_list from to step :=
+ if orb (andb (step >? 0) (from <=? to)) (andb (step <? 0) (to <=? from)) then
+ index_list' from to step (S (Z.abs_nat (from - to)))
+ else [].
+
+Fixpoint foreach_Z' {Vars} from to step n (vars : Vars) (body : Z -> Vars -> Vars) : Vars :=
+ if orb (andb (step >? 0) (from <=? to)) (andb (step <? 0) (to <=? from)) then
+ match n with
+ | O => vars
+ | S n => let vars := body from vars in foreach_Z' (from + step) to step n vars body
+ end
+ else vars.
+
+Definition foreach_Z {Vars} from to step vars body :=
+ foreach_Z' (Vars := Vars) from to step (S (Z.abs_nat (from - to))) vars body.
+
+Fixpoint foreach_Z_up' {Vars} from to step off n `{ArithFact (0 < step)} `{ArithFact (0 <= off)} (vars : Vars) (body : forall (z : Z) `(ArithFact (from <= z <= to)), Vars -> Vars) {struct n} : Vars :=
+ if sumbool_of_bool (from + off <=? to) then
+ match n with
+ | O => vars
+ | S n => let vars := body (from + off) _ vars in foreach_Z_up' from to step (off + step) n vars body
+ end
+ else vars.
+
+Fixpoint foreach_Z_down' {Vars} from to step off n `{ArithFact (0 < step)} `{ArithFact (off <= 0)} (vars : Vars) (body : forall (z : Z) `(ArithFact (to <= z <= from)), Vars -> Vars) {struct n} : Vars :=
+ if sumbool_of_bool (to <=? from + off) then
+ match n with
+ | O => vars
+ | S n => let vars := body (from + off) _ vars in foreach_Z_down' from to step (off - step) n vars body
+ end
+ else vars.
+
+Definition foreach_Z_up {Vars} from to step vars body `{ArithFact (0 < step)} :=
+ foreach_Z_up' (Vars := Vars) from to step 0 (S (Z.abs_nat (from - to))) vars body.
+Definition foreach_Z_down {Vars} from to step vars body `{ArithFact (0 < step)} :=
+ foreach_Z_down' (Vars := Vars) from to step 0 (S (Z.abs_nat (from - to))) vars body.
+
+(*val while : forall vars. vars -> (vars -> bool) -> (vars -> vars) -> vars
+Fixpoint while vars cond body :=
+ if cond vars then while (body vars) cond body else vars
+
+val until : forall vars. vars -> (vars -> bool) -> (vars -> vars) -> vars
+Fixpoint until vars cond body :=
+ let vars := body vars in
+ if cond vars then vars else until (body vars) cond body
+
+
+Definition assert' b msg_opt :=
+ let msg := match msg_opt with
+ | Some msg => msg
+ | None => "unspecified error"
+ end in
+ if b then () else failwith msg
+
+(* convert numbers unsafely to naturals *)
+
+class (ToNatural a) val toNatural : a -> natural end
+(* eta-expanded for Isabelle output, otherwise it breaks *)
+instance (ToNatural Z) let toNatural := (fun n => naturalFromInteger n) end
+instance (ToNatural int) let toNatural := (fun n => naturalFromInt n) end
+instance (ToNatural nat) let toNatural := (fun n => naturalFromNat n) end
+instance (ToNatural natural) let toNatural := (fun n => n) end
+
+Definition toNaturalFiveTup (n1,n2,n3,n4,n5) :=
+ (toNatural n1,
+ toNatural n2,
+ toNatural n3,
+ toNatural n4,
+ toNatural n5)
+
+(* Let the following types be generated by Sail per spec, using either bitlists
+ or machine words as bitvector representation *)
+(*type regfp :=
+ | RFull of (string)
+ | RSlice of (string * Z * Z)
+ | RSliceBit of (string * Z)
+ | RField of (string * string)
+
+type niafp :=
+ | NIAFP_successor
+ | NIAFP_concrete_address of vector bitU
+ | NIAFP_indirect_address
+
+(* only for MIPS *)
+type diafp :=
+ | DIAFP_none
+ | DIAFP_concrete of vector bitU
+ | DIAFP_reg of regfp
+
+Definition regfp_to_reg (reg_info : string -> option string -> (nat * nat * direction * (nat * nat))) := function
+ | RFull name =>
+ let (start,length,direction,_) := reg_info name None in
+ Reg name start length direction
+ | RSlice (name,i,j) =>
+ let i = Z.to_nat i in
+ let j = Z.to_nat j in
+ let (start,length,direction,_) = reg_info name None in
+ let slice = external_slice direction start (i,j) in
+ Reg_slice name start direction slice
+ | RSliceBit (name,i) =>
+ let i = Z.to_nat i in
+ let (start,length,direction,_) = reg_info name None in
+ let slice = external_slice direction start (i,i) in
+ Reg_slice name start direction slice
+ | RField (name,field_name) =>
+ let (start,length,direction,span) = reg_info name (Some field_name) in
+ let slice = external_slice direction start span in
+ Reg_field name start direction field_name slice
+end
+
+Definition niafp_to_nia reginfo = function
+ | NIAFP_successor => NIA_successor
+ | NIAFP_concrete_address v => NIA_concrete_address (address_of_bitv v)
+ | NIAFP_indirect_address => NIA_indirect_address
+end
+
+Definition diafp_to_dia reginfo = function
+ | DIAFP_none => DIA_none
+ | DIAFP_concrete v => DIA_concrete_address (address_of_bitv v)
+ | DIAFP_reg r => DIA_register (regfp_to_reg reginfo r)
+end
+*)
+*)
+
+(* Arithmetic functions which return proofs that match the expected Sail
+ types in smt.sail. *)
+
+Definition ediv_with_eq n m : {o : Z & ArithFact (o = ZEuclid.div n m)} := build_ex (ZEuclid.div n m).
+Definition emod_with_eq n m : {o : Z & ArithFact (o = ZEuclid.modulo n m)} := build_ex (ZEuclid.modulo n m).
+Definition abs_with_eq n : {o : Z & ArithFact (o = Z.abs n)} := build_ex (Z.abs n).
+
+(* Similarly, for ranges (currently in MIPS) *)
+
+Definition eq_range {n m o p} (l : {l & ArithFact (n <= l <= m)}) (r : {r & ArithFact (o <= r <= p)}) : bool :=
+ (projT1 l) =? (projT1 r).
+Definition add_range {n m o p} (l : {l & ArithFact (n <= l <= m)}) (r : {r & ArithFact (o <= r <= p)})
+ : {x & ArithFact (n+o <= x <= m+p)} :=
+ build_ex ((projT1 l) + (projT1 r)).
+Definition sub_range {n m o p} (l : {l & ArithFact (n <= l <= m)}) (r : {r & ArithFact (o <= r <= p)})
+ : {x & ArithFact (n-p <= x <= m-o)} :=
+ build_ex ((projT1 l) - (projT1 r)).
+Definition negate_range {n m} (l : {l : Z & ArithFact (n <= l <= m)})
+ : {x : Z & ArithFact ((- m) <= x <= (- n))} :=
+ build_ex (- (projT1 l)).
+
+Definition min_atom (a : Z) (b : Z) : {c : Z & ArithFact ((c = a \/ c = b) /\ c <= a /\ c <= b)} :=
+ build_ex (Z.min a b).
+Definition max_atom (a : Z) (b : Z) : {c : Z & ArithFact ((c = a \/ c = b) /\ c >= a /\ c >= b)} :=
+ build_ex (Z.max a b).
+
+
+(*** Generic vectors *)
+
+Definition vec (T:Type) (n:Z) := { l : list T & length_list l = n }.
+Definition vec_length {T n} (v : vec T n) := n.
+Definition vec_access_dec {T n} (v : vec T n) m `{ArithFact (0 <= m < n)} : T :=
+ access_list_dec (projT1 v) m.
+Definition vec_access_inc {T n} (v : vec T n) m `{ArithFact (0 <= m < n)} : T :=
+ access_list_inc (projT1 v) m.
+
+Program Definition vec_init {T} (t : T) (n : Z) `{ArithFact (n >= 0)} : vec T n :=
+ existT _ (repeat [t] n) _.
+Next Obligation.
+rewrite repeat_length; auto using fact.
+unfold length_list.
+simpl.
+auto with zarith.
+Qed.
+
+Definition vec_concat {T m n} (v : vec T m) (w : vec T n) : vec T (m + n).
+refine (existT _ (projT1 v ++ projT1 w) _).
+destruct v.
+destruct w.
+simpl.
+unfold length_list in *.
+rewrite <- e, <- e0.
+rewrite app_length.
+rewrite Nat2Z.inj_add.
+reflexivity.
+Defined.
+
+Lemma skipn_length {A n} {l: list A} : (n <= List.length l -> List.length (skipn n l) = List.length l - n)%nat.
+revert l.
+induction n.
+* simpl. auto with arith.
+* intros l H.
+ destruct l.
+ + inversion H.
+ + simpl in H.
+ simpl.
+ rewrite IHn; auto with arith.
+Qed.
+Lemma update_list_inc_length {T} {l:list T} {m x} : 0 <= m < length_list l -> length_list (update_list_inc l m x) = length_list l.
+unfold update_list_inc, list_update, length_list.
+intro H.
+f_equal.
+assert ((0 <= Z.to_nat m < Datatypes.length l)%nat).
+{ destruct H as [H1 H2].
+ split.
+ + change 0%nat with (Z.to_nat 0).
+ apply Z2Nat.inj_le; auto with zarith.
+ + rewrite <- Nat2Z.id.
+ apply Z2Nat.inj_lt; auto with zarith.
+}
+rewrite app_length.
+rewrite firstn_length_le; only 2:omega.
+cbn -[skipn].
+rewrite skipn_length;
+omega.
+Qed.
+
+Program Definition vec_update_dec {T n} (v : vec T n) m t `{ArithFact (0 <= m < n)} : vec T n := existT _ (update_list_dec (projT1 v) m t) _.
+Next Obligation.
+unfold update_list_dec.
+rewrite update_list_inc_length.
++ destruct v. apply e.
++ destruct H.
+ destruct v. simpl (projT1 _). rewrite e.
+ omega.
+Qed.
+
+Program Definition vec_update_inc {T n} (v : vec T n) m t `{ArithFact (0 <= m < n)} : vec T n := existT _ (update_list_inc (projT1 v) m t) _.
+Next Obligation.
+rewrite update_list_inc_length.
++ destruct v. apply e.
++ destruct H.
+ destruct v. simpl (projT1 _). rewrite e.
+ omega.
+Qed.
+
+Program Definition vec_map {S T} (f : S -> T) {n} (v : vec S n) : vec T n := existT _ (List.map f (projT1 v)) _.
+Next Obligation.
+destruct v as [l H].
+cbn.
+unfold length_list.
+rewrite map_length.
+apply H.
+Qed.
+
+Program Definition just_vec {A n} (v : vec (option A) n) : option (vec A n) :=
+ match just_list (projT1 v) with
+ | None => None
+ | Some v' => Some (existT _ v' _)
+ end.
+Next Obligation.
+rewrite <- (just_list_length_Z _ _ Heq_anonymous).
+destruct v.
+assumption.
+Qed.
+
+Definition list_of_vec {A n} (v : vec A n) : list A := projT1 v.
+
+Definition vec_eq_dec {T n} (D : forall x y : T, {x = y} + {x <> y}) (x y : vec T n) :
+ {x = y} + {x <> y}.
+refine (if List.list_eq_dec D (projT1 x) (projT1 y) then left _ else right _).
+* apply eq_sigT_hprop; auto using ZEqdep.UIP.
+* contradict n0. rewrite n0. reflexivity.
+Defined.
+
+Instance Decidable_eq_vec {T : Type} {n} `(DT : forall x y : T, Decidable (x = y)) :
+ forall x y : vec T n, Decidable (x = y) := {
+ Decidable_witness := proj1_sig (bool_of_sumbool (vec_eq_dec (fun x y => generic_dec x y) x y))
+}.
+destruct (vec_eq_dec _ x y); simpl; split; congruence.
+Defined.
+
+Program Definition vec_of_list {A} n (l : list A) : option (vec A n) :=
+ if sumbool_of_bool (n =? length_list l) then Some (existT _ l _) else None.
+Next Obligation.
+symmetry.
+apply Z.eqb_eq.
+assumption.
+Qed.
+
+Definition vec_of_list_len {A} (l : list A) : vec A (length_list l) := existT _ l (eq_refl _).
+
+Definition map_bind {A B} (f : A -> option B) (a : option A) : option B :=
+match a with
+| Some a' => f a'
+| None => None
+end.
+
+Definition sub_nat (x : Z) `{ArithFact (x >= 0)} (y : Z) `{ArithFact (y >= 0)} :
+ {z : Z & ArithFact (z >= 0)} :=
+ let z := x - y in
+ if sumbool_of_bool (z >=? 0) then build_ex z else build_ex 0.
+
+Definition min_nat (x : Z) `{ArithFact (x >= 0)} (y : Z) `{ArithFact (y >= 0)} :
+ {z : Z & ArithFact (z >= 0)} :=
+ build_ex (Z.min x y).
+
+Definition max_nat (x : Z) `{ArithFact (x >= 0)} (y : Z) `{ArithFact (y >= 0)} :
+ {z : Z & ArithFact (z >= 0)} :=
+ build_ex (Z.max x y).
+
+Definition shl_int_8 (x y : Z) `{HE:ArithFact (x = 8)} `{HR:ArithFact (0 <= y <= 3)}: {z : Z & ArithFact (In z [8;16;32;64])}.
+refine (existT _ (shl_int x y) _).
+destruct HE as [HE].
+destruct HR as [HR].
+assert (H : y = 0 \/ y = 1 \/ y = 2 \/ y = 3) by omega.
+constructor.
+intuition (subst; compute; auto).
+Defined.
+
+Definition shl_int_32 (x y : Z) `{HE:ArithFact (x = 32)} `{HR:ArithFact (In y [0;1])}: {z : Z & ArithFact (In z [32;64])}.
+refine (existT _ (shl_int x y) _).
+destruct HE as [HE].
+destruct HR as [[HR1 | [HR2 | []]]];
+subst; compute;
+auto using Build_ArithFact.
+Defined.
+
+Definition shr_int_32 (x y : Z) `{HE:ArithFact (0 <= x <= 31)} `{HR:ArithFact (y = 1)}: {z : Z & ArithFact (0 <= z <= 15)}.
+refine (existT _ (shr_int x y) _).
+destruct HE as [HE].
+destruct HR as [HR];
+subst.
+unfold shr_int.
+rewrite <- Z.div2_spec.
+constructor.
+rewrite Z.div2_div.
+specialize (Z.div_mod x 2).
+specialize (Z.mod_pos_bound x 2).
+generalize (Z.div x 2).
+generalize (x mod 2).
+intros.
+nia.
+Defined.
+
+Lemma shl_8_ge_0 {n} : shl_int 8 n >= 0.
+unfold shl_int.
+apply Z.le_ge.
+apply <- Z.shiftl_nonneg.
+omega.
+Qed.
+Hint Resolve shl_8_ge_0 : sail.
+
+(* This is needed because Sail's internal constraint language doesn't have
+ < and could disappear if we add it... *)
+
+Lemma sail_lt_ge (x y : Z) :
+ x < y <-> y >= x +1.
+omega.
+Qed.
+Hint Resolve sail_lt_ge : sail.
diff --git a/prover_snapshots/coq/lib/sail/_CoqProject b/prover_snapshots/coq/lib/sail/_CoqProject
new file mode 100644
index 0000000..9f5d26b
--- /dev/null
+++ b/prover_snapshots/coq/lib/sail/_CoqProject
@@ -0,0 +1,2 @@
+-R . Sail
+-R ../../../bbv/theories bbv