diff --git a/thys/Berlekamp_Zassenhaus/Finite_Field_Record_Based.thy b/thys/Berlekamp_Zassenhaus/Finite_Field_Record_Based.thy --- a/thys/Berlekamp_Zassenhaus/Finite_Field_Record_Based.thy +++ b/thys/Berlekamp_Zassenhaus/Finite_Field_Record_Based.thy @@ -1,1657 +1,1657 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) subsection \Finite Fields\ text \We provide four implementations for $GF(p)$ -- the field with $p$ elements for some prime $p$ -- one by int, one by integers, one by 32-bit numbers and one 64-bit implementation. Correctness of the implementations is proven by transfer rules to the type-based version of $GF(p)$.\ theory Finite_Field_Record_Based imports Finite_Field Arithmetic_Record_Based Native_Word.Uint32 Native_Word.Uint64 Native_Word.Code_Target_Bits_Int "HOL-Library.Code_Target_Numeral" begin (* mod on standard case which can immediately be mapped to target languages without considering special cases *) definition mod_nonneg_pos :: "integer \ integer \ integer" where "x \ 0 \ y > 0 \ mod_nonneg_pos x y = (x mod y)" code_printing \ \FIXME illusion of partiality\ constant mod_nonneg_pos \ (SML) "IntInf.mod/ ( _,/ _ )" and (Eval) "IntInf.mod/ ( _,/ _ )" and (OCaml) "Z.rem" and (Haskell) "Prelude.mod/ ( _ )/ ( _ )" and (Scala) "!((k: BigInt) => (l: BigInt) =>/ (k '% l))" definition mod_nonneg_pos_int :: "int \ int \ int" where "mod_nonneg_pos_int x y = int_of_integer (mod_nonneg_pos (integer_of_int x) (integer_of_int y))" lemma mod_nonneg_pos_int[simp]: "x \ 0 \ y > 0 \ mod_nonneg_pos_int x y = (x mod y)" unfolding mod_nonneg_pos_int_def using mod_nonneg_pos_def by simp context fixes p :: int begin definition plus_p :: "int \ int \ int" where "plus_p x y \ let z = x + y in if z \ p then z - p else z" definition minus_p :: "int \ int \ int" where "minus_p x y \ if y \ x then x - y else x + p - y" definition uminus_p :: "int \ int" where "uminus_p x = (if x = 0 then 0 else p - x)" definition mult_p :: "int \ int \ int" where "mult_p x y = (mod_nonneg_pos_int (x * y) p)" fun power_p :: "int \ nat \ int" where "power_p x n = (if n = 0 then 1 else let (d,r) = Divides.divmod_nat n 2; rec = power_p (mult_p x x) d in if r = 0 then rec else mult_p rec x)" text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p :: "int \ int" where "inverse_p x = (if x = 0 then 0 else power_p x (nat (p - 2)))" definition divide_p :: "int \ int \ int" where "divide_p x y = mult_p x (inverse_p y)" definition finite_field_ops_int :: "int arith_ops_record" where "finite_field_ops_int \ Arith_Ops_Record 0 1 plus_p mult_p minus_p uminus_p divide_p inverse_p (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) (\ x . x) (\ x . x) (\ x. 0 \ x \ x < p)" end context fixes p :: uint32 begin definition plus_p32 :: "uint32 \ uint32 \ uint32" where "plus_p32 x y \ let z = x + y in if z \ p then z - p else z" definition minus_p32 :: "uint32 \ uint32 \ uint32" where "minus_p32 x y \ if y \ x then x - y else (x + p) - y" definition uminus_p32 :: "uint32 \ uint32" where "uminus_p32 x = (if x = 0 then 0 else p - x)" definition mult_p32 :: "uint32 \ uint32 \ uint32" where "mult_p32 x y = (x * y mod p)" lemma int_of_uint32_shift: "int_of_uint32 (drop_bit k n) = (int_of_uint32 n) div (2 ^ k)" apply transfer apply transfer apply (simp add: take_bit_drop_bit min_def) apply (simp add: drop_bit_eq_div) done lemma int_of_uint32_0_iff: "int_of_uint32 n = 0 \ n = 0" by (transfer, rule uint_0_iff) lemma int_of_uint32_0: "int_of_uint32 0 = 0" unfolding int_of_uint32_0_iff by simp lemma int_of_uint32_ge_0: "int_of_uint32 n \ 0" by (transfer, auto) lemma two_32: "2 ^ LENGTH(32) = (4294967296 :: int)" by simp lemma int_of_uint32_plus: "int_of_uint32 (x + y) = (int_of_uint32 x + int_of_uint32 y) mod 4294967296" by (transfer, unfold uint_word_ariths two_32, rule refl) lemma int_of_uint32_minus: "int_of_uint32 (x - y) = (int_of_uint32 x - int_of_uint32 y) mod 4294967296" by (transfer, unfold uint_word_ariths two_32, rule refl) lemma int_of_uint32_mult: "int_of_uint32 (x * y) = (int_of_uint32 x * int_of_uint32 y) mod 4294967296" by (transfer, unfold uint_word_ariths two_32, rule refl) lemma int_of_uint32_mod: "int_of_uint32 (x mod y) = (int_of_uint32 x mod int_of_uint32 y)" by (transfer, unfold uint_mod two_32, rule refl) lemma int_of_uint32_inv: "0 \ x \ x < 4294967296 \ int_of_uint32 (uint32_of_int x) = x" - by transfer (simp add: take_bit_int_eq_self) + by transfer (simp add: take_bit_int_eq_self unsigned_of_int) context includes bit_operations_syntax begin function power_p32 :: "uint32 \ uint32 \ uint32" where "power_p32 x n = (if n = 0 then 1 else let rec = power_p32 (mult_p32 x x) (drop_bit 1 n) in if n AND 1 = 0 then rec else mult_p32 rec x)" by pat_completeness auto termination proof - { fix n :: uint32 assume "n \ 0" with int_of_uint32_ge_0[of n] int_of_uint32_0_iff[of n] have "int_of_uint32 n > 0" by auto hence "0 < int_of_uint32 n" "int_of_uint32 n div 2 < int_of_uint32 n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_uint32 n))", auto simp: int_of_uint32_shift *) qed end text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p32 :: "uint32 \ uint32" where "inverse_p32 x = (if x = 0 then 0 else power_p32 x (p - 2))" definition divide_p32 :: "uint32 \ uint32 \ uint32" where "divide_p32 x y = mult_p32 x (inverse_p32 y)" definition finite_field_ops32 :: "uint32 arith_ops_record" where "finite_field_ops32 \ Arith_Ops_Record 0 1 plus_p32 mult_p32 minus_p32 uminus_p32 divide_p32 inverse_p32 (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) uint32_of_int int_of_uint32 (\ x. 0 \ x \ x < p)" end lemma shiftr_uint32_code [code_unfold]: "drop_bit 1 x = (uint32_shiftr x 1)" by (simp add: uint32_shiftr_def) (* ******************************************************************************** *) subsubsection \Transfer Relation\ locale mod_ring_locale = fixes p :: int and ty :: "'a :: nontriv itself" assumes p: "p = int CARD('a)" begin lemma nat_p: "nat p = CARD('a)" unfolding p by simp lemma p2: "p \ 2" unfolding p using nontriv[where 'a = 'a] by auto lemma p2_ident: "int (CARD('a) - 2) = p - 2" using p2 unfolding p by simp definition mod_ring_rel :: "int \ 'a mod_ring \ bool" where "mod_ring_rel x x' = (x = to_int_mod_ring x')" (* domain transfer rules *) lemma Domainp_mod_ring_rel [transfer_domain_rule]: "Domainp (mod_ring_rel) = (\ v. v \ {0 ..< p})" proof - { fix v :: int assume *: "0 \ v" "v < p" have "Domainp mod_ring_rel v" proof show "mod_ring_rel v (of_int_mod_ring v)" unfolding mod_ring_rel_def using * p by auto qed } note * = this show ?thesis by (intro ext iffI, insert range_to_int_mod_ring[where 'a = 'a] *, auto simp: mod_ring_rel_def p) qed (* left/right/bi-unique *) lemma bi_unique_mod_ring_rel [transfer_rule]: "bi_unique mod_ring_rel" "left_unique mod_ring_rel" "right_unique mod_ring_rel" unfolding mod_ring_rel_def bi_unique_def left_unique_def right_unique_def by auto (* left/right-total *) lemma right_total_mod_ring_rel [transfer_rule]: "right_total mod_ring_rel" unfolding mod_ring_rel_def right_total_def by simp (* ************************************************************************************ *) subsubsection \Transfer Rules\ (* 0 / 1 *) lemma mod_ring_0[transfer_rule]: "mod_ring_rel 0 0" unfolding mod_ring_rel_def by simp lemma mod_ring_1[transfer_rule]: "mod_ring_rel 1 1" unfolding mod_ring_rel_def by simp (* addition *) lemma plus_p_mod_def: assumes x: "x \ {0 ..< p}" and y: "y \ {0 ..< p}" shows "plus_p p x y = ((x + y) mod p)" proof (cases "p \ x + y") case False thus ?thesis using x y unfolding plus_p_def Let_def by auto next case True from True x y have *: "p > 0" "0 \ x + y - p" "x + y - p < p" by auto from True have id: "plus_p p x y = x + y - p" unfolding plus_p_def by auto show ?thesis unfolding id using * using mod_pos_pos_trivial by fastforce qed lemma mod_ring_plus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (plus_p p) (+)" proof - { fix x y :: "'a mod_ring" have "plus_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x + y)" by (transfer, subst plus_p_mod_def, auto, auto simp: p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* subtraction *) lemma minus_p_mod_def: assumes x: "x \ {0 ..< p}" and y: "y \ {0 ..< p}" shows "minus_p p x y = ((x - y) mod p)" proof (cases "x - y < 0") case False thus ?thesis using x y unfolding minus_p_def Let_def by auto next case True from True x y have *: "p > 0" "0 \ x - y + p" "x - y + p < p" by auto from True have id: "minus_p p x y = x - y + p" unfolding minus_p_def by auto show ?thesis unfolding id using * using mod_pos_pos_trivial by fastforce qed lemma mod_ring_minus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (minus_p p) (-)" proof - { fix x y :: "'a mod_ring" have "minus_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x - y)" by (transfer, subst minus_p_mod_def, auto simp: p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* unary minus *) lemma mod_ring_uminus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (uminus_p p) uminus" proof - { fix x :: "'a mod_ring" have "uminus_p p (to_int_mod_ring x) = to_int_mod_ring (uminus x)" by (transfer, auto simp: uminus_p_def p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* multiplication *) lemma mod_ring_mult[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (mult_p p) ((*))" proof - { fix x y :: "'a mod_ring" have "mult_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x * y)" by (transfer, auto simp: mult_p_def p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* equality *) lemma mod_ring_eq[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> (=)) (=) (=)" by (intro rel_funI, auto simp: mod_ring_rel_def) (* power *) lemma mod_ring_power[transfer_rule]: "(mod_ring_rel ===> (=) ===> mod_ring_rel) (power_p p) (^)" proof (intro rel_funI, clarify, unfold binary_power[symmetric], goal_cases) fix x y n assume xy: "mod_ring_rel x y" from xy show "mod_ring_rel (power_p p x n) (binary_power y n)" proof (induct y n arbitrary: x rule: binary_power.induct) case (1 x n y) note 1(2)[transfer_rule] show ?case proof (cases "n = 0") case True thus ?thesis by (simp add: mod_ring_1) next case False obtain d r where id: "Divides.divmod_nat n 2 = (d,r)" by force let ?int = "power_p p (mult_p p y y) d" let ?gfp = "binary_power (x * x) d" from False have id': "?thesis = (mod_ring_rel (if r = 0 then ?int else mult_p p ?int y) (if r = 0 then ?gfp else ?gfp * x))" unfolding power_p.simps[of _ _ n] binary_power.simps[of _ n] Let_def id split by simp have [transfer_rule]: "mod_ring_rel ?int ?gfp" by (rule 1(1)[OF False refl id[symmetric]], transfer_prover) show ?thesis unfolding id' by transfer_prover qed qed qed declare power_p.simps[simp del] lemma ring_finite_field_ops_int: "ring_ops (finite_field_ops_int p) mod_ring_rel" by (unfold_locales, auto simp: finite_field_ops_int_def bi_unique_mod_ring_rel right_total_mod_ring_rel mod_ring_plus mod_ring_minus mod_ring_uminus mod_ring_mult mod_ring_eq mod_ring_0 mod_ring_1 Domainp_mod_ring_rel) end locale prime_field = mod_ring_locale p ty for p and ty :: "'a :: prime_card itself" begin lemma prime: "prime p" unfolding p using prime_card[where 'a = 'a] by simp (* mod *) lemma mod_ring_mod[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) ((\ x y. if y = 0 then x else 0)) (mod)" proof - { fix x y :: "'a mod_ring" have "(if to_int_mod_ring y = 0 then to_int_mod_ring x else 0) = to_int_mod_ring (x mod y)" unfolding modulo_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* normalize *) lemma mod_ring_normalize[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) ((\ x. if x = 0 then 0 else 1)) normalize" proof - { fix x :: "'a mod_ring" have "(if to_int_mod_ring x = 0 then 0 else 1) = to_int_mod_ring (normalize x)" unfolding normalize_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* unit_factor *) lemma mod_ring_unit_factor[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (\ x. x) unit_factor" proof - { fix x :: "'a mod_ring" have "to_int_mod_ring x = to_int_mod_ring (unit_factor x)" unfolding unit_factor_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* inverse *) lemma mod_ring_inverse[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (inverse_p p) inverse" proof (intro rel_funI) fix x y assume [transfer_rule]: "mod_ring_rel x y" show "mod_ring_rel (inverse_p p x) (inverse y)" unfolding inverse_p_def inverse_mod_ring_def apply (transfer_prover_start) apply (transfer_step)+ apply (unfold p2_ident) apply (rule refl) done qed (* division *) lemma mod_ring_divide[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (divide_p p) (/)" unfolding divide_p_def[abs_def] divide_mod_ring_def[abs_def] inverse_mod_ring_def[symmetric] by transfer_prover lemma mod_ring_rel_unsafe: assumes "x < CARD('a)" shows "mod_ring_rel (int x) (of_nat x)" "0 < x \ of_nat x \ (0 :: 'a mod_ring)" proof - have id: "of_nat x = (of_int (int x) :: 'a mod_ring)" by simp show "mod_ring_rel (int x) (of_nat x)" "0 < x \ of_nat x \ (0 :: 'a mod_ring)" unfolding id unfolding mod_ring_rel_def proof (auto simp add: assms of_int_of_int_mod_ring) assume "0 < x" with assms have "of_int_mod_ring (int x) \ (0 :: 'a mod_ring)" by (metis (no_types) less_imp_of_nat_less less_irrefl of_nat_0_le_iff of_nat_0_less_iff to_int_mod_ring_hom.hom_zero to_int_mod_ring_of_int_mod_ring) thus "of_int_mod_ring (int x) = (0 :: 'a mod_ring) \ False" by blast qed qed lemma finite_field_ops_int: "field_ops (finite_field_ops_int p) mod_ring_rel" by (unfold_locales, auto simp: finite_field_ops_int_def bi_unique_mod_ring_rel right_total_mod_ring_rel mod_ring_divide mod_ring_plus mod_ring_minus mod_ring_uminus mod_ring_inverse mod_ring_mod mod_ring_unit_factor mod_ring_normalize mod_ring_mult mod_ring_eq mod_ring_0 mod_ring_1 Domainp_mod_ring_rel) end text \Once we have proven the soundness of the implementation, we do not care any longer that @{typ "'a mod_ring"} has been defined internally via lifting. Disabling the transfer-rules will hide the internal definition in further applications of transfer.\ lifting_forget mod_ring.lifting text \For soundness of the 32-bit implementation, we mainly prove that this implementation implements the int-based implementation of the mod-ring.\ context mod_ring_locale begin context fixes pp :: "uint32" assumes ppp: "p = int_of_uint32 pp" and small: "p \ 65535" begin lemmas uint32_simps = int_of_uint32_0 int_of_uint32_plus int_of_uint32_minus int_of_uint32_mult definition urel32 :: "uint32 \ int \ bool" where "urel32 x y = (y = int_of_uint32 x \ y < p)" definition mod_ring_rel32 :: "uint32 \ 'a mod_ring \ bool" where "mod_ring_rel32 x y = (\ z. urel32 x z \ mod_ring_rel z y)" lemma urel32_0: "urel32 0 0" unfolding urel32_def using p2 by (simp, transfer, simp) lemma urel32_1: "urel32 1 1" unfolding urel32_def using p2 by (simp, transfer, simp) lemma le_int_of_uint32: "(x \ y) = (int_of_uint32 x \ int_of_uint32 y)" by (transfer, simp add: word_le_def) lemma urel32_plus: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (plus_p32 pp x x') (plus_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" let ?p = "int_of_uint32 pp" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel32_def by auto have le: "(pp \ x + x') = (?p \ ?x + ?x')" unfolding le_int_of_uint32 using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?p \ ?x + ?x'") case True hence True: "(?p \ ?x + ?x') = True" by simp show ?thesis unfolding id using small rel unfolding plus_p32_def plus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?p \ ?x + ?x') = False" by simp show ?thesis unfolding id using small rel unfolding plus_p32_def plus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_minus: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (minus_p32 pp x x') (minus_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel32_def by auto have le: "(x' \ x) = (?x' \ ?x)" unfolding le_int_of_uint32 using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?x' \ ?x") case True hence True: "(?x' \ ?x) = True" by simp show ?thesis unfolding id using small rel unfolding minus_p32_def minus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?x' \ ?x) = False" by simp show ?thesis unfolding id using small rel unfolding minus_p32_def minus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_uminus: assumes "urel32 x y" shows "urel32 (uminus_p32 pp x) (uminus_p p y)" proof - let ?x = "int_of_uint32 x" from assms int_of_uint32_ge_0 have id: "y = ?x" and rel: "0 \ ?x" "?x < p" unfolding urel32_def by auto have le: "(x = 0) = (?x = 0)" unfolding int_of_uint32_0_iff using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?x = 0") case True hence True: "(?x = 0) = True" by simp show ?thesis unfolding id using small rel unfolding uminus_p32_def uminus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?x = 0) = False" by simp show ?thesis unfolding id using small rel unfolding uminus_p32_def uminus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_mult: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (mult_p32 pp x x') (mult_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' < p" unfolding urel32_def by auto from rel have "?x * ?x' < p * p" by (metis mult_strict_mono') also have "\ \ 65536 * 65536" by (rule mult_mono, insert p2 small, auto) finally have le: "?x * ?x' < 4294967296" by simp show ?thesis unfolding id using small rel unfolding mult_p32_def mult_p_def Let_def urel32_def unfolding ppp by (auto simp: uint32_simps, unfold int_of_uint32_mod int_of_uint32_mult, subst mod_pos_pos_trivial[of _ 4294967296], insert le, auto) qed lemma urel32_eq: assumes "urel32 x y" "urel32 x' y'" shows "(x = x') = (y = y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" unfolding urel32_def by auto show ?thesis unfolding id by (transfer, transfer) rule qed lemma urel32_normalize: assumes x: "urel32 x y" shows "urel32 (if x = 0 then 0 else 1) (if y = 0 then 0 else 1)" unfolding urel32_eq[OF x urel32_0] using urel32_0 urel32_1 by auto lemma urel32_mod: assumes x: "urel32 x x'" and y: "urel32 y y'" shows "urel32 (if y = 0 then x else 0) (if y' = 0 then x' else 0)" unfolding urel32_eq[OF y urel32_0] using urel32_0 x by auto lemma urel32_power: "urel32 x x' \ urel32 y (int y') \ urel32 (power_p32 pp x y) (power_p p x' y')" including bit_operations_syntax proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) case (1 x' y' x y) note x = 1(2) note y = 1(3) show ?case proof (cases "y' = 0") case True hence y: "y = 0" using urel32_eq[OF y urel32_0] by auto show ?thesis unfolding y True by (simp add: power_p.simps urel32_1) next case False hence id: "(y = 0) = False" "(y' = 0) = False" using urel32_eq[OF y urel32_0] by auto from y have \int y' = int_of_uint32 y\ \int y' < p\ by (simp_all add: urel32_def) obtain d' r' where dr': "Divides.divmod_nat y' 2 = (d',r')" by force from divmod_nat_def[of y' 2, unfolded dr'] have r': "r' = y' mod 2" and d': "d' = y' div 2" by auto have "urel32 (y AND 1) r'" using \int y' < p\ small apply (simp add: urel32_def and_one_eq r') apply (auto simp add: ppp and_one_eq) apply (simp add: of_nat_mod int_of_uint32.rep_eq modulo_uint32.rep_eq uint_mod \int y' = int_of_uint32 y\) done from urel32_eq[OF this urel32_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel32 (drop_bit 1 y) (int d')" unfolding d' using y unfolding urel32_def using small unfolding ppp apply transfer apply transfer apply (auto simp add: drop_bit_Suc take_bit_int_eq_self) done note IH = 1(1)[OF False refl dr'[symmetric] urel32_mult[OF x x] div] show ?thesis unfolding power_p.simps[of _ _ "y'"] power_p32.simps[of _ _ y] dr' id if_False rem using IH urel32_mult[OF IH x] by (auto simp: Let_def) qed qed lemma urel32_inverse: assumes x: "urel32 x x'" shows "urel32 (inverse_p32 pp x) (inverse_p p x')" proof - have p: "urel32 (pp - 2) (int (nat (p - 2)))" using p2 small unfolding urel32_def unfolding ppp by (simp add: int_of_uint32.rep_eq minus_uint32.rep_eq uint_sub_if') show ?thesis unfolding inverse_p32_def inverse_p_def urel32_eq[OF x urel32_0] using urel32_0 urel32_power[OF x p] by auto qed lemma mod_ring_0_32: "mod_ring_rel32 0 0" using urel32_0 mod_ring_0 unfolding mod_ring_rel32_def by blast lemma mod_ring_1_32: "mod_ring_rel32 1 1" using urel32_1 mod_ring_1 unfolding mod_ring_rel32_def by blast lemma mod_ring_uminus32: "(mod_ring_rel32 ===> mod_ring_rel32) (uminus_p32 pp) uminus" using urel32_uminus mod_ring_uminus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_plus32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (plus_p32 pp) (+)" using urel32_plus mod_ring_plus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_minus32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (minus_p32 pp) (-)" using urel32_minus mod_ring_minus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_mult32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (mult_p32 pp) ((*))" using urel32_mult mod_ring_mult unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_eq32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> (=)) (=) (=)" using urel32_eq mod_ring_eq unfolding mod_ring_rel32_def rel_fun_def by blast lemma urel32_inj: "urel32 x y \ urel32 x z \ y = z" using urel32_eq[of x y x z] by auto lemma urel32_inj': "urel32 x z \ urel32 y z \ x = y" using urel32_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel32: "bi_unique mod_ring_rel32" "left_unique mod_ring_rel32" "right_unique mod_ring_rel32" using bi_unique_mod_ring_rel urel32_inj' unfolding mod_ring_rel32_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel32_def) lemma right_total_mod_ring_rel32: "right_total mod_ring_rel32" unfolding mod_ring_rel32_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel32 (uint32_of_int z) z" unfolding urel32_def using small unfolding ppp by (auto simp: int_of_uint32_inv) with zy show "\ x z. urel32 x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel32: "Domainp mod_ring_rel32 = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel32 x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel32_def proof let ?i = "int_of_uint32" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" using small unfolding ppp by (transfer, auto simp: word_less_def) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel32 a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel32 x (?i x)" unfolding urel32_def using small * unfolding ppp by (transfer, auto simp: word_less_def) qed next assume "\a b. x = a \ (\z. urel32 a z \ mod_ring_rel z b)" then obtain b z where xz: "urel32 x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel32_def using small unfolding ppp by (transfer, auto simp: word_less_def) qed qed lemma ring_finite_field_ops32: "ring_ops (finite_field_ops32 pp) mod_ring_rel32" by (unfold_locales, auto simp: finite_field_ops32_def bi_unique_mod_ring_rel32 right_total_mod_ring_rel32 mod_ring_plus32 mod_ring_minus32 mod_ring_uminus32 mod_ring_mult32 mod_ring_eq32 mod_ring_0_32 mod_ring_1_32 Domainp_mod_ring_rel32) end end context prime_field begin context fixes pp :: "uint32" assumes *: "p = int_of_uint32 pp" "p \ 65535" begin lemma mod_ring_normalize32: "(mod_ring_rel32 ===> mod_ring_rel32) (\x. if x = 0 then 0 else 1) normalize" using urel32_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_mod32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (\x y. if y = 0 then x else 0) (mod)" using urel32_mod[OF *] mod_ring_mod unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor32: "(mod_ring_rel32 ===> mod_ring_rel32) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_inverse32: "(mod_ring_rel32 ===> mod_ring_rel32) (inverse_p32 pp) inverse" using urel32_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_divide32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (divide_p32 pp) (/)" using mod_ring_inverse32 mod_ring_mult32[OF *] unfolding divide_p32_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops32: "field_ops (finite_field_ops32 pp) mod_ring_rel32" by (unfold_locales, insert ring_finite_field_ops32[OF *], auto simp: ring_ops_def finite_field_ops32_def mod_ring_divide32 mod_ring_inverse32 mod_ring_mod32 mod_ring_normalize32) end end (* now there is 64-bit time *) context fixes p :: uint64 begin definition plus_p64 :: "uint64 \ uint64 \ uint64" where "plus_p64 x y \ let z = x + y in if z \ p then z - p else z" definition minus_p64 :: "uint64 \ uint64 \ uint64" where "minus_p64 x y \ if y \ x then x - y else (x + p) - y" definition uminus_p64 :: "uint64 \ uint64" where "uminus_p64 x = (if x = 0 then 0 else p - x)" definition mult_p64 :: "uint64 \ uint64 \ uint64" where "mult_p64 x y = (x * y mod p)" lemma int_of_uint64_shift: "int_of_uint64 (drop_bit k n) = (int_of_uint64 n) div (2 ^ k)" apply transfer apply transfer apply (simp add: take_bit_drop_bit min_def) apply (simp add: drop_bit_eq_div) done lemma int_of_uint64_0_iff: "int_of_uint64 n = 0 \ n = 0" by (transfer, rule uint_0_iff) lemma int_of_uint64_0: "int_of_uint64 0 = 0" unfolding int_of_uint64_0_iff by simp lemma int_of_uint64_ge_0: "int_of_uint64 n \ 0" by (transfer, auto) lemma two_64: "2 ^ LENGTH(64) = (18446744073709551616 :: int)" by simp lemma int_of_uint64_plus: "int_of_uint64 (x + y) = (int_of_uint64 x + int_of_uint64 y) mod 18446744073709551616" by (transfer, unfold uint_word_ariths two_64, rule refl) lemma int_of_uint64_minus: "int_of_uint64 (x - y) = (int_of_uint64 x - int_of_uint64 y) mod 18446744073709551616" by (transfer, unfold uint_word_ariths two_64, rule refl) lemma int_of_uint64_mult: "int_of_uint64 (x * y) = (int_of_uint64 x * int_of_uint64 y) mod 18446744073709551616" by (transfer, unfold uint_word_ariths two_64, rule refl) lemma int_of_uint64_mod: "int_of_uint64 (x mod y) = (int_of_uint64 x mod int_of_uint64 y)" by (transfer, unfold uint_mod two_64, rule refl) lemma int_of_uint64_inv: "0 \ x \ x < 18446744073709551616 \ int_of_uint64 (uint64_of_int x) = x" - by transfer (simp add: take_bit_int_eq_self) + by transfer (simp add: take_bit_int_eq_self unsigned_of_int) context includes bit_operations_syntax begin function power_p64 :: "uint64 \ uint64 \ uint64" where "power_p64 x n = (if n = 0 then 1 else let rec = power_p64 (mult_p64 x x) (drop_bit 1 n) in if n AND 1 = 0 then rec else mult_p64 rec x)" by pat_completeness auto termination proof - { fix n :: uint64 assume "n \ 0" with int_of_uint64_ge_0[of n] int_of_uint64_0_iff[of n] have "int_of_uint64 n > 0" by auto hence "0 < int_of_uint64 n" "int_of_uint64 n div 2 < int_of_uint64 n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_uint64 n))", auto simp: int_of_uint64_shift *) qed end text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p64 :: "uint64 \ uint64" where "inverse_p64 x = (if x = 0 then 0 else power_p64 x (p - 2))" definition divide_p64 :: "uint64 \ uint64 \ uint64" where "divide_p64 x y = mult_p64 x (inverse_p64 y)" definition finite_field_ops64 :: "uint64 arith_ops_record" where "finite_field_ops64 \ Arith_Ops_Record 0 1 plus_p64 mult_p64 minus_p64 uminus_p64 divide_p64 inverse_p64 (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) uint64_of_int int_of_uint64 (\ x. 0 \ x \ x < p)" end lemma shiftr_uint64_code [code_unfold]: "drop_bit 1 x = (uint64_shiftr x 1)" by (simp add: uint64_shiftr_def) text \For soundness of the 64-bit implementation, we mainly prove that this implementation implements the int-based implementation of GF(p).\ context mod_ring_locale begin context fixes pp :: "uint64" assumes ppp: "p = int_of_uint64 pp" and small: "p \ 4294967295" begin lemmas uint64_simps = int_of_uint64_0 int_of_uint64_plus int_of_uint64_minus int_of_uint64_mult definition urel64 :: "uint64 \ int \ bool" where "urel64 x y = (y = int_of_uint64 x \ y < p)" definition mod_ring_rel64 :: "uint64 \ 'a mod_ring \ bool" where "mod_ring_rel64 x y = (\ z. urel64 x z \ mod_ring_rel z y)" lemma urel64_0: "urel64 0 0" unfolding urel64_def using p2 by (simp, transfer, simp) lemma urel64_1: "urel64 1 1" unfolding urel64_def using p2 by (simp, transfer, simp) lemma le_int_of_uint64: "(x \ y) = (int_of_uint64 x \ int_of_uint64 y)" by (transfer, simp add: word_le_def) lemma urel64_plus: assumes "urel64 x y" "urel64 x' y'" shows "urel64 (plus_p64 pp x x') (plus_p p y y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" let ?p = "int_of_uint64 pp" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel64_def by auto have le: "(pp \ x + x') = (?p \ ?x + ?x')" unfolding le_int_of_uint64 using rel small by (auto simp: uint64_simps) show ?thesis proof (cases "?p \ ?x + ?x'") case True hence True: "(?p \ ?x + ?x') = True" by simp show ?thesis unfolding id using small rel unfolding plus_p64_def plus_p_def Let_def urel64_def unfolding ppp le True if_True using True by (auto simp: uint64_simps) next case False hence False: "(?p \ ?x + ?x') = False" by simp show ?thesis unfolding id using small rel unfolding plus_p64_def plus_p_def Let_def urel64_def unfolding ppp le False if_False using False by (auto simp: uint64_simps) qed qed lemma urel64_minus: assumes "urel64 x y" "urel64 x' y'" shows "urel64 (minus_p64 pp x x') (minus_p p y y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel64_def by auto have le: "(x' \ x) = (?x' \ ?x)" unfolding le_int_of_uint64 using rel small by (auto simp: uint64_simps) show ?thesis proof (cases "?x' \ ?x") case True hence True: "(?x' \ ?x) = True" by simp show ?thesis unfolding id using small rel unfolding minus_p64_def minus_p_def Let_def urel64_def unfolding ppp le True if_True using True by (auto simp: uint64_simps) next case False hence False: "(?x' \ ?x) = False" by simp show ?thesis unfolding id using small rel unfolding minus_p64_def minus_p_def Let_def urel64_def unfolding ppp le False if_False using False by (auto simp: uint64_simps) qed qed lemma urel64_uminus: assumes "urel64 x y" shows "urel64 (uminus_p64 pp x) (uminus_p p y)" proof - let ?x = "int_of_uint64 x" from assms int_of_uint64_ge_0 have id: "y = ?x" and rel: "0 \ ?x" "?x < p" unfolding urel64_def by auto have le: "(x = 0) = (?x = 0)" unfolding int_of_uint64_0_iff using rel small by (auto simp: uint64_simps) show ?thesis proof (cases "?x = 0") case True hence True: "(?x = 0) = True" by simp show ?thesis unfolding id using small rel unfolding uminus_p64_def uminus_p_def Let_def urel64_def unfolding ppp le True if_True using True by (auto simp: uint64_simps) next case False hence False: "(?x = 0) = False" by simp show ?thesis unfolding id using small rel unfolding uminus_p64_def uminus_p_def Let_def urel64_def unfolding ppp le False if_False using False by (auto simp: uint64_simps) qed qed lemma urel64_mult: assumes "urel64 x y" "urel64 x' y'" shows "urel64 (mult_p64 pp x x') (mult_p p y y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' < p" unfolding urel64_def by auto from rel have "?x * ?x' < p * p" by (metis mult_strict_mono') also have "\ \ 4294967296 * 4294967296" by (rule mult_mono, insert p2 small, auto) finally have le: "?x * ?x' < 18446744073709551616" by simp show ?thesis unfolding id using small rel unfolding mult_p64_def mult_p_def Let_def urel64_def unfolding ppp by (auto simp: uint64_simps, unfold int_of_uint64_mod int_of_uint64_mult, subst mod_pos_pos_trivial[of _ 18446744073709551616], insert le, auto) qed lemma urel64_eq: assumes "urel64 x y" "urel64 x' y'" shows "(x = x') = (y = y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" unfolding urel64_def by auto show ?thesis unfolding id by (transfer, transfer) rule qed lemma urel64_normalize: assumes x: "urel64 x y" shows "urel64 (if x = 0 then 0 else 1) (if y = 0 then 0 else 1)" unfolding urel64_eq[OF x urel64_0] using urel64_0 urel64_1 by auto lemma urel64_mod: assumes x: "urel64 x x'" and y: "urel64 y y'" shows "urel64 (if y = 0 then x else 0) (if y' = 0 then x' else 0)" unfolding urel64_eq[OF y urel64_0] using urel64_0 x by auto lemma urel64_power: "urel64 x x' \ urel64 y (int y') \ urel64 (power_p64 pp x y) (power_p p x' y')" including bit_operations_syntax proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) case (1 x' y' x y) note x = 1(2) note y = 1(3) show ?case proof (cases "y' = 0") case True hence y: "y = 0" using urel64_eq[OF y urel64_0] by auto show ?thesis unfolding y True by (simp add: power_p.simps urel64_1) next case False hence id: "(y = 0) = False" "(y' = 0) = False" using urel64_eq[OF y urel64_0] by auto from y have \int y' = int_of_uint64 y\ \int y' < p\ by (simp_all add: urel64_def) obtain d' r' where dr': "Divides.divmod_nat y' 2 = (d',r')" by force from divmod_nat_def[of y' 2, unfolded dr'] have r': "r' = y' mod 2" and d': "d' = y' div 2" by auto have "urel64 (y AND 1) r'" using \int y' < p\ small apply (simp add: urel64_def and_one_eq r') apply (auto simp add: ppp and_one_eq) apply (simp add: of_nat_mod int_of_uint64.rep_eq modulo_uint64.rep_eq uint_mod \int y' = int_of_uint64 y\) done from urel64_eq[OF this urel64_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel64 (drop_bit 1 y) (int d')" unfolding d' using y unfolding urel64_def using small unfolding ppp apply transfer apply transfer apply (auto simp add: drop_bit_Suc take_bit_int_eq_self) done note IH = 1(1)[OF False refl dr'[symmetric] urel64_mult[OF x x] div] show ?thesis unfolding power_p.simps[of _ _ "y'"] power_p64.simps[of _ _ y] dr' id if_False rem using IH urel64_mult[OF IH x] by (auto simp: Let_def) qed qed lemma urel64_inverse: assumes x: "urel64 x x'" shows "urel64 (inverse_p64 pp x) (inverse_p p x')" proof - have p: "urel64 (pp - 2) (int (nat (p - 2)))" using p2 small unfolding urel64_def unfolding ppp by (simp add: int_of_uint64.rep_eq minus_uint64.rep_eq uint_sub_if') show ?thesis unfolding inverse_p64_def inverse_p_def urel64_eq[OF x urel64_0] using urel64_0 urel64_power[OF x p] by auto qed lemma mod_ring_0_64: "mod_ring_rel64 0 0" using urel64_0 mod_ring_0 unfolding mod_ring_rel64_def by blast lemma mod_ring_1_64: "mod_ring_rel64 1 1" using urel64_1 mod_ring_1 unfolding mod_ring_rel64_def by blast lemma mod_ring_uminus64: "(mod_ring_rel64 ===> mod_ring_rel64) (uminus_p64 pp) uminus" using urel64_uminus mod_ring_uminus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_plus64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (plus_p64 pp) (+)" using urel64_plus mod_ring_plus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_minus64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (minus_p64 pp) (-)" using urel64_minus mod_ring_minus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_mult64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (mult_p64 pp) ((*))" using urel64_mult mod_ring_mult unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_eq64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> (=)) (=) (=)" using urel64_eq mod_ring_eq unfolding mod_ring_rel64_def rel_fun_def by blast lemma urel64_inj: "urel64 x y \ urel64 x z \ y = z" using urel64_eq[of x y x z] by auto lemma urel64_inj': "urel64 x z \ urel64 y z \ x = y" using urel64_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel64: "bi_unique mod_ring_rel64" "left_unique mod_ring_rel64" "right_unique mod_ring_rel64" using bi_unique_mod_ring_rel urel64_inj' unfolding mod_ring_rel64_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel64_def) lemma right_total_mod_ring_rel64: "right_total mod_ring_rel64" unfolding mod_ring_rel64_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel64 (uint64_of_int z) z" unfolding urel64_def using small unfolding ppp by (auto simp: int_of_uint64_inv) with zy show "\ x z. urel64 x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel64: "Domainp mod_ring_rel64 = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel64 x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel64_def proof let ?i = "int_of_uint64" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" using small unfolding ppp by (transfer, auto simp: word_less_def) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel64 a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel64 x (?i x)" unfolding urel64_def using small * unfolding ppp by (transfer, auto simp: word_less_def) qed next assume "\a b. x = a \ (\z. urel64 a z \ mod_ring_rel z b)" then obtain b z where xz: "urel64 x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel64_def using small unfolding ppp by (transfer, auto simp: word_less_def) qed qed lemma ring_finite_field_ops64: "ring_ops (finite_field_ops64 pp) mod_ring_rel64" by (unfold_locales, auto simp: finite_field_ops64_def bi_unique_mod_ring_rel64 right_total_mod_ring_rel64 mod_ring_plus64 mod_ring_minus64 mod_ring_uminus64 mod_ring_mult64 mod_ring_eq64 mod_ring_0_64 mod_ring_1_64 Domainp_mod_ring_rel64) end end context prime_field begin context fixes pp :: "uint64" assumes *: "p = int_of_uint64 pp" "p \ 4294967295" begin lemma mod_ring_normalize64: "(mod_ring_rel64 ===> mod_ring_rel64) (\x. if x = 0 then 0 else 1) normalize" using urel64_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_mod64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (\x y. if y = 0 then x else 0) (mod)" using urel64_mod[OF *] mod_ring_mod unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor64: "(mod_ring_rel64 ===> mod_ring_rel64) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_inverse64: "(mod_ring_rel64 ===> mod_ring_rel64) (inverse_p64 pp) inverse" using urel64_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_divide64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (divide_p64 pp) (/)" using mod_ring_inverse64 mod_ring_mult64[OF *] unfolding divide_p64_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops64: "field_ops (finite_field_ops64 pp) mod_ring_rel64" by (unfold_locales, insert ring_finite_field_ops64[OF *], auto simp: ring_ops_def finite_field_ops64_def mod_ring_divide64 mod_ring_inverse64 mod_ring_mod64 mod_ring_normalize64) end end (* and a final implementation via integer *) context fixes p :: integer begin definition plus_p_integer :: "integer \ integer \ integer" where "plus_p_integer x y \ let z = x + y in if z \ p then z - p else z" definition minus_p_integer :: "integer \ integer \ integer" where "minus_p_integer x y \ if y \ x then x - y else (x + p) - y" definition uminus_p_integer :: "integer \ integer" where "uminus_p_integer x = (if x = 0 then 0 else p - x)" definition mult_p_integer :: "integer \ integer \ integer" where "mult_p_integer x y = (x * y mod p)" lemma int_of_integer_0_iff: "int_of_integer n = 0 \ n = 0" using integer_eqI by auto lemma int_of_integer_0: "int_of_integer 0 = 0" unfolding int_of_integer_0_iff by simp lemma int_of_integer_plus: "int_of_integer (x + y) = (int_of_integer x + int_of_integer y)" by simp lemma int_of_integer_minus: "int_of_integer (x - y) = (int_of_integer x - int_of_integer y)" by simp lemma int_of_integer_mult: "int_of_integer (x * y) = (int_of_integer x * int_of_integer y)" by simp lemma int_of_integer_mod: "int_of_integer (x mod y) = (int_of_integer x mod int_of_integer y)" by simp lemma int_of_integer_inv: "int_of_integer (integer_of_int x) = x" by simp lemma int_of_integer_shift: "int_of_integer (drop_bit k n) = (int_of_integer n) div (2 ^ k)" by transfer (simp add: int_of_integer_pow shiftr_integer_conv_div_pow2) context includes bit_operations_syntax begin function power_p_integer :: "integer \ integer \ integer" where "power_p_integer x n = (if n \ 0 then 1 else let rec = power_p_integer (mult_p_integer x x) (drop_bit 1 n) in if n AND 1 = 0 then rec else mult_p_integer rec x)" by pat_completeness auto termination proof - { fix n :: integer assume "\ (n \ 0)" hence "n > 0" by auto hence "int_of_integer n > 0" by (simp add: less_integer.rep_eq) hence "0 < int_of_integer n" "int_of_integer n div 2 < int_of_integer n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_integer n))", auto simp: * int_of_integer_shift) qed end text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p_integer :: "integer \ integer" where "inverse_p_integer x = (if x = 0 then 0 else power_p_integer x (p - 2))" definition divide_p_integer :: "integer \ integer \ integer" where "divide_p_integer x y = mult_p_integer x (inverse_p_integer y)" definition finite_field_ops_integer :: "integer arith_ops_record" where "finite_field_ops_integer \ Arith_Ops_Record 0 1 plus_p_integer mult_p_integer minus_p_integer uminus_p_integer divide_p_integer inverse_p_integer (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) integer_of_int int_of_integer (\ x. 0 \ x \ x < p)" end lemma shiftr_integer_code [code_unfold]: "drop_bit 1 x = (integer_shiftr x 1)" unfolding shiftr_integer_code using integer_of_nat_1 by auto text \For soundness of the integer implementation, we mainly prove that this implementation implements the int-based implementation of GF(p).\ context mod_ring_locale begin context fixes pp :: "integer" assumes ppp: "p = int_of_integer pp" begin lemmas integer_simps = int_of_integer_0 int_of_integer_plus int_of_integer_minus int_of_integer_mult definition urel_integer :: "integer \ int \ bool" where "urel_integer x y = (y = int_of_integer x \ y \ 0 \ y < p)" definition mod_ring_rel_integer :: "integer \ 'a mod_ring \ bool" where "mod_ring_rel_integer x y = (\ z. urel_integer x z \ mod_ring_rel z y)" lemma urel_integer_0: "urel_integer 0 0" unfolding urel_integer_def using p2 by simp lemma urel_integer_1: "urel_integer 1 1" unfolding urel_integer_def using p2 by simp lemma le_int_of_integer: "(x \ y) = (int_of_integer x \ int_of_integer y)" by (rule less_eq_integer.rep_eq) lemma urel_integer_plus: assumes "urel_integer x y" "urel_integer x' y'" shows "urel_integer (plus_p_integer pp x x') (plus_p p y y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" let ?p = "int_of_integer pp" from assms have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel_integer_def by auto have le: "(pp \ x + x') = (?p \ ?x + ?x')" unfolding le_int_of_integer using rel by auto show ?thesis proof (cases "?p \ ?x + ?x'") case True hence True: "(?p \ ?x + ?x') = True" by simp show ?thesis unfolding id using rel unfolding plus_p_integer_def plus_p_def Let_def urel_integer_def unfolding ppp le True if_True using True by auto next case False hence False: "(?p \ ?x + ?x') = False" by simp show ?thesis unfolding id using rel unfolding plus_p_integer_def plus_p_def Let_def urel_integer_def unfolding ppp le False if_False using False by auto qed qed lemma urel_integer_minus: assumes "urel_integer x y" "urel_integer x' y'" shows "urel_integer (minus_p_integer pp x x') (minus_p p y y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" from assms have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel_integer_def by auto have le: "(x' \ x) = (?x' \ ?x)" unfolding le_int_of_integer using rel by auto show ?thesis proof (cases "?x' \ ?x") case True hence True: "(?x' \ ?x) = True" by simp show ?thesis unfolding id using rel unfolding minus_p_integer_def minus_p_def Let_def urel_integer_def unfolding ppp le True if_True using True by auto next case False hence False: "(?x' \ ?x) = False" by simp show ?thesis unfolding id using rel unfolding minus_p_integer_def minus_p_def Let_def urel_integer_def unfolding ppp le False if_False using False by auto qed qed lemma urel_integer_uminus: assumes "urel_integer x y" shows "urel_integer (uminus_p_integer pp x) (uminus_p p y)" proof - let ?x = "int_of_integer x" from assms have id: "y = ?x" and rel: "0 \ ?x" "?x < p" unfolding urel_integer_def by auto have le: "(x = 0) = (?x = 0)" unfolding int_of_integer_0_iff using rel by auto show ?thesis proof (cases "?x = 0") case True hence True: "(?x = 0) = True" by simp show ?thesis unfolding id using rel unfolding uminus_p_integer_def uminus_p_def Let_def urel_integer_def unfolding ppp le True if_True using True by auto next case False hence False: "(?x = 0) = False" by simp show ?thesis unfolding id using rel unfolding uminus_p_integer_def uminus_p_def Let_def urel_integer_def unfolding ppp le False if_False using False by auto qed qed lemma pp_pos: "int_of_integer pp > 0" using ppp nontriv[where 'a = 'a] unfolding p by (simp add: less_integer.rep_eq) lemma urel_integer_mult: assumes "urel_integer x y" "urel_integer x' y'" shows "urel_integer (mult_p_integer pp x x') (mult_p p y y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" from assms have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' < p" unfolding urel_integer_def by auto from rel(1,3) have xx: "0 \ ?x * ?x'" by simp show ?thesis unfolding id using rel unfolding mult_p_integer_def mult_p_def Let_def urel_integer_def unfolding ppp mod_nonneg_pos_int[OF xx pp_pos] using xx pp_pos by simp qed lemma urel_integer_eq: assumes "urel_integer x y" "urel_integer x' y'" shows "(x = x') = (y = y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" from assms have id: "y = ?x" "y' = ?x'" unfolding urel_integer_def by auto show ?thesis unfolding id integer_eq_iff .. qed lemma urel_integer_normalize: assumes x: "urel_integer x y" shows "urel_integer (if x = 0 then 0 else 1) (if y = 0 then 0 else 1)" unfolding urel_integer_eq[OF x urel_integer_0] using urel_integer_0 urel_integer_1 by auto lemma urel_integer_mod: assumes x: "urel_integer x x'" and y: "urel_integer y y'" shows "urel_integer (if y = 0 then x else 0) (if y' = 0 then x' else 0)" unfolding urel_integer_eq[OF y urel_integer_0] using urel_integer_0 x by auto lemma urel_integer_power: "urel_integer x x' \ urel_integer y (int y') \ urel_integer (power_p_integer pp x y) (power_p p x' y')" including bit_operations_syntax proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) case (1 x' y' x y) note x = 1(2) note y = 1(3) show ?case proof (cases "y' \ 0") case True hence y: "y = 0" "y' = 0" using urel_integer_eq[OF y urel_integer_0] by auto show ?thesis unfolding y True by (simp add: power_p.simps urel_integer_1) next case False hence id: "(y \ 0) = False" "(y' = 0) = False" using False y by (auto simp add: urel_integer_def not_le) (metis of_int_integer_of of_int_of_nat_eq of_nat_0_less_iff) obtain d' r' where dr': "Divides.divmod_nat y' 2 = (d',r')" by force from divmod_nat_def[of y' 2, unfolded dr'] have r': "r' = y' mod 2" and d': "d' = y' div 2" by auto have aux: "\ y'. int (y' mod 2) = int y' mod 2" by presburger have "urel_integer (y AND 1) r'" unfolding r' using y unfolding urel_integer_def unfolding ppp apply (auto simp add: and_one_eq) apply (simp add: of_nat_mod) done from urel_integer_eq[OF this urel_integer_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel_integer (drop_bit 1 y) (int d')" unfolding d' using y unfolding urel_integer_def unfolding ppp shiftr_integer_conv_div_pow2 by auto from id have "y' \ 0" by auto note IH = 1(1)[OF this refl dr'[symmetric] urel_integer_mult[OF x x] div] show ?thesis unfolding power_p.simps[of _ _ "y'"] power_p_integer.simps[of _ _ y] dr' id if_False rem using IH urel_integer_mult[OF IH x] by (auto simp: Let_def) qed qed lemma urel_integer_inverse: assumes x: "urel_integer x x'" shows "urel_integer (inverse_p_integer pp x) (inverse_p p x')" proof - have p: "urel_integer (pp - 2) (int (nat (p - 2)))" using p2 unfolding urel_integer_def unfolding ppp by auto show ?thesis unfolding inverse_p_integer_def inverse_p_def urel_integer_eq[OF x urel_integer_0] using urel_integer_0 urel_integer_power[OF x p] by auto qed lemma mod_ring_0__integer: "mod_ring_rel_integer 0 0" using urel_integer_0 mod_ring_0 unfolding mod_ring_rel_integer_def by blast lemma mod_ring_1__integer: "mod_ring_rel_integer 1 1" using urel_integer_1 mod_ring_1 unfolding mod_ring_rel_integer_def by blast lemma mod_ring_uminus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (uminus_p_integer pp) uminus" using urel_integer_uminus mod_ring_uminus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_plus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (plus_p_integer pp) (+)" using urel_integer_plus mod_ring_plus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_minus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (minus_p_integer pp) (-)" using urel_integer_minus mod_ring_minus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_mult_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (mult_p_integer pp) ((*))" using urel_integer_mult mod_ring_mult unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_eq_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> (=)) (=) (=)" using urel_integer_eq mod_ring_eq unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma urel_integer_inj: "urel_integer x y \ urel_integer x z \ y = z" using urel_integer_eq[of x y x z] by auto lemma urel_integer_inj': "urel_integer x z \ urel_integer y z \ x = y" using urel_integer_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel_integer: "bi_unique mod_ring_rel_integer" "left_unique mod_ring_rel_integer" "right_unique mod_ring_rel_integer" using bi_unique_mod_ring_rel urel_integer_inj' unfolding mod_ring_rel_integer_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel_integer_def) lemma right_total_mod_ring_rel_integer: "right_total mod_ring_rel_integer" unfolding mod_ring_rel_integer_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel_integer (integer_of_int z) z" unfolding urel_integer_def unfolding ppp by auto with zy show "\ x z. urel_integer x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel_integer: "Domainp mod_ring_rel_integer = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel_integer x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel_integer_def proof let ?i = "int_of_integer" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel_integer a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel_integer x (?i x)" unfolding urel_integer_def using * unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) qed next assume "\a b. x = a \ (\z. urel_integer a z \ mod_ring_rel z b)" then obtain b z where xz: "urel_integer x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel_integer_def unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) qed qed lemma ring_finite_field_ops_integer: "ring_ops (finite_field_ops_integer pp) mod_ring_rel_integer" by (unfold_locales, auto simp: finite_field_ops_integer_def bi_unique_mod_ring_rel_integer right_total_mod_ring_rel_integer mod_ring_plus_integer mod_ring_minus_integer mod_ring_uminus_integer mod_ring_mult_integer mod_ring_eq_integer mod_ring_0__integer mod_ring_1__integer Domainp_mod_ring_rel_integer) end end context prime_field begin context fixes pp :: "integer" assumes *: "p = int_of_integer pp" begin lemma mod_ring_normalize_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (\x. if x = 0 then 0 else 1) normalize" using urel_integer_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_mod_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (\x y. if y = 0 then x else 0) (mod)" using urel_integer_mod[OF *] mod_ring_mod unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_inverse_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (inverse_p_integer pp) inverse" using urel_integer_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_divide_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (divide_p_integer pp) (/)" using mod_ring_inverse_integer mod_ring_mult_integer[OF *] unfolding divide_p_integer_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops_integer: "field_ops (finite_field_ops_integer pp) mod_ring_rel_integer" by (unfold_locales, insert ring_finite_field_ops_integer[OF *], auto simp: ring_ops_def finite_field_ops_integer_def mod_ring_divide_integer mod_ring_inverse_integer mod_ring_mod_integer mod_ring_normalize_integer) end end context prime_field begin (* four implementations of modular integer arithmetic for finite fields *) thm finite_field_ops64 finite_field_ops32 finite_field_ops_integer finite_field_ops_int end context mod_ring_locale begin (* four implementations of modular integer arithmetic for finite rings *) thm ring_finite_field_ops64 ring_finite_field_ops32 ring_finite_field_ops_integer ring_finite_field_ops_int end end diff --git a/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_State.thy b/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_State.thy --- a/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_State.thy +++ b/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_State.thy @@ -1,74 +1,74 @@ section \State for SM\ theory SM_State imports "HOL-Library.Word" "HOL-Library.Multiset" "Word_Lib.Typedef_Morphisms" SM_Syntax begin section \Values\ text \The primitive values are fixed-size signed integers\ type_synonym word_size = 32 \ \Word size\ type_synonym signed = "word_size Word.word" \ \Signed integer\ text \Currently, we only have signed integer values. This may change if we extend the language, and allow, i.e., channel pointers, pids or process references\ type_synonym val = signed \ \Value type\ section \Configurations\ type_synonym valuation = "ident \ val" record local_state = variables :: valuation record global_state = variables :: valuation text \The effect of actions is on focused states\ type_synonym focused_state = "local_state \ global_state" section \Utilities\ abbreviation "word_len \ LENGTH(word_size)" abbreviation "signeds \ sints (LENGTH(word_size))" definition min_signed :: int where "min_signed \ -(2^(word_len - 1))" definition max_signed :: int where "max_signed \ 2^(word_len - 1) - 1" definition signed_of_int :: "int \ signed" where "signed_of_int i \ word_of_int i" definition int_of_signed :: "signed \ int" where "int_of_signed i == sint i" lemma si_inv[simp]: "signed_of_int (int_of_signed i) = i" unfolding signed_of_int_def int_of_signed_def by simp lemma int_of_signed_in_range[simp]: "int_of_signed i \ min_signed" "int_of_signed i \ max_signed" unfolding int_of_signed_def min_signed_def max_signed_def apply - apply (rule sint_ge) using sint_lt[of i] by simp lemma is_inv[simp]: "\ i\min_signed; i\max_signed \ \ (int_of_signed (signed_of_int i)) = i" - by (simp add: signed_take_bit_int_eq_self min_signed_def max_signed_def int_of_signed_def signed_of_int_def) + by (simp add: signed_take_bit_int_eq_self min_signed_def max_signed_def int_of_signed_def signed_of_int_def signed_of_int) primrec val_of_bool :: "bool \ val" where "val_of_bool False = 0" | "val_of_bool True = 1" definition bool_of_val :: "val \ bool" where "bool_of_val v \ v\0" lemma bool_of_val_invs[simp]: "bool_of_val (val_of_bool b) = b" "val_of_bool (bool_of_val v) = (if v=0 then 0 else 1)" unfolding bool_of_val_def by (cases b) auto end diff --git a/thys/Complx/ex/SumArr.thy b/thys/Complx/ex/SumArr.thy --- a/thys/Complx/ex/SumArr.thy +++ b/thys/Complx/ex/SumArr.thy @@ -1,470 +1,470 @@ (* * Copyright 2016, Data61, CSIRO * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(DATA61_BSD) *) section \Case-study\ theory SumArr imports "../OG_Syntax" Word_Lib.Word_32 begin unbundle bit_operations_syntax type_synonym routine = nat type_synonym word32 = "32 word" type_synonym funcs = "string \ nat" datatype faults = Overflow | InvalidMem type_synonym 'a array = "'a list" text \Sumarr computes the combined sum of all the elements of multiple arrays. It does this by running a number of threads in parallel, each computing the sum of elements of one of the arrays, and then adding the result to a global variable gsum shared by all threads. \ record sumarr_state = \ \local variables of threads\ tarr :: "routine \ word32 array" tid :: "routine \ word32" ti :: "routine \ word32" tsum :: "routine \ word32" \ \global variables\ glock :: nat gsum :: word32 gdone :: word32 garr :: "(word32 array) array" \ \ghost variables\ ghost_lock :: "routine \ bool" definition NSUM :: word32 where "NSUM = 10" definition MAXSUM :: word32 where "MAXSUM = 1500" definition array_length :: "'a array \ word32" where "array_length arr \ of_nat (length arr)" definition array_nth :: "'a array \ word32 \'a" where "array_nth arr n \ arr ! unat n" definition array_in_bound :: "'a array \ word32 \ bool" where "array_in_bound arr idx \ unat idx < (length arr)" definition array_nat_sum :: "('a :: len) word array \ nat" where "array_nat_sum arr \ sum_list (map unat arr)" definition "local_sum arr \ of_nat (min (unat MAXSUM) (array_nat_sum arr))" definition "global_sum arr \ sum_list (map local_sum arr)" definition "tarr_inv s i \ length (tarr s i) = unat NSUM \ tarr s i = garr s ! i" abbreviation "sumarr_inv_till_lock s i \ \ bit (gdone s) i \ ((\ (ghost_lock s) (1 - i)) \ ((gdone s = 0 \ gsum s = 0) \ (bit (gdone s) (1 - i) \ gsum s = local_sum (garr s !(1 - i)))))" abbreviation "lock_inv s \ (glock s = fromEnum (ghost_lock s 0) + fromEnum (ghost_lock s 1)) \ (\(ghost_lock s) 0 \ \(ghost_lock s) 1)" abbreviation "garr_inv s i \ (\a b. garr s = [a, b]) \ length (garr s ! (1-i)) = unat NSUM" abbreviation "sumarr_inv s i \ lock_inv s \ tarr_inv s i \ garr_inv s i \ tid s i = (of_nat i + 1)" definition lock :: "routine \ (sumarr_state, funcs, faults) ann_com" where "lock i \ \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \sumarr_inv_till_lock i\ AWAIT \glock = 0 THEN \glock:=1,, \ghost_lock:=\ghost_lock (i:= True) END" definition "sumarr_in_lock1 s i \ \bit (gdone s) i \ ((gdone s = 0 \ gsum s = local_sum (tarr s i)) \ (bit (gdone s) (1 - i) \ \ bit (gdone s) i \ gsum s = global_sum (garr s)))" definition "sumarr_in_lock2 s i \ (bit (gdone s) i \ \ bit (gdone s) (1 - i) \ gsum s = local_sum (tarr s i)) \ (bit (gdone s) i \ bit (gdone s) (1 - i) \ gsum s = global_sum (garr s))" definition unlock :: "routine \ (sumarr_state, funcs, faults) ann_com" where "unlock i \ \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \glock = 1 \ \ghost_lock i \ bit \gdone (unat (\tid i - 1)) \ \sumarr_in_lock2 i \ \\glock := 0,, \ghost_lock:=\ghost_lock (i:= False)\" definition "local_postcond s i \ (\ (ghost_lock s) (1 - i) \ gsum s = (if bit (gdone s) 0 \ bit (gdone s) 1 then global_sum (garr s) else local_sum (garr s ! i))) \ bit (gdone s) i \ \ghost_lock s i" definition sumarr :: "routine \ (sumarr_state, funcs, faults) ann_com" where "sumarr i \ \\sumarr_inv i \ \sumarr_inv_till_lock i\ \tsum:=\tsum(i:=0) ;; \ \tsum i = 0 \ \sumarr_inv i \ \sumarr_inv_till_lock i\ \ti:=\ti(i:=0) ;; TRY \ \tsum i = 0 \ \sumarr_inv i \ \ti i = 0 \ \sumarr_inv_till_lock i\ WHILE \ti i < NSUM INV \ \sumarr_inv i \ \ti i \ NSUM \ \tsum i \ MAXSUM \ \tsum i = local_sum (take (unat (\ti i)) (\tarr i)) \ \sumarr_inv_till_lock i\ DO \ \sumarr_inv i \ \ti i < NSUM \ \tsum i \ MAXSUM \ \tsum i = local_sum (take (unat (\ti i)) (\tarr i)) \ \sumarr_inv_till_lock i\ (InvalidMem, \ array_in_bound (\tarr i) (\ti i) \) \ \ \sumarr_inv i \ \ti i < NSUM \ \tsum i \ MAXSUM \ \tsum i = local_sum (take (unat (\ti i)) (\tarr i)) \ \sumarr_inv_till_lock i\ \tsum:=\tsum(i:=\tsum i + array_nth (\tarr i) (\ti i));; \ \sumarr_inv i \ \ti i < NSUM \ local_sum (take (unat (\ti i)) (\tarr i)) \ MAXSUM \ (\tsum i < MAXSUM \ array_nth (\tarr i) (\ti i) < MAXSUM \ \tsum i = local_sum (take (Suc (unat (\ti i))) (\tarr i))) \ (array_nth (\tarr i) (\ti i) \ MAXSUM \ \tsum i \ MAXSUM\ local_sum (\tarr i) = MAXSUM) \ \sumarr_inv_till_lock i \ (InvalidMem, \ array_in_bound (\tarr i) (\ti i) \) \ \ \sumarr_inv i \ \ti i < NSUM \ (\tsum i < MAXSUM \ array_nth (\tarr i) (\ti i) < MAXSUM \ \tsum i = local_sum (take (Suc (unat (\ti i))) (\tarr i))) \ (array_nth (\tarr i) (\ti i) \ MAXSUM \ \tsum i \ MAXSUM \ local_sum (\tarr i) = MAXSUM) \ \sumarr_inv_till_lock i\ IF array_nth (\tarr i) (\ti i) \ MAXSUM \ \tsum i \ MAXSUM THEN \ \sumarr_inv i \ \ti i < NSUM \ local_sum (\tarr i) = MAXSUM \ \sumarr_inv_till_lock i\ \tsum:=\tsum(i:=MAXSUM);; \ \sumarr_inv i \ \ti i < NSUM \ \tsum i \ MAXSUM \ \tsum i = local_sum (\tarr i) \ \sumarr_inv_till_lock i \ THROW ELSE \ \sumarr_inv i \ \ti i < NSUM \ \tsum i \ MAXSUM \ \tsum i = local_sum (take (Suc (unat (\ti i))) (\tarr i)) \ \sumarr_inv_till_lock i\ SKIP FI;; \ \sumarr_inv i \ \ti i < NSUM \ \tsum i \ MAXSUM \ \tsum i = local_sum (take (Suc (unat (\ti i))) (\tarr i)) \ \sumarr_inv_till_lock i \ \ti:=\ti(i:=\ti i + 1) OD CATCH \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \sumarr_inv_till_lock i\ SKIP END;; \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \sumarr_inv_till_lock i\ SCALL (''lock'', i) 0;; \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \glock = 1 \ \ghost_lock i \ \sumarr_inv_till_lock i \ \gsum:=\gsum + \tsum i ;; \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \glock = 1 \ \ghost_lock i \ \sumarr_in_lock1 i \ \gdone:=(\gdone OR \tid i) ;; \ \sumarr_inv i \ \tsum i = local_sum (\tarr i) \ \glock = 1 \ \ghost_lock i \ bit \gdone (unat (\tid i - 1)) \ \sumarr_in_lock2 i \ SCALL (''unlock'', i) 0" definition precond where "precond s \ (glock s) = 0 \ (gsum s) = 0\ (gdone s) = 0 \ (\a b. garr s = [a, b]) \ (\xs\set (garr s). length xs = unat NSUM) \ (ghost_lock s) 0 = False \ (ghost_lock s) 1 = False" definition postcond where "postcond s \ (gsum s) = global_sum (garr s) \ (\i < 2. bit (gdone s) i)" definition "call_sumarr i \ \length (\garr ! i) = unat NSUM \ \lock_inv \ \garr_inv i \ \sumarr_inv_till_lock i\ CALLX (\s. s\tarr:=(tarr s)(i:=garr s ! i), tid:=(tid s)(i:=of_nat i+1), ti:=(ti s)(i:=undefined), tsum:=(tsum s)(i:=undefined)\) \\sumarr_inv i \ \sumarr_inv_till_lock i\ (''sumarr'', i) 0 (\s t. t\tarr:= (tarr t)(i:=(tarr s) i), tid:=(tid t)(i:=(tid s i)), ti:=(ti t)(i:=(ti s i)), tsum:=(tsum t)(i:=(tsum s i))\) (\_ _. Skip) \\local_postcond i\ \\local_postcond i\ \False\ \False\" definition "\ \ map_of (map (\i. ((''sumarr'', i), com (sumarr i))) [0..<2]) ++ map_of (map (\i. ((''lock'', i), com (lock i))) [0..<2]) ++ map_of (map (\i. ((''unlock'', i), com (unlock i))) [0..<2])" definition "\ \ map_of (map (\i. ((''sumarr'', i), [ann (sumarr i)])) [0..<2]) ++ map_of (map (\i. ((''lock'', i), [ann (lock i)])) [0..<2]) ++ map_of (map (\i. ((''unlock'', i), [ann (unlock i)])) [0..<2])" declare [[goals_limit = 10]] lemma [simp]: "local_sum [] = 0" by (simp add: local_sum_def array_nat_sum_def) lemma MAXSUM_le_plus: "x < MAXSUM \ MAXSUM \ MAXSUM + x" unfolding MAXSUM_def apply (rule word_le_plus[rotated], assumption) apply clarsimp done lemma local_sum_Suc: "\n < length arr; local_sum (take n arr) + arr ! n < MAXSUM; arr ! n < MAXSUM\ \ local_sum (take n arr) + arr ! n = local_sum (take (Suc n) arr)" apply (subst take_Suc_conv_app_nth) apply clarsimp apply (clarsimp simp: local_sum_def array_nat_sum_def ) apply (subst (asm) min_def, clarsimp split: if_splits) apply (clarsimp simp: MAXSUM_le_plus word_not_le[symmetric]) apply (subst min_absorb2) apply (subst of_nat_mono_maybe_le[where 'a=32]) apply (clarsimp simp: MAXSUM_def) apply (clarsimp simp: MAXSUM_def) apply unat_arith apply (clarsimp simp: MAXSUM_def) apply unat_arith apply clarsimp done lemma local_sum_MAXSUM: "k < length arr \ MAXSUM \ arr ! k \ local_sum arr = MAXSUM" apply (clarsimp simp: local_sum_def array_nat_sum_def) apply (rule word_unat.Rep_inverse') apply (rule min_absorb1[symmetric]) apply (subst (asm) word_le_nat_alt) apply (rule le_trans[rotated]) apply (rule elem_le_sum_list) apply simp apply clarsimp done lemma local_sum_MAXSUM': \local_sum arr = MAXSUM\ if \k < length arr\ \MAXSUM \ local_sum (take k arr) + arr ! k\ \local_sum (take k arr) \ MAXSUM\ \arr ! k \ MAXSUM\ proof - define vs u ws where \vs = take k arr\ \u = arr ! k\ \ws = drop (Suc k) arr\ with \k < length arr\ have *: \arr = vs @ u # ws\ and **: \take k arr = vs\ \arr ! k = u\ by (simp_all add: id_take_nth_drop) from that show ?thesis apply (simp add: **) apply (simp add: *) apply (simp add: local_sum_def array_nat_sum_def ac_simps) - find_theorems \min _ (_ + _)\ apply (rule word_unat.Rep_inverse') apply (rule min_absorb1[symmetric]) apply (subst (asm) word_le_nat_alt) apply (subst (asm) unat_plus_simple[THEN iffD1]) apply (rule word_add_le_mono2[where i=0, simplified]) apply (clarsimp simp: MAXSUM_def) apply unat_arith - apply simp_all apply (rule le_trans, assumption) apply (rule add_mono) apply simp_all - apply (meson min.bounded_iff take_bit_nat_less_eq_self trans_le_add1) + apply (subst le_unat_uoi) + apply (rule min.cobounded1) + apply simp done qed lemma word_min_0[simp]: "min (x::'a::len word) 0 = 0" "min 0 (x::'a::len word) = 0" by (simp add:min_def)+ ML \fun TRY' tac i = TRY (tac i)\ lemma imp_disjL_context': "((P \ R) \ (Q \ R)) = ((P \ R) \ (\P \ Q \ R))" by auto lemma map_of_prod_1[simp]: "i < n \ map_of (map (\i. ((p, i), g i)) [0.. p \ q \ (m ++ map_of (map (\i. ((p, i), g i)) [0.. \ (''sumarr'',n) = Some (com (sumarr n))" "n < 2 \ \ (''sumarr'',n) = Some ([ann (sumarr n)])" "n < 2 \ \ (''lock'',n) = Some (com (lock n))" "n < 2 \ \ (''lock'',n) = Some ([ann (lock n)])" "n < 2 \ \ (''unlock'',n) = Some (com (unlock n))" "n < 2 \ \ (''unlock'',n) = Some ([ann (unlock n)])" "[ann (sumarr n)]!0 = ann (sumarr n)" "[ann (lock n)]!0 = ann (lock n)" "[ann (unlock n)]!0 = ann (unlock n)" by (simp add: \_def \_def)+ lemmas sumarr_proc_simp_unfolded = sumarr_proc_simp[unfolded sumarr_def unlock_def lock_def oghoare_simps] lemma oghoare_sumarr: \\, \ |\\<^bsub>/F\<^esub> sumarr i \\local_postcond i\, \False\\ if \i < 2\ proof - from that have \i = 0 \ i = 1\ by auto note sumarr_proc_simp_unfolded[proc_simp add] show ?thesis using that apply - unfolding sumarr_def unlock_def lock_def ann_call_def call_def block_def apply simp apply oghoare (*24*) unfolding tarr_inv_def array_length_def array_nth_def array_in_bound_def sumarr_in_lock1_def sumarr_in_lock2_def apply (tactic "PARALLEL_ALLGOALS ((TRY' o SOLVED') (clarsimp_tac (@{context} addsimps @{thms local_postcond_def global_sum_def ex_in_conv[symmetric]}) THEN_ALL_NEW fast_force_tac (@{context} addSDs @{thms less_2_cases} addIs @{thms local_sum_Suc unat_mono} ) ))") (*4*) using \i = 0 \ i = 1\ apply rule apply (clarsimp simp add: bit_simps even_or_iff) apply (clarsimp simp add: bit_simps even_or_iff) apply clarsimp apply (rule conjI) apply (fastforce intro!: local_sum_Suc unat_mono) apply (subst imp_disjL_context') apply (rule conjI) apply clarsimp apply (erule local_sum_MAXSUM[rotated]) apply unat_arith apply (clarsimp simp: not_le) apply (erule (1) local_sum_MAXSUM'[rotated] ; unat_arith) apply clarsimp apply unat_arith apply (fact that) done qed lemma less_than_two_2[simp]: "i < 2 \ Suc 0 - i < 2" by arith lemma oghoare_call_sumarr: notes sumarr_proc_simp[proc_simp add] shows "i < 2 \ \, \ |\\<^bsub>/F\<^esub> call_sumarr i \\local_postcond i\, \False\" unfolding call_sumarr_def ann_call_def call_def block_def tarr_inv_def apply oghoare (*10*) apply (clarsimp; fail | ((simp only: pre.simps)?, rule oghoare_sumarr))+ apply (clarsimp simp: sumarr_def tarr_inv_def) apply (clarsimp simp: local_postcond_def; fail)+ done lemma less_than_two_inv[simp]: "i < 2 \ j < 2 \ i \ j \ Suc 0 - i = j" by simp lemma inter_aux_call_sumarr [simplified]: notes sumarr_proc_simp_unfolded [proc_simp add] com.simps [oghoare_simps add] bit_simps [simp] shows "i < 2 \ j < 2 \ i \ j \ interfree_aux \ \ F (com (call_sumarr i), (ann (call_sumarr i), \\local_postcond i\, \False\), com (call_sumarr j), ann (call_sumarr j))" unfolding call_sumarr_def ann_call_def call_def block_def tarr_inv_def sumarr_def lock_def unlock_def apply oghoare_interfree_aux (*650*) unfolding tarr_inv_def local_postcond_def sumarr_in_lock1_def sumarr_in_lock2_def by (tactic "PARALLEL_ALLGOALS ( TRY' (remove_single_Bound_mem @{context}) THEN' (TRY' o SOLVED') (clarsimp_tac @{context} THEN_ALL_NEW fast_force_tac (@{context} addSDs @{thms less_2_cases}) ))") (* 2 minutes *) lemma pre_call_sumarr: "i < 2 \ precond x \ x \ pre (ann (call_sumarr i))" unfolding precond_def call_sumarr_def ann_call_def by (fastforce dest: less_2_cases simp: array_length_def) lemma post_call_sumarr: "local_postcond x 0 \ local_postcond x 1 \ postcond x" unfolding postcond_def local_postcond_def by (fastforce dest: less_2_cases split: if_splits) lemma sumarr_correct: "\, \ |\\<^bsub>/F\<^esub> \\precond\ COBEGIN SCHEME [0 \ m < 2] call_sumarr m \\local_postcond m\,\False\ COEND \\postcond\, \False\" apply oghoare (* 5 subgoals *) apply (fastforce simp: pre_call_sumarr) apply (rule oghoare_call_sumarr, simp) apply (clarsimp simp: post_call_sumarr) apply (simp add: inter_aux_call_sumarr) apply clarsimp done end diff --git a/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy b/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy --- a/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy +++ b/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy @@ -1,346 +1,346 @@ (* Author: Fabian Hellauer Fabian Immler *) theory Conversion_IEEE_Float imports "HOL-Library.Float" IEEE_Properties "HOL-Library.Code_Target_Numeral" begin definition "of_finite (x::('e, 'f)float) = (if is_normal x then (Float (normal_mantissa x) (normal_exponent x)) else if is_denormal x then (Float (denormal_mantissa x) (denormal_exponent TYPE(('e, 'f)float))) else 0)" lemma float_val_of_finite: "is_finite x \ of_finite x = valof x" by (induction x) (auto simp: normal_imp_not_denormal of_finite_def) definition is_normal_Float::"('e, 'f)float itself \ Float.float \ bool" where "is_normal_Float x f \ mantissa f \ 0 \ bitlen \mantissa f\ \ fracwidth x + 1 \ - int (bias x) - bitlen \mantissa f\ + 1 < Float.exponent f \ Float.exponent f < 2^(LENGTH('e)) - bitlen \mantissa f\ - bias x" definition is_denormal_Float::"('e, 'f)float itself \ Float.float \ bool" where "is_denormal_Float x f \ mantissa f \ 0 \ bitlen \mantissa f\ \ 1 - Float.exponent f - int (bias x) \ 1 - 2^(LENGTH('e) - 1) - int LENGTH('f) < Float.exponent f" lemmas is_denormal_FloatD = is_denormal_Float_def[THEN iffD1, THEN conjunct1] is_denormal_Float_def[THEN iffD1, THEN conjunct2] definition is_finite_Float::"('e, 'f)float itself \ Float.float \ bool" where "is_finite_Float x f \ is_normal_Float x f \ is_denormal_Float x f \ f = 0" lemma is_finite_Float_eq: "is_finite_Float TYPE(('e, 'f)float) f \ (let e = Float.exponent f; bm = bitlen (abs (mantissa f)) in bm \ Suc LENGTH('f) \ bm \ 2 ^ (LENGTH('e) - 1) - e \ 1 - 2 ^ (LENGTH('e) - 1) - int LENGTH('f) < e)" proof - have *: "(2::int) ^ (LENGTH('e) - Suc 0) - 1 < 2 ^ LENGTH('e)" by (metis Suc_1 diff_le_self lessI linorder_not_less one_less_numeral_iff power_strict_increasing_iff zle_diff1_eq) have **: "1 - 2 ^ (LENGTH('e) - Suc 0) < int LENGTH('f)" by (smt len_gt_0 of_nat_0_less_iff zero_less_power) have ***: "2 ^ (LENGTH('e) - 1) + 1 = 2 ^ LENGTH('e) - int (bias TYPE(('e, 'f) IEEE.float))" by (simp add: bias_def power_Suc[symmetric] of_nat_diff) have rewr: "x \ 2 ^ n - e \ x + e < 2 ^ n + 1" for x::int and n e by auto show ?thesis unfolding *** rewr using * ** unfolding is_finite_Float_def is_normal_Float_def is_denormal_Float_def by (auto simp: Let_def bias_def mantissa_eq_zero_iff of_nat_diff intro: le_less_trans[OF add_right_mono]) qed lift_definition normal_of_Float :: "Float.float \ ('e, 'f)float" is "\x. let m = mantissa x; e = Float.exponent x in (if m > 0 then 0 else 1, word_of_int (e + int (bias TYPE(('e, 'f)float)) + bitlen \m\ - 1), word_of_int (\m\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \m\)) - 2 ^ (LENGTH('f))))" . lemma sign_normal_of_Float:"sign (normal_of_Float x) = (if x > 0 then 0 else 1)" by transfer (auto simp: Let_def mantissa_pos_iff) lemma uint_word_of_int_bitlen_eq: "uint (word_of_int x::'a::len word) = x" if "bitlen x \ LENGTH('a)" "x \ 0" - using that by (simp add: bitlen_le_iff_power take_bit_int_eq_self) + using that by (simp add: bitlen_le_iff_power take_bit_int_eq_self unsigned_of_int) lemma fraction_normal_of_Float:"fraction (normal_of_Float x::('e, 'f)float) = (nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f))" if "is_normal_Float TYPE(('e, 'f)float) x" proof - from that have bmp: "bitlen \mantissa x\ > 0" by (metis abs_of_nonneg bitlen_bounds bitlen_def is_normal_Float_def nat_code(2) of_nat_0_le_iff power.simps(1) zabs_less_one_iff zero_less_abs_iff) have mless: "\mantissa x\ < 2 ^ nat (bitlen \mantissa x\)" using bitlen_bounds by force have lem: "2 ^ nat (bitlen \mantissa x\ - 1) \ \mantissa x\" using bitlen_bounds is_normal_Float_def that zero_less_abs_iff by blast from that have nble: "nat (bitlen \mantissa x\) \ Suc LENGTH('f)" using bitlen_bounds by (auto simp: is_normal_Float_def) have nn: "0 \ \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f)" apply (rule add_le_imp_le_diff) apply (rule order_trans[rotated]) apply (rule mult_right_mono) apply (rule lem, force) unfolding power_add[symmetric] using nble bmp by (auto) have "\mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) < 2 * 2 ^ LENGTH('f)" apply (rule less_le_trans) apply (rule mult_strict_right_mono) apply (rule mless) apply force unfolding power_add[symmetric] power_Suc[symmetric] apply (rule power_increasing) using nble by auto then have "bitlen (\mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f)) \ int LENGTH('f)" unfolding bitlen_le_iff_power by simp then show ?thesis apply (transfer fixing: x) unfolding Let_def split_beta' fst_conv snd_conv uint_nat [symmetric] nat_uint_eq [symmetric] using nn apply (subst uint_word_of_int_bitlen_eq) apply (auto simp: nat_mult_distrib nat_diff_distrib nat_power_eq) done qed lemma exponent_normal_of_Float:"exponent (normal_of_Float x::('e, 'f)float) = nat (Float.exponent x + (bias TYPE(('e, 'f)float)) + bitlen \mantissa x\ - 1)" if "is_normal_Float TYPE(('e, 'f)float) x" using that apply (transfer fixing: x) apply (simp flip: uint_nat nat_uint_eq add: Let_def) apply (auto simp: is_normal_Float_def bitlen_le_iff_power uint_word_of_int_bitlen_eq Let_def) apply transfer apply (simp add: nat_take_bit_eq take_bit_int_eq_self) done lift_definition denormal_of_Float :: "Float.float \ ('e, 'f)float" is "\x. let m = mantissa x; e = Float.exponent x in (if m \ 0 then 0 else 1, 0, word_of_int (\m\ * 2 ^ nat (e + bias TYPE(('e, 'f)float) + fracwidth TYPE(('e, 'f)float) - 1)))" . lemma sign_denormal_of_Float:"sign (denormal_of_Float x) = (if x \ 0 then 0 else 1)" by transfer (auto simp: Let_def mantissa_nonneg_iff) lemma exponent_denormal_of_Float:"exponent (denormal_of_Float x::('e, 'f)float) = 0" by (transfer fixing: x) (auto simp: Let_def) lemma fraction_denormal_of_Float:"fraction (denormal_of_Float x::('e, 'f)float) = (nat \mantissa x\ * 2 ^ nat (Float.exponent x + bias TYPE(('e, 'f)float) + LENGTH('f) - 1))" if "is_denormal_Float TYPE(('e, 'f)float) x" proof - have mless: "\mantissa x\ < 2 ^ nat (bitlen \mantissa x\)" using bitlen_bounds by force have *: "nat (bitlen \mantissa x\) + nat (Float.exponent x + (2 ^ (LENGTH('e) - Suc 0) + int LENGTH('f)) - 2) \ LENGTH('f)" using that by (auto simp: is_denormal_Float_def nat_diff_distrib' le_diff_conv bitlen_nonneg nat_le_iff bias_def nat_add_distrib[symmetric] of_nat_diff) have "\mantissa x\ * 2 ^ nat (Float.exponent x + int (bias TYPE(('e, 'f)float)) + LENGTH('f) - 1) < 2 ^ LENGTH('f)" apply (rule less_le_trans) apply (rule mult_strict_right_mono) apply (rule mless, force) unfolding power_add[symmetric] power_Suc[symmetric] apply (rule power_increasing) apply (auto simp: bias_def) using that * by (auto simp: is_denormal_Float_def algebra_simps of_nat_diff) then show ?thesis apply (transfer fixing: x) apply transfer apply (simp add: Let_def nat_eq_iff take_bit_eq_mod) done qed definition of_finite_Float :: "Float.float \ ('e, 'f) float" where "of_finite_Float x = (if is_normal_Float TYPE(('e, 'f)float) x then normal_of_Float x else if is_denormal_Float TYPE(('e, 'f)float) x then denormal_of_Float x else 0)" lemma valof_normal_of_Float: "valof (normal_of_Float x::('e, 'f)float) = x" if "is_normal_Float TYPE(('e, 'f)float) x" proof - have "valof (normal_of_Float x::('e, 'f)float) = (- 1) ^ sign (normal_of_Float x::('e, 'f)float) * ((1 + real (nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f)) / 2 ^ LENGTH('f)) * 2 powr (bitlen \mantissa x\ - 1)) * 2 powr Float.exponent x" (is "_ = ?s * ?m * ?e") using that by (auto simp: is_normal_Float_def valof_eq fraction_normal_of_Float powr_realpow[symmetric] exponent_normal_of_Float powr_diff powr_add) also have "\mantissa x\ > 0" using that by (auto simp: is_normal_Float_def) have bound: "2 ^ LENGTH('f) \ nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\))" proof - have "(2::nat) ^ LENGTH('f) \ 2 ^ nat (bitlen \mantissa x\ - 1) * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\))" by (simp add: power_add[symmetric]) also have "\ \ nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\))" using bitlen_bounds[of "\mantissa x\"] that by (auto simp: is_normal_Float_def) finally show ?thesis . qed have "?m = abs (mantissa x)" apply (subst of_nat_diff) subgoal using bound by auto subgoal using that by (auto simp: powr_realpow[symmetric] powr_add[symmetric] is_normal_Float_def bitlen_nonneg of_nat_diff divide_simps) done finally show ?thesis by (auto simp: mantissa_exponent sign_normal_of_Float abs_real_def zero_less_mult_iff) qed lemma valof_denormal_of_Float: "valof (denormal_of_Float x::('e, 'f)float) = x" if "is_denormal_Float TYPE(('e, 'f)float) x" proof - have less: "0 < Float.exponent x + (int (bias TYPE(('e, 'f) IEEE.float)) + int LENGTH('f))" using that by (auto simp: is_denormal_Float_def bias_def of_nat_diff) have "valof (denormal_of_Float x::('e, 'f)float) = ((- 1) ^ sign (denormal_of_Float x::('e, 'f)float) * \real_of_int (mantissa x)\) * (2 powr real (nat (Float.exponent x + int (bias TYPE(('e, 'f) IEEE.float)) + int LENGTH('f) - 1)) / (2 powr real (bias TYPE(('e, 'f) IEEE.float)) * 2 powr LENGTH('f)) * 2)" (is "_ = ?m * ?e") by (auto simp: valof_eq exponent_denormal_of_Float fraction_denormal_of_Float that mantissa_exponent powr_realpow[symmetric]) also have "?m = mantissa x" by (auto simp: sign_denormal_of_Float abs_real_def mantissa_neg_iff) also have "?e = 2 powr Float.exponent x" by (auto simp: powr_add[symmetric] divide_simps powr_mult_base less ac_simps) finally show ?thesis by (simp add: mantissa_exponent) qed lemma valof_of_finite_Float: "is_finite_Float (TYPE(('e, 'f) IEEE.float)) x \ valof (of_finite_Float x::('e, 'f)float) = x" by (auto simp: of_finite_Float_def is_finite_Float_def valof_denormal_of_Float valof_normal_of_Float) lemma is_normal_normal_of_Float: "is_normal (normal_of_Float x::('e, 'f)float)" if "is_normal_Float TYPE(('e, 'f)float) x" using that by (auto simp: is_normal_def exponent_normal_of_Float that is_normal_Float_def emax_eq nat_less_iff of_nat_diff) lemma is_denormal_denormal_of_Float: "is_denormal (denormal_of_Float x::('e, 'f)float)" if "is_denormal_Float TYPE(('e, 'f)float) x" using that by (auto simp: is_denormal_def exponent_denormal_of_Float that is_denormal_Float_def emax_eq fraction_denormal_of_Float le_nat_iff bias_def) lemma is_finite_of_finite_Float: "is_finite (of_finite_Float x)" by (auto simp: is_finite_def of_finite_Float_def is_normal_normal_of_Float is_denormal_denormal_of_Float) lemma Float_eq_zero_iff: "Float m e = 0 \ m = 0" by (metis Float.compute_is_float_zero Float_0_eq_0) lemma bitlen_mantissa_Float: shows "bitlen \mantissa (Float m e)\ = (if m = 0 then 0 else bitlen \m\ + e) - Float.exponent (Float m e)" using bitlen_Float[of m e] by auto lemma exponent_Float: shows "Float.exponent (Float m e) = (if m = 0 then 0 else bitlen \m\ + e) - bitlen \mantissa (Float m e)\ " using bitlen_Float[of m e] by auto lemma is_normal_Float_normal: "is_normal_Float TYPE(('e, 'f)float) (Float (normal_mantissa x) (normal_exponent x))" if "is_normal x" for x::"('e, 'f)float" proof - define f where "f = Float (normal_mantissa x) (normal_exponent x)" from that have "f \ 0" by (auto simp: f_def is_normal_def zero_float_def[symmetric] Float_eq_zero_iff normal_mantissa_def add_nonneg_eq_0_iff) from denormalize_shift[OF f_def this] obtain i where i: "normal_mantissa x = mantissa f * 2 ^ i" "normal_exponent x = Float.exponent f - int i" by auto have "mantissa f \ 0" by (auto simp: \f \ 0\ i mantissa_eq_zero_iff Float_eq_zero_iff) moreover have "normal_exponent x \ Float.exponent f" unfolding i by simp then have " bitlen \mantissa f\ \ 1 + int LENGTH('f)" unfolding bitlen_mantissa_Float bitlen_normal_mantissa f_def by auto moreover have "- int (bias TYPE(('e, 'f)float)) - bitlen \mantissa f\ + 1 < Float.exponent f" unfolding bitlen_mantissa_Float bitlen_normal_mantissa f_def using that by (auto simp: mantissa_eq_zero_iff abs_mult bias_def normal_mantissa_def normal_exponent_def is_normal_def emax_eq less_diff_conv add_nonneg_eq_0_iff) moreover have "2 ^ (LENGTH('e) - Suc 0) + - (1::int) * 2 ^ LENGTH('e) \ 0" by simp then have "(2::int) ^ (LENGTH('e) - Suc 0) < 1 + 2 ^ LENGTH('e)" by arith then have "Float.exponent f < 2 ^ LENGTH('e) - bitlen \mantissa f\ - int (bias TYPE(('e, 'f)float))" using normal_exponent_bounds_int[OF that] unfolding bitlen_mantissa_Float bitlen_normal_mantissa f_def by (auto simp: bias_def algebra_simps power_Suc[symmetric] of_nat_diff intro: le_less_trans[OF add_right_mono] normal_exponent_bounds_int[OF that]) ultimately show ?thesis by (auto simp: is_normal_Float_def f_def) qed lemma is_denormal_Float_denormal: "is_denormal_Float TYPE(('e, 'f)float) (Float (denormal_mantissa x) (denormal_exponent TYPE(('e, 'f)float)))" if "is_denormal x" for x::"('e, 'f)float" proof - define f where "f = Float (denormal_mantissa x) (denormal_exponent TYPE(('e, 'f)float))" from that have "f \ 0" by (auto simp: f_def is_denormal_def zero_float_def[symmetric] Float_eq_zero_iff denormal_mantissa_def add_nonneg_eq_0_iff) from denormalize_shift[OF f_def this] obtain i where i: "denormal_mantissa x = mantissa f * 2 ^ i" "denormal_exponent TYPE(('e, 'f)float) = Float.exponent f - int i" by auto have "mantissa f \ 0" by (auto simp: \f \ 0\ i mantissa_eq_zero_iff Float_eq_zero_iff) moreover have "bitlen \mantissa f\ \ 1 - Float.exponent f - int (bias TYPE(('e, 'f) IEEE.float))" using \mantissa f \ 0\ unfolding f_def bitlen_mantissa_Float using bitlen_denormal_mantissa[of x] by (auto simp: denormal_exponent_def) moreover have "2 - 2 ^ (LENGTH('e) - Suc 0) - int LENGTH('f) \ Float.exponent f" (is "?l \ _") proof - have "?l \ denormal_exponent TYPE(('e, 'f)float) + i" using that by (auto simp: is_denormal_def bias_def denormal_exponent_def of_nat_diff) also have "\ = Float.exponent f" unfolding i by auto finally show ?thesis . qed ultimately show ?thesis unfolding is_denormal_Float_def exponent_Float f_def[symmetric] by auto qed lemma is_finite_Float_of_finite: "is_finite_Float TYPE(('e, 'f)float) (of_finite x)" for x::"('e, 'f)float" by (auto simp: is_finite_Float_def of_finite_def is_normal_Float_normal is_denormal_Float_denormal) end diff --git a/thys/Iptables_Semantics/Common/Word_Upto.thy b/thys/Iptables_Semantics/Common/Word_Upto.thy --- a/thys/Iptables_Semantics/Common/Word_Upto.thy +++ b/thys/Iptables_Semantics/Common/Word_Upto.thy @@ -1,206 +1,213 @@ section\Word Upto\ theory Word_Upto imports Main IP_Addresses.Hs_Compat Word_Lib.Word_Lemmas begin text\Enumerate a range of machine words.\ text\enumerate from the back (inefficient)\ function word_upto :: "'a word \ 'a word \ ('a::len) word list" where "word_upto a b = (if a = b then [a] else word_upto a (b - 1) @ [b])" by pat_completeness auto (*by the way: does not terminate practically if b < a; will terminate after it reaches the word wrap-around!*) termination word_upto apply(relation "measure (unat \ uncurry (-) \ prod.swap)") apply(rule wf_measure; fail) apply(simp) apply(subgoal_tac "unat (b - a - 1) < unat (b - a)") apply(simp add: diff_right_commute; fail) apply(rule measure_unat) apply auto done declare word_upto.simps[simp del] text\enumerate from the front (more inefficient)\ function word_upto' :: "'a word \ 'a word \ ('a::len) word list" where "word_upto' a b = (if a = b then [a] else a # word_upto' (a + 1) b)" by pat_completeness auto termination word_upto' apply(relation "measure (\ (a, b). unat (b - a))") apply(rule wf_measure; fail) apply(simp) apply(subgoal_tac "unat (b - a - 1) < unat (b - a)") apply (simp add: diff_diff_add; fail) apply(rule measure_unat) apply auto done declare word_upto'.simps[simp del] lemma word_upto_cons_front[code]: "word_upto a b = word_upto' a b" proof(induction a b rule:word_upto'.induct) case (1 a b) have hlp1: "a \ b \ a # word_upto (a + 1) b = word_upto a b" apply(induction a b rule:word_upto.induct) apply simp apply(subst(1) word_upto.simps) apply(simp) apply safe apply(subst(1) word_upto.simps) apply (simp) apply(subst(1) word_upto.simps) apply (simp; fail) apply(case_tac "a \ b - 1") apply(simp) apply (metis Cons_eq_appendI word_upto.simps) apply(simp) done from 1[symmetric] show ?case apply(cases "a = b") subgoal apply(subst word_upto.simps) apply(subst word_upto'.simps) by(simp) apply(subst word_upto'.simps) by(simp add: hlp1) qed (* Most of the lemmas I show about word_upto hold without a \ b, but I don't need that right now and it's giving me a headache *) lemma word_upto_set_eq: "a \ b \ x \ set (word_upto a b) \ a \ x \ x \ b" proof show "a \ b \ x \ set (word_upto a b) \ a \ x \ x \ b" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst(asm) word_upto.simps) apply(simp; fail) apply(subst(asm) word_upto.simps) apply(simp) apply(erule disjE) apply(simp; fail) proof(goal_cases) case (1 a b) from 1(2-3) have "b \ 0" by force from 1(2,3) have "a \ b - 1" by (simp add: word_le_minus_one_leq) from 1(1)[OF this 1(4)] show ?case by (metis dual_order.trans 1(2,3) less_imp_le measure_unat word_le_0_iff word_le_nat_alt) qed next show "a \ x \ x \ b \ x \ set (word_upto a b)" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst word_upto.simps) apply(simp; force) apply(subst word_upto.simps) apply(simp) apply(case_tac "x = b") apply(simp;fail) proof(goal_cases) case (1 a b) from 1(2-4) have "b \ 0" by force from 1(2,4) have "x \ b - 1" using le_step_down_word by auto from 1(1) this show ?case by simp qed qed lemma word_upto_distinct_hlp: "a \ b \ a \ b \ b \ set (word_upto a (b - 1))" apply(rule ccontr, unfold not_not) apply(subgoal_tac "a \ b - 1") apply(drule iffD1[OF word_upto_set_eq[of a "b -1" b]]) apply(simp add: word_upto.simps) apply(subgoal_tac "b \ 0") apply(meson leD measure_unat word_le_nat_alt) apply(blast intro: iffD1[OF word_le_0_iff]) using le_step_down_word apply blast done lemma distinct_word_upto: "a \ b \ distinct (word_upto a b)" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst word_upto.simps) apply(simp; force) apply(subst word_upto.simps) apply(case_tac "a \ b - 1") apply(simp) apply(rule word_upto_distinct_hlp; simp) apply(simp) apply(rule ccontr) apply (simp add: not_le antisym word_minus_one_le_leq) done lemma word_upto_eq_upto: "s \ e \ e \ unat (max_word :: 'l word) \ word_upto ((of_nat :: nat \ ('l :: len) word) s) (of_nat e) = map of_nat (upt s (Suc e))" proof(induction e) let ?mwon = "of_nat :: nat \ 'l word" let ?mmw = "max_word :: 'l word" case (Suc e) show ?case proof(cases "?mwon s = ?mwon (Suc e)") case True have "s = Suc e" using le_unat_uoi Suc.prems True by metis with True show ?thesis by(subst word_upto.simps) (simp) next case False hence le: "s \ e" using le_SucE Suc.prems by blast have lm: "e \ unat ?mmw" using Suc.prems by simp have sucm: "(of_nat :: nat \ ('l :: len) word) (Suc e) - 1 = of_nat e" using Suc.prems(2) by simp note mIH = Suc.IH[OF le lm] show ?thesis by(subst word_upto.simps) (simp add: False[simplified] Suc.prems mIH sucm) qed qed(simp add: word_upto.simps) lemma word_upto_alt: "(a :: ('l :: len) word) \ b \ word_upto a b = map of_nat (upt (unat a) (Suc (unat b)))" proof - let ?mmw = "- 1 :: 'l word" assume le: "a \ b" hence nle: "unat a \ unat b" by(unat_arith) have lem: "unat b \ unat ?mmw" by (simp add: word_unat_less_le) then show "word_upto a b = map of_nat [unat a.. b then map of_nat (upt (unat a) (Suc (unat b))) else word_upto a b)" using word_upto_alt by metis lemma sorted_word_upto: fixes a b :: "('l :: len) word" assumes "a \ b" shows "sorted (word_upto a b)" proof - define m and n where \m = unat a\ and \n = Suc (unat b)\ - moreover have \sorted (map of_nat [m.. - apply (simp add: sorted_map) - apply (rule sorted_wrt_mono_rel [of _ \(\)\]) - apply simp_all - apply (simp add: le_unat_uoi less_Suc_eq_le n_def word_of_nat_le) - apply transfer - apply simp - apply (subst take_bit_int_eq_self) - apply (simp_all add: le_less_trans) - apply (metis le_unat_uoi of_int_of_nat_eq of_nat_mono uint_word_of_int_eq unat_eq_nat_uint unsigned_of_int) - done + moreover have \sorted_wrt (\x y. (word_of_nat x :: 'l word) \ word_of_nat y) [m.. + proof (rule sorted_wrt_mono_rel [of _ \(\)\]) + show \sorted [m.. + by simp + fix r s + assume \r \ set [m.. \s \ set [m.. \r \ s\ + then have \m \ r\ \s < n\ + by simp_all + then have \take_bit LENGTH('l) s = s\ + by (auto simp add: m_def n_def less_Suc_eq_le unsigned_of_nat dest: le_unat_uoi) + with \r \ s\ show \(word_of_nat r :: 'l word) \ word_of_nat s\ + apply (simp add: of_nat_word_less_eq_iff) + using take_bit_nat_less_eq_self apply (rule order_trans) + apply assumption + done + qed + then have \sorted (map word_of_nat [m.. + by (simp add: sorted_map) ultimately have \sorted (map of_nat [unat a.. by simp with assms show ?thesis by (simp only: word_upto_alt) qed end diff --git a/thys/Native_Word/Code_Target_Word_Base.thy b/thys/Native_Word/Code_Target_Word_Base.thy --- a/thys/Native_Word/Code_Target_Word_Base.thy +++ b/thys/Native_Word/Code_Target_Word_Base.thy @@ -1,424 +1,425 @@ (* Title: Code_Target_Word_Base.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Common base for target language implementations of word types\ theory Code_Target_Word_Base imports "HOL-Library.Word" "Word_Lib.Signed_Division_Word" Bits_Integer begin text \More lemmas\ lemma div_half_nat: fixes x y :: nat assumes "y \ 0" shows "(x div y, x mod y) = (let q = 2 * (x div 2 div y); r = x - q * y in if y \ r then (q + 1, r - y) else (q, r))" proof - let ?q = "2 * (x div 2 div y)" have q: "?q = x div y - x div y mod 2" by(metis div_mult2_eq mult.commute minus_mod_eq_mult_div [symmetric]) let ?r = "x - ?q * y" have r: "?r = x mod y + x div y mod 2 * y" by(simp add: q diff_mult_distrib minus_mod_eq_div_mult [symmetric])(metis diff_diff_cancel mod_less_eq_dividend mod_mult2_eq add.commute mult.commute) show ?thesis proof(cases "y \ x - ?q * y") case True with assms q have "x div y mod 2 \ 0" unfolding r by (metis Nat.add_0_right diff_0_eq_0 diff_Suc_1 le_div_geq mod2_gr_0 mod_div_trivial mult_0 neq0_conv numeral_1_eq_Suc_0 numerals(1)) hence "x div y = ?q + 1" unfolding q by simp moreover hence "x mod y = ?r - y" by simp(metis minus_div_mult_eq_mod [symmetric] diff_commute diff_diff_left mult_Suc) ultimately show ?thesis using True by(simp add: Let_def) next case False hence "x div y mod 2 = 0" unfolding r by(simp add: not_le)(metis Nat.add_0_right assms div_less div_mult_self2 mod_div_trivial mult.commute) hence "x div y = ?q" unfolding q by simp moreover hence "x mod y = ?r" by (metis minus_div_mult_eq_mod [symmetric]) ultimately show ?thesis using False by(simp add: Let_def) qed qed lemma div_half_word: fixes x y :: "'a :: len word" assumes "y \ 0" shows "(x div y, x mod y) = (let q = push_bit 1 (drop_bit 1 x div y); r = x - q * y in if y \ r then (q + 1, r - y) else (q, r))" proof - obtain n where n: "x = of_nat n" "n < 2 ^ LENGTH('a)" by (rule that [of \unat x\]) simp_all moreover obtain m where m: "y = of_nat m" "m < 2 ^ LENGTH('a)" by (rule that [of \unat y\]) simp_all ultimately have [simp]: \unat (of_nat n :: 'a word) = n\ \unat (of_nat m :: 'a word) = m\ by (transfer, simp add: take_bit_of_nat take_bit_nat_eq_self_iff)+ let ?q = "push_bit 1 (drop_bit 1 x div y)" let ?q' = "2 * (n div 2 div m)" have "n div 2 div m < 2 ^ LENGTH('a)" using n by (metis of_nat_inverse unat_lt2p uno_simps(2)) hence q: "?q = of_nat ?q'" using n m - by (auto simp add: drop_bit_eq_div word_arith_nat_div uno_simps take_bit_nat_eq_self) + by (auto simp add: drop_bit_eq_div word_arith_nat_div uno_simps take_bit_nat_eq_self unsigned_of_nat) from assms have "m \ 0" using m by -(rule notI, simp) from n have "2 * (n div 2 div m) < 2 ^ LENGTH('a)" by(metis mult.commute div_mult2_eq minus_mod_eq_mult_div [symmetric] less_imp_diff_less of_nat_inverse unat_lt2p uno_simps(2)) moreover have "2 * (n div 2 div m) * m < 2 ^ LENGTH('a)" using n unfolding div_mult2_eq[symmetric] by(subst (2) mult.commute)(simp add: minus_mod_eq_div_mult [symmetric] diff_mult_distrib minus_mod_eq_mult_div [symmetric] div_mult2_eq) moreover have "2 * (n div 2 div m) * m \ n" by (simp flip: div_mult2_eq ac_simps) ultimately have r: "x - ?q * y = of_nat (n - ?q' * m)" and "y \ x - ?q * y \ of_nat (n - ?q' * m) - y = of_nat (n - ?q' * m - m)" using n m unfolding q apply (simp_all add: of_nat_diff) apply (subst of_nat_diff) - apply (simp_all add: word_le_nat_alt take_bit_nat_eq_self unat_sub_if' unat_word_ariths) + apply (simp_all add: word_le_nat_alt take_bit_nat_eq_self unat_sub_if' unat_word_ariths unsigned_of_nat) done then show ?thesis using n m div_half_nat [OF \m \ 0\, of n] unfolding q - by (simp add: word_le_nat_alt word_div_def word_mod_def Let_def take_bit_nat_eq_self + by (simp add: word_le_nat_alt word_div_def word_mod_def Let_def take_bit_nat_eq_self unsigned_of_nat flip: zdiv_int zmod_int split del: if_split split: if_split_asm) qed lemma word_test_bit_set_bits: "bit (BITS n. f n :: 'a :: len word) n \ n < LENGTH('a) \ f n" by (fact bit_set_bits_word_iff) lemma word_of_int_conv_set_bits: "word_of_int i = (BITS n. bit i n)" by (rule word_eqI) (auto simp add: word_test_bit_set_bits bit_simps) context includes bit_operations_syntax begin lemma word_and_mask_or_conv_and_mask: "bit n index \ (n AND mask index) OR (push_bit index 1) = n AND mask (index + 1)" for n :: \'a::len word\ by(rule word_eqI)(auto simp add: bit_simps) lemma uint_and_mask_or_full: fixes n :: "'a :: len word" assumes "bit n (LENGTH('a) - 1)" and "mask1 = mask (LENGTH('a) - 1)" and "mask2 = push_bit (LENGTH('a) - 1) 1" shows "uint (n AND mask1) OR mask2 = uint n" proof - have "mask2 = uint (push_bit (LENGTH('a) - 1) 1 :: 'a word)" using assms by transfer (simp add: take_bit_push_bit) hence "uint (n AND mask1) OR mask2 = uint (n AND mask1 OR (push_bit (LENGTH('a) - 1) 1 :: 'a word))" by(simp add: uint_or) also have "\ = uint (n AND mask (LENGTH('a) - 1 + 1))" using assms by(simp only: word_and_mask_or_conv_and_mask) also have "\ = uint n" by simp finally show ?thesis . qed end text \Division on @{typ "'a word"} is unsigned, but Scala and OCaml only have signed division and modulus.\ lemmas word_sdiv_def = sdiv_word_def lemmas word_smod_def = smod_word_def lemma [code]: "x sdiv y = (let x' = sint x; y' = sint y; negative = (x' < 0) \ (y' < 0); result = abs x' div abs y' in word_of_int (if negative then -result else result))" for x y :: \'a::len word\ by (simp add: sdiv_word_def signed_divide_int_def sgn_if Let_def not_less not_le) lemma [code]: "x smod y = (let x' = sint x; y' = sint y; negative = (x' < 0); result = abs x' mod abs y' in word_of_int (if negative then -result else result))" for x y :: \'a::len word\ proof - have *: \k mod l = k - k div l * l\ for k l :: int by (simp add: minus_div_mult_eq_mod) show ?thesis by (simp add: smod_word_def signed_modulo_int_def signed_divide_int_def * sgn_if Let_def) qed text \ This algorithm implements unsigned division in terms of signed division. Taken from Hacker's Delight. \ lemma divmod_via_sdivmod: fixes x y :: "'a :: len word" assumes "y \ 0" shows "(x div y, x mod y) = (if push_bit (LENGTH('a) - 1) 1 \ y then if x < y then (0, x) else (1, x - y) else let q = (push_bit 1 (drop_bit 1 x sdiv y)); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" proof(cases "push_bit (LENGTH('a) - 1) 1 \ y") case True note y = this show ?thesis proof(cases "x < y") case True then have "x mod y = x" by transfer simp thus ?thesis using True y by(simp add: word_div_lt_eq_0) next case False obtain n where n: "y = of_nat n" "n < 2 ^ LENGTH('a)" by (rule that [of \unat y\]) simp_all have "unat x < 2 ^ LENGTH('a)" by(rule unat_lt2p) also have "\ = 2 * 2 ^ (LENGTH('a) - 1)" by(metis Suc_pred len_gt_0 power_Suc One_nat_def) also have "\ \ 2 * n" using y n by transfer (simp add: push_bit_of_1 take_bit_eq_mod) finally have div: "x div of_nat n = 1" using False n - by (simp add: word_div_eq_1_iff take_bit_nat_eq_self) + by (simp add: word_div_eq_1_iff take_bit_nat_eq_self unsigned_of_nat) moreover have "x mod y = x - x div y * y" by (simp add: minus_div_mult_eq_mod) with div n have "x mod y = x - y" by simp ultimately show ?thesis using False y n by simp qed next case False note y = this obtain n where n: "x = of_nat n" "n < 2 ^ LENGTH('a)" by (rule that [of \unat x\]) simp_all hence "int n div 2 + 2 ^ (LENGTH('a) - Suc 0) < 2 ^ LENGTH('a)" by (cases \LENGTH('a)\) (auto dest: less_imp_of_nat_less [where ?'a = int]) with y n have "sint (drop_bit 1 x) = uint (drop_bit 1 x)" by (cases \LENGTH('a)\) - (auto simp add: sint_uint drop_bit_eq_div take_bit_nat_eq_self uint_div_distrib signed_take_bit_int_eq_self_iff) + (auto simp add: sint_uint drop_bit_eq_div take_bit_nat_eq_self uint_div_distrib + signed_take_bit_int_eq_self_iff unsigned_of_nat) moreover have "uint y + 2 ^ (LENGTH('a) - Suc 0) < 2 ^ LENGTH('a)" using y by (cases \LENGTH('a)\) (simp_all add: not_le push_bit_of_1 word_less_alt uint_power_lower) then have "sint y = uint y" apply (cases \LENGTH('a)\) apply (auto simp add: sint_uint signed_take_bit_int_eq_self_iff) using uint_ge_0 [of y] by linarith ultimately show ?thesis using y apply (subst div_half_word [OF assms]) apply (simp add: sdiv_word_def signed_divide_int_def flip: uint_div) done qed text \More implementations tailored towards target-language implementations\ context includes integer.lifting begin lift_definition word_of_integer :: "integer \ 'a :: len word" is word_of_int . lemma word_of_integer_code [code]: "word_of_integer n = word_of_int (int_of_integer n)" by(simp add: word_of_integer.rep_eq) end context includes bit_operations_syntax begin lemma word_of_int_code: "uint (word_of_int x :: 'a word) = x AND mask (LENGTH('a :: len))" - by (simp add: take_bit_eq_mask) + by (simp add: unsigned_of_int take_bit_eq_mask) context fixes f :: "nat \ bool" begin definition set_bits_aux :: \nat \ 'a word \ 'a::len word\ where \set_bits_aux n w = push_bit n w OR take_bit n (set_bits f)\ lemma bit_set_bit_aux [bit_simps]: \bit (set_bits_aux n w) m \ m < LENGTH('a) \ (if m < n then f m else bit w (m - n))\ for w :: \'a::len word\ by (auto simp add: bit_simps set_bits_aux_def) corollary set_bits_conv_set_bits_aux: \set_bits f = (set_bits_aux LENGTH('a) 0 :: 'a :: len word)\ by (rule bit_word_eqI) (simp add: bit_simps) lemma set_bits_aux_0 [simp]: \set_bits_aux 0 w = w\ by (simp add: set_bits_aux_def) lemma set_bits_aux_Suc [simp]: \set_bits_aux (Suc n) w = set_bits_aux n (push_bit 1 w OR (if f n then 1 else 0))\ by (rule bit_word_eqI) (auto simp add: bit_simps le_less_Suc_eq mult.commute [of _ 2]) lemma set_bits_aux_simps [code]: \set_bits_aux 0 w = w\ \set_bits_aux (Suc n) w = set_bits_aux n (push_bit 1 w OR (if f n then 1 else 0))\ by simp_all lemma set_bits_aux_rec: \set_bits_aux n w = (if n = 0 then w else let n' = n - 1 in set_bits_aux n' (push_bit 1 w OR (if f n' then 1 else 0)))\ by (cases n) simp_all end lemma word_of_int_via_signed: fixes mask assumes mask_def: "mask = Bit_Operations.mask (LENGTH('a))" and shift_def: "shift = push_bit LENGTH('a) 1" and index_def: "index = LENGTH('a) - 1" and overflow_def:"overflow = push_bit (LENGTH('a) - 1) 1" and least_def: "least = - overflow" shows "(word_of_int i :: 'a :: len word) = (let i' = i AND mask in if bit i' index then if i' - shift < least \ overflow \ i' - shift then arbitrary1 i' else word_of_int (i' - shift) else if i' < least \ overflow \ i' then arbitrary2 i' else word_of_int i')" proof - define i' where "i' = i AND mask" have "shift = mask + 1" unfolding assms by (simp add: mask_eq_exp_minus_1 push_bit_of_1) hence "i' < shift" by (simp add: mask_def i'_def) show ?thesis proof(cases "bit i' index") case True then have unf: "i' = overflow OR i'" apply (simp add: assms i'_def push_bit_of_1 flip: take_bit_eq_mask) apply (rule bit_eqI) apply (auto simp add: bit_take_bit_iff bit_or_iff bit_exp_iff) done have \overflow \ overflow OR i'\ by (simp add: i'_def mask_def or_greater_eq) then have "overflow \ i'" by (subst unf) hence "i' - shift < least \ False" unfolding assms by(cases "LENGTH('a)")(simp_all add: not_less push_bit_of_1) moreover have "overflow \ i' - shift \ False" using \i' < shift\ unfolding assms by(cases "LENGTH('a)")(auto simp add: not_le push_bit_of_1 elim: less_le_trans) moreover have "word_of_int (i' - shift) = (word_of_int i :: 'a word)" using \i' < shift\ by (simp add: i'_def shift_def mask_def push_bit_of_1 word_of_int_eq_iff flip: take_bit_eq_mask) ultimately show ?thesis using True by(simp add: Let_def i'_def) next case False have "i' = i AND Bit_Operations.mask (LENGTH('a) - 1)" apply (rule bit_eqI) apply (use False in \auto simp add: bit_simps assms i'_def\) apply (auto simp add: less_le) done also have "\ \ Bit_Operations.mask (LENGTH('a) - 1)" using AND_upper2 mask_nonnegative_int by blast also have "\ < overflow" by (simp add: mask_int_def push_bit_of_1 overflow_def) also have "least \ 0" unfolding least_def overflow_def by simp have "0 \ i'" by (simp add: i'_def mask_def) hence "least \ i'" using \least \ 0\ by simp moreover have "word_of_int i' = (word_of_int i :: 'a word)" by (simp add: i'_def mask_def of_int_and_eq of_int_mask_eq) ultimately show ?thesis using False by(simp add: Let_def i'_def) qed qed end text \Quickcheck conversion functions\ context includes state_combinator_syntax begin definition qc_random_cnv :: "(natural \ 'a::term_of) \ natural \ Random.seed \ ('a \ (unit \ Code_Evaluation.term)) \ Random.seed" where "qc_random_cnv a_of_natural i = Random.range (i + 1) \\ (\k. Pair ( let n = a_of_natural k in (n, \_. Code_Evaluation.term_of n)))" end definition qc_exhaustive_cnv :: "(natural \ 'a) \ ('a \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "qc_exhaustive_cnv a_of_natural f d = Quickcheck_Exhaustive.exhaustive (%x. f (a_of_natural x)) d" definition qc_full_exhaustive_cnv :: "(natural \ ('a::term_of)) \ ('a \ (unit \ term) \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "qc_full_exhaustive_cnv a_of_natural f d = Quickcheck_Exhaustive.full_exhaustive (%(x, xt). f (a_of_natural x, %_. Code_Evaluation.term_of (a_of_natural x))) d" declare [[quickcheck_narrowing_ghc_options = "-XTypeSynonymInstances"]] definition qc_narrowing_drawn_from :: "'a list \ integer \ _" where "qc_narrowing_drawn_from xs = foldr Quickcheck_Narrowing.sum (map Quickcheck_Narrowing.cons (butlast xs)) (Quickcheck_Narrowing.cons (last xs))" locale quickcheck_narrowing_samples = fixes a_of_integer :: "integer \ 'a \ 'a :: {partial_term_of, term_of}" and zero :: "'a" and tr :: "typerep" begin function narrowing_samples :: "integer \ 'a list" where "narrowing_samples i = (if i > 0 then let (a, a') = a_of_integer i in narrowing_samples (i - 1) @ [a, a'] else [zero])" by pat_completeness auto termination including integer.lifting proof(relation "measure nat_of_integer") fix i :: integer assume "0 < i" thus "(i - 1, i) \ measure nat_of_integer" by simp(transfer, simp) qed simp definition partial_term_of_sample :: "integer \ 'a" where "partial_term_of_sample i = (if i < 0 then undefined else if i = 0 then zero else if i mod 2 = 0 then snd (a_of_integer (i div 2)) else fst (a_of_integer (i div 2 + 1)))" lemma partial_term_of_code: "partial_term_of (ty :: 'a itself) (Quickcheck_Narrowing.Narrowing_variable p t) \ Code_Evaluation.Free (STR ''_'') tr" "partial_term_of (ty :: 'a itself) (Quickcheck_Narrowing.Narrowing_constructor i []) \ Code_Evaluation.term_of (partial_term_of_sample i)" by (rule partial_term_of_anything)+ end lemmas [code] = quickcheck_narrowing_samples.narrowing_samples.simps quickcheck_narrowing_samples.partial_term_of_sample_def text \ The separate code target \SML_word\ collects setups for the code generator that PolyML does not provide. \ setup \Code_Target.add_derived_target ("SML_word", [(Code_ML.target_SML, I)])\ code_identifier code_module Code_Target_Word_Base \ (SML) Word and (Haskell) Word and (OCaml) Word and (Scala) Word end diff --git a/thys/PAC_Checker/PAC_Checker_Relation.thy b/thys/PAC_Checker/PAC_Checker_Relation.thy --- a/thys/PAC_Checker/PAC_Checker_Relation.thy +++ b/thys/PAC_Checker/PAC_Checker_Relation.thy @@ -1,389 +1,389 @@ (* File: PAC_Checker_Relation.thy Author: Mathias Fleury, Daniela Kaufmann, JKU Maintainer: Mathias Fleury, JKU *) theory PAC_Checker_Relation imports PAC_Checker WB_Sort "Native_Word.Uint64" begin section \Various Refinement Relations\ text \When writing this, it was not possible to share the definition with the IsaSAT version.\ definition uint64_nat_rel :: "(uint64 \ nat) set" where \uint64_nat_rel = br nat_of_uint64 (\_. True)\ abbreviation uint64_nat_assn where \uint64_nat_assn \ pure uint64_nat_rel\ instantiation uint32 :: hashable begin definition hashcode_uint32 :: \uint32 \ uint32\ where \hashcode_uint32 n = n\ definition def_hashmap_size_uint32 :: \uint32 itself \ nat\ where \def_hashmap_size_uint32 = (\_. 16)\ \ \same as @{typ nat}\ instance by standard (simp add: def_hashmap_size_uint32_def) end instantiation uint64 :: hashable begin context includes bit_operations_syntax begin definition hashcode_uint64 :: \uint64 \ uint32\ where \hashcode_uint64 n = (uint32_of_nat (nat_of_uint64 ((n) AND ((2 :: uint64)^32 -1))))\ end definition def_hashmap_size_uint64 :: \uint64 itself \ nat\ where \def_hashmap_size_uint64 = (\_. 16)\ \ \same as @{typ nat}\ instance by standard (simp add: def_hashmap_size_uint64_def) end lemma word_nat_of_uint64_Rep_inject[simp]: \nat_of_uint64 ai = nat_of_uint64 bi \ ai = bi\ by transfer (simp add: word_unat_eq_iff) instance uint64 :: heap by standard (auto simp: inj_def exI[of _ nat_of_uint64]) instance uint64 :: semiring_numeral by standard lemma nat_of_uint64_012[simp]: \nat_of_uint64 0 = 0\ \nat_of_uint64 2 = 2\ \nat_of_uint64 1 = 1\ by (simp_all add: nat_of_uint64.rep_eq zero_uint64.rep_eq one_uint64.rep_eq) definition uint64_of_nat_conv where [simp]: \uint64_of_nat_conv (x :: nat) = x\ lemma less_upper_bintrunc_id: \n < 2 ^b \ n \ 0 \ take_bit b n = n\ for n :: int by (rule take_bit_int_eq_self) lemma nat_of_uint64_uint64_of_nat_id: \n < 2^64 \ nat_of_uint64 (uint64_of_nat n) = n\ - by transfer (simp add: take_bit_nat_eq_self) + by transfer (simp add: take_bit_nat_eq_self unsigned_of_nat) lemma [sepref_fr_rules]: \(return o uint64_of_nat, RETURN o uint64_of_nat_conv) \ [\a. a < 2 ^64]\<^sub>a nat_assn\<^sup>k \ uint64_nat_assn\ by sepref_to_hoare (sep_auto simp: uint64_nat_rel_def br_def nat_of_uint64_uint64_of_nat_id) definition string_rel :: \(String.literal \ string) set\ where \string_rel = {(x, y). y = String.explode x}\ abbreviation string_assn :: \string \ String.literal \ assn\ where \string_assn \ pure string_rel\ lemma eq_string_eq: \((=), (=)) \ string_rel \ string_rel \ bool_rel\ by (auto intro!: frefI simp: string_rel_def String.less_literal_def less_than_char_def rel2p_def literal.explode_inject) lemmas eq_string_eq_hnr = eq_string_eq[sepref_import_param] definition string2_rel :: \(string \ string) set\ where \string2_rel \ \Id\list_rel\ abbreviation string2_assn :: \string \ string \ assn\ where \string2_assn \ pure string2_rel\ abbreviation monom_rel where \monom_rel \ \string_rel\list_rel\ abbreviation monom_assn where \monom_assn \ list_assn string_assn\ abbreviation monomial_rel where \monomial_rel \ monom_rel \\<^sub>r int_rel\ abbreviation monomial_assn where \monomial_assn \ monom_assn \\<^sub>a int_assn\ abbreviation poly_rel where \poly_rel \ \monomial_rel\list_rel\ abbreviation poly_assn where \poly_assn \ list_assn monomial_assn\ lemma poly_assn_alt_def: \poly_assn = pure poly_rel\ by (simp add: list_assn_pure_conv) abbreviation polys_assn where \polys_assn \ hm_fmap_assn uint64_nat_assn poly_assn\ lemma string_rel_string_assn: \(\ ((c, a) \ string_rel)) = string_assn a c\ by (auto simp: pure_app_eq) lemma single_valued_string_rel: \single_valued string_rel\ by (auto simp: single_valued_def string_rel_def) lemma IS_LEFT_UNIQUE_string_rel: \IS_LEFT_UNIQUE string_rel\ by (auto simp: IS_LEFT_UNIQUE_def single_valued_def string_rel_def literal.explode_inject) lemma IS_RIGHT_UNIQUE_string_rel: \IS_RIGHT_UNIQUE string_rel\ by (auto simp: single_valued_def string_rel_def literal.explode_inject) lemma single_valued_monom_rel: \single_valued monom_rel\ by (rule list_rel_sv) (auto intro!: frefI simp: string_rel_def rel2p_def single_valued_def p2rel_def) lemma single_valued_monomial_rel: \single_valued monomial_rel\ using single_valued_monom_rel by (auto intro!: frefI simp: rel2p_def single_valued_def p2rel_def) lemma single_valued_monom_rel': \IS_LEFT_UNIQUE monom_rel\ unfolding IS_LEFT_UNIQUE_def inv_list_rel_eq string2_rel_def by (rule list_rel_sv)+ (auto intro!: frefI simp: string_rel_def rel2p_def single_valued_def p2rel_def literal.explode_inject) lemma single_valued_monomial_rel': \IS_LEFT_UNIQUE monomial_rel\ using single_valued_monom_rel' unfolding IS_LEFT_UNIQUE_def inv_list_rel_eq by (auto intro!: frefI simp: rel2p_def single_valued_def p2rel_def) lemma [safe_constraint_rules]: \Sepref_Constraints.CONSTRAINT single_valued string_rel\ \Sepref_Constraints.CONSTRAINT IS_LEFT_UNIQUE string_rel\ by (auto simp: CONSTRAINT_def single_valued_def string_rel_def IS_LEFT_UNIQUE_def literal.explode_inject) lemma eq_string_monom_hnr[sepref_fr_rules]: \(uncurry (return oo (=)), uncurry (RETURN oo (=))) \ monom_assn\<^sup>k *\<^sub>a monom_assn\<^sup>k \\<^sub>a bool_assn\ using single_valued_monom_rel' single_valued_monom_rel unfolding list_assn_pure_conv by sepref_to_hoare (sep_auto simp: list_assn_pure_conv string_rel_string_assn single_valued_def IS_LEFT_UNIQUE_def dest!: mod_starD simp flip: inv_list_rel_eq) definition term_order_rel' where [simp]: \term_order_rel' x y = ((x, y) \ term_order_rel)\ lemma term_order_rel[def_pat_rules]: \(\)$(x,y)$term_order_rel \ term_order_rel'$x$y\ by auto lemma term_order_rel_alt_def: \term_order_rel = lexord (p2rel char.lexordp)\ by (auto simp: p2rel_def char.lexordp_conv_lexord var_order_rel_def intro!: arg_cong[of _ _ lexord]) instantiation char :: linorder begin definition less_char where [symmetric, simp]: "less_char = PAC_Polynomials_Term.less_char" definition less_eq_char where [symmetric, simp]: "less_eq_char = PAC_Polynomials_Term.less_eq_char" instance apply standard using char.linorder_axioms by (auto simp: class.linorder_def class.order_def class.preorder_def less_eq_char_def less_than_char_def class.order_axioms_def class.linorder_axioms_def p2rel_def less_char_def) end instantiation list :: (linorder) linorder begin definition less_list where "less_list = lexordp (<)" definition less_eq_list where "less_eq_list = lexordp_eq" instance proof standard have [dest]: \\x y :: 'a :: linorder list. (x, y) \ lexord {(x, y). x < y} \ lexordp_eq y x \ False\ by (metis lexordp_antisym lexordp_conv_lexord lexordp_eq_conv_lexord) have [simp]: \\x y :: 'a :: linorder list. lexordp_eq x y \ \ lexordp_eq y x \ (x, y) \ lexord {(x, y). x < y}\ using lexordp_conv_lexord lexordp_conv_lexordp_eq by blast show \(x < y) = Restricted_Predicates.strict (\) x y\ \x \ x\ \x \ y \ y \ z \ x \ z\ \x \ y \ y \ x \ x = y\ \x \ y \ y \ x\ for x y z :: \'a :: linorder list\ by (auto simp: less_list_def less_eq_list_def List.lexordp_def lexordp_conv_lexord lexordp_into_lexordp_eq lexordp_antisym antisym_def lexordp_eq_refl lexordp_eq_linear intro: lexordp_eq_trans dest: lexordp_eq_antisym) qed end lemma term_order_rel'_alt_def_lexord: \term_order_rel' x y = ord_class.lexordp x y\ and term_order_rel'_alt_def: \term_order_rel' x y \ x < y\ proof - show \term_order_rel' x y = ord_class.lexordp x y\ \term_order_rel' x y \ x < y\ unfolding less_than_char_of_char[symmetric, abs_def] by (auto simp: lexordp_conv_lexord less_eq_list_def less_list_def lexordp_def var_order_rel_def rel2p_def term_order_rel_alt_def p2rel_def) qed lemma list_rel_list_rel_order_iff: assumes \(a, b) \ \string_rel\list_rel\ \(a', b') \ \string_rel\list_rel\ shows \a < a' \ b < b'\ proof have H: \(a, b) \ \string_rel\list_rel \ (a, cs) \ \string_rel\list_rel \ b = cs\ for cs using single_valued_monom_rel' IS_RIGHT_UNIQUE_string_rel unfolding string2_rel_def by (subst (asm)list_rel_sv_iff[symmetric]) (auto simp: single_valued_def) assume \a < a'\ then consider u u' where \a' = a @ u # u'\ | u aa v w aaa where \a = u @ aa # v\ \a' = u @ aaa # w\ \aa < aaa\ by (subst (asm) less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) then show \b < b'\ proof cases case 1 then show \b < b'\ using assms by (subst less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff dest: H) next case 2 then obtain u' aa' v' w' aaa' where \b = u' @ aa' # v'\ \b' = u' @ aaa' # w'\ \(aa, aa') \ string_rel\ \(aaa, aaa') \ string_rel\ using assms by (smt list_rel_append1 list_rel_split_right_iff single_valued_def single_valued_monom_rel) with \aa < aaa\ have \aa' < aaa'\ by (auto simp: string_rel_def less_literal.rep_eq less_list_def lexordp_conv_lexord lexordp_def char.lexordp_conv_lexord simp flip: less_char_def PAC_Polynomials_Term.less_char_def) then show \b < b'\ using \b = u' @ aa' # v'\ \b' = u' @ aaa' # w'\ by (subst less_list_def) (fastforce simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) qed next have H: \(a, b) \ \string_rel\list_rel \ (a', b) \ \string_rel\list_rel \ a = a'\ for a a' b using single_valued_monom_rel' by (auto simp: single_valued_def IS_LEFT_UNIQUE_def simp flip: inv_list_rel_eq) assume \b < b'\ then consider u u' where \b' = b @ u # u'\ | u aa v w aaa where \b = u @ aa # v\ \b' = u @ aaa # w\ \aa < aaa\ by (subst (asm) less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) then show \a < a'\ proof cases case 1 then show \a < a'\ using assms by (subst less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append2 list_rel_split_left_iff dest: H) next case 2 then obtain u' aa' v' w' aaa' where \a = u' @ aa' # v'\ \a' = u' @ aaa' # w'\ \(aa', aa) \ string_rel\ \(aaa', aaa) \ string_rel\ using assms by (auto simp: lexord_def List.lexordp_def list_rel_append2 list_rel_split_left_iff dest: H) with \aa < aaa\ have \aa' < aaa'\ by (auto simp: string_rel_def less_literal.rep_eq less_list_def lexordp_conv_lexord lexordp_def char.lexordp_conv_lexord simp flip: less_char_def PAC_Polynomials_Term.less_char_def) then show \a < a'\ using \a = u' @ aa' # v'\ \a' = u' @ aaa' # w'\ by (subst less_list_def) (fastforce simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) qed qed lemma string_rel_le[sepref_import_param]: shows \((<), (<)) \ \string_rel\list_rel \ \string_rel\list_rel \ bool_rel\ by (auto intro!: fun_relI simp: list_rel_list_rel_order_iff) (* TODO Move *) lemma [sepref_import_param]: assumes \CONSTRAINT IS_LEFT_UNIQUE R\ \CONSTRAINT IS_RIGHT_UNIQUE R\ shows \(remove1, remove1) \ R \ \R\list_rel \ \R\list_rel\ apply (intro fun_relI) subgoal premises p for x y xs ys using p(2) p(1) assms by (induction xs ys rule: list_rel_induct) (auto simp: IS_LEFT_UNIQUE_def single_valued_def) done instantiation pac_step :: (heap, heap, heap) heap begin instance proof standard obtain f :: \'a \ nat\ where f: \inj f\ by blast obtain g :: \nat \ nat \ nat \ nat \ nat \ nat\ where g: \inj g\ by blast obtain h :: \'b \ nat\ where h: \inj h\ by blast obtain i :: \'c \ nat\ where i: \inj i\ by blast have [iff]: \g a = g b \ a = b\\h a'' = h b'' \ a'' = b''\ \f a' = f b' \ a' = b'\ \i a''' = i b''' \ a''' = b'''\ for a b a' b' a'' b'' a''' b''' using f g h i unfolding inj_def by blast+ let ?f = \\x :: ('a, 'b, 'c) pac_step. g (case x of Add a b c d \ (0, i a, i b, i c, f d) | Del a \ (1, i a, 0, 0, 0) | Mult a b c d \ (2, i a, f b, i c, f d) | Extension a b c \ (3, i a, f c, 0, h b))\ have \inj ?f\ apply (auto simp: inj_def) apply (case_tac x; case_tac y) apply auto done then show \\f :: ('a, 'b, 'c) pac_step \ nat. inj f\ by blast qed end end \ No newline at end of file diff --git a/thys/SPARCv8/SparcModel_MMU/Sparc_Properties.thy b/thys/SPARCv8/SparcModel_MMU/Sparc_Properties.thy --- a/thys/SPARCv8/SparcModel_MMU/Sparc_Properties.thy +++ b/thys/SPARCv8/SparcModel_MMU/Sparc_Properties.thy @@ -1,8588 +1,8589 @@ (* * Copyright 2016, NTU * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * Author: Zhe Hou. *) theory Sparc_Properties imports Main Sparc_Execution begin (*********************************************************************) section\Single step theorem\ (*********************************************************************) text \The following shows that, if the pre-state satisfies certain conditions called \good_context\, there must be a defined post-state after a single step execution.\ method save_restore_proof = ((simp add: save_restore_instr_def), (simp add: Let_def simpler_gets_def bind_def h1_def h2_def), (simp add: case_prod_unfold), (simp add: raise_trap_def simpler_modify_def), (simp add: simpler_gets_def bind_def h1_def h2_def), (simp add: save_retore_sub1_def), (simp add: write_cpu_def simpler_modify_def), (simp add: write_reg_def simpler_modify_def), (simp add: get_curr_win_def), (simp add: simpler_gets_def bind_def h1_def h2_def)) method select_trap_proof0 = ((simp add: select_trap_def exec_gets return_def), (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def), (simp add: write_cpu_tt_def write_cpu_def), (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def), (simp add: return_def simpler_gets_def)) method select_trap_proof1 = ((simp add: select_trap_def exec_gets return_def), (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def), (simp add: write_cpu_tt_def write_cpu_def), (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def), (simp add: return_def simpler_gets_def), (simp add: emp_trap_set_def err_mode_val_def cpu_reg_mod_def)) method dispatch_instr_proof1 = ((simp add: dispatch_instruction_def), (simp add: simpler_gets_def bind_def h1_def h2_def), (simp add: Let_def)) method exe_proof_to_decode = ((simp add: execute_instruction_def), (simp add: exec_gets bind_def h1_def h2_def Let_def return_def), clarsimp, (simp add: simpler_gets_def bind_def h1_def h2_def Let_def simpler_modify_def), (simp add: return_def)) method exe_proof_dispatch_rett = ((simp add: dispatch_instruction_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: rett_instr_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def)) lemma write_cpu_result: "snd (write_cpu w r s) = False" by (simp add: write_cpu_def simpler_modify_def) lemma set_annul_result: "snd (set_annul b s) = False" by (simp add: set_annul_def simpler_modify_def) lemma raise_trap_result : "snd (raise_trap t s) = False" by (simp add: raise_trap_def simpler_modify_def) context includes bit_operations_syntax begin lemma rett_instr_result: "(fst i) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR s) \ 1 \ (((get_S (cpu_reg_val PSR s)))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s)) = 0 \ ((AND) (get_addr (snd i) s) (0b00000000000000000000000000000011::word32)) = 0) \ snd (rett_instr i s) = False" apply (simp add: rett_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (auto simp add: Let_def return_def) done lemma call_instr_result: "(fst i) = call_type CALL \ snd (call_instr i s) = False" apply (simp add: call_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def case_prod_unfold) apply (simp add: write_cpu_def write_reg_def) apply (simp add: get_curr_win_def get_CWP_def) by (simp add: simpler_modify_def simpler_gets_def) lemma branch_instr_result: "(fst i) \ {bicc_type BE,bicc_type BNE,bicc_type BGU, bicc_type BLE,bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA,bicc_type BN} \ snd (branch_instr i s) = False" proof (cases "eval_icc (fst i) (get_icc_N ((cpu_reg s) PSR)) (get_icc_Z ((cpu_reg s) PSR)) (get_icc_V ((cpu_reg s) PSR)) (get_icc_C ((cpu_reg s) PSR)) = 1") case True then have f1: "eval_icc (fst i) (get_icc_N ((cpu_reg s) PSR)) (get_icc_Z ((cpu_reg s) PSR)) (get_icc_V ((cpu_reg s) PSR)) (get_icc_C ((cpu_reg s) PSR)) = 1" by auto then show ?thesis proof (cases "(fst i) = bicc_type BA \ get_operand_flag ((snd i)!0) = 1") case True then show ?thesis using f1 apply (simp add: branch_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: set_annul_def case_prod_unfold) apply (simp add: write_cpu_def simpler_modify_def) by (simp add: return_def) next case False then have f2: "\ (fst i = bicc_type BA \ get_operand_flag (snd i ! 0) = 1)" by auto then show ?thesis using f1 apply (simp add: branch_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: branch_instr_sub1_def) apply (simp add: Let_def) apply auto apply (simp add: write_cpu_def simpler_modify_def) by (simp add: write_cpu_def simpler_modify_def) qed next case False then show ?thesis apply (simp add: branch_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: branch_instr_sub1_def) apply (simp add: Let_def) apply auto apply (simp add: Let_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def set_annul_def simpler_modify_def) by (simp add: write_cpu_def simpler_modify_def) qed lemma nop_instr_result: "(fst i) = nop_type NOP \ snd (nop_instr i s) = False" apply (simp add: nop_instr_def) by (simp add: returnOk_def return_def) lemma sethi_instr_result: "(fst i) = sethi_type SETHI \ snd (sethi_instr i s) = False" apply (simp add: sethi_instr_def) apply (simp add: Let_def) apply (simp add: get_curr_win_def get_CWP_def cpu_reg_val_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_reg_def simpler_modify_def) by (simp add: return_def) lemma jmpl_instr_result: "(fst i) = ctrl_type JMPL \ snd (jmpl_instr i s) = False" apply (simp add: jmpl_instr_def) apply (simp add: get_curr_win_def get_CWP_def cpu_reg_val_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: write_cpu_def simpler_modify_def) by (simp add: raise_trap_def simpler_modify_def) lemma save_restore_instr_result: "(fst i) \ {ctrl_type SAVE,ctrl_type RESTORE} \ snd (save_restore_instr i s) = False" proof (cases "(fst i) = ctrl_type SAVE") case True then show ?thesis by save_restore_proof next case False then show ?thesis by save_restore_proof qed lemma flush_instr_result: "(fst i) = load_store_type FLUSH \ snd (flush_instr i s) = False" apply (simp add: flush_instr_def) by (simp add: simpler_gets_def bind_def h1_def h2_def simpler_modify_def) lemma read_state_reg_instr_result: "(fst i) \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM,sreg_type RDTBR} \ snd (read_state_reg_instr i s) = False" apply (simp add: read_state_reg_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def bind_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: raise_trap_def simpler_modify_def return_def) apply (simp add: bind_def h1_def h2_def) by (simp add: get_curr_win_def simpler_gets_def) lemma write_state_reg_instr_result: "(fst i) \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM,sreg_type WRTBR} \ snd (write_state_reg_instr i s) = False" apply (simp add: write_state_reg_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_modify_def) apply (simp add: raise_trap_def simpler_modify_def return_def) apply (simp add: bind_def h1_def h2_def) apply (simp add: simpler_gets_def) apply (simp add: write_cpu_def simpler_modify_def) by (simp add: get_curr_win_def simpler_gets_def) lemma logical_instr_result: "(fst i) \ {logic_type ANDs,logic_type ANDcc, logic_type ANDN,logic_type ANDNcc,logic_type ORs,logic_type ORcc, logic_type ORN,logic_type XORs,logic_type XNOR} \ snd (logical_instr i s) = False" apply (simp add: logical_instr_def) apply (simp add: Let_def simpler_gets_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: logical_instr_sub1_def) apply (simp add: return_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def) by (simp add: get_curr_win_def simpler_gets_def) lemma shift_instr_result: "(fst i) \ {shift_type SLL,shift_type SRL,shift_type SRA} \ snd (shift_instr i s) = False" apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: get_curr_win_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: return_def) apply (simp add: bind_def h1_def h2_def) by (simp add: write_reg_def simpler_modify_def) method add_sub_instr_proof = ((simp add: Let_def), auto, (simp add: write_reg_def simpler_modify_def), (simp add: simpler_gets_def bind_def), (simp add: get_curr_win_def simpler_gets_def), (simp add: write_reg_def write_cpu_def simpler_modify_def), (simp add: bind_def), (simp add: case_prod_unfold), (simp add: simpler_gets_def), (simp add: get_curr_win_def simpler_gets_def), (simp add: write_reg_def simpler_modify_def), (simp add: simpler_gets_def bind_def), (simp add: get_curr_win_def simpler_gets_def)) lemma add_instr_result: "(fst i) \ {arith_type ADD,arith_type ADDcc,arith_type ADDX} \ snd (add_instr i s) = False" apply (simp add: add_instr_def) apply (simp add: Let_def) apply auto apply (simp add: add_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: add_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: add_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) by (simp add: write_reg_def simpler_modify_def) lemma sub_instr_result: "(fst i) \ {arith_type SUB,arith_type SUBcc, arith_type SUBX} \ snd (sub_instr i s) = False" apply (simp add: sub_instr_def) apply (simp add: Let_def) apply auto apply (simp add: sub_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: sub_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: sub_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) by (simp add: write_reg_def simpler_modify_def) lemma mul_instr_result: "(fst i) \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc} \ snd (mul_instr i s) = False" apply (simp add: mul_instr_def) apply (simp add: Let_def) apply auto apply (simp add: mul_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: write_reg_def write_cpu_def simpler_modify_def) apply (simp add: mul_instr_sub1_def) apply (simp add: simpler_gets_def) apply (simp add: write_cpu_def write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: mul_instr_sub1_def) apply (simp add: simpler_gets_def) apply (simp add: write_cpu_def write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def) by (simp add: get_curr_win_def simpler_gets_def) lemma div_write_new_val_result: "snd (div_write_new_val i result temp_V s) = False" apply (simp add: div_write_new_val_def) apply (simp add: return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) by (simp add: write_cpu_def simpler_modify_def) lemma div_result: "snd (div_comp instr rs1 rd operand2 s) = False" apply (simp add: div_comp_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: get_curr_win_def simpler_gets_def) by (simp add: div_write_new_val_result) lemma div_instr_result: "(fst i) \ {arith_type UDIV,arith_type UDIVcc, arith_type SDIV} \ snd (div_instr i s) = False" apply (simp add: div_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: return_def bind_def) by (simp add: div_result) lemma load_sub2_result: "snd (load_sub2 address asi rd curr_win word0 s) = False" apply (simp add: load_sub2_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def) apply (simp add: write_reg_def simpler_modify_def) by (simp add: simpler_gets_def) lemma load_sub3_result: "snd (load_sub3 instr curr_win rd asi address s) = False" apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: load_sub2_result) by (simp add: raise_trap_def simpler_modify_def) lemma load_sub1_result: "snd (load_sub1 i rd s_val s) = False" apply (simp add: load_sub1_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: get_curr_win_def simpler_gets_def) by (simp add: load_sub3_result) lemma load_instr_result: "(fst i) \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD} \ snd (load_instr i s) = False" apply (simp add: load_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: return_def) by (simp add: load_sub1_result) lemma store_sub2_result: "snd (store_sub2 instr curr_win rd asi address s) = False" apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: return_def) apply (simp add: raise_trap_def simpler_modify_def) by (simp add: bind_def h1_def h2_def) lemma store_sub1_result: "snd (store_sub1 instr rd s_val s) = False" apply (simp add: store_sub1_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: get_curr_win_def) apply (simp add: simpler_gets_def) by (simp add: store_sub2_result) lemma store_instr_result: "(fst i) \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD} \ snd (store_instr i s) = False" apply (simp add: store_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: return_def) by (simp add: store_sub1_result) lemma supported_instr_set: "supported_instruction i = True \ i \ {load_store_type LDSB,load_store_type LDUB,load_store_type LDUBA, load_store_type LDUH,load_store_type LD,load_store_type LDA, load_store_type LDD, load_store_type STB,load_store_type STH,load_store_type ST, load_store_type STA,load_store_type STD, sethi_type SETHI, nop_type NOP, logic_type ANDs,logic_type ANDcc,logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, ctrl_type RETT, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" apply (simp add: supported_instruction_def) by presburger lemma dispatch_instr_result: assumes a1: "supported_instruction (fst i) = True \ (fst i) \ ctrl_type RETT" shows "snd (dispatch_instruction i s) = False" proof (cases "get_trap_set s = {}") case True then have f1: "get_trap_set s = {}" by auto then show ?thesis proof (cases "(fst i) \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (simp add: load_instr_result) next case False then have f2: "(fst i) \ {load_store_type STB,load_store_type STH,load_store_type ST, load_store_type STA,load_store_type STD, sethi_type SETHI, nop_type NOP, logic_type ANDs,logic_type ANDcc,logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using a1 apply (simp add: supported_instruction_def) by presburger then show ?thesis proof (cases "(fst i) \ {load_store_type STB,load_store_type STH, load_store_type ST, load_store_type STA,load_store_type STD}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: store_instr_result) next case False then have f3: "(fst i) \ {sethi_type SETHI, nop_type NOP, logic_type ANDs,logic_type ANDcc,logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f2 by auto then show ?thesis proof (cases "(fst i) = sethi_type SETHI") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (simp add: sethi_instr_result) next case False then have f4: "(fst i) \ {nop_type NOP, logic_type ANDs,logic_type ANDcc,logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f3 by auto then show ?thesis proof (cases "fst i = nop_type NOP") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (simp add: nop_instr_result) next case False then have f5: "(fst i) \ {logic_type ANDs,logic_type ANDcc, logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f4 by auto then show ?thesis proof (cases "(fst i) \ {logic_type ANDs,logic_type ANDcc, logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: logical_instr_result) next case False then have f6: "(fst i) \ {shift_type SLL,shift_type SRL, shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f5 by auto then show ?thesis proof (cases "(fst i) \ {shift_type SLL,shift_type SRL, shift_type SRA}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: shift_instr_result) next case False then have f7: "(fst i) \ {arith_type ADD,arith_type ADDcc, arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f6 by auto then show ?thesis proof (cases "(fst i) \ {arith_type ADD,arith_type ADDcc, arith_type ADDX}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: add_instr_result) next case False then have f8: "(fst i) \ {arith_type SUB,arith_type SUBcc, arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f7 by auto then show ?thesis proof (cases "(fst i) \ {arith_type SUB,arith_type SUBcc, arith_type SUBX}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: sub_instr_result) next case False then have f9: "(fst i) \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f8 by auto then show ?thesis proof (cases "(fst i) \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: mul_instr_result) next case False then have f10: "(fst i) \ {arith_type UDIV,arith_type UDIVcc, arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f9 by auto then show ?thesis proof (cases "(fst i) \ {arith_type UDIV,arith_type UDIVcc, arith_type SDIV}") case True then show ?thesis apply dispatch_instr_proof1 using f1 by (auto simp add: div_instr_result) next case False then have f11: "(fst i) \ {ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f10 by auto then show ?thesis proof (cases "(fst i) \ {ctrl_type SAVE,ctrl_type RESTORE}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: save_restore_instr_result) next case False then have f12: "(fst i) \ {call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f11 by auto then show ?thesis proof (cases "(fst i) = call_type CALL") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: call_instr_result) next case False then have f13: "(fst i) \ {ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f12 by auto then show ?thesis proof (cases "(fst i) = ctrl_type JMPL") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: jmpl_instr_result) next case False then have f14: "(fst i) \ { sreg_type RDY, sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f13 by auto then show ?thesis proof (cases "(fst i) \ {sreg_type RDY, sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: read_state_reg_instr_result) next case False then have f15: "(fst i) \ { sreg_type WRY, sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f14 by auto then show ?thesis proof (cases "(fst i) \ {sreg_type WRY, sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: write_state_reg_instr_result) next case False then have f16: "(fst i) \ { load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f15 by auto then show ?thesis proof (cases "(fst i) = load_store_type FLUSH") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: flush_instr_result) next case False then have f17: "(fst i) \ { bicc_type BE,bicc_type BNE,bicc_type BGU, bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f16 by auto then show ?thesis using f1 proof (cases "(fst i) \ {bicc_type BE, bicc_type BNE,bicc_type BGU, bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA }") case True then show ?thesis using f1 apply dispatch_instr_proof1 apply auto by (auto simp add: branch_instr_result) next case False then have f18: "(fst i) \ {bicc_type BN}" using f17 by auto then show ?thesis using f1 apply dispatch_instr_proof1 apply auto by (auto simp add: branch_instr_result) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed next case False then show ?thesis apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) by (simp add: returnOk_def return_def) qed lemma dispatch_instr_result_rett: assumes a1: "(fst i) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR s) \ 1 \ (((get_S (cpu_reg_val PSR s)))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s)) = 0 \ ((AND) (get_addr (snd i) s) (0b00000000000000000000000000000011::word32)) = 0)" shows "snd (dispatch_instruction i s) = False" proof (cases "get_trap_set s = {}") case True then show ?thesis using a1 apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: rett_instr_result) next case False then show ?thesis using a1 apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: return_def) qed lemma execute_instr_sub1_result: "snd (execute_instr_sub1 i s) = False" proof (cases "get_trap_set s = {} \ (fst i) \ {call_type CALL,ctrl_type RETT, ctrl_type JMPL}") case True then show ?thesis apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply auto by (auto simp add: return_def) next case False then show ?thesis apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) by (auto simp add: return_def) qed lemma next_match : "snd (execute_instruction () s) = False \ NEXT s = Some (snd (fst (execute_instruction () s)))" apply (simp add: NEXT_def) by (simp add: case_prod_unfold) lemma exec_ss1 : "\s'. (execute_instruction () s = (s', False)) \ \s''. (execute_instruction() s = (s'', False))" proof - assume "\s'. (execute_instruction () s = (s', False))" hence "(snd (execute_instruction() s)) = False" by (auto simp add: execute_instruction_def case_prod_unfold) hence "(execute_instruction() s) = ((fst (execute_instruction() s)),False)" by (metis (full_types) prod.collapse) hence "\s''. (execute_instruction() s = (s'', False))" by blast thus ?thesis by assumption qed lemma exec_ss2 : "snd (execute_instruction() s) = False \ snd (execute_instruction () s) = False" proof - assume "snd (execute_instruction() s) = False" hence "snd (execute_instruction () s) = False" by (auto simp add:execute_instruction_def) thus ?thesis by assumption qed lemma good_context_1 : "good_context s \ s' = s \ (get_trap_set s') \ {} \ (reset_trap_val s') = False \ get_ET (cpu_reg_val PSR s') = 0 \ False" proof - assume asm: "good_context s \ s' = s \ (get_trap_set s') \ {} \ (reset_trap_val s') = False \ get_ET (cpu_reg_val PSR s') = 0" then have "(get_trap_set s') \ {} \ (reset_trap_val s') = False \ get_ET (cpu_reg_val PSR s') = 0 \ False" by (simp add: good_context_def get_ET_def cpu_reg_val_def) then show ?thesis using asm by auto qed lemma fetch_instr_result_1 : "\ (\e. fetch_instruction s' = Inl e) \ (\v. fetch_instruction s' = Inr v)" by (meson sumE) lemma fetch_instr_result_2 : "(\v. fetch_instruction s' = Inr v) \ \ (\e. fetch_instruction s' = Inl e)" by force lemma fetch_instr_result_3 : "(\e. fetch_instruction s' = Inl e) \ \ (\v. fetch_instruction s' = Inr v)" by auto lemma decode_instr_result_1 : "\(\v2. ((decode_instruction v1)::(Exception list + instruction)) = Inr v2) \ (\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e)" by (meson sumE) lemma decode_instr_result_2 : "(\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e) \ \(\v2. ((decode_instruction v1)::(Exception list + instruction)) = Inr v2)" by force lemma decode_instr_result_3 : "x = decode_instruction v1 \ y = decode_instruction v2 \ v1 = v2 \ x = y" by auto lemma decode_instr_result_4 : "\ (\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e) \ (\v2. ((decode_instruction v1)::(Exception list + instruction)) = Inr v2)" by (meson sumE) lemma good_context_2 : "good_context (s::(('a::len) sparc_state)) \ fetch_instruction (delayed_pool_write s) = Inr v1 \ \(\v2. (decode_instruction v1::(Exception list + instruction)) = Inr v2) \ False" proof - assume "good_context s \ fetch_instruction (delayed_pool_write s) = Inr v1 \ \(\v2. ((decode_instruction v1)::(Exception list + instruction)) = Inr v2)" hence fact1: "good_context s \ fetch_instruction (delayed_pool_write s) = Inr v1 \ (\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e)" using decode_instr_result_1 by auto hence fact2: "\(\e. fetch_instruction (delayed_pool_write s) = Inl e)" using fetch_instr_result_2 by auto then have "fetch_instruction (delayed_pool_write s) = Inr v1 \ (\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e) \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this fact1 show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto then show ?thesis using fact1 decode_instr_result_3 by (metis (no_types, lifting) good_context_def sum.case(1) sum.case(2)) qed thus ?thesis using fact1 by auto qed lemma good_context_3 : "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ (decode_instruction v1::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = False \ False" proof - assume asm: "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ (decode_instruction v1::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = False" then have "annul_val s'' = False \ supported_instruction (fst v2) = False \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this asm show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto thus ?thesis using asm by (auto simp add: good_context_def) qed thus ?thesis using asm by auto qed lemma good_context_4 : "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) = 0 \ False" proof - assume asm: "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) = 0" then have "(fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) = 0 \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this asm show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto thus ?thesis using asm by (auto simp add: good_context_def) qed thus ?thesis using asm by auto qed lemma good_context_5 : "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) \ 0 \ False" proof - assume asm: "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) \ 0" then have "(fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) \ 0 \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this asm show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto thus ?thesis using asm by (auto simp add: good_context_def) qed thus ?thesis using asm by auto qed lemma good_context_6 : "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) \ 0 \ False" proof - assume asm: "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) \ 0" then have "(fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) \ 0 \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this asm show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto thus ?thesis using asm by (auto simp add: good_context_def) qed thus ?thesis using asm by auto qed lemma good_context_all : "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ (get_trap_set s = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0) \ ((\e. fetch_instruction s'' = Inl e) \ (\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (annul_val s'' = True \ (annul_val s'' = False \ (\v1' v2'. fetch_instruction s'' = Inr v1' \ ((decode_instruction v1')::(Exception list + instruction)) = Inr v2' \ supported_instruction (fst v2') = True) \ ((fst v2) \ ctrl_type RETT \ ((fst v2) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR s'') = 1 \ (get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) = 0))))))))" proof - assume asm: "good_context s \ s'' = delayed_pool_write s" from asm have "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0 \ False" using good_context_1 by blast hence fact1: "(get_trap_set s = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0)" by auto have fact2: "\(\e. fetch_instruction s'' = Inl e) \ \ (\v1. fetch_instruction s'' = Inr v1) \ False" using fetch_instr_result_1 by blast from asm have fact3: "\v1. fetch_instruction s'' = Inr v1 \ \(\v2.((decode_instruction v1)::(Exception list + instruction)) = Inr v2) \ False" using good_context_2 by blast from asm have fact4: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = False \ False" using good_context_3 by blast from asm have fact5: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) = 0 \ False" using good_context_4 by blast from asm have fact6: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) \ 0 \ False" using good_context_5 by blast from asm have fact7: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) \ 0 \ False" using good_context_6 by blast from asm show ?thesis proof (cases "(\e. fetch_instruction s'' = Inl e)") case True then show ?thesis using fact1 by auto next case False then have fact8: "\v1. fetch_instruction s'' = Inr v1 \ (\v2.((decode_instruction v1)::(Exception list + instruction)) = Inr v2)" using fact2 fact3 by auto then show ?thesis proof (cases "annul_val s'' = True") case True then show ?thesis using fact1 fact8 by auto next case False then have fact9: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True" using fact4 fact8 by blast then show ?thesis proof (cases "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) \ ctrl_type RETT") case True then show ?thesis using fact1 fact9 by auto next case False then have fact10: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ (fst v2) = ctrl_type RETT" using fact9 by auto then show ?thesis proof (cases "get_ET (cpu_reg_val PSR s'') = 1") case True then show ?thesis using fact1 fact9 by auto next case False then have fact11: "get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0" using fact10 fact5 by auto then have fact12: "(get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0" using fact10 fact6 by auto then have fact13: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) = 0" using fact10 fact11 fact7 by blast thus ?thesis using fact1 fact10 fact11 fact12 by auto qed qed qed qed qed lemma select_trap_result1 : "(reset_trap_val s) = True \ snd (select_trap() s) = False" apply (simp add: select_trap_def exec_gets return_def) by (simp add: bind_def h1_def h2_def simpler_modify_def) lemma select_trap_result2 : assumes a1: "\(reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0)" shows "snd (select_trap() s) = False" proof (cases "reset_trap_val s = True") case True then show ?thesis using select_trap_result1 by blast next case False then have f1: "reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) \ 0" using a1 by auto then show ?thesis proof (cases "data_store_error \ get_trap_set s") case True then show ?thesis using f1 by select_trap_proof0 next case False then have f2: "data_store_error \ get_trap_set s" by auto then show ?thesis proof (cases "instruction_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 by select_trap_proof0 next case False then have f3: "instruction_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "r_register_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 f3 by select_trap_proof0 next case False then have f4: "r_register_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "instruction_access_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 by select_trap_proof0 next case False then have f5: "instruction_access_exception \ get_trap_set s" by auto then show ?thesis proof (cases "privileged_instruction \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 by select_trap_proof0 next case False then have f6: "privileged_instruction \ get_trap_set s" by auto then show ?thesis proof (cases "illegal_instruction \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 by select_trap_proof0 next case False then have f7: "illegal_instruction \ get_trap_set s" by auto then show ?thesis proof (cases "fp_disabled \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 by select_trap_proof0 next case False then have f8: "fp_disabled \ get_trap_set s" by auto then show ?thesis proof (cases "cp_disabled \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 by select_trap_proof0 next case False then have f9: "cp_disabled \ get_trap_set s" by auto then show ?thesis proof (cases "unimplemented_FLUSH \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 by select_trap_proof0 next case False then have f10: "unimplemented_FLUSH \ get_trap_set s" by auto then show ?thesis proof (cases "window_overflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 by select_trap_proof0 next case False then have f11: "window_overflow \ get_trap_set s" by auto then show ?thesis proof (cases "window_underflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 by select_trap_proof0 next case False then have f12: "window_underflow \ get_trap_set s" by auto then show ?thesis proof (cases "mem_address_not_aligned \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 by select_trap_proof0 next case False then have f13: "mem_address_not_aligned \ get_trap_set s" by auto then show ?thesis proof (cases "fp_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 by select_trap_proof0 next case False then have f14: "fp_exception \ get_trap_set s" by auto then show ?thesis proof (cases "cp_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 by select_trap_proof0 next case False then have f15: "cp_exception \ get_trap_set s" by auto then show ?thesis proof (cases "data_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 by select_trap_proof0 next case False then have f16: "data_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "data_access_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 by select_trap_proof0 next case False then have f17: "data_access_exception \ get_trap_set s" by auto then show ?thesis proof (cases "tag_overflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 by select_trap_proof0 next case False then have f18: "tag_overflow \ get_trap_set s" by auto then show ?thesis proof (cases "division_by_zero \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 by select_trap_proof0 next case False then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 apply (simp add: select_trap_def exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def) apply (simp add: return_def simpler_gets_def) apply (simp add: case_prod_unfold) apply (simp add: return_def) apply (simp add: write_cpu_tt_def write_cpu_def) by (simp add: simpler_gets_def bind_def h1_def h2_def simpler_modify_def) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed lemma emp_trap_set_err_mode : "err_mode_val s = err_mode_val (emp_trap_set s)" by (auto simp add: emp_trap_set_def err_mode_val_def) lemma write_cpu_tt_err_mode : "err_mode_val s = err_mode_val (snd (fst (write_cpu_tt w s)))" apply (simp add: write_cpu_tt_def err_mode_val_def write_cpu_def) apply (simp add: exec_gets return_def) apply (simp add: bind_def simpler_modify_def) by (simp add: cpu_reg_mod_def) lemma select_trap_monad : "snd (select_trap() s) = False \ err_mode_val s = err_mode_val (snd (fst (select_trap () s)))" proof - assume a1: "snd (select_trap() s) = False" then have f0: "reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0 \ False" apply (simp add: select_trap_def exec_gets return_def) apply (simp add: bind_def h1_def h2_def simpler_modify_def) by (simp add: fail_def split_def) then show ?thesis proof (cases "reset_trap_val s = True") case True from a1 f0 this show ?thesis apply (simp add: select_trap_def exec_gets return_def) apply (simp add: bind_def h1_def h2_def simpler_modify_def) by (simp add: emp_trap_set_def err_mode_val_def) next case False then have f1: "reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) \ 0" using f0 by auto then show ?thesis using f1 a1 proof (cases "data_store_error \ get_trap_set s") case True then show ?thesis using f1 a1 by select_trap_proof1 next case False then have f2: "data_store_error \ get_trap_set s" by auto then show ?thesis proof (cases "instruction_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 a1 by select_trap_proof1 next case False then have f3: "instruction_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "r_register_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 f3 a1 by select_trap_proof1 next case False then have f4: "r_register_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "instruction_access_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 a1 by select_trap_proof1 next case False then have f5: "instruction_access_exception \ get_trap_set s" by auto then show ?thesis proof (cases "privileged_instruction \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 a1 by select_trap_proof1 next case False then have f6: "privileged_instruction \ get_trap_set s" by auto then show ?thesis proof (cases "illegal_instruction \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 a1 by select_trap_proof1 next case False then have f7: "illegal_instruction \ get_trap_set s" by auto then show ?thesis proof (cases "fp_disabled \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 a1 by select_trap_proof1 next case False then have f8: "fp_disabled \ get_trap_set s" by auto then show ?thesis proof (cases "cp_disabled \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 a1 by select_trap_proof1 next case False then have f9: "cp_disabled \ get_trap_set s" by auto then show ?thesis proof (cases "unimplemented_FLUSH \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 a1 by select_trap_proof1 next case False then have f10: "unimplemented_FLUSH \ get_trap_set s" by auto then show ?thesis proof (cases "window_overflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 a1 by select_trap_proof1 next case False then have f11: "window_overflow \ get_trap_set s" by auto then show ?thesis proof (cases "window_underflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 a1 by select_trap_proof1 next case False then have f12: "window_underflow \ get_trap_set s" by auto then show ?thesis proof (cases "mem_address_not_aligned \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 a1 by select_trap_proof1 next case False then have f13: "mem_address_not_aligned \ get_trap_set s" by auto then show ?thesis proof (cases "fp_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 a1 by select_trap_proof1 next case False then have f14: "fp_exception \ get_trap_set s" by auto then show ?thesis proof (cases "cp_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 a1 by select_trap_proof1 next case False then have f15: "cp_exception \ get_trap_set s" by auto then show ?thesis proof (cases "data_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 a1 by select_trap_proof1 next case False then have f16: "data_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "data_access_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 a1 by select_trap_proof1 next case False then have f17: "data_access_exception \ get_trap_set s" by auto then show ?thesis proof (cases "tag_overflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 a1 by select_trap_proof1 next case False then have f18: "tag_overflow \ get_trap_set s" by auto then show ?thesis proof (cases "division_by_zero \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 a1 by select_trap_proof1 next case False then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 a1 apply (simp add: select_trap_def exec_gets return_def) apply (simp add: bind_def h1_def h2_def simpler_modify_def) apply (simp add: return_def simpler_gets_def) apply (simp add: emp_trap_set_def err_mode_val_def cpu_reg_mod_def) apply (simp add: case_prod_unfold) apply (simp add: return_def) apply clarsimp apply (simp add: write_cpu_tt_def write_cpu_def write_tt_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def) by (simp add: cpu_reg_val_def cpu_reg_mod_def) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed lemma exe_trap_st_pc_result : "snd (exe_trap_st_pc() s) = False" proof (cases "annul_val s = True") case True then show ?thesis apply (simp add: exe_trap_st_pc_def get_curr_win_def) apply (simp add: exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) by (simp add: set_annul_def write_reg_def simpler_modify_def) next case False then show ?thesis apply (simp add: exe_trap_st_pc_def get_curr_win_def) apply (simp add: exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) by (simp add: write_reg_def simpler_modify_def) qed lemma exe_trap_wr_pc_result : "snd (exe_trap_wr_pc() s) = False" proof (cases "reset_trap_val s = True") case True then show ?thesis apply (simp add: exe_trap_wr_pc_def get_curr_win_def) apply (simp add: exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: simpler_gets_def) apply (simp add: cpu_reg_val_def update_S_def cpu_reg_mod_def reset_trap_val_def) apply (simp add: write_cpu_def simpler_modify_def DetMonad.bind_def h1_def h2_def) apply (simp add: return_def) by (simp add: set_reset_trap_def simpler_modify_def DetMonad.bind_def h1_def h2_def return_def) next case False then show ?thesis apply (simp add: exe_trap_wr_pc_def get_curr_win_def) apply (simp add: exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: simpler_gets_def) apply (simp add: cpu_reg_val_def update_S_def cpu_reg_mod_def reset_trap_val_def) apply (simp add: write_cpu_def simpler_modify_def DetMonad.bind_def h1_def h2_def) by (simp add: return_def) qed lemma execute_trap_result : "\(reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0) \ snd (execute_trap() s) = False" proof - assume "\(reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0)" then have fact1: "snd (select_trap() s) = False" using select_trap_result2 by blast then show ?thesis proof (cases "err_mode_val s = True") case True then show ?thesis using fact1 apply (simp add: execute_trap_def exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (simp add: in_gets return_def select_trap_monad simpler_gets_def) next case False then show ?thesis using fact1 select_trap_monad apply (simp add: execute_trap_def exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def) apply (auto simp add: select_trap_monad) apply (simp add: DetMonad.bind_def h1_def h2_def get_curr_win_def) apply (simp add: get_CWP_def cpu_reg_val_def) apply (simp add: simpler_gets_def return_def write_cpu_def) apply (simp add: simpler_modify_def DetMonad.bind_def h1_def h2_def) apply (simp add: exe_trap_st_pc_result) by (simp add: case_prod_unfold exe_trap_wr_pc_result) qed qed lemma execute_trap_result2 : "\(reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0) \ snd (execute_trap() s) = False" using execute_trap_result by blast lemma exe_instr_all : "good_context (s::(('a::len) sparc_state)) \ snd (execute_instruction() s) = False" proof - assume asm1: "good_context s" let ?s' = "delayed_pool_write s" from asm1 have f1 : "(get_trap_set s = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0) \ ((\e. fetch_instruction ?s' = Inl e) \ (\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (annul_val ?s' = True \ (annul_val ?s' = False \ (\v1' v2'. fetch_instruction ?s' = Inr v1' \ ((decode_instruction v1')::(Exception list + instruction)) = Inr v2' \ supported_instruction (fst v2') = True) \ ((fst v2) \ ctrl_type RETT \ ((fst v2) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR ?s') = 1 \ (get_ET (cpu_reg_val PSR ?s') \ 1 \ (((get_S (cpu_reg_val PSR ?s')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR ?s'))) + 1) mod NWINDOWS)) (cpu_reg_val WIM ?s')) = 0 \ ((AND) (get_addr (snd v2) ?s') (0b00000000000000000000000000000011::word32)) = 0))))))))" using good_context_all by blast from f1 have f2: "get_trap_set s \ {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto show ?thesis proof (cases "get_trap_set s = {}") case True then have f3: "get_trap_set s = {}" by auto then show ?thesis proof (cases "exe_mode_val s = True") case True then have f4: "exe_mode_val s = True" by auto then show ?thesis proof (cases "\e1. fetch_instruction ?s' = Inl e1") case True then show ?thesis using f3 apply exe_proof_to_decode apply (simp add: raise_trap_def simpler_modify_def) by (simp add: bind_def h1_def h2_def return_def) next case False then have f5: "\ v1. fetch_instruction ?s' = Inr v1" using fetch_instr_result_1 by blast then have f6: "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2" using f1 fetch_instr_result_2 by blast then show ?thesis proof (cases "annul_val ?s' = True") case True then show ?thesis using f3 f4 f6 apply exe_proof_to_decode apply (simp add: set_annul_def annul_mod_def simpler_modify_def bind_def h1_def h2_def) apply (simp add: return_def simpler_gets_def) by (simp add: write_cpu_def simpler_modify_def) next case False then have f7: "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (\v1' v2'. fetch_instruction ?s' = Inr v1' \ ((decode_instruction v1')::(Exception list + instruction)) = Inr v2' \ supported_instruction (fst v2') = True) \ annul_val ?s' = False" using f1 f6 fetch_instr_result_2 by auto then have f7': "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ supported_instruction (fst v2) = True \ annul_val ?s' = False" by auto then show ?thesis proof (cases "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) = ctrl_type RETT") case True then have f8: "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) = ctrl_type RETT" by auto then show ?thesis proof (cases "get_trap_set ?s' = {}") case True then have f9: "get_trap_set ?s' = {}" by auto then show ?thesis proof (cases "get_ET (cpu_reg_val PSR ?s') = 1") case True then have f10: "get_ET (cpu_reg_val PSR ?s') = 1" by auto then show ?thesis proof (cases "(((get_S (cpu_reg_val PSR ?s')))::word1) = 0") case True then show ?thesis using f3 f4 f7 f8 f9 f10 apply exe_proof_to_decode apply exe_proof_dispatch_rett apply (simp add: raise_trap_def simpler_modify_def) apply (auto simp add: execute_instr_sub1_result return_def) by (simp add: case_prod_unfold) next case False then show ?thesis using f3 f4 f7 f8 f9 f10 apply exe_proof_to_decode apply exe_proof_dispatch_rett apply (simp add: raise_trap_def simpler_modify_def) apply (auto simp add: execute_instr_sub1_result return_def) by (simp add: case_prod_unfold) qed next case False then have f11: "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val ?s' = False \ (fst v2) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR ?s') \ 1 \ (((get_S (cpu_reg_val PSR ?s')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR ?s'))) + 1) mod NWINDOWS)) (cpu_reg_val WIM ?s')) = 0 \ ((AND) (get_addr (snd v2) ?s') (0b00000000000000000000000000000011::word32)) = 0)" using f1 fetch_instr_result_2 f7' f8 by auto then show ?thesis using f3 f4 proof (cases "get_trap_set ?s' = {}") case True then show ?thesis using f3 f4 f11 apply (simp add: execute_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def simpler_modify_def) apply clarsimp apply (simp add: return_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto apply (simp add: execute_instr_sub1_result) apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: rett_instr_result) next case False then show ?thesis using f3 f4 f11 apply (simp add: execute_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def simpler_modify_def) apply clarsimp apply (simp add: return_def) apply (simp add: bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: execute_instr_sub1_result) apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: return_def) qed qed next case False then show ?thesis using f3 f4 f7 f8 apply exe_proof_to_decode apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) by (auto simp add: execute_instr_sub1_result return_def Let_def) qed next case False \ \Instruction is not \RETT\.\ then have "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) \ ctrl_type RETT" using f7 by auto then have "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) \ ctrl_type RETT \ supported_instruction (fst v2) = True \ annul_val ?s' = False" using f7 by auto then have "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) \ ctrl_type RETT \ supported_instruction (fst v2) = True \ annul_val ?s' = False \ snd (dispatch_instruction v2 ?s') = False" by (auto simp add: dispatch_instr_result) then show ?thesis using f3 f4 apply exe_proof_to_decode apply (simp add: bind_def h1_def h2_def) apply (simp add: case_prod_unfold) by (simp add: execute_instr_sub1_result) qed qed qed next case False then show ?thesis using f3 apply (simp add: execute_instruction_def) by (simp add: exec_gets return_def) qed next case False then have "get_trap_set s \ {} \ ((reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0)" using f2 by auto then show ?thesis apply (simp add: execute_instruction_def exec_gets) by (simp add: execute_trap_result2) qed qed lemma dispatch_fail: "snd (execute_instruction() (s::(('a::len) sparc_state))) = False \ get_trap_set s = {} \ exe_mode_val s \ fetch_instruction (delayed_pool_write s) = Inr v \ ((decode_instruction v)::(Exception list + instruction)) = Inl e \ False" using decode_instr_result_2 apply (simp add: execute_instruction_def) apply (simp add: exec_gets bind_def) apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def return_def) by (simp add: fail_def) lemma no_error : "good_context s \ snd (execute_instruction () s) = False" proof - assume "good_context s" hence "snd (execute_instruction() s) = False" using exe_instr_all by auto hence "snd (execute_instruction () s) = False" by (simp add: exec_ss2) thus ?thesis by assumption qed theorem single_step : "good_context s \ NEXT s = Some (snd (fst (execute_instruction () s)))" by (simp add: no_error next_match) (*********************************************************************) section \Privilege safty\ (*********************************************************************) text \The following shows that, if the pre-state is under user mode, then after a singel step execution, the post-state is aslo under user mode.\ lemma write_cpu_pc_privilege: "s' = snd (fst (write_cpu w PC s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def) by (simp add: cpu_reg_val_def) lemma write_cpu_npc_privilege: "s' = snd (fst (write_cpu w nPC s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def) by (simp add: cpu_reg_val_def) lemma write_cpu_y_privilege: "s' = snd (fst (write_cpu w Y s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def) by (simp add: cpu_reg_val_def) lemma cpu_reg_mod_y_privilege: "s' = cpu_reg_mod w Y s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" by (simp add: cpu_reg_mod_def cpu_reg_val_def) lemma cpu_reg_mod_asr_privilege: "s' = cpu_reg_mod w (ASR r) s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" by (simp add: cpu_reg_mod_def cpu_reg_val_def) lemma global_reg_mod_privilege: "s' = global_reg_mod w1 n w2 s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (induction n arbitrary:s) apply (clarsimp) apply (auto) apply (simp add: Let_def) by (simp add: cpu_reg_val_def) lemma out_reg_mod_privilege: "s' = out_reg_mod a w r s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: out_reg_mod_def Let_def) by (simp add: cpu_reg_val_def) lemma in_reg_mod_privilege: "s' = in_reg_mod a w r s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: in_reg_mod_def Let_def) by (simp add: cpu_reg_val_def) lemma user_reg_mod_privilege: assumes a1: " s' = user_reg_mod d (w::(('a::len) window_size)) r (s::(('a::len) sparc_state)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "r = 0") case True then show ?thesis using a1 by (simp add: user_reg_mod_def) next case False then have f1: "r \ 0" by auto then show ?thesis proof (cases "0 < r \ r < 8") case True then show ?thesis using a1 f1 apply (simp add: user_reg_mod_def) by (auto intro: global_reg_mod_privilege) next case False then have f2: "\(0 < r \ r < 8)" by auto then show ?thesis proof (cases "7 < r \ r < 16") case True then show ?thesis using a1 f1 f2 apply (simp add: user_reg_mod_def) by (auto intro: out_reg_mod_privilege) next case False then have f3: "\ (7 < r \ r < 16)" by auto then show ?thesis proof (cases "15 < r \ r < 24") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: user_reg_mod_def) by (simp add: cpu_reg_val_def) next case False then show ?thesis using a1 f1 f2 f3 apply (simp add: user_reg_mod_def) by (auto intro: in_reg_mod_privilege) qed qed qed qed lemma write_reg_privilege: "s' = snd (fst (write_reg w1 w2 w3 (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_reg_def simpler_modify_def) by (auto intro: user_reg_mod_privilege) lemma set_annul_privilege: "s' = snd (fst (set_annul b s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: set_annul_def simpler_modify_def) apply (simp add: annul_mod_def write_annul_def) by (simp add: cpu_reg_val_def) lemma set_reset_trap_privilege: "s' = snd (fst (set_reset_trap b s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: set_reset_trap_def simpler_modify_def) apply (simp add: reset_trap_mod_def write_annul_def) by (simp add: cpu_reg_val_def) lemma empty_delayed_pool_write_privilege: "get_delayed_pool s = [] \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = delayed_pool_write s \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: delayed_pool_write_def) by (simp add: get_delayed_write_def delayed_write_all_def delayed_pool_rm_list_def) lemma raise_trap_privilege: "(((get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = snd (fst (raise_trap t s)) \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: raise_trap_def) apply (simp add: simpler_modify_def add_trap_set_def) by (simp add: cpu_reg_val_def) lemma write_cpu_tt_privilege: "s' = snd (fst (write_cpu_tt w s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_cpu_tt_def) apply (simp add: exec_gets) apply (simp add: write_cpu_def cpu_reg_mod_def write_tt_def) apply (simp add: simpler_modify_def) by (simp add: cpu_reg_val_def) lemma emp_trap_set_privilege: "s' = emp_trap_set s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: emp_trap_set_def) by (simp add: cpu_reg_val_def) lemma sys_reg_mod_privilege: "s' = sys_reg_mod w r s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: sys_reg_mod_def) by (simp add: cpu_reg_val_def) lemma mem_mod_privilege: assumes a1: "s' = mem_mod a1 a2 v s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "(uint a1) = 8 \ (uint a1) = 10") case True then show ?thesis using a1 apply (simp add: mem_mod_def) apply (simp add: Let_def) by (simp add: cpu_reg_val_def) next case False then have f1: "\((uint a1) = 8 \ (uint a1) = 10)" by auto then show ?thesis proof (cases "(uint a1) = 9 \ (uint a1) = 11") case True then show ?thesis using a1 f1 apply (simp add: mem_mod_def) apply (simp add: Let_def) by (simp add: cpu_reg_val_def) next case False then show ?thesis using a1 f1 apply (simp add: mem_mod_def) by (simp add: cpu_reg_val_def) qed qed lemma mem_mod_w32_privilege: "s' = mem_mod_w32 a1 a2 b d s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: mem_mod_w32_def) apply (simp add: Let_def) by (auto intro: mem_mod_privilege) lemma add_instr_cache_privilege: "s' = add_instr_cache s addr y m \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: add_instr_cache_def) apply (simp add: Let_def) by (simp add: icache_mod_def cpu_reg_val_def) lemma add_data_cache_privilege: "s' = add_data_cache s addr y m \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: add_data_cache_def) apply (simp add: Let_def) by (simp add: dcache_mod_def cpu_reg_val_def) lemma memory_read_privilege: assumes a1: "s' = snd (memory_read asi addr s) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "uint asi = 1") case True then show ?thesis using a1 apply (simp add: memory_read_def) by (simp add: Let_def) next case False then have f1: "uint asi \ 1" by auto then show ?thesis proof (cases "uint asi = 2") case True then show ?thesis using a1 f1 by (simp add: memory_read_def) next case False then have f2: "uint asi \ 2" by auto then show ?thesis proof (cases "uint asi \ {8,9}") case True then have f3: "uint asi \ {8,9}" by auto then show ?thesis proof (cases "load_word_mem s addr asi = None") case True then have f4: "load_word_mem s addr asi = None" by auto then show ?thesis using a1 f1 f2 f3 f4 by (simp add: memory_read_def) next case False then show ?thesis using a1 f1 f2 f3 apply (simp add: memory_read_def) apply auto apply (simp add: add_instr_cache_privilege) by (simp add: add_instr_cache_privilege) qed next case False then have f5: "uint asi \ {8, 9}" by auto then show ?thesis proof (cases "uint asi \ {10,11}") case True then have f6: "uint asi \ {10,11}" by auto then show ?thesis proof (cases "load_word_mem s addr asi = None") case True then have f7: "load_word_mem s addr asi = None" by auto then show ?thesis using a1 f1 f2 f5 f6 f7 by (simp add: memory_read_def) next case False then show ?thesis using a1 f1 f2 f5 f6 apply (simp add: memory_read_def) apply auto apply (simp add: add_data_cache_privilege) by (simp add: add_data_cache_privilege) qed next case False then have f8: "uint asi \ {10,11}" by auto then show ?thesis proof (cases "uint asi = 13") case True then have f9: "uint asi = 13" by auto then show ?thesis proof (cases "read_instr_cache s addr = None") case True then show ?thesis using a1 f1 f2 f5 f8 f9 by (simp add: memory_read_def) next case False then show ?thesis using a1 f1 f2 f5 f8 f9 apply (simp add: memory_read_def) by auto qed next case False then have f10: "uint asi \ 13" by auto then show ?thesis proof (cases "uint asi = 15") case True then show ?thesis using a1 f1 f2 f5 f8 f10 apply (simp add: memory_read_def) apply (cases "read_data_cache s addr = None") by auto next case False then show ?thesis using a1 f1 f2 f5 f8 f10 apply (simp add: memory_read_def) \ \The rest cases are easy.\ by (simp add: Let_def) qed qed qed qed qed qed lemma get_curr_win_privilege: "s' = snd (fst (get_curr_win() s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: get_curr_win_def) by (simp add: simpler_gets_def) lemma load_sub2_privilege: assumes a1: "s' = snd (fst (load_sub2 addr asi r win w s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "fst (memory_read asi (addr + 4) (snd (fst (write_reg w win (r AND 30) s)))) = None") case True then show ?thesis using a1 apply (simp add: load_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (auto intro: raise_trap_privilege write_reg_privilege) next case False then show ?thesis using a1 apply (simp add: load_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) by (auto intro: write_reg_privilege memory_read_privilege) qed lemma load_sub3_privilege: assumes a1: "s' = snd (fst (load_sub3 instr curr_win rd asi address s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "fst (memory_read asi address s) = None") case True then show ?thesis using a1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) by (auto intro: raise_trap_privilege) next case False then have f1: "fst (memory_read asi address s) \ None " by auto then show ?thesis proof (cases "rd \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA \ fst instr = load_store_type LDUH \ fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUBA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSBA)") case True then show ?thesis using a1 f1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) by (auto intro: write_reg_privilege memory_read_privilege) next case False then show ?thesis using a1 f1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply auto apply (simp add: simpler_modify_def bind_def h1_def h2_def) apply (auto intro: load_sub2_privilege memory_read_privilege) apply (simp add: simpler_modify_def bind_def h1_def h2_def) by (auto intro: load_sub2_privilege memory_read_privilege) qed qed lemma load_sub1_privilege: assumes a1: "s' = snd (fst (load_sub1 instr rd s_val s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: load_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto by (auto intro: get_curr_win_privilege raise_trap_privilege load_sub3_privilege) lemma load_instr_privilege: "s' = snd (fst (load_instr i s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: load_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) apply clarsimp by (auto intro: get_curr_win_privilege raise_trap_privilege load_sub1_privilege) lemma store_barrier_pending_mod_privilege: "s' = store_barrier_pending_mod b s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: store_barrier_pending_mod_def) apply (simp add: write_store_barrier_pending_def) by (simp add: cpu_reg_val_def) lemma store_word_mem_privilege: assumes a1: "store_word_mem s addr data byte_mask asi = Some s' \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: store_word_mem_def) apply (case_tac "virt_to_phys addr (mmu s) (mem s) = None") apply auto apply (case_tac "mmu_writable (get_acc_flag b) asi") apply auto by (simp add: mem_mod_w32_privilege) lemma flush_instr_cache_privilege: "(((get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = flush_instr_cache s \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: flush_instr_cache_def) by (simp add: cpu_reg_val_def) lemma flush_data_cache_privilege: "(((get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = flush_data_cache s \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: flush_data_cache_def) by (simp add: cpu_reg_val_def) lemma flush_cache_all_privilege: "(((get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = flush_cache_all s \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: flush_cache_all_def) by (simp add: cpu_reg_val_def) lemma memory_write_asi_privilege: assumes a1: "r = memory_write_asi asi addr byte_mask data s \ r = Some s' \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "uint asi = 1") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto intro: store_word_mem_privilege) next case False then have f1: "uint asi \ 1" by auto then show ?thesis proof (cases "uint asi = 2") case True then have f01: "uint asi = 2" by auto then show ?thesis proof (cases "uint addr = 0") case True then show ?thesis using a1 f1 f01 apply (simp add: memory_write_asi_def) apply (simp add: ccr_flush_def) apply (simp add: Let_def) apply auto apply (metis flush_data_cache_privilege flush_instr_cache_privilege sys_reg_mod_privilege) apply (metis flush_instr_cache_privilege sys_reg_mod_privilege) apply (metis flush_data_cache_privilege sys_reg_mod_privilege) by (simp add: sys_reg_mod_privilege) next case False then show ?thesis using a1 f1 f01 apply (simp add: memory_write_asi_def) apply clarsimp by (metis option.distinct(1) option.sel sys_reg_mod_privilege) qed next case False then have f2: "uint asi \ 2" by auto then show ?thesis proof (cases "uint asi \ {8,9}") case True then show ?thesis using a1 f1 f2 apply (simp add: memory_write_asi_def) using store_word_mem_privilege add_instr_cache_privilege by blast next case False then have f3: "uint asi \ {8,9}" by auto then show ?thesis proof (cases "uint asi \ {10,11}") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: memory_write_asi_def) using store_word_mem_privilege add_data_cache_privilege by blast next case False then have f4: "uint asi \ {10,11}" by auto then show ?thesis proof (cases "uint asi = 13") case True then show ?thesis using a1 f1 f2 f3 f4 apply (simp add: memory_write_asi_def) by (auto simp add: add_instr_cache_privilege) next case False then have f5: "uint asi \ 13" by auto then show ?thesis proof (cases "uint asi = 15") case True then show ?thesis using a1 f1 f2 f3 f4 f5 apply (simp add: memory_write_asi_def) by (auto simp add: add_data_cache_privilege) next case False then have f6: "uint asi \ 15" by auto then show ?thesis proof (cases "uint asi = 16") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto simp add: flush_instr_cache_privilege) next case False then have f7: "uint asi \ 16" by auto then show ?thesis proof (cases "uint asi = 17") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto simp add: flush_data_cache_privilege) next case False then have f8: "uint asi \ 17" by auto then show ?thesis proof (cases "uint asi = 24") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto simp add: flush_cache_all_privilege) next case False then have f9: "uint asi \ 24" by auto then show ?thesis proof (cases "uint asi = 25") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) apply (case_tac "mmu_reg_mod (mmu s) addr data = None") apply auto by (simp add: cpu_reg_val_def) next case False then have f10: "uint asi \ 25" by auto then show ?thesis proof (cases "uint asi = 28") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto simp add: mem_mod_w32_privilege) next case False \ \The remaining cases are easy.\ then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 apply (simp add: memory_write_asi_def) apply (auto simp add: Let_def) apply (case_tac "uint asi = 20 \ uint asi = 21") by auto qed qed qed qed qed qed qed qed qed qed qed lemma memory_write_privilege: assumes a1: "r = memory_write asi addr byte_mask data (s::(('a::len) sparc_state)) \ r = Some s' \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR (s'::(('a::len) sparc_state)))))::word1) = 0" proof - have "\x. Some x \ None" by auto then have "r \ None" using a1 by (simp add: \r = memory_write asi addr byte_mask data s \ r = Some s' \ (get_S (cpu_reg_val PSR s)) = 0\) then have "\s''. r = Some (store_barrier_pending_mod False s'')" using a1 by (metis (no_types, lifting) memory_write_def option.case_eq_if) then have "\s''. s' = store_barrier_pending_mod False s''" using a1 by blast then have "\s''. memory_write_asi asi addr byte_mask data s = Some s'' \ s' = store_barrier_pending_mod False s''" by (metis (no_types, lifting) assms memory_write_def not_None_eq option.case_eq_if option.sel) then show ?thesis using a1 using memory_write_asi_privilege store_barrier_pending_mod_privilege by blast qed lemma store_sub2_privilege: assumes a1: "s' = snd (fst (store_sub2 instr curr_win rd asi address s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "memory_write asi address (st_byte_mask instr address) (st_data0 instr curr_win rd address s) s = None") case True then show ?thesis using a1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (metis fst_conv raise_trap_privilege return_def snd_conv) next case False then have f1: "\(memory_write asi address (st_byte_mask instr address) (st_data0 instr curr_win rd address s) s = None)" by auto then show ?thesis proof (cases "(fst instr) \ {load_store_type STD,load_store_type STDA}") case True then have f2: "(fst instr) \ {load_store_type STD,load_store_type STDA}" by auto then show ?thesis using a1 f1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply (simp add: bind_def case_prod_unfold) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: case_prod_unfold bind_def h1_def h2_def Let_def simpler_modify_def) apply (simp add: simpler_gets_def) apply auto using memory_write_privilege raise_trap_privilege apply blast apply (simp add: simpler_modify_def simpler_gets_def bind_def) apply (meson memory_write_privilege) using memory_write_privilege raise_trap_privilege apply blast by (meson memory_write_privilege) next case False then show ?thesis using a1 f1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply clarsimp apply (simp add: simpler_modify_def return_def) by (auto intro: memory_write_privilege) qed qed lemma store_sub1_privilege: assumes a1: "s' = snd (fst (store_sub1 instr rd s_val (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR (s'::(('a::len) sparc_state)))))::word1) = 0" proof (cases "(fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word1) \ 0") case True then show ?thesis using a1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using get_curr_win_privilege raise_trap_privilege by blast next case False then have f1: "\((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word1) \ 0)" by auto then show ?thesis proof (cases "(fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word2) \ 0") case True then show ?thesis using a1 f1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using get_curr_win_privilege raise_trap_privilege by blast next case False then have f2: "\((fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word2) \ 0)" by auto then show ?thesis proof (cases "(fst instr \ {load_store_type STD,load_store_type STDA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word3) \ 0") case True then show ?thesis using a1 f1 f2 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using get_curr_win_privilege raise_trap_privilege by blast next case False then show ?thesis using a1 f1 f2 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege store_sub2_privilege) qed qed qed lemma store_instr_privilege: assumes a1: "s' = snd (fst (store_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR (s'::(('a::len) sparc_state)))))::word1) = 0" using a1 apply (simp add: store_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) using raise_trap_privilege store_sub1_privilege by blast lemma sethi_instr_privilege: assumes a1: "s' = snd (fst (sethi_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: sethi_instr_def) apply (simp add: Let_def) apply auto apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using get_curr_win_privilege write_reg_privilege apply blast by (simp add: return_def) lemma nop_instr_privilege: assumes a1: "s' = snd (fst (nop_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: nop_instr_def) by (simp add: return_def) lemma ucast_0: "(((get_S w))::word1) = 0 \ get_S w = 0" by simp lemma ucast_02: "get_S w = 0 \ (((get_S w))::word1) = 0" by simp lemma ucast_s: "(((get_S w))::word1) = 0 \ (AND) w (0b00000000000000000000000010000000::word32) = 0" by (simp add: get_S_def split: if_splits) lemma ucast_s2: "(AND) w 0b00000000000000000000000010000000 = 0 \ (((get_S w))::word1) = 0" by (simp add: get_S_def) lemma update_PSR_icc_1: "w' = (AND) w (0b11111111000011111111111111111111::word32) \ (((get_S w))::word1) = 0 \ (((get_S w'))::word1) = 0" by (simp add: get_S_def word_bw_assocs(1)) lemma and_num_1048576_128: "(AND) (0b00000000000100000000000000000000::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma and_num_2097152_128: "(AND) (0b00000000001000000000000000000000::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma and_num_4194304_128: "(AND) (0b00000000010000000000000000000000::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma and_num_8388608_128: "(AND) (0b00000000100000000000000000000000::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma or_and_s: "(AND) w1 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0 \ (AND) ((OR) w1 w2) (0b00000000000000000000000010000000::word32) = 0" by (simp add: word_ao_dist) lemma and_or_s: assumes "(((get_S w1))::word1) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0" shows "(((get_S ((OR) ((AND) w1 (0b11111111000011111111111111111111::word32)) w2)))::word1) = 0" proof - from assms have "w1 AND 128 = 0" using ucast_s by blast then have "(w1 AND 4279238655 OR w2) AND 128 = 0" using assms by (metis word_ao_absorbs(6) word_ao_dist word_bw_comms(2)) then show ?thesis using ucast_s2 by blast qed lemma and_or_or_s: assumes a1: "(((get_S w1))::word1) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w3 (0b00000000000000000000000010000000::word32) = 0" shows "(((get_S ((OR) ((OR) ((AND) w1 (0b11111111000011111111111111111111::word32)) w2) w3)))::word1) = 0" using and_or_s assms or_and_s ucast_s ucast_s2 by blast lemma and_or_or_or_s: assumes a1: "(((get_S w1))::word1) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w3 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w4 (0b00000000000000000000000010000000::word32) = 0" shows "(((get_S ((OR) ((OR) ((OR) ((AND) w1 (0b11111111000011111111111111111111::word32)) w2) w3) w4)))::word1) = 0" using and_or_or_s assms or_and_s ucast_s ucast_s2 by (simp add: and_or_or_s assms or_and_s ucast_s ucast_s2) lemma and_or_or_or_or_s: assumes a1: "(((get_S w1))::word1) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w3 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w4 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w5 (0b00000000000000000000000010000000::word32) = 0" shows "(((get_S ((OR) ((OR) ((OR) ((OR) ((AND) w1 (0b11111111000011111111111111111111::word32)) w2) w3) w4) w5)))::word1) = 0" using and_or_or_or_s assms or_and_s ucast_s ucast_s2 by (simp add: and_or_or_or_s assms or_and_s ucast_s ucast_s2) lemma write_cpu_PSR_icc_privilege: assumes a1: "s' = snd (fst (write_cpu (update_PSR_icc n_val z_val v_val c_val (cpu_reg_val PSR s)) PSR (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: write_cpu_def) apply (simp add: simpler_modify_def) apply (simp add: cpu_reg_mod_def update_PSR_icc_def) apply (simp add: cpu_reg_val_def) apply auto using update_PSR_icc_1 apply blast using update_PSR_icc_1 and_num_1048576_128 and_or_s apply blast using update_PSR_icc_1 and_num_2097152_128 and_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_2097152_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_4194304_128 and_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_4194304_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_2097152_128 and_num_4194304_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_2097152_128 and_num_4194304_128 and_or_or_or_s apply blast using update_PSR_icc_1 and_num_8388608_128 and_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_8388608_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_2097152_128 and_num_8388608_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_2097152_128 and_num_8388608_128 and_or_or_or_s apply blast using update_PSR_icc_1 and_num_4194304_128 and_num_8388608_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_4194304_128 and_num_8388608_128 and_or_or_or_s apply blast using update_PSR_icc_1 and_num_2097152_128 and_num_4194304_128 and_num_8388608_128 and_or_or_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_2097152_128 and_num_4194304_128 and_num_8388608_128 and_or_or_or_or_s by blast lemma and_num_4294967167_128: "(AND) (0b11111111111111111111111101111111::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma s_0_word: "(((get_S ((AND) w (0b11111111111111111111111101111111::word32))))::word1) = 0" apply (simp add: get_S_def) using and_num_4294967167_128 by (simp add: ac_simps) lemma update_PSR_CWP_1: "w' = (AND) w (0b11111111111111111111111111100000::word32) \ (((get_S w))::word1) = 0 \ (((get_S w'))::word1) = 0" by (simp add: get_S_def word_bw_assocs(1)) lemma write_cpu_PSR_CWP_privilege: assumes a1: "s' = snd (fst (write_cpu (update_CWP cwp_val (cpu_reg_val PSR s)) PSR (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: write_cpu_def) apply (simp add: simpler_modify_def) apply (simp add: cpu_reg_mod_def) apply (simp add: update_CWP_def) apply (simp add: Let_def) apply auto apply (simp add: cpu_reg_val_def) using s_0_word by blast lemma logical_instr_sub1_privilege: assumes a1: "s' = snd (fst (logical_instr_sub1 instr_name result (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "instr_name = logic_type ANDcc \ instr_name = logic_type ANDNcc \ instr_name = logic_type ORcc \ instr_name = logic_type ORNcc \ instr_name = logic_type XORcc \ instr_name = logic_type XNORcc") case True then show ?thesis using a1 apply (simp add: logical_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: logical_new_psr_val_def) using write_cpu_PSR_icc_privilege by blast next case False then show ?thesis using a1 apply (simp add: logical_instr_sub1_def) by (simp add: return_def) qed lemma logical_instr_privilege: assumes a1: "s' = snd (fst (logical_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: logical_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply auto apply (meson get_curr_win_privilege logical_instr_sub1_privilege write_reg_privilege) by (meson get_curr_win_privilege logical_instr_sub1_privilege write_reg_privilege) method shift_instr_privilege_proof = ( (simp add: shift_instr_def), (simp add: Let_def), (simp add: simpler_gets_def), (simp add: bind_def h1_def h2_def Let_def case_prod_unfold), auto, (blast intro: get_curr_win_privilege write_reg_privilege), (blast intro: get_curr_win_privilege write_reg_privilege) ) lemma shift_instr_privilege: assumes a1: "s' = snd (fst (shift_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "(fst instr = shift_type SLL) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 by shift_instr_privilege_proof next case False then have f1: "\((fst instr = shift_type SLL) \ (get_operand_w5 ((snd instr)!3) \ 0))" by auto then show ?thesis proof (cases "(fst instr = shift_type SRL) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 f1 by shift_instr_privilege_proof next case False then have f2: "\((fst instr = shift_type SRL) \ (get_operand_w5 ((snd instr)!3) \ 0))" by auto then show ?thesis proof (cases "(fst instr = shift_type SRA) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 f1 f2 by shift_instr_privilege_proof next case False then show ?thesis using a1 f1 f2 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply (simp add: return_def) using get_curr_win_privilege by blast qed qed qed lemma add_instr_sub1_privilege: assumes a1: "s' = snd (fst (add_instr_sub1 instr_name result rs1_val operand2 (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "instr_name = arith_type ADDcc \ instr_name = arith_type ADDXcc") case True then show ?thesis using a1 apply (simp add: add_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: write_cpu_PSR_icc_privilege) next case False then show ?thesis using a1 apply (simp add: add_instr_sub1_def) by (simp add: return_def) qed lemma add_instr_privilege: assumes a1: "s' = snd (fst (add_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: add_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson add_instr_sub1_privilege get_curr_win_privilege write_reg_privilege) lemma sub_instr_sub1_privilege: assumes a1: "s' = snd (fst (sub_instr_sub1 instr_name result rs1_val operand2 (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "instr_name = arith_type SUBcc \ instr_name = arith_type SUBXcc") case True then show ?thesis using a1 apply (simp add: sub_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: write_cpu_PSR_icc_privilege) next case False then show ?thesis using a1 apply (simp add: sub_instr_sub1_def) by (simp add: return_def) qed lemma sub_instr_privilege: assumes a1: "s' = snd (fst (sub_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: sub_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson sub_instr_sub1_privilege get_curr_win_privilege write_reg_privilege) lemma mul_instr_sub1_privilege: assumes a1: "s' = snd (fst (mul_instr_sub1 instr_name result (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "instr_name \ {arith_type SMULcc,arith_type UMULcc}") case True then show ?thesis using a1 apply (simp add: mul_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: write_cpu_PSR_icc_privilege) next case False then show ?thesis using a1 apply (simp add: mul_instr_sub1_def) by (simp add: return_def) qed lemma mul_instr_privilege: assumes a1: "s' = snd (fst (mul_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: mul_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege mul_instr_sub1_privilege write_cpu_y_privilege write_reg_privilege) lemma div_write_new_val_privilege: assumes a1: "s' = snd (fst (div_write_new_val i result temp_V (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "(fst i) \ {arith_type UDIVcc,arith_type SDIVcc}") case True then show ?thesis using a1 apply (simp add: div_write_new_val_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: write_cpu_PSR_icc_privilege) next case False then show ?thesis using a1 apply (simp add: div_write_new_val_def) by (simp add: return_def) qed lemma div_comp_privilege: assumes a1: "s' = snd (fst (div_comp instr rs1 rd operand2 (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: div_comp_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege div_write_new_val_privilege write_reg_privilege) lemma div_instr_privilege: assumes a1: "s' = snd (fst (div_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: div_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: return_def) apply auto using raise_trap_privilege apply blast using div_comp_privilege by blast lemma save_retore_sub1_privilege: assumes a1: "s' = snd (fst (save_retore_sub1 result new_cwp rd (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: save_retore_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using write_cpu_PSR_CWP_privilege write_reg_privilege by blast method save_restore_instr_privilege_proof = ( (simp add: save_restore_instr_def), (simp add: Let_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: case_prod_unfold), auto, (blast intro: get_curr_win_privilege raise_trap_privilege), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def case_prod_unfold), (blast intro: get_curr_win_privilege save_retore_sub1_privilege) ) lemma save_restore_instr_privilege: assumes a1: "s' = snd (fst (save_restore_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "fst instr = ctrl_type SAVE") case True then have f1: "fst instr = ctrl_type SAVE" by auto then show ?thesis using a1 by save_restore_instr_privilege_proof next case False then show ?thesis using a1 by save_restore_instr_privilege_proof qed lemma call_instr_privilege: assumes a1: "s' = snd (fst (call_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: call_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege write_cpu_npc_privilege write_cpu_pc_privilege write_reg_privilege) lemma jmpl_instr_privilege: assumes a1: "s' = snd (fst (jmpl_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: jmpl_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto using get_curr_win_privilege raise_trap_privilege apply blast apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege write_cpu_npc_privilege write_cpu_pc_privilege write_reg_privilege) lemma rett_instr_privilege: assumes a1: "snd (rett_instr i s) = False \ s' = snd (fst (rett_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: rett_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply auto apply (simp add: case_prod_unfold) apply (simp add: return_def) apply (blast intro: raise_trap_privilege) apply (simp add: bind_def h1_def h2_def Let_def) by (simp add: case_prod_unfold fail_def) method read_state_reg_instr_privilege_proof = ( (simp add: read_state_reg_instr_def), (simp add: Let_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: case_prod_unfold) ) lemma read_state_reg_instr_privilege: assumes a1: "s' = snd (fst (read_state_reg_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "(fst instr \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 ((snd instr)!0))))") case True then have "(fst instr \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 ((snd instr)!0)))) \ (((get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s))))))::word1) = 0" by (metis assms get_curr_win_privilege) then show ?thesis using a1 apply read_state_reg_instr_privilege_proof by (blast intro: raise_trap_privilege get_curr_win_privilege) next case False then have f1: "\((fst instr = sreg_type RDPSR \ fst instr = sreg_type RDWIM \ fst instr = sreg_type RDTBR \ fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 (snd instr ! 0))) \ (get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s))))) = 0)" by blast then show ?thesis proof (cases "illegal_instruction_ASR (get_operand_w5 ((snd instr)!0))") case True then show ?thesis using a1 f1 apply read_state_reg_instr_privilege_proof by (simp add: illegal_instruction_ASR_def) next case False then have f2: "\(illegal_instruction_ASR (get_operand_w5 ((snd instr)!0)))" by auto then show ?thesis proof (cases "(get_operand_w5 ((snd instr)!1)) \ 0") case True then have f3: "(get_operand_w5 ((snd instr)!1)) \ 0" by auto then show ?thesis proof (cases "fst instr = sreg_type RDY") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: read_state_reg_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (blast intro: get_curr_win_privilege write_reg_privilege) next case False then have f4: "\(fst instr = sreg_type RDY)" by auto then show ?thesis proof (cases "fst instr = sreg_type RDASR") case True then show ?thesis using a1 f1 f2 f3 f4 apply read_state_reg_instr_privilege_proof apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: get_curr_win_privilege write_reg_privilege) next case False then have f5: "\(fst instr = sreg_type RDASR)" by auto then show ?thesis proof (cases "fst instr = sreg_type RDPSR") case True then show ?thesis using a1 f1 f2 f3 f4 f5 apply read_state_reg_instr_privilege_proof by (blast intro: get_curr_win_privilege write_reg_privilege) next case False then show ?thesis using a1 f1 f2 f3 f4 f5 apply read_state_reg_instr_privilege_proof apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: get_curr_win_privilege write_reg_privilege) qed qed qed next case False then show ?thesis using a1 apply read_state_reg_instr_privilege_proof apply (simp add: return_def) using f1 f2 get_curr_win_privilege by blast qed qed qed method write_state_reg_instr_privilege_proof = ( (simp add: write_state_reg_instr_def), (simp add: Let_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: case_prod_unfold) ) lemma write_state_reg_instr_privilege: assumes a1: "s' = snd (fst (write_state_reg_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "fst instr = sreg_type WRY") case True then show ?thesis using a1 apply write_state_reg_instr_privilege_proof apply (simp add: simpler_modify_def) apply (simp add: delayed_pool_add_def DELAYNUM_def) by (blast intro: cpu_reg_mod_y_privilege get_curr_win_privilege) next case False then have f1: "\(fst instr = sreg_type WRY)" by auto then show ?thesis proof (cases "fst instr = sreg_type WRASR") case True then show ?thesis using a1 f1 apply write_state_reg_instr_privilege_proof apply (simp add: simpler_modify_def) apply auto using illegal_instruction_ASR_def apply blast using illegal_instruction_ASR_def apply blast using illegal_instruction_ASR_def apply blast using raise_trap_privilege get_curr_win_privilege apply blast apply (simp add: simpler_modify_def delayed_pool_add_def DELAYNUM_def) using cpu_reg_mod_asr_privilege get_curr_win_privilege apply blast apply (simp add: simpler_modify_def delayed_pool_add_def DELAYNUM_def) using cpu_reg_mod_asr_privilege get_curr_win_privilege by blast next case False then have f2: "\(fst instr = sreg_type WRASR)" by auto have f3: "get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s)))) = 0" using get_curr_win_privilege a1 by (metis ucast_id) then show ?thesis proof (cases "fst instr = sreg_type WRPSR") case True then show ?thesis using a1 f1 f2 f3 apply write_state_reg_instr_privilege_proof by (metis raise_trap_privilege ucast_0) next case False then have f4: "\(fst instr = sreg_type WRPSR)" by auto then show ?thesis proof (cases "fst instr = sreg_type WRWIM") case True then show ?thesis using a1 f1 f2 f3 f4 apply write_state_reg_instr_privilege_proof by (metis raise_trap_privilege ucast_0) next case False then have f5: "\(fst instr = sreg_type WRWIM)" by auto then show ?thesis using a1 f1 f2 f3 f4 f5 apply write_state_reg_instr_privilege_proof by (metis raise_trap_privilege ucast_0) qed qed qed qed lemma flush_instr_privilege: assumes a1: "s' = snd (fst (flush_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: flush_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def simpler_modify_def) by (auto simp add: flush_cache_all_privilege) lemma branch_instr_privilege: assumes a1: "s' = snd (fst (branch_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: branch_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold return_def) by (meson set_annul_privilege write_cpu_npc_privilege write_cpu_pc_privilege) method dispath_instr_privilege_proof = ( (simp add: dispatch_instruction_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: Let_def) ) lemma dispath_instr_privilege: assumes a1: "snd (dispatch_instruction instr s) = False \ s' = snd (fst (dispatch_instruction instr s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "get_trap_set s = {}") case True then have f1: "get_trap_set s = {}" by auto show ?thesis proof (cases "fst instr \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD}") case True then show ?thesis using a1 f1 apply dispath_instr_privilege_proof by (blast intro: load_instr_privilege) next case False then have f2: "\(fst instr \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD})" by auto then show ?thesis proof (cases "fst instr \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD}") case True then show ?thesis using a1 f1 f2 apply dispath_instr_privilege_proof by (blast intro: store_instr_privilege) next case False then have f3: "\(fst instr \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD})" by auto then show ?thesis proof (cases "fst instr \ {sethi_type SETHI}") case True then show ?thesis using a1 f1 f2 f3 apply dispath_instr_privilege_proof by (blast intro: sethi_instr_privilege) next case False then have f4: "\(fst instr \ {sethi_type SETHI})" by auto then show ?thesis proof (cases "fst instr \ {nop_type NOP}") case True then show ?thesis using a1 f1 f2 f3 f4 apply dispath_instr_privilege_proof by (blast intro: nop_instr_privilege) next case False then have f5: "\(fst instr \ {nop_type NOP})" by auto then show ?thesis proof (cases "fst instr \ {logic_type ANDs,logic_type ANDcc,logic_type ANDN, logic_type ANDNcc,logic_type ORs,logic_type ORcc,logic_type ORN, logic_type XORs,logic_type XNOR}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 apply dispath_instr_privilege_proof by (blast intro: logical_instr_privilege) next case False then have f6: "\(fst instr \ {logic_type ANDs,logic_type ANDcc,logic_type ANDN, logic_type ANDNcc,logic_type ORs,logic_type ORcc,logic_type ORN, logic_type XORs,logic_type XNOR})" by auto show ?thesis proof (cases "fst instr \ {shift_type SLL,shift_type SRL,shift_type SRA}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 apply dispath_instr_privilege_proof by (blast intro: shift_instr_privilege) next case False then have f7: "\(fst instr \ {shift_type SLL,shift_type SRL,shift_type SRA})" by auto then show ?thesis proof (cases "fst instr \ {arith_type ADD,arith_type ADDcc,arith_type ADDX}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 apply dispath_instr_privilege_proof by (blast intro: add_instr_privilege) next case False then have f8: "\(fst instr \ {arith_type ADD,arith_type ADDcc,arith_type ADDX})" by auto then show ?thesis proof (cases "fst instr \ {arith_type SUB,arith_type SUBcc,arith_type SUBX}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 apply dispath_instr_privilege_proof by (blast intro: sub_instr_privilege) next case False then have f9: "\(fst instr \ {arith_type SUB,arith_type SUBcc,arith_type SUBX})" by auto then show ?thesis proof (cases "fst instr \ {arith_type UMUL,arith_type SMUL,arith_type SMULcc}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 apply dispath_instr_privilege_proof by (blast intro: mul_instr_privilege) next case False then have f10: "\(fst instr \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc})" by auto then show ?thesis proof (cases "fst instr \ {arith_type UDIV,arith_type UDIVcc,arith_type SDIV}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 apply dispath_instr_privilege_proof by (blast intro: div_instr_privilege) next case False then have f11: "\(fst instr \ {arith_type UDIV, arith_type UDIVcc,arith_type SDIV})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type SAVE,ctrl_type RESTORE}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 apply dispath_instr_privilege_proof by (blast intro: save_restore_instr_privilege) next case False then have f12: "\(fst instr \ {ctrl_type SAVE,ctrl_type RESTORE})" by auto then show ?thesis proof (cases "fst instr \ {call_type CALL}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 apply dispath_instr_privilege_proof by (blast intro: call_instr_privilege) next case False then have f13: "\(fst instr \ {call_type CALL})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type JMPL}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 apply dispath_instr_privilege_proof by (blast intro: jmpl_instr_privilege) next case False then have f14: "\(fst instr \ {ctrl_type JMPL})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type RETT}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 apply dispath_instr_privilege_proof by (blast intro: rett_instr_privilege) next case False then have f15: "\(fst instr \ {ctrl_type RETT})" by auto then show ?thesis proof (cases "fst instr \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM, sreg_type RDTBR}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 apply dispath_instr_privilege_proof by (blast intro: read_state_reg_instr_privilege) next case False then have f16: "\(fst instr \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM, sreg_type RDTBR})" by auto then show ?thesis proof (cases "fst instr \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM, sreg_type WRTBR}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 apply dispath_instr_privilege_proof by (blast intro: write_state_reg_instr_privilege) next case False then have f17: "\(fst instr \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM, sreg_type WRTBR})" by auto then show ?thesis proof (cases "fst instr \ {load_store_type FLUSH}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 apply dispath_instr_privilege_proof by (blast intro: flush_instr_privilege) next case False then have f18: "\(fst instr \ {load_store_type FLUSH})" by auto then show ?thesis proof (cases "fst instr \ {bicc_type BE,bicc_type BNE, bicc_type BGU,bicc_type BLE,bicc_type BL,bicc_type BGE, bicc_type BNEG,bicc_type BG,bicc_type BCS,bicc_type BLEU, bicc_type BCC,bicc_type BA,bicc_type BN}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 apply dispath_instr_privilege_proof by (blast intro: branch_instr_privilege) next case False then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 apply dispath_instr_privilege_proof by (simp add: fail_def) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed next case False then show ?thesis using a1 apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) by (simp add: return_def) qed lemma execute_instr_sub1_privilege: assumes a1: "snd (execute_instr_sub1 i s) = False \ s' = snd (fst (execute_instr_sub1 i s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "get_trap_set s = {} \ fst i \ {call_type CALL,ctrl_type RETT,ctrl_type JMPL, bicc_type BE,bicc_type BNE,bicc_type BGU, bicc_type BLE,bicc_type BL,bicc_type BGE, bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC, bicc_type BA,bicc_type BN}") case True then show ?thesis using a1 apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold return_def) by (auto intro: write_cpu_pc_privilege write_cpu_npc_privilege) next case False then show ?thesis using a1 apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold return_def) by auto qed text \ Assume that there is no \delayed_write\ and there is no traps to be executed. If an instruction is executed as a user, the privilege will not be changed to supervisor after the execution.\ theorem safe_privilege : assumes a1: "get_delayed_pool s = [] \ get_trap_set s = {} \ snd (execute_instruction() s) = False \ s' = snd (fst (execute_instruction() s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "exe_mode_val s") case True then have f2: "exe_mode_val s = True" by auto then show ?thesis proof (cases "\e. fetch_instruction (delayed_pool_write s) = Inl e") case True then have f3: "\e. fetch_instruction (delayed_pool_write s) = Inl e" by auto then have f4: "\ (\v. fetch_instruction (delayed_pool_write s) = Inr v)" using fetch_instr_result_3 by auto then show ?thesis using a1 f2 f3 raise_trap_result empty_delayed_pool_write_privilege raise_trap_privilege apply (simp add: execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: case_prod_unfold) by (blast intro: empty_delayed_pool_write_privilege raise_trap_privilege) next case False then have f5: "\v. fetch_instruction (delayed_pool_write s) = Inr v" using fetch_instr_result_1 by blast then have f6: "\v. fetch_instruction (delayed_pool_write s) = Inr v \ \ (\e. ((decode_instruction v)::(Exception list + instruction)) = Inl e)" using a1 f2 dispatch_fail by blast then have f7: "\v. fetch_instruction (delayed_pool_write s) = Inr v \ (\v1. ((decode_instruction v)::(Exception list + instruction)) = Inr v1)" using decode_instr_result_4 by auto then show ?thesis proof (cases "annul_val (delayed_pool_write s)") case True then show ?thesis using a1 f2 f7 apply (simp add: execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (auto intro: empty_delayed_pool_write_privilege set_annul_privilege write_cpu_npc_privilege write_cpu_pc_privilege) next case False then show ?thesis using a1 f2 f7 apply (simp add: execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (auto intro: empty_delayed_pool_write_privilege dispath_instr_privilege execute_instr_sub1_privilege) qed qed next case False then show ?thesis using a1 apply (simp add: execute_instruction_def) by (simp add: simpler_gets_def bind_def h1_def h2_def Let_def return_def) qed (*********************************************************************) section \Single step non-interference property.\ (*********************************************************************) definition user_accessible:: "('a::len) sparc_state \ phys_address \ bool" where "user_accessible s pa \ \va p. (virt_to_phys va (mmu s) (mem s)) = Some p \ mmu_readable (get_acc_flag (snd p)) 10 \ (fst p) = pa" \ \Passing \asi = 8\ is the same.\ lemma user_accessible_8: assumes a1: "mmu_readable (get_acc_flag (snd p)) 8" shows "mmu_readable (get_acc_flag (snd p)) 10" using a1 by (simp add: mmu_readable_def) definition mem_equal:: "('a) sparc_state \ ('a) sparc_state \ phys_address \ bool" where "mem_equal s1 s2 pa \ (mem s1) 8 (pa AND 68719476732) = (mem s2) 8 (pa AND 68719476732) \ (mem s1) 8 ((pa AND 68719476732) + 1) = (mem s2) 8 ((pa AND 68719476732) + 1) \ (mem s1) 8 ((pa AND 68719476732) + 2) = (mem s2) 8 ((pa AND 68719476732) + 2) \ (mem s1) 8 ((pa AND 68719476732) + 3) = (mem s2) 8 ((pa AND 68719476732) + 3) \ (mem s1) 9 (pa AND 68719476732) = (mem s2) 9 (pa AND 68719476732) \ (mem s1) 9 ((pa AND 68719476732) + 1) = (mem s2) 9 ((pa AND 68719476732) + 1) \ (mem s1) 9 ((pa AND 68719476732) + 2) = (mem s2) 9 ((pa AND 68719476732) + 2) \ (mem s1) 9 ((pa AND 68719476732) + 3) = (mem s2) 9 ((pa AND 68719476732) + 3) \ (mem s1) 10 (pa AND 68719476732) = (mem s2) 10 (pa AND 68719476732) \ (mem s1) 10 ((pa AND 68719476732) + 1) = (mem s2) 10 ((pa AND 68719476732) + 1) \ (mem s1) 10 ((pa AND 68719476732) + 2) = (mem s2) 10 ((pa AND 68719476732) + 2) \ (mem s1) 10 ((pa AND 68719476732) + 3) = (mem s2) 10 ((pa AND 68719476732) + 3) \ (mem s1) 11 (pa AND 68719476732) = (mem s2) 11 (pa AND 68719476732) \ (mem s1) 11 ((pa AND 68719476732) + 1) = (mem s2) 11 ((pa AND 68719476732) + 1) \ (mem s1) 11 ((pa AND 68719476732) + 2) = (mem s2) 11 ((pa AND 68719476732) + 2) \ (mem s1) 11 ((pa AND 68719476732) + 3) = (mem s2) 11 ((pa AND 68719476732) + 3)" text \\low_equal\ defines the equivalence relation over two sparc states that is an analogy to the \=\<^sub>L\ relation over memory contexts in the traditional non-interference theorem.\ definition low_equal:: "('a::len) sparc_state \ ('a) sparc_state \ bool" where "low_equal s1 s2 \ (cpu_reg s1) = (cpu_reg s2) \ (user_reg s1) = (user_reg s2) \ (sys_reg s1) = (sys_reg s2) \ (\va. (virt_to_phys va (mmu s1) (mem s1)) = (virt_to_phys va (mmu s2) (mem s2))) \ (\pa. (user_accessible s1 pa) \ mem_equal s1 s2 pa) \ (mmu s1) = (mmu s2) \ (state_var s1) = (state_var s2) \ (traps s1) = (traps s2) \ (undef s1) = (undef s2) " lemma low_equal_com: "low_equal s1 s2 \ low_equal s2 s1" apply (simp add: low_equal_def) apply (simp add: mem_equal_def user_accessible_def) by metis lemma non_exe_mode_equal: "exe_mode_val s = False \ get_trap_set s = {} \ Some t = NEXT s \ t = s" apply (simp add: NEXT_def execute_instruction_def) apply auto by (simp add: simpler_gets_def bind_def h1_def h2_def Let_def return_def) lemma exe_mode_low_equal: assumes a1: "low_equal s1 s2" shows " exe_mode_val s1 = exe_mode_val s2" using a1 apply (simp add: low_equal_def) by (simp add: exe_mode_val_def) lemma mem_val_mod_state: "mem_val_alt asi a s = mem_val_alt asi a (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)" apply (simp add: mem_val_alt_def) by (simp add: Let_def) lemma mem_val_w32_mod_state: "mem_val_w32 asi a s = mem_val_w32 asi a (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) by (metis mem_val_mod_state) lemma load_word_mem_mod_state: "load_word_mem s addr asi = load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi" apply (simp add: load_word_mem_def) apply (case_tac "virt_to_phys addr (mmu s) (mem s) = None") apply auto by (auto simp add: mem_val_w32_mod_state) lemma load_word_mem2_mod_state: "fst (case load_word_mem s addr asi of None \ (None, s) | Some w \ (Some w, add_data_cache s addr w 15)) = fst (case load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi of None \ (None, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)) | Some w \ (Some w, add_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr w 15))" proof (cases "load_word_mem s addr asi = None") case True then have "load_word_mem s addr asi = None \ load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi = None" using load_word_mem_mod_state by metis then show ?thesis by auto next case False then have "\w. load_word_mem s addr asi = Some w" by auto then have "\w. load_word_mem s addr asi = Some w \ load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi = Some w" using load_word_mem_mod_state by metis then show ?thesis by auto qed lemma load_word_mem3_mod_state: "fst (case load_word_mem s addr asi of None \ (None, s) | Some w \ (Some w, add_instr_cache s addr w 15)) = fst (case load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi of None \ (None, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)) | Some w \ (Some w, add_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr w 15))" proof (cases "load_word_mem s addr asi = None") case True then have "load_word_mem s addr asi = None \ load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi = None" using load_word_mem_mod_state by metis then show ?thesis by auto next case False then have "\w. load_word_mem s addr asi = Some w" by auto then have "\w. load_word_mem s addr asi = Some w \ load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi = Some w" using load_word_mem_mod_state by metis then show ?thesis by auto qed lemma read_dcache_mod_state: "read_data_cache s addr = read_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr" apply (simp add: read_data_cache_def) by (simp add: dcache_val_def) lemma read_dcache2_mod_state: "fst (case read_data_cache s addr of None \ (None, s) | Some w \ (Some w, s)) = fst (case read_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr of None \ (None, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)) | Some w \ (Some w, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)))" proof (cases "read_data_cache s addr = None") case True then have "read_data_cache s addr = None \ read_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr = None" using read_dcache_mod_state by metis then show ?thesis by auto next case False then have "\w. read_data_cache s addr = Some w" by auto then have "\w. read_data_cache s addr = Some w \ read_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr = Some w" using read_dcache_mod_state by metis then show ?thesis by auto qed lemma read_icache_mod_state: "read_instr_cache s addr = read_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr" apply (simp add: read_instr_cache_def) by (simp add: icache_val_def) lemma read_icache2_mod_state: "fst (case read_instr_cache s addr of None \ (None, s) | Some w \ (Some w, s)) = fst (case read_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr of None \ (None, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)) | Some w \ (Some w, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)))" proof (cases "read_instr_cache s addr = None") case True then have "read_instr_cache s addr = None \ read_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr = None" using read_icache_mod_state by metis then show ?thesis by auto next case False then have "\w. read_instr_cache s addr = Some w" by auto then have "\w. read_instr_cache s addr = Some w \ read_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr = Some w" using read_icache_mod_state by metis then show ?thesis by auto qed lemma mem_read_mod_state: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\))" apply (simp add: memory_read_def) apply (case_tac "uint asi = 1") apply (simp add: Let_def) apply (metis load_word_mem_mod_state option.distinct(1)) apply (case_tac "uint asi = 2") apply (simp add: Let_def) apply (simp add: sys_reg_val_def) apply (case_tac "uint asi \ {8,9}") apply (simp add: Let_def) apply (simp add: load_word_mem3_mod_state) apply (simp add: load_word_mem_mod_state) apply (case_tac "uint asi \ {10,11}") apply (simp add: Let_def) apply (simp add: load_word_mem2_mod_state) apply (simp add: load_word_mem_mod_state) apply (case_tac "uint asi = 13") apply (simp add: Let_def) apply (simp add: read_icache2_mod_state) apply (case_tac "uint asi = 15") apply (simp add: Let_def) apply (simp add: read_dcache2_mod_state) apply (case_tac "uint asi = 25") apply (simp add: Let_def) apply (case_tac "uint asi = 28") apply (simp add: Let_def) apply (simp add: mem_val_w32_mod_state) by (simp add: Let_def) lemma insert_trap_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\traps := new_traps\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := (cpu_reg s), user_reg := (user_reg s), dwrite := (dwrite s), state_var := (state_var s), traps := new_traps, undef := (undef s)\))" using mem_read_mod_state by blast then show ?thesis by auto qed lemma cpu_reg_mod_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg, user_reg := (user_reg s), dwrite := (dwrite s), state_var := (state_var s), traps := (traps s), undef := (undef s)\))" using mem_read_mod_state by blast then show ?thesis by auto qed lemma user_reg_mod_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\user_reg := new_user_reg\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := (cpu_reg s), user_reg := new_user_reg, dwrite := (dwrite s), state_var := (state_var s), traps := (traps s), undef := (undef s)\))" using mem_read_mod_state by blast then show ?thesis by auto qed lemma annul_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\state_var := new_state_var, cpu_reg := new_cpu_reg\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg, user_reg := (user_reg s), dwrite := (dwrite s), state_var := new_state_var, traps := (traps s), undef := (undef s)\))" using mem_read_mod_state by blast then have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg, state_var := new_state_var\))" by auto then show ?thesis by (metis Sparc_State.sparc_state.surjective Sparc_State.sparc_state.update_convs(1) Sparc_State.sparc_state.update_convs(8)) qed lemma state_var_mod_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\state_var := new_state_var\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := (cpu_reg s), user_reg := (user_reg s), dwrite := (dwrite s), state_var := new_state_var, traps := (traps s), undef := (undef s)\))" using mem_read_mod_state by blast then show ?thesis by auto qed lemma mod_state_low_equal: "low_equal s1 s2 \ t1 = (s1\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) \ t2 = (s2\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) \ low_equal t1 t2" apply (simp add: low_equal_def) apply clarsimp apply (simp add: mem_equal_def) by (simp add: user_accessible_def) lemma user_reg_state_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (s1\user_reg := new_user_reg\) \ t2 = (s2\user_reg := new_user_reg\)" shows "low_equal t1 t2" proof - have "low_equal s1 s2 \ t1 = (s1\cpu_reg := (cpu_reg s1), user_reg := new_user_reg, dwrite := (dwrite s1), state_var := (state_var s1), traps := (traps s1), undef := (undef s1)\) \ t2 = (s2\cpu_reg := (cpu_reg s2), user_reg := new_user_reg, dwrite := (dwrite s2), state_var := (state_var s2), traps := (traps s2), undef := (undef s2)\) \ low_equal t1 t2" using mod_state_low_equal apply (simp add: low_equal_def) apply (simp add: user_accessible_def mem_equal_def) by clarsimp then show ?thesis using a1 by clarsimp qed lemma mod_trap_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (s1\traps := new_traps\) \ t2 = (s2\traps := new_traps\)" shows "low_equal t1 t2" proof - have "low_equal s1 s2 \ t1 = (s1\cpu_reg := (cpu_reg s1), user_reg := (user_reg s1), dwrite := (dwrite s1), state_var := (state_var s1), traps := new_traps, undef := (undef s1)\) \ t2 = (s2\cpu_reg := (cpu_reg s2), user_reg := (user_reg s2), dwrite := (dwrite s2), state_var := (state_var s2), traps := new_traps, undef := (undef s2)\) \ low_equal t1 t2" using mod_state_low_equal apply (simp add: low_equal_def) apply (simp add: user_accessible_def mem_equal_def) by clarsimp then show ?thesis using a1 by clarsimp qed lemma state_var_low_equal: "low_equal s1 s2 \ state_var s1 = state_var s2" by (simp add: low_equal_def) lemma state_var2_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (s1\state_var := new_state_var\) \ t2 = (s2\state_var := new_state_var\)" shows "low_equal t1 t2" proof - have "low_equal s1 s2 \ t1 = (s1\cpu_reg := (cpu_reg s1), user_reg := (user_reg s1), dwrite := (dwrite s1), state_var := new_state_var, traps := (traps s1), undef := (undef s1)\) \ t2 = (s2\cpu_reg := (cpu_reg s2), user_reg := (user_reg s2), dwrite := (dwrite s2), state_var := new_state_var, traps := (traps s2), undef := (undef s2)\) \ low_equal t1 t2" using mod_state_low_equal apply (simp add: low_equal_def) apply (simp add: user_accessible_def mem_equal_def) by clarsimp then show ?thesis using a1 by clarsimp qed lemma traps_low_equal: "low_equal s1 s2 \ traps s1 = traps s2" by (simp add: low_equal_def) lemma s_low_equal: "low_equal s1 s2 \ (get_S (cpu_reg_val PSR s1)) = (get_S (cpu_reg_val PSR s2))" by (simp add: low_equal_def cpu_reg_val_def) lemma cpu_reg_val_low_equal: "low_equal s1 s2 \ (cpu_reg_val cr s1) = (cpu_reg_val cr s2)" by (simp add: cpu_reg_val_def low_equal_def) lemma get_curr_win_low_equal: "low_equal s1 s2 \ (fst (fst (get_curr_win () s1))) = (fst (fst (get_curr_win () s2)))" apply (simp add: low_equal_def) apply (simp add: get_curr_win_def cpu_reg_val_def get_CWP_def) by (simp add: simpler_gets_def) lemma get_curr_win2_low_equal: "low_equal s1 s2 \ t1 = (snd (fst (get_curr_win () s1))) \ t2 = (snd (fst (get_curr_win () s2))) \ low_equal t1 t2" apply (simp add: low_equal_def) apply (simp add: get_curr_win_def cpu_reg_val_def get_CWP_def) by (auto simp add: simpler_gets_def) lemma get_curr_win3_low_equal: "low_equal s1 s2 \ (traps (snd (fst (get_curr_win () s1)))) = (traps (snd (fst (get_curr_win () s2))))" using low_equal_def get_curr_win2_low_equal by blast lemma get_addr_low_equal: "low_equal s1 s2 \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) = ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) = ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) = ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1)" apply (simp add: low_equal_def) apply (simp add: get_curr_win_def cpu_reg_val_def get_CWP_def) apply (simp add: simpler_gets_def get_addr_def user_reg_val_def) apply (simp add: Let_def ) apply (simp add: get_CWP_def cpu_reg_val_def get_operand2_def) by (simp add: user_reg_val_def) lemma get_addr2_low_equal: "low_equal s1 s2 \ get_addr (snd instr) (snd (fst (get_curr_win () s1))) = get_addr (snd instr) (snd (fst (get_curr_win () s2)))" apply (simp add: low_equal_def) apply (simp add: get_curr_win_def cpu_reg_val_def get_CWP_def) apply (simp add: simpler_gets_def get_addr_def user_reg_val_def) apply (simp add: Let_def ) apply (simp add: get_CWP_def cpu_reg_val_def get_operand2_def) by (simp add: user_reg_val_def) lemma sys_reg_low_equal: "low_equal s1 s2 \ sys_reg s1 = sys_reg s2" by (simp add: low_equal_def) lemma user_reg_low_equal: "low_equal s1 s2 \ user_reg s1 = user_reg s2" by (simp add: low_equal_def) lemma user_reg_val_low_equal: "low_equal s1 s2 \ user_reg_val win ur s1 = user_reg_val win ur s2" apply (simp add: user_reg_val_def) by (simp add: user_reg_low_equal) lemma get_operand2_low_equal: "low_equal s1 s2 \ get_operand2 op_list s1 = get_operand2 op_list s2" apply (simp add: get_operand2_def) apply (simp add: cpu_reg_val_low_equal) apply auto apply (simp add: user_reg_val_def) using user_reg_low_equal by fastforce lemma mem_val_mod_cache: "mem_val_alt asi a s = mem_val_alt asi a (s\cache := new_cache\)" apply (simp add: mem_val_alt_def) by (simp add: Let_def) lemma mem_val_w32_mod_cache: "mem_val_w32 asi a s = mem_val_w32 asi a (s\cache := new_cache\)" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) by (metis mem_val_mod_cache) lemma load_word_mem_mod_cache: "load_word_mem s addr asi = load_word_mem (s\cache := new_cache\) addr asi" apply (simp add: load_word_mem_def) apply (case_tac "virt_to_phys addr (mmu s) (mem s) = None") apply auto by (simp add: mem_val_w32_mod_cache) lemma memory_read_8_mod_cache: "fst (memory_read 8 addr s) = fst (memory_read 8 addr (s\cache := new_cache\))" apply (simp add: memory_read_def) apply (case_tac "sys_reg s CCR AND 1 \ 0") apply auto apply (simp add: option.case_eq_if load_word_mem_mod_cache) apply (auto intro: load_word_mem_mod_cache) apply (metis load_word_mem_mod_cache option.distinct(1)) by (metis load_word_mem_mod_cache option.distinct(1)) lemma memory_read_10_mod_cache: "fst (memory_read 10 addr s) = fst (memory_read 10 addr (s\cache := new_cache\))" apply (simp add: memory_read_def) apply (case_tac "sys_reg s CCR AND 1 \ 0") apply auto apply (simp add: option.case_eq_if load_word_mem_mod_cache) apply (auto intro: load_word_mem_mod_cache) apply (metis load_word_mem_mod_cache option.distinct(1)) by (metis load_word_mem_mod_cache option.distinct(1)) lemma mem_equal_mod_cache: "mem_equal s1 s2 pa \ mem_equal (s1\cache := new_cache1\) (s2\cache := new_cache2\) pa" by (simp add: mem_equal_def) lemma user_accessible_mod_cache: "user_accessible (s\cache := new_cache\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_equal_mod_user_reg: "mem_equal s1 s2 pa \ mem_equal (s1\user_reg := new_user_reg1\) (s2\user_reg := user_reg2\) pa" by (simp add: mem_equal_def) lemma user_accessible_mod_user_reg: "user_accessible (s\user_reg := new_user_reg\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_equal_mod_cpu_reg: "mem_equal s1 s2 pa \ mem_equal (s1\cpu_reg := new_cpu1\) (s2\cpu_reg := cpu_reg2\) pa" by (simp add: mem_equal_def) lemma user_accessible_mod_cpu_reg: "user_accessible (s\cpu_reg := new_cpu_reg\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_equal_mod_trap: "mem_equal s1 s2 pa \ mem_equal (s1\traps := new_traps1\) (s2\traps := traps2\) pa" by (simp add: mem_equal_def) lemma user_accessible_mod_trap: "user_accessible (s\traps := new_traps\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_equal_annul: "mem_equal s1 s2 pa \ mem_equal (s1\state_var := new_state_var, cpu_reg := new_cpu_reg\) (s2\state_var := new_state_var2, cpu_reg := new_cpu_reg2\) pa" by (simp add: mem_equal_def) lemma user_accessible_annul: "user_accessible (s\state_var := new_state_var, cpu_reg := new_cpu_reg\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_val_alt_10_mem_equal_0: "mem_equal s1 s2 pa \ mem_val_alt 10 (pa AND 68719476732) s1 = mem_val_alt 10 (pa AND 68719476732) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_mem_equal_1: "mem_equal s1 s2 pa \ mem_val_alt 10 ((pa AND 68719476732) + 1) s1 = mem_val_alt 10 ((pa AND 68719476732) + 1) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_mem_equal_2: "mem_equal s1 s2 pa \ mem_val_alt 10 ((pa AND 68719476732) + 2) s1 = mem_val_alt 10 ((pa AND 68719476732) + 2) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_mem_equal_3: "mem_equal s1 s2 pa \ mem_val_alt 10 ((pa AND 68719476732) + 3) s1 = mem_val_alt 10 ((pa AND 68719476732) + 3) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_mem_equal: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 10 (pa AND 68719476732) s1 = mem_val_alt 10 (pa AND 68719476732) s2 \ mem_val_alt 10 ((pa AND 68719476732) + 1) s1 = mem_val_alt 10 ((pa AND 68719476732) + 1) s2 \ mem_val_alt 10 ((pa AND 68719476732) + 2) s1 = mem_val_alt 10 ((pa AND 68719476732) + 2) s2 \ mem_val_alt 10 ((pa AND 68719476732) + 3) s1 = mem_val_alt 10 ((pa AND 68719476732) + 3) s2" using mem_val_alt_10_mem_equal_0 mem_val_alt_10_mem_equal_1 mem_val_alt_10_mem_equal_2 mem_val_alt_10_mem_equal_3 a1 by blast lemma mem_val_w32_10_mem_equal: assumes a1: "mem_equal s1 s2 a" shows "mem_val_w32 10 a s1 = mem_val_w32 10 a s2" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) using mem_val_alt_10_mem_equal a1 apply auto apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce by fastforce lemma mem_val_alt_8_mem_equal_0: "mem_equal s1 s2 pa \ mem_val_alt 8 (pa AND 68719476732) s1 = mem_val_alt 8 (pa AND 68719476732) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_mem_equal_1: "mem_equal s1 s2 pa \ mem_val_alt 8 ((pa AND 68719476732) + 1) s1 = mem_val_alt 8 ((pa AND 68719476732) + 1) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_mem_equal_2: "mem_equal s1 s2 pa \ mem_val_alt 8 ((pa AND 68719476732) + 2) s1 = mem_val_alt 8 ((pa AND 68719476732) + 2) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_mem_equal_3: "mem_equal s1 s2 pa \ mem_val_alt 8 ((pa AND 68719476732) + 3) s1 = mem_val_alt 8 ((pa AND 68719476732) + 3) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_mem_equal: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 8 (pa AND 68719476732) s1 = mem_val_alt 8 (pa AND 68719476732) s2 \ mem_val_alt 8 ((pa AND 68719476732) + 1) s1 = mem_val_alt 8 ((pa AND 68719476732) + 1) s2 \ mem_val_alt 8 ((pa AND 68719476732) + 2) s1 = mem_val_alt 8 ((pa AND 68719476732) + 2) s2 \ mem_val_alt 8 ((pa AND 68719476732) + 3) s1 = mem_val_alt 8 ((pa AND 68719476732) + 3) s2" using mem_val_alt_8_mem_equal_0 mem_val_alt_8_mem_equal_1 mem_val_alt_8_mem_equal_2 mem_val_alt_8_mem_equal_3 a1 by blast lemma mem_val_w32_8_mem_equal: assumes a1: "mem_equal s1 s2 a" shows "mem_val_w32 8 a s1 = mem_val_w32 8 a s2" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) using mem_val_alt_8_mem_equal a1 apply auto apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce by fastforce lemma load_word_mem_10_low_equal: assumes a1: "low_equal s1 s2" shows "load_word_mem s1 address 10 = load_word_mem s2 address 10" using a1 apply (simp add: low_equal_def load_word_mem_def) apply clarsimp apply (case_tac "virt_to_phys address (mmu s2) (mem s2) = None") apply auto apply (simp add: user_accessible_def) using mem_val_w32_10_mem_equal apply blast apply (simp add: user_accessible_def) using mem_val_w32_10_mem_equal by blast lemma load_word_mem_8_low_equal: assumes a1: "low_equal s1 s2" shows "load_word_mem s1 address 8 = load_word_mem s2 address 8" using a1 apply (simp add: low_equal_def load_word_mem_def) apply clarsimp apply (case_tac "virt_to_phys address (mmu s2) (mem s2) = None") apply auto apply (simp add: user_accessible_def) using mem_val_w32_8_mem_equal user_accessible_8 apply fastforce apply (simp add: user_accessible_def) using mem_val_w32_8_mem_equal user_accessible_8 by fastforce lemma mem_read_low_equal: assumes a1: "low_equal s1 s2 \ asi \ {8,10}" shows "fst (memory_read asi address s1) = fst (memory_read asi address s2)" proof (cases "asi = 8") case True then show ?thesis using a1 apply (simp add: low_equal_def) apply (simp add: memory_read_def) using a1 load_word_mem_8_low_equal apply auto apply (simp add: option.case_eq_if) by (simp add: option.case_eq_if) next case False then have "asi = 10" using a1 by auto then show ?thesis using a1 apply (simp add: low_equal_def) apply (simp add: memory_read_def) using a1 load_word_mem_10_low_equal apply auto apply (simp add: option.case_eq_if) by (simp add: option.case_eq_if) qed lemma read_mem_pc_low_equal: assumes a1: "low_equal s1 s2" shows "fst (memory_read 8 (cpu_reg_val PC s1) s1) = fst (memory_read 8 (cpu_reg_val PC s2) s2)" proof - have f2: "cpu_reg_val PC s1 = cpu_reg_val PC s2" using a1 by (simp add: low_equal_def cpu_reg_val_def) then show ?thesis using a1 f2 mem_read_low_equal by auto qed lemma dcache_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = dcache_mod c v s1 \ t2 = dcache_mod c v s2" shows "low_equal t1 t2" using a1 apply (simp add: low_equal_def) apply (simp add: dcache_mod_def) apply auto apply (simp add: user_accessible_mod_cache mem_equal_mod_cache) by (simp add: user_accessible_mod_cache mem_equal_mod_cache) lemma add_data_cache_low_equal: assumes a1: "low_equal s1 s2 \ t1 = add_data_cache s1 address w bm \ t2 = add_data_cache s2 address w bm" shows "low_equal t1 t2" using a1 apply (simp add: add_data_cache_def) apply (case_tac "bm AND 8 >> 3 = 1") apply auto apply (case_tac "bm AND 4 >> 2 = 1") apply auto apply (case_tac "bm AND 2 >> Suc 0 = 1") apply auto apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 2 >> Suc 0 = 1") apply auto apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 4 >> 2 = 1") apply auto apply (case_tac "bm AND 2 >> Suc 0 = 1") apply auto apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 2 >> Suc 0 = 1") apply auto apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) by (meson dcache_mod_low_equal) lemma mem_read2_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (memory_read (10::word8) address s1) \ t2 = snd (memory_read (10::word8) address s2)" shows "low_equal t1 t2" using a1 apply (simp add: memory_read_def) using a1 apply (auto simp add: sys_reg_low_equal mod_2_eq_odd) using a1 apply (simp add: load_word_mem_10_low_equal) apply (auto split: option.splits) using add_data_cache_low_equal apply force using add_data_cache_low_equal apply force done lemma mem_read_delayed_write_low_equal: assumes a1: "low_equal s1 s2 \ get_delayed_pool s1 = [] \ get_delayed_pool s2 = []" shows "fst (memory_read 8 (cpu_reg_val PC (delayed_pool_write s1)) (delayed_pool_write s1)) = fst (memory_read 8 (cpu_reg_val PC (delayed_pool_write s2)) (delayed_pool_write s2))" using a1 apply (simp add: delayed_pool_write_def) apply (simp add: Let_def) apply (simp add: get_delayed_write_def) by (simp add: read_mem_pc_low_equal) lemma global_reg_mod_low_equal: assumes a1: "low_equal s1 s2\ t1 = (global_reg_mod w n rd s1) \ t2 = (global_reg_mod w n rd s2)" shows "low_equal t1 t2" using a1 apply (induction n arbitrary: s1 s2) apply clarsimp apply auto apply (simp add: Let_def) apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal by blast lemma out_reg_mod_low_equal: assumes a1: "low_equal s1 s2\ t1 = (out_reg_mod w curr_win rd s1) \ t2 = (out_reg_mod w curr_win rd s2)" shows "low_equal t1 t2" using a1 apply (simp add: out_reg_mod_def Let_def) apply auto apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal apply fastforce apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal by blast lemma in_reg_mod_low_equal: assumes a1: "low_equal s1 s2\ t1 = (in_reg_mod w curr_win rd s1) \ t2 = (in_reg_mod w curr_win rd s2)" shows "low_equal t1 t2" using a1 apply (simp add: in_reg_mod_def Let_def) apply auto apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal apply fastforce apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal by blast lemma user_reg_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = user_reg_mod w curr_win rd s1 \ t2 = user_reg_mod w curr_win rd s2" shows "low_equal t1 t2" proof (cases "rd = 0") case True then show ?thesis using a1 by (simp add: user_reg_mod_def) next case False then have f1: "rd \ 0" by auto then show ?thesis proof (cases "0 < rd \ rd < 8") case True then show ?thesis using a1 f1 apply (simp add: user_reg_mod_def) using global_reg_mod_low_equal by blast next case False then have f2: "\ (0 < rd \ rd < 8)" by auto then show ?thesis proof (cases "7 < rd \ rd < 16") case True then show ?thesis using a1 f1 f2 apply (simp add: user_reg_mod_def) by (auto intro: out_reg_mod_low_equal) next case False then have f3: "\ (7 < rd \ rd < 16)" by auto then show ?thesis proof (cases "15 < rd \ rd < 24") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: user_reg_mod_def) apply (simp add: low_equal_def) apply clarsimp by (simp add: user_accessible_mod_user_reg mem_equal_mod_user_reg) next case False then show ?thesis using a1 f1 f2 f3 apply (simp add: user_reg_mod_def) by (auto intro: in_reg_mod_low_equal) qed qed qed qed lemma virt_to_phys_low_equal: "low_equal s1 s2 \ virt_to_phys addr (mmu s1) (mem s1) = virt_to_phys addr (mmu s2) (mem s2)" by (auto simp add: low_equal_def) lemma write_reg_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (snd (fst (write_reg w curr_win rd s1))) \ t2 = (snd (fst (write_reg w curr_win rd s2)))" shows "low_equal t1 t2" using a1 apply (simp add: write_reg_def) apply (simp add: simpler_modify_def) by (auto intro: user_reg_mod_low_equal) lemma write_cpu_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (write_cpu w cr s1)) \ t2 = (snd (fst (write_cpu w cr s2)))" shows "low_equal t1 t2" using a1 apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def) apply (simp add: low_equal_def) using user_accessible_mod_cpu_reg mem_equal_mod_cpu_reg by metis lemma cpu_reg_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = cpu_reg_mod w cr s1 \ t2 = cpu_reg_mod w cr s2" shows "low_equal t1 t2" using a1 apply (simp add: cpu_reg_mod_def) apply (simp add: low_equal_def) using user_accessible_mod_cpu_reg mem_equal_mod_cpu_reg by metis lemma load_sub2_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (snd (fst (load_sub2 address 10 rd curr_win w s1))) \ t2 = (snd (fst (load_sub2 address 10 rd curr_win w s2)))" shows "low_equal t1 t2" proof (cases "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = None") case True then have f0: "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = None" by auto have f1: "low_equal (snd (fst (write_reg w curr_win (rd AND 30) s1))) (snd (fst (write_reg w curr_win (rd AND 30) s2)))" using a1 by (auto intro: write_reg_low_equal) then have "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = None \ fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s2))))" using f0 by (blast intro: mem_read_low_equal) then have "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = None \ fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s2)))) = None" by auto then show ?thesis using a1 apply (simp add: load_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) using f1 apply (simp add: traps_low_equal) using f1 by (auto intro: mod_trap_low_equal) next case False then have f2: "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) \ None" by auto have f3: "low_equal (snd (fst (write_reg w curr_win (rd AND 30) s1))) (snd (fst (write_reg w curr_win (rd AND 30) s2)))" using a1 by (auto intro: write_reg_low_equal) then have f4: "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s2))))" using f2 by (blast intro: mem_read_low_equal) then have "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) \ None \ fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s2)))) \ None" using f2 by auto then show ?thesis using a1 apply (simp add: load_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) using f4 apply clarsimp using f3 by (auto intro: mem_read2_low_equal write_reg_low_equal) qed lemma load_sub3_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (load_sub3 instr curr_win rd (10::word8) address s1)) \ t2 = snd (fst (load_sub3 instr curr_win rd (10::word8) address s2))" shows "low_equal t1 t2" proof (cases "fst (memory_read 10 address s1) = None") case True then have "fst (memory_read 10 address s1) = None \ fst (memory_read 10 address s2) = None" using a1 by (auto simp add: mem_read_low_equal) then show ?thesis using a1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (auto simp add: traps_low_equal) by (auto intro: mod_trap_low_equal) next case False then have f1: "fst (memory_read 10 address s1) \ None \ fst (memory_read 10 address s2) \ None" using a1 by (auto simp add: mem_read_low_equal) then show ?thesis proof (cases "rd \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA \ fst instr = load_store_type LDUH \ fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUBA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSBA)") case True then show ?thesis using a1 f1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: mem_read_low_equal) by (meson mem_read2_low_equal write_reg_low_equal) next case False then show ?thesis using a1 f1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: mem_read_low_equal) by (meson load_sub2_low_equal mem_read2_low_equal) qed qed lemma ld_asi_user: "(fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUH \ fst instr = load_store_type LD \ fst instr = load_store_type LDD) \ ld_asi instr 0 = 10" apply (simp add: ld_asi_def) by auto lemma load_sub1_low_equal: assumes a1: "low_equal s1 s2 \ (fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUH \ fst instr = load_store_type LD \ fst instr = load_store_type LDD) \ t1 = snd (fst (load_sub1 instr rd 0 s1)) \ t2 = snd (fst (load_sub1 instr rd 0 s2))" shows "low_equal t1 t2" proof (cases "(fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0") case True then have "((fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0) \ ((fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis using a1 apply (simp add: load_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: get_curr_win3_low_equal) by (auto intro: get_curr_win2_low_equal mod_trap_low_equal) next case False then have f1: "\ ((fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0) \ \ ((fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1) \ 0)" by (metis assms get_addr_low_equal) show ?thesis proof - have "low_equal s1 s2 \ low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by auto then have f2: "low_equal s1 s2 \ low_equal (snd (fst (load_sub3 instr (fst (fst (get_curr_win () s2))) rd 10 (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) (snd (fst (get_curr_win () s1)))))) (snd (fst (load_sub3 instr (fst (fst (get_curr_win () s2))) rd 10 (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) (snd (fst (get_curr_win () s2))))))" using load_sub3_low_equal by blast show ?thesis using a1 unfolding load_sub1_def simpler_gets_def bind_def h1_def h2_def Let_def case_prod_unfold using f1 f2 apply clarsimp by (simp add: get_addr2_low_equal get_curr_win_low_equal ld_asi_user) qed qed lemma load_instr_low_equal: assumes a1: "low_equal s1 s2 \ (fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUBA \ fst instr = load_store_type LDUH \ fst instr = load_store_type LD \ fst instr = load_store_type LDA \ fst instr = load_store_type LDD) \ (((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (load_instr instr s1)) \ t2 = snd (fst (load_instr instr s2))" shows "low_equal t1 t2" proof - have "get_S (cpu_reg_val PSR s1) = 0 \ get_S (cpu_reg_val PSR s2) = 0" using a1 by (simp add: ucast_id) then show ?thesis using a1 apply (simp add: load_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) apply clarsimp apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (simp add: traps_low_equal) by (auto intro: mod_trap_low_equal load_sub1_low_equal) qed lemma st_data0_low_equal: "low_equal s1 s2 \ st_data0 instr curr_win rd addr s1 = st_data0 instr curr_win rd addr s2" apply (simp add: st_data0_def) by (simp add: user_reg_val_def low_equal_def) lemma store_word_mem_low_equal_none: "low_equal s1 s2 \ store_word_mem (add_data_cache s1 addr data bm) addr data bm 10 = None \ store_word_mem (add_data_cache s2 addr data bm) addr data bm 10 = None" apply (simp add: store_word_mem_def) proof - assume a1: "low_equal s1 s2" assume a2: "(case virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) of None \ None | Some pair \ if mmu_writable (get_acc_flag (snd pair)) 10 then Some (mem_mod_w32 10 (fst pair) bm data (add_data_cache s1 addr data bm)) else None) = None" have f3: "(if mmu_writable (get_acc_flag (snd (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s2 addr data bm)) else None) = (case Some (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) of None \ if mmu_writable (get_acc_flag (snd (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None)" by auto obtain pp :: "(word36 \ word8) option \ word36 \ word8" where f4: "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = None \ virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm))))" by (metis (no_types) option.exhaust) have f5: "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))" using a1 by (meson add_data_cache_low_equal virt_to_phys_low_equal) { assume "Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) \ (case Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s1 addr data bm)) else None)" then have "None = (if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s2 addr data bm)) else None)" by fastforce moreover { assume "(if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s2 addr data bm)) else None) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None)" then have "(case Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) of None \ if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None)" using f3 by simp then have "Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) \ (if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None) \ None" proof - have "(case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) = (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) \ Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) \ (if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None) \ None" by simp then show ?thesis using \(case Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) of None \ if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None)\ by force qed moreover { assume "Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))" then have "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) \ Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm))))" using f5 by simp } ultimately have "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) \ Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) \ Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm))))" using a2 by force } ultimately have "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) = None" by fastforce } then have "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) = None" using a2 by force then show "(case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) = None" using f5 f4 by force qed lemma memory_write_asi_low_equal_none: "low_equal s1 s2 \ memory_write_asi 10 addr bm data s1 = None \ memory_write_asi 10 addr bm data s2 = None" apply (simp add: memory_write_asi_def) by (simp add: store_word_mem_low_equal_none) lemma memory_write_low_equal_none: "low_equal s1 s2 \ memory_write 10 addr bm data s1 = None \ memory_write 10 addr bm data s2 = None" apply (simp add: memory_write_def) by (metis map_option_case memory_write_asi_low_equal_none option.map_disc_iff) lemma memory_write_low_equal_none2: "low_equal s1 s2 \ memory_write 10 addr bm data s2 = None \ memory_write 10 addr bm data s1 = None" apply (simp add: memory_write_def) by (metis low_equal_com memory_write_def memory_write_low_equal_none) lemma mem_context_val_9_unchanged: "mem_context_val 9 addr1 (mem s1) = mem_context_val 9 addr1 ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))" apply (simp add: mem_context_val_def) by (simp add: Let_def) lemma mem_context_val_w32_9_unchanged: "mem_context_val_w32 9 addr1 (mem s1) = mem_context_val_w32 9 addr1 ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))" apply (simp add: mem_context_val_w32_def) apply (simp add: Let_def) by (metis mem_context_val_9_unchanged) lemma ptd_lookup_unchanged_4: "ptd_lookup va ptp (mem s1) 4 = ptd_lookup va ptp ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 4" by auto lemma ptd_lookup_unchanged_3: "ptd_lookup va ptp (mem s1) 3 = ptd_lookup va ptp ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 3" proof (cases "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) = None") case True then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) = None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None" using mem_context_val_w32_9_unchanged by metis then show ?thesis by auto next case False then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" using mem_context_val_w32_9_unchanged by metis then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ (\y. (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) = Some y) \ (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))= Some y))" using mem_context_val_w32_9_unchanged by metis then show ?thesis apply auto by (simp add: Let_def) qed lemma ptd_lookup_unchanged_2: "ptd_lookup va ptp (mem s1) 2 = ptd_lookup va ptp ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 2" proof (cases "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) = None") case True then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) = None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None" using mem_context_val_w32_9_unchanged by metis then show ?thesis by auto next case False then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" using mem_context_val_w32_9_unchanged by metis then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ (\y. (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) = Some y) \ (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))= Some y))" using mem_context_val_w32_9_unchanged by metis then show ?thesis apply auto using ptd_lookup_unchanged_3 unfolding Let_def by auto qed lemma ptd_lookup_unchanged_1: "ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1" proof (cases "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) = None") case True then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) = None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None" using mem_context_val_w32_9_unchanged by metis then show ?thesis by auto next case False then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" using mem_context_val_w32_9_unchanged by metis then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ (\y. (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) = Some y) \ (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))= Some y))" using mem_context_val_w32_9_unchanged by metis then show ?thesis apply auto using ptd_lookup_unchanged_2 unfolding Let_def proof - fix y :: word32 have "(y AND 3 \ 0 \ y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = None) \ (y AND 3 = 1 \ y AND 3 \ 2 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8))) \ (y AND 3 = 0 \ (y AND 3 \ 1 \ (y AND 3 \ 0 \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) = None) \ (y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1)) \ (y AND 3 = 1 \ (y AND 3 \ 2 \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) = Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y)) \ (y AND 3 = 2 \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) = None)))) \ (y AND 3 = 1 \ (y AND 3 \ 2 \ (y AND 3 \ 0 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8)) \ (y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y)) \ (y AND 3 = 1 \ y AND 3 = 2 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8)))) \ (y AND 3 = 2 \ y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = None) \ (y AND 3 = 1 \ y AND 3 \ 2 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8))))) \ (\w. mem s1 w = ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) w)" by (metis (no_types) One_nat_def Suc_1 Suc_eq_plus1 ptd_lookup_unchanged_2) then show "(if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) = (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None)" proof - have f1: "2 = Suc 0 + 1" by (metis One_nat_def Suc_1 Suc_eq_plus1) { assume "y AND 3 = 1" moreover { assume "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None)" have "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) = (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None)" by presburger moreover { assume "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1)" then have "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ ptd_lookup va (y AND 4294967292) (mem s1) 2" by (metis One_nat_def Suc_1 Suc_eq_plus1 ptd_lookup_unchanged_2) then have ?thesis using f1 by auto } ultimately have ?thesis by blast } ultimately have ?thesis by blast } then show ?thesis by presburger qed qed qed lemma virt_to_phys_unchanged_sub1: assumes a1: "(let context_table_entry = (v1 >> 11 << 11) OR (v2 AND 511 << 2) in Let (mem_context_val_w32 (word_of_int 9) (ucast context_table_entry) (mem s1)) (case_option None (\lvl1_page_table. ptd_lookup va lvl1_page_table (mem s1) 1))) = (let context_table_entry = (v1 >> 11 << 11) OR (v2 AND 511 << 2) in Let (mem_context_val_w32 (word_of_int 9) (ucast context_table_entry) (mem s2)) (case_option None (\lvl1_page_table. ptd_lookup va lvl1_page_table (mem s2) 1)))" shows "(let context_table_entry = (v1 >> 11 << 11) OR (v2 AND 511 << 2) in Let (mem_context_val_w32 (word_of_int 9) (ucast context_table_entry) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))) (case_option None (\lvl1_page_table. ptd_lookup va lvl1_page_table ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1))) = (let context_table_entry = (v1 >> 11 << 11) OR (v2 AND 511 << 2) in Let (mem_context_val_w32 (word_of_int 9) (ucast context_table_entry) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None)))) (case_option None (\lvl1_page_table. ptd_lookup va lvl1_page_table ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)))" proof - from a1 have "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s1) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)" unfolding Let_def by auto then have "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)" using mem_context_val_w32_9_unchanged by (metis word_numeral_alt) then have "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" using ptd_lookup_unchanged_1 proof - obtain ww :: "word32 option \ word32" where f1: "\z. (z = None \ z = Some (ww z)) \ (z \ None \ (\w. z \ Some w))" by moura then have f2: "(mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))))) \ (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ (\w. mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ Some w))" by blast then have f3: "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) \ None \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s1) 1)" by (metis (no_types) \\val va s1 ptp addr. ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1\ option.case(2) option.simps(4)) have f4: "mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))))) \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))))) \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) = (case Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))))) of None \ None | Some w \ ptd_lookup va w ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" by (metis (no_types) \(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)\ \\val va s1 ptp addr. ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1\ option.case(2)) have f5: "(mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = None \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None)))))) \ (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) \ None \ (\w. mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) \ Some w))" using f1 by blast { assume "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" { assume "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s2) 1) \ None \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s2) 1) \ None" then have "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s2) 1) \ None \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" by (metis (no_types) \(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)\ option.simps(4)) then have ?thesis using f5 f4 f2 by force } then have ?thesis using f5 f3 by (metis (no_types) \(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)\ \\val va s1 ptp addr. ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1\ option.case(2) option.simps(4)) } then show ?thesis by blast qed then show ?thesis unfolding Let_def by auto qed lemma virt_to_phys_unchanged: assumes a1: "(\va. virt_to_phys va (mmu s2) (mem s1) = virt_to_phys va (mmu s2) (mem s2))" shows "(\va. virt_to_phys va (mmu s2) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = virt_to_phys va (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))))" proof (cases "registers (mmu s2) CR AND 1 \ 0") case True then have f1: "registers (mmu s2) CR AND 1 \ 0" by auto then show ?thesis proof (cases "mmu_reg_val (mmu s2) 256 = None") case True then show ?thesis by (simp add: virt_to_phys_def) next case False then have f2: "mmu_reg_val (mmu s2) 256 \ None" by auto then show ?thesis proof (cases "mmu_reg_val (mmu s2) 512 = None") case True then show ?thesis using f1 f2 apply (simp add: virt_to_phys_def) by auto next case False then show ?thesis using f1 f2 a1 apply (simp add: virt_to_phys_def) apply clarify using virt_to_phys_unchanged_sub1 by fastforce qed qed next case False then show ?thesis by (simp add: virt_to_phys_def) qed lemma virt_to_phys_unchanged2_sub1: "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" proof (cases "mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) = None") case True then have "mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) = None \ mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = None" using mem_context_val_w32_9_unchanged by metis then show ?thesis by auto next case False then have "mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) \ None \ (\y. mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) = Some y \ mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some y)" using mem_context_val_w32_9_unchanged by metis then show ?thesis using ptd_lookup_unchanged_1 by fastforce qed lemma virt_to_phys_unchanged2: "virt_to_phys va (mmu s2) (mem s2) = virt_to_phys va (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None)))" proof (cases "registers (mmu s2) CR AND 1 \ 0") case True then have f1: "registers (mmu s2) CR AND 1 \ 0" by auto then show ?thesis proof (cases "mmu_reg_val (mmu s2) 256 = None") case True then show ?thesis by (simp add: virt_to_phys_def) next case False then have f2: "mmu_reg_val (mmu s2) 256 \ None" by auto then show ?thesis proof (cases "mmu_reg_val (mmu s2) 512 = None") case True then show ?thesis using f1 f2 apply (simp add: virt_to_phys_def) by auto next case False then show ?thesis using f1 f2 apply (simp add: virt_to_phys_def) apply clarify unfolding Let_def using virt_to_phys_unchanged2_sub1 by auto qed qed next case False then show ?thesis by (simp add: virt_to_phys_def) qed lemma virt_to_phys_unchanged_low_equal: assumes a1: "low_equal s1 s2" shows "(\va. virt_to_phys va (mmu s2) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = virt_to_phys va (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))))" using a1 apply (simp add: low_equal_def) using virt_to_phys_unchanged by metis lemma mmu_low_equal: "low_equal s1 s2 \ mmu s1 = mmu s2" by (simp add: low_equal_def) lemma mem_val_alt_8_unchanged0: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 8 (pa AND 68719476732) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 (pa AND 68719476732) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_unchanged1: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 8 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_unchanged2: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 8 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_unchanged3: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 8 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_unchanged: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 8 (pa AND 68719476732) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 (pa AND 68719476732) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) \ mem_val_alt 8 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) \ mem_val_alt 8 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) \ mem_val_alt 8 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 8 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" using a1 mem_val_alt_8_unchanged0 mem_val_alt_8_unchanged1 mem_val_alt_8_unchanged2 mem_val_alt_8_unchanged3 by blast lemma mem_val_w32_8_unchanged: assumes a1: "mem_equal s1 s2 a" shows "mem_val_w32 8 a (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_w32 8 a (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) using mem_val_alt_8_unchanged a1 apply auto apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce by fastforce lemma load_word_mem_8_unchanged: assumes a1: "low_equal s1 s2 \ load_word_mem s1 addra 8 = load_word_mem s2 addra 8" shows "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 8" proof (cases "virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None") case True then have "virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None \ virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = None" using a1 apply (auto simp add: mmu_low_equal) using a1 virt_to_phys_unchanged_low_equal by metis then show ?thesis by (simp add: load_word_mem_def) next case False then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some p" using a1 apply (auto simp add: mmu_low_equal) using a1 virt_to_phys_unchanged_low_equal by metis then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s1) (mem s1) = Some p \ virt_to_phys addra (mmu s2) (mem s2) = Some p" using virt_to_phys_unchanged2 by metis then show ?thesis using a1 apply (simp add: load_word_mem_def) apply auto apply (simp add: low_equal_def) apply (simp add: user_accessible_def) using mem_val_w32_8_unchanged a1 user_accessible_8 by (metis snd_conv) qed lemma load_word_mem_select_8: assumes a1: "fst (case load_word_mem s1 addra 8 of None \ (None, s1) | Some w \ (Some w, add_instr_cache s1 addra w 15)) = fst (case load_word_mem s2 addra 8 of None \ (None, s2) | Some w \ (Some w, add_instr_cache s2 addra w 15))" shows "load_word_mem s1 addra 8 = load_word_mem s2 addra 8" using a1 by (metis (mono_tags, lifting) fst_conv not_None_eq option.simps(4) option.simps(5)) lemma memory_read_8_unchanged: assumes a1: "low_equal s1 s2 \ fst (memory_read 8 addra s1) = fst (memory_read 8 addra s2)" shows "fst (memory_read 8 addra (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\)) = fst (memory_read 8 addra (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\))" proof (cases "sys_reg s1 CCR AND 1 = 0") case True then have "sys_reg s1 CCR AND 1 = 0 \ sys_reg s2 CCR AND 1 = 0" using a1 sys_reg_low_equal by fastforce then show ?thesis using a1 apply (simp add: memory_read_def) using load_word_mem_8_unchanged by blast next case False then have f1: "sys_reg s1 CCR AND 1 \ 0 \ sys_reg s2 CCR AND 1 \ 0" using a1 sys_reg_low_equal by fastforce then show ?thesis using a1 proof (cases "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = None") case True then have "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = None \ load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 8 = None" using a1 f1 apply (simp add: memory_read_def) apply clarsimp using load_word_mem_select_8 load_word_mem_8_unchanged by fastforce then show ?thesis by (simp add: memory_read_def) next case False then have "\y. load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = Some y" by auto then have "\y. load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = Some y \ load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 8 = Some y" using a1 f1 apply (simp add: memory_read_def) apply clarsimp using load_word_mem_select_8 load_word_mem_8_unchanged by fastforce then show ?thesis using a1 f1 apply (simp add: memory_read_def) by auto qed qed lemma mem_val_alt_mod: assumes a1: "addr1 \ addr2" shows "mem_val_alt 10 addr1 s = mem_val_alt 10 addr1 (s\mem := (mem s)(10 := mem s 10(addr2 \ val), 11 := (mem s 11)(addr2 := None))\)" using a1 apply (simp add: mem_val_alt_def) by (simp add: Let_def) lemma mem_val_alt_mod2: "mem_val_alt 10 addr (s\mem := (mem s)(10 := mem s 10(addr \ val), 11 := (mem s 11)(addr := None))\) = Some val" by (simp add: mem_val_alt_def) lemma mem_val_alt_10_unchanged0: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 10 (pa AND 68719476732) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 (pa AND 68719476732) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_unchanged1: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 10 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_unchanged2: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 10 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_unchanged3: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 10 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_unchanged: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 10 (pa AND 68719476732) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 (pa AND 68719476732) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) \ mem_val_alt 10 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) \ mem_val_alt 10 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) \ mem_val_alt 10 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_alt 10 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" using a1 mem_val_alt_10_unchanged0 mem_val_alt_10_unchanged1 mem_val_alt_10_unchanged2 mem_val_alt_10_unchanged3 by blast lemma mem_val_w32_10_unchanged: assumes a1: "mem_equal s1 s2 a" shows "mem_val_w32 10 a (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) = mem_val_w32 10 a (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) using mem_val_alt_10_unchanged a1 apply auto apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce by fastforce lemma is_accessible: "low_equal s1 s2 \ virt_to_phys addra (mmu s1) (mem s1) = Some (a, b) \ virt_to_phys addra (mmu s2) (mem s2) = Some (a, b) \ mmu_readable (get_acc_flag b) 10 \ mem_equal s1 s2 a" apply (simp add: low_equal_def) apply (simp add: user_accessible_def) by fastforce lemma load_word_mem_10_unchanged: assumes a1: "low_equal s1 s2 \ load_word_mem s1 addra 10 = load_word_mem s2 addra 10" shows "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 10" proof (cases "virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None") case True then have "virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None \ virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = None" using a1 apply (auto simp add: mmu_low_equal) using a1 virt_to_phys_unchanged_low_equal by metis then show ?thesis by (simp add: load_word_mem_def) next case False then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some p" using a1 apply (auto simp add: mmu_low_equal) using a1 virt_to_phys_unchanged_low_equal by metis then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s1) (mem s1) = Some p \ virt_to_phys addra (mmu s2) (mem s2) = Some p" using virt_to_phys_unchanged2 by metis then show ?thesis using a1 apply (simp add: load_word_mem_def) apply auto apply (simp add: low_equal_def) apply (simp add: user_accessible_def) using mem_val_w32_10_unchanged a1 by metis qed lemma load_word_mem_select_10: assumes a1: "fst (case load_word_mem s1 addra 10 of None \ (None, s1) | Some w \ (Some w, add_data_cache s1 addra w 15)) = fst (case load_word_mem s2 addra 10 of None \ (None, s2) | Some w \ (Some w, add_data_cache s2 addra w 15))" shows "load_word_mem s1 addra 10 = load_word_mem s2 addra 10" using a1 by (metis (mono_tags, lifting) fst_conv not_None_eq option.simps(4) option.simps(5)) lemma memory_read_10_unchanged: assumes a1: "low_equal s1 s2 \ fst (memory_read 10 addra s1) = fst (memory_read 10 addra s2)" shows "fst (memory_read 10 addra (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\)) = fst (memory_read 10 addra (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\))" proof (cases "sys_reg s1 CCR AND 1 = 0") case True then have "sys_reg s1 CCR AND 1 = 0 \ sys_reg s2 CCR AND 1 = 0" using a1 sys_reg_low_equal by fastforce then show ?thesis using a1 apply (simp add: memory_read_def) using load_word_mem_10_unchanged by blast next case False then have f1: "sys_reg s1 CCR AND 1 \ 0 \ sys_reg s2 CCR AND 1 \ 0" using a1 sys_reg_low_equal by fastforce then show ?thesis using a1 proof (cases "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = None") case True then have "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = None \ load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 10 = None" using a1 f1 apply (simp add: memory_read_def) apply clarsimp using load_word_mem_select_10 load_word_mem_10_unchanged by fastforce then show ?thesis by (simp add: memory_read_def) next case False then have "\y. load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = Some y" by auto then have "\y. load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = Some y \ load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 10 = Some y" using a1 f1 apply (simp add: memory_read_def) apply clarsimp using load_word_mem_select_10 load_word_mem_10_unchanged by fastforce then show ?thesis using a1 f1 apply (simp add: memory_read_def) by auto qed qed lemma state_mem_mod_1011_low_equal_sub1: assumes a1: "(\va. virt_to_phys va (mmu s2) (mem s1) = virt_to_phys va (mmu s2) (mem s2)) \ (\pa. (\va b. virt_to_phys va (mmu s2) (mem s2) = Some (pa, b) \ mmu_readable (get_acc_flag b) 10) \ mem_equal s1 s2 pa) \ mmu s1 = mmu s2 \ virt_to_phys va (mmu s2) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (pa, b) \ mmu_readable (get_acc_flag b) 10" shows "mem_equal s1 s2 pa" proof - have "virt_to_phys va (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (pa, b)" using a1 by auto then have "virt_to_phys va (mmu s1) (mem s1) = Some (pa, b)" using virt_to_phys_unchanged2 by metis then have "virt_to_phys va (mmu s2) (mem s2) = Some (pa, b)" using a1 by auto then show ?thesis using a1 by auto qed lemma mem_equal_unchanged: assumes a1: "mem_equal s1 s2 pa" shows "mem_equal (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\) pa" using a1 apply (simp add: mem_equal_def) by auto lemma state_mem_mod_1011_low_equal: assumes a1: "low_equal s1 s2 \ t1 = s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\ \ t2 = s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\" shows "low_equal t1 t2" using a1 apply (simp add: low_equal_def) apply (simp add: user_accessible_def) apply auto apply (simp add: assms virt_to_phys_unchanged_low_equal) using state_mem_mod_1011_low_equal_sub1 mem_equal_unchanged apply metis apply (metis virt_to_phys_unchanged2) using state_mem_mod_1011_low_equal_sub1 mem_equal_unchanged by metis lemma mem_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (mem_mod 10 addr val s1) \ t2 = (mem_mod 10 addr val s2)" shows "low_equal t1 t2" using a1 apply (simp add: mem_mod_def) by (auto intro: state_mem_mod_1011_low_equal) lemma mem_mod_w32_low_equal: assumes a1: "low_equal s1 s2 \ t1 = mem_mod_w32 10 a bm data s1 \ t2 = mem_mod_w32 10 a bm data s2" shows "low_equal t1 t2" using a1 apply (simp add: mem_mod_w32_def) apply (simp add: Let_def) by (meson mem_mod_low_equal) lemma store_word_mem_low_equal: assumes a1: "low_equal s1 s2 \ Some t1 = store_word_mem s1 addr data bm 10 \ Some t2 = store_word_mem s2 addr data bm 10" shows "low_equal t1 t2" using a1 apply (simp add: store_word_mem_def) apply (auto simp add: virt_to_phys_low_equal) apply (case_tac "virt_to_phys addr (mmu s2) (mem s2) = None") apply auto apply (case_tac "mmu_writable (get_acc_flag b) 10") apply auto using mem_mod_w32_low_equal by blast lemma memory_write_asi_low_equal: assumes a1: "low_equal s1 s2 \ Some t1 = memory_write_asi 10 addr bm data s1 \ Some t2 = memory_write_asi 10 addr bm data s2" shows "low_equal t1 t2" using a1 apply (simp add: memory_write_asi_def) by (meson add_data_cache_low_equal store_word_mem_low_equal) lemma store_barrier_pending_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = store_barrier_pending_mod False s1 \ t2 = store_barrier_pending_mod False s2" shows "low_equal t1 t2" using a1 apply (simp add: store_barrier_pending_mod_def) apply clarsimp using a1 apply (auto simp add: state_var_low_equal) by (auto intro: state_var2_low_equal) lemma memory_write_low_equal: assumes a1: "low_equal s1 s2 \ Some t1 = memory_write 10 addr bm data s1 \ Some t2 = memory_write 10 addr bm data s2" shows "low_equal t1 t2" apply (case_tac "memory_write_asi 10 addr bm data s1 = None") using a1 apply (simp add: memory_write_def) apply (case_tac "memory_write_asi 10 addr bm data s2 = None") apply (meson assms low_equal_com memory_write_asi_low_equal_none) using a1 apply (simp add: memory_write_def) apply auto by (metis memory_write_asi_low_equal store_barrier_pending_mod_low_equal) lemma memory_write_low_equal2: assumes a1: "low_equal s1 s2 \ Some t1 = memory_write 10 addr bm data s1" shows "\t2. Some t2 = memory_write 10 addr bm data s2" using a1 apply (simp add: memory_write_def) apply auto by (metis (full_types) memory_write_def memory_write_low_equal_none2 not_None_eq) lemma store_sub2_low_equal_sub1: assumes a1: "low_equal s1 s2 \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s1 = Some y \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = Some ya" shows "low_equal (y\traps := insert data_access_exception (traps y)\) (ya\traps := insert data_access_exception (traps ya)\)" proof - from a1 have f1: "low_equal y ya" using memory_write_low_equal by metis then have "traps y = traps ya" by (simp add: low_equal_def) then show ?thesis using f1 mod_trap_low_equal by fastforce qed lemma store_sub2_low_equal_sub2: assumes a1: "low_equal s1 s2 \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s1 = Some y \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = Some ya \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) y) y = None \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) ya) ya = Some yb" shows "False" proof - from a1 have f1: "low_equal y ya" using memory_write_low_equal by metis then have "(user_reg_val curr_win (rd OR 1) y) = (user_reg_val curr_win (rd OR 1) ya)" by (simp add: low_equal_def user_reg_val_def) then show ?thesis using a1 using f1 memory_write_low_equal_none by fastforce qed lemma store_sub2_low_equal_sub3: assumes a1: "low_equal s1 s2 \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s1 = Some y \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = Some ya \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) y) y = Some yb \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) ya) ya = None" shows "False" proof - from a1 have f1: "low_equal y ya" using memory_write_low_equal by metis then have "(user_reg_val curr_win (rd OR 1) y) = (user_reg_val curr_win (rd OR 1) ya)" by (simp add: low_equal_def user_reg_val_def) then show ?thesis using a1 using f1 memory_write_low_equal_none2 by fastforce qed lemma store_sub2_low_equal_sub4: assumes a1: "low_equal s1 s2 \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s1 = Some y \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = Some ya \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) y) y = Some yb \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) ya) ya = Some yc" shows "low_equal yb yc" proof - from a1 have f1: "low_equal y ya" using memory_write_low_equal by metis then have "(user_reg_val curr_win (rd OR 1) y) = (user_reg_val curr_win (rd OR 1) ya)" by (simp add: low_equal_def user_reg_val_def) then show ?thesis using a1 f1 by (metis memory_write_low_equal) qed lemma store_sub2_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (store_sub2 instr curr_win rd 10 addr s1)) \ t2 = snd (fst (store_sub2 instr curr_win rd 10 addr s2))" shows "low_equal t1 t2" proof (cases "memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s1) s1 = None") case True then have "memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s1) s1 = None \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = None" using a1 by (metis memory_write_low_equal_none st_data0_low_equal) then show ?thesis using a1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold return_def) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) using mod_trap_low_equal traps_low_equal by fastforce next case False then have f1: "memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s1) s1 \ None \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 \ None" using a1 by (metis memory_write_low_equal_none2 st_data0_low_equal) then show ?thesis proof (cases "(fst instr) \ {load_store_type STD,load_store_type STDA}") case True then show ?thesis using a1 f1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply (simp add: bind_def case_prod_unfold) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: case_prod_unfold bind_def h1_def h2_def Let_def simpler_modify_def) apply (simp add: simpler_gets_def) apply auto apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (simp add: st_data0_low_equal) apply (simp add: store_sub2_low_equal_sub1) apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub2 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub3 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub4 apply blast apply (simp add: st_data0_low_equal) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) using store_sub2_low_equal_sub1 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub2 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub3 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub4 by blast next case False then show ?thesis using a1 f1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply (simp add: bind_def case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def) apply (simp add: st_data0_low_equal) using memory_write_low_equal by metis qed qed lemma store_sub1_low_equal: assumes a1: "low_equal s1 s2 \ (fst instr = load_store_type STB \ fst instr = load_store_type STH \ fst instr = load_store_type ST \ fst instr = load_store_type STD) \ t1 = snd (fst (store_sub1 instr rd 0 s1)) \ t2 = snd (fst (store_sub1 instr rd 0 s2))" shows "low_equal t1 t2" proof (cases "(fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0") case True then have "((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0) \ ((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis using a1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: get_curr_win3_low_equal) by (auto intro: get_curr_win2_low_equal mod_trap_low_equal) next case False then have f1: "\ ((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0) \ \ ((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis proof (cases "(fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0") case True then have "(fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0 \ (fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis using a1 f1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: get_curr_win3_low_equal) by (auto intro: get_curr_win2_low_equal mod_trap_low_equal) next case False then have "\((fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0) \ \((fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then have f2: "\((fst instr = load_store_type ST \ fst instr = load_store_type STA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0) \ \((fst instr = load_store_type ST \ fst instr = load_store_type STA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0)" by auto then show ?thesis proof (cases "(fst instr \ {load_store_type STD,load_store_type STDA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0") case True then have "(fst instr \ {load_store_type STD,load_store_type STDA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0 \ (fst instr \ {load_store_type STD,load_store_type STDA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis using a1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply auto apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (simp add: get_curr_win3_low_equal) by (auto intro: get_curr_win2_low_equal mod_trap_low_equal) next case False then have "\ (fst instr \ {load_store_type STD, load_store_type STDA} \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0) \ \ (fst instr \ {load_store_type STD, load_store_type STDA} \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then have f3: "\ ((fst instr = load_store_type STD \ fst instr = load_store_type STDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0) \ \ ((fst instr = load_store_type STD \ fst instr = load_store_type STDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0)" by auto show ?thesis using a1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (unfold case_prod_beta) apply (simp add: f1 f2 f3) apply (simp_all add: st_asi_def) using a1 apply clarsimp apply (simp add: get_curr_win_low_equal get_addr2_low_equal) by (metis store_sub2_low_equal get_curr_win2_low_equal) qed qed qed lemma store_instr_low_equal: assumes a1: "low_equal s1 s2 \ (fst instr = load_store_type STB \ fst instr = load_store_type STH \ fst instr = load_store_type ST \ fst instr = load_store_type STA \ fst instr = load_store_type STD) \ (((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (store_instr instr s1)) \ t2 = snd (fst (store_instr instr s2))" shows "low_equal t1 t2" proof - have "get_S (cpu_reg_val PSR s1) = 0 \ get_S (cpu_reg_val PSR s2) = 0" using a1 by (simp add: ucast_id) then show ?thesis using a1 apply (simp add: store_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) apply clarsimp apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (simp add: traps_low_equal) by (auto intro: mod_trap_low_equal store_sub1_low_equal) qed lemma sethi_low_equal: "low_equal s1 s2 \ t1 = snd (fst (sethi_instr instr s1)) \ t2 = snd (fst (sethi_instr instr s2)) \ low_equal t1 t2" apply (simp add: sethi_instr_def) apply (simp add: Let_def) apply (case_tac "get_operand_w5 (snd instr ! Suc 0) \ 0") apply auto apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: get_curr_win_low_equal) using get_curr_win2_low_equal write_reg_low_equal apply metis by (simp add: return_def) lemma nop_low_equal: "low_equal s1 s2 \ t1 = snd (fst (nop_instr instr s1)) \ t2 = snd (fst (nop_instr instr s2)) \ low_equal t1 t2" apply (simp add: nop_instr_def) by (simp add: return_def) lemma logical_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (logical_instr_sub1 instr_name result s1)) \ t2 = snd (fst (logical_instr_sub1 instr_name result s2))" shows "low_equal t1 t2" proof (cases "instr_name = logic_type ANDcc \ instr_name = logic_type ANDNcc \ instr_name = logic_type ORcc \ instr_name = logic_type ORNcc \ instr_name = logic_type XORcc \ instr_name = logic_type XNORcc") case True then show ?thesis using a1 apply (simp add: logical_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: logical_new_psr_val_def) using write_cpu_low_equal cpu_reg_val_low_equal by fastforce next case False then show ?thesis using a1 apply (simp add: logical_instr_sub1_def) by (simp add: return_def) qed lemma logical_instr_low_equal: "low_equal s1 s2 \ t1 = snd (fst (logical_instr instr s1)) \ t2 = snd (fst (logical_instr instr s2)) \ low_equal t1 t2" apply (simp add: logical_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply auto apply (simp_all add: get_curr_win_low_equal) apply (simp_all add: get_operand2_low_equal) using logical_instr_sub1_low_equal get_operand2_low_equal get_curr_win2_low_equal write_reg_low_equal user_reg_val_low_equal proof - assume a1: "low_equal s1 s2" assume "t2 = snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" assume "t1 = snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))" have "\w wa. user_reg_val w wa (snd (fst (get_curr_win () s2))) = user_reg_val w wa (snd (fst (get_curr_win () s1)))" using a1 by (metis (no_types) get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using a1 by (metis (no_types) get_curr_win2_low_equal logical_instr_sub1_low_equal write_reg_low_equal) next assume a2: "low_equal s1 s2" assume "t1 = snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))" assume "t2 = snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))" have "\w wa. user_reg_val w wa (snd (fst (get_curr_win () s2))) = user_reg_val w wa (snd (fst (get_curr_win () s1)))" using a2 by (metis (no_types) get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" proof - have "low_equal (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" by (meson a2 get_curr_win2_low_equal logical_instr_sub1_low_equal write_reg_low_equal) then show ?thesis using \\wa w. user_reg_val w wa (snd (fst (get_curr_win () s2))) = user_reg_val w wa (snd (fst (get_curr_win () s1)))\ by presburger qed qed lemma shift_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (shift_instr instr s1)) \ t2 = snd (fst (shift_instr instr s2))" shows "low_equal t1 t2" proof (cases "(fst instr = shift_type SLL) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply auto apply (simp_all add: get_curr_win_low_equal) proof - assume a1: "low_equal s1 s2" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) << unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) << unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" have "\w wa wb. low_equal (snd (fst (write_reg w wa wb s1))) (snd (fst (write_reg w wa wb s2)))" using a1 by (metis write_reg_low_equal) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) << unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) << unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a1 by (simp add: get_curr_win_def simpler_gets_def user_reg_val_low_equal) next assume a2: "low_equal s1 s2" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) << unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) << unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" have "\w wa wb. low_equal (snd (fst (write_reg w wa wb s1))) (snd (fst (write_reg w wa wb s2)))" using a2 by (metis write_reg_low_equal) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) << unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) << unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" proof - assume a1: "\w wa wb. low_equal (snd (fst (write_reg w wa wb s1))) (snd (fst (write_reg w wa wb s2)))" have "\u s. fst (get_curr_win u s) = (ucast (get_CWP (cpu_reg_val PSR s))::'a word, s)" by (simp add: get_curr_win_def simpler_gets_def) then show ?thesis using a1 assms user_reg_val_low_equal by fastforce qed qed next case False then have f1: "\((fst instr = shift_type SLL) \ (get_operand_w5 ((snd instr)!3) \ 0))" by auto then show ?thesis proof (cases "(fst instr = shift_type SRL) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 f1 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply auto apply (simp_all add: get_curr_win_low_equal) proof - assume a1: "low_equal s1 s2" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" have "\u s. fst (get_curr_win u s) = (ucast (get_CWP (cpu_reg_val PSR s))::'a word, s)" by (simp add: get_curr_win_def simpler_gets_def) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a1 user_reg_val_low_equal write_reg_low_equal by fastforce next assume a2: "low_equal s1 s2" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" have "\u s. fst (get_curr_win u s) = (ucast (get_CWP (cpu_reg_val PSR s))::'a word, s)" by (simp add: get_curr_win_def simpler_gets_def) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a2 user_reg_val_low_equal write_reg_low_equal by fastforce qed next case False then have f2: "\((fst instr = shift_type SRL) \ (get_operand_w5 ((snd instr)!3) \ 0))" by auto then show ?thesis proof (cases "(fst instr = shift_type SRA) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 f1 f2 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply auto apply (simp_all add: get_curr_win_low_equal) proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >>> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >>> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" have "\w wa. user_reg_val wa w (snd (fst (get_curr_win () s1))) = user_reg_val wa w (snd (fst (get_curr_win () s2)))" using a1 by (meson get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >>> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >>> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a1 by (metis (no_types) get_curr_win2_low_equal write_reg_low_equal) next assume a2: "low_equal s1 s2" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >>> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >>> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" have "\w wa. user_reg_val wa w (snd (fst (get_curr_win () s1))) = user_reg_val wa w (snd (fst (get_curr_win () s2)))" using a2 by (meson get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >>> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >>> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a2 get_curr_win2_low_equal write_reg_low_equal by fastforce qed next case False then show ?thesis using a1 f1 f2 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply (simp add: return_def) using get_curr_win2_low_equal by blast qed qed qed lemma add_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (add_instr_sub1 instr_name result rs1_val operand2 s1)) \ t2 = snd (fst (add_instr_sub1 instr_name result rs1_val operand2 s2))" shows "low_equal t1 t2" proof (cases "instr_name = arith_type ADDcc \ instr_name = arith_type ADDXcc") case True then show ?thesis using a1 apply (simp add: add_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (clarsimp simp add: cpu_reg_val_low_equal) using write_cpu_low_equal by blast next case False then show ?thesis using a1 apply (simp add: add_instr_sub1_def) by (simp add: return_def) qed lemma add_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (add_instr instr s1)) \ t2 = snd (fst (add_instr instr s2))" shows "low_equal t1 t2" proof - have f1: "low_equal s1 s2 \ t1 = snd (fst (add_instr_sub1 (fst instr) (if fst instr = arith_type ADD \ fst instr = arith_type ADDcc then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then if fst instr = arith_type ADD \ fst instr = arith_type ADDcc then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1))))) else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))))) \ t2 = snd (fst (add_instr_sub1 (fst instr) (if fst instr = arith_type ADD \ fst instr = arith_type ADDcc then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then if fst instr = arith_type ADD \ fst instr = arith_type ADDcc then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))) else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" using a1 apply (simp add: add_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: case_prod_unfold) then show ?thesis proof (cases "get_operand_w5 (snd instr ! 3) \ 0") case True then have f2: "get_operand_w5 (snd instr ! 3) \ 0" by auto then show ?thesis proof (cases "fst instr = arith_type ADD \ fst instr = arith_type ADDcc") case True then show ?thesis using f1 f2 apply clarsimp proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))" assume a2: "t2 = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" have f3: "\is. get_operand2 is s1 = get_operand2 is s2" using a1 by (metis get_operand2_low_equal) have f4: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using a1 by (meson get_curr_win_low_equal) have "\s. snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s + get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))) = t2 \ \ low_equal s (snd (fst (get_curr_win () s2)))" using a2 user_reg_val_low_equal by fastforce then show "low_equal (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f4 f3 a2 a1 by (metis (no_types) add_instr_sub1_low_equal get_curr_win2_low_equal write_reg_low_equal) qed next case False then show ?thesis using f1 f2 apply clarsimp proof - assume a1: "low_equal s1 s2" have f2: "\s sa sb w wa wb sc. (\ low_equal s sa \ sb \ snd (fst (write_reg w (wa::'a word) wb s)) \ sc \ snd (fst (write_reg w wa wb sa))) \ low_equal sb sc" by (meson write_reg_low_equal) have f3: "gets (\s. ucast (get_CWP (cpu_reg_val PSR s))::'a word) = get_curr_win ()" by (simp add: get_curr_win_def) then have "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = (fst (get_curr_win () s1), snd (get_curr_win () s1))" by (metis (no_types) prod.collapse simpler_gets_def) then have "(ucast (get_CWP (cpu_reg_val PSR s1)), s1) = fst (get_curr_win () s1) \ \ snd (get_curr_win () s1)" by blast then have f4: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s1)) \ s1 = snd (fst (get_curr_win () s1))" by (metis (no_types) prod.collapse prod.simps(1)) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = (fst (get_curr_win () s2), snd (get_curr_win () s2))" using f3 by (metis (no_types) prod.collapse simpler_gets_def) then have "(ucast (get_CWP (cpu_reg_val PSR s2)), s2) = fst (get_curr_win () s2) \ \ snd (get_curr_win () s2)" by blast then have f5: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis prod.collapse prod.simps(1)) then have f6: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2))) = low_equal s1 s2" using f4 by presburger have f7: "fst (fst (get_curr_win () s1)) = ucast (get_CWP (cpu_reg_val PSR s1))" using f4 by presburger have f8: "cpu_reg_val PSR s1 = cpu_reg_val PSR s2" using a1 by (meson cpu_reg_val_low_equal) have f9: "user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) = user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))" using f6 a1 by (meson user_reg_val_low_equal) have f10: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2))" using f5 by meson have f11: "\s sa is. \ low_equal (s::'a sparc_state) sa \ get_operand2 is s = get_operand2 is sa" using get_operand2_low_equal by blast then have f12: "user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1))))) = user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))" using f9 f8 f5 f4 a1 by auto then have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using f10 f8 f6 f4 f2 a1 by simp then show "low_equal (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f12 f11 f10 f9 f8 f7 a1 add_instr_sub1_low_equal by fastforce qed qed next case False then have f3: "\ get_operand_w5 (snd instr ! 3) \ 0" by auto then show ?thesis proof (cases "fst instr = arith_type ADD \ fst instr = arith_type ADDcc") case True then show ?thesis using f1 f3 apply clarsimp proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))" assume "t2 = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))" have f2: "\is. get_operand2 is s1 = get_operand2 is s2" using a1 by (meson get_operand2_low_equal) have f3: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using a1 by (meson get_curr_win_low_equal) have "\w wa. user_reg_val wa w (snd (fst (get_curr_win () s1))) = user_reg_val wa w (snd (fst (get_curr_win () s2)))" using a1 by (meson get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using f3 f2 a1 by (metis (no_types) add_instr_sub1_low_equal get_curr_win2_low_equal write_reg_low_equal) qed next case False then show ?thesis using f1 f3 apply clarsimp proof - assume a1: "low_equal s1 s2" have f2: "gets (\s. ucast (get_CWP (cpu_reg_val PSR s))::'a word) = get_curr_win ()" by (simp add: get_curr_win_def) then have "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = (fst (get_curr_win () s1), snd (get_curr_win () s1))" by (metis (no_types) prod.collapse simpler_gets_def) then have "(ucast (get_CWP (cpu_reg_val PSR s1)), s1) = fst (get_curr_win () s1) \ \ snd (get_curr_win () s1)" by fastforce then have f3: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s1)) \ s1 = snd (fst (get_curr_win () s1))" by (metis prod.collapse prod.simps(1)) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = (fst (get_curr_win () s2), snd (get_curr_win () s2))" using f2 by (metis (no_types) prod.collapse simpler_gets_def) then have "(ucast (get_CWP (cpu_reg_val PSR s2)), s2) = fst (get_curr_win () s2) \ \ snd (get_curr_win () s2)" by fastforce then have f4: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis (no_types) prod.collapse prod.simps(1)) then have f5: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2))) = low_equal s1 s2" using f3 by presburger have f6: "fst (fst (get_curr_win () s1)) = ucast (get_CWP (cpu_reg_val PSR s1))" using f3 by auto have f7: "cpu_reg_val PSR s1 = cpu_reg_val PSR s2" using a1 by (meson cpu_reg_val_low_equal) have f8: "\s sa w wa. \ low_equal s sa \ user_reg_val (w::'a word) wa s = user_reg_val w wa sa" by (meson user_reg_val_low_equal) have f9: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2))" using f4 by meson have "\s sa is. \ low_equal (s::'a sparc_state) sa \ get_operand2 is s = get_operand2 is sa" using get_operand2_low_equal by blast then have f10: "get_operand2 (snd instr) s1 = get_operand2 (snd instr) s2" using a1 by meson have f11: "cpu_reg_val PSR (snd (fst (get_curr_win () s2))) = cpu_reg_val PSR s1" using f4 a1 by (simp add: cpu_reg_val_low_equal) have f12: "user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))) = 0" by (meson user_reg_val_def) have "user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))) = 0" by (meson user_reg_val_def) then have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))" using f12 f9 f7 f5 f3 a1 write_reg_low_equal by fastforce then have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))) \ snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))))) = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (if get_operand_w5 (snd instr ! Suc 0) = 0 then 0 else user_reg (snd (fst (get_curr_win () s2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))))) \ snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))) = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (if get_operand_w5 (snd instr ! Suc 0) = 0 then 0 else user_reg (snd (fst (get_curr_win () s2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))" using f11 f10 f9 f8 f7 f6 f5 f3 a1 by (simp add: user_reg_val_def) then show "low_equal (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using add_instr_sub1_low_equal by blast qed qed qed qed lemma sub_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (sub_instr_sub1 instr_name result rs1_val operand2 s1)) \ t2 = snd (fst (sub_instr_sub1 instr_name result rs1_val operand2 s2))" shows "low_equal t1 t2" proof (cases "instr_name = arith_type SUBcc \ instr_name = arith_type SUBXcc") case True then show ?thesis using a1 apply (simp add: sub_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (clarsimp simp add: cpu_reg_val_low_equal) using write_cpu_low_equal by blast next case False then show ?thesis using a1 apply (simp add: sub_instr_sub1_def) by (simp add: return_def) qed lemma sub_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (sub_instr instr s1)) \ t2 = snd (fst (sub_instr instr s2))" shows "low_equal t1 t2" proof - have f1: "low_equal s1 s2 \ t1 = snd (fst (sub_instr_sub1 (fst instr) (if fst instr = arith_type SUB \ fst instr = arith_type SUBcc then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then if fst instr = arith_type SUB \ fst instr = arith_type SUBcc then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1))))) else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))))) \ t2 = snd (fst (sub_instr_sub1 (fst instr) (if fst instr = arith_type SUB \ fst instr = arith_type SUBcc then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then if fst instr = arith_type SUB \ fst instr = arith_type SUBcc then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))) else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" using a1 apply (simp add: sub_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: case_prod_unfold) then show ?thesis proof (cases "get_operand_w5 (snd instr ! 3) \ 0") case True then have f2: "get_operand_w5 (snd instr ! 3) \ 0" by auto then show ?thesis proof (cases "fst instr = arith_type SUB \ fst instr = arith_type SUBcc") case True then show ?thesis using f1 f2 apply clarsimp proof - assume a1: "low_equal s1 s2" assume a2: "t1 = snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))" assume a3: "t2 = snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" then have f4: "snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))) = t2" using a1 by (simp add: get_operand2_low_equal) have "\s. \ low_equal (snd (fst (get_curr_win () s1))) s \ snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s - get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))))) = t1" using a2 a1 by (simp add: get_curr_win_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f4 a3 a2 a1 by (metis (no_types) get_curr_win2_low_equal sub_instr_sub1_low_equal write_reg_low_equal) qed next case False then show ?thesis using f1 f2 apply clarsimp proof - assume a1: "low_equal s1 s2" have f2: "fst (get_curr_win () s1) = (ucast (get_CWP (cpu_reg_val PSR s1)), s1)" by (simp add: get_curr_win_def simpler_gets_def) have f3: "cpu_reg_val PSR s1 = cpu_reg_val PSR s2" using a1 by (meson cpu_reg_val_low_equal) then have f4: "user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) = user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) s1" using f2 by simp have f5: "\s sa is. \ low_equal (s::'a sparc_state) sa \ get_operand2 is s = get_operand2 is sa" using get_operand2_low_equal by blast then have f6: "sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) s2) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) = sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))" using f4 a1 by (simp add: user_reg_val_low_equal) have f7: "fst (get_curr_win () s2) = (ucast (get_CWP (cpu_reg_val PSR s2)), s2)" by (simp add: get_curr_win_def simpler_gets_def) then have f8: "user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))) = user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))" using f5 f2 a1 by (simp add: cpu_reg_val_low_equal user_reg_val_low_equal) then have f9: "sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) s2) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))) = sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using f7 by fastforce have "write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! 3)) s2 = write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))" using f8 f7 by simp then have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using f3 f2 a1 by (metis (no_types) prod.sel(1) prod.sel(2) write_reg_low_equal) then show "low_equal (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f9 f6 by (metis (no_types) sub_instr_sub1_low_equal) qed qed next case False then have f3: "\ get_operand_w5 (snd instr ! 3) \ 0" by auto then show ?thesis proof (cases "fst instr = arith_type SUB \ fst instr = arith_type SUBcc") case True then show ?thesis using f1 f3 apply clarsimp proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))" assume "t2 = snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))" have f2: "\is. get_operand2 is s1 = get_operand2 is s2" using a1 get_operand2_low_equal by blast have f3: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using a1 by (meson get_curr_win_low_equal) have "\w wa. user_reg_val wa w (snd (fst (get_curr_win () s1))) = user_reg_val wa w (snd (fst (get_curr_win () s2)))" using a1 by (metis (no_types) get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using f3 f2 a1 by (metis (no_types) get_curr_win2_low_equal sub_instr_sub1_low_equal write_reg_low_equal) qed next case False then show ?thesis using f1 f3 apply clarsimp proof - assume a1: "low_equal s1 s2" have f2: "\s sa sb w wa wb sc. (\ low_equal s sa \ sb \ snd (fst (write_reg w (wa::'a word) wb s)) \ sc \ snd (fst (write_reg w wa wb sa))) \ low_equal sb sc" by (meson write_reg_low_equal) have "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = get_curr_win () s1" by (simp add: get_curr_win_def simpler_gets_def) then have f3: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s1)) \ s1 = snd (fst (get_curr_win () s1))" by (metis (no_types) prod.collapse prod.simps(1)) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = get_curr_win () s2" by (simp add: get_curr_win_def simpler_gets_def) then have f4: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis (no_types) prod.collapse prod.simps(1)) have f5: "\s sa sb sc w wa wb sd. (\ low_equal (s::'a sparc_state) sa \ sb \ snd (fst (sub_instr_sub1 sc w wa wb s)) \ sd \ snd (fst (sub_instr_sub1 sc w wa wb sa))) \ low_equal sb sd" by (meson sub_instr_sub1_low_equal) have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))" using f4 f3 f2 a1 by (simp add: cpu_reg_val_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using f5 f4 f3 a1 by (simp add: cpu_reg_val_low_equal get_operand2_low_equal user_reg_val_low_equal) qed qed qed qed lemma mul_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (mul_instr_sub1 instr_name result s1)) \ t2 = snd (fst (mul_instr_sub1 instr_name result s2))" shows "low_equal t1 t2" proof (cases "instr_name \ {arith_type SMULcc,arith_type UMULcc}") case True then show ?thesis using a1 apply (simp add: mul_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (clarsimp simp add: cpu_reg_val_low_equal) using write_cpu_low_equal by blast next case False then show ?thesis using a1 apply (simp add: mul_instr_sub1_def) by (simp add: return_def) qed lemma mul_instr_low_equal: \low_equal t1 t2\ if \low_equal s1 s2 \ t1 = snd (fst (mul_instr instr s1)) \ t2 = snd (fst (mul_instr instr s2))\ proof - from that have \low_equal s1 s2\ and t1: \t1 = snd (fst (mul_instr instr s1))\ and t2: \t2 = snd (fst (mul_instr instr s2))\ by simp_all have f2: "\s sa sb sc w sd. \ low_equal (s::'a sparc_state) sa \ sb \ snd (fst (mul_instr_sub1 sc w s)) \ sd \ snd (fst (mul_instr_sub1 sc w sa)) \ low_equal sb sd" using mul_instr_sub1_low_equal by blast have f3: "\s sa sb w wa wb sc. \ low_equal s sa \ sb \ snd (fst (write_reg w (wa::'a word) wb s)) \ sc \ snd (fst (write_reg w wa wb sa)) \ low_equal sb sc" by (meson write_reg_low_equal) have f4: "\s sa sb w c sc. \ low_equal (s::'a sparc_state) sa \ sb \ snd (fst (write_cpu w c s)) \ sc \ snd (fst (write_cpu w c sa)) \ low_equal sb sc" by (meson write_cpu_low_equal) have f6: "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = (fst (get_curr_win () s1), snd (get_curr_win () s1))" by (simp add: get_curr_win_def simpler_gets_def) have f7: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using \low_equal s1 s2\ by (meson get_curr_win_low_equal) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = (fst (get_curr_win () s2), snd (get_curr_win () s2))" by (simp add: get_curr_win_def simpler_gets_def) then have f8: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis prod.collapse prod.simps(1)) then have f9: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using f6 \low_equal s1 s2\ by (metis (no_types) prod.collapse prod.simps(1)) have f10: "\s sa w wa. \ low_equal s sa \ user_reg_val (w::'a word) wa s = user_reg_val w wa sa" using user_reg_val_low_equal by blast have f11: "get_operand2 (snd instr) s1 = get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))" using f9 f6 by (metis (no_types) get_operand2_low_equal prod.collapse prod.simps(1)) then have f12: "uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2) = uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)" using f10 f9 f8 f7 by presburger then have f13: "(word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) \ low_equal (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))" using f9 f4 by presburger have "get_operand_w5 (snd instr ! 3) = 0 \ low_equal (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) = write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))" using f10 f7 by force then have f14: "get_operand_w5 (snd instr ! 3) \ 0 \ low_equal (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))) \ \ low_equal (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))" using f3 by metis then have f15: "low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ fst instr \ arith_type UMULcc \ get_operand_w5 (snd instr ! 3) \ 0 \ (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))" using f13 f12 f2 by fastforce have f16: "user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) = user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))" using f10 f9 f7 by presburger { assume "fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "\ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) \ mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))" then have "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ fst instr \ arith_type UMULcc" by fastforce } moreover { assume "mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))) \ mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))" then have "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ fst instr \ arith_type UMULcc" by fastforce } ultimately have "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ fst instr \ arith_type UMULcc" by force } ultimately have "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ fst instr \ arith_type UMULcc" by fastforce } ultimately have "fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))) = write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) = write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))" by blast } moreover { assume "\ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "\ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))))) = snd (fst (mul_instr_sub1 (arith_type UMULcc) (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))))) \ snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))) = snd (fst (mul_instr_sub1 (arith_type UMULcc) (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))" then have "\ low_equal (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))" using f2 by blast moreover { assume "(if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) \ (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))))" moreover { assume "(if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) \ ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))" then have "(if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" by (metis f11 f16 f8) } ultimately have "(if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" by fastforce } ultimately have "fst instr = arith_type UMULcc \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" using f13 f7 f3 by fastforce } moreover { assume "mul_instr_sub1 (arith_type UMULcc) (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))) \ mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))" moreover { assume "(if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) \ ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))" then have "(if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" by (metis f11 f16 f8) } ultimately have "fst instr = arith_type UMULcc \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" by fastforce } ultimately have "fst instr = arith_type UMULcc \ get_operand_w5 (snd instr ! 3) = 0" using f12 by fastforce } moreover { assume "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))" then have "fst instr = arith_type UMULcc \ get_operand_w5 (snd instr ! 3) = 0" by presburger } ultimately have "fst instr = arith_type UMULcc \ get_operand_w5 (snd instr ! 3) = 0 \ get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))" by force moreover { assume "fst instr \ arith_type UMULcc" { assume "fst instr \ arith_type UMULcc \ low_equal (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))" moreover { assume "fst instr \ arith_type UMULcc \ low_equal (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))) \ snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))) = snd (fst (mul_instr_sub1 (arith_type UMUL) (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))" then have "(fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ fst instr \ arith_type UMUL \ fst instr \ arith_type UMULcc \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ get_operand_w5 (snd instr ! 3) = 0" using f2 by presburger } ultimately have "fst instr \ arith_type UMULcc \ (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) \ ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ fst instr \ arith_type UMUL \ fst instr \ arith_type UMULcc \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ get_operand_w5 (snd instr ! 3) = 0" by fastforce } then have "(get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ get_operand_w5 (snd instr ! 3) = 0" using f16 f11 f9 f8 f7 f4 f3 f2 by force } moreover { assume "get_operand_w5 (snd instr ! 3) = 0" moreover { assume "(fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0" moreover { assume "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" then have "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ arith_type UMUL \ arith_type UMULcc \ fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0" by fastforce then have "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0" by force } ultimately have "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ ((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0" by simp } ultimately have "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0 \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" by auto then have "fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0 \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" using f15 by presburger then have "(get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" using f14 f13 f12 f2 by force } ultimately have "(get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" using f16 f14 f11 f9 f8 f4 f2 by fastforce } ultimately have "(get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr = arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMUL \ fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr = arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMUL \ fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" by blast moreover from t1 have \t1 = snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))\ by (simp add: mul_instr_def) (simp add: simpler_gets_def bind_def h1_def h2_def Let_def case_prod_unfold) moreover from t2 have \t2 = snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))\ by (simp add: mul_instr_def) (simp add: simpler_gets_def bind_def h1_def h2_def Let_def case_prod_unfold) ultimately show ?thesis by (simp add: mul_instr_def) qed lemma div_write_new_val_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (div_write_new_val i result temp_V s1)) \ t2 = snd (fst (div_write_new_val i result temp_V s2))" shows "low_equal t1 t2" proof (cases "(fst i) \ {arith_type UDIVcc,arith_type SDIVcc}") case True then show ?thesis using a1 apply (simp add: div_write_new_val_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (clarsimp simp add: cpu_reg_val_low_equal) using write_cpu_low_equal by blast next case False then show ?thesis using a1 apply (simp add: div_write_new_val_def) by (simp add: return_def) qed lemma div_comp_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (div_comp instr rs1 rd operand2 s1)) \ t2 = snd (fst (div_comp instr rs1 rd operand2 s2))" shows "low_equal t1 t2" using a1 apply (simp add: div_comp_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (clarsimp simp add: get_curr_win_low_equal) proof - assume a1: "low_equal s1 s2" have f2: "\s sa sb w wa wb sc. \ low_equal s sa \ sb \ snd (fst (write_reg w (wa::'a word) wb s)) \ sc \ snd (fst (write_reg w wa wb sa)) \ low_equal sb sc" by (meson write_reg_low_equal) have f3: "gets (\s. ucast (get_CWP (cpu_reg_val PSR s))::'a word) = get_curr_win ()" by (simp add: get_curr_win_def) then have "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = (fst (get_curr_win () s1), snd (get_curr_win () s1))" by (metis (no_types) prod.collapse simpler_gets_def) then have f4: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s1)) \ s1 = snd (fst (get_curr_win () s1))" by (metis prod.collapse prod.simps(1)) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = (fst (get_curr_win () s2), snd (get_curr_win () s2))" using f3 by (metis (no_types) prod.collapse simpler_gets_def) then have f5: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis (no_types) prod.collapse prod.simps(1)) then have f6: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using f4 a1 by presburger have f7: "\s sa sb p w wa sc. \ low_equal (s::'a sparc_state) sa \ sb \ snd (fst (div_write_new_val p w wa s)) \ sc \ snd (fst (div_write_new_val p w wa sa)) \ low_equal sb sc" by (meson div_write_new_val_low_equal) have f8: "cpu_reg_val PSR s2 = cpu_reg_val PSR s1" using a1 by (simp add: cpu_reg_val_def low_equal_def) then have "fst (fst (get_curr_win () s2)) = ucast (get_CWP (cpu_reg_val PSR s1))" using f5 by presburger then have f9: "fst (fst (get_curr_win () s2)) = fst (fst (get_curr_win () s1))" using f4 by presburger have f10: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using f8 f5 f4 by presburger have f11: "(word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))::word64) = word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))" using f5 f4 a1 by (metis (no_types) cpu_reg_val_def low_equal_def user_reg_val_low_equal) have f12: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s2))" using f8 f5 by presburger then have "rd = 0 \ (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) = user_reg_val (ucast (get_CWP (cpu_reg_val PSR s1))) 0 (snd (fst (get_curr_win () s1)))" using f6 user_reg_val_low_equal by fastforce then have f13: "rd = 0 \ write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) 0 (snd (fst (get_curr_win () s1))) = write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))" using f12 f10 by presburger have f14: "write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s2))) = write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))" using f12 f11 by auto have "write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s1))) = write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))) \ write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s2))) = write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) \ low_equal (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))))))" using f6 f2 by metis moreover { assume "low_equal (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))))))" then have "low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))" using f11 f9 f7 by metis moreover { assume "low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))))))))) \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))" then have "div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2) = (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) \ rd = 0" by fastforce } ultimately have "div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2) = (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) \ rd = 0 \ (rd \ 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))) \ (rd = 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))))" by fastforce } moreover { assume "write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s1))) \ write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))" then have "div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2) \ (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2))" using f12 f9 by fastforce } moreover { assume "write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s2))) \ write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))" then have "rd = 0" using f14 by presburger } moreover { assume "rd = 0" then have "rd = 0 \ low_equal (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))))))" using f13 f12 f6 f2 by metis then have "rd = 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))" using f11 f9 f7 by metis then have "(rd \ 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))) \ (rd = 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))))" using f10 by fastforce } ultimately show "(rd \ 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))) \ (rd = 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))))" using f9 by fastforce qed lemma div_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (div_instr instr s1)) \ t2 = snd (fst (div_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: div_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: return_def) apply (auto simp add: get_operand2_low_equal) apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (auto simp add: traps_low_equal) apply (blast intro: mod_trap_low_equal) using div_comp_low_equal by blast lemma get_curr_win_traps_low_equal: assumes a1: "low_equal s1 s2" shows "low_equal (snd (fst (get_curr_win () s1)) \traps := insert some_trap (traps (snd (fst (get_curr_win () s1))))\) (snd (fst (get_curr_win () s2)) \traps := insert some_trap (traps (snd (fst (get_curr_win () s2))))\)" proof - from a1 have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by auto then have f2: "(traps (snd (fst (get_curr_win () s1)))) = (traps (snd (fst (get_curr_win () s2))))" using traps_low_equal by auto then show ?thesis using f1 f2 mod_trap_low_equal by fastforce qed lemma save_restore_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (save_retore_sub1 result new_cwp rd s1)) \ t2 = snd (fst (save_retore_sub1 result new_cwp rd s2))" shows "low_equal t1 t2" using a1 apply (simp add: save_retore_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (auto simp add: cpu_reg_val_low_equal) using write_cpu_low_equal write_reg_low_equal by fastforce lemma get_WIM_bit_low_equal: \get_WIM_bit (nat ((uint (fst (fst (get_curr_win () s1))) - 1) mod NWINDOWS)) (cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = get_WIM_bit (nat ((uint (fst (fst (get_curr_win () s2))) - 1) mod NWINDOWS)) (cpu_reg_val WIM (snd (fst (get_curr_win () s2))))\ if \low_equal s1 s2\ proof - from that have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then have f2: "(cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = (cpu_reg_val WIM (snd (fst (get_curr_win () s2))))" using cpu_reg_val_low_equal by auto from that have "(fst (fst (get_curr_win () s1))) = (fst (fst (get_curr_win () s2)))" using get_curr_win_low_equal by auto then show ?thesis using f1 f2 by auto qed lemma get_WIM_bit_low_equal2: \get_WIM_bit (nat ((uint (fst (fst (get_curr_win () s1))) + 1) mod NWINDOWS)) (cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = get_WIM_bit (nat ((uint (fst (fst (get_curr_win () s2))) + 1) mod NWINDOWS)) (cpu_reg_val WIM (snd (fst (get_curr_win () s2))))\ if \low_equal s1 s2\ proof - from that have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then have f2: "(cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = (cpu_reg_val WIM (snd (fst (get_curr_win () s2))))" using cpu_reg_val_low_equal by auto from that have "(fst (fst (get_curr_win () s1))) = (fst (fst (get_curr_win () s2)))" using get_curr_win_low_equal by auto then show ?thesis using f1 f2 by auto qed lemma take_bit_5_mod_NWINDOWS_eq [simp]: \take_bit 5 (k mod NWINDOWS) = k mod NWINDOWS\ by (simp add: NWINDOWS_def take_bit_eq_mod) lemma save_restore_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (save_restore_instr instr s1)) \ t2 = snd (fst (save_restore_instr instr s2))" shows "low_equal t1 t2" proof (cases "fst instr = ctrl_type SAVE") case True then have f1: "fst instr = ctrl_type SAVE" by auto then show ?thesis using a1 - apply (simp add: save_restore_instr_def) - apply (simp add: Let_def) - apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) - apply (simp add: case_prod_unfold) - apply auto - apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) - apply (simp add: get_curr_win_traps_low_equal) + apply (simp add: save_restore_instr_def) + apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) - apply (simp add: get_WIM_bit_low_equal) - apply (simp add: get_WIM_bit_low_equal) - apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) - apply (simp add: get_curr_win_low_equal) - using get_curr_win2_low_equal save_restore_instr_sub1_low_equal get_addr2_low_equal - apply metis - done + apply (simp add: case_prod_unfold) + apply (auto simp add: unsigned_of_int) + apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) + apply (simp add: get_curr_win_traps_low_equal) + apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) + apply (simp add: get_WIM_bit_low_equal) + apply (simp add: get_WIM_bit_low_equal) + apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) + apply (simp add: get_curr_win_low_equal) + using get_curr_win2_low_equal save_restore_instr_sub1_low_equal get_addr2_low_equal + apply metis + done next case False then show ?thesis using a1 - apply (simp add: save_restore_instr_def) - apply (simp add: Let_def) - apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) - apply (simp add: case_prod_unfold) - apply auto - apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) - apply (simp add: get_curr_win_traps_low_equal) + apply (simp add: save_restore_instr_def) + apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) - apply (simp add: get_WIM_bit_low_equal2) - apply (simp add: get_WIM_bit_low_equal2) - apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) - apply (simp add: get_curr_win_low_equal) - using get_curr_win2_low_equal save_restore_instr_sub1_low_equal get_addr2_low_equal - by metis + apply (simp add: case_prod_unfold) + apply (auto simp add: unsigned_of_int) + apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) + apply (simp add: get_curr_win_traps_low_equal) + apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) + apply (simp add: get_WIM_bit_low_equal2) + apply (simp add: get_WIM_bit_low_equal2) + apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) + apply (simp add: get_curr_win_low_equal) + using get_curr_win2_low_equal save_restore_instr_sub1_low_equal get_addr2_low_equal + apply metis + done qed lemma call_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (call_instr instr s1)) \ t2 = snd (fst (call_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: call_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (auto simp add: get_curr_win_low_equal) using cpu_reg_val_low_equal get_curr_win2_low_equal write_cpu_low_equal write_reg_low_equal proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (write_cpu (cpu_reg_val PC (snd (fst (get_curr_win () s1))) + (ucast (get_operand_w30 (snd instr ! 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (get_curr_win () s1)))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 15 (snd (fst (get_curr_win () s1)))))))))))" assume "t2 = snd (fst (write_cpu (cpu_reg_val PC (snd (fst (get_curr_win () s2))) + (ucast (get_operand_w30 (snd instr ! 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (get_curr_win () s2)))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 15 (snd (fst (get_curr_win () s2)))))))))))" have "\c. cpu_reg_val c (snd (fst (get_curr_win () s1))) = cpu_reg_val c (snd (fst (get_curr_win () s2)))" using a1 by (meson cpu_reg_val_low_equal get_curr_win2_low_equal) then show "low_equal (snd (fst (write_cpu (cpu_reg_val PC (snd (fst (get_curr_win () s1))) + (ucast (get_operand_w30 (snd instr ! 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (get_curr_win () s1)))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 15 (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (write_cpu (cpu_reg_val PC (snd (fst (get_curr_win () s2))) + (ucast (get_operand_w30 (snd instr ! 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (get_curr_win () s2)))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 15 (snd (fst (get_curr_win () s2))))))))))))" using a1 by (metis (no_types) get_curr_win2_low_equal write_cpu_low_equal write_reg_low_equal) qed lemma jmpl_instr_low_equal_sub1: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (write_cpu (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))))) \ t2 = snd (fst (write_cpu (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))))" shows "low_equal t1 t2" proof - from a1 have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then have f2: "(cpu_reg_val PC (snd (fst (get_curr_win () s1)))) = (cpu_reg_val PC (snd (fst (get_curr_win () s2))))" using cpu_reg_val_low_equal by blast then have f3: "low_equal (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using f1 write_reg_low_equal by fastforce then have "(cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))) = (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))" using cpu_reg_val_low_equal by auto then have f4: "low_equal (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f3 write_cpu_low_equal by fastforce then show ?thesis using write_cpu_low_equal using assms by blast qed lemma jmpl_instr_low_equal_sub2: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (write_cpu (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))) PC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))))) \ t2 = snd (fst (write_cpu (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))) PC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))))" shows "low_equal t1 t2" proof - from a1 have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then have f2: "(user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) = (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))" using user_reg_val_low_equal by blast then have f3: "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))" using f1 write_reg_low_equal by fastforce then have "(cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))) = (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))" using cpu_reg_val_low_equal by blast then have "low_equal (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))) PC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))) PC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using f1 f2 f3 write_cpu_low_equal by fastforce then show ?thesis using write_cpu_low_equal using assms by blast qed lemma jmpl_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (jmpl_instr instr s1)) \ t2 = snd (fst (jmpl_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: jmpl_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (simp add: get_curr_win_traps_low_equal) apply (simp add: get_addr2_low_equal) apply (simp add: get_addr2_low_equal) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp_all add: get_addr2_low_equal) apply (simp_all add: get_curr_win_low_equal) apply (case_tac "get_operand_w5 (snd instr ! 3) \ 0") apply auto using jmpl_instr_low_equal_sub1 apply blast apply (simp_all add: get_curr_win_low_equal) using jmpl_instr_low_equal_sub2 by blast lemma rett_instr_low_equal: assumes a1: "low_equal s1 s2 \ \ snd (rett_instr instr s1) \ \ snd (rett_instr instr s2) \ (((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (rett_instr instr s1)) \ t2 = snd (fst (rett_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: rett_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply auto apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (simp add: return_def) using mod_trap_low_equal traps_low_equal apply fastforce using cpu_reg_val_low_equal apply fastforce using cpu_reg_val_low_equal apply fastforce apply (simp add: bind_def h1_def h2_def Let_def) by (simp add: case_prod_unfold fail_def) lemma read_state_reg_low_equal: assumes a1: "low_equal s1 s2 \ (((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (read_state_reg_instr instr s1)) \ t2 = snd (fst (read_state_reg_instr instr s2))" shows "low_equal t1 t2" proof (cases "(fst instr \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 ((snd instr)!0))))") case True then have "(fst instr \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 ((snd instr)!0)))) \ (((get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s1))))))::word1) = 0 \ (((get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))))::word1) = 0" by (metis assms get_curr_win_privilege) then show ?thesis using a1 apply (simp add: read_state_reg_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply clarsimp using get_curr_win_traps_low_equal by auto next case False then have f1: "\((fst instr = sreg_type RDPSR \ fst instr = sreg_type RDWIM \ fst instr = sreg_type RDTBR \ fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 (snd instr ! 0))))" by blast then show ?thesis proof (cases "illegal_instruction_ASR (get_operand_w5 ((snd instr)!0))") case True then show ?thesis using a1 f1 apply read_state_reg_instr_privilege_proof by (simp add: illegal_instruction_ASR_def) next case False then have f2: "\(illegal_instruction_ASR (get_operand_w5 ((snd instr)!0)))" by auto then show ?thesis proof (cases "(get_operand_w5 ((snd instr)!1)) \ 0") case True then have f3: "(get_operand_w5 ((snd instr)!1)) \ 0" by auto then show ?thesis proof (cases "fst instr = sreg_type RDY") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: read_state_reg_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (auto simp add: get_curr_win_low_equal) using cpu_reg_val_low_equal get_curr_win2_low_equal write_reg_low_equal proof - assume "low_equal s1 s2" then have "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" by (meson get_curr_win2_low_equal) then show "low_equal (snd (fst (write_reg (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))))))" using cpu_reg_val_low_equal write_reg_low_equal by fastforce qed next case False then have f4: "\(fst instr = sreg_type RDY)" by auto then show ?thesis proof (cases "fst instr = sreg_type RDASR") case True then show ?thesis using a1 f1 f2 f3 f4 apply read_state_reg_instr_privilege_proof apply (clarsimp simp add: get_curr_win_low_equal) using cpu_reg_val_low_equal get_curr_win2_low_equal write_reg_low_equal proof - assume a1: "low_equal s1 s2" then have "cpu_reg_val (ASR (get_operand_w5 (snd instr ! 0))) (snd (fst (get_curr_win () s1))) = cpu_reg_val (ASR (get_operand_w5 (snd instr ! 0))) (snd (fst (get_curr_win () s2)))" by (meson cpu_reg_val_low_equal get_curr_win2_low_equal) then show "low_equal (snd (fst (write_reg (cpu_reg_val (ASR (get_operand_w5 (snd instr ! 0))) (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (cpu_reg_val (ASR (get_operand_w5 (snd instr ! 0))) (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))))))" using a1 by (metis (no_types) get_curr_win2_low_equal write_reg_low_equal) qed next case False then have f5: "\(fst instr = sreg_type RDASR)" by auto then show ?thesis using a1 f1 f2 f3 f4 f5 apply read_state_reg_instr_privilege_proof apply (clarsimp simp add: get_curr_win_low_equal) using cpu_reg_val_low_equal get_curr_win2_low_equal write_reg_low_equal proof - assume a1: "low_equal s1 s2" assume a2: "t1 = snd (fst (write_reg (cpu_reg_val TBR (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (cpu_reg_val TBR (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))))" have "\s. \ low_equal (snd (fst (get_curr_win () s1))) s \ snd (fst (write_reg (cpu_reg_val TBR s) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))))) = t1" using a2 by (simp add: cpu_reg_val_low_equal) then show "low_equal (snd (fst (write_reg (cpu_reg_val TBR (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (cpu_reg_val TBR (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))))))" using a2 a1 by (metis (no_types) get_curr_win2_low_equal write_reg_low_equal) qed qed qed next case False then show ?thesis using a1 f1 f2 apply (simp add: read_state_reg_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply clarsimp apply (simp add: case_prod_unfold) using get_curr_win2_low_equal by auto qed qed qed lemma get_s_get_curr_win: assumes a1: "low_equal s1 s2" shows "get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))) = get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))" proof - from a1 have "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then show ?thesis using cpu_reg_val_low_equal by fastforce qed lemma write_state_reg_low_equal: assumes a1: "low_equal s1 s2 \ (((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (write_state_reg_instr instr s1)) \ t2 = snd (fst (write_state_reg_instr instr s2))" shows "low_equal t1 t2" proof (cases "fst instr = sreg_type WRY") case True then show ?thesis using a1 apply write_state_reg_instr_privilege_proof apply (simp add: simpler_modify_def) apply (simp add: delayed_pool_add_def DELAYNUM_def) apply (auto simp add: get_curr_win_low_equal) using get_curr_win2_low_equal cpu_reg_mod_low_equal user_reg_val_low_equal get_operand2_low_equal proof - assume a1: "low_equal s1 s2" assume "t2 = cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))) Y (snd (fst (get_curr_win () s2)))" assume "t1 = cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1)))) Y (snd (fst (get_curr_win () s1)))" have f2: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using a1 by (meson get_curr_win2_low_equal) then have f3: "\w wa. user_reg_val w wa (snd (fst (get_curr_win () s2))) = user_reg_val w wa (snd (fst (get_curr_win () s1)))" by (simp add: user_reg_val_low_equal) have "\is. get_operand2 is (snd (fst (get_curr_win () s2))) = get_operand2 is (snd (fst (get_curr_win () s1)))" using f2 by (simp add: get_operand2_low_equal) then show "low_equal (cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1)))) Y (snd (fst (get_curr_win () s1)))) (cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))) Y (snd (fst (get_curr_win () s2))))" using f3 f2 by (metis cpu_reg_mod_low_equal) qed next case False then have f1: "\(fst instr = sreg_type WRY)" by auto then show ?thesis proof (cases "fst instr = sreg_type WRASR") case True then have f1_1: "fst instr = sreg_type WRASR" by auto then show ?thesis proof (cases "privileged_ASR (get_operand_w5 (snd instr ! 3)) \ get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))) = 0") case True then show ?thesis using a1 f1 f1_1 apply write_state_reg_instr_privilege_proof apply (clarsimp simp add: get_s_get_curr_win) apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce next case False then have f1_2: "\ (privileged_ASR (get_operand_w5 (snd instr ! 3)) \ get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))) = 0)" by auto then show ?thesis proof (cases "illegal_instruction_ASR (get_operand_w5 (snd instr ! 3))") case True then show ?thesis using a1 f1 f1_1 f1_2 apply write_state_reg_instr_privilege_proof apply (clarsimp simp add: get_s_get_curr_win) apply auto apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal apply fastforce apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce next case False then show ?thesis using a1 f1 f1_1 f1_2 apply write_state_reg_instr_privilege_proof apply (clarsimp simp add: get_s_get_curr_win) apply auto apply (simp add: simpler_modify_def) apply (simp add: delayed_pool_add_def DELAYNUM_def) apply (auto simp add: get_curr_win_low_equal) using get_curr_win2_low_equal cpu_reg_mod_low_equal user_reg_val_low_equal get_operand2_low_equal proof - assume a1: "low_equal s1 s2" assume "t2 = cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))) (ASR (get_operand_w5 (snd instr ! 3))) (snd (fst (get_curr_win () s2)))" assume "t1 = cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1)))) (ASR (get_operand_w5 (snd instr ! 3))) (snd (fst (get_curr_win () s1)))" have "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using a1 by (meson get_curr_win2_low_equal) then show "low_equal (cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1)))) (ASR (get_operand_w5 (snd instr ! 3))) (snd (fst (get_curr_win () s1)))) (cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))) (ASR (get_operand_w5 (snd instr ! 3))) (snd (fst (get_curr_win () s2))))" using cpu_reg_mod_low_equal get_operand2_low_equal user_reg_val_low_equal by fastforce next assume f1: "\ illegal_instruction_ASR (get_operand_w5 (snd instr ! 3))" assume f2: "fst instr = sreg_type WRASR" assume f3: "snd (fst (write_state_reg_instr instr s1)) = snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s1))))) " assume f4: "snd (fst (write_state_reg_instr instr s2)) = snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s2)))))" assume f5: "low_equal s1 s2" assume f6: "(get_S (cpu_reg_val PSR s1)) = 0" assume f7: "(get_S (cpu_reg_val PSR s2)) = 0" assume f8: "t1 = snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s1)))))" assume f9: "t2 = snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s2)))))" assume f10: "get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))) \ 0" assume f11: "(\s1 s2 t1 t2. low_equal s1 s2 \ t1 = snd (fst (get_curr_win () s1)) \ t2 = snd (fst (get_curr_win () s2)) \ low_equal t1 t2)" assume f12: "(\s1 s2 t1 w cr t2. low_equal s1 s2 \ t1 = cpu_reg_mod w cr s1 \ t2 = cpu_reg_mod w cr s2 \ low_equal t1 t2)" assume f13: "(\s1 s2 win ur. low_equal s1 s2 \ user_reg_val win ur s1 = user_reg_val win ur s2)" assume f14: "(\s1 s2 op_list. low_equal s1 s2 \ get_operand2 op_list s1 = get_operand2 op_list s2)" show "low_equal (snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s1)))))) (snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s2))))))" using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 using Sparc_Properties.ucast_0 assms get_curr_win_privilege by blast qed qed qed next case False then have f2: "\(fst instr = sreg_type WRASR)" by auto have f3: "get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))) = 0 \ get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))) = 0" using get_curr_win_privilege a1 by (metis ucast_id) then show ?thesis proof (cases "fst instr = sreg_type WRPSR") case True then show ?thesis using a1 f1 f2 f3 apply write_state_reg_instr_privilege_proof apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce next case False then have f4: "\(fst instr = sreg_type WRPSR)" by auto then show ?thesis proof (cases "fst instr = sreg_type WRWIM") case True then show ?thesis using a1 f1 f2 f3 f4 apply write_state_reg_instr_privilege_proof apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce next case False then have f5: "\(fst instr = sreg_type WRWIM)" by auto then show ?thesis using a1 f1 f2 f3 f4 f5 apply write_state_reg_instr_privilege_proof apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce qed qed qed qed lemma flush_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (flush_instr instr s1)) \ t2 = snd (fst (flush_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: flush_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def simpler_modify_def) apply (simp add: flush_cache_all_def) apply (simp add: low_equal_def) apply (simp add: user_accessible_def) apply (simp add: mem_equal_def) by auto lemma branch_instr_sub1_low_equal: assumes a1: "low_equal s1 s2" shows "branch_instr_sub1 instr_name s1 = branch_instr_sub1 instr_name s2" using a1 apply (simp add: branch_instr_sub1_def) by (simp add: low_equal_def) lemma set_annul_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (set_annul True s1)) \ t2 = snd (fst (set_annul True s2))" shows "low_equal t1 t2" using a1 apply (simp add: set_annul_def) apply (simp add: simpler_modify_def annul_mod_def) using state_var2_low_equal state_var_low_equal by fastforce lemma branch_instr_low_equal_sub0: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1))))) \ t2 = snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2)))))" shows "low_equal t1 t2" proof - from a1 have "low_equal (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1))) (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2)))" using write_cpu_low_equal by blast then show ?thesis using a1 write_cpu_low_equal by blast qed lemma branch_instr_low_equal_sub1: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (set_annul True (snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1)))))))) \ t2 = snd (fst (set_annul True (snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2))))))))" shows "low_equal t1 t2" proof - from a1 have "low_equal (snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1)))))) (snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2))))))" using branch_instr_low_equal_sub0 by blast then show ?thesis using a1 using set_annul_low_equal by blast qed lemma branch_instr_low_equal_sub2: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (set_annul True (snd (fst (write_cpu (cpu_reg_val nPC s2 + 4) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1)))))))) \ t2 = snd (fst (set_annul True (snd (fst (write_cpu (cpu_reg_val nPC s2 + 4) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2))))))))" shows "low_equal t1 t2" proof - from a1 have "low_equal (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1))) (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2)))" using write_cpu_low_equal by blast then have "low_equal (snd (fst (write_cpu (cpu_reg_val nPC s2 + 4) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1)))))) (snd (fst (write_cpu (cpu_reg_val nPC s2 + 4) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2))))))" using write_cpu_low_equal by blast then show ?thesis using a1 using set_annul_low_equal by blast qed lemma branch_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (branch_instr instr s1)) \ t2 = snd (fst (branch_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: branch_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold return_def) apply clarsimp apply (simp add: branch_instr_sub1_low_equal) apply (simp_all add: cpu_reg_val_low_equal) apply (cases "branch_instr_sub1 (fst instr) s2 = 1") apply clarsimp apply (simp add: bind_def h1_def h2_def Let_def) apply (simp_all add: cpu_reg_val_low_equal) apply (simp add: case_prod_unfold) apply (cases "fst instr = bicc_type BA \ get_operand_flag (snd instr ! 0) = 1") apply clarsimp using branch_instr_low_equal_sub1 apply blast apply clarsimp apply (simp add: return_def) using branch_instr_low_equal_sub0 apply fastforce apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (cases "get_operand_flag (snd instr ! 0) = 1") apply clarsimp apply (simp_all add: cpu_reg_val_low_equal) using branch_instr_low_equal_sub2 apply metis apply (simp add: return_def) using write_cpu_low_equal by metis lemma dispath_instr_low_equal: assumes a1: "low_equal s1 s2 \ (((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ \ snd (dispatch_instruction instr s1) \ \ snd (dispatch_instruction instr s2) \ t1 = (snd (fst (dispatch_instruction instr s1))) \ t2 = (snd (fst (dispatch_instruction instr s2)))" shows "low_equal t1 t2" proof (cases "get_trap_set s1 = {}") case True then have f_no_traps: "get_trap_set s1 = {} \ get_trap_set s2 = {}" using a1 by (simp add: low_equal_def get_trap_set_def) then show ?thesis proof (cases "fst instr \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD}") case True then show ?thesis using a1 f_no_traps apply dispath_instr_privilege_proof by (blast intro: load_instr_low_equal) next case False then have f1: "fst instr \ {load_store_type LDSB, load_store_type LDUB, load_store_type LDUBA, load_store_type LDUH, load_store_type LD, load_store_type LDA, load_store_type LDD}" by auto then show ?thesis proof (cases "fst instr \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD}") case True then show ?thesis using a1 f_no_traps f1 apply dispath_instr_privilege_proof using store_instr_low_equal by blast next case False then have f2: "\(fst instr \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD})" by auto then show ?thesis proof (cases "fst instr \ {sethi_type SETHI}") case True then show ?thesis using a1 f_no_traps f1 f2 apply dispath_instr_privilege_proof by (auto intro: sethi_low_equal) next case False then have f3: "\(fst instr \ {sethi_type SETHI})" by auto then show ?thesis proof (cases "fst instr \ {nop_type NOP}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 apply dispath_instr_privilege_proof by (auto intro: nop_low_equal) next case False then have f4: "\(fst instr \ {nop_type NOP})" by auto then show ?thesis proof (cases "fst instr \ {logic_type ANDs,logic_type ANDcc,logic_type ANDN, logic_type ANDNcc,logic_type ORs,logic_type ORcc,logic_type ORN, logic_type XORs,logic_type XNOR}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 apply dispath_instr_privilege_proof using logical_instr_low_equal by blast next case False then have f5: "\(fst instr \ {logic_type ANDs,logic_type ANDcc,logic_type ANDN, logic_type ANDNcc,logic_type ORs,logic_type ORcc,logic_type ORN, logic_type XORs,logic_type XNOR})" by auto then show ?thesis proof (cases "fst instr \ {shift_type SLL,shift_type SRL,shift_type SRA}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 apply dispath_instr_privilege_proof using shift_instr_low_equal by blast next case False then have f6: "\(fst instr \ {shift_type SLL,shift_type SRL,shift_type SRA})" by auto then show ?thesis proof (cases "fst instr \ {arith_type ADD,arith_type ADDcc,arith_type ADDX}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 apply dispath_instr_privilege_proof using add_instr_low_equal by blast next case False then have f7: "\(fst instr \ {arith_type ADD,arith_type ADDcc,arith_type ADDX})" by auto then show ?thesis proof (cases "fst instr \ {arith_type SUB,arith_type SUBcc,arith_type SUBX}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 apply dispath_instr_privilege_proof using sub_instr_low_equal by blast next case False then have f8: "\(fst instr \ {arith_type SUB,arith_type SUBcc,arith_type SUBX})" by auto then show ?thesis proof (cases "fst instr \ {arith_type UMUL,arith_type SMUL,arith_type SMULcc}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 apply dispath_instr_privilege_proof using mul_instr_low_equal by blast next case False then have f9: "\(fst instr \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc})" by auto then show ?thesis proof (cases "fst instr \ {arith_type UDIV,arith_type UDIVcc,arith_type SDIV}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 apply dispath_instr_privilege_proof using div_instr_low_equal by blast next case False then have f10: "\(fst instr \ {arith_type UDIV, arith_type UDIVcc,arith_type SDIV})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type SAVE,ctrl_type RESTORE}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 apply dispath_instr_privilege_proof using save_restore_instr_low_equal by blast next case False then have f11: "\(fst instr \ {ctrl_type SAVE,ctrl_type RESTORE})" by auto then show ?thesis proof (cases "fst instr \ {call_type CALL}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 apply dispath_instr_privilege_proof using call_instr_low_equal by blast next case False then have f12: "\(fst instr \ {call_type CALL})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type JMPL}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 apply dispath_instr_privilege_proof using jmpl_instr_low_equal by blast next case False then have f13: "\(fst instr \ {ctrl_type JMPL})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type RETT}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 apply dispath_instr_privilege_proof using rett_instr_low_equal by blast next case False then have f14: "\(fst instr \ {ctrl_type RETT})" by auto then show ?thesis proof (cases "fst instr \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM, sreg_type RDTBR}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 apply dispath_instr_privilege_proof using read_state_reg_low_equal by blast next case False then have f15: "\(fst instr \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM, sreg_type RDTBR})" by auto then show ?thesis proof (cases "fst instr \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM, sreg_type WRTBR}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 apply dispath_instr_privilege_proof using write_state_reg_low_equal by blast next case False then have f16: "\(fst instr \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM, sreg_type WRTBR})" by auto then show ?thesis proof (cases "fst instr \ {load_store_type FLUSH}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 apply dispath_instr_privilege_proof using flush_instr_low_equal by blast next case False then have f17: "\(fst instr \ {load_store_type FLUSH})" by auto then show ?thesis proof (cases "fst instr \ {bicc_type BE,bicc_type BNE, bicc_type BGU,bicc_type BLE,bicc_type BL,bicc_type BGE, bicc_type BNEG,bicc_type BG,bicc_type BCS,bicc_type BLEU, bicc_type BCC,bicc_type BA,bicc_type BN}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 apply dispath_instr_privilege_proof using branch_instr_low_equal by blast next case False then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 apply dispath_instr_privilege_proof by (simp add: fail_def) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed next case False then have "get_trap_set s1 \ {} \ get_trap_set s2 \ {}" using a1 by (simp add: low_equal_def get_trap_set_def) then show ?thesis using a1 apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) by (simp add: return_def) qed lemma execute_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ \ snd (execute_instr_sub1 instr s1) \ \ snd (execute_instr_sub1 instr s2) \ t1 = (snd (fst (execute_instr_sub1 instr s1))) \ t2 = (snd (fst (execute_instr_sub1 instr s2)))" shows "low_equal t1 t2" proof (cases "get_trap_set s1 = {}") case True then have "get_trap_set s1 = {} \ get_trap_set s2 = {}" using a1 by (simp add: low_equal_def get_trap_set_def) then show ?thesis using a1 apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (case_tac "fst instr \ call_type CALL \ fst instr \ ctrl_type RETT \ fst instr \ ctrl_type JMPL \ fst instr \ bicc_type BE \ fst instr \ bicc_type BNE \ fst instr \ bicc_type BGU \ fst instr \ bicc_type BLE \ fst instr \ bicc_type BL \ fst instr \ bicc_type BGE \ fst instr \ bicc_type BNEG \ fst instr \ bicc_type BG \ fst instr \ bicc_type BCS \ fst instr \ bicc_type BLEU \ fst instr \ bicc_type BCC \ fst instr \ bicc_type BA \ fst instr \ bicc_type BN") apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: low_equal_def) apply (simp add: cpu_reg_val_def write_cpu_def cpu_reg_mod_def) apply (simp add: simpler_modify_def return_def) apply (simp add: user_accessible_mod_cpu_reg mem_equal_mod_cpu_reg) apply clarsimp by (auto simp add: return_def) next case False then have "get_trap_set s1 \ {} \ get_trap_set s2 \ {}" using a1 by (simp add: low_equal_def get_trap_set_def) then show ?thesis using a1 apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: return_def) qed theorem non_interference_step: assumes a1: "(((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ good_context s1 \ get_delayed_pool s1 = [] \ get_trap_set s1 = {} \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ get_delayed_pool s2 = [] \ get_trap_set s2 = {} \ good_context s2 \ low_equal s1 s2" shows "\t1 t2. Some t1 = NEXT s1 \ Some t2 = NEXT s2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2" proof - from a1 have "good_context s1 \ good_context s2" by auto then have "NEXT s1 = Some (snd (fst (execute_instruction () s1))) \ NEXT s2 = Some (snd (fst (execute_instruction () s2)))" by (simp add: single_step) then have "\t1 t2. Some t1 = NEXT s1 \ Some t2 = NEXT s2" by auto then have f0: "snd (execute_instruction() s1) = False \ snd (execute_instruction() s2) = False" by (auto simp add: NEXT_def case_prod_unfold) then have f1: "\t1 t2. Some t1 = NEXT s1 \ Some t2 = NEXT s2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0" using a1 apply (auto simp add: NEXT_def case_prod_unfold) by (auto simp add: safe_privilege) then show ?thesis proof (cases "exe_mode_val s1") case True then have f_exe0: "exe_mode_val s1" by auto then have f_exe: "exe_mode_val s1 \ exe_mode_val s2" proof - have "low_equal s1 s2" using a1 by auto then have "state_var s1 = state_var s2" by (simp add: low_equal_def) then have "exe_mode_val s1 = exe_mode_val s2" by (simp add: exe_mode_val_def) then show ?thesis using f_exe0 by auto qed then show ?thesis proof (cases "\e. fetch_instruction (delayed_pool_write s1) = Inl e") case True then have f_fetch_error: "\e. fetch_instruction (delayed_pool_write s1) = Inl e" by auto then have f_fetch_error2: "(\e. fetch_instruction (delayed_pool_write s1) = Inl e) \ (\e. fetch_instruction (delayed_pool_write s2) = Inl e)" proof - have "cpu_reg s1 = cpu_reg s2" using a1 by (simp add: low_equal_def) then have "cpu_reg_val PC s1 = cpu_reg_val PC s2" by (simp add: cpu_reg_val_def) then have "cpu_reg_val PC s1 = cpu_reg_val PC s2 \ (((get_S (cpu_reg_val PSR (delayed_pool_write s1))))::word1) = 0 \ (((get_S (cpu_reg_val PSR (delayed_pool_write s2))))::word1) = 0" using a1 by (auto simp add: empty_delayed_pool_write_privilege) then show ?thesis using a1 f_fetch_error apply (simp add: fetch_instruction_def) apply (simp add: Let_def) apply clarsimp apply (case_tac "uint (3 AND cpu_reg_val PC (delayed_pool_write s1)) = 0") apply auto apply (case_tac "fst (memory_read 8 (cpu_reg_val PC (delayed_pool_write s1)) (delayed_pool_write s1)) = None") apply auto apply (simp add: case_prod_unfold) using a1 apply (auto simp add: mem_read_delayed_write_low_equal) apply (simp add: case_prod_unfold) using a1 apply (auto simp add: mem_read_delayed_write_low_equal) apply (simp add: delayed_pool_write_def) by (simp add: Let_def get_delayed_write_def) qed then show ?thesis proof (cases "exe_mode_val s1") case True then have "exe_mode_val s1 \ exe_mode_val s2" using exe_mode_low_equal a1 by auto then show ?thesis using f1 apply (simp add: NEXT_def execute_instruction_def) apply (simp add: bind_def h1_def h2_def Let_def simpler_gets_def) using a1 apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) using f_fetch_error2 apply clarsimp apply (simp add: raise_trap_def simpler_modify_def return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: return_def simpler_modify_def) apply (simp add: raise_trap_def simpler_modify_def return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply (simp add: delayed_pool_write_def get_delayed_write_def Let_def) apply (simp add: low_equal_def) apply (simp add: add_trap_set_def) apply (simp add: cpu_reg_val_def) apply clarsimp by (simp add: mem_equal_mod_trap user_accessible_mod_trap) next case False then have "\ (exe_mode_val s1) \ \ (exe_mode_val s2)" using exe_mode_low_equal a1 by auto then show ?thesis using f1 apply (simp add: NEXT_def execute_instruction_def) apply (simp add: bind_def h1_def h2_def Let_def simpler_gets_def) using a1 apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: return_def) qed next case False then have f_fetch_suc: "(\v. fetch_instruction (delayed_pool_write s1) = Inr v)" using fetch_instr_result_1 by auto then have "(\v. fetch_instruction (delayed_pool_write s1) = Inr v \ fetch_instruction (delayed_pool_write s2) = Inr v)" proof - have "cpu_reg s1 = cpu_reg s2" using a1 by (simp add: low_equal_def) then have "cpu_reg_val PC s1 = cpu_reg_val PC s2" by (simp add: cpu_reg_val_def) then have "cpu_reg_val PC s1 = cpu_reg_val PC s2 \ (((get_S (cpu_reg_val PSR (delayed_pool_write s1))))::word1) = 0 \ (((get_S (cpu_reg_val PSR (delayed_pool_write s2))))::word1) = 0" using a1 by (auto simp add: empty_delayed_pool_write_privilege) then show ?thesis using a1 f_fetch_suc apply (simp add: fetch_instruction_def) apply (simp add: Let_def) apply clarsimp apply (case_tac "uint (3 AND cpu_reg_val PC (delayed_pool_write s1)) = 0") apply auto apply (case_tac "fst (memory_read 8 (cpu_reg_val PC (delayed_pool_write s1)) (delayed_pool_write s1)) = None") apply auto apply (simp add: case_prod_unfold) using a1 apply (auto simp add: mem_read_delayed_write_low_equal) apply (simp add: case_prod_unfold) using a1 apply (auto simp add: mem_read_delayed_write_low_equal) apply (simp add: delayed_pool_write_def) by (simp add: Let_def get_delayed_write_def) qed then have "(\v. fetch_instruction (delayed_pool_write s1) = Inr v \ fetch_instruction (delayed_pool_write s2) = Inr v \ \ (\e. (decode_instruction v) = Inl e))" using dispatch_fail f0 a1 f_exe by auto then have f_fetch_dec: "(\v. fetch_instruction (delayed_pool_write s1) = Inr v \ fetch_instruction (delayed_pool_write s2) = Inr v \ (\v1. (decode_instruction v) = Inr v1))" using decode_instr_result_4 by auto then show ?thesis proof (cases "annul_val (delayed_pool_write s1)") case True then have "annul_val (delayed_pool_write s1) \ annul_val (delayed_pool_write s2)" using a1 apply (simp add: low_equal_def) by (simp add: delayed_pool_write_def get_delayed_write_def annul_val_def) then show ?thesis using a1 f1 f_exe f_fetch_dec apply (simp add: NEXT_def execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: write_cpu_def cpu_reg_val_def set_annul_def) apply (simp add: simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: write_cpu_def cpu_reg_val_def set_annul_def) apply (simp add: simpler_modify_def) apply (simp add: cpu_reg_mod_def annul_mod_def) apply (simp add: delayed_pool_write_def get_delayed_write_def) apply (simp add: write_annul_def) apply clarsimp apply (simp add: low_equal_def) apply (simp add: user_accessible_annul mem_equal_annul) by (metis) next case False then have "\ annul_val (delayed_pool_write s1) \ \ annul_val (delayed_pool_write s2)" using a1 apply (simp add: low_equal_def) apply (simp add: delayed_pool_write_def get_delayed_write_def) by (simp add: annul_val_def) then show ?thesis using a1 f1 f_exe f_fetch_dec apply (simp add: NEXT_def execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (case_tac "snd (execute_instr_sub1 (a, b) (snd (fst (dispatch_instruction (a, b) (delayed_pool_write s1))))) \ snd (dispatch_instruction (a, b) (delayed_pool_write s1))") apply auto apply (case_tac "snd (execute_instr_sub1 (a, b) (snd (fst (dispatch_instruction (a, b) (delayed_pool_write s2))))) \ snd (dispatch_instruction (a, b) (delayed_pool_write s2))") apply auto apply (simp add: simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: delayed_pool_write_def get_delayed_write_def) by (meson dispath_instr_low_equal dispath_instr_privilege execute_instr_sub1_low_equal) qed qed next case False then have f_non_exe: "exe_mode_val s1 = False" by auto then have "exe_mode_val s1 = False \ exe_mode_val s2 = False" proof - have "low_equal s1 s2" using a1 by auto then have "state_var s1 = state_var s2" by (simp add: low_equal_def) then have "exe_mode_val s1 = exe_mode_val s2" by (simp add: exe_mode_val_def) then show ?thesis using f_non_exe by auto qed then show ?thesis using f1 a1 apply (simp add: NEXT_def execute_instruction_def) by (simp add: simpler_gets_def bind_def h1_def h2_def Let_def return_def) qed qed function (sequential) SEQ:: "nat \ ('a::len) sparc_state \ ('a) sparc_state option" where "SEQ 0 s = Some s" |"SEQ n s = ( case SEQ (n-1) s of None \ None | Some t \ NEXT t )" by pat_completeness auto termination by lexicographic_order lemma SEQ_suc: "SEQ n s = Some t \ SEQ (Suc n) s = NEXT t" apply (induction n) apply clarsimp by (simp add: option.case_eq_if) definition user_seq_exe:: "nat \ ('a::len) sparc_state \ bool" where "user_seq_exe n s \ \i t. (i \ n \ SEQ i s = Some t) \ (good_context t \ get_delayed_pool t = [] \ get_trap_set t = {})" text \NIA is short for non-interference assumption.\ definition "NIA t1 t2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ good_context t1 \ get_delayed_pool t1 = [] \ get_trap_set t1 = {} \ good_context t2 \ get_delayed_pool t2 = [] \ get_trap_set t2 = {} \ low_equal t1 t2" text \NIC is short for non-interference conclusion.\ definition "NIC t1 t2 \ (\u1 u2. Some u1 = NEXT t1 \ Some u2 = NEXT t2 \ (((get_S (cpu_reg_val PSR u1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR u2)))::word1) = 0 \ low_equal u1 u2)" lemma NIS_short: "\t1 t2. NIA t1 t2 \ NIC t1 t2" apply (simp add: NIA_def NIC_def) using non_interference_step by auto lemma non_interference_induct_case_sub1: assumes a1: "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ NIA t1 t2))" shows "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ NIA t1 t2 \ NIC t1 t2))" using NIS_short using assms by auto lemma non_interference_induct_case: assumes a1: "((\i t. i \ n \ SEQ i s1 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}) \ (\i t. i \ n \ SEQ i s2 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}) \ (\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2))) \ (\i t. i \ Suc n \ SEQ i s1 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}) \ (\i t. i \ Suc n \ SEQ i s2 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {})" shows "\t1. Some t1 = (case SEQ n s1 of None \ None | Some x \ NEXT x) \ (\t2. Some t2 = (case SEQ n s2 of None \ None | Some x \ NEXT x) \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2)" proof - from a1 have f1: "((\i t. i \ n \ SEQ i s1 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}) \ (\i t. i \ n \ SEQ i s2 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}))" by (metis le_SucI) then have f2: "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2))" using a1 by auto then have f3: "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ NIA t1 t2))" using f1 NIA_def by (metis (full_types) dual_order.refl) then have "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ NIA t1 t2 \ NIC t1 t2))" using non_interference_induct_case_sub1 by blast then have "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ ((((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ good_context t1 \ get_delayed_pool t1 = [] \ get_trap_set t1 = {} \ good_context t2 \ get_delayed_pool t2 = [] \ get_trap_set t2 = {} \ low_equal t1 t2) \ (\u1 u2. Some u1 = NEXT t1 \ Some u2 = NEXT t2 \ (((get_S (cpu_reg_val PSR u1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR u2)))::word1) = 0 \ low_equal u1 u2)))" using NIA_def NIC_def by fastforce then show ?thesis by (metis option.simps(5)) qed lemma non_interference_induct_case_sub2: assumes a1: "(user_seq_exe n s1 \ user_seq_exe n s2 \ (\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2))) \ user_seq_exe (Suc n) s1 \ user_seq_exe (Suc n) s2" shows "\t1. Some t1 = (case SEQ n s1 of None \ None | Some x \ NEXT x) \ (\t2. Some t2 = (case SEQ n s2 of None \ None | Some x \ NEXT x) \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2)" using a1 by (simp add: non_interference_induct_case user_seq_exe_def) theorem non_interference: assumes a1: "(((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ good_context s1 \ get_delayed_pool s1 = [] \ get_trap_set s1 = {} \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ get_delayed_pool s2 = [] \ get_trap_set s2 = {} \ good_context s2 \ user_seq_exe n s1 \ user_seq_exe n s2 \ low_equal s1 s2" shows "(\t1 t2. Some t1 = SEQ n s1 \ Some t2 = SEQ n s2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2)" using a1 apply (induction n) apply (simp add: user_seq_exe_def) apply clarsimp by (simp add: non_interference_induct_case_sub2) end end diff --git a/thys/Word_Lib/Aligned.thy b/thys/Word_Lib/Aligned.thy --- a/thys/Word_Lib/Aligned.thy +++ b/thys/Word_Lib/Aligned.thy @@ -1,1331 +1,1331 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Word Alignment" theory Aligned imports "HOL-Library.Word" More_Word Word_EqI begin context includes bit_operations_syntax begin lift_definition is_aligned :: \'a::len word \ nat \ bool\ is \\k n. 2 ^ n dvd take_bit LENGTH('a) k\ by simp lemma is_aligned_iff_udvd: \is_aligned w n \ 2 ^ n udvd w\ by transfer (simp flip: take_bit_eq_0_iff add: min_def) lemma is_aligned_iff_take_bit_eq_0: \is_aligned w n \ take_bit n w = 0\ by (simp add: is_aligned_iff_udvd take_bit_eq_0_iff exp_dvd_iff_exp_udvd) lemma is_aligned_iff_dvd_int: \is_aligned ptr n \ 2 ^ n dvd uint ptr\ by transfer simp lemma is_aligned_iff_dvd_nat: \is_aligned ptr n \ 2 ^ n dvd unat ptr\ proof - have \unat ptr = nat \uint ptr\\ by transfer simp then have \2 ^ n dvd unat ptr \ 2 ^ n dvd uint ptr\ by (simp only: dvd_nat_abs_iff) simp then show ?thesis by (simp add: is_aligned_iff_dvd_int) qed lemma is_aligned_0 [simp]: \is_aligned 0 n\ by transfer simp lemma is_aligned_at_0 [simp]: \is_aligned w 0\ by transfer simp lemma is_aligned_beyond_length: \is_aligned w n \ w = 0\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that apply (simp add: is_aligned_iff_udvd) apply transfer apply auto done lemma is_alignedI [intro?]: \is_aligned x n\ if \x = 2 ^ n * k\ for x :: \'a::len word\ proof (unfold is_aligned_iff_udvd) from that show \2 ^ n udvd x\ using dvd_triv_left exp_dvd_iff_exp_udvd by blast qed lemma is_alignedE: fixes w :: \'a::len word\ assumes \is_aligned w n\ obtains q where \w = 2 ^ n * word_of_nat q\ \q < 2 ^ (LENGTH('a) - n)\ proof (cases \n < LENGTH('a)\) case False with assms have \w = 0\ by (simp add: is_aligned_beyond_length) with that [of 0] show thesis by simp next case True moreover define m where \m = LENGTH('a) - n\ ultimately have l: \LENGTH('a) = n + m\ and \m \ 0\ by simp_all from \n < LENGTH('a)\ have *: \unat (2 ^ n :: 'a word) = 2 ^ n\ by transfer simp from assms have \2 ^ n udvd w\ by (simp add: is_aligned_iff_udvd) then obtain v :: \'a word\ where \unat w = unat (2 ^ n :: 'a word) * unat v\ .. moreover define q where \q = unat v\ ultimately have unat_w: \unat w = 2 ^ n * q\ by (simp add: *) then have \word_of_nat (unat w) = (word_of_nat (2 ^ n * q) :: 'a word)\ by simp then have w: \w = 2 ^ n * word_of_nat q\ by simp moreover have \q < 2 ^ (LENGTH('a) - n)\ proof (rule ccontr) assume \\ q < 2 ^ (LENGTH('a) - n)\ then have \2 ^ (LENGTH('a) - n) \ q\ by simp then have \2 ^ LENGTH('a) \ 2 ^ n * q\ by (simp add: l power_add) with unat_w [symmetric] show False by (metis le_antisym nat_less_le unsigned_less) qed ultimately show thesis using that by blast qed lemma is_alignedE' [elim?]: fixes w :: \'a::len word\ assumes \is_aligned w n\ obtains q where \w = push_bit n (word_of_nat q)\ \q < 2 ^ (LENGTH('a) - n)\ proof - from assms obtain q where \w = 2 ^ n * word_of_nat q\ \q < 2 ^ (LENGTH('a) - n)\ by (rule is_alignedE) then have \w = push_bit n (word_of_nat q)\ by (simp add: push_bit_eq_mult) with that show thesis using \q < 2 ^ (LENGTH('a) - n)\ . qed lemma is_aligned_mask: \is_aligned w n \ w AND mask n = 0\ by (simp add: is_aligned_iff_take_bit_eq_0 take_bit_eq_mask) lemma is_aligned_imp_not_bit: \\ bit w m\ if \is_aligned w n\ and \m < n\ for w :: \'a::len word\ proof - from \is_aligned w n\ obtain q where \w = push_bit n (word_of_nat q)\ \q < 2 ^ (LENGTH('a) - n)\ .. moreover have \\ bit (push_bit n (word_of_nat q :: 'a word)) m\ using \m < n\ by (simp add: bit_simps) ultimately show ?thesis by simp qed lemma is_aligned_weaken: "\ is_aligned w x; x \ y \ \ is_aligned w y" unfolding is_aligned_iff_dvd_nat by (erule dvd_trans [rotated]) (simp add: le_imp_power_dvd) lemma is_alignedE_pre: fixes w::"'a::len word" assumes aligned: "is_aligned w n" shows rl: "\q. w = 2 ^ n * (of_nat q) \ q < 2 ^ (LENGTH('a) - n)" using aligned is_alignedE by blast lemma aligned_add_aligned: fixes x::"'a::len word" assumes aligned1: "is_aligned x n" and aligned2: "is_aligned y m" and lt: "m \ n" shows "is_aligned (x + y) m" proof cases assume nlt: "n < LENGTH('a)" show ?thesis unfolding is_aligned_iff_dvd_nat dvd_def proof - from aligned2 obtain q2 where yv: "y = 2 ^ m * of_nat q2" and q2v: "q2 < 2 ^ (LENGTH('a) - m)" by (auto elim: is_alignedE) from lt obtain k where kv: "m + k = n" by (auto simp: le_iff_add) with aligned1 obtain q1 where xv: "x = 2 ^ (m + k) * of_nat q1" and q1v: "q1 < 2 ^ (LENGTH('a) - (m + k))" by (auto elim: is_alignedE) have l1: "2 ^ (m + k) * q1 < 2 ^ LENGTH('a)" by (rule nat_less_power_trans [OF q1v]) (subst kv, rule order_less_imp_le [OF nlt]) have l2: "2 ^ m * q2 < 2 ^ LENGTH('a)" by (rule nat_less_power_trans [OF q2v], rule order_less_imp_le [OF order_le_less_trans]) fact+ have "x = of_nat (2 ^ (m + k) * q1)" using xv by simp moreover have "y = of_nat (2 ^ m * q2)" using yv by simp ultimately have upls: "unat x + unat y = 2 ^ m * (2 ^ k * q1 + q2)" proof - have f1: "unat x = 2 ^ (m + k) * q1" using q1v unat_mult_power_lem xv by blast have "unat y = 2 ^ m * q2" using q2v unat_mult_power_lem yv by blast then show ?thesis using f1 by (simp add: power_add semiring_normalization_rules(34)) qed (* (2 ^ k * q1 + q2) *) show "\d. unat (x + y) = 2 ^ m * d" proof (cases "unat x + unat y < 2 ^ LENGTH('a)") case True have "unat (x + y) = unat x + unat y" by (subst unat_plus_if', rule if_P) fact also have "\ = 2 ^ m * (2 ^ k * q1 + q2)" by (rule upls) finally show ?thesis .. next case False then have "unat (x + y) = (unat x + unat y) mod 2 ^ LENGTH('a)" by (subst unat_word_ariths(1)) simp also have "\ = (2 ^ m * (2 ^ k * q1 + q2)) mod 2 ^ LENGTH('a)" by (subst upls, rule refl) also have "\ = 2 ^ m * ((2 ^ k * q1 + q2) mod 2 ^ (LENGTH('a) - m))" proof - have "m \ len_of (TYPE('a))" by (meson le_trans less_imp_le_nat lt nlt) then show ?thesis by (metis mult_mod_right ordered_cancel_comm_monoid_diff_class.add_diff_inverse power_add) qed finally show ?thesis .. qed qed next assume "\ n < LENGTH('a)" with assms show ?thesis by (simp add: is_aligned_mask not_less take_bit_eq_mod power_overflow word_arith_nat_defs(7) flip: take_bit_eq_mask) qed corollary aligned_sub_aligned: "\is_aligned (x::'a::len word) n; is_aligned y m; m \ n\ \ is_aligned (x - y) m" apply (simp del: add_uminus_conv_diff add:diff_conv_add_uminus) apply (erule aligned_add_aligned, simp_all) apply (erule is_alignedE) apply (rule_tac k="- of_nat q" in is_alignedI) apply simp done lemma is_aligned_shift: fixes k::"'a::len word" shows "is_aligned (push_bit m k) m" proof cases assume mv: "m < LENGTH('a)" from mv obtain q where mq: "m + q = LENGTH('a)" and "0 < q" by (auto dest: less_imp_add_positive) have "(2::nat) ^ m dvd unat (push_bit m k)" proof have kv: "(unat k div 2 ^ q) * 2 ^ q + unat k mod 2 ^ q = unat k" by (rule div_mult_mod_eq) have "unat (push_bit m k) = unat (2 ^ m * k)" by (simp add: push_bit_eq_mult ac_simps) also have "\ = (2 ^ m * unat k) mod (2 ^ LENGTH('a))" using mv by (simp add: unat_word_ariths(2)) also have "\ = 2 ^ m * (unat k mod 2 ^ q)" by (subst mq [symmetric], subst power_add, subst mod_mult2_eq) simp finally show "unat (push_bit m k) = 2 ^ m * (unat k mod 2 ^ q)" . qed then show ?thesis by (unfold is_aligned_iff_dvd_nat) next assume "\ m < LENGTH('a)" then show ?thesis by (simp add: not_less power_overflow is_aligned_mask word_size) qed lemma word_mod_by_0: "k mod (0::'a::len word) = k" by (simp add: word_arith_nat_mod) lemma aligned_mod_eq_0: fixes p::"'a::len word" assumes al: "is_aligned p sz" shows "p mod 2 ^ sz = 0" proof cases assume szv: "sz < LENGTH('a)" with al show ?thesis unfolding is_aligned_iff_dvd_nat by (simp add: and_mask_dvd_nat p2_gt_0 word_mod_2p_is_mask) next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: is_aligned_mask flip: take_bit_eq_mask take_bit_eq_mod) qed lemma is_aligned_triv: "is_aligned (2 ^ n ::'a::len word) n" by (rule is_alignedI [where k = 1], simp) lemma is_aligned_mult_triv1: "is_aligned (2 ^ n * x ::'a::len word) n" by (rule is_alignedI [OF refl]) lemma is_aligned_mult_triv2: "is_aligned (x * 2 ^ n ::'a::len word) n" by (subst mult.commute, simp add: is_aligned_mult_triv1) lemma word_power_less_0_is_0: fixes x :: "'a::len word" shows "x < a ^ 0 \ x = 0" by simp lemma is_aligned_no_wrap: fixes off :: "'a::len word" fixes ptr :: "'a::len word" assumes al: "is_aligned ptr sz" and off: "off < 2 ^ sz" shows "unat ptr + unat off < 2 ^ LENGTH('a)" proof - have szv: "sz < LENGTH('a)" using off p2_gt_0 word_neq_0_conv by fastforce from al obtain q where ptrq: "ptr = 2 ^ sz * of_nat q" and qv: "q < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis proof (cases "sz = 0") case True then show ?thesis using off ptrq qv by simp next case False then have sne: "0 < sz" .. show ?thesis proof - have uq: "unat (of_nat q ::'a::len word) = q" apply (subst unat_of_nat) apply (rule mod_less) apply (rule order_less_trans [OF qv]) apply (rule power_strict_increasing [OF diff_less [OF sne]]) apply (simp_all) done have uptr: "unat ptr = 2 ^ sz * q" apply (subst ptrq) apply (subst iffD1 [OF unat_mult_lem]) apply (subst unat_power_lower [OF szv]) apply (subst uq) apply (rule nat_less_power_trans [OF qv order_less_imp_le [OF szv]]) apply (subst uq) apply (subst unat_power_lower [OF szv]) apply simp done show "unat ptr + unat off < 2 ^ LENGTH('a)" using szv apply (subst uptr) apply (subst mult.commute, rule nat_add_offset_less [OF _ qv]) apply (rule order_less_le_trans [OF unat_mono [OF off] order_eq_refl]) apply simp_all done qed qed qed lemma is_aligned_no_wrap': fixes ptr :: "'a::len word" assumes al: "is_aligned ptr sz" and off: "off < 2 ^ sz" shows "ptr \ ptr + off" by (subst no_plus_overflow_unat_size, subst word_size, rule is_aligned_no_wrap) fact+ lemma is_aligned_no_overflow': fixes p :: "'a::len word" assumes al: "is_aligned p n" shows "p \ p + (2 ^ n - 1)" proof cases assume "n n ptr \ ptr + 2^sz - 1" by (drule is_aligned_no_overflow') (simp add: field_simps) lemma replicate_not_True: "\n. xs = replicate n False \ True \ set xs" by (induct xs) auto lemma map_zip_replicate_False_xor: "n = length xs \ map (\(x, y). x = (\ y)) (zip xs (replicate n False)) = xs" by (induct xs arbitrary: n, auto) lemma drop_minus_lem: "\ n \ length xs; 0 < n; n' = length xs \ \ drop (n' - n) xs = rev xs ! (n - 1) # drop (Suc (n' - n)) xs" proof (induct xs arbitrary: n n') case Nil then show ?case by simp next case (Cons y ys) from Cons.prems show ?case apply simp apply (cases "n = Suc (length ys)") apply (simp add: nth_append) apply (simp add: Suc_diff_le Cons.hyps nth_append) apply clarsimp apply arith done qed lemma drop_minus: "\ n < length xs; n' = length xs \ \ drop (n' - Suc n) xs = rev xs ! n # drop (n' - n) xs" apply (subst drop_minus_lem) apply simp apply simp apply simp apply simp apply (cases "length xs", simp) apply (simp add: Suc_diff_le) done lemma aligned_add_xor: \(x + 2 ^ n) XOR 2 ^ n = x\ if al: \is_aligned (x::'a::len word) n'\ and le: \n < n'\ proof - have \\ bit x n\ using that by (rule is_aligned_imp_not_bit) then have \x + 2 ^ n = x OR 2 ^ n\ by (subst disjunctive_add) (auto simp add: bit_simps disjunctive_add) moreover have \(x OR 2 ^ n) XOR 2 ^ n = x\ by (rule bit_word_eqI) (auto simp add: bit_simps \\ bit x n\) ultimately show ?thesis by simp qed lemma is_aligned_add_mult_multI: fixes p :: "'a::len word" shows "\is_aligned p m; n \ m; n' = n\ \ is_aligned (p + x * 2 ^ n * z) n'" apply (erule aligned_add_aligned) apply (auto intro: is_alignedI [where k="x*z"]) done lemma is_aligned_add_multI: fixes p :: "'a::len word" shows "\is_aligned p m; n \ m; n' = n\ \ is_aligned (p + x * 2 ^ n) n'" apply (erule aligned_add_aligned) apply (auto intro: is_alignedI [where k="x"]) done lemma is_aligned_no_wrap''': fixes ptr :: "'a::len word" shows"\ is_aligned ptr sz; sz < LENGTH('a); off < 2 ^ sz \ \ unat ptr + off < 2 ^ LENGTH('a)" apply (drule is_aligned_no_wrap[where off="of_nat off"]) apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) - apply (simp add: take_bit_eq_mod) + apply (simp add: take_bit_eq_mod unsigned_of_nat) apply (subst(asm) unat_of_nat_len) apply (erule order_less_trans) apply (erule power_strict_increasing) apply simp apply assumption done lemma is_aligned_get_word_bits: fixes p :: "'a::len word" shows "\ is_aligned p n; \ is_aligned p n; n < LENGTH('a) \ \ P; \ p = 0; n \ LENGTH('a) \ \ P \ \ P" apply (cases "n < LENGTH('a)") apply simp apply simp apply (erule meta_mp) apply (simp add: is_aligned_mask power_add power_overflow not_less flip: take_bit_eq_mask) apply (metis take_bit_length_eq take_bit_of_0 take_bit_tightened) done lemma aligned_small_is_0: "\ is_aligned x n; x < 2 ^ n \ \ x = 0" by (simp add: is_aligned_mask less_mask_eq) corollary is_aligned_less_sz: "\is_aligned a sz; a \ 0\ \ \ a < 2 ^ sz" by (rule notI, drule(1) aligned_small_is_0, erule(1) notE) lemma aligned_at_least_t2n_diff: "\is_aligned x n; is_aligned y n; x < y\ \ x \ y - 2 ^ n" apply (erule is_aligned_get_word_bits[where p=y]) apply (rule ccontr) apply (clarsimp simp: linorder_not_le) apply (subgoal_tac "y - x = 0") apply clarsimp apply (rule aligned_small_is_0) apply (erule(1) aligned_sub_aligned) apply simp apply unat_arith apply simp done lemma is_aligned_no_overflow'': "\is_aligned x n; x + 2 ^ n \ 0\ \ x \ x + 2 ^ n" apply (frule is_aligned_no_overflow') apply (erule order_trans) apply (simp add: field_simps) apply (erule word_sub_1_le) done lemma is_aligned_bitI: \is_aligned p m\ if \\n. n < m \ \ bit p n\ apply (simp add: is_aligned_mask) apply (rule bit_word_eqI) using that apply (auto simp add: bit_simps) done lemma is_aligned_nth [word_eqI_simps]: "is_aligned p m = (\n < m. \ bit p n)" apply (auto intro: is_aligned_bitI simp add: is_aligned_mask bit_eq_iff) apply (auto simp: bit_simps) using bit_imp_le_length not_less apply blast done lemma range_inter: "({a..b} \ {c..d} = {}) = (\x. \(a \ x \ x \ b \ c \ x \ x \ d))" by auto lemma aligned_inter_non_empty: "\ {p..p + (2 ^ n - 1)} \ {p..p + 2 ^ m - 1} = {}; is_aligned p n; is_aligned p m\ \ False" apply (clarsimp simp only: range_inter) apply (erule_tac x=p in allE) apply simp apply (erule impE) apply (erule is_aligned_no_overflow') apply (erule notE) apply (erule is_aligned_no_overflow) done lemma not_aligned_mod_nz: assumes al: "\ is_aligned a n" shows "a mod 2 ^ n \ 0" apply (rule ccontr) using al apply (rule notE) apply simp apply (rule is_alignedI [of _ _ \a div 2 ^ n\]) apply (metis add.right_neutral mult.commute word_mod_div_equality) done lemma nat_add_offset_le: fixes x :: nat assumes yv: "y \ 2 ^ n" and xv: "x < 2 ^ m" and mn: "sz = m + n" shows "x * 2 ^ n + y \ 2 ^ sz" proof (subst mn) from yv obtain qy where "y + qy = 2 ^ n" by (auto simp: le_iff_add) have "x * 2 ^ n + y \ x * 2 ^ n + 2 ^ n" using yv xv by simp also have "\ = (x + 1) * 2 ^ n" by simp also have "\ \ 2 ^ (m + n)" using xv by (subst power_add) (rule mult_le_mono1, simp) finally show "x * 2 ^ n + y \ 2 ^ (m + n)" . qed lemma is_aligned_no_wrap_le: fixes ptr::"'a::len word" assumes al: "is_aligned ptr sz" and szv: "sz < LENGTH('a)" and off: "off \ 2 ^ sz" shows "unat ptr + off \ 2 ^ LENGTH('a)" proof - from al obtain q where ptrq: "ptr = 2 ^ sz * of_nat q" and qv: "q < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis proof (cases "sz = 0") case True then show ?thesis using off ptrq qv by (auto simp add: le_Suc_eq Suc_le_eq) (simp add: le_less) next case False then have sne: "0 < sz" .. show ?thesis proof - have uq: "unat (of_nat q :: 'a word) = q" apply (subst unat_of_nat) apply (rule mod_less) apply (rule order_less_trans [OF qv]) apply (rule power_strict_increasing [OF diff_less [OF sne]]) apply simp_all done have uptr: "unat ptr = 2 ^ sz * q" apply (subst ptrq) apply (subst iffD1 [OF unat_mult_lem]) apply (subst unat_power_lower [OF szv]) apply (subst uq) apply (rule nat_less_power_trans [OF qv order_less_imp_le [OF szv]]) apply (subst uq) apply (subst unat_power_lower [OF szv]) apply simp done show "unat ptr + off \ 2 ^ LENGTH('a)" using szv apply (subst uptr) apply (subst mult.commute, rule nat_add_offset_le [OF off qv]) apply simp done qed qed qed lemma is_aligned_neg_mask: "m \ n \ is_aligned (x AND NOT (mask n)) m" by (rule is_aligned_bitI) (simp add: bit_simps) lemma unat_minus: "unat (- (x :: 'a :: len word)) = (if x = 0 then 0 else 2 ^ size x - unat x)" using unat_sub_if_size[where x="2 ^ size x" and y=x] by (simp add: unat_eq_0 word_size) lemma is_aligned_minus: \is_aligned (- p) n\ if \is_aligned p n\ for p :: \'a::len word\ using that apply (cases \n < LENGTH('a)\) apply (simp_all add: not_less is_aligned_beyond_length) apply transfer apply (simp flip: take_bit_eq_0_iff) apply (subst take_bit_minus [symmetric]) apply simp done lemma add_mask_lower_bits: "\is_aligned (x :: 'a :: len word) n; \n' \ n. n' < LENGTH('a) \ \ bit p n'\ \ x + p AND NOT (mask n) = x" apply (subst word_plus_and_or_coroll) apply (rule word_eqI) apply (clarsimp simp: word_size is_aligned_nth) apply (erule_tac x=na in allE)+ apply (simp add: bit_simps) apply (rule bit_word_eqI) apply (auto simp add: bit_simps not_less word_size) apply (metis is_aligned_nth not_le) done lemma is_aligned_andI1: "is_aligned x n \ is_aligned (x AND y) n" by (simp add: is_aligned_nth bit_simps) lemma is_aligned_andI2: "is_aligned y n \ is_aligned (x AND y) n" by (simp add: is_aligned_nth bit_simps) lemma is_aligned_shiftl: "is_aligned w (n - m) \ is_aligned (push_bit m w) n" by (simp add: is_aligned_nth bit_simps) lemma is_aligned_shiftr: "is_aligned w (n + m) \ is_aligned (drop_bit m w) n" by (simp add: is_aligned_nth bit_simps) lemma is_aligned_shiftl_self: "is_aligned (push_bit n p) n" by (rule is_aligned_shift) lemma is_aligned_neg_mask_eq: "is_aligned p n \ p AND NOT (mask n) = p" apply (rule bit_word_eqI) apply (auto simp add: bit_simps is_aligned_nth) done lemma is_aligned_shiftr_shiftl: "is_aligned w n \ push_bit n (drop_bit n w) = w" apply (rule bit_word_eqI) apply (auto simp add: bit_simps is_aligned_nth) using not_le_imp_less apply blast apply (metis add_diff_inverse_nat) done lemma aligned_shiftr_mask_shiftl: "is_aligned x n \ push_bit n (drop_bit n x AND mask v) = x AND mask (v + n)" apply (rule word_eqI) apply (simp add: word_size bit_simps) apply (subgoal_tac "\m. bit x m \ m \ n") apply auto[1] apply (clarsimp simp: is_aligned_mask) apply (drule_tac x=m in word_eqD) apply (frule test_bit_size) apply (simp add: word_size bit_simps) done lemma mask_zero: "is_aligned x a \ x AND mask a = 0" by (metis is_aligned_mask) lemma is_aligned_neg_mask_eq_concrete: "\ is_aligned p n; msk AND NOT (mask n) = NOT (mask n) \ \ p AND msk = p" by (metis word_bw_assocs(1) word_bw_comms(1) is_aligned_neg_mask_eq) lemma is_aligned_and_not_zero: "\ is_aligned n k; n \ 0 \ \ 2 ^ k \ n" using is_aligned_less_sz leI by blast lemma is_aligned_and_2_to_k: "(n AND 2 ^ k - 1) = 0 \ is_aligned (n :: 'a :: len word) k" by (simp add: is_aligned_mask mask_eq_decr_exp) lemma is_aligned_power2: "b \ a \ is_aligned (2 ^ a) b" by (metis is_aligned_triv is_aligned_weaken) lemma aligned_sub_aligned': "\ is_aligned (a :: 'a :: len word) n; is_aligned b n; n < LENGTH('a) \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma is_aligned_neg_mask_weaken: "\ is_aligned p n; m \ n \ \ p AND NOT (mask m) = p" using is_aligned_neg_mask_eq is_aligned_weaken by blast lemma is_aligned_neg_mask2 [simp]: "is_aligned (a AND NOT (mask n)) n" by (rule is_aligned_bitI) (simp add: bit_simps) lemma is_aligned_0': "is_aligned 0 n" by (fact is_aligned_0) lemma aligned_add_offset_no_wrap: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off < 2 ^ sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis using szv apply (subst xv) apply (subst unat_mult_power_lem[OF kl]) apply (subst mult.commute, rule nat_add_offset_less) apply (rule less_le_trans[OF unat_mono[OF offv, simplified]]) apply (erule eq_imp_le[OF unat_power_lower]) apply (rule kl) apply simp done next assume "\ sz < LENGTH('a)" with offv show ?thesis by (simp add: not_less power_overflow ) qed lemma aligned_add_offset_mod: fixes x :: "('a::len) word" assumes al: "is_aligned x sz" and kv: "k < 2 ^ sz" shows "(x + k) mod 2 ^ sz = k" proof cases assume szv: "sz < LENGTH('a)" have ux: "unat x + unat k < 2 ^ LENGTH('a)" by (rule aligned_add_offset_no_wrap) fact+ show ?thesis using al szv apply (simp flip: take_bit_eq_mod) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) apply (metis assms(2) bit_or_iff is_aligned_mask is_aligned_nth leD less_mask_eq word_and_le1 word_bw_lcs(1) word_neq_0_conv word_plus_and_or_coroll) apply (meson assms(2) leI less_2p_is_upper_bits_unset) apply (metis assms(2) bit_disjunctive_add_iff bit_imp_le_length bit_push_bit_iff is_alignedE' less_2p_is_upper_bits_unset) done next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: not_less power_overflow is_aligned_mask mask_eq_decr_exp word_mod_by_0) qed lemma aligned_neq_into_no_overlap: fixes x :: "'a::len word" assumes neq: "x \ y" and alx: "is_aligned x sz" and aly: "is_aligned y sz" shows "{x .. x + (2 ^ sz - 1)} \ {y .. y + (2 ^ sz - 1)} = {}" proof cases assume szv: "sz < LENGTH('a)" show ?thesis proof (rule equals0I, clarsimp) fix z assume xb: "x \ z" and xt: "z \ x + (2 ^ sz - 1)" and yb: "y \ z" and yt: "z \ y + (2 ^ sz - 1)" have rl: "\(p::'a word) k w. \uint p + uint k < 2 ^ LENGTH('a); w = p + k; w \ p + (2 ^ sz - 1) \ \ k < 2 ^ sz" apply - apply simp apply (subst (asm) add.commute, subst (asm) add.commute, drule word_plus_mcs_4) apply (subst add.commute, subst no_plus_overflow_uint_size) apply transfer apply simp apply (auto simp add: le_less power_2_ge_iff szv) apply (metis le_less_trans mask_eq_decr_exp mask_lt_2pn order_less_imp_le szv) done from xb obtain kx where kx: "z = x + kx" and kxl: "uint x + uint kx < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') from yb obtain ky where ky: "z = y + ky" and kyl: "uint y + uint ky < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') have "x = y" proof - have "kx = z mod 2 ^ sz" proof (subst kx, rule sym, rule aligned_add_offset_mod) show "kx < 2 ^ sz" by (rule rl) fact+ qed fact+ also have "\ = ky" proof (subst ky, rule aligned_add_offset_mod) show "ky < 2 ^ sz" using kyl ky yt by (rule rl) qed fact+ finally have kxky: "kx = ky" . moreover have "x + kx = y + ky" by (simp add: kx [symmetric] ky [symmetric]) ultimately show ?thesis by simp qed then show False using neq by simp qed next assume "\ sz < LENGTH('a)" with neq alx aly have False by (simp add: is_aligned_mask mask_eq_decr_exp power_overflow) then show ?thesis .. qed lemma is_aligned_add_helper: "\ is_aligned p n; d < 2 ^ n \ \ (p + d AND mask n = d) \ (p + d AND (NOT (mask n)) = p)" apply (subst (asm) is_aligned_mask) apply (drule less_mask_eq) apply (rule context_conjI) apply (subst word_plus_and_or_coroll) apply (simp_all flip: take_bit_eq_mask) apply (metis take_bit_eq_mask word_bw_lcs(1) word_log_esimps(1)) apply (metis add.commute add_left_imp_eq take_bit_eq_mask word_plus_and_or_coroll2) done lemmas mask_inner_mask = mask_eqs(1) lemma mask_add_aligned: "is_aligned p n \ (p + q) AND mask n = q AND mask n" apply (simp add: is_aligned_mask) apply (subst mask_inner_mask [symmetric]) apply simp done lemma mask_out_add_aligned: assumes al: "is_aligned p n" shows "p + (q AND NOT (mask n)) = (p + q) AND NOT (mask n)" using mask_add_aligned [OF al] by (simp add: mask_out_sub_mask) lemma is_aligned_add_or: "\is_aligned p n; d < 2 ^ n\ \ p + d = p OR d" apply (subst disjunctive_add, simp_all) apply (clarsimp simp: is_aligned_nth less_2p_is_upper_bits_unset) subgoal for m apply (cases \m < n\) apply (auto simp add: not_less dest: bit_imp_possible_bit) done done lemma not_greatest_aligned: "\ x < y; is_aligned x n; is_aligned y n \ \ x + 2 ^ n \ 0" by (metis NOT_mask add_diff_cancel_right' diff_0 is_aligned_neg_mask_eq not_le word_and_le1) lemma neg_mask_mono_le: "x \ y \ x AND NOT(mask n) \ y AND NOT(mask n)" for x :: "'a :: len word" proof (rule ccontr, simp add: linorder_not_le, cases "n < LENGTH('a)") case False then show "y AND NOT(mask n) < x AND NOT(mask n) \ False" by (simp add: mask_eq_decr_exp linorder_not_less power_overflow) next case True assume a: "x \ y" and b: "y AND NOT(mask n) < x AND NOT(mask n)" have word_bits: "n < LENGTH('a)" by fact have "y \ (y AND NOT(mask n)) + (y AND mask n)" by (simp add: word_plus_and_or_coroll2 add.commute) also have "\ \ (y AND NOT(mask n)) + 2 ^ n" apply (rule word_plus_mono_right) apply (rule order_less_imp_le, rule and_mask_less_size) apply (simp add: word_size word_bits) apply (rule is_aligned_no_overflow'', simp add: is_aligned_neg_mask word_bits) apply (rule not_greatest_aligned, rule b; simp add: is_aligned_neg_mask) done also have "\ \ x AND NOT(mask n)" using b apply (subst add.commute) apply (rule le_plus) apply (rule aligned_at_least_t2n_diff; simp add: is_aligned_neg_mask) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0[rotated]; simp add: is_aligned_neg_mask) done also have "\ \ x" by (rule word_and_le2) also have "x \ y" by fact finally show "False" using b by simp qed lemma and_neg_mask_eq_iff_not_mask_le: "w AND NOT(mask n) = NOT(mask n) \ NOT(mask n) \ w" for w :: \'a::len word\ by (metis eq_iff neg_mask_mono_le word_and_le1 word_and_le2 word_bw_same(1)) lemma neg_mask_le_high_bits [word_eqI_simps]: \NOT (mask n) \ w \ (\i \ {n ..< size w}. bit w i)\ (is \?P \ ?Q\) for w :: \'a::len word\ proof assume ?Q then have \w AND NOT (mask n) = NOT (mask n)\ by (auto simp add: bit_simps word_size intro: bit_word_eqI) then show ?P by (simp add: and_neg_mask_eq_iff_not_mask_le) next assume ?P then have *: \w AND NOT (mask n) = NOT (mask n)\ by (simp add: and_neg_mask_eq_iff_not_mask_le) show \?Q\ proof (rule ccontr) assume \\ (\i\{n.. then obtain m where m: \\ bit w m\ \n \ m\ \m < LENGTH('a)\ by (auto simp add: word_size) from * have \bit (w AND NOT (mask n)) m \ bit (NOT (mask n :: 'a word)) m\ by auto with m show False by (auto simp add: bit_simps) qed qed lemma is_aligned_add_less_t2n: "\is_aligned (p::'a::len word) n; d < 2^n; n \ m; p < 2^m\ \ p + d < 2^m" apply (case_tac "m < LENGTH('a)") apply (subst mask_eq_iff_w2p[symmetric]) apply (simp add: word_size) apply (simp add: is_aligned_add_or word_ao_dist less_mask_eq) apply (subst less_mask_eq) apply (erule order_less_le_trans) apply (erule(1) two_power_increasing) apply simp apply (simp add: power_overflow) done lemma aligned_offset_non_zero: "\ is_aligned x n; y < 2 ^ n; x \ 0 \ \ x + y \ 0" apply (cases "y = 0") apply simp apply (subst word_neq_0_conv) apply (subst gt0_iff_gem1) apply (erule is_aligned_get_word_bits) apply (subst field_simps[symmetric], subst plus_le_left_cancel_nowrap) apply (rule is_aligned_no_wrap') apply simp apply (rule word_leq_le_minus_one) apply simp apply assumption apply (erule (1) is_aligned_no_wrap') apply (simp add: gt0_iff_gem1 [symmetric] word_neq_0_conv) apply simp done lemma is_aligned_over_length: "\ is_aligned p n; LENGTH('a) \ n \ \ (p::'a::len word) = 0" by (simp add: is_aligned_mask mask_over_length) lemma is_aligned_no_overflow_mask: "is_aligned x n \ x \ x + mask n" by (simp add: mask_eq_decr_exp) (erule is_aligned_no_overflow') lemma aligned_mask_step: "\ n' \ n; p' \ p + mask n; is_aligned p n; is_aligned p' n' \ \ (p'::'a::len word) + mask n' \ p + mask n" apply (cases "LENGTH('a) \ n") apply (frule (1) is_aligned_over_length) apply (drule mask_over_length) apply clarsimp apply (simp add: not_le) apply (simp add: word_le_nat_alt unat_plus_simple) apply (subst unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask)+ apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (clarsimp simp: dvd_def is_aligned_iff_dvd_nat) apply (rename_tac k k') apply (thin_tac "unat p = x" for p x)+ apply (subst Suc_le_mono[symmetric]) apply (simp only: Suc_2p_unat_mask) apply (drule le_imp_less_Suc, subst (asm) Suc_2p_unat_mask, assumption) apply (erule (1) power_2_mult_step_le) done lemma is_aligned_mask_offset_unat: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off \ mask sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) from offv szv have offv': "unat off < 2 ^ sz" by (simp add: mask_2pm1 unat_less_power) show ?thesis using szv using al is_aligned_no_wrap''' offv' by blast next assume "\ sz < LENGTH('a)" with al have "x = 0" by (meson is_aligned_get_word_bits) thus ?thesis by simp qed lemma aligned_less_plus_1: "\ is_aligned x n; n > 0 \ \ x < x + 1" apply (rule plus_one_helper2) apply (rule order_refl) apply (clarsimp simp: field_simps) apply (drule arg_cong[where f="\x. x - 1"]) apply (clarsimp simp: is_aligned_mask) apply (drule word_eqD[where x=0]) apply (simp add: even_mask_iff) done lemma aligned_add_offset_less: "\is_aligned x n; is_aligned y n; x < y; z < 2 ^ n\ \ x + z < y" apply (cases "y = 0") apply simp apply (erule is_aligned_get_word_bits[where p=y], simp_all) apply (cases "z = 0", simp_all) apply (drule(2) aligned_at_least_t2n_diff[rotated -1]) apply (drule plus_one_helper2) apply (rule less_is_non_zero_p1) apply (rule aligned_less_plus_1) apply (erule aligned_sub_aligned[OF _ _ order_refl], simp_all add: is_aligned_triv)[1] apply (cases n, simp_all)[1] apply (simp only: trans[OF diff_add_eq diff_diff_eq2[symmetric]]) apply (drule word_less_add_right) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0, erule order_less_trans) apply (clarsimp simp: power_overflow) apply simp apply (erule order_le_less_trans[rotated], rule word_plus_mono_right) apply (erule word_le_minus_one_leq) apply (simp add: is_aligned_no_wrap' is_aligned_no_overflow field_simps) done lemma gap_between_aligned: "\a < (b :: 'a ::len word); is_aligned a n; is_aligned b n; n < LENGTH('a) \ \ a + (2^n - 1) < b" by (simp add: aligned_add_offset_less) lemma is_aligned_add_step_le: "\ is_aligned (a::'a::len word) n; is_aligned b n; a < b; b \ a + mask n \ \ False" apply (simp flip: not_le) apply (erule notE) apply (cases "LENGTH('a) \ n") apply (drule (1) is_aligned_over_length)+ apply (drule mask_over_length) apply clarsimp apply (clarsimp simp: word_le_nat_alt not_less not_le) apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (subst (asm) unat_add_lem' [symmetric]) apply (simp add: is_aligned_mask_offset_unat) apply (metis gap_between_aligned linorder_not_less mask_eq_decr_exp unat_arith_simps(2)) done lemma aligned_add_mask_lessD: "\ x + mask n < y; is_aligned x n \ \ x < y" for y::"'a::len word" by (metis is_aligned_no_overflow' mask_2pm1 order_le_less_trans) lemma aligned_add_mask_less_eq: "\ is_aligned x n; is_aligned y n; n < LENGTH('a) \ \ (x + mask n < y) = (x < y)" for y::"'a::len word" using aligned_add_mask_lessD is_aligned_add_step_le word_le_not_less by blast lemma is_aligned_diff: fixes m :: "'a::len word" assumes alm: "is_aligned m s1" and aln: "is_aligned n s2" and s2wb: "s2 < LENGTH('a)" and nm: "m \ {n .. n + (2 ^ s2 - 1)}" and s1s2: "s1 \ s2" and s10: "0 < s1" (* Probably can be folded into the proof \ *) shows "\q. m - n = of_nat q * 2 ^ s1 \ q < 2 ^ (s2 - s1)" proof - have rl: "\m s. \ m < 2 ^ (LENGTH('a) - s); s < LENGTH('a) \ \ unat ((2::'a word) ^ s * of_nat m) = 2 ^ s * m" proof - fix m :: nat and s assume m: "m < 2 ^ (LENGTH('a) - s)" and s: "s < LENGTH('a)" then have "unat ((of_nat m) :: 'a word) = m" apply (subst unat_of_nat) apply (subst mod_less) apply (erule order_less_le_trans) apply (rule power_increasing) apply simp_all done then show "?thesis m s" using s m apply (subst iffD1 [OF unat_mult_lem]) apply (simp add: nat_less_power_trans)+ done qed have s1wb: "s1 < LENGTH('a)" using s2wb s1s2 by simp from alm obtain mq where mmq: "m = 2 ^ s1 * of_nat mq" and mq: "mq < 2 ^ (LENGTH('a) - s1)" by (auto elim: is_alignedE simp: field_simps) from aln obtain nq where nnq: "n = 2 ^ s2 * of_nat nq" and nq: "nq < 2 ^ (LENGTH('a) - s2)" by (auto elim: is_alignedE simp: field_simps) from s1s2 obtain sq where sq: "s2 = s1 + sq" by (auto simp: le_iff_add) note us1 = rl [OF mq s1wb] note us2 = rl [OF nq s2wb] from nm have "n \ m" by clarsimp then have "(2::'a word) ^ s2 * of_nat nq \ 2 ^ s1 * of_nat mq" using nnq mmq by simp then have "2 ^ s2 * nq \ 2 ^ s1 * mq" using s1wb s2wb by (simp add: word_le_nat_alt us1 us2) then have nqmq: "2 ^ sq * nq \ mq" using sq by (simp add: power_add) have "m - n = 2 ^ s1 * of_nat mq - 2 ^ s2 * of_nat nq" using mmq nnq by simp also have "\ = 2 ^ s1 * of_nat mq - 2 ^ s1 * 2 ^ sq * of_nat nq" using sq by (simp add: power_add) also have "\ = 2 ^ s1 * (of_nat mq - 2 ^ sq * of_nat nq)" by (simp add: field_simps) also have "\ = 2 ^ s1 * of_nat (mq - 2 ^ sq * nq)" using s1wb s2wb us1 us2 nqmq by (simp add: of_nat_diff) finally have mn: "m - n = of_nat (mq - 2 ^ sq * nq) * 2 ^ s1" by simp moreover from nm have "m - n \ 2 ^ s2 - 1" by - (rule word_diff_ls', (simp add: field_simps)+) then have "(2::'a word) ^ s1 * of_nat (mq - 2 ^ sq * nq) < 2 ^ s2" using mn s2wb by (simp add: field_simps) then have "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (s2 - s1)" proof (rule word_power_less_diff) have mm: "mq - 2 ^ sq * nq < 2 ^ (LENGTH('a) - s1)" using mq by simp moreover from s10 have "LENGTH('a) - s1 < LENGTH('a)" by (rule diff_less, simp) ultimately show "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (LENGTH('a) - s1)" using take_bit_nat_less_self_iff [of \LENGTH('a)\ \mq - 2 ^ sq * nq\] - apply (auto simp add: word_less_nat_alt not_le not_less) + apply (auto simp add: word_less_nat_alt not_le not_less unsigned_of_nat) apply (metis take_bit_nat_eq_self_iff) done qed then have "mq - 2 ^ sq * nq < 2 ^ (s2 - s1)" using mq s2wb - apply (simp add: word_less_nat_alt take_bit_eq_mod) + apply (simp add: word_less_nat_alt take_bit_eq_mod unsigned_of_nat) apply (subst (asm) mod_less) apply auto apply (rule order_le_less_trans) apply (rule diff_le_self) apply (erule order_less_le_trans) apply simp done ultimately show ?thesis by auto qed lemma is_aligned_addD1: assumes al1: "is_aligned (x + y) n" and al2: "is_aligned (x::'a::len word) n" shows "is_aligned y n" using al2 proof (rule is_aligned_get_word_bits) assume "x = 0" then show ?thesis using al1 by simp next assume nv: "n < LENGTH('a)" from al1 obtain q1 where xy: "x + y = 2 ^ n * of_nat q1" and "q1 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) moreover from al2 obtain q2 where x: "x = 2 ^ n * of_nat q2" and "q2 < 2 ^ (LENGTH('a) - n)" by (rule is_alignedE) ultimately have "y = 2 ^ n * (of_nat q1 - of_nat q2)" by (simp add: field_simps) then show ?thesis using nv by (simp add: is_aligned_mult_triv1) qed lemmas is_aligned_addD2 = is_aligned_addD1[OF subst[OF add.commute, of "%x. is_aligned x n" for n]] lemma is_aligned_add: "\is_aligned p n; is_aligned q n\ \ is_aligned (p + q) n" by (simp add: is_aligned_mask mask_add_aligned) lemma aligned_shift: "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ drop_bit n (x + y) = drop_bit n y" apply (subst word_plus_and_or_coroll; rule bit_word_eqI) apply (auto simp add: bit_simps is_aligned_nth) apply (metis less_2p_is_upper_bits_unset not_le) apply (metis le_add1 less_2p_is_upper_bits_unset test_bit_bin) done lemma aligned_shift': "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ drop_bit n (y + x) = drop_bit n y" apply (subst word_plus_and_or_coroll; rule bit_word_eqI) apply (auto simp add: bit_simps is_aligned_nth) apply (metis less_2p_is_upper_bits_unset not_le) apply (metis bit_imp_le_length le_add1 less_2p_is_upper_bits_unset) done lemma and_neg_mask_plus_mask_mono: "(p AND NOT (mask n)) + mask n \ p" for p :: \'a::len word\ apply (rule word_le_minus_cancel[where x = "p AND NOT (mask n)"]) apply (clarsimp simp: subtract_mask) using word_and_le1[where a = "mask n" and y = p] apply (clarsimp simp: mask_eq_decr_exp word_le_less_eq) apply (rule is_aligned_no_overflow'[folded mask_2pm1]) apply (clarsimp simp: is_aligned_neg_mask) done lemma word_neg_and_le: "ptr \ (ptr AND NOT (mask n)) + (2 ^ n - 1)" for ptr :: \'a::len word\ by (simp add: and_neg_mask_plus_mask_mono mask_2pm1[symmetric]) lemma is_aligned_sub_helper: "\ is_aligned (p - d) n; d < 2 ^ n \ \ (p AND mask n = d) \ (p AND (NOT (mask n)) = p - d)" by (drule(1) is_aligned_add_helper, simp) lemma is_aligned_after_mask: "\is_aligned k m;m\ n\ \ is_aligned (k AND mask n) m" by (rule is_aligned_andI1) lemma and_mask_plus: "\is_aligned ptr m; m \ n; a < 2 ^ m\ \ ptr + a AND mask n = (ptr AND mask n) + a" apply (rule mask_eqI[where n = m]) apply (simp add:mask_twice min_def) apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct1]) apply (erule is_aligned_after_mask) apply simp apply simp apply simp apply (subgoal_tac "(ptr + a AND mask n) AND NOT (mask m) = (ptr + a AND NOT (mask m) ) AND mask n") apply (simp add:is_aligned_add_helper) apply (subst is_aligned_add_helper[THEN conjunct2]) apply (simp add:is_aligned_after_mask) apply simp apply simp apply (simp add:word_bw_comms word_bw_lcs) done lemma is_aligned_add_not_aligned: "\is_aligned (p::'a::len word) n; \ is_aligned (q::'a::len word) n\ \ \ is_aligned (p + q) n" by (metis is_aligned_addD1) lemma neg_mask_add_aligned: "\ is_aligned p n; q < 2 ^ n \ \ (p + q) AND NOT (mask n) = p AND NOT (mask n)" by (metis is_aligned_add_helper is_aligned_neg_mask_eq) lemma word_add_power_off: fixes a :: "'a :: len word" assumes ak: "a < k" and kw: "k < 2 ^ (LENGTH('a) - m)" and mw: "m < LENGTH('a)" and off: "off < 2 ^ m" shows "(a * 2 ^ m) + off < k * 2 ^ m" proof (cases "m = 0") case True then show ?thesis using off ak by simp next case False from ak have ak1: "a + 1 \ k" by (rule inc_le) then have "(a + 1) * 2 ^ m \ 0" apply - apply (rule word_power_nonzero) apply (erule order_le_less_trans [OF _ kw]) apply (rule mw) apply (rule less_is_non_zero_p1 [OF ak]) done then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw apply - apply (simp add: distrib_right) apply (rule word_plus_strict_mono_right [OF off]) apply (rule is_aligned_no_overflow'') apply (rule is_aligned_mult_triv2) apply assumption done also have "\ \ k * 2 ^ m" using ak1 mw kw False apply - apply (erule word_mult_le_mono1) apply (simp add: p2_gt_0) apply (simp add: word_less_nat_alt) apply (meson nat_mult_power_less_eq zero_less_numeral) done finally show ?thesis . qed lemma offset_not_aligned: "\ is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\ \ \ is_aligned (p + of_nat i) n" apply (erule is_aligned_add_not_aligned) apply transfer apply (auto simp add: is_aligned_iff_udvd) apply (metis (no_types, lifting) le_less_trans less_not_refl2 less_or_eq_imp_le of_nat_eq_0_iff take_bit_eq_0_iff take_bit_nat_eq_self_iff take_bit_of_nat unat_lt2p unat_power_lower) done lemma le_or_mask: "w \ w' \ w OR mask x \ w' OR mask x" for w w' :: \'a::len word\ by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) end end diff --git a/thys/Word_Lib/More_Word.thy b/thys/Word_Lib/More_Word.thy --- a/thys/Word_Lib/More_Word.thy +++ b/thys/Word_Lib/More_Word.thy @@ -1,2543 +1,2546 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section \Lemmas on words\ theory More_Word imports "HOL-Library.Word" More_Arithmetic More_Divides begin context includes bit_operations_syntax begin \ \problem posed by TPHOLs referee: criterion for overflow of addition of signed integers\ lemma sofl_test: \sint x + sint y = sint (x + y) \ drop_bit (size x - 1) ((x + y XOR x) AND (x + y XOR y)) = 0\ for x y :: \'a::len word\ proof - obtain n where n: \LENGTH('a) = Suc n\ by (cases \LENGTH('a)\) simp_all have *: \sint x + sint y + 2 ^ Suc n > signed_take_bit n (sint x + sint y) \ sint x + sint y \ - (2 ^ n)\ \signed_take_bit n (sint x + sint y) > sint x + sint y - 2 ^ Suc n \ 2 ^ n > sint x + sint y\ using signed_take_bit_int_greater_eq [of \sint x + sint y\ n] signed_take_bit_int_less_eq [of n \sint x + sint y\] by (auto intro: ccontr) have \sint x + sint y = sint (x + y) \ (sint (x + y) < 0 \ sint x < 0) \ (sint (x + y) < 0 \ sint y < 0)\ using sint_less [of x] sint_greater_eq [of x] sint_less [of y] sint_greater_eq [of y] signed_take_bit_int_eq_self [of \LENGTH('a) - 1\ \sint x + sint y\] apply (auto simp add: not_less) apply (unfold sint_word_ariths) apply (subst signed_take_bit_int_eq_self) prefer 4 apply (subst signed_take_bit_int_eq_self) prefer 7 apply (subst signed_take_bit_int_eq_self) prefer 10 apply (subst signed_take_bit_int_eq_self) apply (auto simp add: signed_take_bit_int_eq_self signed_take_bit_eq_take_bit_minus take_bit_Suc_from_most n not_less intro!: *) apply (smt (z3) take_bit_nonnegative) apply (smt (z3) take_bit_int_less_exp) apply (smt (z3) take_bit_nonnegative) apply (smt (z3) take_bit_int_less_exp) done then show ?thesis apply (simp only: One_nat_def word_size drop_bit_eq_zero_iff_not_bit_last bit_and_iff bit_xor_iff) apply (simp add: bit_last_iff) done qed lemma unat_power_lower [simp]: "unat ((2::'a::len word) ^ n) = 2 ^ n" if "n < LENGTH('a::len)" using that by transfer simp lemma unat_p2: "n < LENGTH('a :: len) \ unat (2 ^ n :: 'a word) = 2 ^ n" by (fact unat_power_lower) lemma word_div_lt_eq_0: "x < y \ x div y = 0" for x :: "'a :: len word" by transfer simp lemma word_div_eq_1_iff: "n div m = 1 \ n \ m \ unat n < 2 * unat (m :: 'a :: len word)" apply (simp only: word_arith_nat_defs word_le_nat_alt word_of_nat_eq_iff flip: nat_div_eq_Suc_0_iff) apply (simp flip: unat_div unsigned_take_bit_eq) done lemma AND_twice [simp]: "(w AND m) AND m = w AND m" by (fact and.right_idem) lemma word_combine_masks: "w AND m = z \ w AND m' = z' \ w AND (m OR m') = (z OR z')" for w m m' z z' :: \'a::len word\ by (simp add: bit.conj_disj_distrib) lemma p2_gt_0: "(0 < (2 ^ n :: 'a :: len word)) = (n < LENGTH('a))" by (simp add : word_gt_0 not_le) lemma uint_2p_alt: \n < LENGTH('a::len) \ uint ((2::'a::len word) ^ n) = 2 ^ n\ using p2_gt_0 [of n, where ?'a = 'a] by (simp add: uint_2p) lemma p2_eq_0: \(2::'a::len word) ^ n = 0 \ LENGTH('a::len) \ n\ by (fact exp_eq_zero_iff) lemma p2len: \(2 :: 'a word) ^ LENGTH('a::len) = 0\ by simp lemma neg_mask_is_div: "w AND NOT (mask n) = (w div 2^n) * 2^n" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps simp flip: push_bit_eq_mult drop_bit_eq_div) lemma neg_mask_is_div': "n < size w \ w AND NOT (mask n) = ((w div (2 ^ n)) * (2 ^ n))" for w :: \'a::len word\ by (rule neg_mask_is_div) lemma and_mask_arith: "w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps word_size simp flip: push_bit_eq_mult drop_bit_eq_div) lemma and_mask_arith': "0 < n \ w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" for w :: \'a::len word\ by (rule and_mask_arith) lemma mask_2pm1: "mask n = 2 ^ n - (1 :: 'a::len word)" by (fact mask_eq_decr_exp) lemma add_mask_fold: "x + 2 ^ n - 1 = x + mask n" for x :: \'a::len word\ by (simp add: mask_eq_decr_exp) lemma word_and_mask_le_2pm1: "w AND mask n \ 2 ^ n - 1" for w :: \'a::len word\ by (simp add: mask_2pm1[symmetric] word_and_le1) lemma is_aligned_AND_less_0: "u AND mask n = 0 \ v < 2^n \ u AND v = 0" for u v :: \'a::len word\ apply (drule less_mask_eq) apply (simp flip: take_bit_eq_mask) apply (simp add: bit_eq_iff) apply (auto simp add: bit_simps) done lemma and_mask_eq_iff_le_mask: \w AND mask n = w \ w \ mask n\ for w :: \'a::len word\ apply (simp flip: take_bit_eq_mask) apply (cases \n \ LENGTH('a)\; transfer) apply (simp_all add: not_le min_def) apply (simp_all add: mask_eq_exp_minus_1) apply auto apply (metis take_bit_int_less_exp) apply (metis min_def nat_less_le take_bit_int_eq_self_iff take_bit_take_bit) done lemma less_eq_mask_iff_take_bit_eq_self: \w \ mask n \ take_bit n w = w\ for w :: \'a::len word\ by (simp add: and_mask_eq_iff_le_mask take_bit_eq_mask) lemma NOT_eq: "NOT (x :: 'a :: len word) = - x - 1" apply (cut_tac x = "x" in word_add_not) apply (drule add.commute [THEN trans]) apply (drule eq_diff_eq [THEN iffD2]) by simp lemma NOT_mask: "NOT (mask n :: 'a::len word) = - (2 ^ n)" by (simp add : NOT_eq mask_2pm1) lemma le_m1_iff_lt: "(x > (0 :: 'a :: len word)) = ((y \ x - 1) = (y < x))" by uint_arith lemma gt0_iff_gem1: \0 < x \ x - 1 < x\ for x :: \'a::len word\ by (metis add.right_neutral diff_add_cancel less_irrefl measure_unat unat_arith_simps(2) word_neq_0_conv word_sub_less_iff) lemma power_2_ge_iff: \2 ^ n - (1 :: 'a::len word) < 2 ^ n \ n < LENGTH('a)\ using gt0_iff_gem1 p2_gt_0 by blast lemma le_mask_iff_lt_2n: "n < len_of TYPE ('a) = (((w :: 'a :: len word) \ mask n) = (w < 2 ^ n))" unfolding mask_2pm1 by (rule trans [OF p2_gt_0 [THEN sym] le_m1_iff_lt]) lemma mask_lt_2pn: \n < LENGTH('a) \ mask n < (2 :: 'a::len word) ^ n\ by (simp add: mask_eq_exp_minus_1 power_2_ge_iff) lemma word_unat_power: "(2 :: 'a :: len word) ^ n = of_nat (2 ^ n)" by simp lemma of_nat_mono_maybe: assumes xlt: "x < 2 ^ len_of TYPE ('a)" shows "y < x \ of_nat y < (of_nat x :: 'a :: len word)" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (erule order_less_trans [OF _ xlt]) apply (subst mod_less [OF xlt]) apply assumption done lemma word_and_max_word: fixes a::"'a::len word" shows "x = - 1 \ a AND x = a" by simp lemma word_and_full_mask_simp: \x AND mask LENGTH('a) = x\ for x :: \'a::len word\ by (simp add: bit_eq_iff bit_simps) lemma of_int_uint: "of_int (uint x) = x" by (fact word_of_int_uint) corollary word_plus_and_or_coroll: "x AND y = 0 \ x + y = x OR y" for x y :: \'a::len word\ using word_plus_and_or[where x=x and y=y] by simp corollary word_plus_and_or_coroll2: "(x AND w) + (x AND NOT w) = x" for x w :: \'a::len word\ apply (subst disjunctive_add) apply (simp add: bit_simps) apply (simp flip: bit.conj_disj_distrib) done lemma nat_mask_eq: \nat (mask n) = mask n\ by (simp add: nat_eq_iff of_nat_mask_eq) lemma unat_mask_eq: \unat (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ by transfer (simp add: nat_mask_eq) lemma word_plus_mono_left: fixes x :: "'a :: len word" shows "\y \ z; x \ x + z\ \ y + x \ z + x" by unat_arith lemma less_Suc_unat_less_bound: "n < Suc (unat (x :: 'a :: len word)) \ n < 2 ^ LENGTH('a)" by (auto elim!: order_less_le_trans intro: Suc_leI) lemma up_ucast_inj: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ len_of TYPE ('b) \ \ x = (y::'a::len word)" by transfer (simp add: min_def split: if_splits) lemmas ucast_up_inj = up_ucast_inj lemma up_ucast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (ucast x = (ucast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_ucast_inj) lemma no_plus_overflow_neg: "(x :: 'a :: len word) < -y \ x \ x + y" by (metis diff_minus_eq_add less_imp_le sub_wrap_lt) lemma ucast_ucast_eq: "\ ucast x = (ucast (ucast y::'a word)::'c::len word); LENGTH('a) \ LENGTH('b); LENGTH('b) \ LENGTH('c) \ \ x = ucast y" for x :: "'a::len word" and y :: "'b::len word" apply transfer apply (cases \LENGTH('c) = LENGTH('a)\) apply (auto simp add: min_def) done lemma ucast_0_I: "x = 0 \ ucast x = 0" by simp lemma word_add_offset_less: fixes x :: "'a :: len word" assumes yv: "y < 2 ^ n" and xv: "x < 2 ^ m" and mnv: "sz < LENGTH('a :: len)" and xv': "x < 2 ^ (LENGTH('a :: len) - n)" and mn: "sz = m + n" shows "x * 2 ^ n + y < 2 ^ sz" proof (subst mn) from mnv mn have nv: "n < LENGTH('a)" and mv: "m < LENGTH('a)" by auto have uy: "unat y < 2 ^ n" by (rule order_less_le_trans [OF unat_mono [OF yv] order_eq_refl], rule unat_power_lower[OF nv]) have ux: "unat x < 2 ^ m" by (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl], rule unat_power_lower[OF mv]) then show "x * 2 ^ n + y < 2 ^ (m + n)" using ux uy nv mnv xv' apply (subst word_less_nat_alt) apply (subst unat_word_ariths)+ apply (subst mod_less) apply simp apply (subst mult.commute) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv']]) apply (cases "n = 0"; simp) apply (subst unat_power_lower[OF nv]) apply (subst mod_less) apply (erule order_less_le_trans [OF nat_add_offset_less], assumption) apply (rule mn) apply simp apply (simp add: mn mnv) apply (erule nat_add_offset_less; simp) done qed lemma word_less_power_trans: fixes n :: "'a :: len word" assumes nv: "n < 2 ^ (m - k)" and kv: "k \ m" and mv: "m < len_of TYPE ('a)" shows "2 ^ k * n < 2 ^ m" using nv kv mv apply - apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply simp apply (rule nat_less_power_trans) apply (erule order_less_trans [OF unat_mono]) apply simp apply simp apply simp apply (rule nat_less_power_trans) apply (subst unat_power_lower[where 'a = 'a, symmetric]) apply simp apply (erule unat_mono) apply simp done lemma word_less_power_trans2: fixes n :: "'a::len word" shows "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ n * 2 ^ k < 2 ^ m" by (subst field_simps, rule word_less_power_trans) lemma Suc_unat_diff_1: fixes x :: "'a :: len word" assumes lt: "1 \ x" shows "Suc (unat (x - 1)) = unat x" proof - have "0 < unat x" by (rule order_less_le_trans [where y = 1], simp, subst unat_1 [symmetric], rule iffD1 [OF word_le_nat_alt lt]) then show ?thesis by ((subst unat_sub [OF lt])+, simp only: unat_1) qed lemma word_eq_unatI: \v = w\ if \unat v = unat w\ using that by transfer (simp add: nat_eq_iff) lemma word_div_sub: fixes x :: "'a :: len word" assumes yx: "y \ x" and y0: "0 < y" shows "(x - y) div y = x div y - 1" apply (rule word_eq_unatI) apply (subst unat_div) apply (subst unat_sub [OF yx]) apply (subst unat_sub) apply (subst word_le_nat_alt) apply (subst unat_div) apply (subst le_div_geq) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply (subst word_le_nat_alt [symmetric], rule yx) apply simp apply (subst unat_div) apply (subst le_div_geq [OF _ iffD1 [OF word_le_nat_alt yx]]) apply (rule order_le_less_trans [OF _ unat_mono [OF y0]]) apply simp apply simp done lemma word_mult_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k < j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto intro: order_less_subst2 simp: word_less_nat_alt elim: mult_less_mono1) then show ?thesis using ujk knz ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_less_dest: fixes i :: "'a :: len word" assumes ij: "i * k < j * k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem] elim: mult_less_mono1) lemma word_mult_less_cancel: fixes k :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" by (rule iffI [OF word_mult_less_dest [OF _ uik ujk] word_mult_less_mono1 [OF _ knz ujk]]) lemma Suc_div_unat_helper: assumes szv: "sz < LENGTH('a :: len)" and usszv: "us \ sz" shows "2 ^ (sz - us) = Suc (unat (((2::'a :: len word) ^ sz - 1) div 2 ^ us))" proof - note usv = order_le_less_trans [OF usszv szv] from usszv obtain q where qv: "sz = us + q" by (auto simp: le_iff_add) have "Suc (unat (((2:: 'a word) ^ sz - 1) div 2 ^ us)) = (2 ^ us + unat ((2:: 'a word) ^ sz - 1)) div 2 ^ us" apply (subst unat_div unat_power_lower[OF usv])+ apply (subst div_add_self1, simp+) done also have "\ = ((2 ^ us - 1) + 2 ^ sz) div 2 ^ us" using szv by (simp add: unat_minus_one) also have "\ = 2 ^ q + ((2 ^ us - 1) div 2 ^ us)" apply (subst qv) apply (subst power_add) apply (subst div_mult_self2; simp) done also have "\ = 2 ^ (sz - us)" using qv by simp finally show ?thesis .. qed lemma enum_word_nth_eq: \(Enum.enum :: 'a::len word list) ! n = word_of_nat n\ if \n < 2 ^ LENGTH('a)\ for n using that by (simp add: enum_word_def) lemma length_enum_word_eq: \length (Enum.enum :: 'a::len word list) = 2 ^ LENGTH('a)\ by (simp add: enum_word_def) lemma unat_lt2p [iff]: \unat x < 2 ^ LENGTH('a)\ for x :: \'a::len word\ by transfer simp lemma of_nat_unat [simp]: "of_nat \ unat = id" by (rule ext, simp) lemma Suc_unat_minus_one [simp]: "x \ 0 \ Suc (unat (x - 1)) = unat x" by (metis Suc_diff_1 unat_gt_0 unat_minus_one) lemma word_add_le_dest: fixes i :: "'a :: len word" assumes le: "i + k \ j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i \ j" using uik ujk le by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem] elim: add_le_mono1) lemma word_add_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k \ j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: add_le_mono1) then show ?thesis using ujk ij by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_le_mono2: fixes i :: "'a :: len word" shows "\i \ j; unat j + unat k < 2 ^ LENGTH('a)\ \ k + i \ k + j" by (subst field_simps, subst field_simps, erule (1) word_add_le_mono1) lemma word_add_le_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k \ j + k) = (i \ j)" proof assume "i \ j" show "i + k \ j + k" by (rule word_add_le_mono1) fact+ next assume "i + k \ j + k" show "i \ j" by (rule word_add_le_dest) fact+ qed lemma word_add_less_mono1: fixes i :: "'a :: len word" assumes ij: "i < j" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i + k < j + k" proof - from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_less_nat_alt elim: add_less_mono1) then show ?thesis using ujk ij by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem]) qed lemma word_add_less_dest: fixes i :: "'a :: len word" assumes le: "i + k < j + k" and uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "i < j" using uik ujk le by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem] elim: add_less_mono1) lemma word_add_less_iff: fixes i :: "'a :: len word" assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)" shows "(i + k < j + k) = (i < j)" proof assume "i < j" show "i + k < j + k" by (rule word_add_less_mono1) fact+ next assume "i + k < j + k" show "i < j" by (rule word_add_less_dest) fact+ qed lemma word_mult_less_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k < j * k) = (i < j)" using assms by (rule word_mult_less_cancel) lemma word_le_imp_diff_le: fixes n :: "'a::len word" shows "\k \ n; n \ m\ \ n - k \ m" by (auto simp: unat_sub word_le_nat_alt) lemma word_less_imp_diff_less: fixes n :: "'a::len word" shows "\k \ n; n < m\ \ n - k < m" by (clarsimp simp: unat_sub word_less_nat_alt intro!: less_imp_diff_less) lemma word_mult_le_mono1: fixes i :: "'a :: len word" assumes ij: "i \ j" and knz: "0 < k" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "i * k \ j * k" proof - from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)" by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: mult_le_mono1) then show ?thesis using ujk knz ij by (auto simp: word_le_nat_alt iffD1 [OF unat_mult_lem]) qed lemma word_mult_le_iff: fixes i :: "'a :: len word" assumes knz: "0 < k" and uik: "unat i * unat k < 2 ^ len_of TYPE ('a)" and ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)" shows "(i * k \ j * k) = (i \ j)" proof assume "i \ j" show "i * k \ j * k" by (rule word_mult_le_mono1) fact+ next assume p: "i * k \ j * k" have "0 < unat k" using knz by (simp add: word_less_nat_alt) then show "i \ j" using p by (clarsimp simp: word_le_nat_alt iffD1 [OF unat_mult_lem uik] iffD1 [OF unat_mult_lem ujk]) qed lemma word_diff_less: fixes n :: "'a :: len word" shows "\0 < n; 0 < m; n \ m\ \ m - n < m" apply (subst word_less_nat_alt) apply (subst unat_sub) apply assumption apply (rule diff_less) apply (simp_all add: word_less_nat_alt) done lemma word_add_increasing: fixes x :: "'a :: len word" shows "\ p + w \ x; p \ p + w \ \ p \ x" by unat_arith lemma word_random: fixes x :: "'a :: len word" shows "\ p \ p + x'; x \ x' \ \ p \ p + x" by unat_arith lemma word_sub_mono: "\ a \ c; d \ b; a - b \ a; c - d \ c \ \ (a - b) \ (c - d :: 'a :: len word)" by unat_arith lemma power_not_zero: "n < LENGTH('a::len) \ (2 :: 'a word) ^ n \ 0" by (metis p2_gt_0 word_neq_0_conv) lemma word_gt_a_gt_0: "a < n \ (0 :: 'a::len word) < n" apply (case_tac "n = 0") apply clarsimp apply (clarsimp simp: word_neq_0_conv) done lemma word_power_less_1 [simp]: "sz < LENGTH('a::len) \ (2::'a word) ^ sz - 1 < 2 ^ sz" apply (simp add: word_less_nat_alt) apply (subst unat_minus_one) apply simp_all done lemma word_sub_1_le: "x \ 0 \ x - 1 \ (x :: ('a :: len) word)" apply (subst no_ulen_sub) apply simp apply (cases "uint x = 0") apply (simp add: uint_0_iff) apply (insert uint_ge_0[where x=x]) apply arith done lemma push_bit_word_eq_nonzero: \push_bit n w \ 0\ if \w < 2 ^ m\ \m + n < LENGTH('a)\ \w \ 0\ for w :: \'a::len word\ using that apply (simp only: word_neq_0_conv word_less_nat_alt mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (metis add_diff_cancel_right' gr0I gr_implies_not0 less_or_eq_imp_le min_def push_bit_eq_0_iff take_bit_nat_eq_self_iff take_bit_push_bit take_bit_take_bit unsigned_push_bit_eq) done lemma unat_less_power: fixes k :: "'a::len word" assumes szv: "sz < LENGTH('a)" and kv: "k < 2 ^ sz" shows "unat k < 2 ^ sz" using szv unat_mono [OF kv] by simp lemma unat_mult_power_lem: assumes kv: "k < 2 ^ (LENGTH('a::len) - sz)" shows "unat (2 ^ sz * of_nat k :: (('a::len) word)) = 2 ^ sz * k" proof (cases \sz < LENGTH('a)\) case True with assms show ?thesis - by (simp add: unat_word_ariths take_bit_eq_mod mod_simps) + by (simp add: unat_word_ariths take_bit_eq_mod mod_simps unsigned_of_nat) (simp add: take_bit_nat_eq_self_iff nat_less_power_trans flip: take_bit_eq_mod) next case False with assms show ?thesis by simp qed lemma word_plus_mcs_4: "\v + x \ w + x; x \ v + x\ \ v \ (w::'a::len word)" by uint_arith lemma word_plus_mcs_3: "\v \ w; x \ w + x\ \ v + x \ w + (x::'a::len word)" by unat_arith lemma word_le_minus_one_leq: "x < y \ x \ y - 1" for x :: "'a :: len word" by transfer (metis le_less_trans less_irrefl take_bit_decr_eq take_bit_nonnegative zle_diff1_eq) lemma word_less_sub_le[simp]: fixes x :: "'a :: len word" assumes nv: "n < LENGTH('a)" shows "(x \ 2 ^ n - 1) = (x < 2 ^ n)" using le_less_trans word_le_minus_one_leq nv power_2_ge_iff by blast lemma unat_of_nat_len: "x < 2 ^ LENGTH('a) \ unat (of_nat x :: 'a::len word) = x" - by (simp add: take_bit_nat_eq_self_iff) + by (simp add: unsigned_of_nat take_bit_nat_eq_self_iff) lemma unat_of_nat_eq: "x < 2 ^ LENGTH('a) \ unat (of_nat x ::'a::len word) = x" by (rule unat_of_nat_len) lemma unat_eq_of_nat: "n < 2 ^ LENGTH('a) \ (unat (x :: 'a::len word) = n) = (x = of_nat n)" by transfer (auto simp add: take_bit_of_nat nat_eq_iff take_bit_nat_eq_self_iff intro: sym) lemma alignUp_div_helper: fixes a :: "'a::len word" assumes kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" and le: "a \ x" and sz: "n < LENGTH('a)" and anz: "a mod 2 ^ n \ 0" shows "a div 2 ^ n < of_nat k" proof - have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using xk kv sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst unat_power_lower, simp) apply (subst mult.commute) apply (rule nat_less_power_trans) apply simp apply simp done have "unat a div 2 ^ n * 2 ^ n \ unat a" proof - have "unat a = unat a div 2 ^ n * 2 ^ n + unat a mod 2 ^ n" by (simp add: div_mult_mod_eq) also have "\ \ unat a div 2 ^ n * 2 ^ n" using sz anz by (simp add: unat_arith_simps) finally show ?thesis .. qed then have "a div 2 ^ n * 2 ^ n < a" using sz anz apply (subst word_less_nat_alt) apply (subst unat_word_ariths) apply (subst unat_div) apply simp apply (rule order_le_less_trans [OF mod_less_eq_dividend]) apply (erule order_le_neq_trans [OF div_mult_le]) done also from xk le have "\ \ of_nat k * 2 ^ n" by (simp add: field_simps) finally show ?thesis using sz kv apply - apply (erule word_mult_less_dest [OF _ _ kn]) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply (rule unat_lt2p) done qed lemma mask_out_sub_mask: "(x AND NOT (mask n)) = x - (x AND (mask n))" for x :: \'a::len word\ by (simp add: field_simps word_plus_and_or_coroll2) lemma subtract_mask: "p - (p AND mask n) = (p AND NOT (mask n))" "p - (p AND NOT (mask n)) = (p AND mask n)" for p :: \'a::len word\ by (simp add: field_simps word_plus_and_or_coroll2)+ lemma take_bit_word_eq_self_iff: \take_bit n w = w \ n \ LENGTH('a) \ w < 2 ^ n\ for w :: \'a::len word\ using take_bit_int_eq_self_iff [of n \take_bit LENGTH('a) (uint w)\] by (transfer fixing: n) auto lemma word_power_increasing: assumes x: "2 ^ x < (2 ^ y::'a::len word)" "x < LENGTH('a::len)" "y < LENGTH('a::len)" shows "x < y" using x using assms by transfer simp lemma mask_twice: "(x AND mask n) AND mask m = x AND mask (min m n)" for x :: \'a::len word\ by (simp flip: take_bit_eq_mask) lemma plus_one_helper[elim!]: "x < n + (1 :: 'a :: len word) \ x \ n" apply (simp add: word_less_nat_alt word_le_nat_alt field_simps) apply (case_tac "1 + n = 0") apply simp_all apply (subst(asm) unatSuc, assumption) apply arith done lemma plus_one_helper2: "\ x \ n; n + 1 \ 0 \ \ x < n + (1 :: 'a :: len word)" by (simp add: word_less_nat_alt word_le_nat_alt field_simps unatSuc) lemma less_x_plus_1: fixes x :: "'a :: len word" shows "x \ - 1 \ (y < (x + 1)) = (y < x \ y = x)" apply (rule iffI) apply (rule disjCI) apply (drule plus_one_helper) apply simp apply (subgoal_tac "x < x + 1") apply (erule disjE, simp_all) apply (rule plus_one_helper2 [OF order_refl]) apply (rule notI, drule max_word_wrap) apply simp done lemma word_Suc_leq: fixes k::"'a::len word" shows "k \ - 1 \ x < k + 1 \ x \ k" using less_x_plus_1 word_le_less_eq by auto lemma word_Suc_le: fixes k::"'a::len word" shows "x \ - 1 \ x + 1 \ k \ x < k" by (meson not_less word_Suc_leq) lemma word_lessThan_Suc_atMost: \{..< k + 1} = {..k}\ if \k \ - 1\ for k :: \'a::len word\ using that by (simp add: lessThan_def atMost_def word_Suc_leq) lemma word_atLeastLessThan_Suc_atLeastAtMost: \{l ..< u + 1} = {l..u}\ if \u \ - 1\ for l :: \'a::len word\ using that by (simp add: atLeastAtMost_def atLeastLessThan_def word_lessThan_Suc_atMost) lemma word_atLeastAtMost_Suc_greaterThanAtMost: \{m<..u} = {m + 1..u}\ if \m \ - 1\ for m :: \'a::len word\ using that by (simp add: greaterThanAtMost_def greaterThan_def atLeastAtMost_def atLeast_def word_Suc_le) lemma word_atLeastLessThan_Suc_atLeastAtMost_union: fixes l::"'a::len word" assumes "m \ - 1" and "l \ m" and "m \ u" shows "{l..m} \ {m+1..u} = {l..u}" proof - from ivl_disj_un_two(8)[OF assms(2) assms(3)] have "{l..u} = {l..m} \ {m<..u}" by blast with assms show ?thesis by(simp add: word_atLeastAtMost_Suc_greaterThanAtMost) qed lemma max_word_less_eq_iff [simp]: \- 1 \ w \ w = - 1\ for w :: \'a::len word\ by (fact word_order.extremum_unique) lemma word_or_zero: "(a OR b = 0) = (a = 0 \ b = 0)" for a b :: \'a::len word\ by (fact or_eq_0_iff) lemma word_2p_mult_inc: assumes x: "2 * 2 ^ n < (2::'a::len word) * 2 ^ m" assumes suc_n: "Suc n < LENGTH('a::len)" shows "2^n < (2::'a::len word)^m" by (smt suc_n le_less_trans lessI nat_less_le nat_mult_less_cancel_disj p2_gt_0 power_Suc power_Suc unat_power_lower word_less_nat_alt x) lemma power_overflow: "n \ LENGTH('a) \ 2 ^ n = (0 :: 'a::len word)" by simp lemmas extra_sle_sless_unfolds [simp] = word_sle_eq[where a=0 and b=1] word_sle_eq[where a=0 and b="numeral n"] word_sle_eq[where a=1 and b=0] word_sle_eq[where a=1 and b="numeral n"] word_sle_eq[where a="numeral n" and b=0] word_sle_eq[where a="numeral n" and b=1] word_sless_alt[where a=0 and b=1] word_sless_alt[where a=0 and b="numeral n"] word_sless_alt[where a=1 and b=0] word_sless_alt[where a=1 and b="numeral n"] word_sless_alt[where a="numeral n" and b=0] word_sless_alt[where a="numeral n" and b=1] for n lemma word_sint_1: "sint (1::'a::len word) = (if LENGTH('a) = 1 then -1 else 1)" by (fact signed_1) lemma ucast_of_nat: "is_down (ucast :: 'a :: len word \ 'b :: len word) \ ucast (of_nat n :: 'a word) = (of_nat n :: 'b word)" by transfer simp lemma scast_1': "(scast (1::'a::len word) :: 'b::len word) = (word_of_int (signed_take_bit (LENGTH('a::len) - Suc 0) (1::int)))" by transfer simp lemma scast_1: "(scast (1::'a::len word) :: 'b::len word) = (if LENGTH('a) = 1 then -1 else 1)" by (fact signed_1) lemma unat_minus_one_word: "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" apply (simp only: flip: mask_eq_exp_minus_1) apply transfer apply (simp add: take_bit_minus_one_eq_mask nat_mask_eq) done lemmas word_diff_ls'' = word_diff_ls [where xa=x and x=x for x] lemmas word_diff_ls' = word_diff_ls'' [simplified] lemmas word_l_diffs' = word_l_diffs [where xa=x and x=x for x] lemmas word_l_diffs = word_l_diffs' [simplified] lemma two_power_increasing: "\ n \ m; m < LENGTH('a) \ \ (2 :: 'a :: len word) ^ n \ 2 ^ m" by (simp add: word_le_nat_alt) lemma word_leq_le_minus_one: "\ x \ y; x \ 0 \ \ x - 1 < (y :: 'a :: len word)" apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst unat_minus_one) apply assumption apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma neg_mask_combine: "NOT(mask a) AND NOT(mask b) = NOT(mask (max a b) :: 'a::len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma neg_mask_twice: "x AND NOT(mask n) AND NOT(mask m) = x AND NOT(mask (max n m))" for x :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma multiple_mask_trivia: "n \ m \ (x AND NOT(mask n)) + (x AND mask n AND NOT(mask m)) = x AND NOT(mask m)" for x :: \'a::len word\ apply (rule trans[rotated], rule_tac w="mask n" in word_plus_and_or_coroll2) apply (simp add: word_bw_assocs word_bw_comms word_bw_lcs neg_mask_twice max_absorb2) done lemma word_of_nat_less: "\ n < unat x \ \ of_nat n < x" apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) - apply (simp add: take_bit_eq_mod) + apply (simp add: unsigned_of_nat take_bit_eq_mod) done lemma unat_mask: "unat (mask n :: 'a :: len word) = 2 ^ (min n (LENGTH('a))) - 1" apply (subst min.commute) apply (simp add: mask_eq_decr_exp not_less min_def split: if_split_asm) apply (intro conjI impI) apply (simp add: unat_sub_if_size) apply (simp add: power_overflow word_size) apply (simp add: unat_sub_if_size) done lemma mask_over_length: "LENGTH('a) \ n \ mask n = (-1::'a::len word)" by (simp add: mask_eq_decr_exp) lemma Suc_2p_unat_mask: "n < LENGTH('a) \ Suc (2 ^ n * k + unat (mask n :: 'a::len word)) = 2 ^ n * (k+1)" by (simp add: unat_mask) lemma sint_of_nat_ge_zero: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) \ 0" - by (simp add: bit_iff_odd) + by (simp add: bit_iff_odd signed_of_nat) lemma int_eq_sint: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) = int x" apply transfer apply (rule signed_take_bit_int_eq_self) apply simp_all apply (metis negative_zle numeral_power_eq_of_nat_cancel_iff) done lemma sint_of_nat_le: "\ b < 2 ^ (LENGTH('a) - 1); a \ b \ \ sint (of_nat a :: 'a :: len word) \ sint (of_nat b :: 'a :: len word)" apply (cases \LENGTH('a)\) apply simp_all apply transfer apply (subst signed_take_bit_eq_if_positive) apply (simp add: bit_simps) apply (metis bit_take_bit_iff nat_less_le order_less_le_trans take_bit_nat_eq_self_iff) apply (subst signed_take_bit_eq_if_positive) apply (simp add: bit_simps) apply (metis bit_take_bit_iff nat_less_le take_bit_nat_eq_self_iff) apply (simp flip: of_nat_take_bit add: take_bit_nat_eq_self) done lemma word_le_not_less: "((b::'a::len word) \ a) = (\(a < b))" by fastforce lemma less_is_non_zero_p1: fixes a :: "'a :: len word" shows "a < k \ a + 1 \ 0" apply (erule contrapos_pn) apply (drule max_word_wrap) apply (simp add: not_less) done lemma unat_add_lem': "(unat x + unat y < 2 ^ LENGTH('a)) \ (unat (x + y :: 'a :: len word) = unat x + unat y)" by (subst unat_add_lem[symmetric], assumption) lemma word_less_two_pow_divI: "\ (x :: 'a::len word) < 2 ^ (n - m); m \ n; n < LENGTH('a) \ \ x < 2 ^ n div 2 ^ m" apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (subst mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (simp add: power_sub) done lemma word_less_two_pow_divD: "\ (x :: 'a::len word) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (cases "n < LENGTH('a)") apply (cases "m < LENGTH('a)") apply (simp add: word_less_nat_alt) apply (subst(asm) unat_word_ariths) apply (subst(asm) mod_less) apply (rule order_le_less_trans [OF div_le_dividend]) apply (rule unat_lt2p) apply (clarsimp dest!: less_two_pow_divD) apply (simp add: power_overflow) apply (simp add: word_div_def) apply (simp add: power_overflow word_div_def) done lemma of_nat_less_two_pow_div_set: "\ n < LENGTH('a) \ \ {x. x < (2 ^ n div 2 ^ m :: 'a::len word)} = of_nat ` {k. k < 2 ^ n div 2 ^ m}" apply (simp add: image_def) apply (safe dest!: word_less_two_pow_divD less_two_pow_divD intro!: word_less_two_pow_divI) apply (rule_tac x="unat x" in exI) apply (simp add: power_sub[symmetric]) apply (subst unat_power_lower[symmetric, where 'a='a]) apply simp apply (erule unat_mono) apply (subst word_unat_power) apply (rule of_nat_mono_maybe) apply (rule power_strict_increasing) apply simp apply simp apply assumption done lemma ucast_less: "LENGTH('b) < LENGTH('a) \ (ucast (x :: 'b :: len word) :: ('a :: len word)) < 2 ^ LENGTH('b)" by transfer simp lemma ucast_range_less: "LENGTH('a :: len) < LENGTH('b :: len) \ range (ucast :: 'a word \ 'b word) = {x. x < 2 ^ len_of TYPE ('a)}" apply safe apply (erule ucast_less) apply (simp add: image_def) apply (rule_tac x="ucast x" in exI) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) - apply (metis bit_take_bit_iff less_mask_eq not_less take_bit_eq_mask) + apply (metis bit_take_bit_iff take_bit_word_eq_self_iff) done lemma word_power_less_diff: "\2 ^ n * q < (2::'a::len word) ^ m; q < 2 ^ (LENGTH('a) - n)\ \ q < 2 ^ (m - n)" apply (case_tac "m \ LENGTH('a)") apply (simp add: power_overflow) apply (case_tac "n \ LENGTH('a)") apply (simp add: power_overflow) apply (cases "n = 0") apply simp apply (subst word_less_nat_alt) apply (subst unat_power_lower) apply simp apply (rule nat_power_less_diff) apply (simp add: word_less_nat_alt) apply (subst (asm) iffD1 [OF unat_mult_lem]) apply (simp add:nat_less_power_trans) apply simp done lemma word_less_sub_1: "x < (y :: 'a :: len word) \ x \ y - 1" by (fact word_le_minus_one_leq) lemma word_sub_mono2: "\ a + b \ c + d; c \ a; b \ a + b; d \ c + d \ \ b \ (d :: 'a :: len word)" apply (drule(1) word_sub_mono) apply simp apply simp apply simp done lemma word_not_le: "(\ x \ (y :: 'a :: len word)) = (y < x)" by fastforce lemma word_subset_less: "\ {x .. x + r - 1} \ {y .. y + s - 1}; x \ x + r - 1; y \ y + (s :: 'a :: len word) - 1; s \ 0 \ \ r \ s" apply (frule subsetD[where c=x]) apply simp apply (drule subsetD[where c="x + r - 1"]) apply simp apply (clarsimp simp: add_diff_eq[symmetric]) apply (drule(1) word_sub_mono2) apply (simp_all add: olen_add_eqv[symmetric]) apply (erule word_le_minus_cancel) apply (rule ccontr) apply (simp add: word_not_le) done lemma uint_power_lower: "n < LENGTH('a) \ uint (2 ^ n :: 'a :: len word) = (2 ^ n :: int)" by (rule uint_2p_alt) lemma power_le_mono: "\2 ^ n \ (2::'a::len word) ^ m; n < LENGTH('a); m < LENGTH('a)\ \ n \ m" apply (clarsimp simp add: le_less) apply safe apply (simp add: word_less_nat_alt) apply (simp only: uint_arith_simps(3)) apply (drule uint_power_lower)+ apply simp done lemma two_power_eq: "\n < LENGTH('a); m < LENGTH('a)\ \ ((2::'a::len word) ^ n = 2 ^ m) = (n = m)" apply safe apply (rule order_antisym) apply (simp add: power_le_mono[where 'a='a])+ done lemma unat_less_helper: "x < of_nat n \ unat x < n" apply (simp add: word_less_nat_alt) apply (erule order_less_le_trans) - apply (simp add: take_bit_eq_mod) + apply (simp add: take_bit_eq_mod unsigned_of_nat) done lemma nat_uint_less_helper: "nat (uint y) = z \ x < y \ nat (uint x) < z" apply (erule subst) apply (subst unat_eq_nat_uint [symmetric]) apply (subst unat_eq_nat_uint [symmetric]) by (simp add: unat_mono) lemma of_nat_0: "\of_nat n = (0::'a::len word); n < 2 ^ LENGTH('a)\ \ n = 0" - by transfer (simp add: take_bit_eq_mod) + by (auto simp add: word_of_nat_eq_0_iff) lemma of_nat_inj: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)" by (metis unat_of_nat_len) lemma div_to_mult_word_lt: "\ (x :: 'a :: len word) \ y div z \ \ x * z \ y" apply (cases "z = 0") apply simp apply (simp add: word_neq_0_conv) apply (rule order_trans) apply (erule(1) word_mult_le_mono1) apply (simp add: unat_div) apply (rule order_le_less_trans [OF div_mult_le]) apply simp apply (rule word_div_mult_le) done lemma ucast_ucast_mask: "(ucast :: 'a :: len word \ 'b :: len word) (ucast x) = x AND mask (len_of TYPE ('a))" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: ac_simps) done lemma ucast_ucast_len: "\ x < 2 ^ LENGTH('b) \ \ ucast (ucast x::'b::len word) = (x::'a::len word)" apply (subst ucast_ucast_mask) apply (erule less_mask_eq) done lemma ucast_ucast_id: "LENGTH('a) < LENGTH('b) \ ucast (ucast (x::'a::len word)::'b::len word) = x" by (auto intro: ucast_up_ucast_id simp: is_up_def source_size_def target_size_def word_size) lemma unat_ucast: "unat (ucast x :: ('a :: len) word) = unat x mod 2 ^ (LENGTH('a))" proof - have \2 ^ LENGTH('a) = nat (2 ^ LENGTH('a))\ by simp moreover have \unat (ucast x :: 'a word) = unat x mod nat (2 ^ LENGTH('a))\ by transfer (simp flip: nat_mod_distrib take_bit_eq_mod) ultimately show ?thesis by (simp only:) qed lemma ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x < ((ucast (y :: 'a::len word)) :: 'b::len word)) = (x < y)" apply (simp add: word_less_nat_alt unat_ucast) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply (subst mod_less) apply(rule less_le_trans[OF unat_lt2p], simp) apply simp done \ \This weaker version was previously called @{text ucast_less_ucast}. We retain it to support existing proofs.\ lemmas ucast_less_ucast_weak = ucast_less_ucast[OF order.strict_implies_order] lemma unat_Suc2: fixes n :: "'a :: len word" shows "n \ -1 \ unat (n + 1) = Suc (unat n)" apply (subst add.commute, rule unatSuc) apply (subst eq_diff_eq[symmetric], simp add: minus_equation_iff) done lemma word_div_1: "(n :: 'a :: len word) div 1 = n" by (fact bits_div_by_1) lemma word_minus_one_le: "-1 \ (x :: 'a :: len word) = (x = -1)" by (fact word_order.extremum_unique) lemma up_scast_inj: "\ scast x = (scast y :: 'b :: len word); size x \ LENGTH('b) \ \ x = y" apply transfer apply (cases \LENGTH('a)\) apply simp_all apply (metis order_refl take_bit_signed_take_bit take_bit_tightened) done lemma up_scast_inj_eq: "LENGTH('a) \ len_of TYPE ('b) \ (scast x = (scast y::'b::len word)) = (x = (y::'a::len word))" by (fastforce dest: up_scast_inj simp: word_size) lemma word_le_add: fixes x :: "'a :: len word" shows "x \ y \ \n. y = x + of_nat n" by (rule exI [where x = "unat (y - x)"]) simp lemma word_plus_mcs_4': fixes x :: "'a :: len word" shows "\x + v \ x + w; x \ x + v\ \ v \ w" apply (rule word_plus_mcs_4) apply (simp add: add.commute) apply (simp add: add.commute) done lemma unat_eq_1: \unat x = Suc 0 \ x = 1\ by (auto intro!: unsigned_word_eqI [where ?'a = nat]) lemma word_unat_Rep_inject1: \unat x = unat 1 \ x = 1\ by (simp add: unat_eq_1) lemma and_not_mask_twice: "(w AND NOT (mask n)) AND NOT (mask m) = w AND NOT (mask (max m n))" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma word_less_cases: "x < y \ x = y - 1 \ x < y - (1 ::'a::len word)" apply (drule word_less_sub_1) apply (drule order_le_imp_less_or_eq) apply auto done lemma mask_and_mask: "mask a AND mask b = (mask (min a b) :: 'a::len word)" by (simp flip: take_bit_eq_mask ac_simps) lemma mask_eq_0_eq_x: "(x AND w = 0) = (x AND NOT w = x)" for x w :: \'a::len word\ using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma mask_eq_x_eq_0: "(x AND w = x) = (x AND NOT w = 0)" for x w :: \'a::len word\ using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma compl_of_1: "NOT 1 = (-2 :: 'a :: len word)" by (fact not_one_eq) lemma split_word_eq_on_mask: "(x = y) = (x AND m = y AND m \ x AND NOT m = y AND NOT m)" for x y m :: \'a::len word\ apply transfer apply (simp add: bit_eq_iff) apply (auto simp add: bit_simps ac_simps) done lemma word_FF_is_mask: "0xFF = (mask 8 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma word_1FF_is_mask: "0x1FF = (mask 9 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma ucast_of_nat_small: "x < 2 ^ LENGTH('a) \ ucast (of_nat x :: 'a :: len word) = (of_nat x :: 'b :: len word)" apply transfer apply (auto simp add: take_bit_of_nat min_def not_le) apply (metis linorder_not_less min_def take_bit_nat_eq_self take_bit_take_bit) done lemma word_le_make_less: fixes x :: "'a :: len word" shows "y \ -1 \ (x \ y) = (x < (y + 1))" apply safe apply (erule plus_one_helper2) apply (simp add: eq_diff_eq[symmetric]) done lemmas finite_word = finite [where 'a="'a::len word"] lemma word_to_1_set: "{0 ..< (1 :: 'a :: len word)} = {0}" by fastforce lemma word_leq_minus_one_le: fixes x :: "'a::len word" shows "\y \ 0; x \ y - 1 \ \ x < y" using le_m1_iff_lt word_neq_0_conv by blast lemma word_count_from_top: "n \ 0 \ {0 ..< n :: 'a :: len word} = {0 ..< n - 1} \ {n - 1}" apply (rule set_eqI, rule iffI) apply simp apply (drule word_le_minus_one_leq) apply (rule disjCI) apply simp apply simp apply (erule word_leq_minus_one_le) apply fastforce done lemma word_minus_one_le_leq: "\ x - 1 < y \ \ x \ (y :: 'a :: len word)" apply (cases "x = 0") apply simp apply (simp add: word_less_nat_alt word_le_nat_alt) apply (subst(asm) unat_minus_one) apply (simp add: word_less_nat_alt) apply (cases "unat x") apply (simp add: unat_eq_zero) apply arith done lemma word_div_less: "m < n \ m div n = 0" for m :: "'a :: len word" by (simp add: unat_mono word_arith_nat_defs(6)) lemma word_must_wrap: "\ x \ n - 1; n \ x \ \ n = (0 :: 'a :: len word)" using dual_order.trans sub_wrap word_less_1 by blast lemma range_subset_card: "\ {a :: 'a :: len word .. b} \ {c .. d}; b \ a \ \ d \ c \ d - c \ b - a" using word_sub_le word_sub_mono by fastforce lemma less_1_simp: "n - 1 < m = (n \ (m :: 'a :: len word) \ n \ 0)" by unat_arith lemma word_power_mod_div: fixes x :: "'a::len word" shows "\ n < LENGTH('a); m < LENGTH('a)\ \ x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" apply (simp add: word_arith_nat_div unat_mod power_mod_div) apply (subst unat_arith_simps(3)) apply (subst unat_mod) apply (subst unat_of_nat)+ apply (simp add: mod_mod_power min.commute) done lemma word_range_minus_1': fixes a :: "'a :: len word" shows "a \ 0 \ {a - 1<..b} = {a..b}" by (simp add: greaterThanAtMost_def atLeastAtMost_def greaterThan_def atLeast_def less_1_simp) lemma word_range_minus_1: fixes a :: "'a :: len word" shows "b \ 0 \ {a..b - 1} = {a.. 'b :: len word) x" by transfer simp lemma overflow_plus_one_self: "(1 + p \ p) = (p = (-1 :: 'a :: len word))" apply rule apply (rule ccontr) apply (drule plus_one_helper2) apply (rule notI) apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply (simp add: field_simps) apply simp done lemma plus_1_less: "(x + 1 \ (x :: 'a :: len word)) = (x = -1)" apply (rule iffI) apply (rule ccontr) apply (cut_tac plus_one_helper2[where x=x, OF order_refl]) apply simp apply clarsimp apply (drule arg_cong[where f="\x. x - 1"]) apply simp apply simp done lemma pos_mult_pos_ge: "[|x > (0::int); n>=0 |] ==> n * x >= n*1" apply (simp only: mult_left_mono) done lemma word_plus_strict_mono_right: fixes x :: "'a :: len word" shows "\y < z; x \ x + z\ \ x + y < x + z" by unat_arith lemma word_div_mult: "0 < c \ a < b * c \ a div c < b" for a b c :: "'a::len word" by (rule classical) (use div_to_mult_word_lt [of b a c] in \auto simp add: word_less_nat_alt word_le_nat_alt unat_div\) lemma word_less_power_trans_ofnat: "\n < 2 ^ (m - k); k \ m; m < LENGTH('a)\ \ of_nat n * 2 ^ k < (2::'a::len word) ^ m" apply (subst mult.commute) apply (rule word_less_power_trans) - apply (simp_all add: word_less_nat_alt less_le_trans take_bit_eq_mod) + apply (simp_all add: word_less_nat_alt unsigned_of_nat) + using take_bit_nat_less_eq_self + apply (rule le_less_trans) + apply assumption done lemma word_1_le_power: "n < LENGTH('a) \ (1 :: 'a :: len word) \ 2 ^ n" by (rule inc_le[where i=0, simplified], erule iffD2[OF p2_gt_0]) lemma unat_1_0: "1 \ (x::'a::len word) = (0 < unat x)" by (auto simp add: word_le_nat_alt) lemma x_less_2_0_1': fixes x :: "'a::len word" shows "\LENGTH('a) \ 1; x < 2\ \ x = 0 \ x = 1" apply (cases \2 \ LENGTH('a)\) apply simp_all apply transfer apply auto apply (metis add.commute add.right_neutral even_two_times_div_two mod_div_trivial mod_pos_pos_trivial mult.commute mult_zero_left not_less not_take_bit_negative odd_two_times_div_two_succ) done lemmas word_add_le_iff2 = word_add_le_iff [folded no_olen_add_nat] lemma of_nat_power: shows "\ p < 2 ^ x; x < len_of TYPE ('a) \ \ of_nat p < (2 :: 'a :: len word) ^ x" apply (rule order_less_le_trans) apply (rule of_nat_mono_maybe) apply (erule power_strict_increasing) apply simp apply assumption apply (simp add: word_unat_power del: of_nat_power) done lemma of_nat_n_less_equal_power_2: "n < LENGTH('a::len) \ ((of_nat n)::'a word) < 2 ^ n" apply (induct n) apply clarsimp apply clarsimp apply (metis of_nat_power n_less_equal_power_2 of_nat_Suc power_Suc) done lemma eq_mask_less: fixes w :: "'a::len word" assumes eqm: "w = w AND mask n" and sz: "n < len_of TYPE ('a)" shows "w < (2::'a word) ^ n" by (subst eqm, rule and_mask_less' [OF sz]) lemma of_nat_mono_maybe': fixes Y :: "nat" assumes xlt: "x < 2 ^ len_of TYPE ('a)" assumes ylt: "y < 2 ^ len_of TYPE ('a)" shows "(y < x) = (of_nat y < (of_nat x :: 'a :: len word))" apply (subst word_less_nat_alt) apply (subst unat_of_nat)+ apply (subst mod_less) apply (rule ylt) apply (subst mod_less) apply (rule xlt) apply simp done lemma of_nat_mono_maybe_le: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (y \ x) = ((of_nat y :: 'a :: len word) \ of_nat x)" apply (clarsimp simp: le_less) apply (rule disj_cong) apply (rule of_nat_mono_maybe', assumption+) apply auto using of_nat_inj apply blast done lemma mask_AND_NOT_mask: "(w AND NOT (mask n)) AND mask n = 0" for w :: \'a::len word\ by (rule bit_word_eqI) (simp add: bit_simps) lemma AND_NOT_mask_plus_AND_mask_eq: "(w AND NOT (mask n)) + (w AND mask n) = w" for w :: \'a::len word\ apply (subst disjunctive_add) apply (auto simp add: bit_simps) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma mask_eqI: fixes x :: "'a :: len word" assumes m1: "x AND mask n = y AND mask n" and m2: "x AND NOT (mask n) = y AND NOT (mask n)" shows "x = y" proof - have *: \x = x AND mask n OR x AND NOT (mask n)\ for x :: \'a word\ by (rule bit_word_eqI) (auto simp add: bit_simps) from assms * [of x] * [of y] show ?thesis by simp qed lemma neq_0_no_wrap: fixes x :: "'a :: len word" shows "\ x \ x + y; x \ 0 \ \ x + y \ 0" by clarsimp lemma unatSuc2: fixes n :: "'a :: len word" shows "n + 1 \ 0 \ unat (n + 1) = Suc (unat n)" by (simp add: add.commute unatSuc) lemma word_of_nat_le: "n \ unat x \ of_nat n \ x" apply (simp add: word_le_nat_alt unat_of_nat) apply (erule order_trans[rotated]) apply (simp add: take_bit_eq_mod) done lemma word_unat_less_le: "a \ of_nat b \ unat a \ b" by (metis eq_iff le_cases le_unat_uoi word_of_nat_le) lemma mask_Suc_0 : "mask (Suc 0) = (1 :: 'a::len word)" by (simp add: mask_eq_decr_exp) lemma bool_mask': fixes x :: "'a :: len word" shows "2 < LENGTH('a) \ (0 < x AND 1) = (x AND 1 = 1)" by (simp add: and_one_eq mod_2_eq_odd) lemma ucast_ucast_add: fixes x :: "'a :: len word" fixes y :: "'b :: len word" shows "LENGTH('b) \ LENGTH('a) \ ucast (ucast x + y) = x + ucast y" apply transfer apply simp apply (subst (2) take_bit_add [symmetric]) apply (subst take_bit_add [symmetric]) apply simp done lemma lt1_neq0: fixes x :: "'a :: len word" shows "(1 \ x) = (x \ 0)" by unat_arith lemma word_plus_one_nonzero: fixes x :: "'a :: len word" shows "\x \ x + y; y \ 0\ \ x + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (erule word_random) apply (simp add: lt1_neq0) done lemma word_sub_plus_one_nonzero: fixes n :: "'a :: len word" shows "\n' \ n; n' \ 0\ \ (n - n') + 1 \ 0" apply (subst lt1_neq0 [symmetric]) apply (subst olen_add_eqv [symmetric]) apply (rule word_random [where x' = n']) apply simp apply (erule word_sub_le) apply (simp add: lt1_neq0) done lemma word_le_minus_mono_right: fixes x :: "'a :: len word" shows "\ z \ y; y \ x; z \ x \ \ x - y \ x - z" apply (rule word_sub_mono) apply simp apply assumption apply (erule word_sub_le) apply (erule word_sub_le) done lemma word_0_sle_from_less: \0 \s x\ if \x < 2 ^ (LENGTH('a) - 1)\ for x :: \'a::len word\ using that apply transfer apply (cases \LENGTH('a)\) apply simp_all apply (metis bit_take_bit_iff min_def nat_less_le not_less_eq take_bit_int_eq_self_iff take_bit_take_bit) done lemma ucast_sub_ucast: fixes x :: "'a::len word" assumes "y \ x" assumes T: "LENGTH('a) \ LENGTH('b)" shows "ucast (x - y) = (ucast x - ucast y :: 'b::len word)" proof - from T have P: "unat x < 2 ^ LENGTH('b)" "unat y < 2 ^ LENGTH('b)" by (fastforce intro!: less_le_trans[OF unat_lt2p])+ then show ?thesis by (simp add: unat_arith_simps unat_ucast assms[simplified unat_arith_simps]) qed lemma word_1_0: "\a + (1::('a::len) word) \ b; a < of_nat x\ \ a < b" apply transfer apply (subst (asm) take_bit_incr_eq) apply (auto simp add: diff_less_eq) using take_bit_int_less_exp le_less_trans by blast lemma unat_of_nat_less:"\ a < b; unat b = c \ \ a < of_nat c" by fastforce lemma word_le_plus_1: "\ (y::('a::len) word) < y + n; a < n \ \ y + a \ y + a + 1" by unat_arith lemma word_le_plus:"\(a::('a::len) word) < a + b; c < b\ \ a \ a + c" by (metis order_less_imp_le word_random) lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)" apply (cases \LENGTH('a)\) apply simp_all apply transfer apply (simp flip: signed_take_bit_eq_iff_take_bit_eq) done lemma sint_0 [simp]: "(sint x = 0) = (x = 0)" by (fact signed_eq_0_iff) (* It is not always that case that "sint 1 = 1", because of 1-bit word sizes. * This lemma produces the different cases. *) lemma sint_1_cases: P if \\ len_of TYPE ('a::len) = 1; (a::'a word) = 0; sint a = 0 \ \ P\ \\ len_of TYPE ('a) = 1; a = 1; sint (1 :: 'a word) = -1 \ \ P\ \\ len_of TYPE ('a) > 1; sint (1 :: 'a word) = 1 \ \ P\ proof (cases \LENGTH('a) = 1\) case True then have \a = 0 \ a = 1\ by transfer auto with True that show ?thesis by auto next case False with that show ?thesis by (simp add: less_le Suc_le_eq) qed lemma sint_int_min: "sint (- (2 ^ (LENGTH('a) - Suc 0)) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (cases \LENGTH('a)\) apply simp_all apply transfer apply (simp add: signed_take_bit_int_eq_self) done lemma sint_int_max_plus_1: "sint (2 ^ (LENGTH('a) - Suc 0) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))" apply (cases \LENGTH('a)\) apply simp_all apply (subst word_of_int_2p [symmetric]) apply (subst int_word_sint) apply simp done lemma uint_range': \0 \ uint x \ uint x < 2 ^ LENGTH('a)\ for x :: \'a::len word\ by transfer simp lemma sint_of_int_eq: "\ - (2 ^ (LENGTH('a) - 1)) \ x; x < 2 ^ (LENGTH('a) - 1) \ \ sint (of_int x :: ('a::len) word) = x" - by (simp add: signed_take_bit_int_eq_self) + by (simp add: signed_take_bit_int_eq_self signed_of_int) lemma of_int_sint: "of_int (sint a) = a" by simp lemma sint_ucast_eq_uint: "\ \ is_down (ucast :: ('a::len word \ 'b::len word)) \ \ sint ((ucast :: ('a::len word \ 'b::len word)) x) = uint x" apply transfer apply (simp add: signed_take_bit_take_bit) done lemma word_less_nowrapI': "(x :: 'a :: len word) \ z - k \ k \ z \ 0 < k \ x < x + k" by uint_arith lemma mask_plus_1: "mask n + 1 = (2 ^ n :: 'a::len word)" by (clarsimp simp: mask_eq_decr_exp) lemma unat_inj: "inj unat" by (metis eq_iff injI word_le_nat_alt) lemma unat_ucast_upcast: "is_up (ucast :: 'b word \ 'a word) \ unat (ucast x :: ('a::len) word) = unat (x :: ('b::len) word)" unfolding ucast_eq unat_eq_nat_uint apply transfer apply simp done lemma ucast_mono: "\ (x :: 'b :: len word) < y; y < 2 ^ LENGTH('a) \ \ ucast x < ((ucast y) :: 'a :: len word)" apply (simp only: flip: ucast_nat_def) apply (rule of_nat_mono_maybe) apply (rule unat_less_helper) apply simp apply (simp add: word_less_nat_alt) done lemma ucast_mono_le: "\x \ y; y < 2 ^ LENGTH('b)\ \ (ucast (x :: 'a :: len word) :: 'b :: len word) \ ucast y" apply (simp only: flip: ucast_nat_def) apply (subst of_nat_mono_maybe_le[symmetric]) apply (rule unat_less_helper) apply simp apply (rule unat_less_helper) apply (erule le_less_trans) apply (simp_all add: word_le_nat_alt) done lemma ucast_mono_le': "\ unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x \ y \ \ ucast x \ (ucast y :: 'b word)" for x y :: \'a::len word\ by (auto simp: word_less_nat_alt intro: ucast_mono_le) lemma neg_mask_add_mask: "((x:: 'a :: len word) AND NOT (mask n)) + (2 ^ n - 1) = x OR mask n" unfolding mask_2pm1 [symmetric] apply (subst word_plus_and_or_coroll; rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma le_step_down_word:"\(i::('a::len) word) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by unat_arith lemma le_step_down_word_2: fixes x :: "'a::len word" shows "\x \ y; x \ y\ \ x \ y - 1" by (subst (asm) word_le_less_eq, clarsimp, simp add: word_le_minus_one_leq) lemma NOT_mask_AND_mask[simp]: "(w AND mask n) AND NOT (mask n) = 0" by (rule bit_eqI) (simp add: bit_simps) lemma and_and_not[simp]:"(a AND b) AND NOT b = 0" for a b :: \'a::len word\ apply (subst word_bw_assocs(1)) apply clarsimp done lemma ex_mask_1[simp]: "(\x. mask x = (1 :: 'a::len word))" apply (rule_tac x=1 in exI) apply (simp add:mask_eq_decr_exp) done lemma not_switch:"NOT a = x \ a = NOT x" by auto lemma test_bit_eq_iff: "bit u = bit v \ u = v" for u v :: "'a::len word" by (auto intro: bit_eqI simp add: fun_eq_iff) lemma test_bit_size: "bit w n \ n < size w" for w :: "'a::len word" by transfer simp lemma word_eq_iff: "x = y \ (\n?P \ ?Q\) for x y :: "'a::len word" by transfer (auto simp add: bit_eq_iff bit_take_bit_iff) lemma word_eqI: "(\n. n < size u \ bit u n = bit v n) \ u = v" for u :: "'a::len word" by (simp add: word_size word_eq_iff) lemma word_eqD: "u = v \ bit u x = bit v x" for u v :: "'a::len word" by simp lemma test_bit_bin': "bit w n \ n < size w \ bit (uint w) n" by transfer (simp add: bit_take_bit_iff) lemmas test_bit_bin = test_bit_bin' [unfolded word_size] lemma word_test_bit_def: \bit a = bit (uint a)\ by transfer (simp add: fun_eq_iff bit_take_bit_iff) lemmas test_bit_def' = word_test_bit_def [THEN fun_cong] lemma word_test_bit_transfer [transfer_rule]: "(rel_fun pcr_word (rel_fun (=) (=))) (\x n. n < LENGTH('a) \ bit x n) (bit :: 'a::len word \ _)" by transfer_prover lemma test_bit_wi: "bit (word_of_int x :: 'a::len word) n \ n < LENGTH('a) \ bit x n" by transfer simp lemma word_ops_nth_size: "n < size x \ bit (x OR y) n = (bit x n | bit y n) \ bit (x AND y) n = (bit x n \ bit y n) \ bit (x XOR y) n = (bit x n \ bit y n) \ bit (NOT x) n = (\ bit x n)" for x :: "'a::len word" by transfer (simp add: bit_or_iff bit_and_iff bit_xor_iff bit_not_iff) lemma word_ao_nth: "bit (x OR y) n = (bit x n | bit y n) \ bit (x AND y) n = (bit x n \ bit y n)" for x :: "'a::len word" by transfer (auto simp add: bit_or_iff bit_and_iff) lemmas lsb0 = len_gt_0 [THEN word_ops_nth_size [unfolded word_size]] lemma nth_sint: fixes w :: "'a::len word" defines "l \ LENGTH('a)" shows "bit (sint w) n = (if n < l - 1 then bit w n else bit w (l - 1))" unfolding sint_uint l_def by (auto simp: bit_signed_take_bit_iff word_test_bit_def not_less min_def) lemma test_bit_2p: "bit (word_of_int (2 ^ n)::'a::len word) m \ m = n \ m < LENGTH('a)" by transfer (auto simp add: bit_exp_iff) lemma nth_w2p: "bit ((2::'a::len word) ^ n) m \ m = n \ m < LENGTH('a::len)" by transfer (auto simp add: bit_exp_iff) lemma bang_is_le: "bit x m \ 2 ^ m \ x" for x :: "'a::len word" apply (rule xtrans(3)) apply (rule_tac [2] y = "x" in le_word_or2) apply (rule word_eqI) apply (auto simp add: word_ao_nth nth_w2p word_size) done lemmas msb0 = len_gt_0 [THEN diff_Suc_less, THEN word_ops_nth_size [unfolded word_size]] lemmas msb1 = msb0 [where i = 0] lemma test_bit_1 [iff]: "bit (1 :: 'a::len word) n \ n = 0" by transfer (auto simp add: bit_1_iff) lemma nth_0: "\ bit (0 :: 'a::len word) n" by transfer simp lemma nth_minus1: "bit (-1 :: 'a::len word) n \ n < LENGTH('a)" by transfer simp lemma nth_ucast: "bit (ucast w::'a::len word) n = (bit w n \ n < LENGTH('a))" by transfer (simp add: bit_take_bit_iff ac_simps) lemma drop_bit_numeral_bit0_1 [simp]: \drop_bit (Suc 0) (numeral k) = (word_of_int (drop_bit (Suc 0) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ by (metis Word_eq_word_of_int drop_bit_word.abs_eq of_int_numeral) lemma nth_mask: \bit (mask n :: 'a::len word) i \ i < n \ i < size (mask n :: 'a word)\ by (auto simp add: word_size Word.bit_mask_iff) lemma nth_slice: "bit (slice n w :: 'a::len word) m = (bit w (m + n) \ m < LENGTH('a))" apply (auto simp add: bit_simps less_diff_conv dest: bit_imp_le_length) using bit_imp_le_length apply fastforce done lemma test_bit_cat [OF refl]: "wc = word_cat a b \ bit wc n = (n < size wc \ (if n < size b then bit b n else bit a (n - size b)))" apply (simp add: word_size not_less; transfer) apply (auto simp add: bit_concat_bit_iff bit_take_bit_iff) done \ \keep quantifiers for use in simplification\ lemma test_bit_split': "word_split c = (a, b) \ (\n m. bit b n = (n < size b \ bit c n) \ bit a m = (m < size a \ bit c (m + size b)))" by (auto simp add: word_split_bin' bit_unsigned_iff word_size bit_drop_bit_eq ac_simps dest: bit_imp_le_length) lemma test_bit_split: "word_split c = (a, b) \ (\n::nat. bit b n \ n < size b \ bit c n) \ (\m::nat. bit a m \ m < size a \ bit c (m + size b))" by (simp add: test_bit_split') lemma test_bit_split_eq: "word_split c = (a, b) \ ((\n::nat. bit b n = (n < size b \ bit c n)) \ (\m::nat. bit a m = (m < size a \ bit c (m + size b))))" apply (rule_tac iffI) apply (rule_tac conjI) apply (erule test_bit_split [THEN conjunct1]) apply (erule test_bit_split [THEN conjunct2]) apply (case_tac "word_split c") apply (frule test_bit_split) apply (erule trans) apply (fastforce intro!: word_eqI simp add: word_size) done lemma test_bit_rcat: "sw = size (hd wl) \ rc = word_rcat wl \ bit rc n = (n < size rc \ n div sw < size wl \ bit ((rev wl) ! (n div sw)) (n mod sw))" for wl :: "'a::len word list" by (simp add: word_size word_rcat_def rev_map bit_horner_sum_uint_exp_iff bit_simps not_le) lemmas test_bit_cong = arg_cong [where f = "bit", THEN fun_cong] lemma max_test_bit: "bit (- 1::'a::len word) n \ n < LENGTH('a)" by (fact nth_minus1) lemma map_nth_0 [simp]: "map (bit (0::'a::len word)) xs = replicate (length xs) False" by (simp flip: map_replicate_const) lemma word_and_1: "n AND 1 = (if bit n 0 then 1 else 0)" for n :: "_ word" by (rule bit_word_eqI) (auto simp add: bit_and_iff bit_1_iff intro: gr0I) lemma test_bit_1': "bit (1 :: 'a :: len word) n \ 0 < LENGTH('a) \ n = 0" by simp lemma nth_w2p_same: "bit (2^n :: 'a :: len word) n = (n < LENGTH('a))" by (simp add: nth_w2p) lemma word_leI: "(\n. \n < size (u::'a::len word); bit u n \ \ bit (v::'a::len word) n) \ u <= v" apply (rule order_trans [of u \u AND v\ v]) apply (rule eq_refl) apply (rule bit_word_eqI) apply (auto simp add: bit_simps word_and_le1 word_size) done lemma bang_eq: fixes x :: "'a::len word" shows "(x = y) = (\n. bit x n = bit y n)" by (auto intro!: bit_eqI) lemma neg_mask_test_bit: "bit (NOT(mask n) :: 'a :: len word) m = (n \ m \ m < LENGTH('a))" by (auto simp add: bit_simps) lemma upper_bits_unset_is_l2p: \(\n' \ n. n' < LENGTH('a) \ \ bit p n') \ (p < 2 ^ n)\ (is \?P \ ?Q\) if \n < LENGTH('a)\ for p :: "'a :: len word" proof assume ?Q then show ?P by (meson bang_is_le le_less_trans not_le word_power_increasing) next assume ?P have \take_bit n p = p\ proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ show \bit (take_bit n p) q \ bit p q\ proof (cases \q < n\) case True then show ?thesis by (auto simp add: bit_simps) next case False then have \n \ q\ by simp with \?P\ \q < LENGTH('a)\ have \\ bit p q\ by simp then show ?thesis by (simp add: bit_simps) qed qed with that show ?Q using take_bit_word_eq_self_iff [of n p] by auto qed lemma less_2p_is_upper_bits_unset: "p < 2 ^ n \ n < LENGTH('a) \ (\n' \ n. n' < LENGTH('a) \ \ bit p n')" for p :: "'a :: len word" by (meson le_less_trans le_mask_iff_lt_2n upper_bits_unset_is_l2p word_zero_le) lemma test_bit_over: "n \ size (x::'a::len word) \ (bit x n) = False" by transfer auto lemma le_mask_high_bits: "w \ mask n \ (\i \ {n ..< size w}. \ bit w i)" for w :: \'a::len word\ apply (auto simp add: bit_simps word_size less_eq_mask_iff_take_bit_eq_self) apply (metis bit_take_bit_iff leD) apply (metis atLeastLessThan_iff leI take_bit_word_eq_self_iff upper_bits_unset_is_l2p) done lemma test_bit_conj_lt: "(bit x m \ m < LENGTH('a)) = bit x m" for x :: "'a :: len word" using test_bit_bin by blast lemma neg_test_bit: "bit (NOT x) n = (\ bit x n \ n < LENGTH('a))" for x :: "'a::len word" by (cases "n < LENGTH('a)") (auto simp add: test_bit_over word_ops_nth_size word_size) lemma nth_bounded: "\bit (x :: 'a :: len word) n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" apply (rule ccontr) apply (auto simp add: not_less) apply (meson bit_imp_le_length bit_uint_iff less_2p_is_upper_bits_unset test_bit_bin) done lemma and_neq_0_is_nth: \x AND y \ 0 \ bit x n\ if \y = 2 ^ n\ for x y :: \'a::len word\ apply (simp add: bit_eq_iff bit_simps) using that apply (simp add: bit_simps not_le) apply transfer apply auto done lemma nth_is_and_neq_0: "bit (x::'a::len word) n = (x AND 2 ^ n \ 0)" by (subst and_neq_0_is_nth; rule refl) lemma max_word_not_less [simp]: "\ - 1 < x" for x :: \'a::len word\ by (fact word_order.extremum_strict) lemma bit_twiddle_min: "(y::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = min x y" by (rule bit_eqI) (auto simp add: bit_simps) lemma bit_twiddle_max: "(x::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = max x y" by (rule bit_eqI) (auto simp add: bit_simps max_def) lemma swap_with_xor: "\(x::'a::len word) = a XOR b; y = b XOR x; z = x XOR y\ \ z = b \ y = a" by (auto intro: bit_word_eqI simp add: bit_simps) lemma le_mask_imp_and_mask: "(x::'a::len word) \ mask n \ x AND mask n = x" by (metis and_mask_eq_iff_le_mask) lemma or_not_mask_nop: "((x::'a::len word) OR NOT (mask n)) AND mask n = x AND mask n" by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3)) lemma mask_subsume: "\n \ m\ \ ((x::'a::len word) OR y AND mask n) AND NOT (mask m) = x AND NOT (mask m)" by (rule bit_word_eqI) (auto simp add: bit_simps word_size) lemma and_mask_0_iff_le_mask: fixes w :: "'a::len word" shows "(w AND NOT(mask n) = 0) = (w \ mask n)" by (simp add: mask_eq_0_eq_x le_mask_imp_and_mask and_mask_eq_iff_le_mask) lemma mask_twice2: "n \ m \ ((x::'a::len word) AND mask m) AND mask n = x AND mask n" by (metis mask_twice min_def) lemma uint_2_id: "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" by simp lemma div_of_0_id[simp]:"(0::('a::len) word) div n = 0" by (simp add: word_div_def) lemma degenerate_word:"LENGTH('a) = 1 \ (x::('a::len) word) = 0 \ x = 1" by (metis One_nat_def less_irrefl_nat sint_1_cases) lemma div_by_0_word:"(x::('a::len) word) div 0 = 0" by (metis div_0 div_by_0 unat_0 word_arith_nat_defs(6) word_div_1) lemma div_less_dividend_word:"\x \ 0; n \ 1\ \ (x::('a::len) word) div n < x" apply (cases \n = 0\) apply clarsimp apply (simp add:word_neq_0_conv) apply (subst word_arith_nat_div) apply (rule word_of_nat_less) apply (rule div_less_dividend) using unat_eq_zero word_unat_Rep_inject1 apply force apply (simp add:unat_gt_0) done lemma word_less_div: fixes x :: "('a::len) word" and y :: "('a::len) word" shows "x div y = 0 \ y = 0 \ x < y" apply (case_tac "y = 0", clarsimp+) by (metis One_nat_def Suc_le_mono le0 le_div_geq not_less unat_0 unat_div unat_gt_0 word_less_nat_alt zero_less_one) lemma not_degenerate_imp_2_neq_0:"LENGTH('a) > 1 \ (2::('a::len) word) \ 0" by (metis numerals(1) power_not_zero power_zero_numeral) lemma word_overflow:"(x::('a::len) word) + 1 > x \ x + 1 = 0" apply clarsimp by (metis diff_0 eq_diff_eq less_x_plus_1) lemma word_overflow_unat:"unat ((x::('a::len) word) + 1) = unat x + 1 \ x + 1 = 0" by (metis Suc_eq_plus1 add.commute unatSuc) lemma even_word_imp_odd_next:"even (unat (x::('a::len) word)) \ x + 1 = 0 \ odd (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma odd_word_imp_even_next:"odd (unat (x::('a::len) word)) \ x + 1 = 0 \ even (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma overflow_imp_lsb:"(x::('a::len) word) + 1 = 0 \ bit x 0" using even_plus_one_iff [of x] by simp lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = bit x 0" by transfer (simp add: even_nat_iff) lemma of_nat_neq_iff_word: "x mod 2 ^ LENGTH('a) \ y mod 2 ^ LENGTH('a) \ (((of_nat x)::('a::len) word) \ of_nat y) = (x \ y)" apply (rule iffI) apply (case_tac "x = y") apply (subst (asm) of_nat_eq_iff[symmetric]) apply auto apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") apply auto apply (metis unat_of_nat) done lemma lsb_this_or_next: "\ (bit ((x::('a::len) word) + 1) 0) \ bit x 0" by simp lemma mask_or_not_mask: "x AND mask n OR x AND NOT (mask n) = x" for x :: \'a::len word\ apply (subst word_oa_dist, simp) apply (subst word_oa_dist2, simp) done lemma word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" by (metis add.commute add_minus_cancel) lemma revcast_down_us [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (signed_drop_bit n w)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_simps ac_simps) done lemma revcast_down_ss [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (signed_drop_bit n w)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_simps ac_simps) done lemma revcast_down_uu [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (drop_bit n w)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_simps ac_simps) done lemma revcast_down_su [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (drop_bit n w)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_simps ac_simps) done lemma cast_down_rev [OF refl]: "uc = ucast \ source_size uc = target_size uc + n \ uc w = revcast (push_bit n w)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_simps) done lemma revcast_up [OF refl]: "rc = revcast \ source_size rc + n = target_size rc \ rc w = push_bit n (ucast w :: 'a::len word)" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_simps) apply auto apply (metis add.commute add_diff_cancel_right) apply (metis diff_add_inverse2 diff_diff_add) done lemmas rc1 = revcast_up [THEN revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] lemmas rc2 = revcast_down_uu [THEN revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] lemma word_ops_nth: fixes x y :: \'a::len word\ shows word_or_nth: "bit (x OR y) n = (bit x n \ bit y n)" and word_and_nth: "bit (x AND y) n = (bit x n \ bit y n)" and word_xor_nth: "bit (x XOR y) n = (bit x n \ bit y n)" by (simp_all add: bit_simps) lemma word_power_nonzero: "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ \ x * 2 ^ n \ 0" by (metis gr_implies_not0 mult_eq_0_iff nat_mult_power_less_eq numeral_2_eq_2 p2_gt_0 unat_eq_zero unat_less_power unat_mult_lem unat_power_lower word_gt_a_gt_0 zero_less_Suc) lemma less_1_helper: "n \ m \ (n - 1 :: int) < m" by arith lemma div_power_helper: "\ x \ y; y < LENGTH('a) \ \ (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1" apply (simp flip: mask_eq_exp_minus_1 drop_bit_eq_div) apply (rule bit_word_eqI) apply (auto simp add: bit_simps not_le) done lemma max_word_mask: "(- 1 :: 'a::len word) = mask LENGTH('a)" by (fact minus_1_eq_mask) lemmas mask_len_max = max_word_mask[symmetric] lemma mask_out_first_mask_some: "\ x AND NOT (mask n) = y; n \ m \ \ x AND NOT (mask m) = y AND NOT (mask m)" for x y :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps word_size) lemma mask_lower_twice: "n \ m \ (x AND NOT (mask n)) AND NOT (mask m) = x AND NOT (mask m)" for x :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps word_size) lemma mask_lower_twice2: "(a AND NOT (mask n)) AND NOT (mask m) = a AND NOT (mask (max n m))" for a :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma ucast_and_neg_mask: "ucast (x AND NOT (mask n)) = ucast x AND NOT (mask n)" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma ucast_and_mask: "ucast (x AND mask n) = ucast x AND mask n" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma ucast_mask_drop: "LENGTH('a :: len) \ n \ (ucast (x AND mask n) :: 'a word) = ucast x" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma mask_exceed: "n \ LENGTH('a) \ (x::'a::len word) AND NOT (mask n) = 0" by (rule bit_word_eqI) (simp add: bit_simps) lemma word_add_no_overflow:"(x::'a::len word) < - 1 \ x < x + 1" using less_x_plus_1 order_less_le by blast lemma lt_plus_1_le_word: fixes x :: "'a::len word" assumes bound:"n < unat (maxBound::'a word)" shows "x < 1 + of_nat n = (x \ of_nat n)" by (metis add.commute bound max_word_max word_Suc_leq word_not_le word_of_nat_less) lemma unat_ucast_up_simp: fixes x :: "'a::len word" assumes "LENGTH('a) \ LENGTH('b)" shows "unat (ucast x :: 'b::len word) = unat x" apply (rule bit_eqI) using assms apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma unat_ucast_less_no_overflow: "\n < 2 ^ LENGTH('a); unat f < n\ \ (f::('a::len) word) < of_nat n" by (erule (1) order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) simp lemma unat_ucast_less_no_overflow_simp: "n < 2 ^ LENGTH('a) \ (unat f < n) = ((f::('a::len) word) < of_nat n)" using unat_less_helper unat_ucast_less_no_overflow by blast lemma unat_ucast_no_overflow_le: assumes no_overflow: "unat b < (2 :: nat) ^ LENGTH('a)" and upward_cast: "LENGTH('a) < LENGTH('b)" shows "(ucast (f::'a::len word) < (b :: 'b :: len word)) = (unat f < unat b)" proof - have LR: "ucast f < b \ unat f < unat b" apply (rule unat_less_helper) apply (simp add:ucast_nat_def) apply (rule_tac 'b1 = 'b in ucast_less_ucast[OF order.strict_implies_order, THEN iffD1]) apply (rule upward_cast) apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt unat_power_lower[OF upward_cast] no_overflow) done have RL: "unat f < unat b \ ucast f < b" proof- assume ineq: "unat f < unat b" have "ucast (f::'a::len word) < ((ucast (ucast b ::'a::len word)) :: 'b :: len word)" apply (simp add: ucast_less_ucast[OF order.strict_implies_order] upward_cast) apply (simp only: flip: ucast_nat_def) apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq]) done then show ?thesis apply (rule order_less_le_trans) apply (simp add:ucast_ucast_mask word_and_le2) done qed then show ?thesis by (simp add:RL LR iffI) qed lemmas ucast_up_mono = ucast_less_ucast[THEN iffD2] lemma minus_one_word: "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by simp lemma le_2p_upper_bits: "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ \n'\n. n' < LENGTH('a) \ \ bit p n'" by (subst upper_bits_unset_is_l2p; simp) lemma le2p_bits_unset: "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ bit (p::'a::len word) n'" using upper_bits_unset_is_l2p [where p=p] by (cases "n < LENGTH('a)") auto lemma complement_nth_w2p: shows "n' < LENGTH('a) \ bit (NOT (2 ^ n :: 'a::len word)) n' = (n' \ n)" by (fastforce simp: word_ops_nth_size word_size nth_w2p) lemma word_unat_and_lt: "unat x < n \ unat y < n \ unat (x AND y) < n" by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt) lemma word_unat_mask_lt: "m \ size w \ unat ((w::'a::len word) AND mask m) < 2 ^ m" by (rule word_unat_and_lt) (simp add: unat_mask word_size) lemma word_sless_sint_le:"x sint x \ sint y - 1" by (metis word_sless_alt zle_diff1_eq) lemma upper_trivial: fixes x :: "'a::len word" shows "x \ 2 ^ LENGTH('a) - 1 \ x < 2 ^ LENGTH('a) - 1" by (simp add: less_le) lemma constraint_expand: fixes x :: "'a::len word" shows "x \ {y. lower \ y \ y \ upper} = (lower \ x \ x \ upper)" by (rule mem_Collect_eq) lemma card_map_elide: "card ((of_nat :: nat \ 'a::len word) ` {0.. CARD('a::len word)" proof - let ?of_nat = "of_nat :: nat \ 'a word" have "inj_on ?of_nat {i. i < CARD('a word)}" by (rule inj_onI) (simp add: card_word of_nat_inj) moreover have "{0.. {i. i < CARD('a word)}" using that by auto ultimately have "inj_on ?of_nat {0.. CARD('a::len word) \ card ((of_nat::nat \ 'a::len word) ` {0.. LENGTH('a) \ x = ucast y \ ucast x = y" for x :: "'a::len word" and y :: "'b::len word" by transfer simp lemma le_ucast_ucast_le: "x \ ucast y \ ucast x \ y" for x :: "'a::len word" and y :: "'b::len word" by (smt le_unat_uoi linorder_not_less order_less_imp_le ucast_nat_def unat_arith_simps(1)) lemma less_ucast_ucast_less: "LENGTH('b) \ LENGTH('a) \ x < ucast y \ ucast x < y" for x :: "'a::len word" and y :: "'b::len word" by (metis ucast_nat_def unat_mono unat_ucast_up_simp word_of_nat_less) lemma ucast_le_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x \ (ucast y::'b::len word)) = (x \ y)" for x :: "'a::len word" by (simp add: unat_arith_simps(1) unat_ucast_up_simp) lemmas ucast_up_mono_le = ucast_le_ucast[THEN iffD2] lemma ucast_or_distrib: fixes x :: "'a::len word" fixes y :: "'a::len word" shows "(ucast (x OR y) :: ('b::len) word) = ucast x OR ucast y" by (fact unsigned_or_eq) lemma word_exists_nth: "(w::'a::len word) \ 0 \ \i. bit w i" by (auto simp add: bit_eq_iff) lemma max_word_not_0 [simp]: "- 1 \ (0 :: 'a::len word)" by simp lemma unat_max_word_pos[simp]: "0 < unat (- 1 :: 'a::len word)" using unat_gt_0 [of \- 1 :: 'a::len word\] by simp (* Miscellaneous conditional injectivity rules. *) lemma mult_pow2_inj: assumes ws: "m + n \ LENGTH('a)" assumes le: "x \ mask m" "y \ mask m" assumes eq: "x * 2 ^ n = y * (2 ^ n::'a::len word)" shows "x = y" proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ from eq have \push_bit n x = push_bit n y\ by (simp add: push_bit_eq_mult) moreover from le have \take_bit m x = x\ \take_bit m y = y\ by (simp_all add: less_eq_mask_iff_take_bit_eq_self) ultimately have \push_bit n (take_bit m x) = push_bit n (take_bit m y)\ by simp_all with \q < LENGTH('a)\ ws show \bit x q \ bit y q\ apply (simp add: push_bit_take_bit) unfolding bit_eq_iff apply (simp add: bit_simps not_le) apply (metis (full_types) \take_bit m x = x\ \take_bit m y = y\ add.commute add_diff_cancel_right' add_less_cancel_right bit_take_bit_iff le_add2 less_le_trans) done qed lemma word_of_nat_inj: assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)" assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)" shows "x = y" by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x") (auto dest: bounded[THEN of_nat_mono_maybe]) lemma word_of_int_bin_cat_eq_iff: "(word_of_int (concat_bit LENGTH('b) (uint b) (uint a))::'c::len word) = word_of_int (concat_bit LENGTH('b) (uint d) (uint c)) \ b = d \ a = c" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len word" and b::"'b::len word" proof - from that show ?thesis using that concat_bit_eq_iff [of \LENGTH('b)\ \uint b\ \uint a\ \uint d\ \uint c\] apply (simp add: word_of_int_eq_iff take_bit_int_eq_self flip: word_eq_iff_unsigned) apply (simp add: concat_bit_def take_bit_int_eq_self bintr_uint take_bit_push_bit) done qed lemma word_cat_inj: "(word_cat a b::'c::len word) = word_cat c d \ a = c \ b = d" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len word" and b::"'b::len word" using word_of_int_bin_cat_eq_iff [OF that, of b a d c] by (simp add: word_cat_eq' ac_simps) lemma p2_eq_1: "2 ^ n = (1::'a::len word) \ n = 0" proof - have "2 ^ n = (1::'a word) \ n = 0" by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0 power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one) then show ?thesis by auto qed end end diff --git a/thys/Word_Lib/More_Word_Operations.thy b/thys/Word_Lib/More_Word_Operations.thy --- a/thys/Word_Lib/More_Word_Operations.thy +++ b/thys/Word_Lib/More_Word_Operations.thy @@ -1,1023 +1,1064 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section \Misc word operations\ theory More_Word_Operations imports "HOL-Library.Word" Aligned Reversed_Bit_Lists More_Misc Signed_Words Word_Lemmas begin context includes bit_operations_syntax begin definition ptr_add :: "'a :: len word \ nat \ 'a word" where "ptr_add ptr n \ ptr + of_nat n" definition alignUp :: "'a::len word \ nat \ 'a word" where "alignUp x n \ x + 2 ^ n - 1 AND NOT (2 ^ n - 1)" lemma alignUp_unfold: \alignUp w n = (w + mask n) AND NOT (mask n)\ by (simp add: alignUp_def mask_eq_exp_minus_1 add_mask_fold) (* standard notation for blocks of 2^n-1 words, usually aligned; abbreviation so it simplifies directly *) abbreviation mask_range :: "'a::len word \ nat \ 'a word set" where "mask_range p n \ {p .. p + mask n}" definition w2byte :: "'a :: len word \ 8 word" where "w2byte \ ucast" (* Count leading zeros *) definition word_clz :: "'a::len word \ nat" where "word_clz w \ length (takeWhile Not (to_bl w))" (* Count trailing zeros *) definition word_ctz :: "'a::len word \ nat" where "word_ctz w \ length (takeWhile Not (rev (to_bl w)))" +lemma word_ctz_unfold: + \word_ctz w = length (takeWhile (Not \ bit w) [0.. for w :: \'a::len word\ + by (simp add: word_ctz_def rev_to_bl_eq takeWhile_map) + +lemma word_ctz_unfold': + \word_ctz w = Min (insert LENGTH('a) {n. bit w n})\ for w :: \'a::len word\ +proof (cases \\n. bit w n\) + case True + then obtain n where \bit w n\ .. + from \bit w n\ show ?thesis + apply (simp add: word_ctz_unfold) + apply (subst Min_eq_length_takeWhile [symmetric]) + apply (auto simp add: bit_imp_le_length) + apply (subst Min_insert) + apply auto + apply (subst min.absorb2) + apply (subst Min_le_iff) + apply auto + apply (meson bit_imp_le_length order_less_le) + done +next + case False + then have \bit w = bot\ + by auto + then have \word_ctz w = LENGTH('a)\ + by (simp add: word_ctz_def rev_to_bl_eq bot_fun_def map_replicate_const) + with \bit w = bot\ show ?thesis + by simp +qed + lemma word_ctz_le: "word_ctz (w :: ('a::len word)) \ LENGTH('a)" apply (clarsimp simp: word_ctz_def) using length_takeWhile_le apply (rule order_trans) apply simp done lemma word_ctz_less: "w \ 0 \ word_ctz (w :: ('a::len word)) < LENGTH('a)" apply (clarsimp simp: word_ctz_def eq_zero_set_bl) using length_takeWhile_less apply (rule less_le_trans) apply auto done lemma take_bit_word_ctz_eq [simp]: \take_bit LENGTH('a) (word_ctz w) = word_ctz w\ for w :: \'a::len word\ apply (simp add: take_bit_nat_eq_self_iff word_ctz_def to_bl_unfold) using length_takeWhile_le apply (rule le_less_trans) apply simp done lemma word_ctz_not_minus_1: \word_of_nat (word_ctz (w :: 'a :: len word)) \ (- 1 :: 'a::len word)\ if \1 < LENGTH('a)\ proof - note word_ctz_le also from that have \LENGTH('a) < mask LENGTH('a)\ by (simp add: less_mask) finally have \word_ctz w < mask LENGTH('a)\ . then have \word_of_nat (word_ctz w) < (word_of_nat (mask LENGTH('a)) :: 'a word)\ by (simp add: of_nat_word_less_iff) also have \\ = - 1\ by (rule bit_word_eqI) (simp add: bit_simps) finally show ?thesis by simp qed lemma unat_of_nat_ctz_mw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len word) = word_ctz w" - by simp + by (simp add: unsigned_of_nat) lemma unat_of_nat_ctz_smw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len signed word) = word_ctz w" - by simp + by (simp add: unsigned_of_nat) definition word_log2 :: "'a::len word \ nat" where "word_log2 (w::'a::len word) \ size w - 1 - word_clz w" (* Bit population count. Equivalent of __builtin_popcount. *) definition pop_count :: "('a::len) word \ nat" where "pop_count w \ length (filter id (to_bl w))" (* Sign extension from bit n *) definition sign_extend :: "nat \ 'a::len word \ 'a word" where "sign_extend n w \ if bit w n then w OR NOT (mask n) else w AND mask n" lemma sign_extend_eq_signed_take_bit: \sign_extend = signed_take_bit\ proof (rule ext)+ fix n and w :: \'a::len word\ show \sign_extend n w = signed_take_bit n w\ proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ then show \bit (sign_extend n w) q \ bit (signed_take_bit n w) q\ by (auto simp add: bit_signed_take_bit_iff sign_extend_def bit_and_iff bit_or_iff bit_not_iff bit_mask_iff not_less exp_eq_0_imp_not_bit not_le min_def) qed qed definition sign_extended :: "nat \ 'a::len word \ bool" where "sign_extended n w \ \i. n < i \ i < size w \ bit w i = bit w n" lemma ptr_add_0 [simp]: "ptr_add ref 0 = ref " unfolding ptr_add_def by simp lemma pop_count_0[simp]: "pop_count 0 = 0" by (clarsimp simp:pop_count_def) lemma pop_count_1[simp]: "pop_count 1 = 1" by (clarsimp simp:pop_count_def to_bl_1) lemma pop_count_0_imp_0: "(pop_count w = 0) = (w = 0)" apply (rule iffI) apply (clarsimp simp:pop_count_def) apply (subst (asm) filter_empty_conv) apply (clarsimp simp:eq_zero_set_bl) apply fast apply simp done lemma word_log2_zero_eq [simp]: \word_log2 0 = 0\ by (simp add: word_log2_def word_clz_def word_size) lemma word_log2_unfold: \word_log2 w = (if w = 0 then 0 else Max {n. bit w n})\ for w :: \'a::len word\ proof (cases \w = 0\) case True then show ?thesis by simp next case False then obtain r where \bit w r\ by (auto simp add: bit_eq_iff) then have \Max {m. bit w m} = LENGTH('a) - Suc (length (takeWhile (Not \ bit w) (rev [0.. by (subst Max_eq_length_takeWhile [of _ \LENGTH('a)\]) (auto simp add: bit_imp_le_length) then have \word_log2 w = Max {x. bit w x}\ by (simp add: word_log2_def word_clz_def word_size to_bl_unfold rev_map takeWhile_map) with \w \ 0\ show ?thesis by simp qed lemma word_log2_eqI: \word_log2 w = n\ if \w \ 0\ \bit w n\ \\m. bit w m \ m \ n\ for w :: \'a::len word\ proof - from \w \ 0\ have \word_log2 w = Max {n. bit w n}\ by (simp add: word_log2_unfold) also have \Max {n. bit w n} = n\ using that by (auto intro: Max_eqI) finally show ?thesis . qed lemma bit_word_log2: \bit w (word_log2 w)\ if \w \ 0\ proof - from \w \ 0\ have \\r. bit w r\ by (auto intro: bit_eqI) then obtain r where \bit w r\ .. from \w \ 0\ have \word_log2 w = Max {n. bit w n}\ by (simp add: word_log2_unfold) also have \Max {n. bit w n} \ {n. bit w n}\ using \bit w r\ by (subst Max_in) auto finally show ?thesis by simp qed lemma word_log2_maximum: \n \ word_log2 w\ if \bit w n\ proof - have \n \ Max {n. bit w n}\ using that by (auto intro: Max_ge) also from that have \w \ 0\ by force then have \Max {n. bit w n} = word_log2 w\ by (simp add: word_log2_unfold) finally show ?thesis . qed lemma word_log2_nth_same: "w \ 0 \ bit w (word_log2 w)" by (drule bit_word_log2) simp lemma word_log2_nth_not_set: "\ word_log2 w < i ; i < size w \ \ \ bit w i" using word_log2_maximum [of w i] by auto lemma word_log2_highest: assumes a: "bit w i" shows "i \ word_log2 w" using a by (simp add: word_log2_maximum) lemma word_log2_max: "word_log2 w < size w" apply (cases \w = 0\) apply (simp_all add: word_size) apply (drule bit_word_log2) apply (fact bit_imp_le_length) done lemma word_clz_0[simp]: "word_clz (0::'a::len word) = LENGTH('a)" unfolding word_clz_def by simp lemma word_clz_minus_one[simp]: "word_clz (-1::'a::len word) = 0" unfolding word_clz_def by simp lemma is_aligned_alignUp[simp]: "is_aligned (alignUp p n) n" by (simp add: alignUp_def is_aligned_mask mask_eq_decr_exp word_bw_assocs) lemma alignUp_le[simp]: "alignUp p n \ p + 2 ^ n - 1" unfolding alignUp_def by (rule word_and_le2) lemma alignUp_idem: fixes a :: "'a::len word" assumes "is_aligned a n" "n < LENGTH('a)" shows "alignUp a n = a" using assms unfolding alignUp_def by (metis add_cancel_right_right add_diff_eq and_mask_eq_iff_le_mask mask_eq_decr_exp mask_out_add_aligned order_refl word_plus_and_or_coroll2) lemma alignUp_not_aligned_eq: fixes a :: "'a :: len word" assumes al: "\ is_aligned a n" and sz: "n < LENGTH('a)" shows "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" proof - have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) fact+ then have um: "unat (a mod 2 ^ n - 1) div 2 ^ n = 0" using sz by (meson Euclidean_Division.div_eq_0_iff le_m1_iff_lt measure_unat order_less_trans unat_less_power word_less_sub_le word_mod_less_divisor) have "a + 2 ^ n - 1 = (a div 2 ^ n) * 2 ^ n + (a mod 2 ^ n) + 2 ^ n - 1" by (simp add: word_mod_div_equality) also have "\ = (a mod 2 ^ n - 1) + (a div 2 ^ n + 1) * 2 ^ n" by (simp add: field_simps) finally show "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" using sz unfolding alignUp_def apply (subst mask_eq_decr_exp [symmetric]) apply (erule ssubst) apply (subst neg_mask_is_div) apply (simp add: word_arith_nat_div) apply (subst unat_word_ariths(1) unat_word_ariths(2))+ apply (subst uno_simps) apply (subst unat_1) apply (subst mod_add_right_eq) apply simp apply (subst power_mod_div) apply (subst div_mult_self1) apply simp apply (subst um) apply simp apply (subst mod_mod_power) apply simp apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst mult_mod_left) apply (subst power_add [symmetric]) apply simp apply (subst Abs_fnat_hom_1) apply (subst Abs_fnat_hom_add) apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst word_unat.Rep_inverse[symmetric], subst Abs_fnat_hom_mult) apply simp done qed lemma alignUp_ge: fixes a :: "'a :: len word" assumes sz: "n < LENGTH('a)" and nowrap: "alignUp a n \ 0" shows "a \ alignUp a n" proof (cases "is_aligned a n") case True then show ?thesis using sz by (subst alignUp_idem, simp_all) next case False have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis le_add_diff_inverse2 less_mult_imp_div_less order_less_imp_le power_add unsigned_less) have"2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans nat_less_le) moreover have "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using nowrap sz apply - apply (erule contrapos_nn) apply (subst alignUp_not_aligned_eq [OF False sz]) apply (subst unat_arith_simps) apply (subst unat_word_ariths) apply (subst unat_word_ariths) apply simp apply (subst mult_mod_left) apply (simp add: unat_div field_simps power_add[symmetric] mod_mod_power) done ultimately have lt: "2 ^ n * (unat a div 2 ^ n + 1) < 2 ^ LENGTH('a)" by simp have "a = a div 2 ^ n * 2 ^ n + a mod 2 ^ n" by (rule word_mod_div_equality [symmetric]) also have "\ < (a div 2 ^ n + 1) * 2 ^ n" using sz lt apply (simp add: field_simps) apply (rule word_add_less_mono1) apply (rule word_mod_less_divisor) apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (simp add: unat_div) done also have "\ = alignUp a n" by (rule alignUp_not_aligned_eq [symmetric]) fact+ finally show ?thesis by (rule order_less_imp_le) qed lemma alignUp_le_greater_al: fixes x :: "'a :: len word" assumes le: "a \ x" and sz: "n < LENGTH('a)" and al: "is_aligned x n" shows "alignUp a n \ x" proof (cases "is_aligned a n") case True then show ?thesis using sz le by (simp add: alignUp_idem) next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) from al obtain k where xk: "x = 2 ^ n * of_nat k" and kv: "k < 2 ^ (LENGTH('a) - n)" by (auto elim!: is_alignedE) then have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst mult.commute) apply simp apply (rule nat_less_power_trans) apply simp apply simp done have au: "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ also have "\ \ of_nat k * 2 ^ n" proof (rule word_mult_le_mono1 [OF inc_le _ kn]) show "a div 2 ^ n < of_nat k" using kv xk le sz anz by (simp add: alignUp_div_helper) show "(0:: 'a word) < 2 ^ n" using sz by (simp add: p2_gt_0 sz) qed finally show ?thesis using xk by (simp add: field_simps) qed lemma alignUp_is_aligned_nz: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and ax: "a \ x" and az: "a \ 0" shows "alignUp (a::'a :: len word) n \ 0" proof (cases "is_aligned a n") case True then have "alignUp a n = a" using sz by (simp add: alignUp_idem) then show ?thesis using az by simp next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) { assume asm: "alignUp a n = 0" have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis le_add_diff_inverse2 less_mult_imp_div_less order_less_imp_le power_add unsigned_less) have leq: "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans order_less_imp_le) from al obtain k where kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" by (auto elim!: is_alignedE) then have "a div 2 ^ n < of_nat k" using ax sz anz by (rule alignUp_div_helper) then have r: "unat a div 2 ^ n < k" using sz by (simp flip: drop_bit_eq_div unat_drop_bit_eq) (metis leI le_unat_uoi unat_mono) have "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ then have "\ = 0" using asm by simp then have "2 ^ LENGTH('a) dvd 2 ^ n * (unat a div 2 ^ n + 1)" using sz by (simp add: unat_arith_simps ac_simps) (simp add: unat_word_ariths mod_simps mod_eq_0_iff_dvd) with leq have "2 ^ n * (unat a div 2 ^ n + 1) = 2 ^ LENGTH('a)" by (force elim!: le_SucE) then have "unat a div 2 ^ n = 2 ^ LENGTH('a) div 2 ^ n - 1" by (metis (no_types, opaque_lifting) Groups.add_ac(2) add.right_neutral add_diff_cancel_left' div_le_dividend div_mult_self4 gr_implies_not0 le_neq_implies_less power_eq_0_iff zero_neq_numeral) then have "unat a div 2 ^ n = 2 ^ (LENGTH('a) - n) - 1" using sz by (simp add: power_sub) then have "2 ^ (LENGTH('a) - n) - 1 < k" using r by simp then have False using kv by simp } then show ?thesis by clarsimp qed lemma alignUp_ar_helper: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and sub: "{x..x + 2 ^ n - 1} \ {a..b}" and anz: "a \ 0" shows "a \ alignUp a n \ alignUp a n + 2 ^ n - 1 \ b" proof from al have xl: "x \ x + 2 ^ n - 1" by (simp add: is_aligned_no_overflow) from xl sub have ax: "a \ x" by auto show "a \ alignUp a n" proof (rule alignUp_ge) show "alignUp a n \ 0" using al sz ax anz by (rule alignUp_is_aligned_nz) qed fact+ show "alignUp a n + 2 ^ n - 1 \ b" proof (rule order_trans) from xl show tp: "x + 2 ^ n - 1 \ b" using sub by auto from ax have "alignUp a n \ x" by (rule alignUp_le_greater_al) fact+ then have "alignUp a n + (2 ^ n - 1) \ x + (2 ^ n - 1)" using xl al is_aligned_no_overflow' olen_add_eqv word_plus_mcs_3 by blast then show "alignUp a n + 2 ^ n - 1 \ x + 2 ^ n - 1" by (simp add: field_simps) qed qed lemma alignUp_def2: "alignUp a sz = a + 2 ^ sz - 1 AND NOT (mask sz)" by (simp add: alignUp_def flip: mask_eq_decr_exp) lemma alignUp_def3: "alignUp a sz = 2^ sz + (a - 1 AND NOT (mask sz))" by (simp add: alignUp_def2 is_aligned_triv field_simps mask_out_add_aligned) lemma alignUp_plus: "is_aligned w us \ alignUp (w + a) us = w + alignUp a us" by (clarsimp simp: alignUp_def2 mask_out_add_aligned field_simps) lemma alignUp_distance: "alignUp (q :: 'a :: len word) sz - q \ mask sz" by (metis (no_types) add.commute add_diff_cancel_left alignUp_def2 diff_add_cancel mask_2pm1 subtract_mask(2) word_and_le1 word_sub_le_iff) lemma is_aligned_diff_neg_mask: "is_aligned p sz \ (p - q AND NOT (mask sz)) = (p - ((alignUp q sz) AND NOT (mask sz)))" apply (clarsimp simp only:word_and_le2 diff_conv_add_uminus) apply (subst mask_out_add_aligned[symmetric]; simp) apply (simp add: eq_neg_iff_add_eq_0) apply (subst add.commute) apply (simp add: alignUp_distance is_aligned_neg_mask_eq mask_out_add_aligned and_mask_eq_iff_le_mask flip: mask_eq_x_eq_0) done lemma word_clz_max: "word_clz w \ size (w::'a::len word)" unfolding word_clz_def by (metis length_takeWhile_le word_size_bl) lemma word_clz_nonzero_max: fixes w :: "'a::len word" assumes nz: "w \ 0" shows "word_clz w < size (w::'a::len word)" proof - { assume a: "word_clz w = size (w::'a::len word)" hence "length (takeWhile Not (to_bl w)) = length (to_bl w)" by (simp add: word_clz_def word_size) hence allj: "\j\set(to_bl w). \ j" by (metis a length_takeWhile_less less_irrefl_nat word_clz_def) hence "to_bl w = replicate (length (to_bl w)) False" using eq_zero_set_bl nz by fastforce hence "w = 0" by (metis to_bl_0 word_bl.Rep_eqD word_bl_Rep') with nz have False by simp } thus ?thesis using word_clz_max by (fastforce intro: le_neq_trans) qed (* Sign extension from bit n. *) lemma bin_sign_extend_iff [bit_simps]: \bit (sign_extend e w) i \ bit w (min e i)\ if \i < LENGTH('a)\ for w :: \'a::len word\ using that by (simp add: sign_extend_def bit_simps min_def) lemma sign_extend_bitwise_if: "i < size w \ bit (sign_extend e w) i \ (if i < e then bit w i else bit w e)" by (simp add: word_size bit_simps) lemma sign_extend_bitwise_if' [word_eqI_simps]: \i < LENGTH('a) \ bit (sign_extend e w) i \ (if i < e then bit w i else bit w e)\ for w :: \'a::len word\ using sign_extend_bitwise_if [of i w e] by (simp add: word_size) lemma sign_extend_bitwise_disj: "i < size w \ bit (sign_extend e w) i \ i \ e \ bit w i \ e \ i \ bit w e" by (auto simp: sign_extend_bitwise_if) lemma sign_extend_bitwise_cases: "i < size w \ bit (sign_extend e w) i \ (i \ e \ bit w i) \ (e \ i \ bit w e)" by (auto simp: sign_extend_bitwise_if) lemmas sign_extend_bitwise_disj' = sign_extend_bitwise_disj[simplified word_size] lemmas sign_extend_bitwise_cases' = sign_extend_bitwise_cases[simplified word_size] (* Often, it is easier to reason about an operation which does not overwrite the bit which determines which mask operation to apply. *) lemma sign_extend_def': "sign_extend n w = (if bit w n then w OR NOT (mask (Suc n)) else w AND mask (Suc n))" by (rule bit_word_eqI) (auto simp add: bit_simps sign_extend_eq_signed_take_bit min_def less_Suc_eq_le) lemma sign_extended_sign_extend: "sign_extended n (sign_extend n w)" by (clarsimp simp: sign_extended_def word_size sign_extend_bitwise_if) lemma sign_extended_iff_sign_extend: "sign_extended n w \ sign_extend n w = w" apply auto apply (auto simp add: bit_eq_iff) apply (simp_all add: bit_simps sign_extend_eq_signed_take_bit not_le min_def sign_extended_def word_size split: if_splits) - using le_imp_less_or_eq apply auto[1] - apply (metis bit_imp_le_length nat_less_le) - apply (metis Suc_leI Suc_n_not_le_n le_trans nat_less_le) + using le_imp_less_or_eq apply auto done lemma sign_extended_weaken: "sign_extended n w \ n \ m \ sign_extended m w" unfolding sign_extended_def by (cases "n < m") auto lemma sign_extend_sign_extend_eq: "sign_extend m (sign_extend n w) = sign_extend (min m n) w" by (rule bit_word_eqI) (simp add: sign_extend_eq_signed_take_bit bit_simps) lemma sign_extended_high_bits: "\ sign_extended e p; j < size p; e \ i; i < j \ \ bit p i = bit p j" by (drule (1) sign_extended_weaken; simp add: sign_extended_def) lemma sign_extend_eq: "w AND mask (Suc n) = v AND mask (Suc n) \ sign_extend n w = sign_extend n v" by (simp flip: take_bit_eq_mask add: sign_extend_eq_signed_take_bit signed_take_bit_eq_iff_take_bit_eq) lemma sign_extended_add: assumes p: "is_aligned p n" assumes f: "f < 2 ^ n" assumes e: "n \ e" assumes "sign_extended e p" shows "sign_extended e (p + f)" proof (cases "e < size p") case True note and_or = is_aligned_add_or[OF p f] have "\ bit f e" using True e less_2p_is_upper_bits_unset[THEN iffD1, OF f] by (fastforce simp: word_size) hence i: "bit (p + f) e = bit p e" by (simp add: and_or bit_simps) have fm: "f AND mask e = f" by (fastforce intro: subst[where P="\f. f AND mask e = f", OF less_mask_eq[OF f]] simp: mask_twice e) show ?thesis using assms apply (simp add: sign_extended_iff_sign_extend sign_extend_def i) apply (simp add: and_or word_bw_comms[of p f]) apply (clarsimp simp: word_ao_dist fm word_bw_assocs split: if_splits) done next case False thus ?thesis by (simp add: sign_extended_def word_size) qed lemma sign_extended_neq_mask: "\sign_extended n ptr; m \ n\ \ sign_extended n (ptr AND NOT (mask m))" by (fastforce simp: sign_extended_def word_size neg_mask_test_bit bit_simps) definition "limited_and (x :: 'a :: len word) y \ (x AND y = x)" lemma limited_and_eq_0: "\ limited_and x z; y AND NOT z = y \ \ x AND y = 0" unfolding limited_and_def apply (subst arg_cong2[where f="(AND)"]) apply (erule sym)+ apply (simp(no_asm) add: word_bw_assocs word_bw_comms word_bw_lcs) done lemma limited_and_eq_id: "\ limited_and x z; y AND z = z \ \ x AND y = x" unfolding limited_and_def by (erule subst, fastforce simp: word_bw_lcs word_bw_assocs word_bw_comms) lemma lshift_limited_and: "limited_and x z \ limited_and (x << n) (z << n)" unfolding limited_and_def by (metis push_bit_and) lemma rshift_limited_and: "limited_and x z \ limited_and (x >> n) (z >> n)" unfolding limited_and_def by (metis drop_bit_and) lemmas limited_and_simps1 = limited_and_eq_0 limited_and_eq_id lemmas is_aligned_limited_and = is_aligned_neg_mask_eq[unfolded mask_eq_decr_exp, folded limited_and_def] lemmas limited_and_simps = limited_and_simps1 limited_and_simps1[OF is_aligned_limited_and] limited_and_simps1[OF lshift_limited_and] limited_and_simps1[OF rshift_limited_and] limited_and_simps1[OF rshift_limited_and, OF is_aligned_limited_and] not_one_eq definition from_bool :: "bool \ 'a::len word" where "from_bool b \ case b of True \ of_nat 1 | False \ of_nat 0" lemma from_bool_eq: \from_bool = of_bool\ by (simp add: fun_eq_iff from_bool_def) lemma from_bool_0: "(from_bool x = 0) = (\ x)" by (simp add: from_bool_def split: bool.split) lemma from_bool_eq_if': "((if P then 1 else 0) = from_bool Q) = (P = Q)" by (cases Q) (simp_all add: from_bool_def) definition to_bool :: "'a::len word \ bool" where "to_bool \ (\) 0" lemma to_bool_and_1: "to_bool (x AND 1) \ bit x 0" by (simp add: to_bool_def and_one_eq mod_2_eq_odd) lemma to_bool_from_bool [simp]: "to_bool (from_bool r) = r" unfolding from_bool_def to_bool_def by (simp split: bool.splits) lemma from_bool_neq_0 [simp]: "(from_bool b \ 0) = b" by (simp add: from_bool_def split: bool.splits) lemma from_bool_mask_simp [simp]: "(from_bool r :: 'a::len word) AND 1 = from_bool r" unfolding from_bool_def by (clarsimp split: bool.splits) lemma from_bool_1 [simp]: "(from_bool P = 1) = P" by (simp add: from_bool_def split: bool.splits) lemma ge_0_from_bool [simp]: "(0 < from_bool P) = P" by (simp add: from_bool_def split: bool.splits) lemma limited_and_from_bool: "limited_and (from_bool b) 1" by (simp add: from_bool_def limited_and_def split: bool.split) lemma to_bool_1 [simp]: "to_bool 1" by (simp add: to_bool_def) lemma to_bool_0 [simp]: "\to_bool 0" by (simp add: to_bool_def) lemma from_bool_eq_if: "(from_bool Q = (if P then 1 else 0)) = (P = Q)" by (cases Q) (simp_all add: from_bool_def) lemma to_bool_eq_0: "(\ to_bool x) = (x = 0)" by (simp add: to_bool_def) lemma to_bool_neq_0: "(to_bool x) = (x \ 0)" by (simp add: to_bool_def) lemma from_bool_all_helper: "(\bool. from_bool bool = val \ P bool) = ((\bool. from_bool bool = val) \ P (val \ 0))" by (auto simp: from_bool_0) lemma fold_eq_0_to_bool: "(v = 0) = (\ to_bool v)" by (simp add: to_bool_def) lemma from_bool_to_bool_iff: "w = from_bool b \ to_bool w = b \ (w = 0 \ w = 1)" by (cases b) (auto simp: from_bool_def to_bool_def) lemma from_bool_eqI: "from_bool x = from_bool y \ x = y" unfolding from_bool_def by (auto split: bool.splits) lemma neg_mask_in_mask_range: "is_aligned ptr bits \ (ptr' AND NOT(mask bits) = ptr) = (ptr' \ mask_range ptr bits)" apply (erule is_aligned_get_word_bits) apply (rule iffI) apply (drule sym) apply (simp add: word_and_le2) apply (subst word_plus_and_or_coroll, word_eqI_solve) apply (metis bit.disj_ac(2) bit.disj_conj_distrib2 le_word_or2 word_and_max word_or_not) apply clarsimp apply (smt add.right_neutral eq_iff is_aligned_neg_mask_eq mask_out_add_aligned neg_mask_mono_le word_and_not) apply (simp add: power_overflow mask_eq_decr_exp) done lemma aligned_offset_in_range: "\ is_aligned (x :: 'a :: len word) m; y < 2 ^ m; is_aligned p n; n \ m; n < LENGTH('a) \ \ (x + y \ {p .. p + mask n}) = (x \ mask_range p n)" apply (subst disjunctive_add) apply (simp add: bit_simps) apply (erule is_alignedE') apply (auto simp add: bit_simps not_le)[1] apply (metis less_2p_is_upper_bits_unset) apply (simp only: is_aligned_add_or word_ao_dist flip: neg_mask_in_mask_range) apply (subgoal_tac \y AND NOT (mask n) = 0\) apply simp apply (metis (full_types) is_aligned_mask is_aligned_neg_mask less_mask_eq word_bw_comms(1) word_bw_lcs(1)) done lemma mask_range_to_bl': "\ is_aligned (ptr :: 'a :: len word) bits; bits < LENGTH('a) \ \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (rule set_eqI, rule iffI) apply clarsimp apply (subgoal_tac "\y. x = ptr + y \ y < 2 ^ bits") apply clarsimp apply (subst is_aligned_add_conv) apply assumption apply simp apply simp apply (rule_tac x="x - ptr" in exI) apply (simp add: add_diff_eq[symmetric]) apply (simp only: word_less_sub_le[symmetric]) apply (rule word_diff_ls') apply (simp add: field_simps mask_eq_decr_exp) apply assumption apply simp apply (subgoal_tac "\y. y < 2 ^ bits \ to_bl (ptr + y) = to_bl x") apply clarsimp apply (rule conjI) apply (erule(1) is_aligned_no_wrap') apply (simp only: add_diff_eq[symmetric] mask_eq_decr_exp) apply (rule word_plus_mono_right) apply simp apply (erule is_aligned_no_wrap') apply simp apply (rule_tac x="of_bl (drop (LENGTH('a) - bits) (to_bl x))" in exI) apply (rule context_conjI) apply (rule order_less_le_trans [OF of_bl_length]) apply simp apply simp apply (subst is_aligned_add_conv) apply assumption apply simp apply (drule sym) apply (simp add: word_rep_drop) done lemma mask_range_to_bl: "is_aligned (ptr :: 'a :: len word) bits \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (erule is_aligned_get_word_bits) apply (erule(1) mask_range_to_bl') apply (rule set_eqI) apply (simp add: power_overflow mask_eq_decr_exp) done lemma aligned_mask_range_cases: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n' \ \ mask_range p n \ mask_range p' n' = {} \ mask_range p n \ mask_range p' n' \ mask_range p n \ mask_range p' n'" apply (simp add: mask_range_to_bl) apply (rule Meson.disj_comm, rule disjCI) apply auto apply (subgoal_tac "(\n''. LENGTH('a) - n = (LENGTH('a) - n') + n'') \ (\n''. LENGTH('a) - n' = (LENGTH('a) - n) + n'')") apply (fastforce simp: take_add) apply arith done lemma aligned_mask_range_offset_subset: assumes al: "is_aligned (ptr :: 'a :: len word) sz" and al': "is_aligned x sz'" and szv: "sz' \ sz" and xsz: "x < 2 ^ sz" shows "mask_range (ptr+x) sz' \ mask_range ptr sz" using al proof (rule is_aligned_get_word_bits) assume p0: "ptr = 0" and szv': "LENGTH ('a) \ sz" then have "(2 ::'a word) ^ sz = 0" by simp show ?thesis using p0 by (simp add: \2 ^ sz = 0\ mask_eq_decr_exp) next assume szv': "sz < LENGTH('a)" hence blah: "2 ^ (sz - sz') < (2 :: nat) ^ LENGTH('a)" using szv by auto show ?thesis using szv szv' apply auto using al assms(4) is_aligned_no_wrap' apply blast apply (simp only: flip: add_diff_eq add_mask_fold) apply (subst add.assoc, rule word_plus_mono_right) using al' is_aligned_add_less_t2n xsz apply fastforce apply (simp add: field_simps szv al is_aligned_no_overflow) done qed lemma aligned_mask_ranges_disjoint: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n'; p AND NOT(mask n') \ p'; p' AND NOT(mask n) \ p \ \ mask_range p n \ mask_range p' n' = {}" using aligned_mask_range_cases by (auto simp: neg_mask_in_mask_range) lemma aligned_mask_ranges_disjoint2: "\ is_aligned p n; is_aligned ptr bits; n \ m; n < size p; m \ bits; (\y < 2 ^ (n - m). p + (y << m) \ mask_range ptr bits) \ \ mask_range p n \ mask_range ptr bits = {}" apply safe apply (simp only: flip: neg_mask_in_mask_range) apply (drule_tac x="x AND mask n >> m" in spec) apply (erule notE[OF mp]) apply (rule shiftr_less_t2n) apply (simp add: word_size and_mask_less_size) apply (subst disjunctive_add) apply (auto simp add: bit_simps word_size intro!: bit_eqI) done lemma word_clz_sint_upper[simp]: "LENGTH('a) \ 3 \ sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a sword) \ int (LENGTH('a))" using word_clz_max [of w] - apply (simp add: word_size) + apply (simp add: word_size signed_of_nat) apply (subst signed_take_bit_int_eq_self) apply simp_all apply (metis negative_zle of_nat_numeral semiring_1_class.of_nat_power) apply (drule small_powers_of_2) apply (erule le_less_trans) apply simp done lemma word_clz_sint_lower[simp]: "LENGTH('a) \ 3 \ - sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a signed word) \ int (LENGTH('a))" apply (subst sint_eq_uint) using word_clz_max [of w] - apply (simp_all add: word_size) + apply (simp_all add: word_size unsigned_of_nat) apply (rule not_msb_from_less) - apply (simp add: word_less_nat_alt) + apply (simp add: word_less_nat_alt unsigned_of_nat) apply (subst take_bit_nat_eq_self) apply (simp add: le_less_trans) apply (drule small_powers_of_2) apply (erule le_less_trans) apply simp done lemma mask_range_subsetD: "\ p' \ mask_range p n; x' \ mask_range p' n'; n' \ n; is_aligned p n; is_aligned p' n' \ \ x' \ mask_range p n" using aligned_mask_step by fastforce lemma nasty_split_lt: "\ (x :: 'a:: len word) < 2 ^ (m - n); n \ m; m < LENGTH('a::len) \ \ x * 2 ^ n + (2 ^ n - 1) \ 2 ^ m - 1" apply (simp only: add_diff_eq) apply (subst mult_1[symmetric], subst distrib_right[symmetric]) apply (rule word_sub_mono) apply (rule order_trans) apply (rule word_mult_le_mono1) apply (rule inc_le) apply assumption apply (subst word_neq_0_conv[symmetric]) apply (rule power_not_zero) apply simp apply (subst unat_power_lower, simp)+ apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply (subst power_add[symmetric]) apply simp apply simp apply (rule word_sub_1_le) apply (subst mult.commute) apply (subst shiftl_t2n[symmetric]) apply (rule word_shift_nonzero) apply (erule inc_le) apply simp apply (unat_arith) apply (drule word_power_less_1) apply simp done lemma nasty_split_less: "\m \ n; n \ nm; nm < LENGTH('a::len); x < 2 ^ (nm - n)\ \ (x :: 'a word) * 2 ^ n + (2 ^ m - 1) < 2 ^ nm" apply (simp only: word_less_sub_le[symmetric]) apply (rule order_trans [OF _ nasty_split_lt]) apply (rule word_plus_mono_right) apply (rule word_sub_mono) apply (simp add: word_le_nat_alt) apply simp apply (simp add: word_sub_1_le[OF power_not_zero]) apply (simp add: word_sub_1_le[OF power_not_zero]) apply (rule is_aligned_no_wrap') apply (rule is_aligned_mult_triv2) apply simp apply (erule order_le_less_trans, simp) apply simp+ done lemma add_mult_in_mask_range: "\ is_aligned (base :: 'a :: len word) n; n < LENGTH('a); bits \ n; x < 2 ^ (n - bits) \ \ base + x * 2^bits \ mask_range base n" by (simp add: is_aligned_no_wrap' mask_2pm1 nasty_split_lt word_less_power_trans2 word_plus_mono_right) lemma from_to_bool_last_bit: "from_bool (to_bool (x AND 1)) = x AND 1" by (metis from_bool_to_bool_iff word_and_1) lemma sint_ctz: - "LENGTH('a) > 2 - \ 0 \ sint (of_nat (word_ctz (x :: 'a :: len word)) :: 'a signed word) - \ sint (of_nat (word_ctz x) :: 'a signed word) \ int (LENGTH('a))" - apply (subgoal_tac "LENGTH('a) < 2 ^ (LENGTH('a) - 1)") - apply (rule conjI) - apply (metis len_signed order_le_less_trans sint_of_nat_ge_zero word_ctz_le) - apply (metis int_eq_sint len_signed sint_of_nat_le word_ctz_le) - using small_powers_of_2 [of \LENGTH('a)\] by simp + \0 \ sint (of_nat (word_ctz (x :: 'a :: len word)) :: 'a signed word) + \ sint (of_nat (word_ctz x) :: 'a signed word) \ int (LENGTH('a))\ (is \?P \ ?Q\) + if \LENGTH('a) > 2\ +proof + have *: \word_ctz x < 2 ^ (LENGTH('a) - Suc 0)\ + using word_ctz_le apply (rule le_less_trans) + using that small_powers_of_2 [of \LENGTH('a)\] apply simp + done + have \int (word_ctz x) div 2 ^ (LENGTH('a) - Suc 0) = 0\ + apply (rule div_pos_pos_trivial) + apply (simp_all add: *) + done + then show ?P by (simp add: signed_of_nat bit_iff_odd) + show ?Q + apply (auto simp add: signed_of_nat) + apply (subst signed_take_bit_int_eq_self) + apply (auto simp add: word_ctz_le * minus_le_iff [of _ \int (word_ctz x)\]) + apply (rule order.trans [of _ 0]) + apply simp_all + done +qed lemma unat_of_nat_word_log2: "LENGTH('a) < 2 ^ LENGTH('b) \ unat (of_nat (word_log2 (n :: 'a :: len word)) :: 'b :: len word) = word_log2 n" by (metis less_trans unat_of_nat_eq word_log2_max word_size) lemma aligned_mask_diff: "\ is_aligned (dest :: 'a :: len word) bits; is_aligned (ptr :: 'a :: len word) sz; bits \ sz; sz < LENGTH('a); dest < ptr \ \ mask bits + dest < ptr" apply (frule_tac p' = ptr in aligned_mask_range_cases, assumption) apply (elim disjE) apply (drule_tac is_aligned_no_overflow_mask, simp)+ apply (simp add: algebra_split_simps word_le_not_less) apply (drule is_aligned_no_overflow_mask; fastforce) apply (simp add: is_aligned_weaken algebra_split_simps) apply (auto simp add: not_le) using is_aligned_no_overflow_mask leD apply blast apply (meson aligned_add_mask_less_eq is_aligned_weaken le_less_trans) done end end \ No newline at end of file diff --git a/thys/Word_Lib/Rsplit.thy b/thys/Word_Lib/Rsplit.thy --- a/thys/Word_Lib/Rsplit.thy +++ b/thys/Word_Lib/Rsplit.thy @@ -1,168 +1,168 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson and Gerwin Klein, NICTA *) section \Splitting words into lists\ theory Rsplit imports "HOL-Library.Word" More_Word Bits_Int begin definition word_rsplit :: "'a::len word \ 'b::len word list" where "word_rsplit w = map word_of_int (bin_rsplit LENGTH('b) (LENGTH('a), uint w))" lemma word_rsplit_no: "(word_rsplit (numeral bin :: 'b::len word) :: 'a word list) = map word_of_int (bin_rsplit (LENGTH('a::len)) (LENGTH('b), take_bit (LENGTH('b)) (numeral bin)))" by (simp add: word_rsplit_def of_nat_take_bit) lemmas word_rsplit_no_cl [simp] = word_rsplit_no [unfolded bin_rsplitl_def bin_rsplit_l [symmetric]] text \ This odd result arises from the fact that the statement of the result implies that the decoded words are of the same type, and therefore of the same length, as the original word.\ lemma word_rsplit_same: "word_rsplit w = [w]" apply (simp add: word_rsplit_def bin_rsplit_all) apply transfer apply simp done lemma word_rsplit_empty_iff_size: "word_rsplit w = [] \ size w = 0" by (simp add: word_rsplit_def bin_rsplit_def word_size bin_rsplit_aux_simp_alt Let_def split: prod.split) lemma test_bit_rsplit: "sw = word_rsplit w \ m < size (hd sw) \ k < length sw \ bit (rev sw ! k) m = bit w (k * size (hd sw) + m)" for sw :: "'a::len word list" apply (unfold word_rsplit_def word_test_bit_def) apply (rule trans) apply (rule_tac f = "\x. bit x m" in arg_cong) apply (rule nth_map [symmetric]) apply simp apply (rule bin_nth_rsplit) apply simp_all apply (simp add : word_size rev_map) apply (rule trans) defer apply (rule map_ident [THEN fun_cong]) apply (rule refl [THEN map_cong]) - apply simp + apply (simp add: unsigned_of_int take_bit_int_eq_self_iff) using bin_rsplit_size_sign take_bit_int_eq_self_iff by blast lemma test_bit_rsplit_alt: \bit ((word_rsplit w :: 'b::len word list) ! i) m \ bit w ((length (word_rsplit w :: 'b::len word list) - Suc i) * size (hd (word_rsplit w :: 'b::len word list)) + m)\ if \i < length (word_rsplit w :: 'b::len word list)\ \m < size (hd (word_rsplit w :: 'b::len word list))\ \0 < length (word_rsplit w :: 'b::len word list)\ for w :: \'a::len word\ apply (rule trans) apply (rule test_bit_cong) apply (rule rev_nth [of _ \rev (word_rsplit w)\, simplified rev_rev_ident]) apply simp apply (rule that(1)) apply simp apply (rule test_bit_rsplit) apply (rule refl) apply (rule asm_rl) apply (rule that(2)) apply (rule diff_Suc_less) apply (rule that(3)) done lemma word_rsplit_len_indep [OF refl refl refl refl]: "[u,v] = p \ [su,sv] = q \ word_rsplit u = su \ word_rsplit v = sv \ length su = length sv" by (auto simp: word_rsplit_def bin_rsplit_len_indep) lemma length_word_rsplit_size: "n = LENGTH('a::len) \ length (word_rsplit w :: 'a word list) \ m \ size w \ m * n" by (auto simp: word_rsplit_def word_size bin_rsplit_len_le) lemmas length_word_rsplit_lt_size = length_word_rsplit_size [unfolded Not_eq_iff linorder_not_less [symmetric]] lemma length_word_rsplit_exp_size: "n = LENGTH('a::len) \ length (word_rsplit w :: 'a word list) = (size w + n - 1) div n" by (auto simp: word_rsplit_def word_size bin_rsplit_len) lemma length_word_rsplit_even_size: "n = LENGTH('a::len) \ size w = m * n \ length (word_rsplit w :: 'a word list) = m" by (cases \LENGTH('a)\) (simp_all add: length_word_rsplit_exp_size div_nat_eqI) lemmas length_word_rsplit_exp_size' = refl [THEN length_word_rsplit_exp_size] \ \alternative proof of \word_rcat_rsplit\\ lemmas tdle = times_div_less_eq_dividend lemmas dtle = xtrans(4) [OF tdle mult.commute] lemma word_rcat_rsplit: "word_rcat (word_rsplit w) = w" apply (rule word_eqI) apply (clarsimp simp: test_bit_rcat word_size) apply (subst refl [THEN test_bit_rsplit]) apply (simp_all add: word_size refl [THEN length_word_rsplit_size [simplified not_less [symmetric], simplified]]) apply safe apply (erule xtrans(7), rule dtle)+ done lemma size_word_rsplit_rcat_size: "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) \ length (word_rsplit frcw::'a word list) = length ws" for ws :: "'a::len word list" and frcw :: "'b::len word" by (cases \LENGTH('a)\) (simp_all add: word_size length_word_rsplit_exp_size' div_nat_eqI) lemma msrevs: "0 < n \ (k * n + m) div n = m div n + k" "(k * n + m) mod n = m mod n" for n :: nat by (auto simp: add.commute) lemma word_rsplit_rcat_size [OF refl]: "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) \ word_rsplit frcw = ws" for ws :: "'a::len word list" apply (frule size_word_rsplit_rcat_size, assumption) apply (clarsimp simp add : word_size) apply (rule nth_equalityI, assumption) apply clarsimp apply (rule word_eqI [rule_format]) apply (rule trans) apply (rule test_bit_rsplit_alt) apply (clarsimp simp: word_size)+ apply (rule trans) apply (rule test_bit_rcat [OF refl refl]) apply (simp add: word_size) apply (subst rev_nth) apply arith apply (simp add: le0 [THEN [2] xtrans(7), THEN diff_Suc_less]) apply safe apply (simp add: diff_mult_distrib) apply (cases "size ws") apply simp_all done lemma word_rsplit_upt: "\ size x = LENGTH('a :: len) * n; n \ 0 \ \ word_rsplit x = map (\i. ucast (drop_bit (i * LENGTH('a)) x) :: 'a word) (rev [0 ..< n])" apply (subgoal_tac "length (word_rsplit x :: 'a word list) = n") apply (rule nth_equalityI, simp) apply (intro allI word_eqI impI) apply (simp add: test_bit_rsplit_alt word_size) apply (simp add: nth_ucast bit_simps rev_nth field_simps) apply (simp add: length_word_rsplit_exp_size) apply transfer apply (metis (no_types, lifting) Nat.add_diff_assoc Suc_leI add_0_left diff_Suc_less div_less len_gt_0 msrevs(1) mult.commute) done end \ No newline at end of file diff --git a/thys/Word_Lib/Word_32.thy b/thys/Word_Lib/Word_32.thy --- a/thys/Word_Lib/Word_32.thy +++ b/thys/Word_Lib/Word_32.thy @@ -1,348 +1,345 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Words of Length 32" theory Word_32 imports Word_Lemmas Word_Syntax Word_Names Rsplit More_Word_Operations Bitwise begin context includes bit_operations_syntax begin type_synonym word32 = "32 word" lemma len32: "len_of (x :: 32 itself) = 32" by simp type_synonym sword32 = "32 sword" type_synonym machine_word_len = 32 type_synonym machine_word = "machine_word_len word" definition word_bits :: nat where "word_bits = LENGTH(machine_word_len)" text \The following two are numerals so they can be used as nats and words.\ definition word_size_bits :: "'a :: numeral" where "word_size_bits = 2" definition word_size :: "'a :: numeral" where "word_size = 4" lemma word_bits_conv[code]: "word_bits = 32" unfolding word_bits_def by simp lemma word_size_word_size_bits: "(word_size::nat) = 2 ^ word_size_bits" unfolding word_size_def word_size_bits_def by simp lemma word_bits_word_size_conv: "word_bits = word_size * 8" unfolding word_bits_def word_size_def by simp lemma ucast_8_32_inj: "inj (ucast :: 8 word \ 32 word)" by (rule down_ucast_inj) (clarsimp simp: is_down_def target_size source_size) lemma upto_2_helper: "{0..<2 :: 32 word} = {0, 1}" by (safe; simp) unat_arith lemmas upper_bits_unset_is_l2p_32 = upper_bits_unset_is_l2p [where 'a=32, folded word_bits_def] lemmas le_2p_upper_bits_32 = le_2p_upper_bits [where 'a=32, folded word_bits_def] lemmas le2p_bits_unset_32 = le2p_bits_unset[where 'a=32, folded word_bits_def] lemma word_bits_len_of: "len_of TYPE (32) = word_bits" by (simp add: word_bits_conv) lemmas unat_power_lower32' = unat_power_lower[where 'a=32] lemmas unat_power_lower32 [simp] = unat_power_lower32'[unfolded word_bits_len_of] lemmas word32_less_sub_le' = word_less_sub_le[where 'a = 32] lemmas word32_less_sub_le[simp] = word32_less_sub_le' [folded word_bits_def] lemma word_bits_size: "size (w::word32) = word_bits" by (simp add: word_bits_def word_size) lemmas word32_power_less_1' = word_power_less_1[where 'a = 32] lemmas word32_power_less_1[simp] = word32_power_less_1'[folded word_bits_def] lemma of_nat32_0: "\of_nat n = (0::word32); n < 2 ^ word_bits\ \ n = 0" by (erule of_nat_0, simp add: word_bits_def) lemma unat_mask_2_less_4: "unat (p && mask 2 :: word32) < 4" by (rule unat_less_helper) (simp flip: take_bit_eq_mask add: take_bit_eq_mod word_mod_less_divisor) lemmas unat_of_nat32' = unat_of_nat_eq[where 'a=32] lemmas unat_of_nat32 = unat_of_nat32'[unfolded word_bits_len_of] lemmas word_power_nonzero_32 = word_power_nonzero [where 'a=32, folded word_bits_def] lemmas unat_mult_simple = iffD1 [OF unat_mult_lem [where 'a = 32, unfolded word_bits_len_of]] lemmas div_power_helper_32 = div_power_helper [where 'a=32, folded word_bits_def] lemma n_less_word_bits: "(n < word_bits) = (n < 32)" by (simp add: word_bits_def) lemmas of_nat_less_pow_32 = of_nat_power [where 'a=32, folded word_bits_def] lemma lt_word_bits_lt_pow: "sz < word_bits \ sz < 2 ^ word_bits" by (simp add: word_bits_conv) lemma unat_less_word_bits: fixes y :: word32 shows "x < unat y \ x < 2 ^ word_bits" unfolding word_bits_def by (rule order_less_trans [OF _ unat_lt2p]) lemmas unat_mask_word32' = unat_mask[where 'a=32] lemmas unat_mask_word32 = unat_mask_word32'[folded word_bits_def] lemma unat_less_2p_word_bits: "unat (x :: 32 word) < 2 ^ word_bits" apply (simp only: word_bits_def) apply (rule unat_lt2p) done lemma Suc_unat_mask_div: "Suc (unat (mask sz div word_size::word32)) = 2 ^ (min sz word_bits - 2)" proof (cases \sz \ 2\) case False then have \sz \ {0, 1}\ by auto then show ?thesis by (auto simp add: unat_div word_size_def) next case True moreover define n where \n = sz - 2\ ultimately have \sz = n + 2\ by simp moreover have \4 * 2 ^ n = 2 ^ n * (4::nat)\ by (simp add: ac_simps) then have \4 * 2 ^ n - Suc 0 = (2 ^ n - 1) * 4 + 3\ by (simp add: mult_eq_if) ultimately show ?thesis by (simp add: unat_div unat_mask word_size_def word_bits_def min_def) qed lemmas word32_minus_one_le' = word_minus_one_le[where 'a=32] lemmas word32_minus_one_le = word32_minus_one_le'[simplified] lemma ucast_not_helper: fixes a::"8 word" assumes a: "a \ 0xFF" shows "ucast a \ (0xFF::word32)" proof assume "ucast a = (0xFF::word32)" also have "(0xFF::word32) = ucast (0xFF::8 word)" by simp finally show False using a apply - apply (drule up_ucast_inj, simp) apply simp done qed lemma less_4_cases: "(x::word32) < 4 \ x=0 \ x=1 \ x=2 \ x=3" apply clarsimp apply (drule word_less_cases, erule disjE, simp, simp)+ done lemma unat_ucast_8_32: fixes x :: "8 word" shows "unat (ucast x :: word32) = unat x" by transfer simp lemma if_then_1_else_0: "((if P then 1 else 0) = (0 :: word32)) = (\ P)" by simp lemma if_then_0_else_1: "((if P then 0 else 1) = (0 :: word32)) = (P)" by simp lemmas if_then_simps = if_then_0_else_1 if_then_1_else_0 lemma ucast_le_ucast_8_32: "(ucast x \ (ucast y :: word32)) = (x \ (y :: 8 word))" by (simp add: ucast_le_ucast) lemma in_16_range: "0 \ S \ r \ (\x. r + x * (16 :: word32)) ` S" "n - 1 \ S \ (r + (16 * n - 16)) \ (\x :: word32. r + x * 16) ` S" by (clarsimp simp: image_def elim!: bexI[rotated])+ lemma eq_2_32_0: "(2 ^ 32 :: word32) = 0" by simp lemma x_less_2_0_1: fixes x :: word32 shows "x < 2 \ x = 0 \ x = 1" by (rule x_less_2_0_1') auto lemmas mask_32_max_word = max_word_mask [symmetric, where 'a=32, simplified] lemma of_nat32_n_less_equal_power_2: "n < 32 \ ((of_nat n)::32 word) < 2 ^ n" by (rule of_nat_n_less_equal_power_2, clarsimp simp: word_size) lemma word_rsplit_0: "word_rsplit (0 :: word32) = [0, 0, 0, 0 :: 8 word]" by (simp add: word_rsplit_def bin_rsplit_def) lemma unat_ucast_10_32 : fixes x :: "10 word" shows "unat (ucast x :: word32) = unat x" by transfer simp lemma bool_mask [simp]: fixes x :: word32 shows "(0 < x && 1) = (x && 1 = 1)" by (rule bool_mask') auto lemma word32_bounds: "- (2 ^ (size (x :: word32) - 1)) = (-2147483648 :: int)" "((2 ^ (size (x :: word32) - 1)) - 1) = (2147483647 :: int)" "- (2 ^ (size (y :: 32 signed word) - 1)) = (-2147483648 :: int)" "((2 ^ (size (y :: 32 signed word) - 1)) - 1) = (2147483647 :: int)" by (simp_all add: word_size) lemma word_ge_min:"sint (x::32 word) \ -2147483648" by (metis sint_ge word32_bounds(1) word_size) lemmas signed_arith_ineq_checks_to_eq_word32' = signed_arith_ineq_checks_to_eq[where 'a=32] signed_arith_ineq_checks_to_eq[where 'a="32 signed"] lemmas signed_arith_ineq_checks_to_eq_word32 = signed_arith_ineq_checks_to_eq_word32' [unfolded word32_bounds] lemmas signed_mult_eq_checks32_to_64' = signed_mult_eq_checks_double_size[where 'a=32 and 'b=64] signed_mult_eq_checks_double_size[where 'a="32 signed" and 'b=64] lemmas signed_mult_eq_checks32_to_64 = signed_mult_eq_checks32_to_64'[simplified] lemmas sdiv_word32_max' = sdiv_word_max [where 'a=32] sdiv_word_max [where 'a="32 signed"] lemmas sdiv_word32_max = sdiv_word32_max'[simplified word_size, simplified] lemmas sdiv_word32_min' = sdiv_word_min [where 'a=32] sdiv_word_min [where 'a="32 signed"] lemmas sdiv_word32_min = sdiv_word32_min' [simplified word_size, simplified] lemmas sint32_of_int_eq' = sint_of_int_eq [where 'a=32] lemmas sint32_of_int_eq = sint32_of_int_eq' [simplified] lemma ucast_of_nats [simp]: "(ucast (of_nat x :: word32) :: sword32) = (of_nat x)" "(ucast (of_nat x :: word32) :: 16 sword) = (of_nat x)" "(ucast (of_nat x :: word32) :: 8 sword) = (of_nat x)" "(ucast (of_nat x :: 16 word) :: 16 sword) = (of_nat x)" "(ucast (of_nat x :: 16 word) :: 8 sword) = (of_nat x)" "(ucast (of_nat x :: 8 word) :: 8 sword) = (of_nat x)" - by (simp_all add: of_nat_take_bit take_bit_word_eq_self) + by (simp_all add: of_nat_take_bit take_bit_word_eq_self unsigned_of_nat) lemmas signed_shift_guard_simpler_32' = power_strict_increasing_iff[where b="2 :: nat" and y=31] lemmas signed_shift_guard_simpler_32 = signed_shift_guard_simpler_32'[simplified] lemma word32_31_less: "31 < len_of TYPE (32 signed)" "31 > (0 :: nat)" "31 < len_of TYPE (32)" "31 > (0 :: nat)" by auto lemmas signed_shift_guard_to_word_32 = signed_shift_guard_to_word[OF word32_31_less(1-2)] signed_shift_guard_to_word[OF word32_31_less(3-4)] lemma le_step_down_word_3: fixes x :: "32 word" shows "\x \ y; x \ y; y < 2 ^ 32 - 1\ \ x \ y - 1" by (rule le_step_down_word_2, assumption+) lemma shiftr_1: "(x::word32) >> 1 = 0 \ x < 2" by transfer (simp add: take_bit_drop_bit drop_bit_Suc) lemma has_zero_byte: "~~ (((((v::word32) && 0x7f7f7f7f) + 0x7f7f7f7f) || v) || 0x7f7f7f7f) \ 0 \ v && 0xff000000 = 0 \ v && 0xff0000 = 0 \ v && 0xff00 = 0 \ v && 0xff = 0" by word_bitwise auto lemma mask_step_down_32: \\x. mask x = b\ if \b && 1 = 1\ and \\x. x < 32 \ mask x = b >> 1\ for b :: \32word\ proof - from \b && 1 = 1\ have \odd b\ by (auto simp add: mod_2_eq_odd and_one_eq) then have \b mod 2 = 1\ using odd_iff_mod_2_eq_one by blast from \\x. x < 32 \ mask x = b >> 1\ obtain x where \x < 32\ \mask x = b >> 1\ by blast then have \mask x = b div 2\ using shiftr1_is_div_2 [of b] by simp with \b mod 2 = 1\ have \2 * mask x + 1 = 2 * (b div 2) + b mod 2\ by (simp only:) also have \\ = b\ by (simp add: mult_div_mod_eq) finally have \2 * mask x + 1 = b\ . moreover have \mask (Suc x) = 2 * mask x + (1 :: 'a::len word)\ by (simp add: mask_Suc_rec) ultimately show ?thesis by auto qed lemma unat_of_int_32: "\i \ 0; i \2 ^ 31\ \ (unat ((of_int i)::sword32)) = nat i" - unfolding unat_eq_nat_uint - apply (subst eq_nat_nat_iff) - apply (auto simp add: take_bit_int_eq_self) - done + by (simp add: unsigned_of_int nat_take_bit_eq take_bit_nat_eq_self) lemmas word_ctz_not_minus_1_32 = word_ctz_not_minus_1[where 'a=32, simplified] (* Helper for packing then unpacking a 64-bit variable. *) lemma cast_chunk_assemble_id_64[simp]: "(((ucast ((ucast (x::64 word))::32 word))::64 word) || (((ucast ((ucast (x >> 32))::32 word))::64 word) << 32)) = x" by (simp add:cast_chunk_assemble_id) (* Another variant of packing and unpacking a 64-bit variable. *) lemma cast_chunk_assemble_id_64'[simp]: "(((ucast ((scast (x::64 word))::32 word))::64 word) || (((ucast ((scast (x >> 32))::32 word))::64 word) << 32)) = x" by (simp add:cast_chunk_scast_assemble_id) (* Specialisations of down_cast_same for adding to local simpsets. *) lemma cast_down_u64: "(scast::64 word \ 32 word) = (ucast::64 word \ 32 word)" by (subst down_cast_same[symmetric]; simp add:is_down)+ lemma cast_down_s64: "(scast::64 sword \ 32 word) = (ucast::64 sword \ 32 word)" by (subst down_cast_same[symmetric]; simp add:is_down) lemma word32_and_max_simp: \x AND 0xFFFFFFFF = x\ for x :: \32 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) end end diff --git a/thys/Word_Lib/Word_64.thy b/thys/Word_Lib/Word_64.thy --- a/thys/Word_Lib/Word_64.thy +++ b/thys/Word_Lib/Word_64.thy @@ -1,312 +1,309 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Words of Length 64" theory Word_64 imports Word_Lemmas Word_Names Word_Syntax Rsplit More_Word_Operations begin context includes bit_operations_syntax begin lemma len64: "len_of (x :: 64 itself) = 64" by simp type_synonym machine_word_len = 64 type_synonym machine_word = "machine_word_len word" definition word_bits :: nat where "word_bits = LENGTH(machine_word_len)" text \The following two are numerals so they can be used as nats and words.\ definition word_size_bits :: "'a :: numeral" where "word_size_bits = 3" definition word_size :: "'a :: numeral" where "word_size = 8" lemma word_bits_conv[code]: "word_bits = 64" unfolding word_bits_def by simp lemma word_size_word_size_bits: "(word_size::nat) = 2 ^ word_size_bits" unfolding word_size_def word_size_bits_def by simp lemma word_bits_word_size_conv: "word_bits = word_size * 8" unfolding word_bits_def word_size_def by simp lemma ucast_8_64_inj: "inj (ucast :: 8 word \ 64 word)" by (rule down_ucast_inj) (clarsimp simp: is_down_def target_size source_size) lemma upto_2_helper: "{0..<2 :: 64 word} = {0, 1}" by (safe; simp) unat_arith lemmas upper_bits_unset_is_l2p_64 = upper_bits_unset_is_l2p [where 'a=64, folded word_bits_def] lemmas le_2p_upper_bits_64 = le_2p_upper_bits [where 'a=64, folded word_bits_def] lemmas le2p_bits_unset_64 = le2p_bits_unset[where 'a=64, folded word_bits_def] lemma word_bits_len_of: "len_of TYPE (64) = word_bits" by (simp add: word_bits_conv) lemmas unat_power_lower64' = unat_power_lower[where 'a=64] lemmas unat_power_lower64 [simp] = unat_power_lower64'[unfolded word_bits_len_of] lemmas word64_less_sub_le' = word_less_sub_le[where 'a = 64] lemmas word64_less_sub_le[simp] = word64_less_sub_le' [folded word_bits_def] lemma word_bits_size: "size (w::word64) = word_bits" by (simp add: word_bits_def word_size) lemmas word64_power_less_1' = word_power_less_1[where 'a = 64] lemmas word64_power_less_1[simp] = word64_power_less_1'[folded word_bits_def] lemma of_nat64_0: "\of_nat n = (0::word64); n < 2 ^ word_bits\ \ n = 0" by (erule of_nat_0, simp add: word_bits_def) lemma unat_mask_2_less_4: "unat (p && mask 2 :: word64) < 4" by (rule unat_less_helper) (simp flip: take_bit_eq_mask add: take_bit_eq_mod word_mod_less_divisor) lemmas unat_of_nat64' = unat_of_nat_eq[where 'a=64] lemmas unat_of_nat64 = unat_of_nat64'[unfolded word_bits_len_of] lemmas word_power_nonzero_64 = word_power_nonzero [where 'a=64, folded word_bits_def] lemmas unat_mult_simple = iffD1 [OF unat_mult_lem [where 'a = 64, unfolded word_bits_len_of]] lemmas div_power_helper_64 = div_power_helper [where 'a=64, folded word_bits_def] lemma n_less_word_bits: "(n < word_bits) = (n < 64)" by (simp add: word_bits_def) lemmas of_nat_less_pow_64 = of_nat_power [where 'a=64, folded word_bits_def] lemma lt_word_bits_lt_pow: "sz < word_bits \ sz < 2 ^ word_bits" by (simp add: word_bits_conv) lemma unat_less_word_bits: fixes y :: word64 shows "x < unat y \ x < 2 ^ word_bits" unfolding word_bits_def by (rule order_less_trans [OF _ unat_lt2p]) lemmas unat_mask_word64' = unat_mask[where 'a=64] lemmas unat_mask_word64 = unat_mask_word64'[folded word_bits_def] lemma unat_less_2p_word_bits: "unat (x :: 64 word) < 2 ^ word_bits" apply (simp only: word_bits_def) apply (rule unat_lt2p) done lemma Suc_unat_mask_div: "Suc (unat (mask sz div word_size::word64)) = 2 ^ (min sz word_bits - 3)" proof (cases \sz \ 3\) case False then have \sz \ {0, 1, 2}\ by auto then show ?thesis by (auto simp add: unat_div word_size_def unat_mask) next case True moreover define n where \n = sz - 3\ ultimately have \sz = n + 3\ by simp moreover have \2 ^ n * 8 - Suc 0 = (2 ^ n - 1) * 8 + 7\ by (simp add: mult_eq_if) ultimately show ?thesis by (simp add: unat_div unat_mask word_size_def word_bits_def min_def power_add) qed lemmas word64_minus_one_le' = word_minus_one_le[where 'a=64] lemmas word64_minus_one_le = word64_minus_one_le'[simplified] lemma ucast_not_helper: fixes a::"8 word" assumes a: "a \ 0xFF" shows "ucast a \ (0xFF::word64)" proof assume "ucast a = (0xFF::word64)" also have "(0xFF::word64) = ucast (0xFF::8 word)" by simp finally show False using a apply - apply (drule up_ucast_inj, simp) apply simp done qed lemma less_4_cases: "(x::word64) < 4 \ x=0 \ x=1 \ x=2 \ x=3" apply clarsimp apply (drule word_less_cases, erule disjE, simp, simp)+ done lemma if_then_1_else_0: "((if P then 1 else 0) = (0 :: word64)) = (\ P)" by simp lemma if_then_0_else_1: "((if P then 0 else 1) = (0 :: word64)) = (P)" by simp lemmas if_then_simps = if_then_0_else_1 if_then_1_else_0 lemma ucast_le_ucast_8_64: "(ucast x \ (ucast y :: word64)) = (x \ (y :: 8 word))" by (simp add: ucast_le_ucast) lemma in_16_range: "0 \ S \ r \ (\x. r + x * (16 :: word64)) ` S" "n - 1 \ S \ (r + (16 * n - 16)) \ (\x :: word64. r + x * 16) ` S" by (clarsimp simp: image_def elim!: bexI[rotated])+ lemma eq_2_64_0: "(2 ^ 64 :: word64) = 0" by simp lemma x_less_2_0_1: fixes x :: word64 shows "x < 2 \ x = 0 \ x = 1" by (rule x_less_2_0_1') auto lemmas mask_64_max_word = max_word_mask [symmetric, where 'a=64, simplified] lemma of_nat64_n_less_equal_power_2: "n < 64 \ ((of_nat n)::64 word) < 2 ^ n" by (rule of_nat_n_less_equal_power_2, clarsimp simp: word_size) lemma word_rsplit_0: "word_rsplit (0 :: word64) = [0, 0, 0, 0, 0, 0, 0, 0 :: 8 word]" by (simp add: word_rsplit_def bin_rsplit_def) lemma unat_ucast_10_64 : fixes x :: "10 word" shows "unat (ucast x :: word64) = unat x" by transfer simp lemma bool_mask [simp]: fixes x :: word64 shows "(0 < x && 1) = (x && 1 = 1)" by (rule bool_mask') auto lemma word64_bounds: "- (2 ^ (size (x :: word64) - 1)) = (-9223372036854775808 :: int)" "((2 ^ (size (x :: word64) - 1)) - 1) = (9223372036854775807 :: int)" "- (2 ^ (size (y :: 64 signed word) - 1)) = (-9223372036854775808 :: int)" "((2 ^ (size (y :: 64 signed word) - 1)) - 1) = (9223372036854775807 :: int)" by (simp_all add: word_size) lemma word_ge_min:"sint (x::64 word) \ -9223372036854775808" by (metis sint_ge word64_bounds(1) word_size) lemmas signed_arith_ineq_checks_to_eq_word64' = signed_arith_ineq_checks_to_eq[where 'a=64] signed_arith_ineq_checks_to_eq[where 'a="64 signed"] lemmas signed_arith_ineq_checks_to_eq_word64 = signed_arith_ineq_checks_to_eq_word64' [unfolded word64_bounds] lemmas signed_mult_eq_checks64_to_64' = signed_mult_eq_checks_double_size[where 'a=64 and 'b=64] signed_mult_eq_checks_double_size[where 'a="64 signed" and 'b=64] lemmas signed_mult_eq_checks64_to_64 = signed_mult_eq_checks64_to_64'[simplified] lemmas sdiv_word64_max' = sdiv_word_max [where 'a=64] sdiv_word_max [where 'a="64 signed"] lemmas sdiv_word64_max = sdiv_word64_max'[simplified word_size, simplified] lemmas sdiv_word64_min' = sdiv_word_min [where 'a=64] sdiv_word_min [where 'a="64 signed"] lemmas sdiv_word64_min = sdiv_word64_min' [simplified word_size, simplified] lemmas sint64_of_int_eq' = sint_of_int_eq [where 'a=64] lemmas sint64_of_int_eq = sint64_of_int_eq' [simplified] lemma ucast_of_nats [simp]: "(ucast (of_nat x :: word64) :: sword64) = (of_nat x)" "(ucast (of_nat x :: word64) :: 16 sword) = (of_nat x)" "(ucast (of_nat x :: word64) :: 8 sword) = (of_nat x)" - by (simp_all add: of_nat_take_bit take_bit_word_eq_self) + by (simp_all add: of_nat_take_bit take_bit_word_eq_self unsigned_of_nat) lemmas signed_shift_guard_simpler_64' = power_strict_increasing_iff[where b="2 :: nat" and y=31] lemmas signed_shift_guard_simpler_64 = signed_shift_guard_simpler_64'[simplified] lemma word64_31_less: "31 < len_of TYPE (64 signed)" "31 > (0 :: nat)" "31 < len_of TYPE (64)" "31 > (0 :: nat)" by auto lemmas signed_shift_guard_to_word_64 = signed_shift_guard_to_word[OF word64_31_less(1-2)] signed_shift_guard_to_word[OF word64_31_less(3-4)] lemma le_step_down_word_3: fixes x :: "64 word" shows "\x \ y; x \ y; y < 2 ^ 64 - 1\ \ x \ y - 1" by (rule le_step_down_word_2, assumption+) lemma shiftr_1: "(x::word64) >> 1 = 0 \ x < 2" by transfer (simp add: take_bit_drop_bit drop_bit_Suc) lemma mask_step_down_64: \\x. mask x = b\ if \b && 1 = 1\ and \\x. x < 64 \ mask x = b >> 1\ for b :: \64word\ proof - from \b && 1 = 1\ have \odd b\ by (auto simp add: mod_2_eq_odd and_one_eq) then have \b mod 2 = 1\ using odd_iff_mod_2_eq_one by blast from \\x. x < 64 \ mask x = b >> 1\ obtain x where \x < 64\ \mask x = b >> 1\ by blast then have \mask x = b div 2\ using shiftr1_is_div_2 [of b] by simp with \b mod 2 = 1\ have \2 * mask x + 1 = 2 * (b div 2) + b mod 2\ by (simp only:) also have \\ = b\ by (simp add: mult_div_mod_eq) finally have \2 * mask x + 1 = b\ . moreover have \mask (Suc x) = 2 * mask x + (1 :: 'a::len word)\ by (simp add: mask_Suc_rec) ultimately show ?thesis by auto qed lemma unat_of_int_64: "\i \ 0; i \ 2 ^ 63\ \ (unat ((of_int i)::sword64)) = nat i" - unfolding unat_eq_nat_uint - apply (subst eq_nat_nat_iff) - apply (simp_all add: take_bit_int_eq_self) - done + by (simp add: unsigned_of_int nat_take_bit_eq take_bit_nat_eq_self) lemmas word_ctz_not_minus_1_64 = word_ctz_not_minus_1[where 'a=64, simplified] lemma word64_and_max_simp: \x AND 0xFFFFFFFFFFFFFFFF = x\ for x :: \64 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) end end diff --git a/thys/Word_Lib/Word_8.thy b/thys/Word_Lib/Word_8.thy --- a/thys/Word_Lib/Word_8.thy +++ b/thys/Word_Lib/Word_8.thy @@ -1,114 +1,114 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Words of Length 8" theory Word_8 imports More_Word Enumeration_Word Even_More_List Signed_Words Word_Lemmas begin context includes bit_operations_syntax begin lemma len8: "len_of (x :: 8 itself) = 8" by simp lemma word8_and_max_simp: \x AND 0xFF = x\ for x :: \8 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) lemma enum_word8_eq: \enum = [0 :: 8 word, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27, 28, 29, 30, 31, 32, 33, 34, 35, 36, 37, 38, 39, 40, 41, 42, 43, 44, 45, 46, 47, 48, 49, 50, 51, 52, 53, 54, 55, 56, 57, 58, 59, 60, 61, 62, 63, 64, 65, 66, 67, 68, 69, 70, 71, 72, 73, 74, 75, 76, 77, 78, 79, 80, 81, 82, 83, 84, 85, 86, 87, 88, 89, 90, 91, 92, 93, 94, 95, 96, 97, 98, 99, 100, 101, 102, 103, 104, 105, 106, 107, 108, 109, 110, 111, 112, 113, 114, 115, 116, 117, 118, 119, 120, 121, 122, 123, 124, 125, 126, 127, 128, 129, 130, 131, 132, 133, 134, 135, 136, 137, 138, 139, 140, 141, 142, 143, 144, 145, 146, 147, 148, 149, 150, 151, 152, 153, 154, 155, 156, 157, 158, 159, 160, 161, 162, 163, 164, 165, 166, 167, 168, 169, 170, 171, 172, 173, 174, 175, 176, 177, 178, 179, 180, 181, 182, 183, 184, 185, 186, 187, 188, 189, 190, 191, 192, 193, 194, 195, 196, 197, 198, 199, 200, 201, 202, 203, 204, 205, 206, 207, 208, 209, 210, 211, 212, 213, 214, 215, 216, 217, 218, 219, 220, 221, 222, 223, 224, 225, 226, 227, 228, 229, 230, 231, 232, 233, 234, 235, 236, 237, 238, 239, 240, 241, 242, 243, 244, 245, 246, 247, 248, 249, 250, 251, 252, 253, 254, 255]\ (is \?lhs = ?rhs\) proof - have \map unat ?lhs = [0..<256]\ - by (simp add: enum_word_def comp_def take_bit_nat_eq_self map_idem_upt_eq) + by (simp add: enum_word_def comp_def take_bit_nat_eq_self map_idem_upt_eq unsigned_of_nat) also have \\ = map unat ?rhs\ by (simp add: upt_zero_numeral_unfold) finally show ?thesis using unat_inj by (rule map_injective) qed lemma set_enum_word8_def: "(set enum :: 8 word set) = {0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27, 28, 29, 30, 31, 32, 33, 34, 35, 36, 37, 38, 39, 40, 41, 42, 43, 44, 45, 46, 47, 48, 49, 50, 51, 52, 53, 54, 55, 56, 57, 58, 59, 60, 61, 62, 63, 64, 65, 66, 67, 68, 69, 70, 71, 72, 73, 74, 75, 76, 77, 78, 79, 80, 81, 82, 83, 84, 85, 86, 87, 88, 89, 90, 91, 92, 93, 94, 95, 96, 97, 98, 99, 100, 101, 102, 103, 104, 105, 106, 107, 108, 109, 110, 111, 112, 113, 114, 115, 116, 117, 118, 119, 120, 121, 122, 123, 124, 125, 126, 127, 128, 129, 130, 131, 132, 133, 134, 135, 136, 137, 138, 139, 140, 141, 142, 143, 144, 145, 146, 147, 148, 149, 150, 151, 152, 153, 154, 155, 156, 157, 158, 159, 160, 161, 162, 163, 164, 165, 166, 167, 168, 169, 170, 171, 172, 173, 174, 175, 176, 177, 178, 179, 180, 181, 182, 183, 184, 185, 186, 187, 188, 189, 190, 191, 192, 193, 194, 195, 196, 197, 198, 199, 200, 201, 202, 203, 204, 205, 206, 207, 208, 209, 210, 211, 212, 213, 214, 215, 216, 217, 218, 219, 220, 221, 222, 223, 224, 225, 226, 227, 228, 229, 230, 231, 232, 233, 234, 235, 236, 237, 238, 239, 240, 241, 242, 243, 244, 245, 246, 247, 248, 249, 250, 251, 252, 253, 254, 255}" by (simp add: enum_word8_eq) lemma set_strip_insert: "\ x \ insert a S; x \ a \ \ x \ S" by simp lemma word8_exhaust: fixes x :: \8 word\ shows "\x \ 0; x \ 1; x \ 2; x \ 3; x \ 4; x \ 5; x \ 6; x \ 7; x \ 8; x \ 9; x \ 10; x \ 11; x \ 12; x \ 13; x \ 14; x \ 15; x \ 16; x \ 17; x \ 18; x \ 19; x \ 20; x \ 21; x \ 22; x \ 23; x \ 24; x \ 25; x \ 26; x \ 27; x \ 28; x \ 29; x \ 30; x \ 31; x \ 32; x \ 33; x \ 34; x \ 35; x \ 36; x \ 37; x \ 38; x \ 39; x \ 40; x \ 41; x \ 42; x \ 43; x \ 44; x \ 45; x \ 46; x \ 47; x \ 48; x \ 49; x \ 50; x \ 51; x \ 52; x \ 53; x \ 54; x \ 55; x \ 56; x \ 57; x \ 58; x \ 59; x \ 60; x \ 61; x \ 62; x \ 63; x \ 64; x \ 65; x \ 66; x \ 67; x \ 68; x \ 69; x \ 70; x \ 71; x \ 72; x \ 73; x \ 74; x \ 75; x \ 76; x \ 77; x \ 78; x \ 79; x \ 80; x \ 81; x \ 82; x \ 83; x \ 84; x \ 85; x \ 86; x \ 87; x \ 88; x \ 89; x \ 90; x \ 91; x \ 92; x \ 93; x \ 94; x \ 95; x \ 96; x \ 97; x \ 98; x \ 99; x \ 100; x \ 101; x \ 102; x \ 103; x \ 104; x \ 105; x \ 106; x \ 107; x \ 108; x \ 109; x \ 110; x \ 111; x \ 112; x \ 113; x \ 114; x \ 115; x \ 116; x \ 117; x \ 118; x \ 119; x \ 120; x \ 121; x \ 122; x \ 123; x \ 124; x \ 125; x \ 126; x \ 127; x \ 128; x \ 129; x \ 130; x \ 131; x \ 132; x \ 133; x \ 134; x \ 135; x \ 136; x \ 137; x \ 138; x \ 139; x \ 140; x \ 141; x \ 142; x \ 143; x \ 144; x \ 145; x \ 146; x \ 147; x \ 148; x \ 149; x \ 150; x \ 151; x \ 152; x \ 153; x \ 154; x \ 155; x \ 156; x \ 157; x \ 158; x \ 159; x \ 160; x \ 161; x \ 162; x \ 163; x \ 164; x \ 165; x \ 166; x \ 167; x \ 168; x \ 169; x \ 170; x \ 171; x \ 172; x \ 173; x \ 174; x \ 175; x \ 176; x \ 177; x \ 178; x \ 179; x \ 180; x \ 181; x \ 182; x \ 183; x \ 184; x \ 185; x \ 186; x \ 187; x \ 188; x \ 189; x \ 190; x \ 191; x \ 192; x \ 193; x \ 194; x \ 195; x \ 196; x \ 197; x \ 198; x \ 199; x \ 200; x \ 201; x \ 202; x \ 203; x \ 204; x \ 205; x \ 206; x \ 207; x \ 208; x \ 209; x \ 210; x \ 211; x \ 212; x \ 213; x \ 214; x \ 215; x \ 216; x \ 217; x \ 218; x \ 219; x \ 220; x \ 221; x \ 222; x \ 223; x \ 224; x \ 225; x \ 226; x \ 227; x \ 228; x \ 229; x \ 230; x \ 231; x \ 232; x \ 233; x \ 234; x \ 235; x \ 236; x \ 237; x \ 238; x \ 239; x \ 240; x \ 241; x \ 242; x \ 243; x \ 244; x \ 245; x \ 246; x \ 247; x \ 248; x \ 249; x \ 250; x \ 251; x \ 252; x \ 253; x \ 254; x \ 255\ \ P" apply (subgoal_tac "x \ set enum", subst (asm) set_enum_word8_def) apply (drule set_strip_insert, assumption)+ apply (erule emptyE) apply (subst enum_UNIV, rule UNIV_I) done end end diff --git a/thys/Word_Lib/Word_Lemmas.thy b/thys/Word_Lib/Word_Lemmas.thy --- a/thys/Word_Lib/Word_Lemmas.thy +++ b/thys/Word_Lib/Word_Lemmas.thy @@ -1,1944 +1,1940 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Lemmas with Generic Word Length" theory Word_Lemmas imports Type_Syntax Signed_Division_Word Signed_Words More_Word Most_significant_bit Enumeration_Word Aligned Bit_Shifts_Infix_Syntax begin context includes bit_operations_syntax begin lemma ucast_le_ucast_eq: fixes x y :: "'a::len word" assumes x: "x < 2 ^ n" assumes y: "y < 2 ^ n" assumes n: "n = LENGTH('b::len)" shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" apply (rule iffI) apply (cases "LENGTH('b) < LENGTH('a)") apply (subst less_mask_eq[OF x, symmetric]) apply (subst less_mask_eq[OF y, symmetric]) apply (unfold n) apply (subst ucast_ucast_mask[symmetric])+ apply (simp add: ucast_le_ucast)+ apply (erule ucast_mono_le[OF _ y[unfolded n]]) done lemma ucast_zero_is_aligned: \is_aligned w n\ if \UCAST('a::len \ 'b::len) w = 0\ \n \ LENGTH('b)\ proof (rule is_aligned_bitI) fix q assume \q < n\ moreover have \bit (UCAST('a::len \ 'b::len) w) q = bit 0 q\ using that by simp with \q < n\ \n \ LENGTH('b)\ show \\ bit w q\ by (simp add: bit_simps) qed lemma unat_ucast_eq_unat_and_mask: "unat (UCAST('b::len \ 'a::len) w) = unat (w AND mask LENGTH('a))" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: ac_simps) done lemma le_max_word_ucast_id: \UCAST('b \ 'a) (UCAST('a \ 'b) x) = x\ if \x \ UCAST('b::len \ 'a) (- 1)\ for x :: \'a::len word\ proof - from that have a1: \x \ word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1) :: 'b word))\ by simp have f2: "((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (0::int) \ - 1 + 2 ^ LENGTH('b) \ (0::int) \ - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \ (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) = - 1 + 2 ^ LENGTH('b)) = ((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (1::int) \ 2 ^ LENGTH('b) \ 2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)) = 1)" by force have f3: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * ia \ i mod ia = i" using mod_pos_pos_trivial by force have "(1::int) \ 2 ^ LENGTH('b)" by simp then have "2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ len_of TYPE ('b)) = 1" using f3 f2 by blast then have f4: "- (1::int) + 2 ^ LENGTH('b) = (- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)" by linarith have f5: "x \ word_of_int (uint (word_of_int (- 1 + 2 ^ LENGTH('b))::'b word))" using a1 by force have f6: "2 ^ LENGTH('b) + - (1::int) = - 1 + 2 ^ LENGTH('b)" by force have f7: "- (1::int) * 1 = - 1" by auto have "\x0 x1. (x1::int) - x0 = x1 + - 1 * x0" by force then have "x \ 2 ^ LENGTH('b) - 1" using f7 f6 f5 f4 by (metis uint_word_of_int wi_homs(2) word_arith_wis(8) word_of_int_2p) then have \uint x \ uint (2 ^ LENGTH('b) - (1 :: 'a word))\ by (simp add: word_le_def) then have \uint x \ 2 ^ LENGTH('b) - 1\ by (simp add: uint_word_ariths) (metis \1 \ 2 ^ LENGTH('b)\ \uint x \ uint (2 ^ LENGTH('b) - 1)\ linorder_not_less lt2p_lem uint_1 uint_minus_simple_alt uint_power_lower word_le_def zle_diff1_eq) then show ?thesis apply (simp add: unsigned_ucast_eq take_bit_word_eq_self_iff) apply (meson \x \ 2 ^ LENGTH('b) - 1\ not_le word_less_sub_le) done qed lemma uint_shiftr_eq: \uint (w >> n) = uint w div 2 ^ n\ by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit min_def le_less less_diff_conv) lemma bit_shiftl_word_iff [bit_simps]: \bit (w << m) n \ m \ n \ n < LENGTH('a) \ bit w (n - m)\ for w :: \'a::len word\ by (simp add: bit_push_bit_iff not_le) lemma shiftl_def: \w << n = ((*) 2 ^^ n) w\ for w :: \'a::len word\ proof - have \push_bit n = (((*) 2 ^^ n) :: int \ int)\ for n by (induction n) (simp_all add: fun_eq_iff funpow_swap1, simp add: ac_simps) then show ?thesis by transfer simp qed lemma shiftr_def: \w >> n = ((\w. w div 2) ^^ n) w\ for w :: \'a::len word\ proof - have \(\w. w div 2) ^^ n = (drop_bit n :: 'a word \ 'a word)\ by (induction n) (simp_all add: drop_bit_half drop_bit_Suc) then show ?thesis by simp qed lemma bit_shiftr_word_iff: \bit (w >> m) n \ bit w (m + n)\ for w :: \'a::len word\ by (simp add: bit_simps) lemma sshiftr_eq_funpow_sshiftr1: \w >>> n = (signed_drop_bit (Suc 0) ^^ n) w\ apply (rule sym) apply (induction n) apply simp_all done lemma uint_sshiftr_eq: \uint (w >>> n) = take_bit LENGTH('a) (sint w div 2 ^ n)\ for w :: \'a::len word\ by transfer (simp flip: drop_bit_eq_div) lemma sshiftr_0: "0 >>> n = 0" by (fact signed_drop_bit_of_0) lemma sshiftr_n1: "-1 >>> n = -1" by (fact signed_drop_bit_of_minus_1) lemma bit_sshiftr_word_iff: \bit (w >>> m) n \ bit w (if LENGTH('a) - m \ n \ n < LENGTH('a) then LENGTH('a) - 1 else (m + n))\ for w :: \'a::len word\ by (fact bit_signed_drop_bit_iff) lemma nth_sshiftr : "bit (w >>> m) n = (n < size w \ (if n + m \ size w then bit w (size w - 1) else bit w (n + m)))" apply (auto simp add: bit_simps word_size ac_simps not_less) apply (meson bit_imp_le_length bit_shiftr_word_iff leD) done lemma sshiftr_numeral: \(numeral k >>> numeral n :: 'a::len word) = word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (numeral k)))\ by (fact signed_drop_bit_word_numeral) lemma sshiftr_div_2n: "sint (w >>> n) = sint w div 2 ^ n" using sint_signed_drop_bit_eq [of n w] by (simp add: drop_bit_eq_div) lemma mask_eq: \mask n = (1 << n) - (1 :: 'a::len word)\ by (simp add: mask_eq_exp_minus_1 push_bit_of_1) lemma shiftl_0: "(0::'a::len word) << n = 0" by (fact push_bit_of_0) lemma shiftr_0: "(0::'a::len word) >> n = 0" by (fact drop_bit_of_0) lemma nth_shiftl': "bit (w << m) n \ n < size w \ n >= m \ bit w (n - m)" for w :: "'a::len word" by transfer (auto simp add: bit_push_bit_iff) lemmas nth_shiftl = nth_shiftl' [unfolded word_size] lemma nth_shiftr: "bit (w >> m) n = bit w (n + m)" for w :: "'a::len word" by (simp add: bit_simps ac_simps) lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" by (fact uint_shiftr_eq) lemma shiftl_rev: "shiftl w n = word_reverse (shiftr (word_reverse w) n)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma rev_shiftl: "word_reverse w << n = word_reverse (w >> n)" by (simp add: shiftl_rev) lemma shiftr_rev: "w >> n = word_reverse (word_reverse w << n)" by (simp add: rev_shiftl) lemma rev_shiftr: "word_reverse w >> n = word_reverse (w << n)" by (simp add: shiftr_rev) lemmas ucast_up = rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] lemmas ucast_down = rc2 [simplified rev_shiftr revcast_ucast [symmetric]] lemma shiftl_zero_size: "size x \ n \ x << n = 0" for x :: "'a::len word" apply transfer apply (simp add: take_bit_push_bit) done lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" for w :: "'a::len word" by (simp add: push_bit_eq_mult) lemma word_shift_by_2: "x * 4 = (x::'a::len word) << 2" by (simp add: shiftl_t2n) lemma slice_shiftr: "slice n w = ucast (w >> n)" apply (rule bit_word_eqI) apply (cases \n \ LENGTH('b)\) apply (auto simp add: bit_slice_iff bit_ucast_iff bit_shiftr_word_iff ac_simps dest: bit_imp_le_length) done lemma shiftr_zero_size: "size x \ n \ x >> n = 0" for x :: "'a :: len word" by (rule word_eqI) (auto simp add: nth_shiftr dest: test_bit_size) lemma shiftr_x_0: "x >> 0 = x" for x :: "'a::len word" by simp lemma shiftl_x_0: "x << 0 = x" for x :: "'a::len word" by simp lemma shiftl_1: "(1::'a::len word) << n = 2^n" by (fact push_bit_of_1) lemma shiftr_1: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" by simp lemma shiftl0: "x << 0 = (x :: 'a :: len word)" by (fact shiftl_x_0) lemma and_not_mask: "w AND NOT (mask n) = (w >> n) << n" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma and_mask: "w AND mask n = (w << (size w - n)) >> (size w - n)" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps word_size) lemma shiftr_div_2n_w: "n < size w \ w >> n = w div (2^n :: 'a :: len word)" apply (unfold word_div_def) apply (simp add: uint_2p_alt word_size) apply (metis uint_shiftr_eq word_of_int_uint) done lemma le_shiftr: "u \ v \ u >> (n :: nat) \ (v :: 'a :: len word) >> n" apply transfer apply (simp add: take_bit_drop_bit) apply (simp add: drop_bit_eq_div zdiv_mono1) done lemma le_shiftr': "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" apply (metis le_cases le_shiftr verit_la_disequality) done lemma shiftr_mask_le: "n <= m \ mask n >> m = (0 :: 'a::len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma shiftr_mask [simp]: \mask m >> m = (0::'a::len word)\ by (rule shiftr_mask_le) simp lemma le_mask_iff: "(w \ mask n) = (w >> n = 0)" for w :: \'a::len word\ apply safe apply (rule word_le_0_iff [THEN iffD1]) apply (rule xtrans(3)) apply (erule_tac [2] le_shiftr) apply simp apply (rule word_leI) apply (rename_tac n') apply (drule_tac x = "n' - n" in word_eqD) apply (simp add : nth_shiftr word_size bit_simps) apply (case_tac "n <= n'") by auto lemma and_mask_eq_iff_shiftr_0: "(w AND mask n = w) = (w >> n = 0)" for w :: \'a::len word\ apply (unfold test_bit_eq_iff [THEN sym]) apply (rule iffI) apply (rule ext) apply (rule_tac [2] ext) apply (auto simp add : word_ao_nth nth_shiftr) apply (drule arg_cong) apply (drule iffD2) apply assumption apply (simp add : word_ao_nth) prefer 2 apply (simp add : word_size test_bit_bin) apply transfer apply (auto simp add: fun_eq_iff bit_simps) apply (metis add_diff_inverse_nat) done lemma mask_shiftl_decompose: "mask m << n = mask (m + n) AND NOT (mask n :: 'a::len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma shiftl_over_and_dist: fixes a::"'a::len word" shows "(a AND b) << c = (a << c) AND (b << c)" by (fact push_bit_and) lemma shiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >> c = (a >> c) AND (b >> c)" by (fact drop_bit_and) lemma sshiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >>> c = (a >>> c) AND (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemma shiftl_over_or_dist: fixes a::"'a::len word" shows "a OR b << c = (a << c) OR (b << c)" by (fact push_bit_or) lemma shiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >> c = (a >> c) OR (b >> c)" by (fact drop_bit_or) lemma sshiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >>> c = (a >>> c) OR (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemmas shift_over_ao_dists = shiftl_over_or_dist shiftr_over_or_dist sshiftr_over_or_dist shiftl_over_and_dist shiftr_over_and_dist sshiftr_over_and_dist lemma shiftl_shiftl: fixes a::"'a::len word" shows "a << b << c = a << (b + c)" apply(rule word_eqI) apply(auto simp:word_size nth_shiftl add.commute add.left_commute) done lemma shiftr_shiftr: fixes a::"'a::len word" shows "a >> b >> c = a >> (b + c)" apply(rule word_eqI) apply(simp add:word_size nth_shiftr add.left_commute add.commute) done lemma shiftl_shiftr1: fixes a::"'a::len word" shows "c \ b \ a << b >> c = a AND (mask (size a - b)) << (b - c)" apply (rule word_eqI) apply (auto simp add: bit_simps not_le word_size ac_simps) done lemma shiftl_shiftr2: fixes a::"'a::len word" shows "b < c \ a << b >> c = (a >> (c - b)) AND (mask (size a - c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth bit_simps) done lemma shiftr_shiftl1: fixes a::"'a::len word" shows "c \ b \ a >> b << c = (a >> (b - c)) AND (NOT (mask c))" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma shiftr_shiftl2: fixes a::"'a::len word" shows "b < c \ a >> b << c = (a << (c - b)) AND (NOT (mask c))" apply (rule word_eqI) apply (auto simp add: bit_simps not_le word_size ac_simps) done lemmas multi_shift_simps = shiftl_shiftl shiftr_shiftr shiftl_shiftr1 shiftl_shiftr2 shiftr_shiftl1 shiftr_shiftl2 lemma shiftr_mask2: "n \ LENGTH('a) \ (mask n >> m :: ('a :: len) word) = mask (n - m)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma word_shiftl_add_distrib: fixes x :: "'a :: len word" shows "(x + y) << n = (x << n) + (y << n)" by (simp add: shiftl_t2n ring_distribs) lemma mask_shift: "(x AND NOT (mask y)) >> y = x >> y" for x :: \'a::len word\ apply (rule bit_eqI) apply (simp add: bit_and_iff bit_not_iff bit_shiftr_word_iff bit_mask_iff not_le) using bit_imp_le_length apply auto done lemma shiftr_div_2n': "unat (w >> n) = unat w div 2 ^ n" apply (unfold unat_eq_nat_uint) apply (subst shiftr_div_2n) apply (subst nat_div_distrib) apply simp apply (simp add: nat_power_eq) done lemma shiftl_shiftr_id: assumes nv: "n < LENGTH('a)" and xv: "x < 2 ^ (LENGTH('a) - n)" shows "x << n >> n = (x::'a::len word)" apply (simp add: shiftl_t2n) apply (rule word_eq_unatI) apply (subst shiftr_div_2n') apply (cases n) apply simp apply (subst iffD1 [OF unat_mult_lem])+ apply (subst unat_power_lower[OF nv]) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl]) apply (rule unat_power_lower) apply simp apply (subst unat_power_lower[OF nv]) apply simp done lemma ucast_shiftl_eq_0: fixes w :: "'a :: len word" shows "\ n \ LENGTH('b) \ \ ucast (w << n) = (0 :: 'b :: len word)" by transfer (simp add: take_bit_push_bit) lemma word_shift_nonzero: "\ (x::'a::len word) \ 2 ^ m; m + n < LENGTH('a::len); x \ 0\ \ x << n \ 0" apply (simp only: word_neq_0_conv word_less_nat_alt shiftl_t2n mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (subst mod_less) apply (rule order_le_less_trans) apply (erule mult_le_mono2) apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done lemma word_shiftr_lt: fixes w :: "'a::len word" shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))" apply (subst shiftr_div_2n') apply transfer apply (simp flip: drop_bit_eq_div add: drop_bit_nat_eq drop_bit_take_bit) done lemma shiftr_less_t2n': "\ x AND mask (n + m) = x; m < LENGTH('a) \ \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_drop_bit ac_simps) done lemma shiftr_less_t2n: "x < 2 ^ (n + m) \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (rule shiftr_less_t2n') apply (erule less_mask_eq) apply (rule ccontr) apply (simp add: not_less) apply (subst (asm) p2_eq_0[symmetric]) apply (simp add: power_add) done lemma shiftr_eq_0: "n \ LENGTH('a) \ ((w::'a::len word) >> n) = 0" apply (cut_tac shiftr_less_t2n'[of w n 0], simp) apply (simp add: mask_eq_iff) apply (simp add: lt2p_lem) apply simp done lemma shiftl_less_t2n: fixes x :: "'a :: len word" shows "\ x < (2 ^ (m - n)); m < LENGTH('a) \ \ (x << n) < 2 ^ m" apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_push_bit) done lemma shiftl_less_t2n': "(x::'a::len word) < 2 ^ m \ m+n < LENGTH('a) \ x << n < 2 ^ (m + n)" by (rule shiftl_less_t2n) simp_all lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (rule bit_word_eqI) (simp add: bit_simps) lemma signed_shift_guard_to_word: \unat x * 2 ^ y < 2 ^ n \ x = 0 \ x < 1 << n >> y\ if \n < LENGTH('a)\ \0 < n\ for x :: \'a::len word\ proof (cases \x = 0\) case True then show ?thesis by simp next case False then have \unat x \ 0\ by (simp add: unat_eq_0) then have \unat x \ 1\ by simp show ?thesis proof (cases \y < n\) case False then have \n \ y\ by simp then obtain q where \y = n + q\ using le_Suc_ex by blast moreover have \(2 :: nat) ^ n >> n + q \ 1\ by (simp add: drop_bit_eq_div power_add) ultimately show ?thesis using \x \ 0\ \unat x \ 1\ \n < LENGTH('a)\ by (simp add: power_add not_less word_le_nat_alt unat_drop_bit_eq push_bit_of_1) next case True with that have \y < LENGTH('a)\ by simp show ?thesis proof (cases \2 ^ n = unat x * 2 ^ y\) case True moreover have \unat x * 2 ^ y < 2 ^ LENGTH('a)\ using \n < LENGTH('a)\ by (simp flip: True) moreover have \(word_of_nat (2 ^ n) :: 'a word) = word_of_nat (unat x * 2 ^ y)\ using True by simp then have \2 ^ n = x * 2 ^ y\ by simp ultimately show ?thesis using \y < LENGTH('a)\ by (auto simp add: push_bit_of_1 drop_bit_eq_div word_less_nat_alt unat_div unat_word_ariths) next case False with \y < n\ have *: \unat x \ 2 ^ n div 2 ^ y\ by (auto simp flip: power_sub power_add) have \unat x * 2 ^ y < 2 ^ n \ unat x * 2 ^ y \ 2 ^ n\ using False by (simp add: less_le) also have \\ \ unat x \ 2 ^ n div 2 ^ y\ by (simp add: less_eq_div_iff_mult_less_eq) also have \\ \ unat x < 2 ^ n div 2 ^ y\ using * by (simp add: less_le) finally show ?thesis using that \x \ 0\ by (simp flip: push_bit_eq_mult drop_bit_eq_div add: push_bit_of_1 unat_drop_bit_eq word_less_iff_unsigned [where ?'a = nat]) qed qed qed lemma shiftr_not_mask_0: "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) AND NOT (mask m) = 0" by (rule bit_word_eqI) (auto simp add: bit_simps word_size dest: bit_imp_le_length) lemma shiftl_mask_is_0[simp]: "(x << n) AND mask n = 0" for x :: \'a::len word\ by (simp flip: take_bit_eq_mask add: take_bit_push_bit) lemma rshift_sub_mask_eq: "(a >> (size a - b)) AND mask b = a >> (size a - b)" for a :: \'a::len word\ using shiftl_shiftr2[where a=a and b=0 and c="size a - b"] apply (cases "b < size a") apply simp apply (simp add: linorder_not_less mask_eq_decr_exp word_size p2_eq_0[THEN iffD2]) done lemma shiftl_shiftr3: "b \ c \ a << b >> c = (a >> c - b) AND mask (size a - c)" for a :: \'a::len word\ apply (cases "b = c") apply (simp add: shiftl_shiftr1) apply (simp add: shiftl_shiftr2) done lemma and_mask_shiftr_comm: "m \ size w \ (w AND mask m) >> n = (w >> n) AND mask (m-n)" for w :: \'a::len word\ by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) lemma and_mask_shiftl_comm: "m+n \ size w \ (w AND mask m) << n = (w << n) AND mask (m+n)" for w :: \'a::len word\ by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1) lemma le_mask_shiftl_le_mask: "s = m + n \ x \ mask n \ x << m \ mask s" for x :: \'a::len word\ by (simp add: le_mask_iff shiftl_shiftr3) lemma word_and_1_shiftl: "x AND (1 << n) = (if bit x n then (1 << n) else 0)" for x :: "'a :: len word" apply (rule bit_word_eqI; transfer) apply (auto simp add: bit_simps not_le ac_simps) done lemmas word_and_1_shiftls' = word_and_1_shiftl[where n=0] word_and_1_shiftl[where n=1] word_and_1_shiftl[where n=2] lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified] lemma word_and_mask_shiftl: "x AND (mask n << m) = ((x >> m) AND mask n) << m" for x :: \'a::len word\ apply (rule bit_word_eqI; transfer) apply (auto simp add: bit_simps not_le ac_simps) done lemma shift_times_fold: "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)" by (simp add: shiftl_t2n ac_simps power_add) lemma of_bool_nth: "of_bool (bit x v) = (x >> v) AND 1" for x :: \'a::len word\ by (simp add: bit_iff_odd_drop_bit word_and_1) lemma shiftr_mask_eq: "(x >> n) AND mask (size x - n) = x >> n" for x :: "'a :: len word" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_drop_bit) done lemma shiftr_mask_eq': "m = (size x - n) \ (x >> n) AND mask m = x >> n" for x :: "'a :: len word" by (simp add: shiftr_mask_eq) lemma and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x AND y) = 0) = (\ (bit x n))" by (simp add: and_exp_eq_0_iff_not_bit push_bit_of_1) lemma word_shift_zero: "\ x << n = 0; x \ 2^m; m + n < LENGTH('a)\ \ (x::'a::len word) = 0" apply (rule ccontr) apply (drule (2) word_shift_nonzero) apply simp done lemma mask_shift_and_negate[simp]:"(w AND mask n << m) AND NOT (mask n << m) = 0" for w :: \'a::len word\ by (rule bit_word_eqI) (simp add: bit_simps) (* The seL4 bitfield generator produces functions containing mask and shift operations, such that * invoking two of them consecutively can produce something like the following. *) lemma bitfield_op_twice: "(x AND NOT (mask n << m) OR ((y AND mask n) << m)) AND NOT (mask n << m) = x AND NOT (mask n << m)" for x :: \'a::len word\ by (induct n arbitrary: m) (auto simp: word_ao_dist) lemma bitfield_op_twice'': "\NOT a = b << c; \x. b = mask x\ \ (x AND a OR (y AND b << c)) AND a = x AND a" for a b :: \'a::len word\ apply clarsimp apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) apply (clarsimp simp:mask_eq_decr_exp) apply (drule not_switch) apply clarsimp done lemma shiftr1_unfold: "x div 2 = x >> 1" by (simp add: drop_bit_eq_div) lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" by (simp add: drop_bit_eq_div) lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2" by (metis One_nat_def mult_2 mult_2_right one_add_one power_0 power_Suc shiftl_t2n) lemma shiftr1_lt:"x \ 0 \ (x::('a::len) word) >> 1 < x" apply (subst shiftr1_is_div_2) apply (rule div_less_dividend_word) apply simp+ done lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \ x = 0 \ x = 1" apply (subst (asm) shiftr1_is_div_2) apply (drule word_less_div) apply (case_tac "LENGTH('a) = 1") apply (simp add:degenerate_word) apply (erule disjE) apply (subgoal_tac "(2::'a word) \ 0") apply simp apply (rule not_degenerate_imp_2_neq_0) apply (subgoal_tac "LENGTH('a) \ 0") apply arith apply simp apply (rule x_less_2_0_1', simp+) done lemma shiftr1_irrelevant_lsb: "bit (x::('a::len) word) 0 \ x >> 1 = (x + 1) >> 1" apply (cases \LENGTH('a)\; transfer) apply (simp_all add: take_bit_drop_bit) apply (simp add: drop_bit_take_bit drop_bit_Suc) done lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \ x = 0 \ x + 1 = 0" by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow) lemma shiftr1_irrelevant_lsb': "\ (bit (x::('a::len) word) 0) \ x >> 1 = (x + 1) >> 1" by (metis shiftr1_irrelevant_lsb) (* Perhaps this one should be a simp lemma, but it seems a little dangerous. *) lemma cast_chunk_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((ucast (x::'b word))::'a word))::'b word) OR (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n") apply clarsimp apply (subst and_not_mask[symmetric]) apply (subst ucast_ucast_mask) apply (subst word_ao_dist2[symmetric]) apply clarsimp apply (rule ucast_ucast_len) apply (rule shiftr_less_t2n') apply (subst and_mask_eq_iff_le_mask) apply (simp_all add: mask_eq_decr_exp flip: mult_2_right) apply (metis add_diff_cancel_left' len_gt_0 mult_2_right zero_less_diff) done lemma cast_chunk_scast_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((scast (x::'b word))::'a word))::'b word) OR (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)") apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)") apply (simp add:cast_chunk_assemble_id) apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+ done lemma unat_shiftr_less_t2n: fixes x :: "'a :: len word" shows "unat x < 2 ^ (n + m) \ unat (x >> n) < 2 ^ m" by (simp add: shiftr_div_2n' power_add mult.commute less_mult_imp_div_less) lemma ucast_less_shiftl_helper: "\ LENGTH('b) + 2 < LENGTH('a); 2 ^ (LENGTH('b) + 2) \ n\ \ (ucast (x :: 'b::len word) << 2) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done (* negating a mask which has been shifted to the very left *) lemma NOT_mask_shifted_lenword: "NOT (mask len << (LENGTH('a) - len) ::'a::len word) = mask (LENGTH('a) - len)" by (rule bit_word_eqI) (auto simp add: word_size bit_not_iff bit_push_bit_iff bit_mask_iff) (* Comparisons between different word sizes. *) lemma shiftr_less: "(w::'a::len word) < k \ w >> n < k" by (metis div_le_dividend le_less_trans shiftr_div_2n' unat_arith_simps(2)) lemma word_and_notzeroD: "w AND w' \ 0 \ w \ 0 \ w' \ 0" by auto lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" apply (auto simp add: take_bit_word_eq_self_iff word_less_nat_alt simp flip: take_bit_eq_self_iff_drop_bit_eq_0) apply (rule ccontr) apply (simp add: not_le) done lemma of_nat_shiftl: "(of_nat x << n) = (of_nat (x * 2 ^ n) :: ('a::len) word)" proof - have "(of_nat x::'a word) << n = of_nat (2 ^ n) * of_nat x" using shiftl_t2n by (metis word_unat_power) thus ?thesis by simp qed lemma shiftl_1_not_0: "n < LENGTH('a) \ (1::'a::len word) << n \ 0" by (simp add: shiftl_t2n) (* continue sorting out from here *) (* usually: x,y = (len_of TYPE ('a)) *) lemma bitmagic_zeroLast_leq_or1Last: "(a::('a::len) word) AND (mask len << x - len) \ a OR mask (y - len)" by (meson le_word_or2 order_trans word_and_le2) lemma zero_base_lsb_imp_set_eq_as_bit_operation: fixes base ::"'a::len word" assumes valid_prefix: "mask (LENGTH('a) - len) AND base = 0" shows "(base = NOT (mask (LENGTH('a) - len)) AND a) \ (a \ {base .. base OR mask (LENGTH('a) - len)})" proof have helper3: "x OR y = x OR y AND NOT x" for x y ::"'a::len word" by (simp add: word_oa_dist2) from assms show "base = NOT (mask (LENGTH('a) - len)) AND a \ a \ {base..base OR mask (LENGTH('a) - len)}" apply(simp add: word_and_le1) apply(metis helper3 le_word_or2 word_bw_comms(1) word_bw_comms(2)) done next assume "a \ {base..base OR mask (LENGTH('a) - len)}" hence a: "base \ a \ a \ base OR mask (LENGTH('a) - len)" by simp show "base = NOT (mask (LENGTH('a) - len)) AND a" proof - have f2: "\x\<^sub>0. base AND NOT (mask x\<^sub>0) \ a AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f3: "\x\<^sub>0. a AND NOT (mask x\<^sub>0) \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f4: "base = base AND NOT (mask (LENGTH('a) - len))" using valid_prefix by (metis mask_eq_0_eq_x word_bw_comms(1)) hence f5: "\x\<^sub>6. (base OR x\<^sub>6) AND NOT (mask (LENGTH('a) - len)) = base OR x\<^sub>6 AND NOT (mask (LENGTH('a) - len))" using word_ao_dist by (metis) have f6: "\x\<^sub>2 x\<^sub>3. a AND NOT (mask x\<^sub>2) \ x\<^sub>3 \ \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>2) \ x\<^sub>3" using f3 dual_order.trans by auto have "base = (base OR mask (LENGTH('a) - len)) AND NOT (mask (LENGTH('a) - len))" using f5 by auto hence "base = a AND NOT (mask (LENGTH('a) - len))" using f2 f4 f6 by (metis eq_iff) thus "base = NOT (mask (LENGTH('a) - len)) AND a" by (metis word_bw_comms(1)) qed qed lemma of_nat_eq_signed_scast: "(of_nat x = (y :: ('a::len) signed word)) = (of_nat x = (scast y :: 'a word))" by (metis scast_of_nat scast_scast_id(2)) lemma word_aligned_add_no_wrap_bounded: "\ w + 2^n \ x; w + 2^n \ 0; is_aligned w n \ \ (w::'a::len word) < x" by (blast dest: is_aligned_no_overflow le_less_trans word_leq_le_minus_one) lemma mask_Suc: "mask (Suc n) = (2 :: 'a::len word) ^ n + mask n" by (simp add: mask_eq_decr_exp) lemma mask_mono: "sz' \ sz \ mask sz' \ (mask sz :: 'a::len word)" by (simp add: le_mask_iff shiftr_mask_le) lemma aligned_mask_disjoint: "\ is_aligned (a :: 'a :: len word) n; b \ mask n \ \ a AND b = 0" by (metis and_zero_eq is_aligned_mask le_mask_imp_and_mask word_bw_lcs(1)) lemma word_and_or_mask_aligned: "\ is_aligned a n; b \ mask n \ \ a + b = a OR b" by (simp add: aligned_mask_disjoint word_plus_and_or_coroll) lemma word_and_or_mask_aligned2: \is_aligned b n \ a \ mask n \ a + b = a OR b\ using word_and_or_mask_aligned [of b n a] by (simp add: ac_simps) lemma is_aligned_ucastI: "is_aligned w n \ is_aligned (ucast w) n" by (simp add: bit_ucast_iff is_aligned_nth) lemma ucast_le_maskI: "a \ mask n \ UCAST('a::len \ 'b::len) a \ mask n" by (metis and_mask_eq_iff_le_mask ucast_and_mask) lemma ucast_add_mask_aligned: "\ a \ mask n; is_aligned b n \ \ UCAST ('a::len \ 'b::len) (a + b) = ucast a + ucast b" by (metis add.commute is_aligned_ucastI ucast_le_maskI ucast_or_distrib word_and_or_mask_aligned) lemma ucast_shiftl: "LENGTH('b) \ LENGTH ('a) \ UCAST ('a::len \ 'b::len) x << n = ucast (x << n)" by word_eqI_solve lemma ucast_leq_mask: "LENGTH('a) \ n \ ucast (x::'a::len word) \ mask n" apply (simp add: less_eq_mask_iff_take_bit_eq_self) apply transfer apply (simp add: ac_simps) done lemma shiftl_inj: "\ x << n = y << n; x \ mask (LENGTH('a)-n); y \ mask (LENGTH('a)-n) \ \ x = (y :: 'a :: len word)" apply word_eqI apply (rename_tac n') apply (case_tac "LENGTH('a) - n \ n'", simp) by (metis add.commute add.right_neutral diff_add_inverse le_diff_conv linorder_not_less zero_order(1)) lemma distinct_word_add_ucast_shift_inj: \p' = p \ off' = off\ if *: \p + (UCAST('a::len \ 'b::len) off << n) = p' + (ucast off' << n)\ and \is_aligned p n'\ \is_aligned p' n'\ \n' = n + LENGTH('a)\ \n' < LENGTH('b)\ proof - from \n' = n + LENGTH('a)\ have [simp]: \n' - n = LENGTH('a)\ \n + LENGTH('a) = n'\ by simp_all from \is_aligned p n'\ obtain q where p: \p = push_bit n' (word_of_nat q)\ \q < 2 ^ (LENGTH('b) - n')\ by (rule is_alignedE') from \is_aligned p' n'\ obtain q' where p': \p' = push_bit n' (word_of_nat q')\ \q' < 2 ^ (LENGTH('b) - n')\ by (rule is_alignedE') define m :: nat where \m = unat off\ then have off: \off = word_of_nat m\ by simp define m' :: nat where \m' = unat off'\ then have off': \off' = word_of_nat m'\ by simp have \push_bit n' q + take_bit n' (push_bit n m) < 2 ^ LENGTH('b)\ by (metis id_apply is_aligned_no_wrap''' of_nat_eq_id of_nat_push_bit p(1) p(2) take_bit_nat_eq_self_iff take_bit_nat_less_exp take_bit_push_bit that(2) that(5) unsigned_of_nat) moreover have \push_bit n' q' + take_bit n' (push_bit n m') < 2 ^ LENGTH('b)\ by (metis \n' - n = LENGTH('a)\ id_apply is_aligned_no_wrap''' m'_def of_nat_eq_id of_nat_push_bit off' p'(1) p'(2) take_bit_nat_eq_self_iff take_bit_push_bit that(3) that(5) unsigned_of_nat) ultimately have \push_bit n' q + take_bit n' (push_bit n m) = push_bit n' q' + take_bit n' (push_bit n m')\ - using * by (simp add: p p' off off' push_bit_of_nat push_bit_take_bit word_of_nat_inj flip: of_nat_add) + using * by (simp add: p p' off off' push_bit_of_nat push_bit_take_bit word_of_nat_inj unsigned_of_nat flip: of_nat_add) then have \int (push_bit n' q + take_bit n' (push_bit n m)) = int (push_bit n' q' + take_bit n' (push_bit n m'))\ by simp then have \concat_bit n' (int (push_bit n m)) (int q) = concat_bit n' (int (push_bit n m')) (int q')\ by (simp add: of_nat_push_bit of_nat_take_bit concat_bit_eq) then show ?thesis by (simp add: p p' off off' take_bit_of_nat take_bit_push_bit word_of_nat_eq_iff concat_bit_eq_iff) (simp add: push_bit_eq_mult) qed lemma word_upto_Nil: "y < x \ [x .e. y ::'a::len word] = []" by (simp add: upto_enum_red not_le word_less_nat_alt) lemma word_enum_decomp_elem: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y" proof - have "set as \ set [x .e. y] \ a \ set [x .e. y]" using assms by (auto dest: arg_cong[where f=set]) then show ?thesis by auto qed lemma word_enum_prefix: "[x .e. (y ::'a::len word)] = as @ a # bs \ as = (if x < a then [x .e. a - 1] else [])" apply (induct as arbitrary: x; clarsimp) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (simp add: word_upto_Cons_eq) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (clarsimp simp: word_upto_Cons_eq) apply (frule word_enum_decomp_elem) apply clarsimp apply (rule conjI) prefer 2 apply (subst word_Suc_le[symmetric]; clarsimp) apply (drule meta_spec) apply (drule (1) meta_mp) apply clarsimp apply (rule conjI; clarsimp) apply (subst (2) word_upto_Cons_eq) apply unat_arith apply simp done lemma word_enum_decomp_set: "[x .e. (y ::'a::len word)] = as @ a # bs \ a \ set as" by (metis distinct_append distinct_enum_upto' not_distinct_conv_prefix) lemma word_enum_decomp: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y \ a \ set as \ (\z \ set as. x \ z \ z \ y)" proof - from assms have "set as \ set [x .e. y] \ a \ set [x .e. y]" by (auto dest: arg_cong[where f=set]) with word_enum_decomp_set[OF assms] show ?thesis by auto qed lemma of_nat_unat_le_mask_ucast: "\of_nat (unat t) = w; t \ mask LENGTH('a)\ \ t = UCAST('a::len \ 'b::len) w" by (clarsimp simp: ucast_nat_def ucast_ucast_mask simp flip: and_mask_eq_iff_le_mask) lemma less_diff_gt0: "a < b \ (0 :: 'a :: len word) < b - a" by unat_arith lemma unat_plus_gt: "unat ((a :: 'a :: len word) + b) \ unat a + unat b" by (clarsimp simp: unat_plus_if_size) lemma const_less: "\ (a :: 'a :: len word) - 1 < b; a \ b \ \ a < b" by (metis less_1_simp word_le_less_eq) lemma add_mult_aligned_neg_mask: \(x + y * m) AND NOT(mask n) = (x AND NOT(mask n)) + y * m\ if \m AND (2 ^ n - 1) = 0\ for x y m :: \'a::len word\ by (metis (no_types, opaque_lifting) add.assoc add.commute add.right_neutral add_uminus_conv_diff mask_eq_decr_exp mask_eqs(2) mask_eqs(6) mult.commute mult_zero_left subtract_mask(1) that) lemma unat_of_nat_minus_1: "\ n < 2 ^ LENGTH('a); n \ 0 \ \ unat ((of_nat n:: 'a :: len word) - 1) = n - 1" by (simp add: of_nat_diff unat_eq_of_nat) lemma word_eq_zeroI: "a \ a - 1 \ a = 0" for a :: "'a :: len word" by (simp add: word_must_wrap) lemma word_add_format: "(-1 :: 'a :: len word) + b + c = b + (c - 1)" by simp lemma upto_enum_word_nth: "\ i \ j; k \ unat (j - i) \ \ [i .e. j] ! k = i + of_nat k" apply (clarsimp simp: upto_enum_def nth_append) apply (clarsimp simp: word_le_nat_alt[symmetric]) apply (rule conjI, clarsimp) apply (subst toEnum_of_nat, unat_arith) apply unat_arith apply (clarsimp simp: not_less unat_sub[symmetric]) apply unat_arith done lemma upto_enum_step_nth: "\ a \ c; n \ unat ((c - a) div (b - a)) \ \ [a, b .e. c] ! n = a + of_nat n * (b - a)" by (clarsimp simp: upto_enum_step_def not_less[symmetric] upto_enum_word_nth) lemma upto_enum_inc_1_len: "a < - 1 \ [(0 :: 'a :: len word) .e. 1 + a] = [0 .e. a] @ [1 + a]" apply (simp add: upto_enum_word) apply (subgoal_tac "unat (1+a) = 1 + unat a") apply simp apply (subst unat_plus_simple[THEN iffD1]) apply (metis add.commute no_plus_overflow_neg olen_add_eqv) apply unat_arith done lemma neg_mask_add: "y AND mask n = 0 \ x + y AND NOT(mask n) = (x AND NOT(mask n)) + y" for x y :: \'a::len word\ by (clarsimp simp: mask_out_sub_mask mask_eqs(7)[symmetric] mask_twice) lemma shiftr_shiftl_shiftr[simp]: "(x :: 'a :: len word) >> a << a >> a = x >> a" by word_eqI_solve lemma add_right_shift: "\ x AND mask n = 0; y AND mask n = 0; x \ x + y \ \ (x + y :: ('a :: len) word) >> n = (x >> n) + (y >> n)" apply (simp add: no_olen_add_nat is_aligned_mask[symmetric]) apply (simp add: unat_arith_simps shiftr_div_2n' split del: if_split) apply (subst if_P) apply (erule order_le_less_trans[rotated]) apply (simp add: add_mono) apply (simp add: shiftr_div_2n' is_aligned_iff_dvd_nat) done lemma sub_right_shift: "\ x AND mask n = 0; y AND mask n = 0; y \ x \ \ (x - y) >> n = (x >> n :: 'a :: len word) - (y >> n)" using add_right_shift[where x="x - y" and y=y and n=n] by (simp add: aligned_sub_aligned is_aligned_mask[symmetric] word_sub_le) lemma and_and_mask_simple: "y AND mask n = mask n \ (x AND y) AND mask n = x AND mask n" by (simp add: ac_simps) lemma and_and_mask_simple_not: "y AND mask n = 0 \ (x AND y) AND mask n = 0" by (simp add: ac_simps) lemma word_and_le': "b \ c \ (a :: 'a :: len word) AND b \ c" by (metis word_and_le1 order_trans) lemma word_and_less': "b < c \ (a :: 'a :: len word) AND b < c" by transfer simp lemma shiftr_w2p: "x < LENGTH('a) \ 2 ^ x = (2 ^ (LENGTH('a) - 1) >> (LENGTH('a) - 1 - x) :: 'a :: len word)" by word_eqI_solve lemma t2p_shiftr: "\ b \ a; a < LENGTH('a) \ \ (2 :: 'a :: len word) ^ a >> b = 2 ^ (a - b)" by word_eqI_solve lemma scast_1[simp]: "scast (1 :: 'a :: len signed word) = (1 :: 'a word)" by simp lemma unsigned_uminus1 [simp]: \(unsigned (-1::'b::len word)::'c::len word) = mask LENGTH('b)\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma ucast_ucast_mask_eq: "\ UCAST('a::len \ 'b::len) x = y; x AND mask LENGTH('b) = x \ \ x = ucast y" by (drule sym) (simp flip: take_bit_eq_mask add: unsigned_ucast_eq) lemma ucast_up_eq: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ LENGTH ('b) \ \ ucast x = (ucast y::'a::len word)" by word_eqI_solve lemma ucast_up_neq: "\ ucast x \ (ucast y::'b::len word); LENGTH('b) \ LENGTH ('a) \ \ ucast x \ (ucast y::'a::len word)" by (fastforce dest: ucast_up_eq) lemma mask_AND_less_0: "\ x AND mask n = 0; m \ n \ \ x AND mask m = 0" for x :: \'a::len word\ by (metis mask_twice2 word_and_notzeroD) lemma mask_len_id [simp]: "(x :: 'a :: len word) AND mask LENGTH('a) = x" using uint_lt2p [of x] by (simp add: mask_eq_iff) lemma scast_ucast_down_same: "LENGTH('b) \ LENGTH('a) \ SCAST('a \ 'b) = UCAST('a::len \ 'b::len)" by (simp add: down_cast_same is_down) lemma word_aligned_0_sum: "\ a + b = 0; is_aligned (a :: 'a :: len word) n; b \ mask n; n < LENGTH('a) \ \ a = 0 \ b = 0" by (simp add: word_plus_and_or_coroll aligned_mask_disjoint word_or_zero) lemma mask_eq1_nochoice: "\ LENGTH('a) > 1; (x :: 'a :: len word) AND 1 = x \ \ x = 0 \ x = 1" by (metis word_and_1) lemma shiftr_and_eq_shiftl: "(w >> n) AND x = y \ w AND (x << n) = (y << n)" for y :: "'a:: len word" apply (drule sym) apply simp apply (rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma add_mask_lower_bits': "\ len = LENGTH('a); is_aligned (x :: 'a :: len word) n; \n' \ n. n' < len \ \ bit p n' \ \ x + p AND NOT(mask n) = x" using add_mask_lower_bits by auto lemma leq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + high_bits) \ (x >> low_bits) \ mask high_bits" by (simp add: le_mask_iff shiftr_shiftr ac_simps) lemma ucast_ucast_eq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + LENGTH('b)) \ ucast((ucast (x >> low_bits)) :: 'b :: len word) = x >> low_bits" by (meson and_mask_eq_iff_le_mask eq_ucast_ucast_eq not_le_imp_less shiftr_less_t2n' ucast_ucast_len) lemma const_le_unat: "\ b < 2 ^ LENGTH('a); of_nat b \ a \ \ b \ unat (a :: 'a :: len word)" - apply (simp add: word_le_def) - apply (simp only: uint_nat zle_int) - apply transfer - apply (simp add: take_bit_nat_eq_self) - done + by (simp add: word_le_nat_alt unsigned_of_nat take_bit_nat_eq_self) lemma upt_enum_offset_trivial: "\ x < 2 ^ LENGTH('a) - 1 ; n \ unat x \ \ ([(0 :: 'a :: len word) .e. x] ! n) = of_nat n" apply (induct x arbitrary: n) apply simp by (simp add: upto_enum_word_nth) lemma word_le_mask_out_plus_2sz: "x \ (x AND NOT(mask sz)) + 2 ^ sz - 1" for x :: \'a::len word\ by (metis add_diff_eq word_neg_and_le) lemma ucast_add: "ucast (a + (b :: 'a :: len word)) = ucast a + (ucast b :: ('a signed word))" by transfer (simp add: take_bit_add) lemma ucast_minus: "ucast (a - (b :: 'a :: len word)) = ucast a - (ucast b :: ('a signed word))" apply (insert ucast_add[where a=a and b="-b"]) apply (metis (no_types, opaque_lifting) add_diff_eq diff_add_cancel ucast_add) done lemma scast_ucast_add_one [simp]: "scast (ucast (x :: 'a::len word) + (1 :: 'a signed word)) = x + 1" apply (subst ucast_1[symmetric]) apply (subst ucast_add[symmetric]) apply clarsimp done lemma word_and_le_plus_one: "a > 0 \ (x :: 'a :: len word) AND (a - 1) < a" by (simp add: gt0_iff_gem1 word_and_less') lemma unat_of_ucast_then_shift_eq_unat_of_shift[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) >> n) = unat (x >> n)" by (simp add: shiftr_div_2n' unat_ucast_up_simp) lemma unat_of_ucast_then_mask_eq_unat_of_mask[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) AND mask m) = unat (x AND mask m)" by (metis ucast_and_mask unat_ucast_up_simp) lemma shiftr_less_t2n3: "\ (2 :: 'a word) ^ (n + m) = 0; m < LENGTH('a) \ \ (x :: 'a :: len word) >> n < 2 ^ m" by (fastforce intro: shiftr_less_t2n' simp: mask_eq_decr_exp power_overflow) lemma unat_shiftr_le_bound: "\ 2 ^ (LENGTH('a :: len) - n) - 1 \ bnd; 0 < n \ \ unat ((x :: 'a word) >> n) \ bnd" apply transfer apply (simp add: take_bit_drop_bit) apply (simp add: drop_bit_take_bit) apply (rule order_trans) defer apply assumption apply (simp add: nat_le_iff of_nat_diff) done lemma shiftr_eqD: "\ x >> n = y >> n; is_aligned x n; is_aligned y n \ \ x = y" by (metis is_aligned_shiftr_shiftl) lemma word_shiftr_shiftl_shiftr_eq_shiftr: "a \ b \ (x :: 'a :: len word) >> a << b >> b = x >> a" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma of_int_uint_ucast: "of_int (uint (x :: 'a::len word)) = (ucast x :: 'b::len word)" by (fact Word.of_int_uint) lemma mod_mask_drop: "\ m = 2 ^ n; 0 < m; mask n AND msk = mask n \ \ (x mod m) AND msk = x mod m" for x :: \'a::len word\ by (simp add: word_mod_2p_is_mask word_bw_assocs) lemma mask_eq_ucast_eq: "\ x AND mask LENGTH('a) = (x :: ('c :: len word)); LENGTH('a) \ LENGTH('b)\ \ ucast (ucast x :: ('a :: len word)) = (ucast x :: ('b :: len word))" by (metis ucast_and_mask ucast_id ucast_ucast_mask ucast_up_eq) lemma of_nat_less_t2n: "of_nat i < (2 :: ('a :: len) word) ^ n \ n < LENGTH('a) \ unat (of_nat i :: 'a word) < 2 ^ n" by (metis order_less_trans p2_gt_0 unat_less_power word_neq_0_conv) lemma two_power_increasing_less_1: "\ n \ m; m \ LENGTH('a) \ \ (2 :: 'a :: len word) ^ n - 1 \ 2 ^ m - 1" by (metis diff_diff_cancel le_m1_iff_lt less_imp_diff_less p2_gt_0 two_power_increasing word_1_le_power word_le_minus_mono_left word_less_sub_1) lemma word_sub_mono4: "\ y + x \ z + x; y \ y + x; z \ z + x \ \ y \ z" for y :: "'a :: len word" by (simp add: word_add_le_iff2) lemma eq_or_less_helperD: "\ n = unat (2 ^ m - 1 :: 'a :: len word) \ n < unat (2 ^ m - 1 :: 'a word); m < LENGTH('a) \ \ n < 2 ^ m" by (meson le_less_trans nat_less_le unat_less_power word_power_less_1) lemma mask_sub: "n \ m \ mask m - mask n = mask m AND NOT(mask n :: 'a::len word)" by (metis (full_types) and_mask_eq_iff_shiftr_0 mask_out_sub_mask shiftr_mask_le word_bw_comms(1)) lemma neg_mask_diff_bound: "sz'\ sz \ (ptr AND NOT(mask sz')) - (ptr AND NOT(mask sz)) \ 2 ^ sz - 2 ^ sz'" (is "_ \ ?lhs \ ?rhs") for ptr :: \'a::len word\ proof - assume lt: "sz' \ sz" hence "?lhs = ptr AND (mask sz AND NOT(mask sz'))" by (metis add_diff_cancel_left' multiple_mask_trivia) also have "\ \ ?rhs" using lt by (metis (mono_tags) add_diff_eq diff_eq_eq eq_iff mask_2pm1 mask_sub word_and_le') finally show ?thesis by simp qed lemma mask_out_eq_0: "\ idx < 2 ^ sz; sz < LENGTH('a) \ \ (of_nat idx :: 'a :: len word) AND NOT(mask sz) = 0" by (simp add: of_nat_power less_mask_eq mask_eq_0_eq_x) lemma is_aligned_neg_mask_eq': "is_aligned ptr sz = (ptr AND NOT(mask sz) = ptr)" using is_aligned_mask mask_eq_0_eq_x by blast lemma neg_mask_mask_unat: "sz < LENGTH('a) \ unat ((ptr :: 'a :: len word) AND NOT(mask sz)) + unat (ptr AND mask sz) = unat ptr" by (metis AND_NOT_mask_plus_AND_mask_eq unat_plus_simple word_and_le2) lemma unat_pow_le_intro: "LENGTH('a) \ n \ unat (x :: 'a :: len word) < 2 ^ n" by (metis lt2p_lem not_le of_nat_le_iff of_nat_numeral semiring_1_class.of_nat_power uint_nat) lemma unat_shiftl_less_t2n: "\ unat (x :: 'a :: len word) < 2 ^ (m - n); m < LENGTH('a) \ \ unat (x << n) < 2 ^ m" by (metis More_Word.of_nat_power nat_mult_power_less_eq numeral_2_eq_2 of_nat_push_bit push_bit_eq_mult unat_less_power unat_of_nat_len unsigned_less word_eq_unatI zero_less_Suc) lemma unat_is_aligned_add: "\ is_aligned p n; unat d < 2 ^ n \ \ unat (p + d AND mask n) = unat d \ unat (p + d AND NOT(mask n)) = unat p" by (metis add.right_neutral and_mask_eq_iff_le_mask and_not_mask le_mask_iff mask_add_aligned mask_out_add_aligned mult_zero_right shiftl_t2n shiftr_le_0) lemma unat_shiftr_shiftl_mask_zero: "\ c + a \ LENGTH('a) + b ; c < LENGTH('a) \ \ unat (((q :: 'a :: len word) >> a << b) AND NOT(mask c)) = 0" by (fastforce intro: unat_is_aligned_add[where p=0 and n=c, simplified, THEN conjunct2] unat_shiftl_less_t2n unat_shiftr_less_t2n unat_pow_le_intro) lemmas of_nat_ucast = ucast_of_nat[symmetric] lemma shift_then_mask_eq_shift_low_bits: "x \ mask (low_bits + high_bits) \ (x >> low_bits) AND mask high_bits = x >> low_bits" for x :: \'a::len word\ by (simp add: leq_mask_shift le_mask_imp_and_mask) lemma leq_low_bits_iff_zero: "\ x \ mask (low bits + high bits); x >> low_bits = 0 \ \ (x AND mask low_bits = 0) = (x = 0)" for x :: \'a::len word\ using and_mask_eq_iff_shiftr_0 by force lemma unat_less_iff: "\ unat (a :: 'a :: len word) = b; c < 2 ^ LENGTH('a) \ \ (a < of_nat c) = (b < c)" using unat_ucast_less_no_overflow_simp by blast lemma is_aligned_no_overflow3: "\ is_aligned (a :: 'a :: len word) n; n < LENGTH('a); b < 2 ^ n; c \ 2 ^ n; b < c \ \ a + b \ a + (c - 1)" by (meson is_aligned_no_wrap' le_m1_iff_lt not_le word_less_sub_1 word_plus_mono_right) lemma mask_add_aligned_right: "is_aligned p n \ (q + p) AND mask n = q AND mask n" by (simp add: mask_add_aligned add.commute) lemma leq_high_bits_shiftr_low_bits_leq_bits_mask: "x \ mask high_bits \ (x :: 'a :: len word) << low_bits \ mask (low_bits + high_bits)" by (metis le_mask_shiftl_le_mask) lemma word_two_power_neg_ineq: "2 ^ m \ (0 :: 'a word) \ 2 ^ n \ - (2 ^ m :: 'a :: len word)" apply (cases "n < LENGTH('a)"; simp add: power_overflow) apply (cases "m < LENGTH('a)"; simp add: power_overflow) apply (simp add: word_le_nat_alt unat_minus word_size) apply (cases "LENGTH('a)"; simp) apply (simp add: less_Suc_eq_le) apply (drule power_increasing[where a=2 and n=n] power_increasing[where a=2 and n=m], simp)+ apply (drule(1) add_le_mono) apply simp done lemma unat_shiftl_absorb: "\ x \ 2 ^ p; p + k < LENGTH('a) \ \ unat (x :: 'a :: len word) * 2 ^ k = unat (x * 2 ^ k)" by (smt add_diff_cancel_right' add_lessD1 le_add2 le_less_trans mult.commute nat_le_power_trans unat_lt2p unat_mult_lem unat_power_lower word_le_nat_alt) lemma word_plus_mono_right_split: "\ unat ((x :: 'a :: len word) AND mask sz) + unat z < 2 ^ sz; sz < LENGTH('a) \ \ x \ x + z" apply (subgoal_tac "(x AND NOT(mask sz)) + (x AND mask sz) \ (x AND NOT(mask sz)) + ((x AND mask sz) + z)") apply (simp add:word_plus_and_or_coroll2 field_simps) apply (rule word_plus_mono_right) apply (simp add: less_le_trans no_olen_add_nat) using of_nat_power is_aligned_no_wrap' by force lemma mul_not_mask_eq_neg_shiftl: "NOT(mask n :: 'a::len word) = -1 << n" by (simp add: NOT_mask shiftl_t2n) lemma shiftr_mul_not_mask_eq_and_not_mask: "(x >> n) * NOT(mask n) = - (x AND NOT(mask n))" for x :: \'a::len word\ by (metis NOT_mask and_not_mask mult_minus_left semiring_normalization_rules(7) shiftl_t2n) lemma mask_eq_n1_shiftr: "n \ LENGTH('a) \ (mask n :: 'a :: len word) = -1 >> (LENGTH('a) - n)" by (metis diff_diff_cancel eq_refl mask_full shiftr_mask2) lemma is_aligned_mask_out_add_eq: "is_aligned p n \ (p + x) AND NOT(mask n) = p + (x AND NOT(mask n))" by (simp add: mask_out_sub_mask mask_add_aligned) lemmas is_aligned_mask_out_add_eq_sub = is_aligned_mask_out_add_eq[where x="a - b" for a b, simplified field_simps] lemma aligned_bump_down: "is_aligned x n \ (x - 1) AND NOT(mask n) = x - 2 ^ n" by (drule is_aligned_mask_out_add_eq[where x="-1"]) (simp add: NOT_mask) lemma unat_2tp_if: "unat (2 ^ n :: ('a :: len) word) = (if n < LENGTH ('a) then 2 ^ n else 0)" by (split if_split, simp_all add: power_overflow) lemma mask_of_mask: "mask (n::nat) AND mask (m::nat) = (mask (min m n) :: 'a::len word)" by word_eqI_solve lemma unat_signed_ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ unat (ucast (x :: 'a :: len word) :: 'b :: len signed word) = unat x" by (simp add: unat_ucast_up_simp) lemma toEnum_of_ucast: "LENGTH('b) \ LENGTH('a) \ (toEnum (unat (b::'b :: len word))::'a :: len word) = of_nat (unat b)" by (simp add: unat_pow_le_intro) lemmas unat_ucast_mask = unat_ucast_eq_unat_and_mask[where w=a for a] lemma t2n_mask_eq_if: "2 ^ n AND mask m = (if n < m then 2 ^ n else (0 :: 'a::len word))" by (rule word_eqI) (auto simp add: bit_simps) lemma unat_ucast_le: "unat (ucast (x :: 'a :: len word) :: 'b :: len word) \ unat x" by (simp add: ucast_nat_def word_unat_less_le) lemma ucast_le_up_down_iff: "\ LENGTH('a) \ LENGTH('b); (x :: 'b :: len word) \ ucast (- 1 :: 'a :: len word) \ \ (ucast x \ (y :: 'a word)) = (x \ ucast y)" using le_max_word_ucast_id ucast_le_ucast by metis lemma ucast_ucast_mask_shift: "a \ LENGTH('a) + b \ ucast (ucast (p AND mask a >> b) :: 'a :: len word) = p AND mask a >> b" by (metis add.commute le_mask_iff shiftr_mask_le ucast_ucast_eq_mask_shift word_and_le') lemma unat_ucast_mask_shift: "a \ LENGTH('a) + b \ unat (ucast (p AND mask a >> b) :: 'a :: len word) = unat (p AND mask a >> b)" by (metis linear ucast_ucast_mask_shift unat_ucast_up_simp) lemma mask_overlap_zero: "a \ b \ (p AND mask a) AND NOT(mask b) = 0" for p :: \'a::len word\ by (metis NOT_mask_AND_mask mask_lower_twice2 max_def) lemma mask_shifl_overlap_zero: "a + c \ b \ (p AND mask a << c) AND NOT(mask b) = 0" for p :: \'a::len word\ by (metis and_mask_0_iff_le_mask mask_mono mask_shiftl_decompose order_trans shiftl_over_and_dist word_and_le' word_and_le2) lemma mask_overlap_zero': "a \ b \ (p AND NOT(mask a)) AND mask b = 0" for p :: \'a::len word\ using mask_AND_NOT_mask mask_AND_less_0 by blast lemma mask_rshift_mult_eq_rshift_lshift: "((a :: 'a :: len word) >> b) * (1 << c) = (a >> b << c)" by (simp add: shiftl_t2n) lemma shift_alignment: "a \ b \ is_aligned (p >> a << a) b" using is_aligned_shift is_aligned_weaken by blast lemma mask_split_sum_twice: "a \ b \ (p AND NOT(mask a)) + ((p AND mask a) AND NOT(mask b)) + (p AND mask b) = p" for p :: \'a::len word\ by (simp add: add.commute multiple_mask_trivia word_bw_comms(1) word_bw_lcs(1) word_plus_and_or_coroll2) lemma mask_shift_eq_mask_mask: "(p AND mask a >> b << b) = (p AND mask a) AND NOT(mask b)" for p :: \'a::len word\ by (simp add: and_not_mask) lemma mask_shift_sum: "\ a \ b; unat n = unat (p AND mask b) \ \ (p AND NOT(mask a)) + (p AND mask a >> b) * (1 << b) + n = (p :: 'a :: len word)" apply (simp add: push_bit_of_1 flip: push_bit_eq_mult) apply (subst disjunctive_add) apply (auto simp add: bit_simps) apply (smt (z3) AND_NOT_mask_plus_AND_mask_eq and.comm_neutral and.right_idem and_not_mask bit.conj_disj_distrib bit.disj_cancel_right mask_out_first_mask_some word_bw_assocs(1) word_bw_comms(1) word_bw_comms(2) word_eq_unatI) done lemma is_up_compose: "\ is_up uc; is_up uc' \ \ is_up (uc' \ uc)" unfolding is_up_def by (simp add: Word.target_size Word.source_size) lemma of_int_sint_scast: "of_int (sint (x :: 'a :: len word)) = (scast x :: 'b :: len word)" by (fact Word.of_int_sint) lemma scast_of_nat_to_signed [simp]: "scast (of_nat x :: 'a :: len word) = (of_nat x :: 'a signed word)" by (rule bit_word_eqI) (simp add: bit_simps) lemma scast_of_nat_signed_to_unsigned_add: "scast (of_nat x + of_nat y :: 'a :: len signed word) = (of_nat x + of_nat y :: 'a :: len word)" by (metis of_nat_add scast_of_nat) lemma scast_of_nat_unsigned_to_signed_add: "(scast (of_nat x + of_nat y :: 'a :: len word)) = (of_nat x + of_nat y :: 'a :: len signed word)" by (metis Abs_fnat_hom_add scast_of_nat_to_signed) lemma and_mask_cases: fixes x :: "'a :: len word" assumes len: "n < LENGTH('a)" shows "x AND mask n \ of_nat ` set [0 ..< 2 ^ n]" apply (simp flip: take_bit_eq_mask) apply (rule image_eqI [of _ _ \unat (take_bit n x)\]) using len apply simp_all apply transfer apply simp done lemma sint_eq_uint_2pl: "\ (a :: 'a :: len word) < 2 ^ (LENGTH('a) - 1) \ \ sint a = uint a" by (simp add: not_msb_from_less sint_eq_uint word_2p_lem word_size) lemma pow_sub_less: "\ a + b \ LENGTH('a); unat (x :: 'a :: len word) = 2 ^ a \ \ unat (x * 2 ^ b - 1) < 2 ^ (a + b)" by (smt (z3) eq_or_less_helperD le_add2 le_eq_less_or_eq le_trans power_add unat_mult_lem unat_pow_le_intro unat_power_lower word_eq_unatI) lemma sle_le_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a \ b \ \ a <=s b" by (simp add: not_msb_from_less word_sle_msb_le) lemma sless_less_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a < b \ \ a > n = w AND mask (size w - n)" for w :: \'a::len word\ by (cases "n \ size w"; clarsimp simp: word_and_le2 and_mask shiftl_zero_size) lemma aligned_sub_aligned_simple: "\ is_aligned a n; is_aligned b n \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma minus_one_shift: "- (1 << n) = (-1 << n :: 'a::len word)" by (simp flip: mul_not_mask_eq_neg_shiftl minus_exp_eq_not_mask add: push_bit_of_1) lemma ucast_eq_mask: "(UCAST('a::len \ 'b::len) x = UCAST('a \ 'b) y) = (x AND mask LENGTH('b) = y AND mask LENGTH('b))" by transfer (simp flip: take_bit_eq_mask add: ac_simps) context fixes w :: "'a::len word" begin private lemma sbintrunc_uint_ucast: assumes "Suc n = LENGTH('b::len)" shows "signed_take_bit n (uint (ucast w :: 'b word)) = signed_take_bit n (uint w)" by (rule bit_eqI) (use assms in \simp add: bit_simps\) private lemma test_bit_sbintrunc: assumes "i < LENGTH('a)" shows "bit (word_of_int (signed_take_bit n (uint w)) :: 'a word) i = (if n < i then bit w n else bit w i)" using assms by (simp add: bit_simps) private lemma test_bit_sbintrunc_ucast: assumes len_a: "i < LENGTH('a)" shows "bit (word_of_int (signed_take_bit (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) i = (if LENGTH('b::len) \ i then bit w (LENGTH('b) - 1) else bit w i)" using len_a by (auto simp add: sbintrunc_uint_ucast bit_simps) lemma scast_ucast_high_bits: \scast (ucast w :: 'b::len word) = w \ (\ i \ {LENGTH('b) ..< size w}. bit w i = bit w (LENGTH('b) - 1))\ proof (cases \LENGTH('a) \ LENGTH('b)\) case True moreover define m where \m = LENGTH('b) - LENGTH('a)\ ultimately have \LENGTH('b) = m + LENGTH('a)\ by simp then show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (rule bit_word_eqI) apply (simp add: bit_signed_take_bit_iff) done next case False define q where \q = LENGTH('b) - 1\ then have \LENGTH('b) = Suc q\ by simp moreover define m where \m = Suc LENGTH('a) - LENGTH('b)\ with False \LENGTH('b) = Suc q\ have \LENGTH('a) = m + q\ by (simp add: not_le) ultimately show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (transfer fixing: m q) apply (simp add: signed_take_bit_take_bit) apply rule apply (subst bit_eq_iff) apply (simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def) apply (auto simp add: Suc_le_eq) using less_imp_le_nat apply blast using less_imp_le_nat apply blast done qed lemma scast_ucast_mask_compare: "scast (ucast w :: 'b::len word) = w \ (w \ mask (LENGTH('b) - 1) \ NOT(mask (LENGTH('b) - 1)) \ w)" apply (clarsimp simp: le_mask_high_bits neg_mask_le_high_bits scast_ucast_high_bits word_size) apply (rule iffI; clarsimp) apply (rename_tac i j; case_tac "i = LENGTH('b) - 1"; case_tac "j = LENGTH('b) - 1") by auto lemma ucast_less_shiftl_helper': "\ LENGTH('b) + (a::nat) < LENGTH('a); 2 ^ (LENGTH('b) + a) \ n\ \ (ucast (x :: 'b::len word) << a) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done end lemma ucast_ucast_mask2: "is_down (UCAST ('a \ 'b)) \ UCAST ('b::len \ 'c::len) (UCAST ('a::len \ 'b::len) x) = UCAST ('a \ 'c) (x AND mask LENGTH('b))" apply (simp flip: take_bit_eq_mask) apply transfer apply simp done lemma ucast_NOT: "ucast (NOT x) = NOT(ucast x) AND mask (LENGTH('a))" for x::"'a::len word" by word_eqI lemma ucast_NOT_down: "is_down UCAST('a::len \ 'b::len) \ UCAST('a \ 'b) (NOT x) = NOT(UCAST('a \ 'b) x)" by word_eqI lemma upto_enum_step_shift: "\ is_aligned p n \ \ ([p , p + 2 ^ m .e. p + 2 ^ n - 1]) = map ((+) p) [0, 2 ^ m .e. 2 ^ n - 1]" apply (erule is_aligned_get_word_bits) prefer 2 apply (simp add: map_idI) apply (clarsimp simp: upto_enum_step_def) apply (frule is_aligned_no_overflow) apply (simp add: linorder_not_le [symmetric]) done lemma upto_enum_step_shift_red: "\ is_aligned p sz; sz < LENGTH('a); us \ sz \ \ [p :: 'a :: len word, p + 2 ^ us .e. p + 2 ^ sz - 1] = map (\x. p + of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" apply (subst upto_enum_step_shift, assumption) apply (simp add: upto_enum_step_red) done lemma upto_enum_step_subset: "set [x, y .e. z] \ {x .. z}" apply (clarsimp simp: upto_enum_step_def linorder_not_less) apply (drule div_to_mult_word_lt) apply (rule conjI) apply (erule word_random[rotated]) apply simp apply (rule order_trans) apply (erule word_plus_mono_right) apply simp apply simp done lemma ucast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (ucast :: 'a word \ 'b word)" shows "ucast (M a b) = M' (ucast a) (ucast b)" apply (simp only: ucast_eq) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') apply (metis local.distrib local.is_down take_bit_eq_mod ucast_down_wi uint_word_of_int_eq word_of_int_uint) done lemma ucast_down_add: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) + b) = (ucast a + ucast b :: 'b::len word)" by (rule ucast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma ucast_down_minus: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) - b) = (ucast a - ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma ucast_down_mult: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) * b) = (ucast a * ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (scast :: 'a word \ 'b word)" shows "scast (M a b) = M' (scast a) (scast b)" apply (subst (1 2 3) down_cast_same [symmetric]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) apply (rule ucast_distrib [where L=L, OF lift_M lift_M' distrib]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) done lemma scast_down_add: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) + b) = (scast a + scast b :: 'b::len word)" by (rule scast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma scast_down_minus: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) - b) = (scast a - scast b :: 'b::len word)" apply (rule scast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma scast_down_mult: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) * b) = (scast a * scast b :: 'b::len word)" apply (rule scast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_ucast_1: "\ is_down (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_3: "\ is_down (ucast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_4: "\ is_up (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_scast_b: "\ is_up (scast :: 'a word \ 'b word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq sint_up_scast) lemma ucast_scast_1: "\ is_down (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_3: "\ is_down (scast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_4: "\ is_up (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis down_cast_same scast_eq sint_up_scast) lemma ucast_ucast_a: "\ is_down (ucast :: 'b word \ 'c word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma ucast_ucast_b: "\ is_up (ucast :: 'a word \ 'b word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis ucast_up_ucast) lemma scast_scast_a: "\ is_down (scast :: 'b word \ 'c word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" apply (simp only: scast_eq) apply (metis down_cast_same is_up_down scast_eq ucast_down_wi) done lemma scast_down_wi [OF refl]: "uc = scast \ is_down uc \ uc (word_of_int x) = word_of_int x" by (metis down_cast_same is_up_down ucast_down_wi) lemmas cast_simps = is_down is_up scast_down_add scast_down_minus scast_down_mult ucast_down_add ucast_down_minus ucast_down_mult scast_ucast_1 scast_ucast_3 scast_ucast_4 ucast_scast_1 ucast_scast_3 ucast_scast_4 ucast_ucast_a ucast_ucast_b scast_scast_a scast_scast_b ucast_down_wi scast_down_wi ucast_of_nat scast_of_nat uint_up_ucast sint_up_scast up_scast_surj up_ucast_surj lemma sdiv_word_max: "(sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word) < (2 ^ (size a - 1))) = ((a \ - (2 ^ (size a - 1)) \ (b \ -1)))" (is "?lhs = (\ ?a_int_min \ \ ?b_minus1)") proof (rule classical) assume not_thesis: "\ ?thesis" have not_zero: "b \ 0" using not_thesis by (clarsimp) let ?range = \{- (2 ^ (size a - 1))..<2 ^ (size a - 1)} :: int set\ have result_range: "sint a sdiv sint b \ ?range \ {2 ^ (size a - 1)}" using sdiv_word_min [of a b] sdiv_word_max [of a b] by auto have result_range_overflow: "(sint a sdiv sint b = 2 ^ (size a - 1)) = (?a_int_min \ ?b_minus1)" apply (rule iffI [rotated]) apply (clarsimp simp: signed_divide_int_def sgn_if word_size sint_int_min) apply (rule classical) apply (case_tac "?a_int_min") apply (clarsimp simp: word_size sint_int_min) apply (metis diff_0_right int_sdiv_negated_is_minus1 minus_diff_eq minus_int_code(2) power_eq_0_iff sint_minus1 zero_neq_numeral) apply (subgoal_tac "abs (sint a) < 2 ^ (size a - 1)") apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (clarsimp simp: word_size) apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply auto apply (cases \size a\) apply simp_all apply (smt (z3) One_nat_def diff_Suc_1 signed_word_eqI sint_int_min sint_range_size wsst_TYs(3)) done have result_range_simple: "(sint a sdiv sint b \ ?range) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: word_size sint_int_min) done show ?thesis apply (rule UnE [OF result_range result_range_simple]) apply simp apply (clarsimp simp: word_size) using result_range_overflow apply (clarsimp simp: word_size) done qed lemmas sdiv_word_min' = sdiv_word_min [simplified word_size, simplified] lemmas sdiv_word_max' = sdiv_word_max [simplified word_size, simplified] lemma signed_arith_ineq_checks_to_eq: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) = (sint a + sint b = sint (a + b ))" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) = (sint a - sint b = sint (a - b))" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) = ((- sint a) = sint (- a))" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) = (sint a * sint b = sint (a * b))" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) = (sint a sdiv sint b = sint (a sdiv b))" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) = (sint a smod sint b = sint (a smod b))" by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith signed_take_bit_int_eq_self_iff intro: sym dest: sym) lemma signed_arith_sint: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) \ sint (a + b) = (sint a + sint b)" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) \ sint (a - b) = (sint a - sint b)" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) \ sint (- a) = (- sint a)" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) \ sint (a * b) = (sint a * sint b)" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) \ sint (a sdiv b) = (sint a sdiv sint b)" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) \ sint (a smod b) = (sint a smod sint b)" by (subst (asm) signed_arith_ineq_checks_to_eq; simp)+ end end