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,1640 +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) +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')" -proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) +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) +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')" -proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) +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')" -proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) +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/Buchi_Complementation/Complementation_Final.thy b/thys/Buchi_Complementation/Complementation_Final.thy --- a/thys/Buchi_Complementation/Complementation_Final.thy +++ b/thys/Buchi_Complementation/Complementation_Final.thy @@ -1,182 +1,182 @@ section \Final Instantiation of Algorithms Related to Complementation\ theory Complementation_Final imports "Complementation_Implement" "Formula" "Transition_Systems_and_Automata.NBA_Translate" "Transition_Systems_and_Automata.NGBA_Algorithms" "HOL-Library.Multiset" begin subsection \Syntax\ (* TODO: this syntax has unnecessarily high inner binding strength, requiring extra parentheses the regular let syntax correctly uses inner binding strength 0: ("(2_ =/ _)" 10) *) no_syntax "_do_let" :: "[pttrn, 'a] \ do_bind" ("(2let _ =/ _)" [1000, 13] 13) syntax "_do_let" :: "[pttrn, 'a] \ do_bind" ("(2let _ =/ _)" 13) subsection \Hashcodes on Complement States\ definition "hci k \ uint32_of_nat k * 1103515245 + 12345" definition "hc \ \ (p, q, b). hci p + hci q * 31 + (if b then 1 else 0)" - definition "list_hash xs \ fold ((XOR) \ hc) xs 0" + definition "list_hash xs \ fold (xor \ hc) xs 0" lemma list_hash_eq: assumes "distinct xs" "distinct ys" "set xs = set ys" shows "list_hash xs = list_hash ys" proof - have "mset (remdups xs) = mset (remdups ys)" using assms(3) using set_eq_iff_mset_remdups_eq by blast then have "mset xs = mset ys" using assms(1, 2) by (simp add: distinct_remdups_id) - have "fold ((XOR) \ hc) xs = fold ((XOR) \ hc) ys" + have "fold (xor \ hc) xs = fold (xor \ hc) ys" apply (rule fold_multiset_equiv) apply (simp_all add: fun_eq_iff ac_simps) using \mset xs = mset ys\ . then show ?thesis unfolding list_hash_def by simp qed definition state_hash :: "nat \ Complementation_Implement.state \ nat" where "state_hash n p \ nat_of_hashcode (list_hash p) mod n" lemma state_hash_bounded_hashcode[autoref_ga_rules]: "is_bounded_hashcode state_rel (gen_equals (Gen_Map.gen_ball (foldli \ list_map_to_list)) (list_map_lookup (=)) (prod_eq (=) (\))) state_hash" proof show [param]: "(gen_equals (Gen_Map.gen_ball (foldli \ list_map_to_list)) (list_map_lookup (=)) (prod_eq (=) (\)), (=)) \ state_rel \ state_rel \ bool_rel" by autoref show "state_hash n xs = state_hash n ys" if "xs \ Domain state_rel" "ys \ Domain state_rel" "gen_equals (Gen_Map.gen_ball (foldli \ list_map_to_list)) (list_map_lookup (=)) (prod_eq (=) (=)) xs ys" for xs ys n proof - have 1: "distinct (map fst xs)" "distinct (map fst ys)" using that(1, 2) unfolding list_map_rel_def list_map_invar_def by (auto simp: in_br_conv) have 2: "distinct xs" "distinct ys" using 1 by (auto intro: distinct_mapI) have 3: "(xs, map_of xs) \ state_rel" "(ys, map_of ys) \ state_rel" using 1 unfolding list_map_rel_def list_map_invar_def by (auto simp: in_br_conv) have 4: "(gen_equals (Gen_Map.gen_ball (foldli \ list_map_to_list)) (list_map_lookup (=)) (prod_eq (=) (\)) xs ys, map_of xs = map_of ys) \ bool_rel" using 3 by parametricity have 5: "map_to_set (map_of xs) = map_to_set (map_of ys)" using that(3) 4 by simp have 6: "set xs = set ys" using map_to_set_map_of 1 5 by blast show "state_hash n xs = state_hash n ys" unfolding state_hash_def using list_hash_eq 2 6 by metis qed show "state_hash n x < n" if "1 < n" for n x using that unfolding state_hash_def by simp qed subsection \Complementation\ schematic_goal complement_impl: assumes [simp]: "finite (NBA.nodes A)" assumes [autoref_rules]: "(Ai, A) \ \Id, nat_rel\ nbai_nba_rel" shows "(?f :: ?'c, op_translate (complement_4 A)) \ ?R" by (autoref_monadic (plain)) concrete_definition complement_impl uses complement_impl theorem complement_impl_correct: assumes "finite (NBA.nodes A)" assumes "(Ai, A) \ \Id, nat_rel\ nbai_nba_rel" shows "NBA.language (nbae_nba (nbaei_nbae (complement_impl Ai))) = streams (nba.alphabet A) - NBA.language A" using op_translate_language[OF complement_impl.refine[OF assms]] using complement_4_correct[OF assms(1)] by simp subsection \Language Subset\ definition [simp]: "op_language_subset A B \ NBA.language A \ NBA.language B" lemmas [autoref_op_pat] = op_language_subset_def[symmetric] schematic_goal language_subset_impl: assumes [simp]: "finite (NBA.nodes B)" assumes [autoref_rules]: "(Ai, A) \ \Id, nat_rel\ nbai_nba_rel" assumes [autoref_rules]: "(Bi, B) \ \Id, nat_rel\ nbai_nba_rel" shows "(?f :: ?'c, do { let AB' = intersect' A (complement_4 B); ASSERT (finite (NGBA.nodes AB')); RETURN (NGBA.language AB' = {}) }) \ ?R" by (autoref_monadic (plain)) concrete_definition language_subset_impl uses language_subset_impl lemma language_subset_impl_refine[autoref_rules]: assumes "SIDE_PRECOND (finite (NBA.nodes A))" assumes "SIDE_PRECOND (finite (NBA.nodes B))" assumes "SIDE_PRECOND (nba.alphabet A \ nba.alphabet B)" assumes "(Ai, A) \ \Id, nat_rel\ nbai_nba_rel" assumes "(Bi, B) \ \Id, nat_rel\ nbai_nba_rel" shows "(language_subset_impl Ai Bi, (OP op_language_subset ::: \Id, nat_rel\ nbai_nba_rel \ \Id, nat_rel\ nbai_nba_rel \ bool_rel) $ A $ B) \ bool_rel" proof - have "(RETURN (language_subset_impl Ai Bi), do { let AB' = intersect' A (complement_4 B); ASSERT (finite (NGBA.nodes AB')); RETURN (NGBA.language AB' = {}) }) \ \bool_rel\ nres_rel" using language_subset_impl.refine assms(2, 4, 5) unfolding autoref_tag_defs by this also have "(do { let AB' = intersect' A (complement_4 B); ASSERT (finite (NGBA.nodes AB')); RETURN (NGBA.language AB' = {}) }, RETURN (NBA.language A \ NBA.language B)) \ \bool_rel\ nres_rel" proof refine_vcg show "finite (NGBA.nodes (intersect' A (complement_4 B)))" using assms(1, 2) by auto have 1: "NBA.language A \ streams (nba.alphabet B)" using nba.language_alphabet streams_mono2 assms(3) unfolding autoref_tag_defs by blast have 2: "NBA.language (complement_4 B) = streams (nba.alphabet B) - NBA.language B" using complement_4_correct assms(2) by auto show "(NGBA.language (intersect' A (complement_4 B)) = {}, NBA.language A \ NBA.language B) \ bool_rel" using 1 2 by auto qed finally show ?thesis using RETURN_nres_relD unfolding nres_rel_comp by force qed subsection \Language Equality\ definition [simp]: "op_language_equal A B \ NBA.language A = NBA.language B" lemmas [autoref_op_pat] = op_language_equal_def[symmetric] schematic_goal language_equal_impl: assumes [simp]: "finite (NBA.nodes A)" assumes [simp]: "finite (NBA.nodes B)" assumes [simp]: "nba.alphabet A = nba.alphabet B" assumes [autoref_rules]: "(Ai, A) \ \Id, nat_rel\ nbai_nba_rel" assumes [autoref_rules]: "(Bi, B) \ \Id, nat_rel\ nbai_nba_rel" shows "(?f :: ?'c, NBA.language A \ NBA.language B \ NBA.language B \ NBA.language A) \ ?R" by autoref concrete_definition language_equal_impl uses language_equal_impl lemma language_equal_impl_refine[autoref_rules]: assumes "SIDE_PRECOND (finite (NBA.nodes A))" assumes "SIDE_PRECOND (finite (NBA.nodes B))" assumes "SIDE_PRECOND (nba.alphabet A = nba.alphabet B)" assumes "(Ai, A) \ \Id, nat_rel\ nbai_nba_rel" assumes "(Bi, B) \ \Id, nat_rel\ nbai_nba_rel" shows "(language_equal_impl Ai Bi, (OP op_language_equal ::: \Id, nat_rel\ nbai_nba_rel \ \Id, nat_rel\ nbai_nba_rel \ bool_rel) $ A $ B) \ bool_rel" using language_equal_impl.refine[OF assms[unfolded autoref_tag_defs]] by auto schematic_goal product_impl: assumes [simp]: "finite (NBA.nodes B)" assumes [autoref_rules]: "(Ai, A) \ \Id, nat_rel\ nbai_nba_rel" assumes [autoref_rules]: "(Bi, B) \ \Id, nat_rel\ nbai_nba_rel" shows "(?f :: ?'c, do { let AB' = intersect A (complement_4 B); ASSERT (finite (NBA.nodes AB')); op_translate AB' }) \ ?R" by (autoref_monadic (plain)) concrete_definition product_impl uses product_impl (* TODO: possible optimizations: - introduce op_map_map operation for maps instead of manually iterating via FOREACH - consolidate various binds and maps in expand_map_get_7 *) export_code Set.empty Set.insert Set.member "Inf :: 'a set set \ 'a set" "Sup :: 'a set set \ 'a set" image Pow set nat_of_integer integer_of_nat Variable Negation Conjunction Disjunction satisfies map_formula nbaei alphabetei initialei transitionei acceptingei nbae_nba_impl complement_impl language_equal_impl product_impl in SML module_name Complementation file_prefix Complementation end \ No newline at end of file diff --git a/thys/CAVA_LTL_Modelchecker/SM/Impl/SM_Ample_Impl.thy b/thys/CAVA_LTL_Modelchecker/SM/Impl/SM_Ample_Impl.thy --- a/thys/CAVA_LTL_Modelchecker/SM/Impl/SM_Ample_Impl.thy +++ b/thys/CAVA_LTL_Modelchecker/SM/Impl/SM_Ample_Impl.thy @@ -1,1412 +1,1413 @@ theory SM_Ample_Impl imports SM_Datastructures "../Analysis/SM_POR" "HOL-Library.RBT_Mapping" begin text \ Given a sticky program, we implement a valid ample-function by picking the transitions from the first PID that qualifies as an ample-set\ context visible_prog begin definition ga_gample :: "(cmdc\cmdc) set \ pid_global_config \ global_action set" where "ga_gample sticky_E gc \ let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc; as = find_min_idx_f (\lc. let c = local_config.command lc; ls = local_config.state lc; as = set (comp.succ_impl c) in if (\(a,_)\as. read_globals a = {} \ write_globals a = {}) then let as = {(a,c')\as. la_en' (ls,gs) a} in if as={} \ (\(a,c')\as. (c,c')\sticky_E) then None else Some (c,as) else None ) lcs in case as of Some (pid,c,as) \ (\(a,c'). (pid,c,a,c'))`as | None \ ga_gen gc" definition "ga_ample sticky_E \ stutter_extend_en (ga_gample sticky_E)" lemma ga_gample_subset: "ga_gample sticky_E gc \ ga_gen gc" unfolding ga_gample_def ga_gen_def by (auto simp: comp.astep_impl_def split: option.splits if_split_asm dest!: find_min_idx_f_SomeD) lemma ga_gample_empty_iff: "ga_gample sticky_E gc = {} \ ga_gen gc = {}" unfolding ga_gample_def ga_gen_def by (fastforce simp: comp.astep_impl_def split: option.splits if_split_asm dest!: find_min_idx_f_SomeD) lemma ga_ample_None_iff: "None \ ga_ample sticky_E gc \ None \ ga_en gc" by (auto simp: ga_ample_def ga_en_def stutter_extend_en_def ga_gample_empty_iff) lemma ga_ample_neq_en_eq: assumes "ga_ample sticky_E gc \ ga_en gc" shows "ga_ample sticky_E gc = Some`ga_gample sticky_E gc \ ga_en gc = Some`ga_gen gc" using assms ga_gample_subset unfolding ga_ample_def ga_en_def stutter_extend_en_def by (auto split: if_split_asm simp: ga_gample_empty_iff) lemma pid_pac_alt: "pid_pac pid = insert None (Some`{(pid',cac). pid'=pid})" apply (auto simp: pid_pac_def split: option.splits) apply (case_tac x) apply auto done end context sticky_prog begin sublocale ample_prog prog is_vis_var sticky_E "ga_ample sticky_E" proof - have aux: "Some ` A = Some ` B \ pid_pac pid \ A = {(pid', cac). (pid', cac) \ B \ pid' = pid}" for A B pid proof assume 1: "Some ` A = Some ` B \ pid_pac pid" show "A = {(pid', cac). (pid', cac) \ B \ pid' = pid}" using 1 apply (auto simp: pid_pac_alt) unfolding image_def apply auto using 1 by blast next show "Some ` A = Some ` B \ pid_pac pid" if "A = {(pid', cac). (pid', cac) \ B \ pid' = pid}" using that by (auto simp: pid_pac_alt) qed show "ample_prog prog is_vis_var sticky_E (ga_ample sticky_E)" apply unfold_locales apply (frule ga_ample_neq_en_eq, clarsimp) apply (intro conjI) apply (auto simp: ga_gample_def sticky_ga_def dest!: find_min_idx_f_SomeD split: option.splits if_split_asm) [] apply (fastforce simp: ga_gample_def sticky_ga_def dest!: find_min_idx_f_SomeD split: option.splits if_split_asm) [] unfolding aux (* TODO: Clean up! *) apply (simp add: inj_image_eq_iff pid_pac_def) apply (thin_tac "ga_ample c a = b" for a b c) apply (simp add: ga_gample_def split: option.splits prod.splits) apply (thin_tac "a \ ga_gen b" for a b) apply (drule find_min_idx_f_SomeD) apply clarsimp apply (rename_tac gc pid c as) apply (rule_tac x=pid in exI) apply (auto split: if_split_asm prod.splits simp: ga_gen_def comp.astep_impl_def) unfolding is_ample_static_def comp.astep_impl_def apply force done qed end text \ The following locale expresses a correct ample-reduction on the level of subset and stuttering-equivalences of traces \ locale hl_reduction = visible_prog + ample: sa "\ g_V = UNIV, g_E = step, g_V0 = {pid_init_gc}, sa_L = pid_interp_gc \" for step + assumes ample_accept_subset: "ample.accept w \ lv.sa.accept w" assumes ample_accept_stuttering: "lv.sa.accept w \ \w'. w\w' \ ample.accept w'" locale hl_ample_reduction = visible_prog + ample: sa "\ g_V = UNIV, g_E = rel_of_enex (ample, ga_ex), g_V0 = {pid_init_gc}, sa_L = pid_interp_gc \" for ample + assumes ample_accept_subset: "ample.accept w \ lv.sa.accept w" assumes ample_accept_stuttering: "lv.sa.accept w \ \w'. w\w' \ ample.accept w'" begin sublocale hl_reduction prog is_vis_var "rel_of_enex (ample, ga_ex)" apply unfold_locales using ample_accept_subset ample_accept_stuttering by auto end context sticky_prog begin sublocale hl_ample: hl_ample_reduction prog is_vis_var "ga_ample sticky_E" apply unfold_locales unfolding ample_rel_def[symmetric] unfolding reduced_automaton_def[symmetric] apply simp apply simp using ample_accept_subset apply simp apply (erule ample_accept_stuttering) apply blast done end text \Next, we implement the selection of a set of sticky edges\ context visible_prog begin definition "cr_ample \ do { sticky_E \ find_fas_init cfgc_G vis_edges; RETURN (ga_ample (sticky_E)) }" lemma cr_ample_reduction: "cr_ample \ SPEC (hl_ample_reduction prog is_vis_var)" unfolding cr_ample_def apply (refine_vcg order_trans[OF find_fas_init_correct]) apply unfold_locales[] apply simp proof clarify fix sticky_E assume "is_fas cfgc_G sticky_E" "vis_edges \ sticky_E" then interpret sticky_prog prog is_vis_var sticky_E by unfold_locales show "hl_ample_reduction prog is_vis_var (ga_ample sticky_E)" by unfold_locales qed end text \We derive an implementation for finding the feedback arcs, and an efficient implementation for ga_gample.\ text \We replace finding the sticky edges by an efficient algorithm\ context visible_prog begin definition ga_gample_m :: "_ \ _ \ pid_global_config \ global_action set" where "ga_gample_m mem sticky_E gc \ let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc; as = find_min_idx_f (\lc. let c = local_config.command lc; ls = local_config.state lc; as = set (comp.succ_impl c) in if (\(a,_)\as. read_globals a = {} \ write_globals a = {}) then let as = {(a,c')\as. la_en' (ls,gs) a} in if as={} \ (\(a,c')\as. mem (c,c') sticky_E) then None else Some (c,as) else None ) lcs in case as of Some (pid,c,as) \ (\(a,c'). (pid,c,a,c'))`as | None \ ga_gen gc" lemma ga_gample_m_refine: fixes Rs :: "('si\'s) set" shows "(ga_gample_m, ga_gample_m) \ (Id\\<^sub>rId \ Rs \ bool_rel) \ Rs \ Id \ \Id\set_rel" proof (intro fun_relI, simp) fix m :: "cmdc\cmdc \ 's \ bool" and m' :: "cmdc\cmdc \ 'si \ bool" and se se' gc assume "(m',m)\Id \ Rs \ bool_rel" and "(se',se)\Rs" hence "\c c'. (m' (c,c') se', m (c,c') se)\Id" by parametricity simp_all hence A: "\c c'. m' (c,c') se' = m (c,c') se" by simp show "ga_gample_m m' se' gc = ga_gample_m m se gc" unfolding ga_gample_m_def by (subst A) (rule refl) qed lemma ga_gample_eq_gample_m: "ga_gample = ga_gample_m (\)" unfolding ga_gample_def[abs_def] ga_gample_m_def[abs_def] by auto lemma stutter_extend_en_refine: "(stutter_extend_en, stutter_extend_en) \ (R \ \Id\set_rel) \ R \ \\Id\option_rel\set_rel" unfolding stutter_extend_en_def[abs_def] apply parametricity by auto lemma is_vis_var_refine: "(is_vis_var, is_vis_var) \ Id \ bool_rel" by simp lemma vis_edges_refine: "(\x. x\vis_edges,vis_edges)\\Id\\<^sub>rId\fun_set_rel" by (simp add: fun_set_rel_def br_def) lemma [autoref_op_pat_def]: "vis_edges \ Autoref_Tagging.OP vis_edges" "ga_gample_m \ Autoref_Tagging.OP ga_gample_m" "cfgc_G \ Autoref_Tagging.OP cfgc_G" by simp_all schematic_goal cr_ample_impl1_aux: notes [autoref_rules] = ga_gample_m_refine stutter_extend_en_refine is_vis_var_refine cfgc_G_impl_refine vis_edges_refine IdI[of prog] shows "(RETURN (?c::?'c),cr_ample)\?R" unfolding cr_ample_def ga_ample_def ga_gample_eq_gample_m using [[autoref_trace_failed_id, autoref_trace_intf_unif]] by (autoref_monadic (plain, trace)) concrete_definition cr_ample_impl1 uses cr_ample_impl1_aux lemma cr_ample_impl1_reduction: "hl_ample_reduction prog is_vis_var cr_ample_impl1" proof - note cr_ample_impl1.refine[THEN nres_relD] also note cr_ample_reduction finally show ?thesis by simp qed end text \The efficient implementation of @{const visible_prog.ga_gample} uses the @{const collecti_index} function\ lemma collecti_index_cong: assumes C: "\i. i f i (l!i) = f' i (l'!i)" assumes [simp]: "l=l'" shows "collecti_index f l = collecti_index f' l'" proof - { fix i0 s0 assume "\iix. x\set l \ f x = f' x" assumes "l=l'" shows "find_min_idx_f f l = find_min_idx_f f' l'" unfolding assms(2)[symmetric] using assms(1) apply (induction l) apply (auto split: option.split) done lemma find_min_idx_impl_collecti: "( case find_min_idx_f f l of Some (i,r) \ {i} \ s1 i r | None \ {(i,r) | i r. i r\s2 i (l!i)} ) = ( collecti_index (\i x. case f x of Some r \ (True, s1 i r) | None \ (False, s2 i x) ) l )" (is "?l=?r" is "_ = collecti_index ?fc l") proof (cases "\i None") define i where "i \ LEAST i. i < length l \ f (l ! i) \ None" have C_ALT: "(\i None) \ (\i. i < length l \ fst (?fc i (l ! i)))" by (auto split: option.split) have i_alt: "i = (LEAST i. i < length l \ fst (?fc i (l ! i)))" unfolding i_def apply (fo_rule arg_cong) by (auto split: option.split) { case True from LeastI_ex[OF True] obtain r where [simp]: "i s1 i r" by simp from collecti_index_collect[of ?fc l, folded i_alt] True C_ALT have REQ: "?r = {i} \ snd (?fc i (l!i))" by simp show "?l=?r" unfolding LEQ unfolding REQ by auto } { case False from False find_min_idx_f_LEAST_eq[of f l] have "find_min_idx_f f l = None" by simp hence LEQ: "?l = {(i,r) | i r. i r\s2 i (l!i)}" by simp from False C_ALT collecti_index_collect[of ?fc l] have REQ: "?r = {(i, x) |i x. i < length l \ x \ snd (?fc i (l ! i))}" by auto show "?l=?r" unfolding LEQ unfolding REQ using False by auto } qed lemma find_min_idx_impl_collecti_scheme: assumes "\x. s11 x = (case x of (i,r) \ {i} \ iss i r)" assumes "s12 = {(i,r) | i r. i r\isn i (l!i)}" assumes "\i. i f2 i (l!i) = (case f (l!i) of Some r \ (True, iss i r) | None \ (False, isn i (l!i)))" shows "(case find_min_idx_f f l of Some x \ s11 x | None \ s12) = (collecti_index f2 l)" apply (simp only: assms(1,2,3) cong: collecti_index_cong) apply (subst find_min_idx_impl_collecti) apply simp done schematic_goal stutter_extend_en_list_aux: "(?c, stutter_extend_en) \ (R \ \S\list_set_rel) \ R \ \\S\option_rel\list_set_rel" unfolding stutter_extend_en_def[abs_def] apply (autoref) done concrete_definition stutter_extend_en_list uses stutter_extend_en_list_aux lemma succ_of_enex_impl_aux: "{s'. \a\en s. s'=ex a s} = (\a. ex a s)`en s" by auto schematic_goal succ_of_enex_list_aux: "(?c, succ_of_enex) \ (Id \ \Id\list_set_rel) \\<^sub>r (Id \ Id \ Id) \ Id \ \Id\list_set_rel" unfolding succ_of_enex_def[abs_def] unfolding succ_of_enex_impl_aux by (autoref (trace, keep_goal)) concrete_definition succ_of_enex_list uses succ_of_enex_list_aux schematic_goal pred_of_enex_list_aux: "(?c, pred_of_enex) \ (Id \ \Id\list_set_rel) \\<^sub>r (Id \ Id \ Id) \ Id \ Id \ bool_rel" unfolding pred_of_enex_def[abs_def] by (autoref (trace, keep_goal)) concrete_definition pred_of_enex_list uses pred_of_enex_list_aux (* TODO: can this be defined using autoref? *) definition "rel_of_enex_list enex \ rel_of_pred (pred_of_enex_list enex)" lemma rel_of_enex_list_refine: "(rel_of_enex_list, rel_of_enex) \ (Id \ \Id\list_set_rel) \\<^sub>r (Id \ Id \ Id) \ \Id \\<^sub>r Id\ set_rel" proof - have 1: "rel_of_enex_list = rel_of_pred \ pred_of_enex_list" using rel_of_enex_list_def by force have 2: "rel_of_enex = rel_of_pred \ pred_of_enex" by auto have 3: "(rel_of_pred, rel_of_pred) \ (Id \ Id \ bool_rel) \ \Id \\<^sub>r Id\set_rel" by auto show ?thesis unfolding 1 2 by (parametricity add: 3 pred_of_enex_list.refine) qed context visible_prog begin definition "ga_gample_mi2 mem sticky_E (gc::pid_global_config) = ( let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc in collecti_index (\_ lc. let c = local_config.command lc; ls = local_config.state lc; as = set (comp.succ_impl c); s={(a,c')\as. la_en' (ls,gs) a}; ample = s\{} \ (\(a,_)\as. read_globals a = {} \ write_globals a = {}) \ (\(a,c')\s. \ mem (c,c') sticky_E) in (ample,{c}\s) ) lcs )" lemma ga_gample_mi2_impl: "ga_gample_m mem sticky_E gc = ga_gample_mi2 mem sticky_E gc" unfolding ga_gample_m_def ga_gample_mi2_def apply simp apply (rule find_min_idx_impl_collecti_scheme[where iss="\pid (c,S). (\(a,c'). (c,a,c'))`S" and isn="\pid lc. {(c,a,c'). cfgc c a c' \ c = cmd_of_pid gc pid \ la_en' (local_config.state lc, pid_global_config.state gc) a}"] ) apply auto [] unfolding ga_gen_def apply auto [] apply (auto simp: comp.astep_impl_def) done definition "cr_ample_impl2 \ let sticky = find_fas_init_code (=) bounded_hashcode_nat (def_hashmap_size TYPE(nat)) cfgc_G_impl (\x. x \ vis_edges) in stutter_extend_en (ga_gample_mi2 (\x f. f x) sticky)" definition "ample_step_impl2 \ rel_of_enex (cr_ample_impl2,ga_ex)" lemma cr_ample_impl2_reduction: "hl_reduction prog is_vis_var ample_step_impl2" proof - from cr_ample_impl1_reduction interpret hl_ample_reduction prog is_vis_var cr_ample_impl1 . have 1: "ample_step_impl2 = rel_of_enex (cr_ample_impl1, ga_ex)" unfolding ample_step_impl2_def cr_ample_impl1_def cr_ample_impl2_def unfolding ga_gample_mi2_impl[abs_def] by simp show ?thesis apply unfold_locales unfolding 1 apply simp apply simp apply (blast intro: ample_accept_subset ample_accept_stuttering)+ done qed end text \Implementing the State\ type_synonym valuation_impl = "(ident,uint32) mapping" definition "vi_\ \ map_option Rep_uint32 \\ Mapping.rep" abbreviation "vi_rel \ br vi_\ (\_. True)" abbreviation "val_rel \ br Rep_uint32 (\_. True)" lemma vi_rel_rew: "(x = vi_\ y) \ (y,x)\vi_rel" "(vi_\ y = x) \ (y,x)\vi_rel" by (auto simp: br_def) export_code Mapping.lookup Mapping.update Mapping.empty checking SML lift_definition vi_empty :: "valuation_impl" is "Map.empty::valuation" . lemma [code]: "vi_empty = Mapping.empty" by transfer simp lemma vi_empty_correct: "vi_\ vi_empty = Map.empty" unfolding vi_\_def by transfer auto lift_definition vi_lookup :: "valuation_impl \ ident \ uint32" is "\vi::valuation. \x. vi x" . lemma [code]: "vi_lookup = Mapping.lookup" apply (intro ext) unfolding vi_lookup_def Mapping.lookup_def map_fun_def[abs_def] o_def id_apply option.map_comp Rep_uint32_inverse option.map_ident by simp lemma vi_lookup_correct: "vi_lookup v x = map_option Abs_uint32 (vi_\ v x)" unfolding vi_\_def apply transfer apply (simp add: option.map_comp o_def option.map_ident) done lift_definition vi_update :: "ident \ uint32 \ valuation_impl \ valuation_impl" is "\x v. \vi::valuation. vi(x\v)" . lemma [code]: "vi_update = Mapping.update" by transfer simp lemma vi_update_correct: "vi_\ (vi_update k v m) = vi_\ m (k\Rep_uint32 v)" unfolding vi_\_def by transfer simp export_code vi_lookup vi_update vi_empty checking SML record local_state_impl = variables :: valuation_impl record global_state_impl = variables :: valuation_impl type_synonym focused_state_impl = "local_state_impl \ global_state_impl" type_synonym local_config_impl = "(cmdc,local_state_impl) Gen_Scheduler.local_config" type_synonym pid_global_config_impl = "(cmdc,local_state_impl,global_state_impl) Pid_Scheduler.pid_global_config" definition "local_state_rel \ { (\ local_state_impl.variables = vi \, \ local_state.variables = v \) | vi v. v = vi_\ vi }" definition "global_state_rel \ { (\ global_state_impl.variables = vi \, \ global_state.variables = v \) | vi v. v = vi_\ vi }" abbreviation "focused_state_rel \ local_state_rel \\<^sub>r global_state_rel" definition "local_config_rel \ { (\ local_config.command = ci, local_config.state=lsi \, \ local_config.command = c, local_config.state=ls \) | ci c lsi ls. (ci::cmdc,c)\Id \ (lsi,ls)\local_state_rel }" definition "global_config_rel \ { (\ pid_global_config.processes = psi, pid_global_config.state = gsi \, \ pid_global_config.processes = ps, pid_global_config.state = gs \) | psi ps gsi gs. (psi,ps)\\local_config_rel\list_rel \ (gsi,gs)\global_state_rel }" definition "init_valuation_impl vd \ fold (\x v. vi_update x 0 v) vd (vi_empty)" lemma init_valuation_impl: "(init_valuation_impl, init_valuation) \ \Id\list_rel \ vi_rel" proof - { fix vd m have "m ++ init_valuation vd = fold (\x v. v(x\0)) vd m" apply (rule sym) apply (induction vd arbitrary: m) unfolding init_valuation_def[abs_def] apply (auto intro!: ext simp: Map.map_add_def) done } note aux=this have E1: "\vd. init_valuation vd = fold (\x v. v(x\0)) vd Map.empty" by (subst aux[symmetric]) simp have E2: "\vdi vd. (vdi,vd)\\Id\list_rel \ (init_valuation_impl vdi, fold (\x v. v(x\0)) vd Map.empty) \ vi_rel" unfolding init_valuation_impl_def apply parametricity apply (auto simp: vi_update_correct vi_empty_correct br_def zero_uint32.rep_eq) done show ?thesis apply (intro fun_relI) apply (subst E1) apply (rule E2) by auto qed context cprog begin definition init_pc_impl :: "proc_decl \ local_config_impl" where "init_pc_impl pd \ \ local_config.command = comp.\ (proc_decl.body pd), local_config.state = \ local_state_impl.variables = init_valuation_impl (proc_decl.local_vars pd) \ \" lemma init_pc_impl: "(init_pc_impl, init_pc) \ Id \ local_config_rel" apply (rule fun_relI) unfolding init_pc_impl_def init_pc_def local_config_rel_def local_state_rel_def apply (simp del: pair_in_Id_conv, intro conjI) apply simp apply (simp only: vi_rel_rew) apply (parametricity add: init_valuation_impl) apply simp done definition pid_init_gc_impl :: "pid_global_config_impl" where "pid_init_gc_impl \ \ pid_global_config.processes = (map init_pc_impl (program.processes prog)), pid_global_config.state = \ global_state_impl.variables = init_valuation_impl (program.global_vars prog) \ \" lemma pid_init_gc_impl: "(pid_init_gc_impl, pid_init_gc) \ global_config_rel" unfolding pid_init_gc_impl_def pid_init_gc_def global_config_rel_def global_state_rel_def apply (simp del: pair_in_Id_conv, intro conjI) apply (parametricity add: init_pc_impl) apply simp apply (simp only: vi_rel_rew) apply (parametricity add: init_valuation_impl) apply simp done end lift_definition val_of_bool_impl :: "bool \ uint32" is val_of_bool . lemma [code]: "val_of_bool_impl b = (if b then 1 else 0)" by transfer auto lift_definition bool_of_val_impl :: "uint32 \ bool" is bool_of_val . lemma [code]: "bool_of_val_impl v = (v\0)" by transfer (auto simp: bool_of_val_def) definition "set_of_rel R \ {(a,b) . R b a}" definition "rel_of_set S \ \b a. (a,b)\S" lemma [simp]: "(a,b)\set_of_rel R \ R b a" "rel_of_set S b a \ (a,b)\S" unfolding rel_of_set_def set_of_rel_def by auto lemma [simp]: "set_of_rel (rel_of_set S) = S" "rel_of_set (set_of_rel R) = R" unfolding rel_of_set_def set_of_rel_def by auto lemma rel_fun_2set: "rel_fun A B = rel_of_set (set_of_rel A \ set_of_rel B)" unfolding rel_fun_def fun_rel_def unfolding rel_of_set_def set_of_rel_def by (auto) lemma [simp]: "set_of_rel cr_uint32 = val_rel" "set_of_rel (=) = Id" by (auto simp: br_def cr_uint32_def) lemma bool_of_val_impl: "(bool_of_val_impl, bool_of_val) \ val_rel \ bool_rel" using bool_of_val_impl.transfer by (simp add: rel_fun_2set) lemma smod_by_div_abs: "a smod b = a - a sdiv b * b" for a b :: \'a::len word\ by (subst (2) sdiv_smod_id [of a b, symmetric]) simp lift_definition sdiv_impl :: "uint32 \ uint32 \ uint32" is "(sdiv)" . lift_definition smod_impl :: "uint32 \ uint32 \ uint32" is "(smod)" . (* TODO: Is there a more efficient way? *) lemma [code]: "sdiv_impl x y = Abs_uint32' (Rep_uint32' x sdiv Rep_uint32' y)" apply transfer by simp lemma [code]: "smod_impl a b = a - sdiv_impl a b * b" by transfer (simp add: smod_by_div_abs) - + +context + includes bit_operations_syntax +begin primrec eval_bin_op_impl_aux :: "bin_op \ uint32 \ uint32 \ uint32" where "eval_bin_op_impl_aux bo_plus v1 v2 = v1+v2" | "eval_bin_op_impl_aux bo_minus v1 v2 = v1-v2" | "eval_bin_op_impl_aux bo_mul v1 v2 = v1*v2" | "eval_bin_op_impl_aux bo_div v1 v2 = sdiv_impl v1 v2" | "eval_bin_op_impl_aux bo_mod v1 v2 = smod_impl v1 v2" | "eval_bin_op_impl_aux bo_less v1 v2 = val_of_bool_impl (v1 < v2)" | "eval_bin_op_impl_aux bo_less_eq v1 v2 = val_of_bool_impl (v1 \ v2)" | "eval_bin_op_impl_aux bo_eq v1 v2 = val_of_bool_impl (v1 = v2)" | "eval_bin_op_impl_aux bo_and v1 v2 = v1 AND v2" | "eval_bin_op_impl_aux bo_or v1 v2 = v1 OR v2" | "eval_bin_op_impl_aux bo_xor v1 v2 = v1 XOR v2" +end + lift_definition eval_bin_op_impl :: "bin_op \ uint32 \ uint32 \ uint32" is eval_bin_op . -print_theorems lemma [code]: "eval_bin_op_impl bop v1 v2 = eval_bin_op_impl_aux bop v1 v2" apply (cases bop) - apply simp_all - apply (transfer, simp)+ - done - - + apply (transfer; simp)+ + done primrec eval_un_op_impl_aux :: "un_op \ uint32 \ uint32" where "eval_un_op_impl_aux uo_minus v = -v" | "eval_un_op_impl_aux uo_not v = NOT v" lift_definition eval_un_op_impl :: "un_op \ uint32 \ uint32" is eval_un_op . lemma [code]: "eval_un_op_impl uop v = eval_un_op_impl_aux uop v" apply (cases uop) apply simp_all apply (transfer, simp)+ done fun eval_exp_impl :: "exp \ focused_state_impl \ uint32" where "eval_exp_impl (e_var x) (ls,gs) = do { let lv = local_state_impl.variables ls; let gv = global_state_impl.variables gs; case vi_lookup lv x of Some v \ Some v | None \ (case vi_lookup gv x of Some v \ Some v | None \ None) }" | "eval_exp_impl (e_localvar x) (ls,gs) = vi_lookup (local_state_impl.variables ls) x" | "eval_exp_impl (e_globalvar x) (ls,gs) = vi_lookup (global_state_impl.variables gs) x" | "eval_exp_impl (e_const n) fs = do { assert_option (n\min_signed \ n\max_signed); Some (uint32_of_int n) }" | "eval_exp_impl (e_bin bop e1 e2) fs = do { v1\eval_exp_impl e1 fs; v2\eval_exp_impl e2 fs; Some (eval_bin_op_impl bop v1 v2) }" | "eval_exp_impl (e_un uop e) fs = do { v\eval_exp_impl e fs; Some (eval_un_op_impl uop v) }" (* TODO: Move *) lemma map_option_bind: "map_option f (m\g) = m \ (map_option f o g)" by (auto split: Option.bind_split) lemma val_option_rel_as_eq: "(a,b)\\val_rel\option_rel \ map_option Rep_uint32 a = b" unfolding br_def apply (cases a) apply (cases b, simp_all) apply (cases b, auto) done lemma eval_exp_impl: "(eval_exp_impl, eval_exp) \ Id \ focused_state_rel \ \val_rel\option_rel" proof - { fix e ls gs lsi gsi assume "(lsi,ls)\local_state_rel" "(gsi,gs)\global_state_rel" hence "map_option Rep_uint32 (eval_exp_impl e (lsi,gsi)) = eval_exp e (ls,gs)" apply (induction e) apply (auto simp: local_state_rel_def global_state_rel_def br_def simp: vi_lookup_correct vi_update_correct simp: Abs_uint32_inverse[simplified] uint32_of_int.rep_eq signed_of_int_def simp: option.map_comp option.map_ident o_def map_option_bind simp: eval_bin_op_impl.rep_eq eval_un_op_impl.rep_eq split: option.splits) apply (drule sym) apply (drule sym) apply (simp split: Option.bind_split) apply (drule sym) apply (simp split: Option.bind_split) done } thus ?thesis by (auto simp: val_option_rel_as_eq) qed text \Enabledness and effects of actions\ primrec la_en_impl :: "focused_state_impl \ action \ bool" where "la_en_impl fs (AAssign e _ _) = do { v \ eval_exp_impl e fs; Some (bool_of_val_impl v)}" | "la_en_impl fs (AAssign_local e _ _) = do { v \ eval_exp_impl e fs; Some (bool_of_val_impl v)}" | "la_en_impl fs (AAssign_global e _ _) = do { v \ eval_exp_impl e fs; Some (bool_of_val_impl v)}" | "la_en_impl fs (ATest e) = do { v \ eval_exp_impl e fs; Some (bool_of_val_impl v)}" | "la_en_impl _ (ASkip) = Some True" lemma param_rec_action[param]: "(rec_action, rec_action) \ (Id \ Id \ Id \ A) \ (Id \ Id \ Id \ A) \ (Id \ Id \ Id \ A) \ (Id \ A) \ A \ Id \ A" apply (intro fun_relI) subgoal for fi1 f1 fi2 f2 fi3 f3 fi4 f4 fi5 f5 ai a apply simp apply (induction a) apply simp_all apply (parametricity, simp_all?)+ done done (* TODO: Move *) lemma param_option_bind[param]: "(Option.bind, Option.bind) \ \A\option_rel \ (A \ \B\option_rel) \ \B\option_rel" unfolding Option.bind_def by parametricity lemma la_en_impl: "(la_en_impl,la_en) \ focused_state_rel \ Id \ \Id\option_rel" unfolding la_en_impl_def la_en_def by (parametricity add: eval_exp_impl bool_of_val_impl) definition "la_en'_impl fs a \ case la_en_impl fs a of Some b \ b | None \ False" lemma la_en'_impl: "(la_en'_impl,la_en') \ focused_state_rel \ Id \ bool_rel" unfolding la_en'_impl_def[abs_def] la_en'_def[abs_def] by (parametricity add: la_en_impl) abbreviation "exists_var_impl v x \ vi_lookup v x \ None" fun la_ex_impl :: "focused_state_impl \ action \ focused_state_impl" where "la_ex_impl (ls,gs) (AAssign ge x e) = do { v \ eval_exp_impl ge (ls,gs); assert_option (bool_of_val_impl v); v \ eval_exp_impl e (ls,gs); if exists_var_impl (local_state_impl.variables ls) x then do { let ls = local_state_impl.variables_update (vi_update x v) ls; Some (ls,gs) } else do { assert_option (exists_var_impl (global_state_impl.variables gs) x); let gs = global_state_impl.variables_update (vi_update x v) gs; Some (ls,gs) } }" | "la_ex_impl (ls,gs) (AAssign_local ge x e) = do { v \ eval_exp_impl ge (ls,gs); assert_option (bool_of_val_impl v); v \ eval_exp_impl e (ls,gs); assert_option (exists_var_impl (local_state_impl.variables ls) x); let ls = local_state_impl.variables_update (vi_update x v) ls; Some (ls,gs) }" | "la_ex_impl (ls,gs) (AAssign_global ge x e) = do { v \ eval_exp_impl ge (ls,gs); assert_option (bool_of_val_impl v); v \ eval_exp_impl e (ls,gs); assert_option (exists_var_impl (global_state_impl.variables gs) x); let gs = global_state_impl.variables_update (vi_update x v) gs; Some (ls,gs) }" | "la_ex_impl fs (ATest e) = do { v \ eval_exp_impl e fs; assert_option (bool_of_val_impl v); Some fs }" | "la_ex_impl fs ASkip = Some fs" (* TODO: Move *) lemma param_assert_option[param]: "(assert_option, assert_option) \ bool_rel \ \unit_rel\option_rel" unfolding assert_option_def by parametricity lemma la_ex_impl: "(la_ex_impl, la_ex) \ focused_state_rel \ Id \ \focused_state_rel\option_rel" proof (intro fun_relI) fix fsi fs and ai a :: action assume "(fsi,fs)\focused_state_rel" "(ai,a)\Id" thus "(la_ex_impl fsi ai, la_ex fs a)\\focused_state_rel\option_rel" apply (cases fs, cases fsi, clarsimp) apply (cases a) apply simp_all apply (parametricity add: eval_exp_impl bool_of_val_impl) apply (auto simp: local_state_rel_def global_state_rel_def simp: vi_lookup_correct vi_update_correct simp: br_def intro: domI) [7] apply (parametricity add: eval_exp_impl bool_of_val_impl) apply (auto simp: local_state_rel_def global_state_rel_def simp: vi_lookup_correct vi_update_correct br_def intro: domI) [9] apply (parametricity add: eval_exp_impl bool_of_val_impl) apply (auto simp: local_state_rel_def global_state_rel_def simp: vi_lookup_correct vi_update_correct br_def intro: domI) [4] apply (parametricity add: eval_exp_impl bool_of_val_impl) apply simp done qed definition "la_ex'_impl fs a \ case la_ex_impl fs a of Some fs' \ fs' | None \ fs" lemma la_ex'_impl: "(la_ex'_impl,la_ex') \ focused_state_rel \ Id \ focused_state_rel" unfolding la_ex'_impl_def[abs_def] la_ex'_def[abs_def] by (parametricity add: la_ex_impl) lemma param_collecti_index: "(collecti_index, collecti_index) \ (nat_rel \ B \ bool_rel \\<^sub>r \Id\set_rel) \ \B\list_rel \ \nat_rel \\<^sub>r Id\set_rel" unfolding collecti_index'_def apply parametricity apply auto done export_code la_ex_impl la_en_impl checking SML context visible_prog begin definition "ga_gample_mi3 mem sticky_E (gc::pid_global_config_impl) = ( let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc in collecti_index (\_ lc. let c = local_config.command lc; ls = local_config.state lc; as = set (comp.succ_impl c); s={(a,c')\as. la_en'_impl (ls,gs) a}; ample = s\{} \ (\(a,_)\as. read_globals a = {} \ write_globals a = {}) \ (\(a,c')\s. \ mem (c,c') sticky_E) in (ample,{c}\s) ) lcs )" lemma ga_gample_mi3_refine: "(ga_gample_mi3, ga_gample_mi2)\ Id \ Id \ global_config_rel \ \nat_rel \\<^sub>r Id\set_rel" proof - have [param]: "(comp.succ_impl,comp.succ_impl)\Id\\Id\\<^sub>rId\list_rel" by simp from la_en'_impl have aux1: "\fsi fs a. (fsi,fs)\focused_state_rel \ la_en'_impl fsi a = la_en' fs a" apply (rule_tac IdD) apply parametricity by auto show ?thesis using [[goals_limit=1]] apply (intro fun_relI) unfolding ga_gample_mi3_def ga_gample_mi2_def apply (parametricity add: param_collecti_index la_en'_impl) apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply simp using aux1 apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply simp_all done qed definition (in cprog) "is_vis_edge is_vis_var \ \(c,c'). \(a,cc')\set (comp.succ_impl c). c'=cc' \ (\v\write_globals a. is_vis_var v)" lemma is_vis_edge_correct: "is_vis_edge is_vis_var = (\x. x\vis_edges)" unfolding is_vis_edge_def vis_edges_def by (auto intro!: ext simp: comp.astep_impl_def) definition "cr_ample_impl3 \ let sticky = find_fas_init_code (=) bounded_hashcode_nat (def_hashmap_size TYPE(nat)) cfgc_G_impl (\x. x \ vis_edges) in stutter_extend_en (ga_gample_mi3 (\x f. f x) sticky)" lemma cr_ample_impl3_refine: "(cr_ample_impl3, cr_ample_impl2) \ global_config_rel \ \\nat_rel\\<^sub>rId\option_rel\set_rel" proof - have [param]: "\R S. S=Id \ (stutter_extend_en, stutter_extend_en) \ (R \ \S\set_rel) \ R \ \\S\option_rel\set_rel" apply (simp only: ) by (rule stutter_extend_en_refine) show ?thesis unfolding cr_ample_impl3_def cr_ample_impl2_def is_vis_edge_correct apply (parametricity add: stutter_extend_en_refine ga_gample_mi3_refine) apply (rule IdI) by auto qed (* TODO: The current Pid_Scheduler-locale is focused on refinement. We cannot use it to get an instantiation of ga_gex_impl.*) definition (in -) ga_gex_impl :: "global_action \ pid_global_config_impl \ pid_global_config_impl" where "ga_gex_impl ga gc \ let (pid,c,a,c') = ga; fs = fs_of_pid gc pid; (ls',gs') = la_ex'_impl fs a; gc' = \ pid_global_config.processes = (pid_global_config.processes gc) [pid := \ local_config.command = c', local_config.state = ls'\ ], pid_global_config.state = gs'\ in gc'" lemma lc_of_pid_refine: assumes V: "pid_valid gc pid" assumes [param]: "(pidi,pid)\nat_rel" assumes GCI: "(gci,gc)\global_config_rel" shows "(lc_of_pid gci pidi, lc_of_pid gc pid)\local_config_rel" apply parametricity apply fact using GCI by (auto simp: global_config_rel_def) lemma ls_of_pid_refine: assumes "pid_valid gc pid" assumes "(pidi,pid)\nat_rel" assumes "(gci,gc)\global_config_rel" shows "(ls_of_pid gci pidi, ls_of_pid gc pid)\local_state_rel" using lc_of_pid_refine[OF assms] unfolding local_config_rel_def by auto lemma ga_gex_impl: assumes V: "pid_valid gc pid" assumes [param]: "(pidi,pid)\Id" "(ci,c)\Id" "(ai,a)\Id" and CI'I: "(ci',c')\Id" and GCI[param]: "(gci,gc)\global_config_rel" shows "(ga_gex_impl (pidi,ci,ai,ci') gci, ga_gex (pid,c,a,c') gc) \ global_config_rel" unfolding ga_gex_impl_def[abs_def] ga_gex_def[abs_def] apply simp apply (parametricity add: la_ex'_impl ls_of_pid_refine[OF V]) using GCI CI'I apply (auto simp: global_config_rel_def local_config_rel_def) [] apply parametricity apply auto [] using GCI apply (auto simp: global_config_rel_def) [] done definition (in -) "ga_ex_impl \ stutter_extend_ex ga_gex_impl" fun ga_valid where "ga_valid gc None = True" | "ga_valid gc (Some (pid,c,a,c')) = pid_valid gc pid" lemma ga_ex_impl: assumes V: "ga_valid gc ga" assumes P: "(gai,ga)\Id" "(gci,gc)\global_config_rel" shows "(ga_ex_impl gai gci, ga_ex ga gc) \ global_config_rel" using assms apply (cases "(gc,ga)" rule: ga_valid.cases) using P apply (simp add: ga_ex_impl_def) apply (simp add: ga_ex_impl_def ga_ex_def) apply (parametricity add: ga_gex_impl) by auto definition "ample_step_impl3 \ rel_of_enex (cr_ample_impl3,ga_ex_impl)" definition (in -) pid_interp_gc_impl :: "(ident \ bool) \ pid_global_config_impl \ exp set" where "pid_interp_gc_impl is_vis_var gci \ sm_props ( vi_\ (global_state_impl.variables (pid_global_config.state gci)) |` Collect is_vis_var )" lemma pid_interp_gc_impl: "(pid_interp_gc_impl is_vis_var, pid_interp_gc) \ global_config_rel \ Id" apply rule unfolding pid_interp_gc_impl_def pid_interp_gc_def apply (auto simp: global_config_rel_def global_state_rel_def interp_gs_def sm_props_def br_def) done lemma cr_ample_impl3_sim_aux: assumes GCI: "(gci, gc) \ global_config_rel" assumes GA: "ga \ cr_ample_impl3 gci" shows "\gc'. ga\cr_ample_impl2 gc \ gc'=ga_ex ga gc \ (ga_ex_impl ga gci, gc')\global_config_rel" proof (clarsimp, safe) show G1: "ga\cr_ample_impl2 gc" using cr_ample_impl3_refine[param_fo, OF GCI] GA by auto { (* TODO: We are re-proving the validity of the action, due to the lack of a clean refinement argumentation *) fix mem sticky_E gc pid c assume "(pid,c)\ga_gample_m mem sticky_E gc" hence "pid_valid gc pid" unfolding ga_gample_m_def apply (auto split: option.splits) apply (auto simp: ga_gen_def) [] apply (auto dest: find_min_idx_f_SomeD) done } note aux=this from G1 have "ga_valid gc ga" unfolding cr_ample_impl2_def stutter_extend_en_def apply (simp split: if_split_asm) unfolding ga_gample_mi2_impl[symmetric] apply (auto dest: aux) done thus "(ga_ex_impl ga gci, ga_ex ga gc)\global_config_rel" by (parametricity add: ga_ex_impl GCI IdI[of ga]) qed lemma cr_ample_impl3_sim_aux2: assumes GCI: "(gci, gc) \ global_config_rel" assumes GA: "ga \ cr_ample_impl2 gc" shows "ga\cr_ample_impl3 gci \ (ga_ex_impl ga gci, ga_ex ga gc) \ global_config_rel" proof show G1: "ga\cr_ample_impl3 gci" using cr_ample_impl3_refine[param_fo, OF GCI] GA by auto from cr_ample_impl3_sim_aux[OF GCI G1] show "(ga_ex_impl ga gci, ga_ex ga gc) \ global_config_rel" by auto qed sublocale impl3: sa "\ g_V = UNIV, g_E = ample_step_impl3, g_V0 = {pid_init_gc_impl}, sa_L = pid_interp_gc_impl is_vis_var \" by unfold_locales auto sublocale impl3_sim: lbisimulation global_config_rel "\ g_V = UNIV, g_E = ample_step_impl3, g_V0 = {pid_init_gc_impl}, sa_L = pid_interp_gc_impl is_vis_var \" "\ g_V = UNIV, g_E = ample_step_impl2, g_V0 = {pid_init_gc}, sa_L = pid_interp_gc \" apply unfold_locales apply simp apply simp apply simp using pid_init_gc_impl apply (auto dest: fun_relD) [] unfolding ample_step_impl2_def ample_step_impl3_def apply (auto dest: cr_ample_impl3_sim_aux) [] using pid_interp_gc_impl apply (auto dest: fun_relD) [] apply simp using pid_init_gc_impl apply (auto dest: fun_relD) [] defer using pid_interp_gc_impl apply (auto dest: fun_relD) [] apply auto apply (auto dest: cr_ample_impl3_sim_aux2) [] done text \Correctness of impl3\ (* TODO: Locale hl_reduction seems very special, and we cannot use it here*) lemma impl3_accept_subset: "impl3.accept w \ lv.sa.accept w" using impl3_sim.accept_bisim cr_ample_impl2_reduction hl_reduction.ample_accept_subset by simp lemma impl3_accept_stuttering: "lv.sa.accept w \ \w'. w\w' \ impl3.accept w'" using impl3_sim.accept_bisim cr_ample_impl2_reduction hl_reduction.ample_accept_stuttering by simp text \Next, we go from sets of actions to lists of actions\ definition (in cprog) "ga_gample_mil4 mem sticky_Ei (gc::pid_global_config_impl) = ( let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc in collecti_index_list (\_ lc. let c = local_config.command lc; ls = local_config.state lc; as = comp.succ_impl c; s = filter (la_en'_impl (ls,gs) o fst) as; ample = s\[] \ (\(a,_)\set as. read_globals a = {} \ write_globals a = {}) \ (\(a,c')\set s. \mem (c,c') sticky_Ei) in (ample,map (Pair c) s) ) lcs )" lemma ga_gample_mil4_refine: "(ga_gample_mil4, ga_gample_mi3) \ (Id\\<^sub>rId \ Id \ bool_rel) \ Id \ Id \ \nat_rel \\<^sub>r Id\list_set_rel" apply (intro fun_relI) unfolding ga_gample_mil4_def ga_gample_mi3_def apply (parametricity) apply simp apply (rule IdI) apply simp apply (rule IdI) apply (intro fun_relI) apply (rule collecti_index_list_refine[param_fo]) apply (intro fun_relI) apply (simp add: list_set_rel_def br_def distinct_map comp.succ_impl_invar) apply (auto simp: filter_empty_conv) apply force done definition (in cprog) "cr_ample_impl4 is_vis_var \ let sticky = find_fas_init_code (=) bounded_hashcode_nat (def_hashmap_size TYPE(nat)) cfgc_G_impl (is_vis_edge is_vis_var) in stutter_extend_en_list (ga_gample_mil4 (\x f. f x) sticky)" lemma cr_ample_impl4_refine: "(cr_ample_impl4 is_vis_var, cr_ample_impl3)\Id \ \\nat_rel \\<^sub>r Id\option_rel\list_set_rel" unfolding cr_ample_impl4_def cr_ample_impl3_def apply (parametricity add: stutter_extend_en_list.refine ga_gample_mil4_refine) apply (simp_all add: is_vis_edge_correct) done text \Finally, we combine the ample-implementation and the executable implementation to get a step function\ definition (in cprog) "ample_step_impl4 is_vis_var \ rel_of_enex_list (cr_ample_impl4 is_vis_var,ga_ex_impl)" lemma ample_step_impl4: "(ample_step_impl4 is_vis_var, ample_step_impl3) \ \Id \\<^sub>r Id\ set_rel" unfolding ample_step_impl4_def ample_step_impl3_def rel_of_pred_def apply (parametricity add: rel_of_enex_list_refine cr_ample_impl4_refine[simplified]) by auto sublocale impl4: sa "\ g_V = UNIV, g_E = ample_step_impl4 is_vis_var, g_V0 = {pid_init_gc_impl}, sa_L = pid_interp_gc_impl is_vis_var \" by unfold_locales auto lemma impl4_accept_subset: "impl4.accept w \ lv.sa.accept w" using impl3_accept_subset ample_step_impl4 by simp lemma impl4_accept_stuttering: "lv.sa.accept w \ \w'. w\w' \ impl4.accept w'" using impl3_accept_stuttering ample_step_impl4 by simp lemma (in -) pred_of_succ_of_enex_list: fixes enex :: "('c \ 'a list) \ ('a \ 'c \ 'c)" shows "pred_of_succ (set o succ_of_enex_list enex) = pred_of_enex_list enex" proof - { fix x and l have "glist_member (=) x l \ x\set l" by (induction l) auto } note [simp] = this { fix l :: "'a list" and g and l0 have "set (foldli l (\_. True) (\x. glist_insert (=) (g x)) l0) = set l0 \ g`set l" by (induction l arbitrary: l0) (auto simp: glist_insert_def) } note [simp] = this { fix l::"'a list" and P i have "foldli l Not (\x _. P x) i \ i \ (\x\set l. P x)" by (induction l arbitrary: i) auto } note [simp] = this show ?thesis unfolding succ_of_enex_list_def[abs_def] pred_of_enex_list_def[abs_def] by (auto simp: pred_of_succ_def gen_image_def gen_bex_def intro!: ext) qed lemma (in -) rel_of_succ_of_enex_list: fixes enex :: "('c \ 'a list) \ ('a \ 'c \ 'c)" shows "rel_of_succ (set o succ_of_enex_list enex) = rel_of_enex_list enex" unfolding rel_of_enex_list_def unfolding pred_of_succ_of_enex_list[symmetric] by simp definition (in cprog) "impl4_succ is_vis_var \ succ_of_enex_list (cr_ample_impl4 is_vis_var,ga_ex_impl)" lemma impl4_succ_pred: "rel_of_succ (set o (impl4_succ is_vis_var)) = ample_step_impl4 is_vis_var" unfolding impl4_succ_def ample_step_impl4_def by (simp add: rel_of_succ_of_enex_list) end export_code ga_ex_impl checking SML definition "ccfg_E_succ_impl mst \ remdups o map snd o ccfg_succ_impl mst" lemma (in cprog) ccfg_E_succ_impl: "ccfg_E_succ_impl (comp_info_of prog) = cfgc_E_succ" apply (intro ext) unfolding ccfg_E_succ_impl_def cfgc_E_succ_def by (simp add: ccfg_succ_impl) definition init_pc_impl_impl :: "comp_info \ proc_decl \ local_config_impl" where "init_pc_impl_impl mst pd \ \ local_config.command = comp_\_impl mst (proc_decl.body pd), local_config.state = \ local_state_impl.variables = init_valuation_impl (proc_decl.local_vars pd) \ \" lemma (in cprog) init_pc_impl_impl: "init_pc_impl_impl (comp_info_of prog) = init_pc_impl" apply (intro ext) unfolding init_pc_impl_impl_def init_pc_impl_def by (simp add: comp_\_impl) definition pid_init_gc_impl_impl :: "program \ comp_info \ pid_global_config_impl" where "pid_init_gc_impl_impl prog mst \ \ pid_global_config.processes = (map (init_pc_impl_impl mst) (program.processes prog)), pid_global_config.state = \ global_state_impl.variables = init_valuation_impl (program.global_vars prog) \ \" lemma (in cprog) pid_init_gc_impl_impl: "pid_init_gc_impl_impl prog (comp_info_of prog) = pid_init_gc_impl" unfolding pid_init_gc_impl_impl_def pid_init_gc_impl_def by (simp add: init_pc_impl_impl) definition "ccfg_V0_impl prog mst \ map (comp_\_impl mst) (cfg_V0_list prog)" lemma (in cprog) ccfg_V0_impl: "ccfg_V0_impl prog (comp_info_of prog) = cfgc_V0_list" unfolding ccfg_V0_impl_def cfgc_V0_list_def by (simp add: comp_\_impl) definition "ccfg_G_impl_impl prog mst \ \ gi_V = \ _. True, gi_E = ccfg_E_succ_impl mst, gi_V0 = ccfg_V0_impl prog mst \" lemma (in cprog) ccfg_G_impl_impl: "ccfg_G_impl_impl prog (comp_info_of prog) = cfgc_G_impl" unfolding ccfg_G_impl_impl_def cfgc_G_impl_def unfolding ccfg_E_succ_impl ccfg_V0_impl by rule definition "is_vis_edge_impl mst is_vis_var \ \(c,c'). \(a,cc')\set (ccfg_succ_impl mst c). c'=cc' \ (\v\write_globals a. is_vis_var v)" lemma (in cprog) is_vis_edge_impl: "is_vis_edge_impl (comp_info_of prog) = is_vis_edge" apply (intro ext) unfolding is_vis_edge_impl_def is_vis_edge_def by (simp add: ccfg_succ_impl) definition "ga_gample_mil4_impl mst mem sticky_Ei (gc::pid_global_config_impl) = ( let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc in collecti_index_list (\_ lc. let c = local_config.command lc; ls = local_config.state lc; as = ccfg_succ_impl mst c; s = filter (la_en'_impl (ls,gs) o fst) as; ample = s\[] \ (\(a,_)\set as. read_globals a = {} \ write_globals a = {}) \ (\(a,c')\set s. \mem (c,c') sticky_Ei) in (ample,map (Pair c) s) ) lcs )" lemma (in cprog) ga_gample_mil4_impl: "ga_gample_mil4_impl (comp_info_of prog) = ga_gample_mil4" apply (intro ext) unfolding ga_gample_mil4_impl_def ga_gample_mil4_def by (simp add: ccfg_succ_impl) definition "cr_ample_impl4_impl prog mst is_vis_var \ let sticky = find_fas_init_code (=) bounded_hashcode_nat (def_hashmap_size TYPE(nat)) (ccfg_G_impl_impl prog mst) (is_vis_edge_impl mst is_vis_var) in stutter_extend_en_list (ga_gample_mil4_impl mst (\x f. f x) sticky)" lemma (in cprog) cr_ample_impl4_impl: "cr_ample_impl4_impl prog (comp_info_of prog) = cr_ample_impl4" apply (intro ext) unfolding cr_ample_impl4_impl_def cr_ample_impl4_def by (simp add: ga_gample_mil4_impl ccfg_G_impl_impl is_vis_edge_impl) definition "impl4_succ_impl prog mst is_vis_var \ succ_of_enex_list (cr_ample_impl4_impl prog mst is_vis_var,ga_ex_impl)" lemma (in cprog) impl4_succ_impl: "impl4_succ_impl prog (comp_info_of prog) = impl4_succ" apply (intro ext) unfolding impl4_succ_impl_def impl4_succ_def by (simp add: cr_ample_impl4_impl) definition "ample_step_impl4_impl prog mst is_vis_var \ rel_of_enex_list (cr_ample_impl4_impl prog mst is_vis_var,ga_ex_impl)" lemma (in cprog) ample_step_impl4_impl: "ample_step_impl4_impl prog (comp_info_of prog) = ample_step_impl4" apply (intro ext) unfolding ample_step_impl4_impl_def ample_step_impl4_def by (simp add: cr_ample_impl4_impl) export_code impl4_succ_impl pid_init_gc_impl_impl comp_info_of checking SML end diff --git a/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_Semantics.thy b/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_Semantics.thy --- a/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_Semantics.thy +++ b/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_Semantics.thy @@ -1,151 +1,158 @@ section \Semantics of SM\ theory SM_Semantics imports SM_State SM_Cfg Gen_Scheduler begin section \Evaluation of Expressions\ subsection \Basic operations\ text \Attention: Silently overflows, using 2's complement. This should match Java-semantics. In C, signed overflow is undefined!\ + + context + includes bit_operations_syntax + begin + primrec eval_bin_op :: "bin_op \ val \ val \ val" where "eval_bin_op bo_plus v1 v2 = v1+v2" | "eval_bin_op bo_minus v1 v2 = v1-v2" | "eval_bin_op bo_mul v1 v2 = v1*v2" | "eval_bin_op bo_div v1 v2 = v1 sdiv v2" | "eval_bin_op bo_mod v1 v2 = v1 smod v2" | "eval_bin_op bo_less v1 v2 = val_of_bool (v1 < v2)" | "eval_bin_op bo_less_eq v1 v2 = val_of_bool (v1 \ v2)" | "eval_bin_op bo_eq v1 v2 = val_of_bool (v1 = v2)" | "eval_bin_op bo_and v1 v2 = v1 AND v2" | "eval_bin_op bo_or v1 v2 = v1 OR v2" | "eval_bin_op bo_xor v1 v2 = v1 XOR v2" primrec eval_un_op :: "un_op \ val \ val" where "eval_un_op uo_minus v = -v" \ \Attention: Silently overflows. @{term "-min_int = min_int"}!\ | "eval_un_op uo_not v = NOT v" - + + end + subsection \Expressions\ abbreviation exists_var :: "valuation \ ident \ bool" where "exists_var s x \ x \ dom s" abbreviation update_var :: "ident \ val \ valuation \ valuation" where "update_var x v s \ s(x\v)" fun eval_exp :: "exp \ focused_state \ val" where "eval_exp (e_var x) (ls,gs) = do { let lv = local_state.variables ls; let gv = global_state.variables gs; case lv x of Some v \ Some v | None \ (case gv x of Some v \ Some v | None \ None) }" | "eval_exp (e_localvar x) (ls,gs) = local_state.variables ls x" | "eval_exp (e_globalvar x) (ls,gs) = global_state.variables gs x" | "eval_exp (e_const n) fs = do { assert_option (n\min_signed \ n\max_signed); Some (signed_of_int n) }" | "eval_exp (e_bin bop e1 e2) fs = do { v1\eval_exp e1 fs; v2\eval_exp e2 fs; Some (eval_bin_op bop v1 v2) }" | "eval_exp (e_un uop e) fs = do { v\eval_exp e fs; Some (eval_un_op uop v) }" subsection \Local Actions\ text \Enabledness and effects of actions\ primrec la_en :: "focused_state \ action \ bool" where "la_en fs (AAssign c _ _) = do { v \ eval_exp c fs; Some (bool_of_val v)}" | "la_en fs (AAssign_local c _ _) = do { v \ eval_exp c fs; Some (bool_of_val v)}" | "la_en fs (AAssign_global c _ _) = do { v \ eval_exp c fs; Some (bool_of_val v)}" | "la_en fs (ATest e) = do { v \ eval_exp e fs; Some (bool_of_val v)}" | "la_en _ (ASkip) = Some True" text \Extension (for run, handshake): Let action have scheduler effect\ fun la_ex :: "focused_state \ action \ focused_state" where "la_ex (ls,gs) (AAssign c x e) = do { u \ eval_exp c (ls,gs); assert_option (bool_of_val u); v \ eval_exp e (ls,gs); if exists_var (local_state.variables ls) x then do { let ls = local_state.variables_update (update_var x v) ls; Some (ls,gs) } else do { assert_option (exists_var (global_state.variables gs) x); let gs = global_state.variables_update (update_var x v) gs; Some (ls,gs) } }" | "la_ex (ls,gs) (AAssign_local c x e) = do { u \ eval_exp c (ls,gs); assert_option (bool_of_val u); v \ eval_exp e (ls,gs); assert_option (exists_var (local_state.variables ls) x); let ls = local_state.variables_update (update_var x v) ls; Some (ls,gs) }" | "la_ex (ls,gs) (AAssign_global c x e) = do { u \ eval_exp c (ls,gs); assert_option (bool_of_val u); v \ eval_exp e (ls,gs); assert_option (exists_var (global_state.variables gs) x); let gs = global_state.variables_update (update_var x v) gs; Some (ls,gs) }" | "la_ex fs (ATest e) = do { v \ eval_exp e fs; assert_option (bool_of_val v); Some fs }" | "la_ex fs ASkip = Some fs" subsection "Scheduling" type_synonym local_config = "(cmd,local_state) local_config" type_synonym global_config = "(cmd,local_state,global_state) global_config" interpretation li: Gen_Scheduler cfg la_en la_ex . subsection \Initial State\ definition init_valuation :: "var_decl list \ valuation" where "init_valuation vd x \ if x \ set vd then Some 0 else None" lemma init_valuation_eq_Some[simp]: "init_valuation vd x = Some v \ x\set vd \ v=0" unfolding init_valuation_def by auto lemma init_valuation_eq_None[simp]: "init_valuation vd x = None \ x\set vd" unfolding init_valuation_def by auto definition init_pc :: "proc_decl \ local_config" where "init_pc pd \ \ local_config.command = proc_decl.body pd, local_config.state = \ local_state.variables = init_valuation (proc_decl.local_vars pd) \ \" definition init_gc :: "program \ global_config" where "init_gc prog \ \ global_config.processes = mset (map init_pc (program.processes prog)), global_config.state = \ global_state.variables = init_valuation (program.global_vars prog) \ \" subsection "Semantics Reference Point" interpretation li: Gen_Scheduler_linit cfg la_en la_ex "{init_gc prog}" global_config.state for prog . abbreviation "ref_is_run \ li.sa.is_run" abbreviation "ref_accept \ li.sa.accept" end diff --git a/thys/CakeML/generated/CakeML/SemanticPrimitives.thy b/thys/CakeML/generated/CakeML/SemanticPrimitives.thy --- a/thys/CakeML/generated/CakeML/SemanticPrimitives.thy +++ b/thys/CakeML/generated/CakeML/SemanticPrimitives.thy @@ -1,998 +1,998 @@ chapter \Generated by Lem from \semantics/semanticPrimitives.lem\.\ theory "SemanticPrimitives" imports Main "HOL-Library.Datatype_Records" "LEM.Lem_pervasives" "LEM.Lem_list_extra" "LEM.Lem_string" "Lib" "Namespace" "Ast" "Ffi" "FpSem" "LEM.Lem_string_extra" begin \ \\open import Pervasives\\ \ \\open import Lib\\ \ \\import List_extra\\ \ \\import String\\ \ \\import String_extra\\ \ \\open import Ast\\ \ \\open import Namespace\\ \ \\open import Ffi\\ \ \\open import FpSem\\ \ \\ The type that a constructor builds is either a named datatype or an exception. * For exceptions, we also keep the module that the exception was declared in. \\ datatype tid_or_exn = TypeId " (modN, typeN) id0 " | TypeExn " (modN, conN) id0 " \ \\val type_defs_to_new_tdecs : list modN -> type_def -> set tid_or_exn\\ definition type_defs_to_new_tdecs :: "(string)list \((tvarN)list*string*(conN*(t)list)list)list \(tid_or_exn)set " where " type_defs_to_new_tdecs mn tdefs = ( List.set (List.map ( \x . (case x of (tvs,tn,ctors) => TypeId (mk_id mn tn) )) tdefs))" datatype_record 'v sem_env = v ::" (modN, varN, 'v) namespace " c ::" (modN, conN, (nat * tid_or_exn)) namespace " \ \\ Value forms \\ datatype v = Litv " lit " \ \\ Constructor application. \\ | Conv " (conN * tid_or_exn)option " " v list " \ \\ Function closures The environment is used for the free variables in the function \\ | Closure " v sem_env " " varN " " exp0 " \ \\ Function closure for recursive functions * See Closure and Letrec above * The last variable name indicates which function from the mutually * recursive bundle this closure value represents \\ | Recclosure " v sem_env " " (varN * varN * exp0) list " " varN " | Loc " nat " | Vectorv " v list " type_synonym env_ctor =" (modN, conN, (nat * tid_or_exn)) namespace " type_synonym env_val =" (modN, varN, v) namespace " definition Bindv :: " v " where " Bindv = ( Conv (Some((''Bind''),TypeExn(Short(''Bind'')))) [])" \ \\ The result of evaluation \\ datatype abort = Rtype_error | Rtimeout_error datatype 'a error_result = Rraise " 'a " \ \\ Should only be a value of type exn \\ | Rabort " abort " datatype( 'a, 'b) result = Rval " 'a " | Rerr " 'b error_result " \ \\ Stores \\ datatype 'a store_v = \ \\ A ref cell \\ Refv " 'a " \ \\ A byte array \\ | W8array " 8 word list " \ \\ An array of values \\ | Varray " 'a list " \ \\val store_v_same_type : forall 'a. store_v 'a -> store_v 'a -> bool\\ definition store_v_same_type :: " 'a store_v \ 'a store_v \ bool " where " store_v_same_type v1 v2 = ( (case (v1,v2) of (Refv _, Refv _) => True | (W8array _,W8array _) => True | (Varray _,Varray _) => True | _ => False ))" \ \\ The nth item in the list is the value at location n \\ type_synonym 'a store =" ( 'a store_v) list " \ \\val empty_store : forall 'a. store 'a\\ definition empty_store :: "('a store_v)list " where " empty_store = ( [])" \ \\val store_lookup : forall 'a. nat -> store 'a -> maybe (store_v 'a)\\ definition store_lookup :: " nat \('a store_v)list \('a store_v)option " where " store_lookup l st = ( if l < List.length st then Some (List.nth st l) else None )" \ \\val store_alloc : forall 'a. store_v 'a -> store 'a -> store 'a * nat\\ definition store_alloc :: " 'a store_v \('a store_v)list \('a store_v)list*nat " where " store_alloc v2 st = ( ((st @ [v2]), List.length st))" \ \\val store_assign : forall 'a. nat -> store_v 'a -> store 'a -> maybe (store 'a)\\ definition store_assign :: " nat \ 'a store_v \('a store_v)list \(('a store_v)list)option " where " store_assign n v2 st = ( if (n < List.length st) \ store_v_same_type (List.nth st n) v2 then Some (List.list_update st n v2) else None )" datatype_record 'ffi state = clock ::" nat " refs ::" v store " ffi ::" 'ffi ffi_state " defined_types ::" tid_or_exn set " defined_mods ::" ( modN list) set " \ \\ Other primitives \\ \ \\ Check that a constructor is properly applied \\ \ \\val do_con_check : env_ctor -> maybe (id modN conN) -> nat -> bool\\ fun do_con_check :: "((string),(string),(nat*tid_or_exn))namespace \(((string),(string))id0)option \ nat \ bool " where " do_con_check cenv None l = ( True )" |" do_con_check cenv (Some n) l = ( (case nsLookup cenv n of None => False | Some (l',ns) => l = l' ))" \ \\val build_conv : env_ctor -> maybe (id modN conN) -> list v -> maybe v\\ fun build_conv :: "((string),(string),(nat*tid_or_exn))namespace \(((string),(string))id0)option \(v)list \(v)option " where " build_conv envC None vs = ( Some (Conv None vs))" |" build_conv envC (Some id1) vs = ( (case nsLookup envC id1 of None => None | Some (len,t1) => Some (Conv (Some (id_to_n id1, t1)) vs) ))" \ \\val lit_same_type : lit -> lit -> bool\\ definition lit_same_type :: " lit \ lit \ bool " where " lit_same_type l1 l2 = ( (case (l1,l2) of (IntLit _, IntLit _) => True | (Char _, Char _) => True | (StrLit _, StrLit _) => True | (Word8 _, Word8 _) => True | (Word64 _, Word64 _) => True | _ => False ))" datatype 'a match_result = No_match | Match_type_error | Match " 'a " \ \\val same_tid : tid_or_exn -> tid_or_exn -> bool\\ fun same_tid :: " tid_or_exn \ tid_or_exn \ bool " where " same_tid (TypeId tn1) (TypeId tn2) = ( tn1 = tn2 )" |" same_tid (TypeExn _) (TypeExn _) = ( True )" |" same_tid _ _ = ( False )" \ \\val same_ctor : conN * tid_or_exn -> conN * tid_or_exn -> bool\\ fun same_ctor :: " string*tid_or_exn \ string*tid_or_exn \ bool " where " same_ctor (cn1, TypeExn mn1) (cn2, TypeExn mn2) = ( (cn1 = cn2) \ (mn1 = mn2))" |" same_ctor (cn1, _) (cn2, _) = ( cn1 = cn2 )" \ \\val ctor_same_type : maybe (conN * tid_or_exn) -> maybe (conN * tid_or_exn) -> bool\\ definition ctor_same_type :: "(string*tid_or_exn)option \(string*tid_or_exn)option \ bool " where " ctor_same_type c1 c2 = ( (case (c1,c2) of (None, None) => True | (Some (_,t1), Some (_,t2)) => same_tid t1 t2 | _ => False ))" \ \\ A big-step pattern matcher. If the value matches the pattern, return an * environment with the pattern variables bound to the corresponding sub-terms * of the value; this environment extends the environment given as an argument. * No_match is returned when there is no match, but any constructors * encountered in determining the match failure are applied to the correct * number of arguments, and constructors in corresponding positions in the * pattern and value come from the same type. Match_type_error is returned * when one of these conditions is violated \\ \ \\val pmatch : env_ctor -> store v -> pat -> v -> alist varN v -> match_result (alist varN v)\\ function (sequential,domintros) pmatch_list :: "((string),(string),(nat*tid_or_exn))namespace \((v)store_v)list \(pat)list \(v)list \(string*v)list \((string*v)list)match_result " and pmatch :: "((string),(string),(nat*tid_or_exn))namespace \((v)store_v)list \ pat \ v \(string*v)list \((string*v)list)match_result " where " pmatch envC s Pany v' env = ( Match env )" |" pmatch envC s (Pvar x) v' env = ( Match ((x,v')# env))" |" pmatch envC s (Plit l) (Litv l') env = ( if l = l' then Match env else if lit_same_type l l' then No_match else Match_type_error )" |" pmatch envC s (Pcon (Some n) ps) (Conv (Some (n', t')) vs) env = ( (case nsLookup envC n of Some (l, t1) => if same_tid t1 t' \ (List.length ps = l) then if same_ctor (id_to_n n, t1) (n',t') then if List.length vs = l then pmatch_list envC s ps vs env else Match_type_error else No_match else Match_type_error | _ => Match_type_error ))" |" pmatch envC s (Pcon None ps) (Conv None vs) env = ( if List.length ps = List.length vs then pmatch_list envC s ps vs env else Match_type_error )" |" pmatch envC s (Pref p) (Loc lnum) env = ( (case store_lookup lnum s of Some (Refv v2) => pmatch envC s p v2 env | Some _ => Match_type_error | None => Match_type_error ))" |" pmatch envC s (Ptannot p t1) v2 env = ( pmatch envC s p v2 env )" |" pmatch envC _ _ _ env = ( Match_type_error )" |" pmatch_list envC s [] [] env = ( Match env )" |" pmatch_list envC s (p # ps) (v2 # vs) env = ( (case pmatch envC s p v2 env of No_match => No_match | Match_type_error => Match_type_error | Match env' => pmatch_list envC s ps vs env' ))" |" pmatch_list envC s _ _ env = ( Match_type_error )" by pat_completeness auto \ \\ Bind each function of a mutually recursive set of functions to its closure \\ \ \\val build_rec_env : list (varN * varN * exp) -> sem_env v -> env_val -> env_val\\ definition build_rec_env :: "(varN*varN*exp0)list \(v)sem_env \((string),(string),(v))namespace \((string),(string),(v))namespace " where " build_rec_env funs cl_env add_to_env = ( List.foldr ( \x . (case x of (f,x,e) => \ env' . nsBind f (Recclosure cl_env funs f) env' )) funs add_to_env )" \ \\ Lookup in the list of mutually recursive functions \\ \ \\val find_recfun : forall 'a 'b. varN -> list (varN * 'a * 'b) -> maybe ('a * 'b)\\ fun find_recfun :: " string \(string*'a*'b)list \('a*'b)option " where " find_recfun n ([]) = ( None )" |" find_recfun n ((f,x,e) # funs) = ( if f = n then Some (x,e) else find_recfun n funs )" datatype eq_result = Eq_val " bool " | Eq_type_error \ \\val do_eq : v -> v -> eq_result\\ function (sequential,domintros) do_eq_list :: "(v)list \(v)list \ eq_result " and do_eq :: " v \ v \ eq_result " where " do_eq (Litv l1) (Litv l2) = ( if lit_same_type l1 l2 then Eq_val (l1 = l2) else Eq_type_error )" |" do_eq (Loc l1) (Loc l2) = ( Eq_val (l1 = l2))" |" do_eq (Conv cn1 vs1) (Conv cn2 vs2) = ( if (cn1 = cn2) \ (List.length vs1 = List.length vs2) then do_eq_list vs1 vs2 else if ctor_same_type cn1 cn2 then Eq_val False else Eq_type_error )" |" do_eq (Vectorv vs1) (Vectorv vs2) = ( if List.length vs1 = List.length vs2 then do_eq_list vs1 vs2 else Eq_val False )" |" do_eq (Closure _ _ _) (Closure _ _ _) = ( Eq_val True )" |" do_eq (Closure _ _ _) (Recclosure _ _ _) = ( Eq_val True )" |" do_eq (Recclosure _ _ _) (Closure _ _ _) = ( Eq_val True )" |" do_eq (Recclosure _ _ _) (Recclosure _ _ _) = ( Eq_val True )" |" do_eq _ _ = ( Eq_type_error )" |" do_eq_list [] [] = ( Eq_val True )" |" do_eq_list (v1 # vs1) (v2 # vs2) = ( (case do_eq v1 v2 of Eq_type_error => Eq_type_error | Eq_val r => if \ r then Eq_val False else do_eq_list vs1 vs2 ))" |" do_eq_list _ _ = ( Eq_val False )" by pat_completeness auto \ \\val prim_exn : conN -> v\\ definition prim_exn :: " string \ v " where " prim_exn cn = ( Conv (Some (cn, TypeExn (Short cn))) [])" \ \\ Do an application \\ \ \\val do_opapp : list v -> maybe (sem_env v * exp)\\ fun do_opapp :: "(v)list \((v)sem_env*exp0)option " where " do_opapp ([Closure env n e, v2]) = ( Some (( env (| v := (nsBind n v2(v env)) |)), e))" |" do_opapp ([Recclosure env funs n, v2]) = ( if allDistinct (List.map ( \x . (case x of (f,x,e) => f )) funs) then (case find_recfun n funs of Some (n,e) => Some (( env (| v := (nsBind n v2 (build_rec_env funs env(v env))) |)), e) | None => None ) else None )" |" do_opapp _ = ( None )" \ \\ If a value represents a list, get that list. Otherwise return Nothing \\ \ \\val v_to_list : v -> maybe (list v)\\ function (sequential,domintros) v_to_list :: " v \((v)list)option " where " v_to_list (Conv (Some (cn, TypeId (Short tn))) []) = ( if (cn = (''nil'')) \ (tn = (''list'')) then Some [] else None )" |" v_to_list (Conv (Some (cn,TypeId (Short tn))) [v1,v2]) = ( if (cn = (''::'')) \ (tn = (''list'')) then (case v_to_list v2 of Some vs => Some (v1 # vs) | None => None ) else None )" |" v_to_list _ = ( None )" by pat_completeness auto \ \\val v_to_char_list : v -> maybe (list char)\\ function (sequential,domintros) v_to_char_list :: " v \((char)list)option " where " v_to_char_list (Conv (Some (cn, TypeId (Short tn))) []) = ( if (cn = (''nil'')) \ (tn = (''list'')) then Some [] else None )" |" v_to_char_list (Conv (Some (cn,TypeId (Short tn))) [Litv (Char c2),v2]) = ( if (cn = (''::'')) \ (tn = (''list'')) then (case v_to_char_list v2 of Some cs => Some (c2 # cs) | None => None ) else None )" |" v_to_char_list _ = ( None )" by pat_completeness auto \ \\val vs_to_string : list v -> maybe string\\ function (sequential,domintros) vs_to_string :: "(v)list \(string)option " where " vs_to_string [] = ( Some (''''))" |" vs_to_string (Litv(StrLit s1)# vs) = ( (case vs_to_string vs of Some s2 => Some (s1 @ s2) | _ => None ))" |" vs_to_string _ = ( None )" by pat_completeness auto \ \\val copy_array : forall 'a. list 'a * integer -> integer -> maybe (list 'a * integer) -> maybe (list 'a)\\ fun copy_array :: " 'a list*int \ int \('a list*int)option \('a list)option " where " copy_array (src,srcoff) len d = ( if (srcoff <( 0 :: int)) \ ((len <( 0 :: int)) \ (List.length src < nat (abs ( (srcoff + len))))) then None else (let copied = (List.take (nat (abs ( len))) (List.drop (nat (abs ( srcoff))) src)) in (case d of Some (dst,dstoff) => if (dstoff <( 0 :: int)) \ (List.length dst < nat (abs ( (dstoff + len)))) then None else Some ((List.take (nat (abs ( dstoff))) dst @ copied) @ List.drop (nat (abs ( (dstoff + len)))) dst) | None => Some copied )))" \ \\val ws_to_chars : list word8 -> list char\\ definition ws_to_chars :: "(8 word)list \(char)list " where " ws_to_chars ws = ( List.map (\ w . (%n. char_of (n::nat))(unat w)) ws )" \ \\val chars_to_ws : list char -> list word8\\ definition chars_to_ws :: "(char)list \(8 word)list " where " chars_to_ws cs = ( List.map (\ c2 . word_of_int(int(of_char c2))) cs )" \ \\val opn_lookup : opn -> integer -> integer -> integer\\ fun opn_lookup :: " opn \ int \ int \ int " where " opn_lookup Plus = ( (+))" |" opn_lookup Minus = ( (-))" |" opn_lookup Times = ( (*))" |" opn_lookup Divide = ( (div))" |" opn_lookup Modulo = ( (mod))" \ \\val opb_lookup : opb -> integer -> integer -> bool\\ fun opb_lookup :: " opb \ int \ int \ bool " where " opb_lookup Lt = ( (<))" |" opb_lookup Gt = ( (>))" |" opb_lookup Leq = ( (\))" |" opb_lookup Geq = ( (\))" \ \\val opw8_lookup : opw -> word8 -> word8 -> word8\\ fun opw8_lookup :: " opw \ 8 word \ 8 word \ 8 word " where - " opw8_lookup Andw = ( (AND) )" -|" opw8_lookup Orw = ( (OR) )" -|" opw8_lookup Xor = ( (XOR) )" + " opw8_lookup Andw = ( Bit_Operations.and )" +|" opw8_lookup Orw = ( Bit_Operations.or )" +|" opw8_lookup Xor = ( Bit_Operations.xor )" |" opw8_lookup Add = ( Groups.plus )" |" opw8_lookup Sub = ( Groups.minus )" \ \\val opw64_lookup : opw -> word64 -> word64 -> word64\\ fun opw64_lookup :: " opw \ 64 word \ 64 word \ 64 word " where - " opw64_lookup Andw = ( (AND) )" -|" opw64_lookup Orw = ( (OR) )" -|" opw64_lookup Xor = ( (XOR) )" + " opw64_lookup Andw = ( Bit_Operations.and )" +|" opw64_lookup Orw = ( Bit_Operations.or )" +|" opw64_lookup Xor = ( Bit_Operations.xor )" |" opw64_lookup Add = ( Groups.plus )" |" opw64_lookup Sub = ( Groups.minus )" \ \\val shift8_lookup : shift -> word8 -> nat -> word8\\ fun shift8_lookup :: " shift \ 8 word \ nat \ 8 word " where " shift8_lookup Lsl = ( shiftl )" |" shift8_lookup Lsr = ( shiftr )" |" shift8_lookup Asr = ( sshiftr )" |" shift8_lookup Ror = ( (% a b. word_rotr b a) )" \ \\val shift64_lookup : shift -> word64 -> nat -> word64\\ fun shift64_lookup :: " shift \ 64 word \ nat \ 64 word " where " shift64_lookup Lsl = ( shiftl )" |" shift64_lookup Lsr = ( shiftr )" |" shift64_lookup Asr = ( sshiftr )" |" shift64_lookup Ror = ( (% a b. word_rotr b a) )" \ \\val Boolv : bool -> v\\ definition Boolv :: " bool \ v " where " Boolv b = ( if b then Conv (Some ((''true''), TypeId (Short (''bool'')))) [] else Conv (Some ((''false''), TypeId (Short (''bool'')))) [])" datatype exp_or_val = Exp " exp0 " | Val " v " type_synonym( 'ffi, 'v) store_ffi =" 'v store * 'ffi ffi_state " \ \\val do_app : forall 'ffi. store_ffi 'ffi v -> op -> list v -> maybe (store_ffi 'ffi v * result v v)\\ fun do_app :: "((v)store_v)list*'ffi ffi_state \ op0 \(v)list \((((v)store_v)list*'ffi ffi_state)*((v),(v))result)option " where " do_app ((s:: v store),(t1:: 'ffi ffi_state)) op1 vs = ( (case (op1, vs) of (Opn op1, [Litv (IntLit n1), Litv (IntLit n2)]) => if ((op1 = Divide) \ (op1 = Modulo)) \ (n2 =( 0 :: int)) then Some ((s,t1), Rerr (Rraise (prim_exn (''Div'')))) else Some ((s,t1), Rval (Litv (IntLit (opn_lookup op1 n1 n2)))) | (Opb op1, [Litv (IntLit n1), Litv (IntLit n2)]) => Some ((s,t1), Rval (Boolv (opb_lookup op1 n1 n2))) | (Opw W8 op1, [Litv (Word8 w1), Litv (Word8 w2)]) => Some ((s,t1), Rval (Litv (Word8 (opw8_lookup op1 w1 w2)))) | (Opw W64 op1, [Litv (Word64 w1), Litv (Word64 w2)]) => Some ((s,t1), Rval (Litv (Word64 (opw64_lookup op1 w1 w2)))) | (FP_bop bop, [Litv (Word64 w1), Litv (Word64 w2)]) => Some ((s,t1),Rval (Litv (Word64 (fp_bop bop w1 w2)))) | (FP_uop uop, [Litv (Word64 w)]) => Some ((s,t1),Rval (Litv (Word64 (fp_uop uop w)))) | (FP_cmp cmp, [Litv (Word64 w1), Litv (Word64 w2)]) => Some ((s,t1),Rval (Boolv (fp_cmp cmp w1 w2))) | (Shift W8 op1 n, [Litv (Word8 w)]) => Some ((s,t1), Rval (Litv (Word8 (shift8_lookup op1 w n)))) | (Shift W64 op1 n, [Litv (Word64 w)]) => Some ((s,t1), Rval (Litv (Word64 (shift64_lookup op1 w n)))) | (Equality, [v1, v2]) => (case do_eq v1 v2 of Eq_type_error => None | Eq_val b => Some ((s,t1), Rval (Boolv b)) ) | (Opassign, [Loc lnum, v2]) => (case store_assign lnum (Refv v2) s of Some s' => Some ((s',t1), Rval (Conv None [])) | None => None ) | (Opref, [v2]) => (let (s',n) = (store_alloc (Refv v2) s) in Some ((s',t1), Rval (Loc n))) | (Opderef, [Loc n]) => (case store_lookup n s of Some (Refv v2) => Some ((s,t1),Rval v2) | _ => None ) | (Aw8alloc, [Litv (IntLit n), Litv (Word8 w)]) => if n <( 0 :: int) then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else (let (s',lnum) = (store_alloc (W8array (List.replicate (nat (abs ( n))) w)) s) in Some ((s',t1), Rval (Loc lnum))) | (Aw8sub, [Loc lnum, Litv (IntLit i)]) => (case store_lookup lnum s of Some (W8array ws) => if i <( 0 :: int) then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else (let n = (nat (abs ( i))) in if n \ List.length ws then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else Some ((s,t1), Rval (Litv (Word8 (List.nth ws n))))) | _ => None ) | (Aw8length, [Loc n]) => (case store_lookup n s of Some (W8array ws) => Some ((s,t1),Rval (Litv(IntLit(int(List.length ws))))) | _ => None ) | (Aw8update, [Loc lnum, Litv(IntLit i), Litv(Word8 w)]) => (case store_lookup lnum s of Some (W8array ws) => if i <( 0 :: int) then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else (let n = (nat (abs ( i))) in if n \ List.length ws then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else (case store_assign lnum (W8array (List.list_update ws n w)) s of None => None | Some s' => Some ((s',t1), Rval (Conv None [])) )) | _ => None ) | (WordFromInt W8, [Litv(IntLit i)]) => Some ((s,t1), Rval (Litv (Word8 (word_of_int i)))) | (WordFromInt W64, [Litv(IntLit i)]) => Some ((s,t1), Rval (Litv (Word64 (word_of_int i)))) | (WordToInt W8, [Litv (Word8 w)]) => Some ((s,t1), Rval (Litv (IntLit (int(unat w))))) | (WordToInt W64, [Litv (Word64 w)]) => Some ((s,t1), Rval (Litv (IntLit (int(unat w))))) | (CopyStrStr, [Litv(StrLit str),Litv(IntLit off),Litv(IntLit len)]) => Some ((s,t1), (case copy_array ( str,off) len None of None => Rerr (Rraise (prim_exn (''Subscript''))) | Some cs => Rval (Litv(StrLit((cs)))) )) | (CopyStrAw8, [Litv(StrLit str),Litv(IntLit off),Litv(IntLit len), Loc dst,Litv(IntLit dstoff)]) => (case store_lookup dst s of Some (W8array ws) => (case copy_array ( str,off) len (Some(ws_to_chars ws,dstoff)) of None => Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) | Some cs => (case store_assign dst (W8array (chars_to_ws cs)) s of Some s' => Some ((s',t1), Rval (Conv None [])) | _ => None ) ) | _ => None ) | (CopyAw8Str, [Loc src,Litv(IntLit off),Litv(IntLit len)]) => (case store_lookup src s of Some (W8array ws) => Some ((s,t1), (case copy_array (ws,off) len None of None => Rerr (Rraise (prim_exn (''Subscript''))) | Some ws => Rval (Litv(StrLit((ws_to_chars ws)))) )) | _ => None ) | (CopyAw8Aw8, [Loc src,Litv(IntLit off),Litv(IntLit len), Loc dst,Litv(IntLit dstoff)]) => (case (store_lookup src s, store_lookup dst s) of (Some (W8array ws), Some (W8array ds)) => (case copy_array (ws,off) len (Some(ds,dstoff)) of None => Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) | Some ws => (case store_assign dst (W8array ws) s of Some s' => Some ((s',t1), Rval (Conv None [])) | _ => None ) ) | _ => None ) | (Ord, [Litv (Char c2)]) => Some ((s,t1), Rval (Litv(IntLit(int(of_char c2))))) | (Chr, [Litv (IntLit i)]) => Some ((s,t1), (if (i <( 0 :: int)) \ (i >( 255 :: int)) then Rerr (Rraise (prim_exn (''Chr''))) else Rval (Litv(Char((%n. char_of (n::nat))(nat (abs ( i)))))))) | (Chopb op1, [Litv (Char c1), Litv (Char c2)]) => Some ((s,t1), Rval (Boolv (opb_lookup op1 (int(of_char c1)) (int(of_char c2))))) | (Implode, [v2]) => (case v_to_char_list v2 of Some ls => Some ((s,t1), Rval (Litv (StrLit ( ls)))) | None => None ) | (Strsub, [Litv (StrLit str), Litv (IntLit i)]) => if i <( 0 :: int) then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else (let n = (nat (abs ( i))) in if n \ List.length str then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else Some ((s,t1), Rval (Litv (Char (List.nth ( str) n))))) | (Strlen, [Litv (StrLit str)]) => Some ((s,t1), Rval (Litv(IntLit(int(List.length str))))) | (Strcat, [v2]) => (case v_to_list v2 of Some vs => (case vs_to_string vs of Some str => Some ((s,t1), Rval (Litv(StrLit str))) | _ => None ) | _ => None ) | (VfromList, [v2]) => (case v_to_list v2 of Some vs => Some ((s,t1), Rval (Vectorv vs)) | None => None ) | (Vsub, [Vectorv vs, Litv (IntLit i)]) => if i <( 0 :: int) then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else (let n = (nat (abs ( i))) in if n \ List.length vs then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else Some ((s,t1), Rval (List.nth vs n))) | (Vlength, [Vectorv vs]) => Some ((s,t1), Rval (Litv (IntLit (int (List.length vs))))) | (Aalloc, [Litv (IntLit n), v2]) => if n <( 0 :: int) then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else (let (s',lnum) = (store_alloc (Varray (List.replicate (nat (abs ( n))) v2)) s) in Some ((s',t1), Rval (Loc lnum))) | (AallocEmpty, [Conv None []]) => (let (s',lnum) = (store_alloc (Varray []) s) in Some ((s',t1), Rval (Loc lnum))) | (Asub, [Loc lnum, Litv (IntLit i)]) => (case store_lookup lnum s of Some (Varray vs) => if i <( 0 :: int) then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else (let n = (nat (abs ( i))) in if n \ List.length vs then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else Some ((s,t1), Rval (List.nth vs n))) | _ => None ) | (Alength, [Loc n]) => (case store_lookup n s of Some (Varray ws) => Some ((s,t1),Rval (Litv(IntLit(int(List.length ws))))) | _ => None ) | (Aupdate, [Loc lnum, Litv (IntLit i), v2]) => (case store_lookup lnum s of Some (Varray vs) => if i <( 0 :: int) then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else (let n = (nat (abs ( i))) in if n \ List.length vs then Some ((s,t1), Rerr (Rraise (prim_exn (''Subscript'')))) else (case store_assign lnum (Varray (List.list_update vs n v2)) s of None => None | Some s' => Some ((s',t1), Rval (Conv None [])) )) | _ => None ) | (ConfigGC, [Litv (IntLit i), Litv (IntLit j)]) => Some ((s,t1), Rval (Conv None [])) | (FFI n, [Litv(StrLit conf), Loc lnum]) => (case store_lookup lnum s of Some (W8array ws) => (case call_FFI t1 n (List.map (\ c2 . of_nat(of_char c2)) ( conf)) ws of (t', ws') => (case store_assign lnum (W8array ws') s of Some s' => Some ((s', t'), Rval (Conv None [])) | None => None ) ) | _ => None ) | _ => None ))" \ \\ Do a logical operation \\ \ \\val do_log : lop -> v -> exp -> maybe exp_or_val\\ fun do_log :: " lop \ v \ exp0 \(exp_or_val)option " where " do_log And v2 e = ( (case v2 of Litv _ => None | Conv m l2 => (case m of None => None | Some p => (case p of (s1,t1) => if(s1 = (''true'')) then ((case t1 of TypeId i => (case i of Short s2 => if(s2 = (''bool'')) then ((case l2 of [] => Some (Exp e) | _ => None )) else None | Long _ _ => None ) | TypeExn _ => None )) else ( if(s1 = (''false'')) then ((case t1 of TypeId i2 => (case i2 of Short s4 => if(s4 = (''bool'')) then ((case l2 of [] => Some (Val v2) | _ => None )) else None | Long _ _ => None ) | TypeExn _ => None )) else None) ) ) | Closure _ _ _ => None | Recclosure _ _ _ => None | Loc _ => None | Vectorv _ => None ) )" |" do_log Or v2 e = ( (case v2 of Litv _ => None | Conv m0 l6 => (case m0 of None => None | Some p0 => (case p0 of (s8,t0) => if(s8 = (''false'')) then ((case t0 of TypeId i5 => (case i5 of Short s9 => if(s9 = (''bool'')) then ((case l6 of [] => Some (Exp e) | _ => None )) else None | Long _ _ => None ) | TypeExn _ => None )) else ( if(s8 = (''true'')) then ((case t0 of TypeId i8 => (case i8 of Short s11 => if(s11 = (''bool'')) then ((case l6 of [] => Some (Val v2) | _ => None )) else None | Long _ _ => None ) | TypeExn _ => None )) else None) ) ) | Closure _ _ _ => None | Recclosure _ _ _ => None | Loc _ => None | Vectorv _ => None ) )" \ \\ Do an if-then-else \\ \ \\val do_if : v -> exp -> exp -> maybe exp\\ definition do_if :: " v \ exp0 \ exp0 \(exp0)option " where " do_if v2 e1 e2 = ( if v2 = (Boolv True) then Some e1 else if v2 = (Boolv False) then Some e2 else None )" \ \\ Semantic helpers for definitions \\ \ \\ Build a constructor environment for the type definition tds \\ \ \\val build_tdefs : list modN -> list (list tvarN * typeN * list (conN * list t)) -> env_ctor\\ definition build_tdefs :: "(string)list \((tvarN)list*string*(string*(t)list)list)list \((string),(string),(nat*tid_or_exn))namespace " where " build_tdefs mn tds = ( alist_to_ns (List.rev (List.concat (List.map ( \x . (case x of (tvs, tn, condefs) => List.map ( \x . (case x of (conN, ts) => (conN, (List.length ts, TypeId (mk_id mn tn))) )) condefs )) tds))))" \ \\ Checks that no constructor is defined twice in a type \\ \ \\val check_dup_ctors : list (list tvarN * typeN * list (conN * list t)) -> bool\\ definition check_dup_ctors :: "((tvarN)list*string*(string*(t)list)list)list \ bool " where " check_dup_ctors tds = ( Lem_list.allDistinct ((let x2 = ([]) in List.foldr (\x . (case x of (tvs, tn, condefs) => \ x2 . List.foldr (\x . (case x of (n, ts) => \ x2 . if True then n # x2 else x2 )) condefs x2 )) tds x2)))" \ \\val combine_dec_result : forall 'a. sem_env v -> result (sem_env v) 'a -> result (sem_env v) 'a\\ fun combine_dec_result :: "(v)sem_env \(((v)sem_env),'a)result \(((v)sem_env),'a)result " where " combine_dec_result env (Rerr e) = ( Rerr e )" |" combine_dec_result env (Rval env') = ( Rval (| v = (nsAppend(v env')(v env)), c = (nsAppend(c env')(c env)) |) )" \ \\val extend_dec_env : sem_env v -> sem_env v -> sem_env v\\ definition extend_dec_env :: "(v)sem_env \(v)sem_env \(v)sem_env " where " extend_dec_env new_env env = ( (| v = (nsAppend(v new_env)(v env)), c = (nsAppend(c new_env)(c env)) |) )" \ \\val decs_to_types : list dec -> list typeN\\ definition decs_to_types :: "(dec)list \(string)list " where " decs_to_types ds = ( List.concat (List.map (\ d . (case d of Dtype locs tds => List.map ( \x . (case x of (tvs,tn,ctors) => tn )) tds | _ => [] )) ds))" \ \\val no_dup_types : list dec -> bool\\ definition no_dup_types :: "(dec)list \ bool " where " no_dup_types ds = ( Lem_list.allDistinct (decs_to_types ds))" \ \\val prog_to_mods : list top -> list (list modN)\\ definition prog_to_mods :: "(top0)list \((string)list)list " where " prog_to_mods tops = ( List.concat (List.map (\ top1 . (case top1 of Tmod mn _ _ => [[mn]] | _ => [] )) tops))" \ \\val no_dup_mods : list top -> set (list modN) -> bool\\ definition no_dup_mods :: "(top0)list \((modN)list)set \ bool " where " no_dup_mods tops defined_mods2 = ( Lem_list.allDistinct (prog_to_mods tops) \ disjnt (List.set (prog_to_mods tops)) defined_mods2 )" \ \\val prog_to_top_types : list top -> list typeN\\ definition prog_to_top_types :: "(top0)list \(string)list " where " prog_to_top_types tops = ( List.concat (List.map (\ top1 . (case top1 of Tdec d => decs_to_types [d] | _ => [] )) tops))" \ \\val no_dup_top_types : list top -> set tid_or_exn -> bool\\ definition no_dup_top_types :: "(top0)list \(tid_or_exn)set \ bool " where " no_dup_top_types tops defined_types2 = ( Lem_list.allDistinct (prog_to_top_types tops) \ disjnt (List.set (List.map (\ tn . TypeId (Short tn)) (prog_to_top_types tops))) defined_types2 )" end diff --git a/thys/CakeML/generated/Lem_word.thy b/thys/CakeML/generated/Lem_word.thy --- a/thys/CakeML/generated/Lem_word.thy +++ b/thys/CakeML/generated/Lem_word.thy @@ -1,1024 +1,1024 @@ chapter \Generated by Lem from \word.lem\.\ theory "Lem_word" imports Main "Lem_bool" "Lem_maybe" "Lem_num" "Lem_basic_classes" "Lem_list" "HOL-Library.Word" begin \ \\open import Bool Maybe Num Basic_classes List\\ \ \\open import {isabelle} `HOL-Library.Word`\\ \ \\open import {hol} `wordsTheory` `wordsLib`\\ \ \\ ========================================================================== \\ \ \\ Define general purpose word, i.e. sequences of bits of arbitrary length \\ \ \\ ========================================================================== \\ datatype bitSequence = BitSeq " nat option " " \ \\ length of the sequence, Nothing means infinite length \\ bool " " bool \ \\ sign of the word, used to fill up after concrete value is exhausted \\ list " \ \\ the initial part of the sequence, least significant bit first \\ \ \\val bitSeqEq : bitSequence -> bitSequence -> bool\\ \ \\val boolListFrombitSeq : nat -> bitSequence -> list bool\\ fun boolListFrombitSeqAux :: " nat \ 'a \ 'a list \ 'a list " where " boolListFrombitSeqAux n s bl = ( if n =( 0 :: nat) then [] else (case bl of [] => List.replicate n s | b # bl' => b # (boolListFrombitSeqAux (n-( 1 :: nat)) s bl') ))" fun boolListFrombitSeq :: " nat \ bitSequence \(bool)list " where " boolListFrombitSeq n (BitSeq _ s bl) = ( boolListFrombitSeqAux n s bl )" \ \\val bitSeqFromBoolList : list bool -> maybe bitSequence\\ definition bitSeqFromBoolList :: "(bool)list \(bitSequence)option " where " bitSeqFromBoolList bl = ( (case dest_init bl of None => None | Some (bl', s) => Some (BitSeq (Some (List.length bl)) s bl') ))" \ \\ cleans up the representation of a bitSequence without changing its semantics \\ \ \\val cleanBitSeq : bitSequence -> bitSequence\\ fun cleanBitSeq :: " bitSequence \ bitSequence " where " cleanBitSeq (BitSeq len s bl) = ( (case len of None => (BitSeq len s (List.rev (dropWhile ((\) s) (List.rev bl)))) | Some n => (BitSeq len s (List.rev (dropWhile ((\) s) (List.rev (List.take (n-( 1 :: nat)) bl))))) ))" \ \\val bitSeqTestBit : bitSequence -> nat -> maybe bool\\ fun bitSeqTestBit :: " bitSequence \ nat \(bool)option " where " bitSeqTestBit (BitSeq None s bl) pos = ( if pos < List.length bl then index bl pos else Some s )" |" bitSeqTestBit (BitSeq(Some l) s bl) pos = ( if (pos \ l) then None else if ((pos = (l -( 1 :: nat))) \ (pos \ List.length bl)) then Some s else index bl pos )" \ \\val bitSeqSetBit : bitSequence -> nat -> bool -> bitSequence\\ fun bitSeqSetBit :: " bitSequence \ nat \ bool \ bitSequence " where " bitSeqSetBit (BitSeq len s bl) pos v = ( (let bl' = (if (pos < List.length bl) then bl else bl @ List.replicate pos s) in (let bl'' = (List.list_update bl' pos v) in (let bs' = (BitSeq len s bl'') in cleanBitSeq bs'))))" \ \\val resizeBitSeq : maybe nat -> bitSequence -> bitSequence\\ definition resizeBitSeq :: "(nat)option \ bitSequence \ bitSequence " where " resizeBitSeq new_len bs = ( (case cleanBitSeq bs of (BitSeq len s bl) => (let shorten_opt = ((case (new_len, len) of (None, _) => None | (Some l1, None) => Some l1 | (Some l1, Some l2) => if (l1 < l2) then Some l1 else None )) in (case shorten_opt of None => BitSeq new_len s bl | Some l1 => ( (let bl' = (List.take l1 (bl @ [s])) in (case dest_init bl' of None => (BitSeq len s bl) \ \\ do nothing if size 0 is requested \\ | Some (bl'', s') => cleanBitSeq (BitSeq new_len s' bl'') ))) )) ) )" \ \\val bitSeqNot : bitSequence -> bitSequence\\ fun bitSeqNot :: " bitSequence \ bitSequence " where " bitSeqNot (BitSeq len s bl) = ( BitSeq len (\ s) (List.map (\ x. \ x) bl))" \ \\val bitSeqBinop : (bool -> bool -> bool) -> bitSequence -> bitSequence -> bitSequence\\ \ \\val bitSeqBinopAux : (bool -> bool -> bool) -> bool -> list bool -> bool -> list bool -> list bool\\ fun bitSeqBinopAux :: "(bool \ bool \ bool)\ bool \(bool)list \ bool \(bool)list \(bool)list " where " bitSeqBinopAux binop s1 ([]) s2 ([]) = ( [])" |" bitSeqBinopAux binop s1 (b1 # bl1') s2 ([]) = ( (binop b1 s2) # bitSeqBinopAux binop s1 bl1' s2 [])" |" bitSeqBinopAux binop s1 ([]) s2 (b2 # bl2') = ( (binop s1 b2) # bitSeqBinopAux binop s1 [] s2 bl2' )" |" bitSeqBinopAux binop s1 (b1 # bl1') s2 (b2 # bl2') = ( (binop b1 b2) # bitSeqBinopAux binop s1 bl1' s2 bl2' )" definition bitSeqBinop :: "(bool \ bool \ bool)\ bitSequence \ bitSequence \ bitSequence " where " bitSeqBinop binop bs1 bs2 = ( ( (case cleanBitSeq bs1 of (BitSeq len1 s1 bl1) => (case cleanBitSeq bs2 of (BitSeq len2 s2 bl2) => (let len = ((case (len1, len2) of (Some l1, Some l2) => Some (max l1 l2) | _ => None )) in (let s = (binop s1 s2) in (let bl = (bitSeqBinopAux binop s1 bl1 s2 bl2) in cleanBitSeq (BitSeq len s bl)))) ) ) ))" definition bitSeqAnd :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqAnd = ( bitSeqBinop (\))" definition bitSeqOr :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqOr = ( bitSeqBinop (\))" definition bitSeqXor :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqXor = ( bitSeqBinop (\ b1 b2. \ (b1 \ b2)))" \ \\val bitSeqShiftLeft : bitSequence -> nat -> bitSequence\\ fun bitSeqShiftLeft :: " bitSequence \ nat \ bitSequence " where " bitSeqShiftLeft (BitSeq len s bl) n = ( cleanBitSeq (BitSeq len s (List.replicate n False @ bl)))" \ \\val bitSeqArithmeticShiftRight : bitSequence -> nat -> bitSequence\\ definition bitSeqArithmeticShiftRight :: " bitSequence \ nat \ bitSequence " where " bitSeqArithmeticShiftRight bs n = ( (case cleanBitSeq bs of (BitSeq len s bl) => cleanBitSeq (BitSeq len s (List.drop n bl)) ) )" \ \\val bitSeqLogicalShiftRight : bitSequence -> nat -> bitSequence\\ definition bitSeqLogicalShiftRight :: " bitSequence \ nat \ bitSequence " where " bitSeqLogicalShiftRight bs n = ( if (n =( 0 :: nat)) then cleanBitSeq bs else (case cleanBitSeq bs of (BitSeq len s bl) => (case len of None => cleanBitSeq (BitSeq len s (List.drop n bl)) | Some l => cleanBitSeq (BitSeq len False ((List.drop n bl) @ List.replicate l s)) ) ) )" \ \\ integerFromBoolList sign bl creates an integer from a list of bits (least significant bit first) and an explicitly given sign bit. It uses two's complement encoding. \\ \ \\val integerFromBoolList : (bool * list bool) -> integer\\ fun integerFromBoolListAux :: " int \(bool)list \ int " where " integerFromBoolListAux (acc1 :: int) (([]) :: bool list) = ( acc1 )" |" integerFromBoolListAux (acc1 :: int) ((True # bl') :: bool list) = ( integerFromBoolListAux ((acc1 *( 2 :: int)) +( 1 :: int)) bl' )" |" integerFromBoolListAux (acc1 :: int) ((False # bl') :: bool list) = ( integerFromBoolListAux (acc1 *( 2 :: int)) bl' )" fun integerFromBoolList :: " bool*(bool)list \ int " where " integerFromBoolList (sign, bl) = ( if sign then - (integerFromBoolListAux(( 0 :: int)) (List.rev (List.map (\ x. \ x) bl)) +( 1 :: int)) else integerFromBoolListAux(( 0 :: int)) (List.rev bl))" \ \\ [boolListFromInteger i] creates a sign bit and a list of booleans from an integer. The len_opt tells it when to stop.\\ \ \\val boolListFromInteger : integer -> bool * list bool\\ fun boolListFromNatural :: "(bool)list \ nat \(bool)list " where " boolListFromNatural acc1 (remainder :: nat) = ( if (remainder >( 0 :: nat)) then (boolListFromNatural (((remainder mod( 2 :: nat)) =( 1 :: nat)) # acc1) (remainder div( 2 :: nat))) else List.rev acc1 )" definition boolListFromInteger :: " int \ bool*(bool)list " where " boolListFromInteger (i :: int) = ( if (i <( 0 :: int)) then (True, List.map (\ x. \ x) (boolListFromNatural [] (nat (abs (- (i +( 1 :: int))))))) else (False, boolListFromNatural [] (nat (abs i))))" \ \\ [bitSeqFromInteger len_opt i] encodes [i] as a bitsequence with [len_opt] bits. If there are not enough bits, truncation happens \\ \ \\val bitSeqFromInteger : maybe nat -> integer -> bitSequence\\ definition bitSeqFromInteger :: "(nat)option \ int \ bitSequence " where " bitSeqFromInteger len_opt i = ( (let (s, bl) = (boolListFromInteger i) in resizeBitSeq len_opt (BitSeq None s bl)))" \ \\val integerFromBitSeq : bitSequence -> integer\\ definition integerFromBitSeq :: " bitSequence \ int " where " integerFromBitSeq bs = ( (case cleanBitSeq bs of (BitSeq len s bl) => integerFromBoolList (s, bl) ) )" \ \\ Now we can via translation to integers map arithmetic operations to bitSequences \\ \ \\val bitSeqArithUnaryOp : (integer -> integer) -> bitSequence -> bitSequence\\ definition bitSeqArithUnaryOp :: "(int \ int)\ bitSequence \ bitSequence " where " bitSeqArithUnaryOp uop bs = ( (case bs of (BitSeq len _ _) => bitSeqFromInteger len (uop (integerFromBitSeq bs)) ) )" \ \\val bitSeqArithBinOp : (integer -> integer -> integer) -> bitSequence -> bitSequence -> bitSequence\\ definition bitSeqArithBinOp :: "(int \ int \ int)\ bitSequence \ bitSequence \ bitSequence " where " bitSeqArithBinOp binop bs1 bs2 = ( (case bs1 of (BitSeq len1 _ _) => (case bs2 of (BitSeq len2 _ _) => (let len = ((case (len1, len2) of (Some l1, Some l2) => Some (max l1 l2) | _ => None )) in bitSeqFromInteger len (binop (integerFromBitSeq bs1) (integerFromBitSeq bs2))) ) ) )" \ \\val bitSeqArithBinTest : forall 'a. (integer -> integer -> 'a) -> bitSequence -> bitSequence -> 'a\\ definition bitSeqArithBinTest :: "(int \ int \ 'a)\ bitSequence \ bitSequence \ 'a " where " bitSeqArithBinTest binop bs1 bs2 = ( binop (integerFromBitSeq bs1) (integerFromBitSeq bs2))" \ \\ now instantiate the number interface for bit-sequences \\ \ \\val bitSeqFromNumeral : numeral -> bitSequence\\ \ \\val bitSeqLess : bitSequence -> bitSequence -> bool\\ definition bitSeqLess :: " bitSequence \ bitSequence \ bool " where " bitSeqLess bs1 bs2 = ( bitSeqArithBinTest (<) bs1 bs2 )" \ \\val bitSeqLessEqual : bitSequence -> bitSequence -> bool\\ definition bitSeqLessEqual :: " bitSequence \ bitSequence \ bool " where " bitSeqLessEqual bs1 bs2 = ( bitSeqArithBinTest (\) bs1 bs2 )" \ \\val bitSeqGreater : bitSequence -> bitSequence -> bool\\ definition bitSeqGreater :: " bitSequence \ bitSequence \ bool " where " bitSeqGreater bs1 bs2 = ( bitSeqArithBinTest (>) bs1 bs2 )" \ \\val bitSeqGreaterEqual : bitSequence -> bitSequence -> bool\\ definition bitSeqGreaterEqual :: " bitSequence \ bitSequence \ bool " where " bitSeqGreaterEqual bs1 bs2 = ( bitSeqArithBinTest (\) bs1 bs2 )" \ \\val bitSeqCompare : bitSequence -> bitSequence -> ordering\\ definition bitSeqCompare :: " bitSequence \ bitSequence \ ordering " where " bitSeqCompare bs1 bs2 = ( bitSeqArithBinTest (genericCompare (<) (=)) bs1 bs2 )" definition instance_Basic_classes_Ord_Word_bitSequence_dict :: "(bitSequence)Ord_class " where " instance_Basic_classes_Ord_Word_bitSequence_dict = ((| compare_method = bitSeqCompare, isLess_method = bitSeqLess, isLessEqual_method = bitSeqLessEqual, isGreater_method = bitSeqGreater, isGreaterEqual_method = bitSeqGreaterEqual |) )" \ \\ arithmetic negation, don't mix up with bitwise negation \\ \ \\val bitSeqNegate : bitSequence -> bitSequence\\ definition bitSeqNegate :: " bitSequence \ bitSequence " where " bitSeqNegate bs = ( bitSeqArithUnaryOp (\ i. - i) bs )" definition instance_Num_NumNegate_Word_bitSequence_dict :: "(bitSequence)NumNegate_class " where " instance_Num_NumNegate_Word_bitSequence_dict = ((| numNegate_method = bitSeqNegate |) )" \ \\val bitSeqAdd : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqAdd :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqAdd bs1 bs2 = ( bitSeqArithBinOp (+) bs1 bs2 )" definition instance_Num_NumAdd_Word_bitSequence_dict :: "(bitSequence)NumAdd_class " where " instance_Num_NumAdd_Word_bitSequence_dict = ((| numAdd_method = bitSeqAdd |) )" \ \\val bitSeqMinus : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMinus :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMinus bs1 bs2 = ( bitSeqArithBinOp (-) bs1 bs2 )" definition instance_Num_NumMinus_Word_bitSequence_dict :: "(bitSequence)NumMinus_class " where " instance_Num_NumMinus_Word_bitSequence_dict = ((| numMinus_method = bitSeqMinus |) )" \ \\val bitSeqSucc : bitSequence -> bitSequence\\ definition bitSeqSucc :: " bitSequence \ bitSequence " where " bitSeqSucc bs = ( bitSeqArithUnaryOp (\ n. n +( 1 :: int)) bs )" definition instance_Num_NumSucc_Word_bitSequence_dict :: "(bitSequence)NumSucc_class " where " instance_Num_NumSucc_Word_bitSequence_dict = ((| succ_method = bitSeqSucc |) )" \ \\val bitSeqPred : bitSequence -> bitSequence\\ definition bitSeqPred :: " bitSequence \ bitSequence " where " bitSeqPred bs = ( bitSeqArithUnaryOp (\ n. n -( 1 :: int)) bs )" definition instance_Num_NumPred_Word_bitSequence_dict :: "(bitSequence)NumPred_class " where " instance_Num_NumPred_Word_bitSequence_dict = ((| pred_method = bitSeqPred |) )" \ \\val bitSeqMult : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMult :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMult bs1 bs2 = ( bitSeqArithBinOp (*) bs1 bs2 )" definition instance_Num_NumMult_Word_bitSequence_dict :: "(bitSequence)NumMult_class " where " instance_Num_NumMult_Word_bitSequence_dict = ((| numMult_method = bitSeqMult |) )" \ \\val bitSeqPow : bitSequence -> nat -> bitSequence\\ definition bitSeqPow :: " bitSequence \ nat \ bitSequence " where " bitSeqPow bs n = ( bitSeqArithUnaryOp (\ i . i ^ n) bs )" definition instance_Num_NumPow_Word_bitSequence_dict :: "(bitSequence)NumPow_class " where " instance_Num_NumPow_Word_bitSequence_dict = ((| numPow_method = bitSeqPow |) )" \ \\val bitSeqDiv : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqDiv :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqDiv bs1 bs2 = ( bitSeqArithBinOp (div) bs1 bs2 )" definition instance_Num_NumIntegerDivision_Word_bitSequence_dict :: "(bitSequence)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_Word_bitSequence_dict = ((| div_method = bitSeqDiv |) )" definition instance_Num_NumDivision_Word_bitSequence_dict :: "(bitSequence)NumDivision_class " where " instance_Num_NumDivision_Word_bitSequence_dict = ((| numDivision_method = bitSeqDiv |) )" \ \\val bitSeqMod : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMod :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMod bs1 bs2 = ( bitSeqArithBinOp (mod) bs1 bs2 )" definition instance_Num_NumRemainder_Word_bitSequence_dict :: "(bitSequence)NumRemainder_class " where " instance_Num_NumRemainder_Word_bitSequence_dict = ((| mod_method = bitSeqMod |) )" \ \\val bitSeqMin : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMin :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMin bs1 bs2 = ( bitSeqArithBinOp min bs1 bs2 )" \ \\val bitSeqMax : bitSequence -> bitSequence -> bitSequence\\ definition bitSeqMax :: " bitSequence \ bitSequence \ bitSequence " where " bitSeqMax bs1 bs2 = ( bitSeqArithBinOp max bs1 bs2 )" definition instance_Basic_classes_OrdMaxMin_Word_bitSequence_dict :: "(bitSequence)OrdMaxMin_class " where " instance_Basic_classes_OrdMaxMin_Word_bitSequence_dict = ((| max_method = bitSeqMax, min_method = bitSeqMin |) )" \ \\ ========================================================================== \\ \ \\ Interface for bitoperations \\ \ \\ ========================================================================== \\ record 'a WordNot_class= lnot_method ::" 'a \ 'a " record 'a WordAnd_class= land_method ::" 'a \ 'a \ 'a " record 'a WordOr_class= lor_method ::" 'a \ 'a \ 'a " record 'a WordXor_class= lxor_method ::" 'a \ 'a \ 'a " record 'a WordLsl_class= lsl_method ::" 'a \ nat \ 'a " record 'a WordLsr_class= lsr_method ::" 'a \ nat \ 'a " record 'a WordAsr_class= asr_method ::" 'a \ nat \ 'a " \ \\ ----------------------- \\ \ \\ bitSequence \\ \ \\ ----------------------- \\ definition instance_Word_WordNot_Word_bitSequence_dict :: "(bitSequence)WordNot_class " where " instance_Word_WordNot_Word_bitSequence_dict = ((| lnot_method = bitSeqNot |) )" definition instance_Word_WordAnd_Word_bitSequence_dict :: "(bitSequence)WordAnd_class " where " instance_Word_WordAnd_Word_bitSequence_dict = ((| land_method = bitSeqAnd |) )" definition instance_Word_WordOr_Word_bitSequence_dict :: "(bitSequence)WordOr_class " where " instance_Word_WordOr_Word_bitSequence_dict = ((| lor_method = bitSeqOr |) )" definition instance_Word_WordXor_Word_bitSequence_dict :: "(bitSequence)WordXor_class " where " instance_Word_WordXor_Word_bitSequence_dict = ((| lxor_method = bitSeqXor |) )" definition instance_Word_WordLsl_Word_bitSequence_dict :: "(bitSequence)WordLsl_class " where " instance_Word_WordLsl_Word_bitSequence_dict = ((| lsl_method = bitSeqShiftLeft |) )" definition instance_Word_WordLsr_Word_bitSequence_dict :: "(bitSequence)WordLsr_class " where " instance_Word_WordLsr_Word_bitSequence_dict = ((| lsr_method = bitSeqLogicalShiftRight |) )" definition instance_Word_WordAsr_Word_bitSequence_dict :: "(bitSequence)WordAsr_class " where " instance_Word_WordAsr_Word_bitSequence_dict = ((| asr_method = bitSeqArithmeticShiftRight |) )" \ \\ ----------------------- \\ \ \\ int32 \\ \ \\ ----------------------- \\ \ \\val int32Lnot : int32 -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordNot_Num_int32_dict :: "( 32 word)WordNot_class " where " instance_Word_WordNot_Num_int32_dict = ((| - lnot_method = (\ w. (NOT w))|) )" + lnot_method = NOT|) )" \ \\val int32Lor : int32 -> int32 -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordOr_Num_int32_dict :: "( 32 word)WordOr_class " where " instance_Word_WordOr_Num_int32_dict = ((| - lor_method = (OR)|) )" + lor_method = Bit_Operations.or|) )" \ \\val int32Lxor : int32 -> int32 -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordXor_Num_int32_dict :: "( 32 word)WordXor_class " where " instance_Word_WordXor_Num_int32_dict = ((| - lxor_method = (XOR)|) )" + lxor_method = Bit_Operations.xor|) )" \ \\val int32Land : int32 -> int32 -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordAnd_Num_int32_dict :: "( 32 word)WordAnd_class " where " instance_Word_WordAnd_Num_int32_dict = ((| - land_method = (AND)|) )" + land_method = Bit_Operations.and|) )" \ \\val int32Lsl : int32 -> nat -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordLsl_Num_int32_dict :: "( 32 word)WordLsl_class " where " instance_Word_WordLsl_Num_int32_dict = ((| lsl_method = (\w n. push_bit n w)|) )" \ \\val int32Lsr : int32 -> nat -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordLsr_Num_int32_dict :: "( 32 word)WordLsr_class " where " instance_Word_WordLsr_Num_int32_dict = ((| lsr_method = (\w n. drop_bit n w)|) )" \ \\val int32Asr : int32 -> nat -> int32\\ \ \\ XXX: fix \\ definition instance_Word_WordAsr_Num_int32_dict :: "( 32 word)WordAsr_class " where " instance_Word_WordAsr_Num_int32_dict = ((| asr_method = (\w n. signed_drop_bit n w)|) )" \ \\ ----------------------- \\ \ \\ int64 \\ \ \\ ----------------------- \\ \ \\val int64Lnot : int64 -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordNot_Num_int64_dict :: "( 64 word)WordNot_class " where " instance_Word_WordNot_Num_int64_dict = ((| - lnot_method = (\ w. (NOT w))|) )" + lnot_method = NOT|) )" \ \\val int64Lor : int64 -> int64 -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordOr_Num_int64_dict :: "( 64 word)WordOr_class " where " instance_Word_WordOr_Num_int64_dict = ((| - lor_method = (OR)|) )" + lor_method = Bit_Operations.or|) )" \ \\val int64Lxor : int64 -> int64 -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordXor_Num_int64_dict :: "( 64 word)WordXor_class " where " instance_Word_WordXor_Num_int64_dict = ((| - lxor_method = (XOR)|) )" + lxor_method = Bit_Operations.xor|) )" \ \\val int64Land : int64 -> int64 -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordAnd_Num_int64_dict :: "( 64 word)WordAnd_class " where " instance_Word_WordAnd_Num_int64_dict = ((| - land_method = (AND)|) )" + land_method = Bit_Operations.and|) )" \ \\val int64Lsl : int64 -> nat -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordLsl_Num_int64_dict :: "( 64 word)WordLsl_class " where " instance_Word_WordLsl_Num_int64_dict = ((| lsl_method = (\w n. push_bit n w)|) )" \ \\val int64Lsr : int64 -> nat -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordLsr_Num_int64_dict :: "( 64 word)WordLsr_class " where " instance_Word_WordLsr_Num_int64_dict = ((| lsr_method = (\w n. drop_bit n w)|) )" \ \\val int64Asr : int64 -> nat -> int64\\ \ \\ XXX: fix \\ definition instance_Word_WordAsr_Num_int64_dict :: "( 64 word)WordAsr_class " where " instance_Word_WordAsr_Num_int64_dict = ((| asr_method = (\w n. signed_drop_bit n w)|) )" \ \\ ----------------------- \\ \ \\ Words via bit sequences \\ \ \\ ----------------------- \\ \ \\val defaultLnot : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> 'a\\ definition defaultLnot :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ 'a " where " defaultLnot fromBitSeq toBitSeq x = ( fromBitSeq (bitSeqNegate (toBitSeq x)))" \ \\val defaultLand : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> 'a -> 'a\\ definition defaultLand :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ 'a \ 'a " where " defaultLand fromBitSeq toBitSeq x1 x2 = ( fromBitSeq (bitSeqAnd (toBitSeq x1) (toBitSeq x2)))" \ \\val defaultLor : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> 'a -> 'a\\ definition defaultLor :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ 'a \ 'a " where " defaultLor fromBitSeq toBitSeq x1 x2 = ( fromBitSeq (bitSeqOr (toBitSeq x1) (toBitSeq x2)))" \ \\val defaultLxor : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> 'a -> 'a\\ definition defaultLxor :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ 'a \ 'a " where " defaultLxor fromBitSeq toBitSeq x1 x2 = ( fromBitSeq (bitSeqXor (toBitSeq x1) (toBitSeq x2)))" \ \\val defaultLsl : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> nat -> 'a\\ definition defaultLsl :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ nat \ 'a " where " defaultLsl fromBitSeq toBitSeq x n = ( fromBitSeq (bitSeqShiftLeft (toBitSeq x) n))" \ \\val defaultLsr : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> nat -> 'a\\ definition defaultLsr :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ nat \ 'a " where " defaultLsr fromBitSeq toBitSeq x n = ( fromBitSeq (bitSeqLogicalShiftRight (toBitSeq x) n))" \ \\val defaultAsr : forall 'a. (bitSequence -> 'a) -> ('a -> bitSequence) -> 'a -> nat -> 'a\\ definition defaultAsr :: "(bitSequence \ 'a)\('a \ bitSequence)\ 'a \ nat \ 'a " where " defaultAsr fromBitSeq toBitSeq x n = ( fromBitSeq (bitSeqArithmeticShiftRight (toBitSeq x) n))" \ \\ ----------------------- \\ \ \\ integer \\ \ \\ ----------------------- \\ \ \\val integerLnot : integer -> integer\\ definition integerLnot :: " int \ int " where " integerLnot i = ( - (i +( 1 :: int)))" definition instance_Word_WordNot_Num_integer_dict :: "(int)WordNot_class " where " instance_Word_WordNot_Num_integer_dict = ((| lnot_method = integerLnot |) )" \ \\val integerLor : integer -> integer -> integer\\ definition integerLor :: " int \ int \ int " where " integerLor i1 i2 = ( defaultLor integerFromBitSeq (bitSeqFromInteger None) i1 i2 )" definition instance_Word_WordOr_Num_integer_dict :: "(int)WordOr_class " where " instance_Word_WordOr_Num_integer_dict = ((| lor_method = integerLor |) )" \ \\val integerLxor : integer -> integer -> integer\\ definition integerLxor :: " int \ int \ int " where " integerLxor i1 i2 = ( defaultLxor integerFromBitSeq (bitSeqFromInteger None) i1 i2 )" definition instance_Word_WordXor_Num_integer_dict :: "(int)WordXor_class " where " instance_Word_WordXor_Num_integer_dict = ((| lxor_method = integerLxor |) )" \ \\val integerLand : integer -> integer -> integer\\ definition integerLand :: " int \ int \ int " where " integerLand i1 i2 = ( defaultLand integerFromBitSeq (bitSeqFromInteger None) i1 i2 )" definition instance_Word_WordAnd_Num_integer_dict :: "(int)WordAnd_class " where " instance_Word_WordAnd_Num_integer_dict = ((| land_method = integerLand |) )" \ \\val integerLsl : integer -> nat -> integer\\ definition integerLsl :: " int \ nat \ int " where " integerLsl i n = ( defaultLsl integerFromBitSeq (bitSeqFromInteger None) i n )" definition instance_Word_WordLsl_Num_integer_dict :: "(int)WordLsl_class " where " instance_Word_WordLsl_Num_integer_dict = ((| lsl_method = integerLsl |) )" \ \\val integerAsr : integer -> nat -> integer\\ definition integerAsr :: " int \ nat \ int " where " integerAsr i n = ( defaultAsr integerFromBitSeq (bitSeqFromInteger None) i n )" definition instance_Word_WordLsr_Num_integer_dict :: "(int)WordLsr_class " where " instance_Word_WordLsr_Num_integer_dict = ((| lsr_method = integerAsr |) )" definition instance_Word_WordAsr_Num_integer_dict :: "(int)WordAsr_class " where " instance_Word_WordAsr_Num_integer_dict = ((| asr_method = integerAsr |) )" \ \\ ----------------------- \\ \ \\ int \\ \ \\ ----------------------- \\ \ \\ sometimes it is convenient to be able to perform bit-operations on ints. However, since int is not well-defined (it has different size on different systems), it should be used very carefully and only for operations that don't depend on the bitwidth of int \\ \ \\val intFromBitSeq : bitSequence -> int\\ definition intFromBitSeq :: " bitSequence \ int " where " intFromBitSeq bs = ( (integerFromBitSeq (resizeBitSeq (Some(( 31 :: nat))) bs)))" \ \\val bitSeqFromInt : int -> bitSequence\\ definition bitSeqFromInt :: " int \ bitSequence " where " bitSeqFromInt i = ( bitSeqFromInteger (Some(( 31 :: nat))) ( i))" \ \\val intLnot : int -> int\\ definition intLnot :: " int \ int " where " intLnot i = ( - (i +( 1 :: int)))" definition instance_Word_WordNot_Num_int_dict :: "(int)WordNot_class " where " instance_Word_WordNot_Num_int_dict = ((| lnot_method = intLnot |) )" \ \\val intLor : int -> int -> int\\ definition intLor :: " int \ int \ int " where " intLor i1 i2 = ( defaultLor intFromBitSeq bitSeqFromInt i1 i2 )" definition instance_Word_WordOr_Num_int_dict :: "(int)WordOr_class " where " instance_Word_WordOr_Num_int_dict = ((| lor_method = intLor |) )" \ \\val intLxor : int -> int -> int\\ definition intLxor :: " int \ int \ int " where " intLxor i1 i2 = ( defaultLxor intFromBitSeq bitSeqFromInt i1 i2 )" definition instance_Word_WordXor_Num_int_dict :: "(int)WordXor_class " where " instance_Word_WordXor_Num_int_dict = ((| lxor_method = intLxor |) )" \ \\val intLand : int -> int -> int\\ definition intLand :: " int \ int \ int " where " intLand i1 i2 = ( defaultLand intFromBitSeq bitSeqFromInt i1 i2 )" definition instance_Word_WordAnd_Num_int_dict :: "(int)WordAnd_class " where " instance_Word_WordAnd_Num_int_dict = ((| land_method = intLand |) )" \ \\val intLsl : int -> nat -> int\\ definition intLsl :: " int \ nat \ int " where " intLsl i n = ( defaultLsl intFromBitSeq bitSeqFromInt i n )" definition instance_Word_WordLsl_Num_int_dict :: "(int)WordLsl_class " where " instance_Word_WordLsl_Num_int_dict = ((| lsl_method = intLsl |) )" \ \\val intAsr : int -> nat -> int\\ definition intAsr :: " int \ nat \ int " where " intAsr i n = ( defaultAsr intFromBitSeq bitSeqFromInt i n )" definition instance_Word_WordAsr_Num_int_dict :: "(int)WordAsr_class " where " instance_Word_WordAsr_Num_int_dict = ((| asr_method = intAsr |) )" \ \\ ----------------------- \\ \ \\ natural \\ \ \\ ----------------------- \\ \ \\ some operations work also on positive numbers \\ \ \\val naturalFromBitSeq : bitSequence -> natural\\ definition naturalFromBitSeq :: " bitSequence \ nat " where " naturalFromBitSeq bs = ( nat (abs (integerFromBitSeq bs)))" \ \\val bitSeqFromNatural : maybe nat -> natural -> bitSequence\\ definition bitSeqFromNatural :: "(nat)option \ nat \ bitSequence " where " bitSeqFromNatural len n = ( bitSeqFromInteger len (int n))" \ \\val naturalLor : natural -> natural -> natural\\ definition naturalLor :: " nat \ nat \ nat " where " naturalLor i1 i2 = ( defaultLor naturalFromBitSeq (bitSeqFromNatural None) i1 i2 )" definition instance_Word_WordOr_Num_natural_dict :: "(nat)WordOr_class " where " instance_Word_WordOr_Num_natural_dict = ((| lor_method = naturalLor |) )" \ \\val naturalLxor : natural -> natural -> natural\\ definition naturalLxor :: " nat \ nat \ nat " where " naturalLxor i1 i2 = ( defaultLxor naturalFromBitSeq (bitSeqFromNatural None) i1 i2 )" definition instance_Word_WordXor_Num_natural_dict :: "(nat)WordXor_class " where " instance_Word_WordXor_Num_natural_dict = ((| lxor_method = naturalLxor |) )" \ \\val naturalLand : natural -> natural -> natural\\ definition naturalLand :: " nat \ nat \ nat " where " naturalLand i1 i2 = ( defaultLand naturalFromBitSeq (bitSeqFromNatural None) i1 i2 )" definition instance_Word_WordAnd_Num_natural_dict :: "(nat)WordAnd_class " where " instance_Word_WordAnd_Num_natural_dict = ((| land_method = naturalLand |) )" \ \\val naturalLsl : natural -> nat -> natural\\ definition naturalLsl :: " nat \ nat \ nat " where " naturalLsl i n = ( defaultLsl naturalFromBitSeq (bitSeqFromNatural None) i n )" definition instance_Word_WordLsl_Num_natural_dict :: "(nat)WordLsl_class " where " instance_Word_WordLsl_Num_natural_dict = ((| lsl_method = naturalLsl |) )" \ \\val naturalAsr : natural -> nat -> natural\\ definition naturalAsr :: " nat \ nat \ nat " where " naturalAsr i n = ( defaultAsr naturalFromBitSeq (bitSeqFromNatural None) i n )" definition instance_Word_WordLsr_Num_natural_dict :: "(nat)WordLsr_class " where " instance_Word_WordLsr_Num_natural_dict = ((| lsr_method = naturalAsr |) )" definition instance_Word_WordAsr_Num_natural_dict :: "(nat)WordAsr_class " where " instance_Word_WordAsr_Num_natural_dict = ((| asr_method = naturalAsr |) )" \ \\ ----------------------- \\ \ \\ nat \\ \ \\ ----------------------- \\ \ \\ sometimes it is convenient to be able to perform bit-operations on nats. However, since nat is not well-defined (it has different size on different systems), it should be used very carefully and only for operations that don't depend on the bitwidth of nat \\ \ \\val natFromBitSeq : bitSequence -> nat\\ definition natFromBitSeq :: " bitSequence \ nat " where " natFromBitSeq bs = ( (naturalFromBitSeq (resizeBitSeq (Some(( 31 :: nat))) bs)))" \ \\val bitSeqFromNat : nat -> bitSequence\\ definition bitSeqFromNat :: " nat \ bitSequence " where " bitSeqFromNat i = ( bitSeqFromNatural (Some(( 31 :: nat))) ( i))" \ \\val natLor : nat -> nat -> nat\\ definition natLor :: " nat \ nat \ nat " where " natLor i1 i2 = ( defaultLor natFromBitSeq bitSeqFromNat i1 i2 )" definition instance_Word_WordOr_nat_dict :: "(nat)WordOr_class " where " instance_Word_WordOr_nat_dict = ((| lor_method = natLor |) )" \ \\val natLxor : nat -> nat -> nat\\ definition natLxor :: " nat \ nat \ nat " where " natLxor i1 i2 = ( defaultLxor natFromBitSeq bitSeqFromNat i1 i2 )" definition instance_Word_WordXor_nat_dict :: "(nat)WordXor_class " where " instance_Word_WordXor_nat_dict = ((| lxor_method = natLxor |) )" \ \\val natLand : nat -> nat -> nat\\ definition natLand :: " nat \ nat \ nat " where " natLand i1 i2 = ( defaultLand natFromBitSeq bitSeqFromNat i1 i2 )" definition instance_Word_WordAnd_nat_dict :: "(nat)WordAnd_class " where " instance_Word_WordAnd_nat_dict = ((| land_method = natLand |) )" \ \\val natLsl : nat -> nat -> nat\\ definition natLsl :: " nat \ nat \ nat " where " natLsl i n = ( defaultLsl natFromBitSeq bitSeqFromNat i n )" definition instance_Word_WordLsl_nat_dict :: "(nat)WordLsl_class " where " instance_Word_WordLsl_nat_dict = ((| lsl_method = natLsl |) )" \ \\val natAsr : nat -> nat -> nat\\ definition natAsr :: " nat \ nat \ nat " where " natAsr i n = ( defaultAsr natFromBitSeq bitSeqFromNat i n )" definition instance_Word_WordAsr_nat_dict :: "(nat)WordAsr_class " where " instance_Word_WordAsr_nat_dict = ((| asr_method = natAsr |) )" end diff --git a/thys/Collections/GenCF/Impl/Impl_Bit_Set.thy b/thys/Collections/GenCF/Impl/Impl_Bit_Set.thy --- a/thys/Collections/GenCF/Impl/Impl_Bit_Set.thy +++ b/thys/Collections/GenCF/Impl/Impl_Bit_Set.thy @@ -1,382 +1,384 @@ section "Bitvector based Sets of Naturals" theory Impl_Bit_Set imports "../../Iterator/Iterator" "../Intf/Intf_Set" Native_Word.Bits_Integer begin text \ Based on the Native-Word library, using bit-operations on arbitrary precision integers. Fast for sets of small numbers, direct and fast implementations of equal, union, inter, diff. Note: On Poly/ML 5.5.1, bit-operations on arbitrary precision integers are rather inefficient. Use MLton instead, here they are efficiently implemented. \ type_synonym bitset = integer definition bs_\ :: "bitset \ nat set" where "bs_\ s \ { n . bit s n}" -context includes integer.lifting begin + context + includes integer.lifting bit_operations_syntax + begin definition bs_empty :: "unit \ bitset" where "bs_empty \ \_. 0" lemma bs_empty_correct: "bs_\ (bs_empty ()) = {}" unfolding bs_\_def bs_empty_def apply transfer by auto definition bs_isEmpty :: "bitset \ bool" where "bs_isEmpty s \ s=0" lemma bs_isEmpty_correct: "bs_isEmpty s \ bs_\ s = {}" unfolding bs_isEmpty_def bs_\_def by transfer (auto simp: bit_eq_iff) term set_bit definition bs_insert :: "nat \ bitset \ bitset" where "bs_insert i s \ set_bit s i True" lemma bs_insert_correct: "bs_\ (bs_insert i s) = insert i (bs_\ s)" unfolding bs_\_def bs_insert_def by transfer (auto simp add: bit_simps) definition bs_delete :: "nat \ bitset \ bitset" where "bs_delete i s \ set_bit s i False" lemma bs_delete_correct: "bs_\ (bs_delete i s) = (bs_\ s) - {i}" unfolding bs_\_def bs_delete_def by transfer (auto simp add: bit_simps split: if_splits) definition bs_mem :: "nat \ bitset \ bool" where "bs_mem i s \ bit s i" lemma bs_mem_correct: "bs_mem i s \ i\bs_\ s" unfolding bs_mem_def bs_\_def by transfer auto definition bs_eq :: "bitset \ bitset \ bool" where "bs_eq s1 s2 \ (s1=s2)" lemma bs_eq_correct: "bs_eq s1 s2 \ bs_\ s1 = bs_\ s2" unfolding bs_eq_def bs_\_def including integer.lifting by transfer (simp add: bit_eq_iff set_eq_iff) definition bs_subset_eq :: "bitset \ bitset \ bool" where "bs_subset_eq s1 s2 \ s1 AND NOT s2 = 0" lemma bs_subset_eq_correct: "bs_subset_eq s1 s2 \ bs_\ s1 \ bs_\ s2" unfolding bs_\_def bs_subset_eq_def by transfer (simp add: bit_eq_iff, auto simp add: bit_simps) definition bs_disjoint :: "bitset \ bitset \ bool" where "bs_disjoint s1 s2 \ s1 AND s2 = 0" lemma bs_disjoint_correct: "bs_disjoint s1 s2 \ bs_\ s1 \ bs_\ s2 = {}" unfolding bs_\_def bs_disjoint_def by transfer (simp add: bit_eq_iff, auto simp add: bit_simps) definition bs_union :: "bitset \ bitset \ bitset" where "bs_union s1 s2 = s1 OR s2" lemma bs_union_correct: "bs_\ (bs_union s1 s2) = bs_\ s1 \ bs_\ s2" unfolding bs_\_def bs_union_def by transfer (simp add: bit_eq_iff, auto simp add: bit_simps) definition bs_inter :: "bitset \ bitset \ bitset" where "bs_inter s1 s2 = s1 AND s2" lemma bs_inter_correct: "bs_\ (bs_inter s1 s2) = bs_\ s1 \ bs_\ s2" unfolding bs_\_def bs_inter_def by transfer (simp add: bit_eq_iff, auto simp add: bit_simps) definition bs_diff :: "bitset \ bitset \ bitset" where "bs_diff s1 s2 = s1 AND NOT s2" lemma bs_diff_correct: "bs_\ (bs_diff s1 s2) = bs_\ s1 - bs_\ s2" unfolding bs_\_def bs_diff_def by transfer (simp add: bit_eq_iff, auto simp add: bit_simps) definition bs_UNIV :: "unit \ bitset" where "bs_UNIV \ \_. -1" lemma bs_UNIV_correct: "bs_\ (bs_UNIV ()) = UNIV" unfolding bs_\_def bs_UNIV_def by transfer (auto) definition bs_complement :: "bitset \ bitset" where "bs_complement s = NOT s" lemma bs_complement_correct: "bs_\ (bs_complement s) = - bs_\ s" unfolding bs_\_def bs_complement_def by transfer (simp add: bit_eq_iff, auto simp add: bit_simps) end lemmas bs_correct[simp] = bs_empty_correct bs_isEmpty_correct bs_insert_correct bs_delete_correct bs_mem_correct bs_eq_correct bs_subset_eq_correct bs_disjoint_correct bs_union_correct bs_inter_correct bs_diff_correct bs_UNIV_correct bs_complement_correct subsection \Autoref Setup\ definition bs_set_rel_def_internal: "bs_set_rel Rk \ if Rk=nat_rel then br bs_\ (\_. True) else {}" lemma bs_set_rel_def: "\nat_rel\bs_set_rel \ br bs_\ (\_. True)" unfolding bs_set_rel_def_internal relAPP_def by simp lemmas [autoref_rel_intf] = REL_INTFI[of "bs_set_rel" i_set] lemma bs_set_rel_sv[relator_props]: "single_valued (\nat_rel\bs_set_rel)" unfolding bs_set_rel_def by auto term bs_empty lemma [autoref_rules]: "(bs_empty (),{})\\nat_rel\bs_set_rel" by (auto simp: bs_set_rel_def br_def) lemma [autoref_rules]: "(bs_UNIV (),UNIV)\\nat_rel\bs_set_rel" by (auto simp: bs_set_rel_def br_def) lemma [autoref_rules]: "(bs_isEmpty,op_set_isEmpty)\\nat_rel\bs_set_rel \ bool_rel" by (auto simp: bs_set_rel_def br_def) term insert lemma [autoref_rules]: "(bs_insert,insert)\nat_rel \ \nat_rel\bs_set_rel \ \nat_rel\bs_set_rel" by (auto simp: bs_set_rel_def br_def) term op_set_delete lemma [autoref_rules]: "(bs_delete,op_set_delete)\nat_rel \ \nat_rel\bs_set_rel \ \nat_rel\bs_set_rel" by (auto simp: bs_set_rel_def br_def) lemma [autoref_rules]: "(bs_mem,(\))\nat_rel \ \nat_rel\bs_set_rel \ bool_rel" by (auto simp: bs_set_rel_def br_def) lemma [autoref_rules]: "(bs_eq,(=))\\nat_rel\bs_set_rel \ \nat_rel\bs_set_rel \ bool_rel" by (auto simp: bs_set_rel_def br_def) lemma [autoref_rules]: "(bs_subset_eq,(\))\\nat_rel\bs_set_rel \ \nat_rel\bs_set_rel \ bool_rel" by (auto simp: bs_set_rel_def br_def) lemma [autoref_rules]: "(bs_union,(\))\\nat_rel\bs_set_rel \ \nat_rel\bs_set_rel \ \nat_rel\bs_set_rel" by (auto simp: bs_set_rel_def br_def) lemma [autoref_rules]: "(bs_inter,(\))\\nat_rel\bs_set_rel \ \nat_rel\bs_set_rel \ \nat_rel\bs_set_rel" by (auto simp: bs_set_rel_def br_def) lemma [autoref_rules]: "(bs_diff,(-))\\nat_rel\bs_set_rel \ \nat_rel\bs_set_rel \ \nat_rel\bs_set_rel" by (auto simp: bs_set_rel_def br_def) lemma [autoref_rules]: "(bs_complement,uminus)\\nat_rel\bs_set_rel \ \nat_rel\bs_set_rel" by (auto simp: bs_set_rel_def br_def) lemma [autoref_rules]: "(bs_disjoint,op_set_disjoint)\\nat_rel\bs_set_rel \ \nat_rel\bs_set_rel \ bool_rel" by (auto simp: bs_set_rel_def br_def) export_code bs_empty bs_isEmpty bs_insert bs_delete bs_mem bs_eq bs_subset_eq bs_disjoint bs_union bs_inter bs_diff bs_UNIV bs_complement in SML (* TODO: Iterator definition "maxbi s \ GREATEST i. s!!i" lemma cmp_BIT_append_conv[simp]: "i < i BIT b \ ((i\0 \ b=1) \ i>0)" by (cases b) (auto simp: Bit_B0 Bit_B1) lemma BIT_append_cmp_conv[simp]: "i BIT b < i \ ((i<0 \ (i=-1 \ b=0)))" by (cases b) (auto simp: Bit_B0 Bit_B1) lemma BIT_append_eq[simp]: fixes i :: int shows "i BIT b = i \ (i=0 \ b=0) \ (i=-1 \ b=1)" by (cases b) (auto simp: Bit_B0 Bit_B1) lemma int_no_bits_eq_zero[simp]: fixes s::int shows "(\i. \s!!i) \ s=0" apply clarsimp by (metis bin_eqI bin_nth_code(1)) lemma int_obtain_bit: fixes s::int assumes "s\0" obtains i where "s!!i" by (metis assms int_no_bits_eq_zero) lemma int_bit_bound: fixes s::int assumes "s\0" and "s!!i" shows "i \ Bits_Integer.log2 s" proof (rule ccontr) assume "\i\Bits_Integer.log2 s" hence "i>Bits_Integer.log2 s" by simp hence "i - 1 \ Bits_Integer.log2 s" by simp hence "s AND bin_mask (i - 1) = s" by (simp add: int_and_mask `s\0`) hence "\ (s!!i)" by clarsimp (metis Nat.diff_le_self bin_nth_mask bin_nth_ops(1) leD) thus False using `s!!i` .. qed lemma int_bit_bound': fixes s::int assumes "s\0" and "s!!i" shows "i < Bits_Integer.log2 s + 1" using assms int_bit_bound by smt lemma int_obtain_bit_pos: fixes s::int assumes "s>0" obtains i where "s!!i" "i < Bits_Integer.log2 s + 1" by (metis assms int_bit_bound' int_no_bits_eq_zero less_imp_le less_irrefl) lemma maxbi_set: fixes s::int shows "s>0 \ s!!maxbi s" unfolding maxbi_def apply (rule int_obtain_bit_pos, assumption) apply (rule GreatestI_nat, assumption) apply (intro allI impI) apply (rule int_bit_bound'[rotated], assumption) by auto lemma maxbi_max: fixes s::int shows "i>maxbi s \ \ s!!i" oops function get_maxbi :: "nat \ int \ nat" where "get_maxbi n s = (let b = 1<s then get_maxbi (n+1) s else n )" by pat_completeness auto termination apply (rule "termination"[of "measure (\(n,s). nat (s + 1 - (1< bitset \ ('\ \ bool) \ (nat \ '\ \ '\) \ '\ \ '\" where "bs_iterate_aux i s c f \ = ( if s < 1 << i then \ else if \c \ then \ else if test_bit s i then bs_iterate_aux (i+1) s c f (f i \) else bs_iterate_aux (i+1) s c f \ )" definition bs_iteratei :: "bitset \ (nat,'\) set_iterator" where "bs_iteratei s = bs_iterate_aux 0 s" definition bs_set_rel_def_internal: "bs_set_rel Rk \ if Rk=nat_rel then br bs_\ (\_. True) else {}" lemma bs_set_rel_def: "\nat_rel\bs_set_rel \ br bs_\ (\_. True)" unfolding bs_set_rel_def_internal relAPP_def by simp definition "bs_to_list \ it_to_list bs_iteratei" lemma "(1::int)<0" shows "s < 1< Bits_Integer.log2 s \ i" using assms proof (induct i arbitrary: s) case 0 thus ?case by auto next case (Suc i) note GE=`0\s` show ?case proof assume "s < 1 << Suc i" have "s \ (s >> 1) BIT 1" hence "(s >> 1) < (1<_. True) (\x l. l @ [x]) [])" } apply auto lemma "set (bs_to_list s) = bs_\ s" lemma autoref_iam_is_iterator[autoref_ga_rules]: shows "is_set_to_list nat_rel bs_set_rel bs_to_list" unfolding is_set_to_list_def is_set_to_sorted_list_def apply clarsimp unfolding it_to_sorted_list_def apply (refine_rcg refine_vcg) apply (simp_all add: bs_set_rel_def br_def) proof (clarsimp) definition "iterate s c f \ \ let i=0; b=0; (_,_,s) = while in end" *) end diff --git a/thys/Collections/GenCF/Impl/Impl_Uv_Set.thy b/thys/Collections/GenCF/Impl/Impl_Uv_Set.thy --- a/thys/Collections/GenCF/Impl/Impl_Uv_Set.thy +++ b/thys/Collections/GenCF/Impl/Impl_Uv_Set.thy @@ -1,396 +1,401 @@ theory Impl_Uv_Set imports "../../Iterator/Iterator" "../Intf/Intf_Set" Native_Word.Uint begin subsection \Bit-Vectors as Lists of Words\ subsubsection \Lookup\ primrec lookup :: "nat \ ('a::len) word list \ bool" where "lookup _ [] \ False" | "lookup n (w#ws) \ (if n ( if n < LENGTH('a) * length w1 then lookup n w1 else lookup (n - LENGTH('a) * length w1) w2)" by (induction w1 arbitrary: n) auto lemma lookup_zeroes[simp]: "lookup i (replicate n (0::'a::len word)) = False" by (induction n arbitrary: i) auto lemma lookup_out_of_bound: fixes uv :: "'a::len word list" assumes "\ i < LENGTH('a::len) * length uv" shows "\ lookup i uv" using assms by (induction uv arbitrary: i) auto subsubsection \Empty\ definition empty :: "'a::len word list" where "empty = []" subsubsection \Set and Reset Bit\ function single_bit :: "nat \ ('a::len) word list" where "single_bit n = ( if (n i = n" apply (induction n arbitrary: i rule: single_bit.induct) apply (subst single_bit.simps) apply (auto simp add: bit_simps) done primrec set_bit :: "nat \ 'a::len word list \ 'a::len word list" where "set_bit i [] = single_bit i" | "set_bit i (w#ws) = ( if i (lookup i ws \ i=j)" apply (induction ws arbitrary: i j) apply (auto simp add: word_size Bit_Operations.bit_set_bit_iff) done primrec reset_bit :: "nat \ 'a::len word list \ 'a::len word list" where "reset_bit i [] = []" | "reset_bit i (w#ws) = ( if i (lookup i ws \ i\j)" apply (induction ws arbitrary: i j) apply (auto simp: word_size bit_unset_bit_iff) done subsubsection \Binary Operations\ + context + includes bit_operations_syntax + begin + definition is_bin_op_impl :: "(bool\bool\bool) \ ('a::len word \ 'a::len word \ 'a::len word) \ bool" where "is_bin_op_impl f g \ (\w v. \i f (bit w i) (bit v i))" definition "is_strict_bin_op_impl f g \ is_bin_op_impl f g \ f False False = False" fun binary :: "('a::len word \ 'a::len word \ 'a::len word) \ 'a::len word list \ 'a::len word list \ 'a::len word list" where "binary f [] [] = []" | "binary f [] (w#ws) = f 0 w # binary f [] ws" | "binary f (v#vs) [] = f v 0 # binary f vs []" | "binary f (v#vs) (w#ws) = f v w # binary f vs ws" lemma binary_lookup: assumes "is_strict_bin_op_impl f g" shows "lookup i (binary g ws vs) \ f (lookup i ws) (lookup i vs)" using assms apply (induction g ws vs arbitrary: i rule: binary.induct) apply (auto simp: is_strict_bin_op_impl_def is_bin_op_impl_def) done subsection \Abstraction to Sets of Naturals\ definition "\ uv \ {n. lookup n uv}" lemma memb_correct: "lookup i ws \ i\\ ws" by (auto simp: \_def) lemma empty_correct: "\ empty = {}" by (simp add: \_def empty_def) lemma single_bit_correct: "\ (single_bit n) = {n}" by (simp add: \_def) lemma insert_correct: "\ (set_bit i ws) = Set.insert i (\ ws)" by (auto simp add: \_def) lemma delete_correct: "\ (reset_bit i ws) = (\ ws) - {i}" by (auto simp add: \_def) lemma binary_correct: assumes "is_strict_bin_op_impl f g" shows "\ (binary g ws vs) = { i . f (i\\ ws) (i\\ vs) }" unfolding \_def by (auto simp add: binary_lookup[OF assms]) fun union :: "'a::len word list \ 'a::len word list \ 'a::len word list" where "union [] ws = ws" | "union vs [] = vs" | "union (v#vs) (w#ws) = (v OR w) # union vs ws" lemma union_lookup[simp]: fixes vs :: "'a::len word list" shows "lookup i (union vs ws) \ lookup i vs \ lookup i ws" apply (induction vs ws arbitrary: i rule: union.induct) apply (auto simp: word_ao_nth) done lemma union_correct: "\ (union ws vs) = \ ws \ \ vs" by (auto simp add: \_def) fun inter :: "'a::len word list \ 'a::len word list \ 'a::len word list" where "inter [] ws = []" | "inter vs [] = []" | "inter (v#vs) (w#ws) = (v AND w) # inter vs ws" lemma inter_lookup[simp]: fixes vs :: "'a::len word list" shows "lookup i (inter vs ws) \ lookup i vs \ lookup i ws" apply (induction vs ws arbitrary: i rule: inter.induct) apply (auto simp: word_ao_nth) done lemma inter_correct: "\ (inter ws vs) = \ ws \ \ vs" by (auto simp add: \_def) fun diff :: "'a::len word list \ 'a::len word list \ 'a::len word list" where "diff [] ws = []" | "diff vs [] = vs" | "diff (v#vs) (w#ws) = (v AND NOT w) # diff vs ws" lemma diff_lookup[simp]: fixes vs :: "'a::len word list" shows "lookup i (diff vs ws) \ lookup i vs - lookup i ws" apply (induction vs ws arbitrary: i rule: diff.induct) apply (auto simp: word_ops_nth_size word_size) done lemma diff_correct: "\ (diff ws vs) = \ ws - \ vs" by (auto simp add: \_def) fun zeroes :: "'a::len word list \ bool" where "zeroes [] \ True" | "zeroes (w#ws) \ w=0 \ zeroes ws" lemma zeroes_lookup: "zeroes ws \ (\i. \lookup i ws)" apply (induction ws) apply (auto simp: word_eq_iff) by (metis diff_add_inverse2 not_add_less2) lemma isEmpty_correct: "zeroes ws \ \ ws = {}" by (auto simp: zeroes_lookup \_def) fun equal :: "'a::len word list \ 'a::len word list \ bool" where "equal [] [] \ True" | "equal [] ws \ zeroes ws" | "equal vs [] \ zeroes vs" | "equal (v#vs) (w#ws) \ v=w \ equal vs ws" lemma equal_lookup: fixes vs ws :: "'a::len word list" shows "equal vs ws \ (\i. lookup i vs = lookup i ws)" proof (induction vs ws rule: equal.induct) fix v w and vs ws :: "'a::len word list" assume IH: "equal vs ws = (\i. lookup i vs = lookup i ws)" show "equal (v # vs) (w # ws) = (\i. lookup i (v # vs) = lookup i (w # ws))" proof (rule iffI, rule allI) fix i assume "equal (v#vs) (w#ws)" thus "lookup i (v # vs) = lookup i (w # ws)" by (auto simp: IH) next assume "\i. lookup i (v # vs) = lookup i (w # ws)" thus "equal (v # vs) (w # ws)" apply (auto simp: word_eq_iff IH) apply metis apply metis apply (drule_tac x="i + LENGTH('a)" in spec) apply auto [] apply (drule_tac x="i + LENGTH('a)" in spec) apply auto [] done qed qed (auto simp: zeroes_lookup) lemma equal_correct: "equal vs ws \ \ vs = \ ws" by (auto simp: \_def equal_lookup) fun subseteq :: "'a::len word list \ 'a::len word list \ bool" where "subseteq [] ws \ True" | "subseteq vs [] \ zeroes vs" | "subseteq (v#vs) (w#ws) \ (v AND NOT w = 0) \ subseteq vs ws" lemma subseteq_lookup: fixes vs ws :: "'a::len word list" shows "subseteq vs ws \ (\i. lookup i vs \ lookup i ws)" apply (induction vs ws rule: subseteq.induct) apply simp apply (auto simp: zeroes_lookup) [] apply (auto simp: word_ops_nth_size word_size word_eq_iff) by (metis diff_add_inverse2 not_add_less2) lemma subseteq_correct: "subseteq vs ws \ \ vs \ \ ws" by (auto simp: \_def subseteq_lookup) fun subset :: "'a::len word list \ 'a::len word list \ bool" where "subset [] ws \ \zeroes ws" | "subset vs [] \ False" | "subset (v#vs) (w#ws) \ (if v=w then subset vs ws else subseteq (v#vs) (w#ws))" lemma subset_lookup: fixes vs ws :: "'a::len word list" shows "subset vs ws \ ((\i. lookup i vs \ lookup i ws) \ (\i. \lookup i vs \ lookup i ws))" apply (induction vs ws rule: subset.induct) apply (simp add: zeroes_lookup) apply (simp add: zeroes_lookup) [] apply (simp add: subseteq_lookup) apply safe apply simp_all apply (auto simp: word_ops_nth_size word_size word_eq_iff) done lemma subset_correct: "subset vs ws \ \ vs \ \ ws" by (auto simp: \_def subset_lookup) fun disjoint :: "'a::len word list \ 'a::len word list \ bool" where "disjoint [] ws \ True" | "disjoint vs [] \ True" | "disjoint (v#vs) (w#ws) \ (v AND w = 0) \ disjoint vs ws" lemma disjoint_lookup: fixes vs ws :: "'a::len word list" shows "disjoint vs ws \ (\i. \(lookup i vs \ lookup i ws))" apply (induction vs ws rule: disjoint.induct) apply simp apply simp apply (auto simp: word_ops_nth_size word_size word_eq_iff) by (metis diff_add_inverse2 not_add_less2) lemma disjoint_correct: "disjoint vs ws \ \ vs \ \ ws = {}" by (auto simp: \_def disjoint_lookup) subsection \Lifting to Uint\ type_synonym uint_vector = "uint list" lift_definition uv_\ :: "uint_vector \ nat set" is \ . lift_definition uv_lookup :: "nat \ uint_vector \ bool" is lookup . lift_definition uv_empty :: "uint_vector" is empty . lift_definition uv_single_bit :: "nat \ uint_vector" is single_bit . lift_definition uv_set_bit :: "nat \ uint_vector \ uint_vector" is set_bit . lift_definition uv_reset_bit :: "nat \ uint_vector \ uint_vector" is reset_bit . lift_definition uv_union :: "uint_vector \ uint_vector \ uint_vector" is union . lift_definition uv_inter :: "uint_vector \ uint_vector \ uint_vector" is inter . lift_definition uv_diff :: "uint_vector \ uint_vector \ uint_vector" is diff . lift_definition uv_zeroes :: "uint_vector \ bool" is zeroes . lift_definition uv_equal :: "uint_vector \ uint_vector \ bool" is equal . lift_definition uv_subseteq :: "uint_vector \ uint_vector \ bool" is subseteq . lift_definition uv_subset :: "uint_vector \ uint_vector \ bool" is subset . lift_definition uv_disjoint :: "uint_vector \ uint_vector \ bool" is disjoint . lemmas uv_memb_correct = memb_correct[where 'a=dflt_size, Transfer.transferred] lemmas uv_empty_correct = empty_correct[where 'a=dflt_size, Transfer.transferred] lemmas uv_single_bit_correct = single_bit_correct[where 'a=dflt_size, Transfer.transferred] lemmas uv_insert_correct = insert_correct[where 'a=dflt_size, Transfer.transferred] lemmas uv_delete_correct = delete_correct[where 'a=dflt_size, Transfer.transferred] lemmas uv_union_correct = union_correct[where 'a=dflt_size, Transfer.transferred] lemmas uv_inter_correct = inter_correct[where 'a=dflt_size, Transfer.transferred] lemmas uv_diff_correct = diff_correct[where 'a=dflt_size, Transfer.transferred] lemmas uv_isEmpty_correct = isEmpty_correct[where 'a=dflt_size, Transfer.transferred] lemmas uv_equal_correct = equal_correct[where 'a=dflt_size, Transfer.transferred] lemmas uv_subseteq_correct = subseteq_correct[where 'a=dflt_size, Transfer.transferred] lemmas uv_subset_correct = subset_correct[where 'a=dflt_size, Transfer.transferred] lemmas uv_disjoint_correct = disjoint_correct[where 'a=dflt_size, Transfer.transferred] lemmas [where 'a=dflt_size, Transfer.transferred, code] = lookup.simps empty_def single_bit.simps set_bit.simps reset_bit.simps union.simps inter.simps diff.simps zeroes.simps equal.simps subseteq.simps subset.simps disjoint.simps + end hide_const (open) \ lookup empty single_bit set_bit reset_bit union inter diff zeroes equal subseteq subset disjoint -subsection \Autoref Setup\ + subsection \Autoref Setup\ definition uv_set_rel_def_internal: "uv_set_rel Rk \ if Rk=nat_rel then br uv_\ (\_. True) else {}" lemma uv_set_rel_def: "\nat_rel\uv_set_rel \ br uv_\ (\_. True)" unfolding uv_set_rel_def_internal relAPP_def by simp lemmas [autoref_rel_intf] = REL_INTFI[of "uv_set_rel" i_set] lemma uv_set_rel_sv[relator_props]: "single_valued (\nat_rel\uv_set_rel)" unfolding uv_set_rel_def by auto lemma uv_autoref[autoref_rules,param]: "(uv_lookup,(\)) \ nat_rel \ \nat_rel\uv_set_rel \ bool_rel" "(uv_empty,{}) \ \nat_rel\uv_set_rel" "(uv_set_bit,insert) \ nat_rel \ \nat_rel\uv_set_rel \ \nat_rel\uv_set_rel" "(uv_reset_bit,op_set_delete) \ nat_rel \ \nat_rel\uv_set_rel \ \nat_rel\uv_set_rel" "(uv_union,(\)) \ \nat_rel\uv_set_rel \ \nat_rel\uv_set_rel \ \nat_rel\uv_set_rel" "(uv_inter,(\)) \ \nat_rel\uv_set_rel \ \nat_rel\uv_set_rel \ \nat_rel\uv_set_rel" "(uv_diff,(-)) \ \nat_rel\uv_set_rel \ \nat_rel\uv_set_rel \ \nat_rel\uv_set_rel" "(uv_zeroes,op_set_isEmpty) \ \nat_rel\uv_set_rel \ bool_rel" "(uv_equal,(=)) \ \nat_rel\uv_set_rel \ \nat_rel\uv_set_rel \ bool_rel" "(uv_subseteq,(\)) \ \nat_rel\uv_set_rel \ \nat_rel\uv_set_rel \ bool_rel" "(uv_subset,(\)) \ \nat_rel\uv_set_rel \ \nat_rel\uv_set_rel \ bool_rel" "(uv_disjoint,op_set_disjoint) \ \nat_rel\uv_set_rel \ \nat_rel\uv_set_rel \ bool_rel" by (auto simp: uv_set_rel_def br_def simp: uv_memb_correct uv_empty_correct uv_insert_correct uv_delete_correct simp: uv_union_correct uv_inter_correct uv_diff_correct uv_isEmpty_correct simp: uv_equal_correct uv_subseteq_correct uv_subset_correct uv_disjoint_correct) export_code uv_lookup uv_empty uv_single_bit uv_set_bit uv_reset_bit uv_union uv_inter uv_diff uv_zeroes uv_equal uv_subseteq uv_subset uv_disjoint checking SML Scala Haskell? OCaml? 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,468 +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) 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/IP_Addresses/IP_Address.thy b/thys/IP_Addresses/IP_Address.thy --- a/thys/IP_Addresses/IP_Address.thy +++ b/thys/IP_Addresses/IP_Address.thy @@ -1,415 +1,421 @@ (* Title: IP_Address.thy Authors: Cornelius Diekmann *) theory IP_Address imports Word_Lib.Word_Lemmas Word_Lib.Word_Syntax Word_Lib.Reversed_Bit_Lists Hs_Compat WordInterval begin section \Modelling IP Adresses\ text\An IP address is basically an unsigned integer. We model IP addresses of arbitrary lengths. We will write @{typ "'i::len word"} for IP addresses of length @{term "LENGTH('i::len)"}. We use the convention to write @{typ 'i} whenever we mean IP addresses instead of generic words. When we will later have theorems with several polymorphic types in it (e.g. arbitrarily extensible packets), this notation makes it easier to spot that type @{typ 'i} is for IP addresses. The files @{file \IPv4.thy\} @{file \IPv6.thy\} concrete this for IPv4 and IPv6.\ text\The maximum IP address\ definition max_ip_addr :: "'i::len word" where "max_ip_addr \ of_nat ((2^(len_of(TYPE('i)))) - 1)" lemma max_ip_addr_max_word: "max_ip_addr = - 1" by (simp only: max_ip_addr_def of_nat_mask_eq flip: mask_eq_exp_minus_1) simp lemma max_ip_addr_max: "\a. a \ max_ip_addr" by(simp add: max_ip_addr_max_word) lemma range_0_max_UNIV: "UNIV = {0 .. max_ip_addr}" (*not in the simp set, for a reason*) by(simp add: max_ip_addr_max_word) fastforce lemma "size (x::'i::len word) = len_of(TYPE('i))" by(simp add:word_size) subsection\Sets of IP Addresses\ + + context + includes bit_operations_syntax + begin + (*Warning, not executable!*) text\Specifying sets with network masks: 192.168.0.0 255.255.255.0\ definition ipset_from_netmask::"'i::len word \ 'i::len word \ 'i::len word set" where "ipset_from_netmask addr netmask \ let network_prefix = (addr AND netmask) in {network_prefix .. network_prefix OR (NOT netmask)}" text\Example (pseudo syntax): @{const ipset_from_netmask} \192.168.1.129 255.255.255.0\ = \{192.168.1.0 .. 192.168.1.255}\\ text\A network mask of all ones (i.e. @{term "(- 1)::'i::len word"}).\ lemma ipset_from_netmask_minusone: "ipset_from_netmask ip (- 1) = {ip}" by (simp add: ipset_from_netmask_def) lemma ipset_from_netmask_maxword: "ipset_from_netmask ip (- 1) = {ip}" by (simp add: ipset_from_netmask_def) lemma ipset_from_netmask_zero: "ipset_from_netmask ip 0 = UNIV" by (auto simp add: ipset_from_netmask_def) text\Specifying sets in Classless Inter-domain Routing (CIDR) notation: 192.168.0.0/24\ definition ipset_from_cidr ::"'i::len word \ nat \ 'i::len word set" where "ipset_from_cidr addr pflength \ ipset_from_netmask addr ((mask pflength) << (len_of(TYPE('i)) - pflength))" text\Example (pseudo syntax): @{const ipset_from_cidr} \192.168.1.129 24\ = \{192.168.1.0 .. 192.168.1.255}\\ (*does this simplify stuff?*) lemma "(case ipcidr of (base, len) \ ipset_from_cidr base len) = uncurry ipset_from_cidr ipcidr" by(simp add: uncurry_case_stmt) lemma ipset_from_cidr_0: "ipset_from_cidr ip 0 = UNIV" by(auto simp add: ipset_from_cidr_def ipset_from_netmask_def Let_def) text\A prefix length of word size gives back the singleton set with the IP address. Example: \192.168.1.2/32 = {192.168.1.2}\\ lemma ipset_from_cidr_wordlength: fixes ip :: "'i::len word" shows "ipset_from_cidr ip (LENGTH('i)) = {ip}" by (simp add: ipset_from_cidr_def ipset_from_netmask_def) text\Alternative definition: Considering words as bit lists:\ lemma ipset_from_cidr_bl: fixes addr :: "'i::len word" shows "ipset_from_cidr addr pflength \ ipset_from_netmask addr (of_bl ((replicate pflength True) @ (replicate ((len_of(TYPE('i))) - pflength)) False))" by(simp add: ipset_from_cidr_def mask_bl shiftl_of_bl) lemma ipset_from_cidr_alt: fixes pre :: "'i::len word" shows "ipset_from_cidr pre len = {pre AND (mask len << LENGTH('i) - len) .. pre OR mask (LENGTH('i) - len)}" apply(simp add: ipset_from_cidr_def ipset_from_netmask_def Let_def) apply(simp add: word_oa_dist) apply(simp add: NOT_mask_shifted_lenword) done lemma ipset_from_cidr_alt2: fixes base ::"'i::len word" shows "ipset_from_cidr base len = ipset_from_netmask base (NOT (mask (LENGTH('i) - len)))" apply(simp add: ipset_from_cidr_def) using NOT_mask_shifted_lenword by(metis word_not_not) text\In CIDR notation, we cannot express the empty set.\ lemma ipset_from_cidr_not_empty: "ipset_from_cidr base len \ {}" by(simp add: ipset_from_cidr_alt bitmagic_zeroLast_leq_or1Last) text\Though we can write 192.168.1.2/24, we say that 192.168.0.0/24 is well-formed.\ lemma ipset_from_cidr_base_wellforemd: fixes base:: "'i::len word" assumes "mask (LENGTH('i) - l) AND base = 0" shows "ipset_from_cidr base l = {base .. base OR mask (LENGTH('i) - l)}" proof - have maskshift_eq_not_mask_generic: "((mask l << LENGTH('i) - l) :: 'i::len word) = NOT (mask (LENGTH('i) - l))" using NOT_mask_shifted_lenword by (metis word_not_not) have *: "base AND NOT (mask (LENGTH('i) - l)) = base" unfolding mask_eq_0_eq_x[symmetric] using assms word_bw_comms(1)[of base] by simp hence **: "base AND NOT (mask (LENGTH('i) - l)) OR mask (LENGTH('i) - l) = base OR mask (LENGTH('i) - l)" by simp have "ipset_from_netmask base (NOT (mask (LENGTH('i) - l))) = {base .. base || mask (LENGTH('i) - l)}" by(simp add: ipset_from_netmask_def Let_def ** *) thus ?thesis by(simp add: ipset_from_cidr_def maskshift_eq_not_mask_generic) qed lemma ipset_from_cidr_large_pfxlen: fixes ip:: "'i::len word" assumes "n \ LENGTH('i)" shows "ipset_from_cidr ip n = {ip}" proof - have obviously: "mask (LENGTH('i) - n) = 0" by (simp add: assms) show ?thesis apply(subst ipset_from_cidr_base_wellforemd) subgoal using assms by simp by (simp add: obviously) qed lemma ipset_from_netmask_base_mask_consume: fixes base :: "'i::len word" shows "ipset_from_netmask (base AND NOT (mask (LENGTH('i) - m))) (NOT (mask (LENGTH('i) - m))) = ipset_from_netmask base (NOT (mask (LENGTH('i) - m)))" unfolding ipset_from_netmask_def by(simp) text\Another definition of CIDR notation: All IP address which are equal on the first @{term "len - n"} bits\ definition ip_cidr_set :: "'i::len word \ nat \ 'i word set" where "ip_cidr_set i r \ {j . i AND NOT (mask (LENGTH('i) - r)) = j AND NOT (mask (LENGTH('i) - r))}" text\The definitions are equal\ lemma ipset_from_cidr_eq_ip_cidr_set: fixes base::"'i::len word" shows "ipset_from_cidr base len = ip_cidr_set base len" proof - have maskshift_eq_not_mask_generic: "((mask len << LENGTH('a) - len) :: 'a::len word) = NOT (mask (LENGTH('a) - len))" using NOT_mask_shifted_lenword by (metis word_not_not) have 1: "mask (len - m) AND base AND NOT (mask (len - m)) = 0" for len m and base::"'i::len word" by(simp add: word_bw_lcs) have 2: "mask (LENGTH('i) - len) AND pfxm_p = 0 \ (a \ ipset_from_netmask pfxm_p (NOT (mask (LENGTH('i) - len)))) \ (pfxm_p = NOT (mask (LENGTH('i) - len)) AND a)" for a::"'i::len word" and pfxm_p apply(subst ipset_from_cidr_alt2[symmetric]) apply(subst zero_base_lsb_imp_set_eq_as_bit_operation) apply(simp; fail) apply(subst ipset_from_cidr_base_wellforemd) apply(simp; fail) apply(simp) done from 2[OF 1, of _ base] have "(x \ ipset_from_netmask base (~~ (mask (LENGTH('i) - len)))) \ (base && ~~ (mask (LENGTH('i) - len)) = x && ~~ (mask (LENGTH('i) - len)))" for x apply(simp add: ipset_from_netmask_base_mask_consume) unfolding word_bw_comms(1)[of _ " ~~ (mask (LENGTH('i) - len))"] by simp then show ?thesis unfolding ip_cidr_set_def ipset_from_cidr_def by(auto simp add: maskshift_eq_not_mask_generic) qed lemma ip_cidr_set_change_base: "j \ ip_cidr_set i r \ ip_cidr_set j r = ip_cidr_set i r" by (auto simp: ip_cidr_set_def) subsection\IP Addresses as WordIntervals\ text\The nice thing is: @{typ "'i wordinterval"}s are executable.\ definition iprange_single :: "'i::len word \ 'i wordinterval" where "iprange_single ip \ WordInterval ip ip" fun iprange_interval :: "('i::len word \ 'i::len word) \ 'i wordinterval" where "iprange_interval (ip_start, ip_end) = WordInterval ip_start ip_end" declare iprange_interval.simps[simp del] lemma iprange_interval_uncurry: "iprange_interval ipcidr = uncurry WordInterval ipcidr" by(cases ipcidr) (simp add: iprange_interval.simps) lemma "wordinterval_to_set (iprange_single ip) = {ip}" by(simp add: iprange_single_def) lemma "wordinterval_to_set (iprange_interval (ip1, ip2)) = {ip1 .. ip2}" by(simp add: iprange_interval.simps) text\Now we can use the set operations on @{typ "'i::len wordinterval"}s\ term wordinterval_to_set term wordinterval_element term wordinterval_union term wordinterval_empty term wordinterval_setminus term wordinterval_UNIV term wordinterval_invert term wordinterval_intersection term wordinterval_subset term wordinterval_eq subsection\IP Addresses in CIDR Notation\ text\We want to convert IP addresses in CIDR notation to intervals. We already have @{const ipset_from_cidr}, which gives back a non-executable set. We want to convert to something we can store in an @{typ "'i wordinterval"}.\ fun ipcidr_to_interval_start :: "('i::len word \ nat) \ 'i::len word" where "ipcidr_to_interval_start (pre, len) = ( let netmask = (mask len) << (LENGTH('i) - len); network_prefix = (pre AND netmask) in network_prefix)" fun ipcidr_to_interval_end :: "('i::len word \ nat) \ 'i::len word" where "ipcidr_to_interval_end (pre, len) = ( let netmask = (mask len) << (LENGTH('i) - len); network_prefix = (pre AND netmask) in network_prefix OR (NOT netmask))" definition ipcidr_to_interval :: "('i::len word \ nat) \ ('i word \ 'i word)" where "ipcidr_to_interval cidr \ (ipcidr_to_interval_start cidr, ipcidr_to_interval_end cidr)" lemma ipset_from_cidr_ipcidr_to_interval: "ipset_from_cidr base len = {ipcidr_to_interval_start (base,len) .. ipcidr_to_interval_end (base,len)}" by(simp add: Let_def ipcidr_to_interval_def ipset_from_cidr_def ipset_from_netmask_def) declare ipcidr_to_interval_start.simps[simp del] ipcidr_to_interval_end.simps[simp del] lemma ipcidr_to_interval: "ipcidr_to_interval (base, len) = (s,e) \ ipset_from_cidr base len = {s .. e}" by (simp add: ipcidr_to_interval_def ipset_from_cidr_ipcidr_to_interval) definition ipcidr_tuple_to_wordinterval :: "('i::len word \ nat) \ 'i wordinterval" where "ipcidr_tuple_to_wordinterval iprng \ iprange_interval (ipcidr_to_interval iprng)" lemma wordinterval_to_set_ipcidr_tuple_to_wordinterval: "wordinterval_to_set (ipcidr_tuple_to_wordinterval (b, m)) = ipset_from_cidr b m" unfolding ipcidr_tuple_to_wordinterval_def ipset_from_cidr_ipcidr_to_interval ipcidr_to_interval_def by(simp add: iprange_interval.simps) lemma wordinterval_to_set_ipcidr_tuple_to_wordinterval_uncurry: "wordinterval_to_set (ipcidr_tuple_to_wordinterval ipcidr) = uncurry ipset_from_cidr ipcidr" by(cases ipcidr, simp add: wordinterval_to_set_ipcidr_tuple_to_wordinterval) definition ipcidr_union_set :: "('i::len word \ nat) set \ ('i word) set" where "ipcidr_union_set ips \ \(base, len) \ ips. ipset_from_cidr base len" lemma ipcidr_union_set_uncurry: "ipcidr_union_set ips = (\ ipcidr \ ips. uncurry ipset_from_cidr ipcidr)" by(simp add: ipcidr_union_set_def uncurry_case_stmt) subsection\Clever Operations on IP Addresses in CIDR Notation\ text\Intersecting two intervals may result in a new interval. Example: \{1..10} \ {5..20} = {5..10}\ Intersecting two IP address ranges represented as CIDR ranges results either in the empty set or the smaller of the two ranges. It will never create a new range. \ context begin (*contributed by Lars Noschinski*) private lemma less_and_not_mask_eq: fixes i :: "('a :: len) word" assumes "r2 \ r1" "i && ~~ (mask r2) = x && ~~ (mask r2)" shows "i && ~~ (mask r1) = x && ~~ (mask r1)" proof - have "i AND NOT (mask r1) = (i && ~~ (mask r2)) && ~~ (mask r1)" (is "_ = ?w && _") using \r2 \ r1\ by (simp add: and_not_mask_twice max_def) also have "?w = x && ~~ (mask r2)" by fact also have "\ && ~~ (mask r1) = x && ~~ (mask r1)" using \r2 \ r1\ by (simp add: and_not_mask_twice max_def) finally show ?thesis . qed lemma ip_cidr_set_less: fixes i :: "'i::len word" shows "r1 \ r2 \ ip_cidr_set i r2 \ ip_cidr_set i r1" unfolding ip_cidr_set_def apply auto apply (rule less_and_not_mask_eq[where ?r2.0="LENGTH('i) - r2"]) apply auto done private lemma ip_cidr_set_intersect_subset_helper: fixes i1 r1 i2 r2 assumes disj: "ip_cidr_set i1 r1 \ ip_cidr_set i2 r2 \ {}" and "r1 \ r2" shows "ip_cidr_set i2 r2 \ ip_cidr_set i1 r1" proof - from disj obtain j where "j \ ip_cidr_set i1 r1" "j \ ip_cidr_set i2 r2" by auto with \r1 \ r2\ have "j \ ip_cidr_set j r1" "j \ ip_cidr_set j r1" using ip_cidr_set_change_base ip_cidr_set_less by blast+ show "ip_cidr_set i2 r2 \ ip_cidr_set i1 r1" proof fix i assume "i \ ip_cidr_set i2 r2" with \j \ ip_cidr_set i2 r2\ have "i \ ip_cidr_set j r2" using ip_cidr_set_change_base by auto also have "ip_cidr_set j r2 \ ip_cidr_set j r1" using \r1 \ r2\ ip_cidr_set_less by blast also have "\ = ip_cidr_set i1 r1" using \j \ ip_cidr_set i1 r1\ ip_cidr_set_change_base by blast finally show "i \ ip_cidr_set i1 r1" . qed qed lemma ip_cidr_set_notsubset_empty_inter: "\ ip_cidr_set i1 r1 \ ip_cidr_set i2 r2 \ \ ip_cidr_set i2 r2 \ ip_cidr_set i1 r1 \ ip_cidr_set i1 r1 \ ip_cidr_set i2 r2 = {}" apply(cases "r1 \ r2") subgoal using ip_cidr_set_intersect_subset_helper by blast apply(cases "r2 \ r1") subgoal using ip_cidr_set_intersect_subset_helper by blast by(simp) end lemma ip_cidr_intersect: "\ ipset_from_cidr b2 m2 \ ipset_from_cidr b1 m1 \ \ ipset_from_cidr b1 m1 \ ipset_from_cidr b2 m2 \ ipset_from_cidr b1 m1 \ ipset_from_cidr b2 m2 = {}" apply(simp add: ipset_from_cidr_eq_ip_cidr_set) using ip_cidr_set_notsubset_empty_inter by blast text\Computing the intersection of two IP address ranges in CIDR notation\ fun ipcidr_conjunct :: "('i::len word \ nat) \ ('i word \ nat) \ ('i word \ nat) option" where "ipcidr_conjunct (base1, m1) (base2, m2) = ( if ipset_from_cidr base1 m1 \ ipset_from_cidr base2 m2 = {} then None else if ipset_from_cidr base1 m1 \ ipset_from_cidr base2 m2 then Some (base1, m1) else Some (base2, m2) )" text\Intersecting with an address with prefix length zero always yields a non-empty result.\ lemma ipcidr_conjunct_any: "ipcidr_conjunct a (x,0) \ None" "ipcidr_conjunct (y,0) b \ None" apply(cases a, simp add: ipset_from_cidr_0 ipset_from_cidr_not_empty) by(cases b, simp add: ipset_from_cidr_0 ipset_from_cidr_not_empty) lemma ipcidr_conjunct_correct: "(case ipcidr_conjunct (b1, m1) (b2, m2) of Some (bx, mx) \ ipset_from_cidr bx mx | None \ {}) = (ipset_from_cidr b1 m1) \ (ipset_from_cidr b2 m2)" apply(simp split: if_split_asm) using ip_cidr_intersect by fast declare ipcidr_conjunct.simps[simp del] subsection\Code Equations\ text\Executable definition using word intervals\ lemma ipcidr_conjunct_word[code_unfold]: "ipcidr_conjunct ips1 ips2 = ( if wordinterval_empty (wordinterval_intersection (ipcidr_tuple_to_wordinterval ips1) (ipcidr_tuple_to_wordinterval ips2)) then None else if wordinterval_subset (ipcidr_tuple_to_wordinterval ips1) (ipcidr_tuple_to_wordinterval ips2) then Some ips1 else Some ips2 )" apply(simp) apply(cases ips1, cases ips2, rename_tac b1 m1 b2 m2, simp) apply(auto simp add: wordinterval_to_set_ipcidr_tuple_to_wordinterval ipcidr_conjunct.simps split: if_split_asm) done (*with the code_unfold lemma before, this works!*) lemma "ipcidr_conjunct (0::32 word,0) (8,1) = Some (8, 1)" by eval export_code ipcidr_conjunct checking SML text\making element check executable\ lemma addr_in_ipset_from_netmask_code[code_unfold]: "addr \ (ipset_from_netmask base netmask) \ (base AND netmask) \ addr \ addr \ (base AND netmask) OR (NOT netmask)" by(simp add: ipset_from_netmask_def Let_def) lemma addr_in_ipset_from_cidr_code[code_unfold]: "(addr::'i::len word) \ (ipset_from_cidr pre len) \ (pre AND ((mask len) << (LENGTH('i) - len))) \ addr \ addr \ pre OR (mask (LENGTH('i) - len))" unfolding ipset_from_cidr_alt by simp + end end diff --git a/thys/IP_Addresses/IPv4.thy b/thys/IP_Addresses/IPv4.thy --- a/thys/IP_Addresses/IPv4.thy +++ b/thys/IP_Addresses/IPv4.thy @@ -1,260 +1,267 @@ (* Title: IPv4.thy Authors: Cornelius Diekmann, Julius Michaelis *) theory IPv4 imports IP_Address NumberWang_IPv4 (* include "HOL-Library.Code_Target_Nat" if you need to work with actual numbers.*) begin lemma take_bit_word_beyond_length_eq: \take_bit n w = w\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that by transfer simp section \IPv4 Adresses\ text\An IPv4 address is basically a 32 bit unsigned integer.\ type_synonym ipv4addr = "32 word" text\Conversion between natural numbers and IPv4 adresses\ definition nat_of_ipv4addr :: "ipv4addr \ nat" where "nat_of_ipv4addr a = unat a" definition ipv4addr_of_nat :: "nat \ ipv4addr" where "ipv4addr_of_nat n = of_nat n" text\The maximum IPv4 addres\ definition max_ipv4_addr :: "ipv4addr" where "max_ipv4_addr \ ipv4addr_of_nat ((2^32) - 1)" lemma max_ipv4_addr_number: "max_ipv4_addr = 4294967295" unfolding max_ipv4_addr_def ipv4addr_of_nat_def by(simp) lemma "max_ipv4_addr = 0b11111111111111111111111111111111" by(fact max_ipv4_addr_number) lemma max_ipv4_addr_max_word: "max_ipv4_addr = - 1" by(simp add: max_ipv4_addr_number) lemma max_ipv4_addr_max[simp]: "\a. a \ max_ipv4_addr" by(simp add: max_ipv4_addr_max_word) lemma UNIV_ipv4addrset: "UNIV = {0 .. max_ipv4_addr}" (*not in the simp set, for a reason*) by(simp add: max_ipv4_addr_max_word) fastforce text\identity functions\ lemma nat_of_ipv4addr_ipv4addr_of_nat_mod: "nat_of_ipv4addr (ipv4addr_of_nat n) = n mod 2^32" by (simp add: ipv4addr_of_nat_def nat_of_ipv4addr_def unat_of_nat take_bit_eq_mod) lemma nat_of_ipv4addr_ipv4addr_of_nat: "\ n \ nat_of_ipv4addr max_ipv4_addr \ \ nat_of_ipv4addr (ipv4addr_of_nat n) = n" by (simp add: nat_of_ipv4addr_ipv4addr_of_nat_mod max_ipv4_addr_def) lemma ipv4addr_of_nat_nat_of_ipv4addr: "ipv4addr_of_nat (nat_of_ipv4addr addr) = addr" by(simp add: ipv4addr_of_nat_def nat_of_ipv4addr_def) -subsection\Representing IPv4 Adresses (Syntax)\ + subsection\Representing IPv4 Adresses (Syntax)\ + + context + includes bit_operations_syntax + begin + fun ipv4addr_of_dotdecimal :: "nat \ nat \ nat \ nat \ ipv4addr" where "ipv4addr_of_dotdecimal (a,b,c,d) = ipv4addr_of_nat (d + 256 * c + 65536 * b + 16777216 * a )" fun dotdecimal_of_ipv4addr :: "ipv4addr \ nat \ nat \ nat \ nat" where "dotdecimal_of_ipv4addr a = (nat_of_ipv4addr ((a >> 24) AND 0xFF), nat_of_ipv4addr ((a >> 16) AND 0xFF), nat_of_ipv4addr ((a >> 8) AND 0xFF), nat_of_ipv4addr (a AND 0xff))" declare ipv4addr_of_dotdecimal.simps[simp del] declare dotdecimal_of_ipv4addr.simps[simp del] text\Examples:\ lemma "ipv4addr_of_dotdecimal (192, 168, 0, 1) = 3232235521" by(simp add: ipv4addr_of_dotdecimal.simps ipv4addr_of_nat_def) (*could be solved by eval, but needs "HOL-Library.Code_Target_Nat"*) lemma "dotdecimal_of_ipv4addr 3232235521 = (192, 168, 0, 1)" by(simp add: dotdecimal_of_ipv4addr.simps nat_of_ipv4addr_def) text\a different notation for @{term ipv4addr_of_dotdecimal}\ lemma ipv4addr_of_dotdecimal_bit: "ipv4addr_of_dotdecimal (a,b,c,d) = (ipv4addr_of_nat a << 24) + (ipv4addr_of_nat b << 16) + (ipv4addr_of_nat c << 8) + ipv4addr_of_nat d" proof - have a: "(ipv4addr_of_nat a) << 24 = ipv4addr_of_nat (a * 16777216)" by(simp add: ipv4addr_of_nat_def shiftl_t2n) have b: "(ipv4addr_of_nat b) << 16 = ipv4addr_of_nat (b * 65536)" by(simp add: ipv4addr_of_nat_def shiftl_t2n) have c: "(ipv4addr_of_nat c) << 8 = ipv4addr_of_nat (c * 256)" by(simp add: ipv4addr_of_nat_def shiftl_t2n) have ipv4addr_of_nat_suc: "\x. ipv4addr_of_nat (Suc x) = word_succ (ipv4addr_of_nat (x))" by(simp add: ipv4addr_of_nat_def, metis Abs_fnat_hom_Suc of_nat_Suc) { fix x y have "ipv4addr_of_nat x + ipv4addr_of_nat y = ipv4addr_of_nat (x+y)" apply(induction x arbitrary: y) apply(simp add: ipv4addr_of_nat_def; fail) by(simp add: ipv4addr_of_nat_suc word_succ_p1) } from this a b c show ?thesis apply(simp add: ipv4addr_of_dotdecimal.simps) apply(rule arg_cong[where f=ipv4addr_of_nat]) apply(thin_tac _)+ by presburger qed lemma size_ipv4addr: "size (x::ipv4addr) = 32" by(simp add:word_size) lemma dotdecimal_of_ipv4addr_ipv4addr_of_dotdecimal: "\ a < 256; b < 256; c < 256; d < 256 \ \ dotdecimal_of_ipv4addr (ipv4addr_of_dotdecimal (a,b,c,d)) = (a,b,c,d)" proof - assume "a < 256" and "b < 256" and "c < 256" and "d < 256" note assms= \a < 256\ \b < 256\ \c < 256\ \d < 256\ hence a: "nat_of_ipv4addr ((ipv4addr_of_nat (d + 256 * c + 65536 * b + 16777216 * a) >> 24) AND mask 8) = a" apply (simp add: ipv4addr_of_nat_def nat_of_ipv4addr_def) apply transfer apply (simp add: drop_bit_take_bit nat_take_bit_eq flip: take_bit_eq_mask) apply (simp add: drop_bit_eq_div take_bit_eq_mod) done have ipv4addr_of_nat_AND_mask8: "(ipv4addr_of_nat a) AND mask 8 = (ipv4addr_of_nat (a mod 256))" for a apply (simp add: ipv4addr_of_nat_def) apply transfer apply (simp flip: take_bit_eq_mask) apply (simp add: take_bit_eq_mod of_nat_mod) done from assms have b: "nat_of_ipv4addr ((ipv4addr_of_nat (d + 256 * c + 65536 * b + 16777216 * a) >> 16) AND mask 8) = b" apply (simp add: ipv4addr_of_nat_def nat_of_ipv4addr_def) apply transfer apply (simp add: drop_bit_take_bit flip: take_bit_eq_mask) using div65536 apply (simp add: drop_bit_eq_div take_bit_eq_mod) done from assms have c: "nat_of_ipv4addr ((ipv4addr_of_nat (d + 256 * c + 65536 * b + 16777216 * a) >> 8) AND mask 8) = c" apply (simp add: ipv4addr_of_nat_def nat_of_ipv4addr_def) apply transfer apply (simp add: drop_bit_take_bit flip: take_bit_eq_mask) using div256 apply (simp add: drop_bit_eq_div take_bit_eq_mod) done from \d < 256\ have d: "nat_of_ipv4addr (ipv4addr_of_nat (d + 256 * c + 65536 * b + 16777216 * a) AND mask 8) = d" apply (simp add: ipv4addr_of_nat_AND_mask8 ipv4addr_of_nat_def nat_of_ipv4addr_def) apply transfer apply (simp flip: take_bit_eq_mask) apply (simp add: take_bit_eq_mod nat_mod_distrib nat_add_distrib nat_mult_distrib mod256) done from a b c d show ?thesis apply (simp add: ipv4addr_of_dotdecimal.simps dotdecimal_of_ipv4addr.simps) apply (simp add: mask_eq push_bit_of_1) done qed lemma ipv4addr_of_dotdecimal_dotdecimal_of_ipv4addr: "(ipv4addr_of_dotdecimal (dotdecimal_of_ipv4addr ip)) = ip" proof - have ip_and_mask8_bl_drop24: "(ip::ipv4addr) AND mask 8 = of_bl (drop 24 (to_bl ip))" by(simp add: of_drop_to_bl size_ipv4addr) have List_rev_drop_geqn: "length x \ n \ (take n (rev x)) = rev (drop (length x - n) x)" for x :: "'a list" and n by(simp add: List.rev_drop) have and_mask_bl_take: "length x \ n \ ((of_bl x) AND mask n) = (of_bl (rev (take n (rev (x)))))" for x n by(simp add: List_rev_drop_geqn of_bl_drop) have ipv4addr_and_255: "x AND 255 = x AND mask 8" for x :: ipv4addr by (simp add: mask_eq push_bit_of_1) have bit_equality: "((ip >> 24) AND 0xFF << 24) + ((ip >> 16) AND 0xFF << 16) + ((ip >> 8) AND 0xFF << 8) + (ip AND 0xFF) = of_bl (take 8 (to_bl ip) @ take 8 (drop 8 (to_bl ip)) @ take 8 (drop 16 (to_bl ip)) @ drop 24 (to_bl ip))" apply(simp add: ipv4addr_and_255) apply(simp add: shiftr_slice) apply(simp add: slice_take' size_ipv4addr and_mask_bl_take flip: push_bit_and) apply(simp add: List_rev_drop_geqn) apply(simp add: drop_take) apply(simp add: shiftl_of_bl) apply(simp add: of_bl_append) apply(simp add: ip_and_mask8_bl_drop24) done have blip_split: "\ blip. length blip = 32 \ blip = (take 8 blip) @ (take 8 (drop 8 blip)) @ (take 8 (drop 16 blip)) @ (take 8 (drop 24 blip))" by(rename_tac blip,case_tac blip,simp_all)+ (*I'm so sorry for this ...*) have "ipv4addr_of_dotdecimal (dotdecimal_of_ipv4addr ip) = of_bl (to_bl ip)" apply (subst blip_split) apply simp apply (simp add: ipv4addr_of_dotdecimal_bit dotdecimal_of_ipv4addr.simps) apply (simp add: ipv4addr_of_nat_nat_of_ipv4addr) apply (simp flip: bit_equality) done thus ?thesis using word_bl.Rep_inverse[symmetric] by simp qed lemma ipv4addr_of_dotdecimal_eqE: "\ ipv4addr_of_dotdecimal (a,b,c,d) = ipv4addr_of_dotdecimal (e,f,g,h); a < 256; b < 256; c < 256; d < 256; e < 256; f < 256; g < 256; h < 256 \ \ a = e \ b = f \ c = g \ d = h" by (metis Pair_inject dotdecimal_of_ipv4addr_ipv4addr_of_dotdecimal) subsection\IP Ranges: Examples\ lemma "(UNIV :: ipv4addr set) = {0 .. max_ipv4_addr}" by(simp add: UNIV_ipv4addrset) lemma "(42::ipv4addr) \ UNIV" by(simp) (*Warning, not executable!*) lemma "ipset_from_netmask (ipv4addr_of_dotdecimal (192,168,0,42)) (ipv4addr_of_dotdecimal (255,255,0,0)) = {ipv4addr_of_dotdecimal (192,168,0,0) .. ipv4addr_of_dotdecimal (192,168,255,255)}" by(simp add: ipset_from_netmask_def ipv4addr_of_dotdecimal.simps ipv4addr_of_nat_def) lemma "ipset_from_netmask (ipv4addr_of_dotdecimal (192,168,0,42)) (ipv4addr_of_dotdecimal (0,0,0,0)) = UNIV" by(simp add: UNIV_ipv4addrset ipset_from_netmask_def ipv4addr_of_dotdecimal.simps ipv4addr_of_nat_def max_ipv4_addr_max_word) text\192.168.0.0/24\ lemma fixes addr :: ipv4addr shows "ipset_from_cidr addr pflength = ipset_from_netmask addr ((mask pflength) << (32 - pflength))" by(simp add: ipset_from_cidr_def) lemma "ipset_from_cidr (ipv4addr_of_dotdecimal (192,168,0,42)) 16 = {ipv4addr_of_dotdecimal (192,168,0,0) .. ipv4addr_of_dotdecimal (192,168,255,255)}" by(simp add: ipset_from_cidr_alt mask_eq ipv4addr_of_dotdecimal.simps ipv4addr_of_nat_def push_bit_of_1) lemma "ip \ (ipset_from_cidr (ipv4addr_of_dotdecimal (0, 0, 0, 0)) 0)" by(simp add: ipset_from_cidr_0) lemma ipv4set_from_cidr_32: fixes addr :: ipv4addr shows "ipset_from_cidr addr 32 = {addr}" by (simp add: ipset_from_cidr_alt take_bit_word_beyond_length_eq flip: take_bit_eq_mask) lemma fixes pre :: ipv4addr shows "ipset_from_cidr pre len = {(pre AND ((mask len) << (32 - len))) .. pre OR (mask (32 - len))}" by (simp add: ipset_from_cidr_alt ipset_from_cidr_def) text\making element check executable\ lemma addr_in_ipv4set_from_netmask_code[code_unfold]: fixes addr :: ipv4addr shows "addr \ (ipset_from_netmask base netmask) \ (base AND netmask) \ addr \ addr \ (base AND netmask) OR (NOT netmask)" by (simp add: addr_in_ipset_from_netmask_code) lemma addr_in_ipv4set_from_cidr_code[code_unfold]: fixes addr :: ipv4addr shows "addr \ (ipset_from_cidr pre len) \ (pre AND ((mask len) << (32 - len))) \ addr \ addr \ pre OR (mask (32 - len))" by(simp add: addr_in_ipset_from_cidr_code) (*small numbers because we didn't load Code_Target_Nat. Should work by eval*) lemma "ipv4addr_of_dotdecimal (192,168,42,8) \ (ipset_from_cidr (ipv4addr_of_dotdecimal (192,168,0,0)) 16)" by (simp add: ipv4addr_of_nat_def ipset_from_cidr_def ipv4addr_of_dotdecimal.simps ipset_from_netmask_def mask_eq_exp_minus_1 word_le_def take_bit_minus_one_eq_mask) definition ipv4range_UNIV :: "32 wordinterval" where "ipv4range_UNIV \ wordinterval_UNIV" lemma ipv4range_UNIV_set_eq: "wordinterval_to_set ipv4range_UNIV = UNIV" by(simp only: ipv4range_UNIV_def wordinterval_UNIV_set_eq) thm iffD1[OF wordinterval_eq_set_eq] (*TODO: probably the following is a good idea?*) (* declare iffD1[OF wordinterval_eq_set_eq, cong] *) text\This \LENGTH('a)\ is 32 for IPv4 addresses.\ lemma ipv4cidr_to_interval_simps[code_unfold]: "ipcidr_to_interval ((pre::ipv4addr), len) = ( let netmask = (mask len) << (32 - len); network_prefix = (pre AND netmask) in (network_prefix, network_prefix OR (NOT netmask)))" by(simp add: ipcidr_to_interval_def Let_def ipcidr_to_interval_start.simps ipcidr_to_interval_end.simps) + end + end diff --git a/thys/IP_Addresses/IPv6.thy b/thys/IP_Addresses/IPv6.thy --- a/thys/IP_Addresses/IPv6.thy +++ b/thys/IP_Addresses/IPv6.thy @@ -1,958 +1,964 @@ (* Title: IPv6.thy Authors: Cornelius Diekmann *) theory IPv6 imports IP_Address NumberWang_IPv6 (* include "HOL-Library.Code_Target_Nat" if you need to work with actual numbers.*) begin section \IPv6 Addresses\ text\An IPv6 address is basically a 128 bit unsigned integer. RFC 4291, Section 2.\ type_synonym ipv6addr = "128 word" text\Conversion between natural numbers and IPv6 adresses\ definition nat_of_ipv6addr :: "ipv6addr \ nat" where "nat_of_ipv6addr a = unat a" definition ipv6addr_of_nat :: "nat \ ipv6addr" where "ipv6addr_of_nat n = of_nat n" -lemma "ipv6addr_of_nat n = word_of_int (int n)" + lemma "ipv6addr_of_nat n = word_of_int (int n)" by(simp add: ipv6addr_of_nat_def) text\The maximum IPv6 address\ definition max_ipv6_addr :: "ipv6addr" where "max_ipv6_addr \ ipv6addr_of_nat ((2^128) - 1)" lemma max_ipv6_addr_number: "max_ipv6_addr = 0xFFFFFFFFFFFFFFFFFFFFFFFFFFFFFFFF" unfolding max_ipv6_addr_def ipv6addr_of_nat_def by(simp) lemma "max_ipv6_addr = 340282366920938463463374607431768211455" by(fact max_ipv6_addr_number) lemma max_ipv6_addr_max_word: "max_ipv6_addr = - 1" by(simp add: max_ipv6_addr_number) lemma max_ipv6_addr_max: "\a. a \ max_ipv6_addr" by(simp add: max_ipv6_addr_max_word) lemma UNIV_ipv6addrset: "UNIV = {0 .. max_ipv6_addr}" (*not in the simp set, for a reason*) by(simp add: max_ipv6_addr_max_word) fastforce text\identity functions\ lemma nat_of_ipv6addr_ipv6addr_of_nat_mod: "nat_of_ipv6addr (ipv6addr_of_nat n) = n mod 2^128" by (simp add: ipv6addr_of_nat_def nat_of_ipv6addr_def unat_of_nat take_bit_eq_mod) lemma nat_of_ipv6addr_ipv6addr_of_nat: "n \ nat_of_ipv6addr max_ipv6_addr \ nat_of_ipv6addr (ipv6addr_of_nat n) = n" by (simp add: nat_of_ipv6addr_ipv6addr_of_nat_mod max_ipv6_addr_def) lemma ipv6addr_of_nat_nat_of_ipv6addr: "ipv6addr_of_nat (nat_of_ipv6addr addr) = addr" by(simp add: ipv6addr_of_nat_def nat_of_ipv6addr_def) -subsection\Syntax of IPv6 Adresses\ + subsection\Syntax of IPv6 Adresses\ text\RFC 4291, Section 2.2.: Text Representation of Addresses\ text\Quoting the RFC (note: errata exists):\ text_raw\ \begin{verbatim} 1. The preferred form is x:x:x:x:x:x:x:x, where the 'x's are one to four hexadecimal digits of the eight 16-bit pieces of the address. Examples: ABCD:EF01:2345:6789:ABCD:EF01:2345:6789 2001:DB8:0:0:8:800:200C:417A \end{verbatim} \ datatype ipv6addr_syntax = IPv6AddrPreferred "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" text_raw\ \begin{verbatim} 2. [...] In order to make writing addresses containing zero bits easier, a special syntax is available to compress the zeros. The use of "::" indicates one or more groups of 16 bits of zeros. The "::" can only appear once in an address. The "::" can also be used to compress leading or trailing zeros in an address. For example, the following addresses 2001:DB8:0:0:8:800:200C:417A a unicast address FF01:0:0:0:0:0:0:101 a multicast address 0:0:0:0:0:0:0:1 the loopback address 0:0:0:0:0:0:0:0 the unspecified address may be represented as 2001:DB8::8:800:200C:417A a unicast address FF01::101 a multicast address ::1 the loopback address :: the unspecified address \end{verbatim} \ (*datatype may take some minutes to load*) datatype ipv6addr_syntax_compressed = \ \using @{typ unit} for the omission @{text "::"}. Naming convention of the datatype: The first number is the position where the omission occurs. The second number is the length of the specified address pieces. I.e. `8 minus the second number' pieces are omitted.\ IPv6AddrCompressed1_0 unit | IPv6AddrCompressed1_1 unit "16 word" | IPv6AddrCompressed1_2 unit "16 word" "16 word" | IPv6AddrCompressed1_3 unit "16 word" "16 word" "16 word" | IPv6AddrCompressed1_4 unit "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed1_5 unit "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed1_6 unit "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed1_7 unit "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed2_1 "16 word" unit | IPv6AddrCompressed2_2 "16 word" unit "16 word" | IPv6AddrCompressed2_3 "16 word" unit "16 word" "16 word" | IPv6AddrCompressed2_4 "16 word" unit "16 word" "16 word" "16 word" | IPv6AddrCompressed2_5 "16 word" unit "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed2_6 "16 word" unit "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed2_7 "16 word" unit "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed3_2 "16 word" "16 word" unit | IPv6AddrCompressed3_3 "16 word" "16 word" unit "16 word" | IPv6AddrCompressed3_4 "16 word" "16 word" unit "16 word" "16 word" | IPv6AddrCompressed3_5 "16 word" "16 word" unit "16 word" "16 word" "16 word" | IPv6AddrCompressed3_6 "16 word" "16 word" unit "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed3_7 "16 word" "16 word" unit "16 word" "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed4_3 "16 word" "16 word" "16 word" unit | IPv6AddrCompressed4_4 "16 word" "16 word" "16 word" unit "16 word" | IPv6AddrCompressed4_5 "16 word" "16 word" "16 word" unit "16 word" "16 word" | IPv6AddrCompressed4_6 "16 word" "16 word" "16 word" unit "16 word" "16 word" "16 word" | IPv6AddrCompressed4_7 "16 word" "16 word" "16 word" unit "16 word" "16 word" "16 word" "16 word" | IPv6AddrCompressed5_4 "16 word" "16 word" "16 word" "16 word" unit | IPv6AddrCompressed5_5 "16 word" "16 word" "16 word" "16 word" unit "16 word" | IPv6AddrCompressed5_6 "16 word" "16 word" "16 word" "16 word" unit "16 word" "16 word" | IPv6AddrCompressed5_7 "16 word" "16 word" "16 word" "16 word" unit "16 word" "16 word" "16 word" | IPv6AddrCompressed6_5 "16 word" "16 word" "16 word" "16 word" "16 word" unit | IPv6AddrCompressed6_6 "16 word" "16 word" "16 word" "16 word" "16 word" unit "16 word" | IPv6AddrCompressed6_7 "16 word" "16 word" "16 word" "16 word" "16 word" unit "16 word" "16 word" | IPv6AddrCompressed7_6 "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" unit | IPv6AddrCompressed7_7 "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" unit "16 word" | IPv6AddrCompressed8_7 "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" "16 word" unit (*RFC 5952: """ 4. A Recommendation for IPv6 Text Representation 4.2.2. Handling One 16-Bit 0 Field The symbol "::" MUST NOT be used to shorten just one 16-bit 0 field. For example, the representation 2001:db8:0:1:1:1:1:1 is correct, but 2001:db8::1:1:1:1:1 is not correct. """ So we could remove all IPv6AddrCompressed*_7 constructors. But these are `recommendations', we might still see these non-recommended definitions. "[...] all implementations must accept and be able to handle any legitimate RFC 4291 format." *) (*More convenient parser helper function for compressed IPv6 addresses: Input list (from parser): Some 16word \ address piece None \ omission '::' Basically, the parser must only do the following (python syntax): split the string which is an ipv6 address at ':' map empty string to None map everything else to Some (string_to_16word str) sanitize empty strings at the start and the end (see toString and parser theories) Example: "1:2:3".split(":") = ['1', '2', '3'] ":2:3:4".split(":") = ['', '2', '3', '4'] ":2::3".split(":") = ['', '2', '', '3'] "1:2:3:".split(":") = ['1', '2', '3', ''] *) definition parse_ipv6_address_compressed :: "((16 word) option) list \ ipv6addr_syntax_compressed option" where "parse_ipv6_address_compressed as = (case as of [None] \ Some (IPv6AddrCompressed1_0 ()) | [None, Some a] \ Some (IPv6AddrCompressed1_1 () a) | [None, Some a, Some b] \ Some (IPv6AddrCompressed1_2 () a b) | [None, Some a, Some b, Some c] \ Some (IPv6AddrCompressed1_3 () a b c) | [None, Some a, Some b, Some c, Some d] \ Some (IPv6AddrCompressed1_4 () a b c d) | [None, Some a, Some b, Some c, Some d, Some e] \ Some (IPv6AddrCompressed1_5 () a b c d e) | [None, Some a, Some b, Some c, Some d, Some e, Some f] \ Some (IPv6AddrCompressed1_6 () a b c d e f) | [None, Some a, Some b, Some c, Some d, Some e, Some f, Some g] \ Some (IPv6AddrCompressed1_7 () a b c d e f g) | [Some a, None] \ Some (IPv6AddrCompressed2_1 a ()) | [Some a, None, Some b] \ Some (IPv6AddrCompressed2_2 a () b) | [Some a, None, Some b, Some c] \ Some (IPv6AddrCompressed2_3 a () b c) | [Some a, None, Some b, Some c, Some d] \ Some (IPv6AddrCompressed2_4 a () b c d) | [Some a, None, Some b, Some c, Some d, Some e] \ Some (IPv6AddrCompressed2_5 a () b c d e) | [Some a, None, Some b, Some c, Some d, Some e, Some f] \ Some (IPv6AddrCompressed2_6 a () b c d e f) | [Some a, None, Some b, Some c, Some d, Some e, Some f, Some g] \ Some (IPv6AddrCompressed2_7 a () b c d e f g) | [Some a, Some b, None] \ Some (IPv6AddrCompressed3_2 a b ()) | [Some a, Some b, None, Some c] \ Some (IPv6AddrCompressed3_3 a b () c) | [Some a, Some b, None, Some c, Some d] \ Some (IPv6AddrCompressed3_4 a b () c d) | [Some a, Some b, None, Some c, Some d, Some e] \ Some (IPv6AddrCompressed3_5 a b () c d e) | [Some a, Some b, None, Some c, Some d, Some e, Some f] \ Some (IPv6AddrCompressed3_6 a b () c d e f) | [Some a, Some b, None, Some c, Some d, Some e, Some f, Some g] \ Some (IPv6AddrCompressed3_7 a b () c d e f g) | [Some a, Some b, Some c, None] \ Some (IPv6AddrCompressed4_3 a b c ()) | [Some a, Some b, Some c, None, Some d] \ Some (IPv6AddrCompressed4_4 a b c () d) | [Some a, Some b, Some c, None, Some d, Some e] \ Some (IPv6AddrCompressed4_5 a b c () d e) | [Some a, Some b, Some c, None, Some d, Some e, Some f] \ Some (IPv6AddrCompressed4_6 a b c () d e f) | [Some a, Some b, Some c, None, Some d, Some e, Some f, Some g] \ Some (IPv6AddrCompressed4_7 a b c () d e f g) | [Some a, Some b, Some c, Some d, None] \ Some (IPv6AddrCompressed5_4 a b c d ()) | [Some a, Some b, Some c, Some d, None, Some e] \ Some (IPv6AddrCompressed5_5 a b c d () e) | [Some a, Some b, Some c, Some d, None, Some e, Some f] \ Some (IPv6AddrCompressed5_6 a b c d () e f) | [Some a, Some b, Some c, Some d, None, Some e, Some f, Some g] \ Some (IPv6AddrCompressed5_7 a b c d () e f g) | [Some a, Some b, Some c, Some d, Some e, None] \ Some (IPv6AddrCompressed6_5 a b c d e ()) | [Some a, Some b, Some c, Some d, Some e, None, Some f] \ Some (IPv6AddrCompressed6_6 a b c d e () f) | [Some a, Some b, Some c, Some d, Some e, None, Some f, Some g] \ Some (IPv6AddrCompressed6_7 a b c d e () f g) | [Some a, Some b, Some c, Some d, Some e, Some f, None] \ Some (IPv6AddrCompressed7_6 a b c d e f ()) | [Some a, Some b, Some c, Some d, Some e, Some f, None, Some g] \ Some (IPv6AddrCompressed7_7 a b c d e f () g) | [Some a, Some b, Some c, Some d, Some e, Some f, Some g, None] \ Some (IPv6AddrCompressed8_7 a b c d e f g ()) | _ \ None \ \invalid ipv6 copressed address.\ )" fun ipv6addr_syntax_compressed_to_list :: "ipv6addr_syntax_compressed \ ((16 word) option) list" where "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_0 _) = [None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_1 () a) = [None, Some a]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_2 () a b) = [None, Some a, Some b]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_3 () a b c) = [None, Some a, Some b, Some c]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_4 () a b c d) = [None, Some a, Some b, Some c, Some d]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_5 () a b c d e) = [None, Some a, Some b, Some c, Some d, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_6 () a b c d e f) = [None, Some a, Some b, Some c, Some d, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed1_7 () a b c d e f g) = [None, Some a, Some b, Some c, Some d, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_1 a ()) = [Some a, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_2 a () b) = [Some a, None, Some b]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_3 a () b c) = [Some a, None, Some b, Some c]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_4 a () b c d) = [Some a, None, Some b, Some c, Some d]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_5 a () b c d e) = [Some a, None, Some b, Some c, Some d, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_6 a () b c d e f) = [Some a, None, Some b, Some c, Some d, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed2_7 a () b c d e f g) = [Some a, None, Some b, Some c, Some d, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_2 a b ()) = [Some a, Some b, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_3 a b () c) = [Some a, Some b, None, Some c]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_4 a b () c d) = [Some a, Some b, None, Some c, Some d]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_5 a b () c d e) = [Some a, Some b, None, Some c, Some d, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_6 a b () c d e f) = [Some a, Some b, None, Some c, Some d, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed3_7 a b () c d e f g) = [Some a, Some b, None, Some c, Some d, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_3 a b c ()) = [Some a, Some b, Some c, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_4 a b c () d) = [Some a, Some b, Some c, None, Some d]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_5 a b c () d e) = [Some a, Some b, Some c, None, Some d, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_6 a b c () d e f) = [Some a, Some b, Some c, None, Some d, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed4_7 a b c () d e f g) = [Some a, Some b, Some c, None, Some d, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed5_4 a b c d ()) = [Some a, Some b, Some c, Some d, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed5_5 a b c d () e) = [Some a, Some b, Some c, Some d, None, Some e]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed5_6 a b c d () e f) = [Some a, Some b, Some c, Some d, None, Some e, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed5_7 a b c d () e f g) = [Some a, Some b, Some c, Some d, None, Some e, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed6_5 a b c d e ()) = [Some a, Some b, Some c, Some d, Some e, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed6_6 a b c d e () f) = [Some a, Some b, Some c, Some d, Some e, None, Some f]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed6_7 a b c d e () f g) = [Some a, Some b, Some c, Some d, Some e, None, Some f, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed7_6 a b c d e f ()) = [Some a, Some b, Some c, Some d, Some e, Some f, None]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed7_7 a b c d e f () g) = [Some a, Some b, Some c, Some d, Some e, Some f, None, Some g]" | "ipv6addr_syntax_compressed_to_list (IPv6AddrCompressed8_7 a b c d e f g ()) = [Some a, Some b, Some c, Some d, Some e, Some f, Some g, None]" (*for all ipv6_syntax, there is a corresponding list representation*) lemma parse_ipv6_address_compressed_exists: obtains ss where "parse_ipv6_address_compressed ss = Some ipv6_syntax" proof define ss where "ss = ipv6addr_syntax_compressed_to_list ipv6_syntax" thus "parse_ipv6_address_compressed ss = Some ipv6_syntax" by (cases ipv6_syntax; simp add: parse_ipv6_address_compressed_def) qed lemma parse_ipv6_address_compressed_identity: "parse_ipv6_address_compressed (ipv6addr_syntax_compressed_to_list (ipv6_syntax)) = Some ipv6_syntax" by(cases ipv6_syntax; simp add: parse_ipv6_address_compressed_def) lemma parse_ipv6_address_compressed_someE: assumes "parse_ipv6_address_compressed as = Some ipv6" obtains "as = [None]" "ipv6 = (IPv6AddrCompressed1_0 ())" | a where "as = [None, Some a]" "ipv6 = (IPv6AddrCompressed1_1 () a)" | a b where "as = [None, Some a, Some b]" "ipv6 = (IPv6AddrCompressed1_2 () a b)" | a b c where "as = [None, Some a, Some b, Some c]" "ipv6 = (IPv6AddrCompressed1_3 () a b c)" | a b c d where "as = [None, Some a, Some b, Some c, Some d]" "ipv6 = (IPv6AddrCompressed1_4 () a b c d)" | a b c d e where "as = [None, Some a, Some b, Some c, Some d, Some e]" "ipv6 = (IPv6AddrCompressed1_5 () a b c d e)" | a b c d e f where "as = [None, Some a, Some b, Some c, Some d, Some e, Some f]" "ipv6 = (IPv6AddrCompressed1_6 () a b c d e f)" | a b c d e f g where "as = [None, Some a, Some b, Some c, Some d, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed1_7 () a b c d e f g)" | a where "as = [Some a, None]" "ipv6 = (IPv6AddrCompressed2_1 a ())" | a b where "as = [Some a, None, Some b]" "ipv6 = (IPv6AddrCompressed2_2 a () b)" | a b c where "as = [Some a, None, Some b, Some c]" "ipv6 = (IPv6AddrCompressed2_3 a () b c)" | a b c d where "as = [Some a, None, Some b, Some c, Some d]" "ipv6 = (IPv6AddrCompressed2_4 a () b c d)" | a b c d e where "as = [Some a, None, Some b, Some c, Some d, Some e]" "ipv6 = (IPv6AddrCompressed2_5 a () b c d e)" | a b c d e f where "as = [Some a, None, Some b, Some c, Some d, Some e, Some f]" "ipv6 = (IPv6AddrCompressed2_6 a () b c d e f)" | a b c d e f g where "as = [Some a, None, Some b, Some c, Some d, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed2_7 a () b c d e f g)" | a b where "as = [Some a, Some b, None]" "ipv6 = (IPv6AddrCompressed3_2 a b ())" | a b c where "as = [Some a, Some b, None, Some c]" "ipv6 = (IPv6AddrCompressed3_3 a b () c)" | a b c d where "as = [Some a, Some b, None, Some c, Some d]" "ipv6 = (IPv6AddrCompressed3_4 a b () c d)" | a b c d e where "as = [Some a, Some b, None, Some c, Some d, Some e]" "ipv6 = (IPv6AddrCompressed3_5 a b () c d e)" | a b c d e f where "as = [Some a, Some b, None, Some c, Some d, Some e, Some f]" "ipv6 = (IPv6AddrCompressed3_6 a b () c d e f)" | a b c d e f g where "as = [Some a, Some b, None, Some c, Some d, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed3_7 a b () c d e f g)" | a b c where "as = [Some a, Some b, Some c, None]" "ipv6 = (IPv6AddrCompressed4_3 a b c ())" | a b c d where "as = [Some a, Some b, Some c, None, Some d]" "ipv6 = (IPv6AddrCompressed4_4 a b c () d)" | a b c d e where "as = [Some a, Some b, Some c, None, Some d, Some e]" "ipv6 = (IPv6AddrCompressed4_5 a b c () d e)" | a b c d e f where "as = [Some a, Some b, Some c, None, Some d, Some e, Some f]" "ipv6 = (IPv6AddrCompressed4_6 a b c () d e f)" | a b c d e f g where "as = [Some a, Some b, Some c, None, Some d, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed4_7 a b c () d e f g)" | a b c d where "as = [Some a, Some b, Some c, Some d, None]" "ipv6 = (IPv6AddrCompressed5_4 a b c d ())" | a b c d e where "as = [Some a, Some b, Some c, Some d, None, Some e]" "ipv6 = (IPv6AddrCompressed5_5 a b c d () e)" | a b c d e f where "as = [Some a, Some b, Some c, Some d, None, Some e, Some f]" "ipv6 = (IPv6AddrCompressed5_6 a b c d () e f)" | a b c d e f g where "as = [Some a, Some b, Some c, Some d, None, Some e, Some f, Some g]" "ipv6 = (IPv6AddrCompressed5_7 a b c d () e f g)" | a b c d e where "as = [Some a, Some b, Some c, Some d, Some e, None]" "ipv6 = (IPv6AddrCompressed6_5 a b c d e ())" | a b c d e f where "as = [Some a, Some b, Some c, Some d, Some e, None, Some f]" "ipv6 = (IPv6AddrCompressed6_6 a b c d e () f)" | a b c d e f g where "as = [Some a, Some b, Some c, Some d, Some e, None, Some f, Some g]" "ipv6 = (IPv6AddrCompressed6_7 a b c d e () f g)" | a b c d e f where "as = [Some a, Some b, Some c, Some d, Some e, Some f, None]" "ipv6 = (IPv6AddrCompressed7_6 a b c d e f ())" | a b c d e f g where "as = [Some a, Some b, Some c, Some d, Some e, Some f, None, Some g]" "ipv6 = (IPv6AddrCompressed7_7 a b c d e f () g)" | a b c d e f g where "as = [Some a, Some b, Some c, Some d, Some e, Some f, Some g, None]" "ipv6 = (IPv6AddrCompressed8_7 a b c d e f g ())" using assms unfolding parse_ipv6_address_compressed_def by (auto split: list.split_asm option.split_asm) (* takes a minute *) lemma parse_ipv6_address_compressed_identity2: "ipv6addr_syntax_compressed_to_list ipv6_syntax = ls \ (parse_ipv6_address_compressed ls) = Some ipv6_syntax" (is "?lhs = ?rhs") proof assume ?rhs thus ?lhs by (auto elim: parse_ipv6_address_compressed_someE) next assume ?lhs thus ?rhs by (cases ipv6_syntax) (auto simp: parse_ipv6_address_compressed_def) qed text\Valid IPv6 compressed notation: \<^item> at most one omission \<^item> at most 7 pieces \ lemma RFC_4291_format: "parse_ipv6_address_compressed as \ None \ length (filter (\p. p = None) as) = 1 \ length (filter (\p. p \ None) as) \ 7" (is "?lhs = ?rhs") proof assume ?lhs then obtain addr where "parse_ipv6_address_compressed as = Some addr" by blast thus ?rhs by (elim parse_ipv6_address_compressed_someE; simp) next assume ?rhs thus ?lhs unfolding parse_ipv6_address_compressed_def by (auto split: option.split list.split if_split_asm) qed text_raw\ \begin{verbatim} 3. An alternative form that is sometimes more convenient when dealing with a mixed environment of IPv4 and IPv6 nodes is x:x:x:x:x:x:d.d.d.d, where the 'x's are the hexadecimal values of the six high-order 16-bit pieces of the address, and the 'd's are the decimal values of the four low-order 8-bit pieces of the address (standard IPv4 representation). Examples: 0:0:0:0:0:0:13.1.68.3 0:0:0:0:0:FFFF:129.144.52.38 or in compressed form: ::13.1.68.3 ::FFFF:129.144.52.38 \end{verbatim} This is currently not supported by our library! \ (*TODO*) (*TODO: oh boy, they can also be compressed*) -subsection\Semantics\ + subsection\Semantics\ + + context + includes bit_operations_syntax + begin + fun ipv6preferred_to_int :: "ipv6addr_syntax \ ipv6addr" where "ipv6preferred_to_int (IPv6AddrPreferred a b c d e f g h) = (ucast a << (16 * 7)) OR (ucast b << (16 * 6)) OR (ucast c << (16 * 5)) OR (ucast d << (16 * 4)) OR (ucast e << (16 * 3)) OR (ucast f << (16 * 2)) OR (ucast g << (16 * 1)) OR (ucast h << (16 * 0))" lemma "ipv6preferred_to_int (IPv6AddrPreferred 0x2001 0xDB8 0x0 0x0 0x8 0x800 0x200C 0x417A) = 42540766411282592856906245548098208122" by eval lemma "ipv6preferred_to_int (IPv6AddrPreferred 0xFF01 0x0 0x0 0x0 0x0 0x0 0x0 0x101) = 338958331222012082418099330867817087233" by eval declare ipv6preferred_to_int.simps[simp del] - definition int_to_ipv6preferred :: "ipv6addr \ ipv6addr_syntax" where "int_to_ipv6preferred i = IPv6AddrPreferred (ucast ((i AND 0xFFFF0000000000000000000000000000) >> 16*7)) (ucast ((i AND 0xFFFF000000000000000000000000) >> 16*6)) (ucast ((i AND 0xFFFF00000000000000000000) >> 16*5)) (ucast ((i AND 0xFFFF0000000000000000) >> 16*4)) (ucast ((i AND 0xFFFF000000000000) >> 16*3)) (ucast ((i AND 0xFFFF00000000) >> 16*2)) (ucast ((i AND 0xFFFF0000) >> 16*1)) (ucast ((i AND 0xFFFF)))" lemma "int_to_ipv6preferred 42540766411282592856906245548098208122 = IPv6AddrPreferred 0x2001 0xDB8 0x0 0x0 0x8 0x800 0x200C 0x417A" by eval lemma word128_masks_ipv6pieces: "(0xFFFF0000000000000000000000000000::ipv6addr) = (mask 16) << 112" "(0xFFFF000000000000000000000000::ipv6addr) = (mask 16) << 96" "(0xFFFF00000000000000000000::ipv6addr) = (mask 16) << 80" "(0xFFFF0000000000000000::ipv6addr) = (mask 16) << 64" "(0xFFFF000000000000::ipv6addr) = (mask 16) << 48" "(0xFFFF00000000::ipv6addr) = (mask 16) << 32" "(0xFFFF0000::ipv6addr) = (mask 16) << 16" "(0xFFFF::ipv6addr) = (mask 16)" by (simp_all add: mask_eq push_bit_of_1) text\Correctness: round trip property one\ lemma ipv6preferred_to_int_int_to_ipv6preferred: "ipv6preferred_to_int (int_to_ipv6preferred ip) = ip" proof - have and_mask_shift_helper: "w AND (mask m << n) >> n << n = w AND (mask m << n)" for m n::nat and w::ipv6addr by (metis is_aligned_shift is_aligned_shiftr_shiftl shiftr_and_eq_shiftl) have ucast_ipv6_piece_rule: "length (dropWhile Not (to_bl w)) \ 16 \ (ucast::16 word \ 128 word) ((ucast::128 word \ 16 word) w) = w" for w::ipv6addr by(rule ucast_short_ucast_long_ingoreLeadingZero) (simp_all) have ucast_ipv6_piece: "16 \ 128 - n \ (ucast::16 word \ 128 word) ((ucast::128 word \ 16 word) (w AND (mask 16 << n) >> n)) << n = w AND (mask 16 << n)" for w::ipv6addr and n::nat apply(subst ucast_ipv6_piece_rule) apply(rule length_drop_mask_inner) apply(simp; fail) apply(subst and_mask_shift_helper) apply simp done have ucast16_ucast128_masks_highest_bits: "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF0000000000000000000000000000 >> 112)) << 112) = (ip AND 0xFFFF0000000000000000000000000000)" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF000000000000000000000000 >> 96)) << 96) = ip AND 0xFFFF000000000000000000000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF00000000000000000000 >> 80)) << 80) = ip AND 0xFFFF00000000000000000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF0000000000000000 >> 64)) << 64) = ip AND 0xFFFF0000000000000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF000000000000 >> 48)) << 48) = ip AND 0xFFFF000000000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF00000000 >> 32)) << 32) = ip AND 0xFFFF00000000" "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF0000 >> 16)) << 16) = ip AND 0xFFFF0000" apply (simp_all add: word128_masks_ipv6pieces ucast_ipv6_piece and_mask2 word_size flip: take_bit_eq_mask) apply (simp_all add: bit_eq_iff) apply (auto simp add: bit_simps) done have ucast16_ucast128_masks_highest_bits0: "(ucast ((ucast::ipv6addr \ 16 word) (ip AND 0xFFFF))) = ip AND 0xFFFF" apply(subst word128_masks_ipv6pieces)+ apply(subst ucast_short_ucast_long_ingoreLeadingZero) apply simp_all by (simp add: length_drop_mask) have mask_len_word:"n = (LENGTH('a)) \ w AND mask n = w" for n and w::"'a::len word" by (simp add: mask_eq_iff) have ipv6addr_16word_pieces_compose_or: "ip && (mask 16 << 112) || ip && (mask 16 << 96) || ip && (mask 16 << 80) || ip && (mask 16 << 64) || ip && (mask 16 << 48) || ip && (mask 16 << 32) || ip && (mask 16 << 16) || ip && mask 16 = ip" apply(subst word_ao_dist2[symmetric])+ apply(simp add: mask_eq push_bit_of_1) apply(subst mask128) apply(rule mask_len_word) apply simp done show ?thesis apply (simp add: ipv6preferred_to_int.simps int_to_ipv6preferred_def) apply (simp add: ucast16_ucast128_masks_highest_bits ucast16_ucast128_masks_highest_bits0) apply (simp add: word128_masks_ipv6pieces ucast_ucast_mask) using ipv6addr_16word_pieces_compose_or apply (simp flip: push_bit_and add: shiftr_and_eq_shiftl) done qed text\Correctness: round trip property two\ lemma int_to_ipv6preferred_ipv6preferred_to_int: "int_to_ipv6preferred (ipv6preferred_to_int ip) = ip" proof - note ucast_shift_simps=helper_masked_ucast_generic helper_masked_ucast_reverse_generic helper_masked_ucast_generic[where n=0, simplified] helper_masked_ucast_equal_generic note ucast_simps=helper_masked_ucast_reverse_generic[where m=0, simplified] helper_masked_ucast_equal_generic[where n=0, simplified] show ?thesis apply (cases ip, rename_tac a b c d e f g h) apply (simp add: ipv6preferred_to_int.simps int_to_ipv6preferred_def) apply (simp add: word128_masks_ipv6pieces) apply (simp add: word_ao_dist ucast_shift_simps ucast_simps) apply (simp add: unsigned_or_eq) apply (simp flip: take_bit_eq_mask add: unsigned_take_bit_eq take_bit_word_eq_self) apply (simp add: shiftl_shiftr1 shiftl_shiftr2) apply (simp flip: take_bit_eq_mask push_bit_and add: word_size) apply (simp add: unsigned_take_bit_eq take_bit_word_eq_self) apply (simp flip: unsigned_take_bit_eq) apply (simp add: unsigned_ucast_eq) apply (simp add: unsigned_push_bit_eq take_bit_word_eq_self) apply (simp flip: ucast_drop_bit_eq) done qed text\compressed to preferred format\ fun ipv6addr_c2p :: "ipv6addr_syntax_compressed \ ipv6addr_syntax" where "ipv6addr_c2p (IPv6AddrCompressed1_0 ()) = IPv6AddrPreferred 0 0 0 0 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed1_1 () h) = IPv6AddrPreferred 0 0 0 0 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed1_2 () g h) = IPv6AddrPreferred 0 0 0 0 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed1_3 () f g h) = IPv6AddrPreferred 0 0 0 0 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed1_4 () e f g h) = IPv6AddrPreferred 0 0 0 0 e f g h" | "ipv6addr_c2p (IPv6AddrCompressed1_5 () d e f g h) = IPv6AddrPreferred 0 0 0 d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed1_6 () c d e f g h) = IPv6AddrPreferred 0 0 c d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed1_7 () b c d e f g h) = IPv6AddrPreferred 0 b c d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed2_1 a ()) = IPv6AddrPreferred a 0 0 0 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed2_2 a () h) = IPv6AddrPreferred a 0 0 0 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed2_3 a () g h) = IPv6AddrPreferred a 0 0 0 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed2_4 a () f g h) = IPv6AddrPreferred a 0 0 0 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed2_5 a () e f g h) = IPv6AddrPreferred a 0 0 0 e f g h" | "ipv6addr_c2p (IPv6AddrCompressed2_6 a () d e f g h) = IPv6AddrPreferred a 0 0 d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed2_7 a () c d e f g h) = IPv6AddrPreferred a 0 c d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed3_2 a b ()) = IPv6AddrPreferred a b 0 0 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed3_3 a b () h) = IPv6AddrPreferred a b 0 0 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed3_4 a b () g h) = IPv6AddrPreferred a b 0 0 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed3_5 a b () f g h) = IPv6AddrPreferred a b 0 0 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed3_6 a b () e f g h) = IPv6AddrPreferred a b 0 0 e f g h" | "ipv6addr_c2p (IPv6AddrCompressed3_7 a b () d e f g h) = IPv6AddrPreferred a b 0 d e f g h" | "ipv6addr_c2p (IPv6AddrCompressed4_3 a b c ()) = IPv6AddrPreferred a b c 0 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed4_4 a b c () h) = IPv6AddrPreferred a b c 0 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed4_5 a b c () g h) = IPv6AddrPreferred a b c 0 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed4_6 a b c () f g h) = IPv6AddrPreferred a b c 0 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed4_7 a b c () e f g h) = IPv6AddrPreferred a b c 0 e f g h" | "ipv6addr_c2p (IPv6AddrCompressed5_4 a b c d ()) = IPv6AddrPreferred a b c d 0 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed5_5 a b c d () h) = IPv6AddrPreferred a b c d 0 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed5_6 a b c d () g h) = IPv6AddrPreferred a b c d 0 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed5_7 a b c d () f g h) = IPv6AddrPreferred a b c d 0 f g h" | "ipv6addr_c2p (IPv6AddrCompressed6_5 a b c d e ()) = IPv6AddrPreferred a b c d e 0 0 0" | "ipv6addr_c2p (IPv6AddrCompressed6_6 a b c d e () h) = IPv6AddrPreferred a b c d e 0 0 h" | "ipv6addr_c2p (IPv6AddrCompressed6_7 a b c d e () g h) = IPv6AddrPreferred a b c d e 0 g h" | "ipv6addr_c2p (IPv6AddrCompressed7_6 a b c d e f ()) = IPv6AddrPreferred a b c d e f 0 0" | "ipv6addr_c2p (IPv6AddrCompressed7_7 a b c d e f () h) = IPv6AddrPreferred a b c d e f 0 h" | "ipv6addr_c2p (IPv6AddrCompressed8_7 a b c d e f g ()) = IPv6AddrPreferred a b c d e f g 0" definition ipv6_unparsed_compressed_to_preferred :: "((16 word) option) list \ ipv6addr_syntax option" where "ipv6_unparsed_compressed_to_preferred ls = ( if length (filter (\p. p = None) ls) \ 1 \ length (filter (\p. p \ None) ls) > 7 then None else let before_omission = map the (takeWhile (\x. x \ None) ls); after_omission = map the (drop 1 (dropWhile (\x. x \ None) ls)); num_omissions = 8 - (length before_omission + length after_omission); expanded = before_omission @ (replicate num_omissions 0) @ after_omission in case expanded of [a,b,c,d,e,f,g,h] \ Some (IPv6AddrPreferred a b c d e f g h) | _ \ None )" lemma "ipv6_unparsed_compressed_to_preferred [Some 0x2001, Some 0xDB8, None, Some 0x8, Some 0x800, Some 0x200C, Some 0x417A] = Some (IPv6AddrPreferred 0x2001 0xDB8 0 0 8 0x800 0x200C 0x417A)" by eval lemma "ipv6_unparsed_compressed_to_preferred [None] = Some (IPv6AddrPreferred 0 0 0 0 0 0 0 0)" by eval lemma "ipv6_unparsed_compressed_to_preferred [] = None" by eval lemma ipv6_unparsed_compressed_to_preferred_identity1: "ipv6_unparsed_compressed_to_preferred (ipv6addr_syntax_compressed_to_list ipv6compressed) = Some ipv6prferred \ ipv6addr_c2p ipv6compressed = ipv6prferred" by (cases ipv6compressed) (simp_all add: ipv6_unparsed_compressed_to_preferred_def numeral_eq_Suc) (*1s*) lemma ipv6_unparsed_compressed_to_preferred_identity2: "ipv6_unparsed_compressed_to_preferred ls = Some ipv6prferred \ (\ipv6compressed. parse_ipv6_address_compressed ls = Some ipv6compressed \ ipv6addr_c2p ipv6compressed = ipv6prferred)" apply(rule iffI) apply(subgoal_tac "parse_ipv6_address_compressed ls \ None") prefer 2 apply(subst RFC_4291_format) apply(simp add: ipv6_unparsed_compressed_to_preferred_def split: if_split_asm; fail) apply(simp) apply(erule exE, rename_tac ipv6compressed) apply(rule_tac x="ipv6compressed" in exI) apply(simp) apply(subgoal_tac "(ipv6addr_syntax_compressed_to_list ipv6compressed = ls)") prefer 2 using parse_ipv6_address_compressed_identity2 apply presburger using ipv6_unparsed_compressed_to_preferred_identity1 apply blast apply(erule exE, rename_tac ipv6compressed) apply(subgoal_tac "(ipv6addr_syntax_compressed_to_list ipv6compressed = ls)") prefer 2 using parse_ipv6_address_compressed_identity2 apply presburger using ipv6_unparsed_compressed_to_preferred_identity1 apply blast done - + + end subsection\IPv6 Pretty Printing (converting to compressed format)\ text_raw\ RFC5952: \begin{verbatim} 4. A Recommendation for IPv6 Text Representation A recommendation for a canonical text representation format of IPv6 addresses is presented in this section. The recommendation in this document is one that complies fully with [RFC4291], is implemented by various operating systems, and is human friendly. The recommendation in this section SHOULD be followed by systems when generating an address to be represented as text, but all implementations MUST accept and be able to handle any legitimate [RFC4291] format. It is advised that humans also follow these recommendations when spelling an address. 4.1. Handling Leading Zeros in a 16-Bit Field Leading zeros MUST be suppressed. For example, 2001:0db8::0001 is not acceptable and must be represented as 2001:db8::1. A single 16- bit 0000 field MUST be represented as 0. 4.2. "::" Usage 4.2.1. Shorten as Much as Possible The use of the symbol "::" MUST be used to its maximum capability. For example, 2001:db8:0:0:0:0:2:1 must be shortened to 2001:db8::2:1. Likewise, 2001:db8::0:1 is not acceptable, because the symbol "::" could have been used to produce a shorter representation 2001:db8::1. 4.2.2. Handling One 16-Bit 0 Field The symbol "::" MUST NOT be used to shorten just one 16-bit 0 field. For example, the representation 2001:db8:0:1:1:1:1:1 is correct, but 2001:db8::1:1:1:1:1 is not correct. 4.2.3. Choice in Placement of "::" When there is an alternative choice in the placement of a "::", the longest run of consecutive 16-bit 0 fields MUST be shortened (i.e., the sequence with three consecutive zero fields is shortened in 2001: 0:0:1:0:0:0:1). When the length of the consecutive 16-bit 0 fields are equal (i.e., 2001:db8:0:0:1:0:0:1), the first sequence of zero bits MUST be shortened. For example, 2001:db8::1:0:0:1 is correct representation. 4.3. Lowercase The characters "a", "b", "c", "d", "e", and "f" in an IPv6 address MUST be represented in lowercase. \end{verbatim} \ text\See @{file \IP_Address_toString.thy\} for examples and test cases.\ context begin private function goup_by_zeros :: "16 word list \ 16 word list list" where "goup_by_zeros [] = []" | "goup_by_zeros (x#xs) = ( if x = 0 then takeWhile (\x. x = 0) (x#xs) # (goup_by_zeros (dropWhile (\x. x = 0) xs)) else [x]#(goup_by_zeros xs))" by(pat_completeness, auto) termination goup_by_zeros apply(relation "measure (\xs. length xs)") apply(simp_all) by (simp add: le_imp_less_Suc length_dropWhile_le) private lemma "goup_by_zeros [0,1,2,3,0,0,0,0,3,4,0,0,0,2,0,0,2,0,3,0] = [[0], [1], [2], [3], [0, 0, 0, 0], [3], [4], [0, 0, 0], [2], [0, 0], [2], [0], [3], [0]]" by eval private lemma "concat (goup_by_zeros ls) = ls" by(induction ls rule:goup_by_zeros.induct) simp+ private lemma "[] \ set (goup_by_zeros ls)" by(induction ls rule:goup_by_zeros.induct) simp+ private primrec List_replace1 :: "'a \ 'a \ 'a list \ 'a list" where "List_replace1 _ _ [] = []" | "List_replace1 a b (x#xs) = (if a = x then b#xs else x#List_replace1 a b xs)" private lemma "List_replace1 a a ls = ls" by(induction ls) simp_all private lemma "a \ set ls \ List_replace1 a b ls = ls" by(induction ls) simp_all private lemma "a \ set ls \ b \ set (List_replace1 a b ls)" apply(induction ls) apply(simp) apply(simp) by blast private fun List_explode :: "'a list list \ ('a option) list" where "List_explode [] = []" | "List_explode ([]#xs) = None#List_explode xs" | "List_explode (xs1#xs2) = map Some xs1@List_explode xs2" private lemma "List_explode [[0::int], [2,3], [], [3,4]] = [Some 0, Some 2, Some 3, None, Some 3, Some 4]" by eval private lemma List_explode_def: "List_explode xss = concat (map (\xs. if xs = [] then [None] else map Some xs) xss)" by(induction xss rule: List_explode.induct) simp+ private lemma List_explode_no_empty: "[] \ set xss \ List_explode xss = map Some (concat xss)" by(induction xss rule: List_explode.induct) simp+ private lemma List_explode_replace1: "[] \ set xss \ foo \ set xss \ List_explode (List_replace1 foo [] xss) = map Some (concat (takeWhile (\xs. xs \ foo) xss)) @ [None] @ map Some (concat (tl (dropWhile (\xs. xs \ foo) xss)))" apply(induction xss rule: List_explode.induct) apply(simp; fail) apply(simp; fail) apply(simp) apply safe apply(simp_all add: List_explode_no_empty) done fun ipv6_preferred_to_compressed :: "ipv6addr_syntax \ ((16 word) option) list" where "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = ( let lss = goup_by_zeros [a,b,c,d,e,f,g,h]; max_zero_seq = foldr (\xs. max (length xs)) lss 0; shortened = if max_zero_seq > 1 then List_replace1 (replicate max_zero_seq 0) [] lss else lss in List_explode shortened )" declare ipv6_preferred_to_compressed.simps[simp del] private lemma foldr_max_length: "foldr (\xs. max (length xs)) lss n = fold max (map length lss) n" apply(subst List.foldr_fold) apply fastforce apply(induction lss arbitrary: n) apply(simp; fail) apply(simp) done private lemma List_explode_goup_by_zeros: "List_explode (goup_by_zeros xs) = map Some xs" apply(induction xs rule: goup_by_zeros.induct) apply(simp; fail) apply(simp) apply(safe) apply(simp) by (metis map_append takeWhile_dropWhile_id) private definition "max_zero_streak xs \ foldr (\xs. max (length xs)) (goup_by_zeros xs) 0" private lemma max_zero_streak_def2: "max_zero_streak xs = fold max (map length (goup_by_zeros xs)) 0" unfolding max_zero_streak_def by(simp add: foldr_max_length) private lemma ipv6_preferred_to_compressed_pull_out_if: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = ( if max_zero_streak [a,b,c,d,e,f,g,h] > 1 then List_explode (List_replace1 (replicate (max_zero_streak [a,b,c,d,e,f,g,h]) 0) [] (goup_by_zeros [a,b,c,d,e,f,g,h])) else map Some [a,b,c,d,e,f,g,h] )" apply (simp only: ipv6_preferred_to_compressed.simps Let_def max_zero_streak_def List_explode_goup_by_zeros) using List_explode_goup_by_zeros by presburger private lemma "ipv6_preferred_to_compressed (IPv6AddrPreferred 0 0 0 0 0 0 0 0) = [None]" by eval private lemma "ipv6_preferred_to_compressed (IPv6AddrPreferred 0x2001 0xDB8 0 0 8 0x800 0x200C 0x417A) = [Some 0x2001, Some 0xDB8, None, Some 8, Some 0x800, Some 0x200C, Some 0x417A]" by eval private lemma "ipv6_preferred_to_compressed (IPv6AddrPreferred 0x2001 0xDB8 0 3 8 0x800 0x200C 0x417A) = [Some 0x2001, Some 0xDB8, Some 0, Some 3, Some 8, Some 0x800, Some 0x200C, Some 0x417A]" by eval (*the output should even conform to RFC5952, ...*) lemma ipv6_preferred_to_compressed_RFC_4291_format: "ipv6_preferred_to_compressed ip = as \ length (filter (\p. p = None) as) = 0 \ length as = 8 \ length (filter (\p. p = None) as) = 1 \ length (filter (\p. p \ None) as) \ 7" apply(cases ip) apply(simp add: ipv6_preferred_to_compressed_pull_out_if) apply(simp only: split: if_split_asm) subgoal for a b c d e f g h apply(rule disjI2) apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(auto simp add: max_zero_streak_def) (*1min*) subgoal apply(rule disjI1) apply(simp) by force done \ \Idea for the following proof:\ private lemma "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = None#xs \ xs = map Some (dropWhile (\x. x=0) [a,b,c,d,e,f,g,h])" apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*20s*) lemma ipv6_preferred_to_compressed: assumes "ipv6_unparsed_compressed_to_preferred (ipv6_preferred_to_compressed ip) = Some ip'" shows "ip = ip'" proof - from assms have 1: "\ipv6compressed. parse_ipv6_address_compressed (ipv6_preferred_to_compressed ip) = Some ipv6compressed \ ipv6addr_c2p ipv6compressed = ip'" using ipv6_unparsed_compressed_to_preferred_identity2 by simp obtain a b c d e f g h where ip: "ip = IPv6AddrPreferred a b c d e f g h" by(cases ip) have ipv6_preferred_to_compressed_None1: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = None#xs \ (map Some (dropWhile (\x. x=0) [a,b,c,d,e,f,g,h]) = xs \ (IPv6AddrPreferred a b c d e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None2: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#None#xs \ (map Some (dropWhile (\x. x=0) [b,c,d,e,f,g,h]) = xs \ (IPv6AddrPreferred a' b c d e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None3: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#None#xs \ (map Some (dropWhile (\x. x=0) [c,d,e,f,g,h]) = xs \ (IPv6AddrPreferred a' b' c d e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None4: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#None#xs \ (map Some (dropWhile (\x. x=0) [d,e,f,g,h]) = xs \ (IPv6AddrPreferred a' b' c' d e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None5: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#(Some d')#None#xs \ (map Some (dropWhile (\x. x=0) [e,f,g,h]) = xs \ (IPv6AddrPreferred a' b' c' d' e f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' d' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None6: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#(Some d')#(Some e')#None#xs \ (map Some (dropWhile (\x. x=0) [f,g,h]) = xs \ (IPv6AddrPreferred a' b' c' d' e' f g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' d' e' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None7: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#(Some d')#(Some e')#(Some f')#None#xs \ (map Some (dropWhile (\x. x=0) [g,h]) = xs \ (IPv6AddrPreferred a' b' c' d' e' f' g h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' d' e' f' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have ipv6_preferred_to_compressed_None8: "ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h) = (Some a')#(Some b')#(Some c')#(Some d')#(Some e')#(Some f')#(Some g')#None#xs \ (map Some (dropWhile (\x. x=0) [h]) = xs \ (IPv6AddrPreferred a' b' c' d' e' f' g' h) = ip') \ (IPv6AddrPreferred a b c d e f g h) = ip'" for xs a' b' c' d' e' f' g' apply(case_tac "a=0",case_tac [!] "b=0",case_tac [!] "c=0",case_tac [!] "d=0", case_tac [!] "e=0",case_tac [!] "f=0",case_tac [!] "g=0",case_tac [!] "h=0") by(simp_all add: ipv6_preferred_to_compressed_pull_out_if max_zero_streak_def) (*5s*) have 2: "parse_ipv6_address_compressed (ipv6_preferred_to_compressed (IPv6AddrPreferred a b c d e f g h)) = Some ipv6compressed \ ipv6addr_c2p ipv6compressed = ip' \ IPv6AddrPreferred a b c d e f g h = ip'" for ipv6compressed apply(erule parse_ipv6_address_compressed_someE) apply(simp_all) apply(erule ipv6_preferred_to_compressed_None1, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None2, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None3, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None4, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None5, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None6, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None7, simp split: if_split_asm)+ apply(erule ipv6_preferred_to_compressed_None8, simp split: if_split_asm) done from 1 2 ip show ?thesis by(elim exE conjE, simp) qed + end end diff --git a/thys/IP_Addresses/NumberWang_IPv6.thy b/thys/IP_Addresses/NumberWang_IPv6.thy --- a/thys/IP_Addresses/NumberWang_IPv6.thy +++ b/thys/IP_Addresses/NumberWang_IPv6.thy @@ -1,230 +1,236 @@ theory NumberWang_IPv6 imports Word_Lib.Word_Lemmas Word_Lib.Word_Syntax Word_Lib.Reversed_Bit_Lists begin section\Helper Lemmas for Low-Level Operations on Machine Words\ text\Needed for IPv6 Syntax\ lemma length_drop_bl: "length (dropWhile Not (to_bl (of_bl bs))) \ length bs" proof - have length_takeWhile_Not_replicate_False: "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)" for n ls by(subst takeWhile_append2) simp+ show ?thesis by(simp add: word_rep_drop dropWhile_eq_drop length_takeWhile_Not_replicate_False) qed lemma bl_drop_leading_zeros: "(of_bl:: bool list \ 'a::len word) (dropWhile Not bs) = (of_bl:: bool list \ 'a::len word) bs" by(induction bs) simp_all lemma bl_length_drop_bound: assumes "length (dropWhile Not bs) \ n" shows "length (dropWhile Not (to_bl ((of_bl:: bool list \ 'a::len word) bs))) \ n" proof - have bl_length_drop_twice: "length (dropWhile Not (to_bl ((of_bl:: bool list \ 'a::len word) (dropWhile Not bs)))) = length (dropWhile Not (to_bl ((of_bl:: bool list \ 'a::len word) bs)))" by(simp add: bl_drop_leading_zeros) from length_drop_bl have *: "length (dropWhile Not (to_bl ((of_bl:: bool list \ 'a::len word) bs))) \ length (dropWhile Not bs)" apply(rule dual_order.trans) apply(subst bl_length_drop_twice) .. show ?thesis apply(rule order.trans, rule *) using assms by(simp) qed +context + includes bit_operations_syntax +begin + lemma length_drop_mask_outer: fixes ip::"'a::len word" shows "LENGTH('a) - n' = len \ length (dropWhile Not (to_bl (ip AND (mask n << n') >> n'))) \ len" apply(subst word_and_mask_shiftl) apply(subst shiftl_shiftr1) apply(simp; fail) apply(simp) apply(subst and_mask) apply(simp add: word_size) apply(simp add: length_drop_mask) done + lemma length_drop_mask_inner: fixes ip::"'a::len word" shows "n \ LENGTH('a) - n' \ length (dropWhile Not (to_bl (ip AND (mask n << n') >> n'))) \ n" apply(subst word_and_mask_shiftl) apply(subst shiftl_shiftr3) apply(simp; fail) apply(simp) apply(simp add: word_size) apply(simp add: mask_twice) apply(simp add: length_drop_mask) done +end lemma mask128: "0xFFFFFFFFFFFFFFFFFFFFFFFFFFFFFFFF = (mask 128 :: 'a::len word)" by (simp add: mask_eq push_bit_of_1) (*-------------- things for ipv6 syntax round trip property two ------------------*) (*n small, m large*) lemma helper_masked_ucast_generic: fixes b::"16 word" assumes "n + 16 \ m" and "m < 128" shows "((ucast:: 16 word \ 128 word) b << n) && (mask 16 << m) = 0" proof - have "x < 2 ^ (m - n)" if mnh2: "x < 0x10000" for x::"128 word" proof - from assms(1) have mnh3: "16 \ m - n" by fastforce have power_2_16_nat: "(16::nat) \ n \ (65535::nat) < 2 ^ n" if a:"16 \ n"for n proof - have power2_rule: "a \ b \ (2::nat)^a \ 2 ^ b" for a b by fastforce show ?thesis apply(subgoal_tac "65536 \ 2 ^ n") apply(subst Nat.less_eq_Suc_le) apply(simp; fail) apply(subgoal_tac "(65536::nat) = 2^16") subgoal using power2_rule \16 \ n\ by presburger by(simp) qed have "65536 = unat (65536::128 word)" by auto moreover from mnh2 have "unat x < unat (65536::128 word)" by(rule Word.unat_mono) ultimately have x: "unat x < 65536" by simp with mnh3 have "unat x < 2 ^ (m - n)" using power_2_16_nat [of \m - n\] by simp with assms(2) show ?thesis by(subst word_less_nat_alt) simp qed hence mnhelper2: "(of_bl::bool list \ 128 word) (to_bl b) < 2 ^ (m - n)" apply(subgoal_tac "(of_bl::bool list \ 128 word) (to_bl b) < 2^(LENGTH(16))") apply(simp; fail) by(rule of_bl_length_less) simp+ have mnhelper3: "(of_bl::bool list \ 128 word) (to_bl b) * 2 ^ n < 2 ^ m" apply(rule div_lt_mult) apply(rule word_less_two_pow_divI) using assms by(simp_all add: mnhelper2 p2_gt_0) from assms show ?thesis apply(subst ucast_bl)+ apply(subst shiftl_of_bl) apply(subst of_bl_append) apply simp apply(subst word_and_mask_shiftl) apply(subst shiftr_div_2n_w) subgoal by(simp add: word_size; fail) apply(subst word_div_less) subgoal by(rule mnhelper3) apply simp done qed lemma unat_of_bl_128_16_less_helper: fixes b::"16 word" shows "unat ((of_bl::bool list \ 128 word) (to_bl b)) < 2^16" proof - from word_bl_Rep' have 1: "length (to_bl b) = 16" by simp have "unat ((of_bl::bool list \ 128 word) (to_bl b)) < 2^(length (to_bl b))" by(fact unat_of_bl_length) with 1 show ?thesis by auto qed lemma unat_of_bl_128_16_le_helper: "unat ((of_bl:: bool list \ 128 word) (to_bl (b::16 word))) \ 65535" proof - from unat_of_bl_128_16_less_helper[of b] have "unat ((of_bl:: bool list \ 128 word) (to_bl b)) < 65536" by simp from Suc_leI[OF this] show ?thesis by simp qed (*reverse*) lemma helper_masked_ucast_reverse_generic: fixes b::"16 word" assumes "m + 16 \ n" and "n \ 128 - 16" shows "((ucast:: 16 word \ 128 word) b << n) && (mask 16 << m) = 0" proof - have power_less_128_helper: "2 ^ n * unat ((of_bl::bool list \ 128 word) (to_bl b)) < 2 ^ LENGTH(128)" if n: "n \ 128 - 16" for n proof - have help_mult: "n \ l \ 2 ^ n * x < 2 ^ l \ x < 2 ^ (l - n)" for x::nat and l by (simp add: nat_mult_power_less_eq semiring_normalization_rules(7)) from n show ?thesis apply(subst help_mult) subgoal by (simp) apply(rule order_less_le_trans) apply(rule unat_of_bl_128_16_less_helper) apply(rule Power.power_increasing) apply(simp_all) done qed have *: "2 ^ m * (2 ^ (n - m) * unat ((of_bl::bool list \ 128 word) (to_bl b))) = 2 ^ n * unat ((of_bl::bool list \ 128 word) (to_bl b))" proof(cases "unat ((of_bl::bool list \ 128 word) (to_bl b)) = 0") case True thus ?thesis by simp next case False have help_mult: "x \ 0 \ b * (c * x) = a * (x::nat) \ b * c = a" for x a b c by simp from assms show ?thesis apply(subst help_mult[OF False]) apply(subst Power.monoid_mult_class.power_add[symmetric]) apply(simp) done qed from assms have "unat ((2 ^ n)::128 word) * unat ((of_bl::bool list \ 128 word) (to_bl b)) mod 2 ^ LENGTH(128) = 2 ^ m * (2 ^ (n - m) * unat ((of_bl::bool list \ 128 word) (to_bl b)) mod 2 ^ LENGTH(128))" apply(subst nat_mod_eq') subgoal apply(subst unat_power_lower) subgoal by(simp; fail) subgoal by (rule power_less_128_helper) simp done apply(subst nat_mod_eq') subgoal by(rule power_less_128_helper) simp apply(subst unat_power_lower) apply(simp; fail) apply(simp only: *) done hence ex_k: "\k. unat ((2 ^ n)::128 word) * unat ((of_bl::bool list \ 128 word) (to_bl b)) mod 2 ^ LENGTH(128) = 2 ^ m * k" by blast hence aligned: "is_aligned ((of_bl::bool list \ 128 word) (to_bl b) << n) m" unfolding is_aligned_iff_dvd_nat unfolding dvd_def unfolding shiftl_t2n unfolding Word.unat_word_ariths(2) by assumption from assms have of_bl_to_bl_shift_mask: "((of_bl::bool list \ 128 word) (to_bl b) << n) && mask (16 + m) = 0" using is_aligned_mask is_aligned_shiftl by force (*sledgehammer*) show ?thesis apply(subst ucast_bl)+ apply(subst word_and_mask_shiftl) apply(subst aligned_shiftr_mask_shiftl) subgoal by (fact aligned) subgoal by (fact of_bl_to_bl_shift_mask) done qed lemma helper_masked_ucast_equal_generic: fixes b::"16 word" assumes "n \ 128 - 16" shows "ucast (((ucast:: 16 word \ 128 word) b << n) && (mask 16 << n) >> n) = b" proof - have ucast_mask: "(ucast:: 16 word \ 128 word) b && mask 16 = ucast b" by transfer (simp flip: take_bit_eq_mask) from assms have "ucast (((ucast:: 16 word \ 128 word) b && mask (128 - n) && mask 16) && mask (128 - n)) = b" by (auto simp add: bit_simps word_size intro!: bit_word_eqI) thus ?thesis apply(subst word_and_mask_shiftl) apply(subst shiftl_shiftr3) apply(simp; fail) apply(simp) apply(subst shiftl_shiftr3) apply(simp_all add: word_size and.assoc) done qed end diff --git a/thys/IP_Addresses/Prefix_Match.thy b/thys/IP_Addresses/Prefix_Match.thy --- a/thys/IP_Addresses/Prefix_Match.thy +++ b/thys/IP_Addresses/Prefix_Match.thy @@ -1,235 +1,242 @@ (* Title: Prefix_Match.thy Authors: Julius Michaelis, Cornelius Diekmann *) theory Prefix_Match imports IP_Address begin section\Prefix Match\ text\ The main difference between the prefix match defined here and CIDR notation is a validity constraint imposed on prefix matches. For example, 192.168.42.42/16 is valid CIDR notation whereas for a prefix match, it must be 192.168.0.0/16. I.e. the last bits of the prefix must be set to zero. \ context notes [[typedef_overloaded]] begin datatype 'a prefix_match = PrefixMatch (pfxm_prefix: "'a::len word") (pfxm_length: nat) end definition pfxm_mask :: "'a prefix_match \ 'a::len word" where "pfxm_mask x \ mask (len_of (TYPE('a)) - pfxm_length x)" +context + includes bit_operations_syntax +begin + definition valid_prefix :: "('a::len) prefix_match \ bool" where "valid_prefix pf = ((pfxm_mask pf) AND pfxm_prefix pf = 0)" text\Note that @{const valid_prefix} looks very elegant as a definition. However, it hides something nasty:\ lemma "valid_prefix (PrefixMatch (0::32 word) 42)" by eval text\When zeroing all least significant bits which exceed the @{const pfxm_length}, you get a @{const valid_prefix}\ lemma mk_valid_prefix: fixes base::"'a::len word" shows "valid_prefix (PrefixMatch (base AND NOT (mask (len_of TYPE ('a) - len))) len)" proof - have "mask (len - m) AND base AND NOT (mask (len - m)) = 0" for m len and base::"'a::len word" by(simp add: word_bw_lcs) thus ?thesis by(simp add: valid_prefix_def pfxm_mask_def pfxm_length_def pfxm_prefix_def) qed +end text\The type @{typ "'a prefix_match"} usually requires @{const valid_prefix}. When we allow working on arbitrary IPs in CIDR notation, we will use the type @{typ "('i::len word \ nat)"} directly.\ lemma valid_prefix_00: "valid_prefix (PrefixMatch 0 0)" by (simp add: valid_prefix_def) definition prefix_match_to_CIDR :: "('i::len) prefix_match \ ('i word \ nat)" where "prefix_match_to_CIDR pfx \ (pfxm_prefix pfx, pfxm_length pfx)" lemma prefix_match_to_CIDR_def2: "prefix_match_to_CIDR = (\pfx. (pfxm_prefix pfx, pfxm_length pfx))" unfolding prefix_match_to_CIDR_def fun_eq_iff by simp definition "prefix_match_dtor m \ (case m of PrefixMatch p l \ (p,l))" text\Some more or less random linear order on prefixes. Only used for serialization at the time of this writing.\ instantiation prefix_match :: (len) linorder begin definition "a \ b \ (if pfxm_length a = pfxm_length b then pfxm_prefix a \ pfxm_prefix b else pfxm_length a > pfxm_length b)" definition "a < b \ (a \ b \ (if pfxm_length a = pfxm_length b then pfxm_prefix a \ pfxm_prefix b else pfxm_length a > pfxm_length b))" instance by standard (auto simp: less_eq_prefix_match_def less_prefix_match_def prefix_match.expand split: if_splits) end lemma "sorted_list_of_set {PrefixMatch 0 32 :: 32 prefix_match, PrefixMatch 42 32, PrefixMatch 0 0, PrefixMatch 0 1, PrefixMatch 12 31} = [PrefixMatch 0 32, PrefixMatch 0x2A 32, PrefixMatch 0xC 31, PrefixMatch 0 1, PrefixMatch 0 0]" by eval context + includes bit_operations_syntax begin private lemma valid_prefix_E: "valid_prefix pf \ ((pfxm_mask pf) AND pfxm_prefix pf = 0)" unfolding valid_prefix_def . private lemma valid_prefix_alt: fixes p::"'a::len prefix_match" shows "valid_prefix p = (pfxm_prefix p AND (2 ^ ((len_of TYPE ('a)) - pfxm_length p) - 1) = 0)" unfolding valid_prefix_def unfolding mask_eq using word_bw_comms(1) arg_cong[where f = "\x. (pfxm_prefix p AND x - 1 = 0)"] shiftl_1 unfolding pfxm_prefix_def pfxm_mask_def mask_eq apply (cases p) apply (simp add: ac_simps push_bit_of_1) done subsection\Address Semantics\ text\Matching on a @{typ "'a::len prefix_match"}. Think of routing tables.\ definition prefix_match_semantics where "prefix_match_semantics m a \ pfxm_prefix m = NOT (pfxm_mask m) AND a" lemma same_length_prefixes_distinct: "valid_prefix pfx1 \ valid_prefix pfx2 \ pfx1 \ pfx2 \ pfxm_length pfx1 = pfxm_length pfx2 \ prefix_match_semantics pfx1 w \ prefix_match_semantics pfx2 w \ False" by (simp add: pfxm_mask_def prefix_match.expand prefix_match_semantics_def) subsection\Relation between prefix and set\ definition prefix_to_wordset :: "'a::len prefix_match \ 'a word set" where "prefix_to_wordset pfx = {pfxm_prefix pfx .. pfxm_prefix pfx OR pfxm_mask pfx}" private lemma pfx_not_empty: "valid_prefix pfx \ prefix_to_wordset pfx \ {}" unfolding valid_prefix_def prefix_to_wordset_def by(simp add: le_word_or2) lemma zero_prefix_match_all: "valid_prefix m \ pfxm_length m = 0 \ prefix_match_semantics m ip" by(simp add: pfxm_mask_def mask_2pm1 valid_prefix_alt prefix_match_semantics_def) lemma prefix_to_wordset_subset_ipset_from_cidr: "prefix_to_wordset pfx \ ipset_from_cidr (pfxm_prefix pfx) (pfxm_length pfx)" apply(rule subsetI) apply(simp add: prefix_to_wordset_def addr_in_ipset_from_cidr_code) + apply(intro impI conjI) apply (metis (erased, opaque_lifting) order_trans word_and_le2) apply(simp add: pfxm_mask_def) done subsection\Equivalence Proofs\ theorem prefix_match_semantics_wordset: assumes "valid_prefix pfx" shows "prefix_match_semantics pfx a \ a \ prefix_to_wordset pfx" using assms unfolding valid_prefix_def pfxm_mask_def prefix_match_semantics_def prefix_to_wordset_def apply(cases pfx, rename_tac base len) apply(simp) apply(drule_tac base=base and len=len and a=a in zero_base_lsb_imp_set_eq_as_bit_operation) by (simp) private lemma valid_prefix_ipset_from_netmask_ipset_from_cidr: shows "ipset_from_netmask (pfxm_prefix pfx) (NOT (pfxm_mask pfx)) = ipset_from_cidr (pfxm_prefix pfx) (pfxm_length pfx)" apply(cases pfx) apply(simp add: ipset_from_cidr_alt2 pfxm_mask_def) done lemma prefix_match_semantics_ipset_from_netmask: assumes "valid_prefix pfx" shows "prefix_match_semantics pfx a \ a \ ipset_from_netmask (pfxm_prefix pfx) (NOT (pfxm_mask pfx))" unfolding prefix_match_semantics_wordset[OF assms] unfolding valid_prefix_ipset_from_netmask_ipset_from_cidr unfolding prefix_to_wordset_def apply(subst ipset_from_cidr_base_wellforemd) subgoal using assms by(simp add: valid_prefix_def pfxm_mask_def) by(simp add: pfxm_mask_def) lemma prefix_match_semantics_ipset_from_netmask2: assumes "valid_prefix pfx" shows "prefix_match_semantics pfx (a :: 'i::len word) \ a \ ipset_from_cidr (pfxm_prefix pfx) (pfxm_length pfx)" unfolding prefix_match_semantics_ipset_from_netmask[OF assms] pfxm_mask_def ipset_from_cidr_def by (metis (full_types) NOT_mask_shifted_lenword word_not_not) lemma prefix_to_wordset_ipset_from_cidr: assumes "valid_prefix (pfx::'a::len prefix_match)" shows "prefix_to_wordset pfx = ipset_from_cidr (pfxm_prefix pfx) (pfxm_length pfx)" proof - have helper3: "(x::'a::len word) OR y = x OR y AND NOT x" for x y by (simp add: word_oa_dist2) have prefix_match_semantics_ipset_from_netmask: "(prefix_to_wordset pfx) = ipset_from_netmask (pfxm_prefix pfx) (NOT (pfxm_mask pfx))" unfolding prefix_to_wordset_def ipset_from_netmask_def Let_def using assms by (clarsimp dest!: valid_prefix_E) (metis bit.conj_commute mask_eq_0_eq_x) have "((mask len)::'a::len word) << LENGTH('a) - len = ~~ (mask (LENGTH('a) - len))" for len using NOT_mask_shifted_lenword by (metis word_not_not) from this[of "(pfxm_length pfx)"] have mask_def2_symmetric: "((mask (pfxm_length pfx)::'a::len word) << LENGTH('a) - pfxm_length pfx) = NOT (pfxm_mask pfx)" unfolding pfxm_mask_def by simp have ipset_from_netmask_prefix: "ipset_from_netmask (pfxm_prefix pfx) (NOT (pfxm_mask pfx)) = ipset_from_cidr (pfxm_prefix pfx) (pfxm_length pfx)" unfolding ipset_from_netmask_def ipset_from_cidr_alt unfolding pfxm_mask_def[symmetric] unfolding mask_def2_symmetric apply(simp) unfolding Let_def using assms[unfolded valid_prefix_def] by (metis helper3 word_bw_comms(2)) show ?thesis by (metis ipset_from_netmask_prefix local.prefix_match_semantics_ipset_from_netmask) qed definition prefix_to_wordinterval :: "'a::len prefix_match \ 'a wordinterval" where "prefix_to_wordinterval pfx \ WordInterval (pfxm_prefix pfx) (pfxm_prefix pfx OR pfxm_mask pfx)" lemma prefix_to_wordinterval_set_eq[simp]: "wordinterval_to_set (prefix_to_wordinterval pfx) = prefix_to_wordset pfx" unfolding prefix_to_wordinterval_def prefix_to_wordset_def by simp lemma prefix_to_wordinterval_def2: "prefix_to_wordinterval pfx = iprange_interval ((pfxm_prefix pfx), (pfxm_prefix pfx OR pfxm_mask pfx))" unfolding iprange_interval.simps prefix_to_wordinterval_def by simp corollary prefix_to_wordinterval_ipset_from_cidr: "valid_prefix pfx \ wordinterval_to_set (prefix_to_wordinterval pfx) = ipset_from_cidr (pfxm_prefix pfx) (pfxm_length pfx)" using prefix_to_wordset_ipset_from_cidr prefix_to_wordinterval_set_eq by auto end lemma prefix_never_empty: fixes d:: "'a::len prefix_match" shows"\ wordinterval_empty (prefix_to_wordinterval d)" by (simp add: le_word_or2 prefix_to_wordinterval_def) text\Getting a lowest element\ lemma ipset_from_cidr_lowest: "a \ ipset_from_cidr a n" using ip_cidr_set_def ipset_from_cidr_eq_ip_cidr_set by blast (*this is why I call the previous lemma 'lowest'*) lemma "valid_prefix (PrefixMatch a n) \ is_lowest_element a (ipset_from_cidr a n)" apply(simp add: is_lowest_element_def ipset_from_cidr_lowest) apply(simp add: ipset_from_cidr_eq_ip_cidr_set ip_cidr_set_def) apply(simp add: valid_prefix_def pfxm_mask_def) by (metis diff_zero eq_iff mask_out_sub_mask word_and_le2 word_bw_comms(1)) end diff --git a/thys/Iptables_Semantics/Primitive_Matchers/IpAddresses.thy b/thys/Iptables_Semantics/Primitive_Matchers/IpAddresses.thy --- a/thys/Iptables_Semantics/Primitive_Matchers/IpAddresses.thy +++ b/thys/Iptables_Semantics/Primitive_Matchers/IpAddresses.thy @@ -1,198 +1,203 @@ theory IpAddresses imports IP_Addresses.IP_Address_toString IP_Addresses.CIDR_Split "../Common/WordInterval_Lists" begin \ \Misc\ (*we dont't have an empty ip space, but a space which only contains the 0 address. We will use the option type to denote the empty space in some functions.*) lemma "ipset_from_cidr (ipv4addr_of_dotdecimal (0, 0, 0, 0)) 33 = {0}" by(simp add: ipv4addr_of_dotdecimal.simps ipv4addr_of_nat_def ipset_from_cidr_large_pfxlen) (*helper we use for spoofing protection specification*) definition all_but_those_ips :: "('i::len word \ nat) list \ ('i word \ nat) list" where "all_but_those_ips cidrips = cidr_split (wordinterval_invert (l2wi (map ipcidr_to_interval cidrips)))" lemma all_but_those_ips: "ipcidr_union_set (set (all_but_those_ips cidrips)) = UNIV - (\ (ip,n) \ set cidrips. ipset_from_cidr ip n)" apply(simp add: ) unfolding ipcidr_union_set_uncurry all_but_those_ips_def apply(simp add: cidr_split_prefix) apply(simp add: l2wi) apply(simp add: ipcidr_to_interval_def) using ipset_from_cidr_ipcidr_to_interval by blast section\IPv4 Addresses\ subsection\IPv4 Addresses in IPTables Notation (how we parse it)\ context notes [[typedef_overloaded]] begin datatype 'i ipt_iprange = \ \Singleton IP Address\ IpAddr "'i::len word" \ \CIDR notation: addr/xx\ | IpAddrNetmask "'i word" nat \ \-m iprange --src-range a.b.c.d-e.f.g.h\ | IpAddrRange "'i word" "'i word" (*the range is inclusive*) end fun ipt_iprange_to_set :: "'i::len ipt_iprange \ 'i word set" where "ipt_iprange_to_set (IpAddrNetmask base m) = ipset_from_cidr base m" | "ipt_iprange_to_set (IpAddr ip) = { ip }" | "ipt_iprange_to_set (IpAddrRange ip1 ip2) = { ip1 .. ip2 }" text\@{term ipt_iprange_to_set} can only represent an empty set if it is an empty range.\ lemma ipt_iprange_to_set_nonempty: "ipt_iprange_to_set ip = {} \ (\ip1 ip2. ip = IpAddrRange ip1 ip2 \ ip1 > ip2)" apply(cases ip) apply(simp; fail) apply(simp add: ipset_from_cidr_alt bitmagic_zeroLast_leq_or1Last; fail) apply(simp add:linorder_not_le; fail) done text\maybe this is necessary as code equation?\ + context + includes bit_operations_syntax + begin + lemma element_ipt_iprange_to_set[code_unfold]: "(addr::'i::len word) \ ipt_iprange_to_set X = ( case X of (IpAddrNetmask pre len) \ (pre AND ((mask len) << (len_of (TYPE('i)) - len))) \ addr \ addr \ pre OR (mask (len_of (TYPE('i)) - len)) | IpAddr ip \ (addr = ip) | IpAddrRange ip1 ip2 \ ip1 \ addr \ ip2 \ addr)" apply(cases X) apply(simp; fail) apply(simp add: ipset_from_cidr_alt; fail) apply(simp; fail) done + end lemma ipt_iprange_to_set_uncurry_IpAddrNetmask: "ipt_iprange_to_set (uncurry IpAddrNetmask a) = uncurry ipset_from_cidr a" by(simp split: uncurry_splits) text\IP address ranges to \(start, end)\ notation\ fun ipt_iprange_to_interval :: "'i::len ipt_iprange \ ('i word \ 'i word)" where "ipt_iprange_to_interval (IpAddr addr) = (addr, addr)" | "ipt_iprange_to_interval (IpAddrNetmask pre len) = ipcidr_to_interval (pre, len)" | "ipt_iprange_to_interval (IpAddrRange ip1 ip2) = (ip1, ip2)" lemma ipt_iprange_to_interval: "ipt_iprange_to_interval ip = (s,e) \ {s .. e} = ipt_iprange_to_set ip" apply(cases ip) apply(auto simp add: ipcidr_to_interval) done text\A list of IP address ranges to a @{typ "'i::len wordinterval"}. The nice thing is: the usual set operations are defined on this type. We can use the existing function @{const l2wi_intersect} if we want the intersection of the supplied list\ lemma "wordinterval_to_set (l2wi_intersect (map ipt_iprange_to_interval ips)) = (\ ip \ set ips. ipt_iprange_to_set ip)" apply(simp add: l2wi_intersect) using ipt_iprange_to_interval by blast text\We can use @{const l2wi} if we want the union of the supplied list\ lemma "wordinterval_to_set (l2wi (map ipt_iprange_to_interval ips)) = (\ ip \ set ips. ipt_iprange_to_set ip)" apply(simp add: l2wi) using ipt_iprange_to_interval by blast text\A list of (negated) IP address to a @{typ "'i::len wordinterval"}.\ definition ipt_iprange_negation_type_to_br_intersect :: "'i::len ipt_iprange negation_type list \ 'i wordinterval" where "ipt_iprange_negation_type_to_br_intersect l = l2wi_negation_type_intersect (NegPos_map ipt_iprange_to_interval l)" lemma ipt_iprange_negation_type_to_br_intersect: "wordinterval_to_set (ipt_iprange_negation_type_to_br_intersect l) = (\ ip \ set (getPos l). ipt_iprange_to_set ip) - (\ ip \ set (getNeg l). ipt_iprange_to_set ip)" apply(simp add: ipt_iprange_negation_type_to_br_intersect_def l2wi_negation_type_intersect NegPos_map_simps) using ipt_iprange_to_interval by blast text\The @{typ "'i::len wordinterval"} can be translated back into a list of IP ranges. If a list of intervals is enough, we can use @{const wi2l}. If we need it in @{typ "'i::len ipt_iprange"}, we can use this function.\ definition wi_2_cidr_ipt_iprange_list :: "'i::len wordinterval \ 'i ipt_iprange list" where "wi_2_cidr_ipt_iprange_list r = map (uncurry IpAddrNetmask) (cidr_split r)" lemma wi_2_cidr_ipt_iprange_list: "(\ ip \ set (wi_2_cidr_ipt_iprange_list r). ipt_iprange_to_set ip) = wordinterval_to_set r" proof - have "(\ ip \ set (wi_2_cidr_ipt_iprange_list r). ipt_iprange_to_set ip) = (\x\set (cidr_split r). uncurry ipset_from_cidr x)" unfolding wi_2_cidr_ipt_iprange_list_def by force thus ?thesis using cidr_split_prefix by metis qed text\For example, this allows the following transformation\ definition ipt_iprange_compress :: "'i::len ipt_iprange negation_type list \ 'i ipt_iprange list" where "ipt_iprange_compress = wi_2_cidr_ipt_iprange_list \ ipt_iprange_negation_type_to_br_intersect" lemma ipt_iprange_compress: "(\ ip \ set (ipt_iprange_compress l). ipt_iprange_to_set ip) = (\ ip \ set (getPos l). ipt_iprange_to_set ip) - (\ ip \ set (getNeg l). ipt_iprange_to_set ip)" by (metis wi_2_cidr_ipt_iprange_list comp_apply ipt_iprange_compress_def ipt_iprange_negation_type_to_br_intersect) definition normalized_cidr_ip :: "'i::len ipt_iprange \ bool" where "normalized_cidr_ip ip \ case ip of IpAddrNetmask _ _ \ True | _ \ False" lemma wi_2_cidr_ipt_iprange_list_normalized_IpAddrNetmask: "\a'\set (wi_2_cidr_ipt_iprange_list as). normalized_cidr_ip a'" apply(clarify) apply(simp add: wi_2_cidr_ipt_iprange_list_def normalized_cidr_ip_def) by force lemma ipt_iprange_compress_normalized_IpAddrNetmask: "\a'\set (ipt_iprange_compress as). normalized_cidr_ip a'" by(simp add: ipt_iprange_compress_def wi_2_cidr_ipt_iprange_list_normalized_IpAddrNetmask) definition ipt_iprange_to_cidr :: "'i::len ipt_iprange \ ('i word \ nat) list" where "ipt_iprange_to_cidr ips = cidr_split (iprange_interval (ipt_iprange_to_interval ips))" lemma ipt_ipvange_to_cidr: "ipcidr_union_set (set (ipt_iprange_to_cidr ips)) = (ipt_iprange_to_set ips)" apply(simp add: ipt_iprange_to_cidr_def) apply(simp add: ipcidr_union_set_uncurry) apply(case_tac "(ipt_iprange_to_interval ips)") apply(simp add: ipt_iprange_to_interval cidr_split_prefix_single) done (* actually, these are toString pretty printing helpers*) definition interval_to_wi_to_ipt_iprange :: "'i::len word \ 'i word \ 'i ipt_iprange" where "interval_to_wi_to_ipt_iprange s e \ if s = e then IpAddr s else case cidr_split (WordInterval s e) of [(ip,nmask)] \ IpAddrNetmask ip nmask | _ \ IpAddrRange s e" lemma interval_to_wi_to_ipt_ipv4range: "ipt_iprange_to_set (interval_to_wi_to_ipt_iprange s e) = {s..e}" proof - from cidr_split_prefix_single[of s e] have "cidr_split (WordInterval s e) = [(a, b)] \ ipset_from_cidr a b = {s..e}" for a b by(simp add: iprange_interval.simps) thus ?thesis by(simp add: interval_to_wi_to_ipt_iprange_def split: list.split) qed fun wi_to_ipt_iprange :: "'i::len wordinterval \ 'i ipt_iprange list" where "wi_to_ipt_iprange (WordInterval s e) = (if s > e then [] else [interval_to_wi_to_ipt_iprange s e])" | "wi_to_ipt_iprange (RangeUnion a b) = wi_to_ipt_iprange a @ wi_to_ipt_iprange b" lemma wi_to_ipt_ipv4range: "\(set (map ipt_iprange_to_set (wi_to_ipt_iprange wi))) = wordinterval_to_set wi" apply(induction wi) apply(simp add: interval_to_wi_to_ipt_ipv4range) apply(simp) done end diff --git a/thys/JinjaThreads/Common/BinOp.thy b/thys/JinjaThreads/Common/BinOp.thy --- a/thys/JinjaThreads/Common/BinOp.thy +++ b/thys/JinjaThreads/Common/BinOp.thy @@ -1,589 +1,603 @@ (* Title: JinjaThreads/Common/BinOp.thy Author: Andreas Lochbihler *) section \Binary Operators\ theory BinOp imports WellForm "Word_Lib.Bit_Shifts_Infix_Syntax" begin datatype bop = \ \names of binary operations\ Eq | NotEq | LessThan | LessOrEqual | GreaterThan | GreaterOrEqual | Add | Subtract | Mult | Div | Mod | BinAnd | BinOr | BinXor | ShiftLeft | ShiftRightZeros | ShiftRightSigned subsection\The semantics of binary operators\ +context + includes bit_operations_syntax +begin + type_synonym 'addr binop_ret = "'addr val + 'addr" \ \a value or the address of an exception\ fun binop_LessThan :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_LessThan (Intg i1) (Intg i2) = Some (Inl (Bool (i1 'addr val \ 'addr binop_ret option" where "binop_LessOrEqual (Intg i1) (Intg i2) = Some (Inl (Bool (i1 <=s i2)))" | "binop_LessOrEqual v1 v2 = None" fun binop_GreaterThan :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_GreaterThan (Intg i1) (Intg i2) = Some (Inl (Bool (i2 'addr val \ 'addr binop_ret option" where "binop_GreaterOrEqual (Intg i1) (Intg i2) = Some (Inl (Bool (i2 <=s i1)))" | "binop_GreaterOrEqual v1 v2 = None" fun binop_Add :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_Add (Intg i1) (Intg i2) = Some (Inl (Intg (i1 + i2)))" | "binop_Add v1 v2 = None" fun binop_Subtract :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_Subtract (Intg i1) (Intg i2) = Some (Inl (Intg (i1 - i2)))" | "binop_Subtract v1 v2 = None" fun binop_Mult :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_Mult (Intg i1) (Intg i2) = Some (Inl (Intg (i1 * i2)))" | "binop_Mult v1 v2 = None" fun binop_BinAnd :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_BinAnd (Intg i1) (Intg i2) = Some (Inl (Intg (i1 AND i2)))" | "binop_BinAnd (Bool b1) (Bool b2) = Some (Inl (Bool (b1 \ b2)))" | "binop_BinAnd v1 v2 = None" fun binop_BinOr :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_BinOr (Intg i1) (Intg i2) = Some (Inl (Intg (i1 OR i2)))" | "binop_BinOr (Bool b1) (Bool b2) = Some (Inl (Bool (b1 \ b2)))" | "binop_BinOr v1 v2 = None" fun binop_BinXor :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_BinXor (Intg i1) (Intg i2) = Some (Inl (Intg (i1 XOR i2)))" | "binop_BinXor (Bool b1) (Bool b2) = Some (Inl (Bool (b1 \ b2)))" | "binop_BinXor v1 v2 = None" fun binop_ShiftLeft :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_ShiftLeft (Intg i1) (Intg i2) = Some (Inl (Intg (i1 << unat (i2 AND 0x1f))))" | "binop_ShiftLeft v1 v2 = None" fun binop_ShiftRightZeros :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_ShiftRightZeros (Intg i1) (Intg i2) = Some (Inl (Intg (i1 >> unat (i2 AND 0x1f))))" | "binop_ShiftRightZeros v1 v2 = None" fun binop_ShiftRightSigned :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_ShiftRightSigned (Intg i1) (Intg i2) = Some (Inl (Intg (i1 >>> unat (i2 AND 0x1f))))" | "binop_ShiftRightSigned v1 v2 = None" text \ Division on @{typ "'a word"} is unsigned, but JLS specifies signed division. \ definition word_sdiv :: "'a :: len word \ 'a word \ 'a word" (infixl "sdiv" 70) where [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))" definition word_smod :: "'a :: len word \ 'a word \ 'a word" (infixl "smod" 70) where [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))" declare word_sdiv_def [simp] word_smod_def [simp] lemma sdiv_smod_id: "(a sdiv b) * b + (a smod b) = a" proof - have F5: "\u::'a word. - (- u) = u" by simp have F7: "\v u::'a word. u + v = v + u" by (simp add: ac_simps) have F8: "\(w::'a word) (v::int) u::int. word_of_int u + word_of_int v * w = word_of_int (u + v * sint w)" by simp have "\u. u = - sint b \ word_of_int (sint a mod u + - (- u * (sint a div u))) = a" using F5 by simp hence "word_of_int (sint a mod - sint b + - (sint b * (sint a div - sint b))) = a" by (metis equation_minus_iff) hence "word_of_int (sint a mod - sint b) + word_of_int (- (sint a div - sint b)) * b = a" using F8 by (simp add: ac_simps) hence eq: "word_of_int (- (sint a div - sint b)) * b + word_of_int (sint a mod - sint b) = a" using F7 by simp show ?thesis proof(cases "sint a < 0") case True note a = this show ?thesis proof(cases "sint b < 0") case True with a show ?thesis by simp (metis F7 F8 eq minus_equation_iff minus_mult_minus mod_div_mult_eq) next case False from eq have "word_of_int (- (- sint a div sint b)) * b + word_of_int (- (- sint a mod sint b)) = a" by (metis div_minus_right mod_minus_right) with a False show ?thesis by simp qed next case False note a = this show ?thesis proof(cases "sint b < 0") case True with a eq show ?thesis by simp next case False with a show ?thesis by (simp add: F7 F8) qed qed qed +end + notepad begin have " 5 sdiv ( 3 :: word32) = 1" and " 5 smod ( 3 :: word32) = 2" and " 5 sdiv (-3 :: word32) = -1" and " 5 smod (-3 :: word32) = 2" and "(-5) sdiv ( 3 :: word32) = -1" and "(-5) smod ( 3 :: word32) = -2" and "(-5) sdiv (-3 :: word32) = 1" and "(-5) smod (-3 :: word32) = -2" and "-2147483648 sdiv 1 = (-2147483648 :: word32)" by eval+ end -context heap_base begin +context heap_base +begin fun binop_Mod :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_Mod (Intg i1) (Intg i2) = Some (if i2 = 0 then Inr (addr_of_sys_xcpt ArithmeticException) else Inl (Intg (i1 smod i2)))" | "binop_Mod v1 v2 = None" fun binop_Div :: "'addr val \ 'addr val \ 'addr binop_ret option" where "binop_Div (Intg i1) (Intg i2) = Some (if i2 = 0 then Inr (addr_of_sys_xcpt ArithmeticException) else Inl (Intg (i1 sdiv i2)))" | "binop_Div v1 v2 = None" primrec binop :: "bop \ 'addr val \ 'addr val \ 'addr binop_ret option" where "binop Eq v1 v2 = Some (Inl (Bool (v1 = v2)))" | "binop NotEq v1 v2 = Some (Inl (Bool (v1 \ v2)))" | "binop LessThan = binop_LessThan" | "binop LessOrEqual = binop_LessOrEqual" | "binop GreaterThan = binop_GreaterThan" | "binop GreaterOrEqual = binop_GreaterOrEqual" | "binop Add = binop_Add" | "binop Subtract = binop_Subtract" | "binop Mult = binop_Mult" | "binop Mod = binop_Mod" | "binop Div = binop_Div" | "binop BinAnd = binop_BinAnd" | "binop BinOr = binop_BinOr" | "binop BinXor = binop_BinXor" | "binop ShiftLeft = binop_ShiftLeft" | "binop ShiftRightZeros = binop_ShiftRightZeros" | "binop ShiftRightSigned = binop_ShiftRightSigned" end +context + includes bit_operations_syntax +begin + lemma [simp]: "(binop_LessThan v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Bool (i1 (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Bool (i1 <=s i2)))" by(cases "(v1, v2)" rule: binop_LessOrEqual.cases) auto lemma [simp]: "(binop_GreaterThan v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Bool (i2 (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Bool (i2 <=s i1)))" by(cases "(v1, v2)" rule: binop_GreaterOrEqual.cases) auto lemma [simp]: "(binop_Add v\<^sub>1 v\<^sub>2 = Some va) \ (\i\<^sub>1 i\<^sub>2. v\<^sub>1 = Intg i\<^sub>1 \ v\<^sub>2 = Intg i\<^sub>2 \ va = Inl (Intg (i\<^sub>1+i\<^sub>2)))" by(cases "(v\<^sub>1, v\<^sub>2)" rule: binop_Add.cases) auto lemma [simp]: "(binop_Subtract v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 - i2)))" by(cases "(v1, v2)" rule: binop_Subtract.cases) auto lemma [simp]: "(binop_Mult v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 * i2)))" by(cases "(v1, v2)" rule: binop_Mult.cases) auto lemma [simp]: "(binop_BinAnd v1 v2 = Some va) \ (\b1 b2. v1 = Bool b1 \ v2 = Bool b2 \ va = Inl (Bool (b1 \ b2))) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 AND i2)))" by(cases "(v1, v2)" rule: binop_BinAnd.cases) auto lemma [simp]: "(binop_BinOr v1 v2 = Some va) \ (\b1 b2. v1 = Bool b1 \ v2 = Bool b2 \ va = Inl (Bool (b1 \ b2))) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 OR i2)))" by(cases "(v1, v2)" rule: binop_BinOr.cases) auto lemma [simp]: "(binop_BinXor v1 v2 = Some va) \ (\b1 b2. v1 = Bool b1 \ v2 = Bool b2 \ va = Inl (Bool (b1 \ b2))) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 XOR i2)))" by(cases "(v1, v2)" rule: binop_BinXor.cases) auto lemma [simp]: "(binop_ShiftLeft v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 << unat (i2 AND 0x1f))))" by(cases "(v1, v2)" rule: binop_ShiftLeft.cases) auto lemma [simp]: "(binop_ShiftRightZeros v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 >> unat (i2 AND 0x1f))))" by(cases "(v1, v2)" rule: binop_ShiftRightZeros.cases) auto lemma [simp]: "(binop_ShiftRightSigned v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = Inl (Intg (i1 >>> unat (i2 AND 0x1f))))" by(cases "(v1, v2)" rule: binop_ShiftRightSigned.cases) auto -context heap_base begin +end + +context heap_base +begin lemma [simp]: "(binop_Mod v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = (if i2 = 0 then Inr (addr_of_sys_xcpt ArithmeticException) else Inl (Intg(i1 smod i2))))" by(cases "(v1, v2)" rule: binop_Mod.cases) auto lemma [simp]: "(binop_Div v1 v2 = Some va) \ (\i1 i2. v1 = Intg i1 \ v2 = Intg i2 \ va = (if i2 = 0 then Inr (addr_of_sys_xcpt ArithmeticException) else Inl (Intg(i1 sdiv i2))))" by(cases "(v1, v2)" rule: binop_Div.cases) auto end subsection \Typing for binary operators\ inductive WT_binop :: "'m prog \ ty \ bop \ ty \ ty \ bool" ("_ \ _\_\_ :: _" [51,0,0,0,51] 50) where WT_binop_Eq: "P \ T1 \ T2 \ P \ T2 \ T1 \ P \ T1\Eq\T2 :: Boolean" | WT_binop_NotEq: "P \ T1 \ T2 \ P \ T2 \ T1 \ P \ T1\NotEq\T2 :: Boolean" | WT_binop_LessThan: "P \ Integer\LessThan\Integer :: Boolean" | WT_binop_LessOrEqual: "P \ Integer\LessOrEqual\Integer :: Boolean" | WT_binop_GreaterThan: "P \ Integer\GreaterThan\Integer :: Boolean" | WT_binop_GreaterOrEqual: "P \ Integer\GreaterOrEqual\Integer :: Boolean" | WT_binop_Add: "P \ Integer\Add\Integer :: Integer" | WT_binop_Subtract: "P \ Integer\Subtract\Integer :: Integer" | WT_binop_Mult: "P \ Integer\Mult\Integer :: Integer" | WT_binop_Div: "P \ Integer\Div\Integer :: Integer" | WT_binop_Mod: "P \ Integer\Mod\Integer :: Integer" | WT_binop_BinAnd_Bool: "P \ Boolean\BinAnd\Boolean :: Boolean" | WT_binop_BinAnd_Int: "P \ Integer\BinAnd\Integer :: Integer" | WT_binop_BinOr_Bool: "P \ Boolean\BinOr\Boolean :: Boolean" | WT_binop_BinOr_Int: "P \ Integer\BinOr\Integer :: Integer" | WT_binop_BinXor_Bool: "P \ Boolean\BinXor\Boolean :: Boolean" | WT_binop_BinXor_Int: "P \ Integer\BinXor\Integer :: Integer" | WT_binop_ShiftLeft: "P \ Integer\ShiftLeft\Integer :: Integer" | WT_binop_ShiftRightZeros: "P \ Integer\ShiftRightZeros\Integer :: Integer" | WT_binop_ShiftRightSigned: "P \ Integer\ShiftRightSigned\Integer :: Integer" lemma WT_binopI [intro]: "P \ T1 \ T2 \ P \ T2 \ T1 \ P \ T1\Eq\T2 :: Boolean" "P \ T1 \ T2 \ P \ T2 \ T1 \ P \ T1\NotEq\T2 :: Boolean" "bop = Add \ bop = Subtract \ bop = Mult \ bop = Mod \ bop = Div \ bop = BinAnd \ bop = BinOr \ bop = BinXor \ bop = ShiftLeft \ bop = ShiftRightZeros \ bop = ShiftRightSigned \ P \ Integer\bop\Integer :: Integer" "bop = LessThan \ bop = LessOrEqual \ bop = GreaterThan \ bop = GreaterOrEqual \ P \ Integer\bop\Integer :: Boolean" "bop = BinAnd \ bop = BinOr \ bop = BinXor \ P \ Boolean\bop\Boolean :: Boolean" by(auto intro: WT_binop.intros) inductive_cases [elim]: "P \ T1\Eq\T2 :: T" "P \ T1\NotEq\T2 :: T" "P \ T1\LessThan\T2 :: T" "P \ T1\LessOrEqual\T2 :: T" "P \ T1\GreaterThan\T2 :: T" "P \ T1\GreaterOrEqual\T2 :: T" "P \ T1\Add\T2 :: T" "P \ T1\Subtract\T2 :: T" "P \ T1\Mult\T2 :: T" "P \ T1\Div\T2 :: T" "P \ T1\Mod\T2 :: T" "P \ T1\BinAnd\T2 :: T" "P \ T1\BinOr\T2 :: T" "P \ T1\BinXor\T2 :: T" "P \ T1\ShiftLeft\T2 :: T" "P \ T1\ShiftRightZeros\T2 :: T" "P \ T1\ShiftRightSigned\T2 :: T" lemma WT_binop_fun: "\ P \ T1\bop\T2 :: T; P \ T1\bop\T2 :: T' \ \ T = T'" by(cases bop)(auto) lemma WT_binop_is_type: "\ P \ T1\bop\T2 :: T; is_type P T1; is_type P T2 \ \ is_type P T" by(cases bop) auto inductive WTrt_binop :: "'m prog \ ty \ bop \ ty \ ty \ bool" ("_ \ _\_\_ : _" [51,0,0,0,51] 50) where WTrt_binop_Eq: "P \ T1\Eq\T2 : Boolean" | WTrt_binop_NotEq: "P \ T1\NotEq\T2 : Boolean" | WTrt_binop_LessThan: "P \ Integer\LessThan\Integer : Boolean" | WTrt_binop_LessOrEqual: "P \ Integer\LessOrEqual\Integer : Boolean" | WTrt_binop_GreaterThan: "P \ Integer\GreaterThan\Integer : Boolean" | WTrt_binop_GreaterOrEqual: "P \ Integer\GreaterOrEqual\Integer : Boolean" | WTrt_binop_Add: "P \ Integer\Add\Integer : Integer" | WTrt_binop_Subtract: "P \ Integer\Subtract\Integer : Integer" | WTrt_binop_Mult: "P \ Integer\Mult\Integer : Integer" | WTrt_binop_Div: "P \ Integer\Div\Integer : Integer" | WTrt_binop_Mod: "P \ Integer\Mod\Integer : Integer" | WTrt_binop_BinAnd_Bool: "P \ Boolean\BinAnd\Boolean : Boolean" | WTrt_binop_BinAnd_Int: "P \ Integer\BinAnd\Integer : Integer" | WTrt_binop_BinOr_Bool: "P \ Boolean\BinOr\Boolean : Boolean" | WTrt_binop_BinOr_Int: "P \ Integer\BinOr\Integer : Integer" | WTrt_binop_BinXor_Bool: "P \ Boolean\BinXor\Boolean : Boolean" | WTrt_binop_BinXor_Int: "P \ Integer\BinXor\Integer : Integer" | WTrt_binop_ShiftLeft: "P \ Integer\ShiftLeft\Integer : Integer" | WTrt_binop_ShiftRightZeros: "P \ Integer\ShiftRightZeros\Integer : Integer" | WTrt_binop_ShiftRightSigned: "P \ Integer\ShiftRightSigned\Integer : Integer" lemma WTrt_binopI [intro]: "P \ T1\Eq\T2 : Boolean" "P \ T1\NotEq\T2 : Boolean" "bop = Add \ bop = Subtract \ bop = Mult \ bop = Div \ bop = Mod \ bop = BinAnd \ bop = BinOr \ bop = BinXor \ bop = ShiftLeft \ bop = ShiftRightZeros \ bop = ShiftRightSigned \ P \ Integer\bop\Integer : Integer" "bop = LessThan \ bop = LessOrEqual \ bop = GreaterThan \ bop = GreaterOrEqual \ P \ Integer\bop\Integer : Boolean" "bop = BinAnd \ bop = BinOr \ bop = BinXor \ P \ Boolean\bop\Boolean : Boolean" by(auto intro: WTrt_binop.intros) inductive_cases WTrt_binop_cases [elim]: "P \ T1\Eq\T2 : T" "P \ T1\NotEq\T2 : T" "P \ T1\LessThan\T2 : T" "P \ T1\LessOrEqual\T2 : T" "P \ T1\GreaterThan\T2 : T" "P \ T1\GreaterOrEqual\T2 : T" "P \ T1\Add\T2 : T" "P \ T1\Subtract\T2 : T" "P \ T1\Mult\T2 : T" "P \ T1\Div\T2 : T" "P \ T1\Mod\T2 : T" "P \ T1\BinAnd\T2 : T" "P \ T1\BinOr\T2 : T" "P \ T1\BinXor\T2 : T" "P \ T1\ShiftLeft\T2 : T" "P \ T1\ShiftRightZeros\T2 : T" "P \ T1\ShiftRightSigned\T2 : T" inductive_simps WTrt_binop_simps [simp]: "P \ T1\Eq\T2 : T" "P \ T1\NotEq\T2 : T" "P \ T1\LessThan\T2 : T" "P \ T1\LessOrEqual\T2 : T" "P \ T1\GreaterThan\T2 : T" "P \ T1\GreaterOrEqual\T2 : T" "P \ T1\Add\T2 : T" "P \ T1\Subtract\T2 : T" "P \ T1\Mult\T2 : T" "P \ T1\Div\T2 : T" "P \ T1\Mod\T2 : T" "P \ T1\BinAnd\T2 : T" "P \ T1\BinOr\T2 : T" "P \ T1\BinXor\T2 : T" "P \ T1\ShiftLeft\T2 : T" "P \ T1\ShiftRightZeros\T2 : T" "P \ T1\ShiftRightSigned\T2 : T" fun binop_relevant_class :: "bop \ 'm prog \ cname \ bool" where "binop_relevant_class Div = (\P C. P \ ArithmeticException \\<^sup>* C )" | "binop_relevant_class Mod = (\P C. P \ ArithmeticException \\<^sup>* C )" | "binop_relevant_class _ = (\P C. False)" lemma WT_binop_WTrt_binop: "P \ T1\bop\T2 :: T \ P \ T1\bop\T2 : T" by(auto elim: WT_binop.cases) context heap begin lemma binop_progress: "\ typeof\<^bsub>h\<^esub> v1 = \T1\; typeof\<^bsub>h\<^esub> v2 = \T2\; P \ T1\bop\T2 : T \ \ \va. binop bop v1 v2 = \va\" by(cases bop)(auto del: disjCI split del: if_split) lemma binop_type: assumes wf: "wf_prog wf_md P" and pre: "preallocated h" and type: "typeof\<^bsub>h\<^esub> v1 = \T1\" "typeof\<^bsub>h\<^esub> v2 = \T2\" "P \ T1\bop\T2 : T" shows "binop bop v1 v2 = \Inl v\ \ P,h \ v :\ T" and "binop bop v1 v2 = \Inr a\ \ P,h \ Addr a :\ Class Throwable" using type apply(case_tac [!] bop) apply(auto split: if_split_asm simp add: conf_def wf_preallocatedD[OF wf pre]) done lemma binop_relevant_class: assumes wf: "wf_prog wf_md P" and pre: "preallocated h" and bop: "binop bop v1 v2 = \Inr a\" and sup: "P \ cname_of h a \\<^sup>* C" shows "binop_relevant_class bop P C" using assms by(cases bop)(auto split: if_split_asm) end lemma WTrt_binop_fun: "\ P \ T1\bop\T2 : T; P \ T1\bop\T2 : T' \ \ T = T'" by(cases bop)(auto) lemma WTrt_binop_THE [simp]: "P \ T1\bop\T2 : T \ The (WTrt_binop P T1 bop T2) = T" by(auto dest: WTrt_binop_fun) lemma WTrt_binop_widen_mono: "\ P \ T1\bop\T2 : T; P \ T1' \ T1; P \ T2' \ T2 \ \ \T'. P \ T1'\bop\T2' : T' \ P \ T' \ T" by(cases bop)(auto elim!: WTrt_binop_cases) lemma WTrt_binop_is_type: "\ P \ T1\bop\T2 : T; is_type P T1; is_type P T2 \ \ is_type P T" by(cases bop) auto subsection \Code generator setup\ lemmas [code] = heap_base.binop_Div.simps heap_base.binop_Mod.simps heap_base.binop.simps code_pred (modes: i \ i \ i \ i \ o \ bool, i \ i \ i \ i \ i \ bool) WT_binop . code_pred (modes: i \ i \ i \ i \ o \ bool, i \ i \ i \ i \ i \ bool) WTrt_binop . lemma eval_WTrt_binop_i_i_i_i_o: "Predicate.eval (WTrt_binop_i_i_i_i_o P T1 bop T2) T \ P \ T1\bop\T2 : T" by(auto elim: WTrt_binop_i_i_i_i_oE intro: WTrt_binop_i_i_i_i_oI) lemma the_WTrt_binop_code: "(THE T. P \ T1\bop\T2 : T) = Predicate.the (WTrt_binop_i_i_i_i_o P T1 bop T2)" by(simp add: Predicate.the_def eval_WTrt_binop_i_i_i_i_o) end diff --git a/thys/Native_Word/Bits_Integer.thy b/thys/Native_Word/Bits_Integer.thy --- a/thys/Native_Word/Bits_Integer.thy +++ b/thys/Native_Word/Bits_Integer.thy @@ -1,667 +1,680 @@ (* Title: Bits_Integer.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Bit operations for target language integers\ theory Bits_Integer imports "Word_Lib.Bit_Comprehension" Code_Int_Integer_Conversion Code_Symbolic_Bits_Int begin lemmas [transfer_rule] = identity_quotient fun_quotient Quotient_integer[folded integer.pcr_cr_eq] lemma undefined_transfer: assumes "Quotient R Abs Rep T" shows "T (Rep undefined) undefined" using assms unfolding Quotient_alt_def by blast bundle undefined_transfer = undefined_transfer[transfer_rule] section \More lemmas about @{typ integer}s\ context includes integer.lifting begin lemma bitval_integer_transfer [transfer_rule]: "(rel_fun (=) pcr_integer) of_bool of_bool" by(auto simp add: of_bool_def integer.pcr_cr_eq cr_integer_def) lemma integer_of_nat_less_0_conv [simp]: "\ integer_of_nat n < 0" by(transfer) simp lemma int_of_integer_pow: "int_of_integer (x ^ n) = int_of_integer x ^ n" by(induct n) simp_all lemma pow_integer_transfer [transfer_rule]: "(rel_fun pcr_integer (rel_fun (=) pcr_integer)) (^) (^)" by(auto 4 3 simp add: integer.pcr_cr_eq cr_integer_def int_of_integer_pow) lemma sub1_lt_0_iff [simp]: "Code_Numeral.sub n num.One < 0 \ False" by(cases n)(simp_all add: Code_Numeral.sub_code) lemma nat_of_integer_numeral [simp]: "nat_of_integer (numeral n) = numeral n" by transfer simp lemma nat_of_integer_sub1_conv_pred_numeral [simp]: "nat_of_integer (Code_Numeral.sub n num.One) = pred_numeral n" by(cases n)(simp_all add: Code_Numeral.sub_code) lemma nat_of_integer_1 [simp]: "nat_of_integer 1 = 1" by transfer simp lemma dup_1 [simp]: "Code_Numeral.dup 1 = 2" by transfer simp section \Bit operations on @{typ integer}\ text \Bit operations on @{typ integer} are the same as on @{typ int}\ lift_definition bin_rest_integer :: "integer \ integer" is \\k . k div 2\ . lift_definition bin_last_integer :: "integer \ bool" is odd . lift_definition Bit_integer :: "integer \ bool \ integer" is \\k b. of_bool b + 2 * k\ . end instantiation integer :: lsb begin context includes integer.lifting begin lift_definition lsb_integer :: "integer \ bool" is lsb . instance by (standard; transfer) (fact lsb_odd) end end instantiation integer :: msb begin context includes integer.lifting begin lift_definition msb_integer :: "integer \ bool" is msb . instance .. end end instantiation integer :: set_bit begin context includes integer.lifting begin lift_definition set_bit_integer :: "integer \ nat \ bool \ integer" is set_bit . instance apply standard apply transfer apply (simp add: bit_simps) done end end abbreviation (input) wf_set_bits_integer where "wf_set_bits_integer \ wf_set_bits_int" section \Target language implementations\ text \ Unfortunately, this is not straightforward, because these API functions have different signatures and preconditions on the parameters: \begin{description} \item[Standard ML] Shifts in IntInf are given as word, but not IntInf. \item[Haskell] In the Data.Bits.Bits type class, shifts and bit indices are given as Int rather than Integer. \end{description} Additional constants take only parameters of type @{typ integer} rather than @{typ nat} and check the preconditions as far as possible (e.g., being non-negative) in a portable way. Manual implementations inside code\_printing perform the remaining range checks and convert these @{typ integer}s into the right type. For normalisation by evaluation, we derive custom code equations, because NBE does not know these code\_printing serialisations and would otherwise loop. \ code_identifier code_module Bits_Integer \ (SML) Bits_Int and (OCaml) Bits_Int and (Haskell) Bits_Int and (Scala) Bits_Int code_printing code_module Bits_Integer \ (SML) \structure Bits_Integer : sig val set_bit : IntInf.int -> IntInf.int -> bool -> IntInf.int val shiftl : IntInf.int -> IntInf.int -> IntInf.int val shiftr : IntInf.int -> IntInf.int -> IntInf.int val test_bit : IntInf.int -> IntInf.int -> bool end = struct val maxWord = IntInf.pow (2, Word.wordSize); fun set_bit x n b = if n < maxWord then if b then IntInf.orb (x, IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n))) else IntInf.andb (x, IntInf.notb (IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n)))) else raise (Fail ("Bit index too large: " ^ IntInf.toString n)); fun shiftl x n = if n < maxWord then IntInf.<< (x, Word.fromLargeInt (IntInf.toLarge n)) else raise (Fail ("Shift operand too large: " ^ IntInf.toString n)); fun shiftr x n = if n < maxWord then IntInf.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) else raise (Fail ("Shift operand too large: " ^ IntInf.toString n)); fun test_bit x n = if n < maxWord then IntInf.andb (x, IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n))) <> 0 else raise (Fail ("Bit index too large: " ^ IntInf.toString n)); end; (*struct Bits_Integer*)\ code_reserved SML Bits_Integer code_printing code_module Bits_Integer \ (OCaml) \module Bits_Integer : sig val shiftl : Z.t -> Z.t -> Z.t val shiftr : Z.t -> Z.t -> Z.t val test_bit : Z.t -> Z.t -> bool end = struct (* We do not need an explicit range checks here, because Big_int.int_of_big_int raises Failure if the argument does not fit into an int. *) let shiftl x n = Z.shift_left x (Z.to_int n);; let shiftr x n = Z.shift_right x (Z.to_int n);; let test_bit x n = Z.testbit x (Z.to_int n);; end;; (*struct Bits_Integer*)\ code_reserved OCaml Bits_Integer code_printing code_module Data_Bits \ (Haskell) \ module Data_Bits where { import qualified Data.Bits; {- The ...Bounded functions assume that the Integer argument for the shift or bit index fits into an Int, is non-negative and (for types of fixed bit width) less than bitSize -} infixl 7 .&.; infixl 6 `xor`; infixl 5 .|.; (.&.) :: Data.Bits.Bits a => a -> a -> a; (.&.) = (Data.Bits..&.); xor :: Data.Bits.Bits a => a -> a -> a; xor = Data.Bits.xor; (.|.) :: Data.Bits.Bits a => a -> a -> a; (.|.) = (Data.Bits..|.); complement :: Data.Bits.Bits a => a -> a; complement = Data.Bits.complement; testBitUnbounded :: Data.Bits.Bits a => a -> Integer -> Bool; testBitUnbounded x b | b <= toInteger (Prelude.maxBound :: Int) = Data.Bits.testBit x (fromInteger b) | otherwise = error ("Bit index too large: " ++ show b) ; testBitBounded :: Data.Bits.Bits a => a -> Integer -> Bool; testBitBounded x b = Data.Bits.testBit x (fromInteger b); setBitUnbounded :: Data.Bits.Bits a => a -> Integer -> Bool -> a; setBitUnbounded x n b | n <= toInteger (Prelude.maxBound :: Int) = if b then Data.Bits.setBit x (fromInteger n) else Data.Bits.clearBit x (fromInteger n) | otherwise = error ("Bit index too large: " ++ show n) ; setBitBounded :: Data.Bits.Bits a => a -> Integer -> Bool -> a; setBitBounded x n True = Data.Bits.setBit x (fromInteger n); setBitBounded x n False = Data.Bits.clearBit x (fromInteger n); shiftlUnbounded :: Data.Bits.Bits a => a -> Integer -> a; shiftlUnbounded x n | n <= toInteger (Prelude.maxBound :: Int) = Data.Bits.shiftL x (fromInteger n) | otherwise = error ("Shift operand too large: " ++ show n) ; shiftlBounded :: Data.Bits.Bits a => a -> Integer -> a; shiftlBounded x n = Data.Bits.shiftL x (fromInteger n); shiftrUnbounded :: Data.Bits.Bits a => a -> Integer -> a; shiftrUnbounded x n | n <= toInteger (Prelude.maxBound :: Int) = Data.Bits.shiftR x (fromInteger n) | otherwise = error ("Shift operand too large: " ++ show n) ; shiftrBounded :: (Ord a, Data.Bits.Bits a) => a -> Integer -> a; shiftrBounded x n = Data.Bits.shiftR x (fromInteger n); }\ and \ \@{theory HOL.Quickcheck_Narrowing} maps @{typ integer} to Haskell's Prelude.Int type instead of Integer. For compatibility with the Haskell target, we nevertheless provide bounded and unbounded functions.\ (Haskell_Quickcheck) \ module Data_Bits where { import qualified Data.Bits; {- The functions assume that the Int argument for the shift or bit index is non-negative and (for types of fixed bit width) less than bitSize -} infixl 7 .&.; infixl 6 `xor`; infixl 5 .|.; (.&.) :: Data.Bits.Bits a => a -> a -> a; (.&.) = (Data.Bits..&.); xor :: Data.Bits.Bits a => a -> a -> a; xor = Data.Bits.xor; (.|.) :: Data.Bits.Bits a => a -> a -> a; (.|.) = (Data.Bits..|.); complement :: Data.Bits.Bits a => a -> a; complement = Data.Bits.complement; testBitUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool; testBitUnbounded = Data.Bits.testBit; testBitBounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool; testBitBounded = Data.Bits.testBit; setBitUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool -> a; setBitUnbounded x n True = Data.Bits.setBit x n; setBitUnbounded x n False = Data.Bits.clearBit x n; setBitBounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool -> a; setBitBounded x n True = Data.Bits.setBit x n; setBitBounded x n False = Data.Bits.clearBit x n; shiftlUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> a; shiftlUnbounded = Data.Bits.shiftL; shiftlBounded :: Data.Bits.Bits a => a -> Prelude.Int -> a; shiftlBounded = Data.Bits.shiftL; shiftrUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> a; shiftrUnbounded = Data.Bits.shiftR; shiftrBounded :: (Ord a, Data.Bits.Bits a) => a -> Prelude.Int -> a; shiftrBounded = Data.Bits.shiftR; }\ code_reserved Haskell Data_Bits code_printing code_module Bits_Integer \ (Scala) \object Bits_Integer { def setBit(x: BigInt, n: BigInt, b: Boolean) : BigInt = if (n.isValidInt) if (b) x.setBit(n.toInt) else x.clearBit(n.toInt) else sys.error("Bit index too large: " + n.toString) def shiftl(x: BigInt, n: BigInt) : BigInt = if (n.isValidInt) x << n.toInt else sys.error("Shift index too large: " + n.toString) def shiftr(x: BigInt, n: BigInt) : BigInt = if (n.isValidInt) x << n.toInt else sys.error("Shift index too large: " + n.toString) def testBit(x: BigInt, n: BigInt) : Boolean = if (n.isValidInt) x.testBit(n.toInt) else sys.error("Bit index too large: " + n.toString) } /* object Bits_Integer */\ code_printing - constant "(AND) :: integer \ integer \ integer" \ + constant "Bit_Operations.and :: integer \ integer \ integer" \ (SML) "IntInf.andb ((_),/ (_))" and (OCaml) "Z.logand" and (Haskell) "((Data'_Bits..&.) :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "((Data'_Bits..&.) :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) infixl 3 "&" -| constant "(OR) :: integer \ integer \ integer" \ +| constant "Bit_Operations.or :: integer \ integer \ integer" \ (SML) "IntInf.orb ((_),/ (_))" and (OCaml) "Z.logor" and (Haskell) "((Data'_Bits..|.) :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "((Data'_Bits..|.) :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) infixl 1 "|" -| constant "(XOR) :: integer \ integer \ integer" \ +| constant "Bit_Operations.xor :: integer \ integer \ integer" \ (SML) "IntInf.xorb ((_),/ (_))" and (OCaml) "Z.logxor" and (Haskell) "(Data'_Bits.xor :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.xor :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) infixl 2 "^" | constant "NOT :: integer \ integer" \ (SML) "IntInf.notb" and (OCaml) "Z.lognot" and (Haskell) "(Data'_Bits.complement :: Integer -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.complement :: Prelude.Int -> Prelude.Int)" and (Scala) "_.unary'_~" code_printing constant bin_rest_integer \ (SML) "IntInf.div ((_), 2)" and (OCaml) "Z.shift'_right/ _/ 1" and (Haskell) "(Data'_Bits.shiftrUnbounded _ 1 :: Integer)" and (Haskell_Quickcheck) "(Data'_Bits.shiftrUnbounded _ 1 :: Prelude.Int)" and (Scala) "_ >> 1" context -includes integer.lifting + includes integer.lifting bit_operations_syntax begin lemma bitNOT_integer_code [code]: fixes i :: integer shows "NOT i = - i - 1" by transfer (simp add: not_int_def) lemma bin_rest_integer_code [code nbe]: "bin_rest_integer i = i div 2" by transfer rule lemma bin_last_integer_code [code]: "bin_last_integer i \ i AND 1 \ 0" by transfer (simp add: and_one_eq odd_iff_mod_2_eq_one) lemma bin_last_integer_nbe [code nbe]: "bin_last_integer i \ i mod 2 \ 0" by transfer (simp add: odd_iff_mod_2_eq_one) lemma bitval_bin_last_integer [code_unfold]: "of_bool (bin_last_integer i) = i AND 1" by transfer (simp add: and_one_eq mod_2_eq_odd) end definition integer_test_bit :: "integer \ integer \ bool" where "integer_test_bit x n = (if n < 0 then undefined x n else bit x (nat_of_integer n))" declare [[code drop: \bit :: integer \ nat \ bool\]] lemma bit_integer_code [code]: "bit x n \ integer_test_bit x (integer_of_nat n)" by (simp add: integer_test_bit_def) lemma integer_test_bit_code [code]: "integer_test_bit x (Code_Numeral.Neg n) = undefined x (Code_Numeral.Neg n)" "integer_test_bit 0 0 = False" "integer_test_bit 0 (Code_Numeral.Pos n) = False" "integer_test_bit (Code_Numeral.Pos num.One) 0 = True" "integer_test_bit (Code_Numeral.Pos (num.Bit0 n)) 0 = False" "integer_test_bit (Code_Numeral.Pos (num.Bit1 n)) 0 = True" "integer_test_bit (Code_Numeral.Pos num.One) (Code_Numeral.Pos n') = False" "integer_test_bit (Code_Numeral.Pos (num.Bit0 n)) (Code_Numeral.Pos n') = integer_test_bit (Code_Numeral.Pos n) (Code_Numeral.sub n' num.One)" "integer_test_bit (Code_Numeral.Pos (num.Bit1 n)) (Code_Numeral.Pos n') = integer_test_bit (Code_Numeral.Pos n) (Code_Numeral.sub n' num.One)" "integer_test_bit (Code_Numeral.Neg num.One) 0 = True" "integer_test_bit (Code_Numeral.Neg (num.Bit0 n)) 0 = False" "integer_test_bit (Code_Numeral.Neg (num.Bit1 n)) 0 = True" "integer_test_bit (Code_Numeral.Neg num.One) (Code_Numeral.Pos n') = True" "integer_test_bit (Code_Numeral.Neg (num.Bit0 n)) (Code_Numeral.Pos n') = integer_test_bit (Code_Numeral.Neg n) (Code_Numeral.sub n' num.One)" "integer_test_bit (Code_Numeral.Neg (num.Bit1 n)) (Code_Numeral.Pos n') = integer_test_bit (Code_Numeral.Neg (n + num.One)) (Code_Numeral.sub n' num.One)" by (simp_all add: integer_test_bit_def bit_integer_def flip: bit_not_int_iff') code_printing constant integer_test_bit \ (SML) "Bits'_Integer.test'_bit" and (OCaml) "Bits'_Integer.test'_bit" and (Haskell) "(Data'_Bits.testBitUnbounded :: Integer -> Integer -> Bool)" and (Haskell_Quickcheck) "(Data'_Bits.testBitUnbounded :: Prelude.Int -> Prelude.Int -> Bool)" and (Scala) "Bits'_Integer.testBit" context -includes integer.lifting + includes integer.lifting bit_operations_syntax begin lemma lsb_integer_code [code]: fixes x :: integer shows "lsb x = bit x 0" by transfer(simp add: lsb_int_def) definition integer_set_bit :: "integer \ integer \ bool \ integer" where [code del]: "integer_set_bit x n b = (if n < 0 then undefined x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_integer_code [code]: "set_bit x i b = integer_set_bit x (integer_of_nat i) b" by(simp add: integer_set_bit_def) lemma set_bit_integer_conv_masks: fixes x :: integer shows "set_bit x i b = (if b then x OR (push_bit i 1) else x AND NOT (push_bit i 1))" by transfer (simp add: int_set_bit_False_conv_NAND int_set_bit_True_conv_OR) end code_printing constant integer_set_bit \ (SML) "Bits'_Integer.set'_bit" and (Haskell) "(Data'_Bits.setBitUnbounded :: Integer -> Integer -> Bool -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.setBitUnbounded :: Prelude.Int -> Prelude.Int -> Bool -> Prelude.Int)" and (Scala) "Bits'_Integer.setBit" text \ OCaml.Big\_int does not have a method for changing an individual bit, so we emulate that with masks. We prefer an Isabelle implementation, because this then takes care of the signs for AND and OR. \ + +context + includes bit_operations_syntax +begin + lemma integer_set_bit_code [code]: "integer_set_bit x n b = (if n < 0 then undefined x n b else if b then x OR (push_bit (nat_of_integer n) 1) else x AND NOT (push_bit (nat_of_integer n) 1))" by (auto simp add: integer_set_bit_def not_less set_bit_eq set_bit_def unset_bit_def) +end + definition integer_shiftl :: "integer \ integer \ integer" where [code del]: "integer_shiftl x n = (if n < 0 then undefined x n else push_bit (nat_of_integer n) x)" declare [[code drop: \push_bit :: nat \ integer \ integer\]] lemma shiftl_integer_code [code]: fixes x :: integer shows "push_bit n x = integer_shiftl x (integer_of_nat n)" by(auto simp add: integer_shiftl_def) context includes integer.lifting begin lemma shiftl_integer_conv_mult_pow2: fixes x :: integer shows "push_bit n x = x * 2 ^ n" by (fact push_bit_eq_mult) lemma integer_shiftl_code [code]: "integer_shiftl x (Code_Numeral.Neg n) = undefined x (Code_Numeral.Neg n)" "integer_shiftl x 0 = x" "integer_shiftl x (Code_Numeral.Pos n) = integer_shiftl (Code_Numeral.dup x) (Code_Numeral.sub n num.One)" "integer_shiftl 0 (Code_Numeral.Pos n) = 0" apply (simp_all add: integer_shiftl_def numeral_eq_Suc) apply transfer apply (simp add: ac_simps) done end code_printing constant integer_shiftl \ (SML) "Bits'_Integer.shiftl" and (OCaml) "Bits'_Integer.shiftl" and (Haskell) "(Data'_Bits.shiftlUnbounded :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.shiftlUnbounded :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) "Bits'_Integer.shiftl" definition integer_shiftr :: "integer \ integer \ integer" where [code del]: "integer_shiftr x n = (if n < 0 then undefined x n else drop_bit (nat_of_integer n) x)" declare [[code drop: \drop_bit :: nat \ integer \ integer\]] lemma shiftr_integer_conv_div_pow2: includes integer.lifting fixes x :: integer shows "drop_bit n x = x div 2 ^ n" by (fact drop_bit_eq_div) lemma shiftr_integer_code [code]: fixes x :: integer shows "drop_bit n x = integer_shiftr x (integer_of_nat n)" by(auto simp add: integer_shiftr_def) code_printing constant integer_shiftr \ (SML) "Bits'_Integer.shiftr" and (OCaml) "Bits'_Integer.shiftr" and (Haskell) "(Data'_Bits.shiftrUnbounded :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.shiftrUnbounded :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) "Bits'_Integer.shiftr" lemma integer_shiftr_code [code]: includes integer.lifting shows "integer_shiftr x (Code_Numeral.Neg n) = undefined x (Code_Numeral.Neg n)" "integer_shiftr x 0 = x" "integer_shiftr 0 (Code_Numeral.Pos n) = 0" "integer_shiftr (Code_Numeral.Pos num.One) (Code_Numeral.Pos n) = 0" "integer_shiftr (Code_Numeral.Pos (num.Bit0 n')) (Code_Numeral.Pos n) = integer_shiftr (Code_Numeral.Pos n') (Code_Numeral.sub n num.One)" "integer_shiftr (Code_Numeral.Pos (num.Bit1 n')) (Code_Numeral.Pos n) = integer_shiftr (Code_Numeral.Pos n') (Code_Numeral.sub n num.One)" "integer_shiftr (Code_Numeral.Neg num.One) (Code_Numeral.Pos n) = -1" "integer_shiftr (Code_Numeral.Neg (num.Bit0 n')) (Code_Numeral.Pos n) = integer_shiftr (Code_Numeral.Neg n') (Code_Numeral.sub n num.One)" "integer_shiftr (Code_Numeral.Neg (num.Bit1 n')) (Code_Numeral.Pos n) = integer_shiftr (Code_Numeral.Neg (Num.inc n')) (Code_Numeral.sub n num.One)" apply (simp_all add: integer_shiftr_def numeral_eq_Suc drop_bit_Suc) apply transfer apply simp apply transfer apply simp apply transfer apply (simp add: add_One) done context includes integer.lifting begin lemma Bit_integer_code [code]: "Bit_integer i False = push_bit 1 i" "Bit_integer i True = push_bit 1 i + 1" by (transfer; simp)+ lemma msb_integer_code [code]: "msb (x :: integer) \ x < 0" by transfer (simp add: msb_int_def) end context -includes integer.lifting natural.lifting + includes integer.lifting natural.lifting bit_operations_syntax begin lemma bitAND_integer_unfold [code]: "x AND y = (if x = 0 then 0 else if x = - 1 then y else Bit_integer (bin_rest_integer x AND bin_rest_integer y) (bin_last_integer x \ bin_last_integer y))" by transfer (auto simp add: algebra_simps and_int_rec [of _ \_ * 2\] and_int_rec [of \_ * 2\] and_int_rec [of \1 + _ * 2\] elim!: evenE oddE) lemma bitOR_integer_unfold [code]: "x OR y = (if x = 0 then y else if x = - 1 then - 1 else Bit_integer (bin_rest_integer x OR bin_rest_integer y) (bin_last_integer x \ bin_last_integer y))" by transfer (auto simp add: algebra_simps or_int_rec [of _ \_ * 2\] or_int_rec [of _ \1 + _ * 2\] or_int_rec [of \1 + _ * 2\] elim!: evenE oddE) lemma bitXOR_integer_unfold [code]: "x XOR y = (if x = 0 then y else if x = - 1 then NOT y else Bit_integer (bin_rest_integer x XOR bin_rest_integer y) (\ bin_last_integer x \ bin_last_integer y))" by transfer (auto simp add: algebra_simps xor_int_rec [of _ \_ * 2\] xor_int_rec [of \_ * 2\] xor_int_rec [of \1 + _ * 2\] elim!: evenE oddE) end section \Test code generator setup\ +context + includes bit_operations_syntax +begin + definition bit_integer_test :: "bool" where "bit_integer_test = (([ -1 AND 3, 1 AND -3, 3 AND 5, -3 AND (- 5) , -3 OR 1, 1 OR -3, 3 OR 5, -3 OR (- 5) , NOT 1, NOT (- 3) , -1 XOR 3, 1 XOR (- 3), 3 XOR 5, -5 XOR (- 3) , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , push_bit 2 1, push_bit 3 (- 1) , drop_bit 3 100, drop_bit 3 (- 100)] :: integer list) = [ 3, 1, 1, -7 , -3, -3, 7, -1 , -2, 2 , -4, -4, 6, 6 , 21, -1, 4, -7 , 4, -8 , 12, -13] \ [ bit (5 :: integer) 4, bit (5 :: integer) 2, bit (-5 :: integer) 4, bit (-5 :: integer) 2 , lsb (5 :: integer), lsb (4 :: integer), lsb (-1 :: integer), lsb (-2 :: integer), msb (5 :: integer), msb (0 :: integer), msb (-1 :: integer), msb (-2 :: integer)] = [ False, True, True, False, True, False, True, False, False, False, True, True])" export_code bit_integer_test checking SML Haskell? Haskell_Quickcheck? OCaml? Scala notepad begin have bit_integer_test by eval have bit_integer_test by normalization have bit_integer_test by code_simp end ML_val \val true = @{code bit_integer_test}\ lemma "x AND y = x OR (y :: integer)" quickcheck[random, expect=counterexample] quickcheck[exhaustive, expect=counterexample] oops lemma "(x :: integer) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] oops lemma "(f :: integer \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) +end + hide_const bit_integer_test hide_fact bit_integer_test_def end diff --git a/thys/Native_Word/Code_Symbolic_Bits_Int.thy b/thys/Native_Word/Code_Symbolic_Bits_Int.thy --- a/thys/Native_Word/Code_Symbolic_Bits_Int.thy +++ b/thys/Native_Word/Code_Symbolic_Bits_Int.thy @@ -1,119 +1,125 @@ (* Title: Code_Symbolic_Bits_Int.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Symbolic implementation of bit operations on int\ theory Code_Symbolic_Bits_Int imports "Word_Lib.Least_significant_bit" "Word_Lib.Generic_set_bit" "Word_Lib.Bit_Comprehension" begin section \Implementations of bit operations on \<^typ>\int\ operating on symbolic representation\ +context + includes bit_operations_syntax +begin + lemma test_bit_int_code [code]: "bit (0::int) n = False" "bit (Int.Neg num.One) n = True" "bit (Int.Pos num.One) 0 = True" "bit (Int.Pos (num.Bit0 m)) 0 = False" "bit (Int.Pos (num.Bit1 m)) 0 = True" "bit (Int.Neg (num.Bit0 m)) 0 = False" "bit (Int.Neg (num.Bit1 m)) 0 = True" "bit (Int.Pos num.One) (Suc n) = False" "bit (Int.Pos (num.Bit0 m)) (Suc n) = bit (Int.Pos m) n" "bit (Int.Pos (num.Bit1 m)) (Suc n) = bit (Int.Pos m) n" "bit (Int.Neg (num.Bit0 m)) (Suc n) = bit (Int.Neg m) n" "bit (Int.Neg (num.Bit1 m)) (Suc n) = bit (Int.Neg (Num.inc m)) n" by (simp_all add: Num.add_One bit_Suc) lemma int_not_code [code]: "NOT (0 :: int) = -1" "NOT (Int.Pos n) = Int.Neg (Num.inc n)" "NOT (Int.Neg n) = Num.sub n num.One" by (simp_all add: Num.add_One not_int_def) lemma int_and_code [code]: fixes i j :: int shows "0 AND j = 0" "i AND 0 = 0" "Int.Pos n AND Int.Pos m = (case and_num n m of None \ 0 | Some n' \ Int.Pos n')" "Int.Neg n AND Int.Neg m = NOT (Num.sub n num.One OR Num.sub m num.One)" "Int.Pos n AND Int.Neg num.One = Int.Pos n" "Int.Pos n AND Int.Neg (num.Bit0 m) = Num.sub (or_not_num_neg (Num.BitM m) n) num.One" "Int.Pos n AND Int.Neg (num.Bit1 m) = Num.sub (or_not_num_neg (num.Bit0 m) n) num.One" "Int.Neg num.One AND Int.Pos m = Int.Pos m" "Int.Neg (num.Bit0 n) AND Int.Pos m = Num.sub (or_not_num_neg (Num.BitM n) m) num.One" "Int.Neg (num.Bit1 n) AND Int.Pos m = Num.sub (or_not_num_neg (num.Bit0 n) m) num.One" by (simp_all add: and_num_eq_None_iff and_num_eq_Some_iff sub_one_eq_not_neg numeral_or_not_num_eq ac_simps split: option.split) lemma int_or_code [code]: fixes i j :: int shows "0 OR j = j" "i OR 0 = i" "Int.Pos n OR Int.Pos m = Int.Pos (or_num n m)" "Int.Neg n OR Int.Neg m = NOT (Num.sub n num.One AND Num.sub m num.One)" "Int.Pos n OR Int.Neg num.One = Int.Neg num.One" "Int.Pos n OR Int.Neg (num.Bit0 m) = (case and_not_num (Num.BitM m) n of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" "Int.Pos n OR Int.Neg (num.Bit1 m) = (case and_not_num (num.Bit0 m) n of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" "Int.Neg num.One OR Int.Pos m = Int.Neg num.One" "Int.Neg (num.Bit0 n) OR Int.Pos m = (case and_not_num (Num.BitM n) m of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" "Int.Neg (num.Bit1 n) OR Int.Pos m = (case and_not_num (num.Bit0 n) m of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" apply (simp_all add: and_not_num_eq_None_iff and_not_num_eq_Some_iff numeral_or_num_eq sub_one_eq_not_neg add_One ac_simps split: option.split) apply (simp_all add: or_eq_not_not_and minus_numeral_inc_eq) done lemma int_xor_code [code]: fixes i j :: int shows "0 XOR j = j" "i XOR 0 = i" "Int.Pos n XOR Int.Pos m = (case xor_num n m of None \ 0 | Some n' \ Int.Pos n')" "Int.Neg n XOR Int.Neg m = Num.sub n num.One XOR Num.sub m num.One" "Int.Neg n XOR Int.Pos m = NOT (Num.sub n num.One XOR Int.Pos m)" "Int.Pos n XOR Int.Neg m = NOT (Int.Pos n XOR Num.sub m num.One)" by (simp_all add: xor_num_eq_None_iff xor_num_eq_Some_iff sub_one_eq_not_neg split: option.split) lemma bin_rest_code: "i div 2 = drop_bit 1 i" for i :: int by (simp add: drop_bit_eq_div) lemma set_bits_code [code]: "set_bits = Code.abort (STR ''set_bits is unsupported on type int'') (\_. set_bits :: _ \ int)" by simp lemma fixes i :: int shows int_set_bit_True_conv_OR [code]: "Generic_set_bit.set_bit i n True = i OR push_bit n 1" and int_set_bit_False_conv_NAND [code]: "Generic_set_bit.set_bit i n False = i AND NOT (push_bit n 1)" and int_set_bit_conv_ops: "Generic_set_bit.set_bit i n b = (if b then i OR (push_bit n 1) else i AND NOT (push_bit n 1))" by (simp_all add: bit_eq_iff) (auto simp add: bit_simps) declare [[code drop: \drop_bit :: nat \ int \ int\]] lemma drop_bit_int_code [code]: fixes i :: int shows "drop_bit 0 i = i" "drop_bit (Suc n) 0 = (0 :: int)" "drop_bit (Suc n) (Int.Pos num.One) = 0" "drop_bit (Suc n) (Int.Pos (num.Bit0 m)) = drop_bit n (Int.Pos m)" "drop_bit (Suc n) (Int.Pos (num.Bit1 m)) = drop_bit n (Int.Pos m)" "drop_bit (Suc n) (Int.Neg num.One) = - 1" "drop_bit (Suc n) (Int.Neg (num.Bit0 m)) = drop_bit n (Int.Neg m)" "drop_bit (Suc n) (Int.Neg (num.Bit1 m)) = drop_bit n (Int.Neg (Num.inc m))" by (simp_all add: drop_bit_Suc add_One) declare [[code drop: \push_bit :: nat \ int \ int\]] lemma push_bit_int_code [code]: "push_bit 0 i = i" "push_bit (Suc n) i = push_bit n (Int.dup i)" by (simp_all add: ac_simps) lemma int_lsb_code [code]: "lsb (0 :: int) = False" "lsb (Int.Pos num.One) = True" "lsb (Int.Pos (num.Bit0 w)) = False" "lsb (Int.Pos (num.Bit1 w)) = True" "lsb (Int.Neg num.One) = True" "lsb (Int.Neg (num.Bit0 w)) = False" "lsb (Int.Neg (num.Bit1 w)) = True" by simp_all end + +end diff --git a/thys/Native_Word/Code_Target_Bits_Int.thy b/thys/Native_Word/Code_Target_Bits_Int.thy --- a/thys/Native_Word/Code_Target_Bits_Int.thy +++ b/thys/Native_Word/Code_Target_Bits_Int.thy @@ -1,94 +1,100 @@ (* Title: Code_Target_Bits_Int.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Implementation of bit operations on int by target language operations\ theory Code_Target_Bits_Int imports Bits_Integer "HOL-Library.Code_Target_Int" begin +context + includes bit_operations_syntax +begin + declare [[code drop: "(AND) :: int \ _" "(OR) :: int \ _" "(XOR) :: int \ _" "NOT :: int \ _" "lsb :: int \ _" "set_bit :: int \ _" "bit :: int \ _" "push_bit :: _ \ int \ _" "drop_bit :: _ \ int \ _" int_of_integer_symbolic ]] lemma [code_unfold]: \of_bool (odd i) = i AND 1\ for i :: int by (simp add: and_one_eq mod2_eq_if) lemma [code_unfold]: \bit x n \ x AND (push_bit n 1) \ 0\ for x :: int by (fact bit_iff_and_push_bit_not_eq_0) context includes integer.lifting begin lemma bit_int_code [code]: "bit (int_of_integer x) n = bit x n" by transfer simp lemma and_int_code [code]: "int_of_integer i AND int_of_integer j = int_of_integer (i AND j)" by transfer simp lemma or_int_code [code]: "int_of_integer i OR int_of_integer j = int_of_integer (i OR j)" by transfer simp lemma xor_int_code [code]: "int_of_integer i XOR int_of_integer j = int_of_integer (i XOR j)" by transfer simp lemma not_int_code [code]: "NOT (int_of_integer i) = int_of_integer (NOT i)" by transfer simp lemma push_bit_int_code [code]: \push_bit n (int_of_integer x) = int_of_integer (push_bit n x)\ by transfer simp lemma drop_bit_int_code [code]: \drop_bit n (int_of_integer x) = int_of_integer (drop_bit n x)\ by transfer simp lemma take_bit_int_code [code]: \take_bit n (int_of_integer x) = int_of_integer (take_bit n x)\ by transfer simp lemma lsb_int_code [code]: "lsb (int_of_integer x) = lsb x" by transfer simp lemma set_bit_int_code [code]: "set_bit (int_of_integer x) n b = int_of_integer (set_bit x n b)" by transfer simp lemma int_of_integer_symbolic_code [code]: "int_of_integer_symbolic = int_of_integer" by (simp add: int_of_integer_symbolic_def) context begin qualified definition even :: \int \ bool\ where [code_abbrev]: \even = Parity.even\ end lemma [code]: \Code_Target_Bits_Int.even i \ i AND 1 = 0\ by (simp add: Code_Target_Bits_Int.even_def even_iff_mod_2_eq_zero and_one_eq) lemma bin_rest_code: "int_of_integer i div 2 = int_of_integer (bin_rest_integer i)" by transfer simp end end + +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,409 +1,424 @@ (* 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) 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) 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 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) 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) 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 + 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) + 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) -context fixes f :: "nat \ bool" begin +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/Native_Word/Native_Word_Test.thy b/thys/Native_Word/Native_Word_Test.thy --- a/thys/Native_Word/Native_Word_Test.thy +++ b/thys/Native_Word/Native_Word_Test.thy @@ -1,487 +1,514 @@ (* Title: Native_Word_Test.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Test cases\ theory Native_Word_Test imports Uint64 Uint32 Uint16 Uint8 Uint Native_Cast_Uint "HOL-Library.Code_Test" "Word_Lib.Bit_Shifts_Infix_Syntax" begin section \Tests for @{typ uint32}\ +context + includes bit_operations_syntax +begin + abbreviation (input) sshiftr_uint32 (infixl ">>>" 55) where \w >>> n \ signed_drop_bit_uint32 n w\ definition test_uint32 where "test_uint32 \ (([ 0x100000001, -1, -4294967291, 0xFFFFFFFF, 0x12345678 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + (- 6), 0xFFFFFFFFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x12345678 * 0x87654321 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 32 True, set_bit 5 32 False, set_bit (- 5) 32 True, set_bit (- 5) 32 False , 1 << 2, -1 << 3, 1 << 32, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 32, -100 >> 32 , 100 >>> 3, -100 >>> 3, 100 >>> 32, -100 >>> 32] :: uint32 list) = [ 1, 4294967295, 5, 4294967295, 305419896 , 18 , 126 , 108 , 4294967205 , 11, 1, 4294967295, 4294967285, 0 , 2, 4294967294 , 15, 4294967281, 20, 1891143032 , 1, 1431655763, 0, 0 , 2, 2, 4294967291, 5 , 21, 4294967295, 4, 4294967289 , 5, 5, 4294967291, 4294967291 , 4, 4294967288, 0, 1 , 12, 536870899, 0, 0 , 12, 4294967283, 0, 4294967295]) \ ([ (0x5 :: uint32) = 0x5, (0x5 :: uint32) = 0x6 , (0x5 :: uint32) < 0x5, (0x5 :: uint32) < 0x6, (-5 :: uint32) < 6, (6 :: uint32) < -5 , (0x5 :: uint32) \ 0x5, (0x5 :: uint32) \ 0x4, (-5 :: uint32) \ 6, (6 :: uint32) \ -5 , (0x7FFFFFFF :: uint32) < 0x80000000, (0xFFFFFFFF :: uint32) < 0, (0x80000000 :: uint32) < 0x7FFFFFFF , bit (0x7FFFFFFF :: uint32) 0, bit (0x7FFFFFFF :: uint32) 31, bit (0x80000000 :: uint32) 31, bit (0x80000000 :: uint32) 32 ] = [ True, False , False, True, False, True , True, False, False, True , True, False, False , True, False, True, False ]) \ ([integer_of_uint32 0, integer_of_uint32 0x7FFFFFFF, integer_of_uint32 0x80000000, integer_of_uint32 0xAAAAAAAA] = [0, 0x7FFFFFFF, 0x80000000, 0xAAAAAAAA])" -no_notation sshiftr_uint32 (infixl ">>>" 55) - export_code test_uint32 checking SML Haskell? OCaml? Scala notepad begin have test_uint32 by eval have test_uint32 by code_simp have test_uint32 by normalization end definition test_uint32' :: uint32 where "test_uint32' = 0 + 10 - 14 * 3 div 6 mod 3 << 3 >> 2" ML \val 0wx12 = @{code test_uint32'}\ lemma "x AND y = x OR (y :: uint32)" quickcheck[random, expect=counterexample] quickcheck[exhaustive, expect=counterexample] oops lemma "(x :: uint32) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] by transfer simp lemma "(f :: uint32 \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) +end + + section \Tests for @{typ uint16}\ +context + includes bit_operations_syntax +begin + abbreviation (input) sshiftr_uint16 (infixl ">>>" 55) where \w >>> n \ signed_drop_bit_uint16 n w\ definition test_uint16 where "test_uint16 \ (([ 0x10001, -1, -65535, 0xFFFF, 0x1234 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + -6, 0xFFFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x1234 * 0x8765 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 32 True, set_bit 5 32 False, set_bit (- 5) 32 True, set_bit (- 5) 32 False , 1 << 2, -1 << 3, 1 << 16, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 16, -100 >> 16 , 100 >>> 3, -100 >>> 3, 100 >>> 16, -100 >>> 16] :: uint16 list) = [ 1, 65535, 1, 65535, 4660 , 18 , 126 , 108 , 65445 , 11, 1, 65535, 65525, 0 , 2, 65534 , 15, 65521, 20, 39556 , 1, 21843, 0, 0 , 2, 2, 65531, 5 , 21, 65535, 4, 65529 , 5, 5, 65531, 65531 , 4, 65528, 0, 1 , 12, 8179, 0, 0 , 12, 65523, 0, 65535]) \ ([ (0x5 :: uint16) = 0x5, (0x5 :: uint16) = 0x6 , (0x5 :: uint16) < 0x5, (0x5 :: uint16) < 0x6, (-5 :: uint16) < 6, (6 :: uint16) < -5 , (0x5 :: uint16) \ 0x5, (0x5 :: uint16) \ 0x4, (-5 :: uint16) \ 6, (6 :: uint16) \ -5 , (0x7FFF :: uint16) < 0x8000, (0xFFFF :: uint16) < 0, (0x8000 :: uint16) < 0x7FFF , bit (0x7FFF :: uint16) 0, bit (0x7FFF :: uint16) 15, bit (0x8000 :: uint16) 15, bit (0x8000 :: uint16) 16 ] = [ True, False , False, True, False, True , True, False, False, True , True, False, False , True, False, True, False ]) \ ([integer_of_uint16 0, integer_of_uint16 0x7FFF, integer_of_uint16 0x8000, integer_of_uint16 0xAAAA] = [0, 0x7FFF, 0x8000, 0xAAAA])" export_code test_uint16 checking Haskell? Scala export_code test_uint16 checking SML_word notepad begin have test_uint16 by code_simp have test_uint16 by normalization end lemma "(x :: uint16) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] by transfer simp lemma "(f :: uint16 \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) +end + + section \Tests for @{typ uint8}\ +context + includes bit_operations_syntax +begin + abbreviation (input) sshiftr_uint8 (infixl ">>>" 55) where \w >>> n \ signed_drop_bit_uint8 n w\ definition test_uint8 where "test_uint8 \ (([ 0x101, -1, -255, 0xFF, 0x12 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + -6, 0xFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x12 * 0x87 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 32 True, set_bit 5 32 False, set_bit (- 5) 32 True, set_bit (- 5) 32 False , 1 << 2, -1 << 3, 1 << 8, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 8, -100 >> 8 , 100 >>> 3, -100 >>> 3, 100 >>> 8, -100 >>> 8] :: uint8 list) = [ 1, 255, 1, 255, 18 , 18 , 126 , 108 , 165 , 11, 1, 255, 245, 0 , 2, 254 , 15, 241, 20, 126 , 1, 83, 0, 0 , 2, 2, 251, 5 , 21, 255, 4, 249 , 5, 5, 251, 251 , 4, 248, 0, 1 , 12, 19, 0, 0 , 12, 243, 0, 255]) \ ([ (0x5 :: uint8) = 0x5, (0x5 :: uint8) = 0x6 , (0x5 :: uint8) < 0x5, (0x5 :: uint8) < 0x6, (-5 :: uint8) < 6, (6 :: uint8) < -5 , (0x5 :: uint8) \ 0x5, (0x5 :: uint8) \ 0x4, (-5 :: uint8) \ 6, (6 :: uint8) \ -5 , (0x7F :: uint8) < 0x80, (0xFF :: uint8) < 0, (0x80 :: uint8) < 0x7F , bit (0x7F :: uint8) 0, bit (0x7F :: uint8) 7, bit (0x80 :: uint8) 7, bit (0x80 :: uint8) 8 ] = [ True, False , False, True, False, True , True, False, False, True , True, False, False , True, False, True, False ]) \ ([integer_of_uint8 0, integer_of_uint8 0x7F, integer_of_uint8 0x80, integer_of_uint8 0xAA] = [0, 0x7F, 0x80, 0xAA])" -no_notation sshiftr_uint8 (infixl ">>>" 55) - export_code test_uint8 checking SML Haskell? Scala notepad begin have test_uint8 by eval have test_uint8 by code_simp have test_uint8 by normalization end ML_val \val true = @{code test_uint8}\ definition test_uint8' :: uint8 where "test_uint8' = 0 + 10 - 14 * 3 div 6 mod 3 << 3 >> 2" ML \val 0wx12 = @{code test_uint8'}\ lemma "x AND y = x OR (y :: uint8)" quickcheck[random, expect=counterexample] quickcheck[exhaustive, expect=counterexample] oops lemma "(x :: uint8) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] by transfer simp lemma "(f :: uint8 \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) + section \Tests for @{typ "uint"}\ +context + includes bit_operations_syntax +begin + abbreviation (input) sshiftr_uint (infixl ">>>" 55) where \w >>> n \ signed_drop_bit_uint n w\ definition "test_uint \ let test_list1 = (let HS = uint_of_int (2 ^ (dflt_size - 1)) in ([ HS + HS + 1, -1, -HS - HS + 5, HS + (HS - 1), 0x12 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + -6, HS + (HS - 1) + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x12345678 * 0x87654321] @ (if dflt_size > 4 then [ 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 dflt_size True, set_bit 5 dflt_size False, set_bit (- 5) dflt_size True, set_bit (- 5) dflt_size False , 1 << 2, -1 << 3, push_bit dflt_size 1, 1 << 0 , 31 >> 3, -1 >> 3, 31 >> dflt_size, -1 >> dflt_size , 15 >>> 2, -1 >>> 3, 15 >>> dflt_size, -1 >>> dflt_size] else []) :: uint list)); test_list2 = (let S = wivs_shift in ([ 1, -1, -S + 5, S - 1, 0x12 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + -6, 0 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x12345678 * 0x87654321] @ (if dflt_size > 4 then [ 5 div 3, (S - 5) div 3, (S - 5) div (S - 3), 5 div (S - 3) , 5 mod 3, (S - 5) mod 3, (S - 5) mod (S - 3), 5 mod (S - 3) , set_bit 5 4 True, -1, set_bit 5 0 False, -7 , 5, 5, -5, -5 , 4, -8, 0, 1 , 3, (S >> 3) - 1, 0, 0 , 3, (S >> 1) + (S >> 1) - 1, 0, -1] else []) :: int list)); test_list_c1 = (let HS = uint_of_int ((2^(dflt_size - 1))) in [ (0x5 :: uint) = 0x5, (0x5 :: uint) = 0x6 , (0x5 :: uint) < 0x5, (0x5 :: uint) < 0x6, (-5 :: uint) < 6, (6 :: uint) < -5 , (0x5 :: uint) \ 0x5, (0x5 :: uint) \ 0x4, (-5 :: uint) \ 6, (6 :: uint) \ -5 , (HS - 1) < HS, (HS + HS - 1) < 0, HS < HS - 1 , bit (HS - 1) 0, bit (HS - 1 :: uint) (dflt_size - 1), bit (HS :: uint) (dflt_size - 1), bit (HS :: uint) dflt_size ]); test_list_c2 = [ True, False , False, dflt_size\2, dflt_size=3, dflt_size\3 , True, False, dflt_size=3, dflt_size\3 , True, False, False , dflt_size\1, False, True, False ] in test_list1 = map uint_of_int test_list2 \ test_list_c1 = test_list_c2" -no_notation sshiftr_uint (infixl ">>>" 55) - export_code test_uint checking SML Haskell? OCaml? Scala lemma "test_uint" quickcheck[exhaustive, expect=no_counterexample] oops \ \FIXME: prove correctness of test by reflective means (not yet supported)\ lemma "x AND y = x OR (y :: uint)" quickcheck[random, expect=counterexample] quickcheck[exhaustive, expect=counterexample] oops lemma "(x :: uint) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] by transfer simp lemma "(f :: uint \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) -section \ Tests for @{typ uint64} \ + +section \Tests for @{typ uint64} \ + +context + includes bit_operations_syntax +begin abbreviation (input) sshiftr_uint64 (infixl ">>>" 55) where \w >>> n \ signed_drop_bit_uint64 n w\ definition test_uint64 where "test_uint64 \ (([ 0x10000000000000001, -1, -9223372036854775808, 0xFFFFFFFFFFFFFFFF, 0x1234567890ABCDEF , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + (- 6), 0xFFFFFFFFFFFFFFFFFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x1234567890ABCDEF * 0xFEDCBA0987654321 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 64 True, set_bit 5 64 False, set_bit (- 5) 64 True, set_bit (- 5) 64 False , 1 << 2, -1 << 3, 1 << 64, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 64, -100 >> 64 , 100 >>> 3, -100 >>> 3, 100 >>> 64, -100 >>> 64] :: uint64 list) = [ 1, 18446744073709551615, 9223372036854775808, 18446744073709551615, 1311768467294899695 , 18 , 126 , 108 , 18446744073709551525 , 11, 1, 18446744073709551615, 18446744073709551605, 0 , 2, 18446744073709551614 , 15, 18446744073709551601, 20, 14000077364136384719 , 1, 6148914691236517203, 0, 0 , 2, 2, 18446744073709551611, 5 , 21, 18446744073709551615, 4, 18446744073709551609 , 5, 5, 18446744073709551611, 18446744073709551611 , 4, 18446744073709551608, 0, 1 , 12, 2305843009213693939, 0, 0 , 12, 18446744073709551603, 0, 18446744073709551615]) \ ([ (0x5 :: uint64) = 0x5, (0x5 :: uint64) = 0x6 , (0x5 :: uint64) < 0x5, (0x5 :: uint64) < 0x6, (-5 :: uint64) < 6, (6 :: uint64) < -5 , (0x5 :: uint64) \ 0x5, (0x5 :: uint64) \ 0x4, (-5 :: uint64) \ 6, (6 :: uint64) \ -5 , (0x7FFFFFFFFFFFFFFF :: uint64) < 0x8000000000000000, (0xFFFFFFFFFFFFFFFF :: uint64) < 0, (0x8000000000000000 :: uint64) < 0x7FFFFFFFFFFFFFFF , bit (0x7FFFFFFFFFFFFFFF :: uint64) 0, bit (0x7FFFFFFFFFFFFFFF :: uint64) 63, bit (0x8000000000000000 :: uint64) 63, bit (0x8000000000000000 :: uint64) 64 ] = [ True, False , False, True, False, True , True, False, False, True , True, False, False , True, False, True, False ]) \ ([integer_of_uint64 0, integer_of_uint64 0x7FFFFFFFFFFFFFFF, integer_of_uint64 0x8000000000000000, integer_of_uint64 0xAAAAAAAAAAAAAAAA] = [0, 0x7FFFFFFFFFFFFFFF, 0x8000000000000000, 0xAAAAAAAAAAAAAAAA])" value [nbe] "[0x10000000000000001, -1, -9223372036854775808, 0xFFFFFFFFFFFFFFFF, 0x1234567890ABCDEF , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + (- 6), 0xFFFFFFFFFFFFFFFFFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x1234567890ABCDEF * 0xFEDCBA0987654321 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 64 True, set_bit 5 64 False, set_bit (- 5) 64 True, set_bit (- 5) 64 False , 1 << 2, -1 << 3, 1 << 64, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 64, -100 >> 64 , 100 >>> 3, -100 >>> 3, 100 >>> 64, -100 >>> 64] :: uint64 list" -no_notation sshiftr_uint64 (infixl ">>>" 55) - export_code test_uint64 checking SML Haskell? OCaml? Scala notepad begin have test_uint64 by eval have test_uint64 by code_simp have test_uint64 by normalization end ML_val \val true = @{code test_uint64}\ definition test_uint64' :: uint64 where "test_uint64' = 0 + 10 - 14 * 3 div 6 mod 3 << 3 >> 2" +end + + section \Tests for casts\ definition test_casts :: bool where "test_casts \ map uint8_of_uint32 [10, 0, 0xFE, 0xFFFFFFFF] = [10, 0, 0xFE, 0xFF] \ map uint8_of_uint64 [10, 0, 0xFE, 0xFFFFFFFFFFFFFFFF] = [10, 0, 0xFE, 0xFF] \ map uint32_of_uint8 [10, 0, 0xFF] = [10, 0, 0xFF] \ map uint64_of_uint8 [10, 0, 0xFF] = [10, 0, 0xFF]" definition test_casts' :: bool where "test_casts' \ map uint8_of_uint16 [10, 0, 0xFE, 0xFFFF] = [10, 0, 0xFE, 0xFF] \ map uint16_of_uint8 [10, 0, 0xFF] = [10, 0, 0xFF] \ map uint16_of_uint32 [10, 0, 0xFFFE, 0xFFFFFFFF] = [10, 0, 0xFFFE, 0xFFFF] \ map uint16_of_uint64 [10, 0, 0xFFFE, 0xFFFFFFFFFFFFFFFF] = [10, 0, 0xFFFE, 0xFFFF] \ map uint32_of_uint16 [10, 0, 0xFFFF] = [10, 0, 0xFFFF] \ map uint64_of_uint16 [10, 0, 0xFFFF] = [10, 0, 0xFFFF]" definition test_casts'' :: bool where "test_casts'' \ map uint32_of_uint64 [10, 0, 0xFFFFFFFE, 0xFFFFFFFFFFFFFFFF] = [10, 0, 0xFFFFFFFE, 0xFFFFFFFF] \ map uint64_of_uint32 [10, 0, 0xFFFFFFFF] = [10, 0, 0xFFFFFFFF]" export_code test_casts test_casts'' checking SML Haskell? Scala export_code test_casts'' checking OCaml? export_code test_casts' checking Haskell? Scala notepad begin have test_casts by eval have test_casts by normalization have test_casts by code_simp have test_casts' by normalization have test_casts' by code_simp have test_casts'' by eval have test_casts'' by normalization have test_casts'' by code_simp end ML \ val true = @{code test_casts} val true = @{code test_casts''} \ definition test_casts_uint :: bool where "test_casts_uint \ map uint_of_uint32 ([0, 10] @ (if dflt_size < 32 then [push_bit (dflt_size - 1) 1, 0xFFFFFFFF] else [0xFFFFFFFF])) = [0, 10] @ (if dflt_size < 32 then [push_bit (dflt_size - 1) 1, (push_bit dflt_size 1) - 1] else [0xFFFFFFFF]) \ map uint32_of_uint [0, 10, if dflt_size < 32 then push_bit (dflt_size - 1) 1 else 0xFFFFFFFF] = [0, 10, if dflt_size < 32 then push_bit (dflt_size - 1) 1 else 0xFFFFFFFF] \ map uint_of_uint64 [0, 10, push_bit (dflt_size - 1) 1, 0xFFFFFFFFFFFFFFFF] = [0, 10, push_bit (dflt_size - 1) 1, (push_bit dflt_size 1) - 1] \ map uint64_of_uint [0, 10, push_bit (dflt_size - 1) 1] = [0, 10, push_bit (dflt_size - 1) 1]" definition test_casts_uint' :: bool where "test_casts_uint' \ map uint_of_uint16 [0, 10, 0xFFFF] = [0, 10, 0xFFFF] \ map uint16_of_uint [0, 10, 0xFFFF] = [0, 10, 0xFFFF]" definition test_casts_uint'' :: bool where "test_casts_uint'' \ map uint_of_uint8 [0, 10, 0xFF] = [0, 10, 0xFF] \ map uint8_of_uint [0, 10, 0xFF] = [0, 10, 0xFF]" end + +end + +end diff --git a/thys/Native_Word/Uint.thy b/thys/Native_Word/Uint.thy --- a/thys/Native_Word/Uint.thy +++ b/thys/Native_Word/Uint.thy @@ -1,821 +1,824 @@ (* Title: Uint.thy Author: Peter Lammich, TU Munich Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of default size\ theory Uint imports Code_Target_Word_Base Word_Type_Copies begin text \ This theory provides access to words in the target languages of the code generator whose bit width is the default of the target language. To that end, the type \uint\ models words of width \dflt_size\, but \dflt_size\ is known only to be positive. Usage restrictions: Default-size words (type \uint\) cannot be used for evaluation, because the results depend on the particular choice of word size in the target language and implementation. Symbolic evaluation has not yet been set up for \uint\. \ text \The default size type\ typedecl dflt_size instantiation dflt_size :: typerep begin definition "typerep_class.typerep \ \_ :: dflt_size itself. Typerep.Typerep (STR ''Uint.dflt_size'') []" instance .. end consts dflt_size_aux :: "nat" specification (dflt_size_aux) dflt_size_aux_g0: "dflt_size_aux > 0" by auto hide_fact dflt_size_aux_def instantiation dflt_size :: len begin definition "len_of_dflt_size (_ :: dflt_size itself) \ dflt_size_aux" instance by(intro_classes)(simp add: len_of_dflt_size_def dflt_size_aux_g0) end abbreviation "dflt_size \ len_of (TYPE (dflt_size))" context includes integer.lifting begin lift_definition dflt_size_integer :: integer is "int dflt_size" . declare dflt_size_integer_def[code del] \ \The code generator will substitute a machine-dependent value for this constant\ lemma dflt_size_by_int[code]: "dflt_size = nat_of_integer dflt_size_integer" by transfer simp lemma dflt_size[simp]: "dflt_size > 0" "dflt_size \ Suc 0" "\ dflt_size < Suc 0" using len_gt_0[where 'a=dflt_size] by (simp_all del: len_gt_0) end section \Type definition and primitive operations\ typedef uint = \UNIV :: dflt_size word set\ .. global_interpretation uint: word_type_copy Abs_uint Rep_uint using type_definition_uint by (rule word_type_copy.intro) setup_lifting type_definition_uint declare uint.of_word_of [code abstype] declare Quotient_uint [transfer_rule] instantiation uint :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint :: uint is 0 . lift_definition one_uint :: uint is 1 . lift_definition plus_uint :: \uint \ uint \ uint\ is \(+)\ . lift_definition uminus_uint :: \uint \ uint\ is uminus . lift_definition minus_uint :: \uint \ uint \ uint\ is \(-)\ . lift_definition times_uint :: \uint \ uint \ uint\ is \(*)\ . lift_definition divide_uint :: \uint \ uint \ uint\ is \(div)\ . lift_definition modulo_uint :: \uint \ uint \ uint\ is \(mod)\ . lift_definition equal_uint :: \uint \ uint \ bool\ is \HOL.equal\ . lift_definition less_eq_uint :: \uint \ uint \ bool\ is \(\)\ . lift_definition less_uint :: \uint \ uint \ bool\ is \(<)\ . global_interpretation uint: word_type_copy_ring Abs_uint Rep_uint by standard (fact zero_uint.rep_eq one_uint.rep_eq plus_uint.rep_eq uminus_uint.rep_eq minus_uint.rep_eq times_uint.rep_eq divide_uint.rep_eq modulo_uint.rep_eq equal_uint.rep_eq less_eq_uint.rep_eq less_uint.rep_eq)+ instance proof - show \OFCLASS(uint, comm_ring_1_class)\ by (rule uint.of_class_comm_ring_1) show \OFCLASS(uint, semiring_modulo_class)\ by (fact uint.of_class_semiring_modulo) show \OFCLASS(uint, equal_class)\ by (fact uint.of_class_equal) show \OFCLASS(uint, linorder_class)\ by (fact uint.of_class_linorder) qed end instantiation uint :: ring_bit_operations begin lift_definition bit_uint :: \uint \ nat \ bool\ is bit . lift_definition not_uint :: \uint \ uint\ is NOT . -lift_definition and_uint :: \uint \ uint \ uint\ is \(AND)\ . -lift_definition or_uint :: \uint \ uint \ uint\ is \(OR)\ . -lift_definition xor_uint :: \uint \ uint \ uint\ is \(XOR)\ . +lift_definition and_uint :: \uint \ uint \ uint\ is \Bit_Operations.and\ . +lift_definition or_uint :: \uint \ uint \ uint\ is \Bit_Operations.or\ . +lift_definition xor_uint :: \uint \ uint \ uint\ is \Bit_Operations.xor\ . lift_definition mask_uint :: \nat \ uint\ is mask . lift_definition push_bit_uint :: \nat \ uint \ uint\ is push_bit . lift_definition drop_bit_uint :: \nat \ uint \ uint\ is drop_bit . lift_definition signed_drop_bit_uint :: \nat \ uint \ uint\ is signed_drop_bit . lift_definition take_bit_uint :: \nat \ uint \ uint\ is take_bit . lift_definition set_bit_uint :: \nat \ uint \ uint\ is Bit_Operations.set_bit . lift_definition unset_bit_uint :: \nat \ uint \ uint\ is unset_bit . lift_definition flip_bit_uint :: \nat \ uint \ uint\ is flip_bit . global_interpretation uint: word_type_copy_bits Abs_uint Rep_uint signed_drop_bit_uint by standard (fact bit_uint.rep_eq not_uint.rep_eq and_uint.rep_eq or_uint.rep_eq xor_uint.rep_eq mask_uint.rep_eq push_bit_uint.rep_eq drop_bit_uint.rep_eq signed_drop_bit_uint.rep_eq take_bit_uint.rep_eq set_bit_uint.rep_eq unset_bit_uint.rep_eq flip_bit_uint.rep_eq)+ instance by (fact uint.of_class_ring_bit_operations) end lift_definition uint_of_nat :: \nat \ uint\ is word_of_nat . lift_definition nat_of_uint :: \uint \ nat\ is unat . lift_definition uint_of_int :: \int \ uint\ is word_of_int . lift_definition int_of_uint :: \uint \ int\ is uint . context includes integer.lifting begin lift_definition Uint :: \integer \ uint\ is word_of_int . lift_definition integer_of_uint :: \uint \ integer\ is uint . end global_interpretation uint: word_type_copy_more Abs_uint Rep_uint signed_drop_bit_uint uint_of_nat nat_of_uint uint_of_int int_of_uint Uint integer_of_uint apply standard apply (simp_all add: uint_of_nat.rep_eq nat_of_uint.rep_eq uint_of_int.rep_eq int_of_uint.rep_eq Uint.rep_eq integer_of_uint.rep_eq integer_eq_iff) done instantiation uint :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint :: \uint \ nat\ is size . lift_definition msb_uint :: \uint \ bool\ is msb . lift_definition lsb_uint :: \uint \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint :: \uint \ nat \ bool \ uint\ where set_bit_uint_eq: \set_bit_uint a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint_transfer [transfer_rule]: \(cr_uint ===> (=) ===> (\) ===> cr_uint) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint_eq [abs_def]) transfer_prover end lift_definition set_bits_uint :: \(nat \ bool) \ uint\ is set_bits . lift_definition set_bits_aux_uint :: \(nat \ bool) \ nat \ uint \ uint\ is set_bits_aux . global_interpretation uint: word_type_copy_misc Abs_uint Rep_uint signed_drop_bit_uint uint_of_nat nat_of_uint uint_of_int int_of_uint Uint integer_of_uint dflt_size set_bits_aux_uint by (standard; transfer) simp_all instance using uint.of_class_bit_comprehension uint.of_class_set_bit uint.of_class_lsb by simp_all standard end section \Code setup\ code_printing code_module Uint \ (SML) \ structure Uint : sig val set_bit : Word.word -> IntInf.int -> bool -> Word.word val shiftl : Word.word -> IntInf.int -> Word.word val shiftr : Word.word -> IntInf.int -> Word.word val shiftr_signed : Word.word -> IntInf.int -> Word.word val test_bit : Word.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word.orb (x, mask) else Word.andb (x, Word.notb mask) end fun shiftl x n = Word.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word.andb (x, Word.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word.fromInt 0 end; (* struct Uint *)\ code_reserved SML Uint code_printing code_module Uint \ (Haskell) \module Uint(Int, Word, dflt_size) where import qualified Prelude import Data.Int(Int) import Data.Word(Word) import qualified Data.Bits dflt_size :: Prelude.Integer dflt_size = Prelude.toInteger (bitSize_aux (0::Word)) where bitSize_aux :: (Data.Bits.Bits a, Prelude.Bounded a) => a -> Int bitSize_aux = Data.Bits.bitSize\ and (Haskell_Quickcheck) \module Uint(Int, Word, dflt_size) where import qualified Prelude import Data.Int(Int) import Data.Word(Word) import qualified Data.Bits dflt_size :: Prelude.Int dflt_size = bitSize_aux (0::Word) where bitSize_aux :: (Data.Bits.Bits a, Prelude.Bounded a) => a -> Int bitSize_aux = Data.Bits.bitSize \ code_reserved Haskell Uint dflt_size text \ OCaml and Scala provide only signed bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module "Uint" \ (OCaml) \module Uint : sig type t = int val dflt_size : Z.t val less : t -> t -> bool val less_eq : t -> t -> bool val set_bit : t -> Z.t -> bool -> t val shiftl : t -> Z.t -> t val shiftr : t -> Z.t -> t val shiftr_signed : t -> Z.t -> t val test_bit : t -> Z.t -> bool val int_mask : int val int32_mask : int32 val int64_mask : int64 end = struct type t = int let dflt_size = Z.of_int Sys.int_size;; (* negative numbers have their highest bit set, so they are greater than positive ones *) let less x y = if x<0 then y<0 && x 0;; let int_mask = if Sys.int_size < 32 then lnot 0 else 0xFFFFFFFF;; let int32_mask = if Sys.int_size < 32 then Int32.pred (Int32.shift_left Int32.one Sys.int_size) else Int32.of_string "0xFFFFFFFF";; let int64_mask = if Sys.int_size < 64 then Int64.pred (Int64.shift_left Int64.one Sys.int_size) else Int64.of_string "0xFFFFFFFFFFFFFFFF";; end;; (*struct Uint*)\ code_reserved OCaml Uint code_printing code_module Uint \ (Scala) \object Uint { def dflt_size : BigInt = BigInt(32) def less(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Int, n: BigInt, b: Boolean) : Int = if (b) x | (1 << n.intValue) else x & (1 << n.intValue).unary_~ def shiftl(x: Int, n: BigInt) : Int = x << n.intValue def shiftr(x: Int, n: BigInt) : Int = x >>> n.intValue def shiftr_signed(x: Int, n: BigInt) : Int = x >> n.intValue def test_bit(x: Int, n: BigInt) : Boolean = (x & (1 << n.intValue)) != 0 } /* object Uint */\ code_reserved Scala Uint text \ OCaml's conversion from Big\_int to int demands that the value fits into a signed integer. The following justifies the implementation. \ -context includes integer.lifting begin +context + includes integer.lifting bit_operations_syntax +begin + definition wivs_mask :: int where "wivs_mask = 2^ dflt_size - 1" lift_definition wivs_mask_integer :: integer is wivs_mask . lemma [code]: "wivs_mask_integer = 2 ^ dflt_size - 1" by transfer (simp add: wivs_mask_def) definition wivs_shift :: int where "wivs_shift = 2 ^ dflt_size" lift_definition wivs_shift_integer :: integer is wivs_shift . lemma [code]: "wivs_shift_integer = 2 ^ dflt_size" by transfer (simp add: wivs_shift_def) definition wivs_index :: nat where "wivs_index == dflt_size - 1" lift_definition wivs_index_integer :: integer is "int wivs_index". lemma wivs_index_integer_code[code]: "wivs_index_integer = dflt_size_integer - 1" by transfer (simp add: wivs_index_def of_nat_diff) definition wivs_overflow :: int where "wivs_overflow == 2^ (dflt_size - 1)" lift_definition wivs_overflow_integer :: integer is wivs_overflow . lemma [code]: "wivs_overflow_integer = 2 ^ (dflt_size - 1)" by transfer (simp add: wivs_overflow_def) definition wivs_least :: int where "wivs_least == - wivs_overflow" lift_definition wivs_least_integer :: integer is wivs_least . lemma [code]: "wivs_least_integer = - (2 ^ (dflt_size - 1))" by transfer (simp add: wivs_overflow_def wivs_least_def) definition Uint_signed :: "integer \ uint" where "Uint_signed i = (if i < wivs_least_integer \ wivs_overflow_integer \ i then undefined Uint i else Uint i)" lemma Uint_code [code]: "Uint i = (let i' = i AND wivs_mask_integer in if bit i' wivs_index then Uint_signed (i' - wivs_shift_integer) else Uint_signed i')" including undefined_transfer unfolding Uint_signed_def apply transfer apply (subst word_of_int_via_signed) apply (auto simp add: push_bit_of_1 mask_eq_exp_minus_1 word_of_int_via_signed wivs_mask_def wivs_index_def wivs_overflow_def wivs_least_def wivs_shift_def) done lemma Uint_signed_code [code]: "Rep_uint (Uint_signed i) = (if i < wivs_least_integer \ i \ wivs_overflow_integer then Rep_uint (undefined Uint i) else word_of_int (int_of_integer_symbolic i))" unfolding Uint_signed_def Uint_def int_of_integer_symbolic_def word_of_integer_def by(simp add: Abs_uint_inverse) end text \ Avoid @{term Abs_uint} in generated code, use @{term Rep_uint'} instead. The symbolic implementations for code\_simp use @{term Rep_uint}. The new destructor @{term Rep_uint'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint} ([code abstract] equations for @{typ uint} may use @{term Rep_uint} because these instances will be folded away.) \ definition Rep_uint' where [simp]: "Rep_uint' = Rep_uint" lemma Rep_uint'_code [code]: "Rep_uint' x = (BITS n. bit x n)" unfolding Rep_uint'_def by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint' :: "dflt_size word \ uint" is "\x :: dflt_size word. x" . lemma Abs_uint'_code [code]: "Abs_uint' x = Uint (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint \ _"]] lemma term_of_uint_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint.uint.Abs_uint'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR (STR ''Uint.dflt_size'') []], TR (STR ''Uint.uint'') []])) (term_of_class.term_of (Rep_uint' x))" by(simp add: term_of_anything) text \Important: We must prevent the reflection oracle (eval-tac) to use our machine-dependent type. \ code_printing type_constructor uint \ (SML) "Word.word" and (Haskell) "Uint.Word" and (OCaml) "Uint.t" and (Scala) "Int" and (Eval) "*** \"Error: Machine dependent type\" ***" and (Quickcheck) "Word.word" | constant dflt_size_integer \ (SML) "(IntInf.fromLarge (Int.toLarge Word.wordSize))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.wordSize" and (Haskell) "Uint.dflt'_size" and (OCaml) "Uint.dflt'_size" and (Scala) "Uint.dflt'_size" | constant Uint \ (SML) "Word.fromLargeInt (IntInf.toLarge _)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.fromInt" and (Haskell) "(Prelude.fromInteger _ :: Uint.Word)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint.Word)" and (Scala) "_.intValue" | constant Uint_signed \ (OCaml) "Z.to'_int" | constant "0 :: uint" \ (SML) "(Word.fromInt 0)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "(Word.fromInt 0)" and (Haskell) "(0 :: Uint.Word)" and (OCaml) "0" and (Scala) "0" | constant "1 :: uint" \ (SML) "(Word.fromInt 1)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "(Word.fromInt 1)" and (Haskell) "(1 :: Uint.Word)" and (OCaml) "1" and (Scala) "1" | constant "plus :: uint \ _ " \ (SML) "Word.+ ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.+ ((_), (_))" and (Haskell) infixl 6 "+" and (OCaml) "Pervasives.(+)" and (Scala) infixl 7 "+" | constant "uminus :: uint \ _" \ (SML) "Word.~" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.~" and (Haskell) "negate" and (OCaml) "Pervasives.(~-)" and (Scala) "!(- _)" | constant "minus :: uint \ _" \ (SML) "Word.- ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.- ((_), (_))" and (Haskell) infixl 6 "-" and (OCaml) "Pervasives.(-)" and (Scala) infixl 7 "-" | constant "times :: uint \ _ \ _" \ (SML) "Word.* ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.* ((_), (_))" and (Haskell) infixl 7 "*" and (OCaml) "Pervasives.( * )" and (Scala) infixl 8 "*" | constant "HOL.equal :: uint \ _ \ bool" \ (SML) "!((_ : Word.word) = _)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "!((_ : Word.word) = _)" and (Haskell) infix 4 "==" and (OCaml) "(Pervasives.(=):Uint.t -> Uint.t -> bool)" and (Scala) infixl 5 "==" | class_instance uint :: equal \ (Haskell) - | constant "less_eq :: uint \ _ \ bool" \ (SML) "Word.<= ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.<= ((_), (_))" and (Haskell) infix 4 "<=" and (OCaml) "Uint.less'_eq" and (Scala) "Uint.less'_eq" | constant "less :: uint \ _ \ bool" \ (SML) "Word.< ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.< ((_), (_))" and (Haskell) infix 4 "<" and (OCaml) "Uint.less" and (Scala) "Uint.less" | constant "NOT :: uint \ _" \ (SML) "Word.notb" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.notb" and (Haskell) "Data'_Bits.complement" and (OCaml) "Pervasives.lnot" and (Scala) "_.unary'_~" -| constant "(AND) :: uint \ _" \ +| constant "Bit_Operations.and :: uint \ _" \ (SML) "Word.andb ((_),/ (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (OCaml) "Pervasives.(land)" and (Scala) infixl 3 "&" -| constant "(OR) :: uint \ _" \ +| constant "Bit_Operations.or :: uint \ _" \ (SML) "Word.orb ((_),/ (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (OCaml) "Pervasives.(lor)" and (Scala) infixl 1 "|" -| constant "(XOR) :: uint \ _" \ +| constant "Bit_Operations.xor :: uint \ _" \ (SML) "Word.xorb ((_),/ (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.xorb ((_),/ (_))" and (Haskell) "Data'_Bits.xor" and (OCaml) "Pervasives.(lxor)" and (Scala) infixl 2 "^" definition uint_divmod :: "uint \ uint \ uint \ uint" where "uint_divmod x y = (if y = 0 then (undefined ((div) :: uint \ _) x (0 :: uint), undefined ((mod) :: uint \ _) x (0 :: uint)) else (x div y, x mod y))" definition uint_div :: "uint \ uint \ uint" where "uint_div x y = fst (uint_divmod x y)" definition uint_mod :: "uint \ uint \ uint" where "uint_mod x y = snd (uint_divmod x y)" lemma div_uint_code [code]: "x div y = (if y = 0 then 0 else uint_div x y)" including undefined_transfer unfolding uint_divmod_def uint_div_def by transfer(simp add: word_div_def) lemma mod_uint_code [code]: "x mod y = (if y = 0 then x else uint_mod x y)" including undefined_transfer unfolding uint_mod_def uint_divmod_def by transfer(simp add: word_mod_def) definition uint_sdiv :: "uint \ uint \ uint" where [code del]: "uint_sdiv x y = (if y = 0 then undefined ((div) :: uint \ _) x (0 :: uint) else Abs_uint (Rep_uint x sdiv Rep_uint y))" definition div0_uint :: "uint \ uint" where [code del]: "div0_uint x = undefined ((div) :: uint \ _) x (0 :: uint)" declare [[code abort: div0_uint]] definition mod0_uint :: "uint \ uint" where [code del]: "mod0_uint x = undefined ((mod) :: uint \ _) x (0 :: uint)" declare [[code abort: mod0_uint]] definition wivs_overflow_uint :: uint where "wivs_overflow_uint \ push_bit (dflt_size - 1) 1" lemma Rep_uint_wivs_overflow_uint_eq: \Rep_uint wivs_overflow_uint = 2 ^ (dflt_size - Suc 0)\ by (simp add: wivs_overflow_uint_def one_uint.rep_eq push_bit_uint.rep_eq uint.word_of_power push_bit_eq_mult) lemma wivs_overflow_uint_greater_eq_0: \wivs_overflow_uint > 0\ apply (simp add: less_uint.rep_eq zero_uint.rep_eq Rep_uint_wivs_overflow_uint_eq) apply transfer apply (simp add: take_bit_push_bit push_bit_eq_mult) done lemma uint_divmod_code [code]: "uint_divmod x y = (if wivs_overflow_uint \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint x, mod0_uint x) else let q = push_bit 1 (uint_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" proof (cases \y = 0\) case True moreover have \x \ 0\ by transfer simp moreover note wivs_overflow_uint_greater_eq_0 ultimately show ?thesis by (auto simp add: uint_divmod_def div0_uint_def mod0_uint_def not_less) next case False then show ?thesis including undefined_transfer unfolding uint_divmod_def uint_sdiv_def div0_uint_def mod0_uint_def wivs_overflow_uint_def apply transfer apply (simp add: divmod_via_sdivmod push_bit_of_1) done qed lemma uint_sdiv_code [code]: "Rep_uint (uint_sdiv x y) = (if y = 0 then Rep_uint (undefined ((div) :: uint \ _) x (0 :: uint)) else Rep_uint x sdiv Rep_uint y)" unfolding uint_sdiv_def by(simp add: Abs_uint_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint_divmod_code} computes both with division only. \ code_printing constant uint_div \ (SML) "Word.div ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.div ((_), (_))" and (Haskell) "Prelude.div" | constant uint_mod \ (SML) "Word.mod ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.mod ((_), (_))" and (Haskell) "Prelude.mod" | constant uint_divmod \ (Haskell) "divmod" | constant uint_sdiv \ (OCaml) "Pervasives.('/)" and (Scala) "_ '/ _" definition uint_test_bit :: "uint \ integer \ bool" where [code del]: "uint_test_bit x n = (if n < 0 \ dflt_size_integer \ n then undefined (bit :: uint \ _) x n else bit x (nat_of_integer n))" lemma test_bit_uint_code [code]: "bit x n \ n < dflt_size \ uint_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint_test_bit_def by (transfer, simp, transfer, simp) lemma uint_test_bit_code [code]: "uint_test_bit w n = (if n < 0 \ dflt_size_integer \ n then undefined (bit :: uint \ _) w n else bit (Rep_uint w) (nat_of_integer n))" unfolding uint_test_bit_def by(simp add: bit_uint.rep_eq) code_printing constant uint_test_bit \ (SML) "Uint.test'_bit" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (OCaml) "Uint.test'_bit" and (Scala) "Uint.test'_bit" definition uint_set_bit :: "uint \ integer \ bool \ uint" where [code del]: "uint_set_bit x n b = (if n < 0 \ dflt_size_integer \ n then undefined (set_bit :: uint \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint_code [code]: "set_bit x n b = (if n < dflt_size then uint_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint_set_bit_def by (transfer) (auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint_set_bit_code [code]: "Rep_uint (uint_set_bit w n b) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined (set_bit :: uint \ _) w n b) else set_bit (Rep_uint w) (nat_of_integer n) b)" including undefined_transfer integer.lifting unfolding uint_set_bit_def by transfer simp code_printing constant uint_set_bit \ (SML) "Uint.set'_bit" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (OCaml) "Uint.set'_bit" and (Scala) "Uint.set'_bit" definition uint_shiftl :: "uint \ integer \ uint" where [code del]: "uint_shiftl x n = (if n < 0 \ dflt_size_integer \ n then undefined (push_bit :: nat \ uint \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint_code [code]: "push_bit n x = (if n < dflt_size then uint_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint_shiftl_def by (transfer fixing: n) simp lemma uint_shiftl_code [code]: "Rep_uint (uint_shiftl w n) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined (push_bit :: nat \ uint \ _) w n) else push_bit (nat_of_integer n) (Rep_uint w))" including undefined_transfer integer.lifting unfolding uint_shiftl_def by transfer simp code_printing constant uint_shiftl \ (SML) "Uint.shiftl" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (OCaml) "Uint.shiftl" and (Scala) "Uint.shiftl" definition uint_shiftr :: "uint \ integer \ uint" where [code del]: "uint_shiftr x n = (if n < 0 \ dflt_size_integer \ n then undefined (drop_bit :: nat \ uint \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint_code [code]: "drop_bit n x = (if n < dflt_size then uint_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint_shiftr_def by (transfer fixing: n) simp lemma uint_shiftr_code [code]: "Rep_uint (uint_shiftr w n) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined (drop_bit :: nat \ uint \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint w))" including undefined_transfer integer.lifting unfolding uint_shiftr_def by transfer simp code_printing constant uint_shiftr \ (SML) "Uint.shiftr" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (OCaml) "Uint.shiftr" and (Scala) "Uint.shiftr" definition uint_sshiftr :: "uint \ integer \ uint" where [code del]: "uint_sshiftr x n = (if n < 0 \ dflt_size_integer \ n then undefined signed_drop_bit_uint n x else signed_drop_bit_uint (nat_of_integer n) x)" lemma sshiftr_uint_code [code]: "signed_drop_bit_uint n x = (if n < dflt_size then uint_sshiftr x (integer_of_nat n) else if bit x wivs_index then -1 else 0)" including undefined_transfer integer.lifting unfolding uint_sshiftr_def by transfer(simp add: not_less signed_drop_bit_beyond word_size wivs_index_def) lemma uint_sshiftr_code [code]: "Rep_uint (uint_sshiftr w n) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined signed_drop_bit_uint n w) else signed_drop_bit (nat_of_integer n) (Rep_uint w))" including undefined_transfer integer.lifting unfolding uint_sshiftr_def by transfer simp code_printing constant uint_sshiftr \ (SML) "Uint.shiftr'_signed" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint.Int) _)) :: Uint.Word)" and (OCaml) "Uint.shiftr'_signed" and (Scala) "Uint.shiftr'_signed" lemma uint_msb_test_bit: "msb x \ bit (x :: uint) wivs_index" by transfer (simp add: msb_word_iff_bit wivs_index_def) lemma msb_uint_code [code]: "msb x \ uint_test_bit x wivs_index_integer" apply(simp add: uint_test_bit_def uint_msb_test_bit wivs_index_integer_code dflt_size_integer_def wivs_index_def) by (metis (full_types) One_nat_def dflt_size(2) less_iff_diff_less_0 nat_of_integer_of_nat of_nat_1 of_nat_diff of_nat_less_0_iff wivs_index_def) lemma uint_of_int_code [code]: "uint_of_int i = (BITS n. bit i n)" by transfer (simp add: word_of_int_conv_set_bits) section \Quickcheck setup\ definition uint_of_natural :: "natural \ uint" where "uint_of_natural x \ Uint (integer_of_natural x)" instantiation uint :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint \ qc_random_cnv uint_of_natural" definition "exhaustive_uint \ qc_exhaustive_cnv uint_of_natural" definition "full_exhaustive_uint \ qc_full_exhaustive_cnv uint_of_natural" instance .. end instantiation uint :: narrowing begin interpretation quickcheck_narrowing_samples "\i. (Uint i, Uint (- i))" "0" "Typerep.Typerep (STR ''Uint.uint'') []" . definition "narrowing_uint d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint itself \ _"]] lemmas partial_term_of_uint [code] = partial_term_of_code instance .. end end diff --git a/thys/Native_Word/Uint16.thy b/thys/Native_Word/Uint16.thy --- a/thys/Native_Word/Uint16.thy +++ b/thys/Native_Word/Uint16.thy @@ -1,532 +1,532 @@ (* Title: Uint16.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 16 bits\ theory Uint16 imports Code_Target_Word_Base Word_Type_Copies begin text \ Restriction for ML code generation: This theory assumes that the ML system provides a Word16 implementation (mlton does, but PolyML 5.5 does not). Therefore, the code setup lives in the target \SML_word\ rather than \SML\. This ensures that code generation still works as long as \uint16\ is not involved. For the target \SML\ itself, no special code generation for this type is set up. Nevertheless, it should work by emulation via \<^typ>\16 word\ if the theory \Code_Target_Bits_Int\ is imported. Restriction for OCaml code generation: OCaml does not provide an int16 type, so no special code generation for this type is set up. \ section \Type definition and primitive operations\ typedef uint16 = \UNIV :: 16 word set\ .. global_interpretation uint16: word_type_copy Abs_uint16 Rep_uint16 using type_definition_uint16 by (rule word_type_copy.intro) setup_lifting type_definition_uint16 declare uint16.of_word_of [code abstype] declare Quotient_uint16 [transfer_rule] instantiation uint16 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint16 :: uint16 is 0 . lift_definition one_uint16 :: uint16 is 1 . lift_definition plus_uint16 :: \uint16 \ uint16 \ uint16\ is \(+)\ . lift_definition uminus_uint16 :: \uint16 \ uint16\ is uminus . lift_definition minus_uint16 :: \uint16 \ uint16 \ uint16\ is \(-)\ . lift_definition times_uint16 :: \uint16 \ uint16 \ uint16\ is \(*)\ . lift_definition divide_uint16 :: \uint16 \ uint16 \ uint16\ is \(div)\ . lift_definition modulo_uint16 :: \uint16 \ uint16 \ uint16\ is \(mod)\ . lift_definition equal_uint16 :: \uint16 \ uint16 \ bool\ is \HOL.equal\ . lift_definition less_eq_uint16 :: \uint16 \ uint16 \ bool\ is \(\)\ . lift_definition less_uint16 :: \uint16 \ uint16 \ bool\ is \(<)\ . global_interpretation uint16: word_type_copy_ring Abs_uint16 Rep_uint16 by standard (fact zero_uint16.rep_eq one_uint16.rep_eq plus_uint16.rep_eq uminus_uint16.rep_eq minus_uint16.rep_eq times_uint16.rep_eq divide_uint16.rep_eq modulo_uint16.rep_eq equal_uint16.rep_eq less_eq_uint16.rep_eq less_uint16.rep_eq)+ instance proof - show \OFCLASS(uint16, comm_ring_1_class)\ by (rule uint16.of_class_comm_ring_1) show \OFCLASS(uint16, semiring_modulo_class)\ by (fact uint16.of_class_semiring_modulo) show \OFCLASS(uint16, equal_class)\ by (fact uint16.of_class_equal) show \OFCLASS(uint16, linorder_class)\ by (fact uint16.of_class_linorder) qed end instantiation uint16 :: ring_bit_operations begin lift_definition bit_uint16 :: \uint16 \ nat \ bool\ is bit . lift_definition not_uint16 :: \uint16 \ uint16\ is NOT . -lift_definition and_uint16 :: \uint16 \ uint16 \ uint16\ is \(AND)\ . -lift_definition or_uint16 :: \uint16 \ uint16 \ uint16\ is \(OR)\ . -lift_definition xor_uint16 :: \uint16 \ uint16 \ uint16\ is \(XOR)\ . +lift_definition and_uint16 :: \uint16 \ uint16 \ uint16\ is \Bit_Operations.and\ . +lift_definition or_uint16 :: \uint16 \ uint16 \ uint16\ is \Bit_Operations.or\ . +lift_definition xor_uint16 :: \uint16 \ uint16 \ uint16\ is \Bit_Operations.xor\ . lift_definition mask_uint16 :: \nat \ uint16\ is mask . lift_definition push_bit_uint16 :: \nat \ uint16 \ uint16\ is push_bit . lift_definition drop_bit_uint16 :: \nat \ uint16 \ uint16\ is drop_bit . lift_definition signed_drop_bit_uint16 :: \nat \ uint16 \ uint16\ is signed_drop_bit . lift_definition take_bit_uint16 :: \nat \ uint16 \ uint16\ is take_bit . lift_definition set_bit_uint16 :: \nat \ uint16 \ uint16\ is Bit_Operations.set_bit . lift_definition unset_bit_uint16 :: \nat \ uint16 \ uint16\ is unset_bit . lift_definition flip_bit_uint16 :: \nat \ uint16 \ uint16\ is flip_bit . global_interpretation uint16: word_type_copy_bits Abs_uint16 Rep_uint16 signed_drop_bit_uint16 by standard (fact bit_uint16.rep_eq not_uint16.rep_eq and_uint16.rep_eq or_uint16.rep_eq xor_uint16.rep_eq mask_uint16.rep_eq push_bit_uint16.rep_eq drop_bit_uint16.rep_eq signed_drop_bit_uint16.rep_eq take_bit_uint16.rep_eq set_bit_uint16.rep_eq unset_bit_uint16.rep_eq flip_bit_uint16.rep_eq)+ instance by (fact uint16.of_class_ring_bit_operations) end lift_definition uint16_of_nat :: \nat \ uint16\ is word_of_nat . lift_definition nat_of_uint16 :: \uint16 \ nat\ is unat . lift_definition uint16_of_int :: \int \ uint16\ is word_of_int . lift_definition int_of_uint16 :: \uint16 \ int\ is uint . context includes integer.lifting begin lift_definition Uint16 :: \integer \ uint16\ is word_of_int . lift_definition integer_of_uint16 :: \uint16 \ integer\ is uint . end global_interpretation uint16: word_type_copy_more Abs_uint16 Rep_uint16 signed_drop_bit_uint16 uint16_of_nat nat_of_uint16 uint16_of_int int_of_uint16 Uint16 integer_of_uint16 apply standard apply (simp_all add: uint16_of_nat.rep_eq nat_of_uint16.rep_eq uint16_of_int.rep_eq int_of_uint16.rep_eq Uint16.rep_eq integer_of_uint16.rep_eq integer_eq_iff) done instantiation uint16 :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint16 :: \uint16 \ nat\ is size . lift_definition msb_uint16 :: \uint16 \ bool\ is msb . lift_definition lsb_uint16 :: \uint16 \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint16 :: \uint16 \ nat \ bool \ uint16\ where set_bit_uint16_eq: \set_bit_uint16 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint16_transfer [transfer_rule]: \(cr_uint16 ===> (=) ===> (\) ===> cr_uint16) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint16_eq [abs_def]) transfer_prover end lift_definition set_bits_uint16 :: \(nat \ bool) \ uint16\ is set_bits . lift_definition set_bits_aux_uint16 :: \(nat \ bool) \ nat \ uint16 \ uint16\ is set_bits_aux . global_interpretation uint16: word_type_copy_misc Abs_uint16 Rep_uint16 signed_drop_bit_uint16 uint16_of_nat nat_of_uint16 uint16_of_int int_of_uint16 Uint16 integer_of_uint16 16 set_bits_aux_uint16 by (standard; transfer) simp_all instance using uint16.of_class_bit_comprehension uint16.of_class_set_bit uint16.of_class_lsb by simp_all standard end section \Code setup\ code_printing code_module Uint16 \ (SML_word) \(* Test that words can handle numbers between 0 and 15 *) val _ = if 4 <= Word.wordSize then () else raise (Fail ("wordSize less than 4")); structure Uint16 : sig val set_bit : Word16.word -> IntInf.int -> bool -> Word16.word val shiftl : Word16.word -> IntInf.int -> Word16.word val shiftr : Word16.word -> IntInf.int -> Word16.word val shiftr_signed : Word16.word -> IntInf.int -> Word16.word val test_bit : Word16.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word16.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word16.orb (x, mask) else Word16.andb (x, Word16.notb mask) end fun shiftl x n = Word16.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word16.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word16.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word16.andb (x, Word16.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word16.fromInt 0 end; (* struct Uint16 *)\ code_reserved SML_word Uint16 code_printing code_module Uint16 \ (Haskell) \module Uint16(Int16, Word16) where import Data.Int(Int16) import Data.Word(Word16)\ code_reserved Haskell Uint16 text \Scala provides unsigned 16-bit numbers as Char.\ code_printing code_module Uint16 \ (Scala) \object Uint16 { def set_bit(x: scala.Char, n: BigInt, b: Boolean) : scala.Char = if (b) (x | (1.toChar << n.intValue)).toChar else (x & (1.toChar << n.intValue).unary_~).toChar def shiftl(x: scala.Char, n: BigInt) : scala.Char = (x << n.intValue).toChar def shiftr(x: scala.Char, n: BigInt) : scala.Char = (x >>> n.intValue).toChar def shiftr_signed(x: scala.Char, n: BigInt) : scala.Char = (x.toShort >> n.intValue).toChar def test_bit(x: scala.Char, n: BigInt) : Boolean = (x & (1.toChar << n.intValue)) != 0 } /* object Uint16 */\ code_reserved Scala Uint16 text \ Avoid @{term Abs_uint16} in generated code, use @{term Rep_uint16'} instead. The symbolic implementations for code\_simp use @{term Rep_uint16}. The new destructor @{term Rep_uint16'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint16} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint16} ([code abstract] equations for @{typ uint16} may use @{term Rep_uint16} because these instances will be folded away.) To convert @{typ "16 word"} values into @{typ uint16}, use @{term "Abs_uint16'"}. \ definition Rep_uint16' where [simp]: "Rep_uint16' = Rep_uint16" lemma Rep_uint16'_transfer [transfer_rule]: "rel_fun cr_uint16 (=) (\x. x) Rep_uint16'" unfolding Rep_uint16'_def by(rule uint16.rep_transfer) lemma Rep_uint16'_code [code]: "Rep_uint16' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint16' :: "16 word \ uint16" is "\x :: 16 word. x" . lemma Abs_uint16'_code [code]: "Abs_uint16' x = Uint16 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint16 \ _"]] lemma term_of_uint16_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint16.uint16.Abs_uint16'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]]], TR (STR ''Uint16.uint16'') []])) (term_of_class.term_of (Rep_uint16' x))" by(simp add: term_of_anything) lemma Uint16_code [code]: "Rep_uint16 (Uint16 i) = word_of_int (int_of_integer_symbolic i)" unfolding Uint16_def int_of_integer_symbolic_def by(simp add: Abs_uint16_inverse) code_printing type_constructor uint16 \ (SML_word) "Word16.word" and (Haskell) "Uint16.Word16" and (Scala) "scala.Char" | constant Uint16 \ (SML_word) "Word16.fromLargeInt (IntInf.toLarge _)" and (Haskell) "(Prelude.fromInteger _ :: Uint16.Word16)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint16.Word16)" and (Scala) "_.charValue" | constant "0 :: uint16" \ (SML_word) "(Word16.fromInt 0)" and (Haskell) "(0 :: Uint16.Word16)" and (Scala) "0" | constant "1 :: uint16" \ (SML_word) "(Word16.fromInt 1)" and (Haskell) "(1 :: Uint16.Word16)" and (Scala) "1" | constant "plus :: uint16 \ _ \ _" \ (SML_word) "Word16.+ ((_), (_))" and (Haskell) infixl 6 "+" and (Scala) "(_ +/ _).toChar" | constant "uminus :: uint16 \ _" \ (SML_word) "Word16.~" and (Haskell) "negate" and (Scala) "(- _).toChar" | constant "minus :: uint16 \ _" \ (SML_word) "Word16.- ((_), (_))" and (Haskell) infixl 6 "-" and (Scala) "(_ -/ _).toChar" | constant "times :: uint16 \ _ \ _" \ (SML_word) "Word16.* ((_), (_))" and (Haskell) infixl 7 "*" and (Scala) "(_ */ _).toChar" | constant "HOL.equal :: uint16 \ _ \ bool" \ (SML_word) "!((_ : Word16.word) = _)" and (Haskell) infix 4 "==" and (Scala) infixl 5 "==" | class_instance uint16 :: equal \ (Haskell) - | constant "less_eq :: uint16 \ _ \ bool" \ (SML_word) "Word16.<= ((_), (_))" and (Haskell) infix 4 "<=" and (Scala) infixl 4 "<=" | constant "less :: uint16 \ _ \ bool" \ (SML_word) "Word16.< ((_), (_))" and (Haskell) infix 4 "<" and (Scala) infixl 4 "<" | constant "NOT :: uint16 \ _" \ (SML_word) "Word16.notb" and (Haskell) "Data'_Bits.complement" and (Scala) "_.unary'_~.toChar" -| constant "(AND) :: uint16 \ _" \ +| constant "Bit_Operations.and :: uint16 \ _" \ (SML_word) "Word16.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (Scala) "(_ & _).toChar" -| constant "(OR) :: uint16 \ _" \ +| constant "Bit_Operations.or :: uint16 \ _" \ (SML_word) "Word16.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (Scala) "(_ | _).toChar" -| constant "(XOR) :: uint16 \ _" \ +| constant "Bit_Operations.xor :: uint16 \ _" \ (SML_word) "Word16.xorb ((_),/ (_))" and (Haskell) "Data'_Bits.xor" and (Scala) "(_ ^ _).toChar" definition uint16_div :: "uint16 \ uint16 \ uint16" where "uint16_div x y = (if y = 0 then undefined ((div) :: uint16 \ _) x (0 :: uint16) else x div y)" definition uint16_mod :: "uint16 \ uint16 \ uint16" where "uint16_mod x y = (if y = 0 then undefined ((mod) :: uint16 \ _) x (0 :: uint16) else x mod y)" context includes undefined_transfer begin lemma div_uint16_code [code]: "x div y = (if y = 0 then 0 else uint16_div x y)" unfolding uint16_div_def by transfer (simp add: word_div_def) lemma mod_uint16_code [code]: "x mod y = (if y = 0 then x else uint16_mod x y)" unfolding uint16_mod_def by transfer (simp add: word_mod_def) lemma uint16_div_code [code]: "Rep_uint16 (uint16_div x y) = (if y = 0 then Rep_uint16 (undefined ((div) :: uint16 \ _) x (0 :: uint16)) else Rep_uint16 x div Rep_uint16 y)" unfolding uint16_div_def by transfer simp lemma uint16_mod_code [code]: "Rep_uint16 (uint16_mod x y) = (if y = 0 then Rep_uint16 (undefined ((mod) :: uint16 \ _) x (0 :: uint16)) else Rep_uint16 x mod Rep_uint16 y)" unfolding uint16_mod_def by transfer simp end code_printing constant uint16_div \ (SML_word) "Word16.div ((_), (_))" and (Haskell) "Prelude.div" and (Scala) "(_ '/ _).toChar" | constant uint16_mod \ (SML_word) "Word16.mod ((_), (_))" and (Haskell) "Prelude.mod" and (Scala) "(_ % _).toChar" definition uint16_test_bit :: "uint16 \ integer \ bool" where [code del]: "uint16_test_bit x n = (if n < 0 \ 15 < n then undefined (bit :: uint16 \ _) x n else bit x (nat_of_integer n))" lemma test_bit_uint16_code [code]: "bit x n \ n < 16 \ uint16_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint16_test_bit_def by (transfer, simp, transfer, simp) lemma uint16_test_bit_code [code]: "uint16_test_bit w n = (if n < 0 \ 15 < n then undefined (bit :: uint16 \ _) w n else bit (Rep_uint16 w) (nat_of_integer n))" unfolding uint16_test_bit_def by (simp add: bit_uint16.rep_eq) code_printing constant uint16_test_bit \ (SML_word) "Uint16.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (Scala) "Uint16.test'_bit" definition uint16_set_bit :: "uint16 \ integer \ bool \ uint16" where [code del]: "uint16_set_bit x n b = (if n < 0 \ 15 < n then undefined (set_bit :: uint16 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint16_code [code]: "set_bit x n b = (if n < 16 then uint16_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint16_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint16_set_bit_code [code]: "Rep_uint16 (uint16_set_bit w n b) = (if n < 0 \ 15 < n then Rep_uint16 (undefined (set_bit :: uint16 \ _) w n b) else set_bit (Rep_uint16 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint16_set_bit_def by transfer simp code_printing constant uint16_set_bit \ (SML_word) "Uint16.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (Scala) "Uint16.set'_bit" definition uint16_shiftl :: "uint16 \ integer \ uint16" where [code del]: "uint16_shiftl x n = (if n < 0 \ 16 \ n then undefined (push_bit :: nat \ uint16 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint16_code [code]: "push_bit n x = (if n < 16 then uint16_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint16_shiftl_def by transfer simp lemma uint16_shiftl_code [code]: "Rep_uint16 (uint16_shiftl w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined (push_bit :: nat \ uint16 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint16 w))" including undefined_transfer unfolding uint16_shiftl_def by transfer simp code_printing constant uint16_shiftl \ (SML_word) "Uint16.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (Scala) "Uint16.shiftl" definition uint16_shiftr :: "uint16 \ integer \ uint16" where [code del]: "uint16_shiftr x n = (if n < 0 \ 16 \ n then undefined (drop_bit :: nat \ uint16 \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint16_code [code]: "drop_bit n x = (if n < 16 then uint16_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint16_shiftr_def by transfer simp lemma uint16_shiftr_code [code]: "Rep_uint16 (uint16_shiftr w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined (drop_bit :: nat \ uint16 \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint16 w))" including undefined_transfer unfolding uint16_shiftr_def by transfer simp code_printing constant uint16_shiftr \ (SML_word) "Uint16.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (Scala) "Uint16.shiftr" definition uint16_sshiftr :: "uint16 \ integer \ uint16" where [code del]: "uint16_sshiftr x n = (if n < 0 \ 16 \ n then undefined signed_drop_bit_uint16 n x else signed_drop_bit_uint16 (nat_of_integer n) x)" lemma sshiftr_uint16_code [code]: "signed_drop_bit_uint16 n x = (if n < 16 then uint16_sshiftr x (integer_of_nat n) else if bit x 15 then -1 else 0)" including undefined_transfer integer.lifting unfolding uint16_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond word_size) lemma uint16_sshiftr_code [code]: "Rep_uint16 (uint16_sshiftr w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined signed_drop_bit_uint16 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint16 w))" including undefined_transfer unfolding uint16_sshiftr_def by transfer simp code_printing constant uint16_sshiftr \ (SML_word) "Uint16.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint16.Int16) _)) :: Uint16.Word16)" and (Scala) "Uint16.shiftr'_signed" lemma uint16_msb_test_bit: "msb x \ bit (x :: uint16) 15" by transfer (simp add: msb_word_iff_bit) lemma msb_uint16_code [code]: "msb x \ uint16_test_bit x 15" by (simp add: uint16_test_bit_def uint16_msb_test_bit) lemma uint16_of_int_code [code]: "uint16_of_int i = Uint16 (integer_of_int i)" including integer.lifting by transfer simp lemma int_of_uint16_code [code]: "int_of_uint16 x = int_of_integer (integer_of_uint16 x)" by (simp add: int_of_uint16.rep_eq integer_of_uint16_def) lemma nat_of_uint16_code [code]: "nat_of_uint16 x = nat_of_integer (integer_of_uint16 x)" unfolding integer_of_uint16_def including integer.lifting by transfer simp lemma integer_of_uint16_code [code]: "integer_of_uint16 n = integer_of_int (uint (Rep_uint16' n))" unfolding integer_of_uint16_def by transfer auto code_printing constant "integer_of_uint16" \ (SML_word) "Word16.toInt _ : IntInf.int" and (Haskell) "Prelude.toInteger" and (Scala) "BigInt" section \Quickcheck setup\ definition uint16_of_natural :: "natural \ uint16" where "uint16_of_natural x \ Uint16 (integer_of_natural x)" instantiation uint16 :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint16 \ qc_random_cnv uint16_of_natural" definition "exhaustive_uint16 \ qc_exhaustive_cnv uint16_of_natural" definition "full_exhaustive_uint16 \ qc_full_exhaustive_cnv uint16_of_natural" instance .. end instantiation uint16 :: narrowing begin interpretation quickcheck_narrowing_samples "\i. let x = Uint16 i in (x, 0xFFFF - x)" "0" "Typerep.Typerep (STR ''Uint16.uint16'') []" . definition "narrowing_uint16 d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint16 itself \ _"]] lemmas partial_term_of_uint16 [code] = partial_term_of_code instance .. end end diff --git a/thys/Native_Word/Uint32.thy b/thys/Native_Word/Uint32.thy --- a/thys/Native_Word/Uint32.thy +++ b/thys/Native_Word/Uint32.thy @@ -1,674 +1,685 @@ (* Title: Uint32.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 32 bits\ theory Uint32 imports Code_Target_Word_Base Word_Type_Copies begin section \Type definition and primitive operations\ typedef uint32 = \UNIV :: 32 word set\ .. global_interpretation uint32: word_type_copy Abs_uint32 Rep_uint32 using type_definition_uint32 by (rule word_type_copy.intro) setup_lifting type_definition_uint32 declare uint32.of_word_of [code abstype] declare Quotient_uint32 [transfer_rule] instantiation uint32 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint32 :: uint32 is 0 . lift_definition one_uint32 :: uint32 is 1 . lift_definition plus_uint32 :: \uint32 \ uint32 \ uint32\ is \(+)\ . lift_definition uminus_uint32 :: \uint32 \ uint32\ is uminus . lift_definition minus_uint32 :: \uint32 \ uint32 \ uint32\ is \(-)\ . lift_definition times_uint32 :: \uint32 \ uint32 \ uint32\ is \(*)\ . lift_definition divide_uint32 :: \uint32 \ uint32 \ uint32\ is \(div)\ . lift_definition modulo_uint32 :: \uint32 \ uint32 \ uint32\ is \(mod)\ . lift_definition equal_uint32 :: \uint32 \ uint32 \ bool\ is \HOL.equal\ . lift_definition less_eq_uint32 :: \uint32 \ uint32 \ bool\ is \(\)\ . lift_definition less_uint32 :: \uint32 \ uint32 \ bool\ is \(<)\ . global_interpretation uint32: word_type_copy_ring Abs_uint32 Rep_uint32 by standard (fact zero_uint32.rep_eq one_uint32.rep_eq plus_uint32.rep_eq uminus_uint32.rep_eq minus_uint32.rep_eq times_uint32.rep_eq divide_uint32.rep_eq modulo_uint32.rep_eq equal_uint32.rep_eq less_eq_uint32.rep_eq less_uint32.rep_eq)+ instance proof - show \OFCLASS(uint32, comm_ring_1_class)\ by (rule uint32.of_class_comm_ring_1) show \OFCLASS(uint32, semiring_modulo_class)\ by (fact uint32.of_class_semiring_modulo) show \OFCLASS(uint32, equal_class)\ by (fact uint32.of_class_equal) show \OFCLASS(uint32, linorder_class)\ by (fact uint32.of_class_linorder) qed end instantiation uint32 :: ring_bit_operations begin lift_definition bit_uint32 :: \uint32 \ nat \ bool\ is bit . lift_definition not_uint32 :: \uint32 \ uint32\ is NOT . -lift_definition and_uint32 :: \uint32 \ uint32 \ uint32\ is \(AND)\ . -lift_definition or_uint32 :: \uint32 \ uint32 \ uint32\ is \(OR)\ . -lift_definition xor_uint32 :: \uint32 \ uint32 \ uint32\ is \(XOR)\ . +lift_definition and_uint32 :: \uint32 \ uint32 \ uint32\ is \Bit_Operations.and\ . +lift_definition or_uint32 :: \uint32 \ uint32 \ uint32\ is \Bit_Operations.or\ . +lift_definition xor_uint32 :: \uint32 \ uint32 \ uint32\ is \Bit_Operations.xor\ . lift_definition mask_uint32 :: \nat \ uint32\ is mask . lift_definition push_bit_uint32 :: \nat \ uint32 \ uint32\ is push_bit . lift_definition drop_bit_uint32 :: \nat \ uint32 \ uint32\ is drop_bit . lift_definition signed_drop_bit_uint32 :: \nat \ uint32 \ uint32\ is signed_drop_bit . lift_definition take_bit_uint32 :: \nat \ uint32 \ uint32\ is take_bit . lift_definition set_bit_uint32 :: \nat \ uint32 \ uint32\ is Bit_Operations.set_bit . lift_definition unset_bit_uint32 :: \nat \ uint32 \ uint32\ is unset_bit . lift_definition flip_bit_uint32 :: \nat \ uint32 \ uint32\ is flip_bit . global_interpretation uint32: word_type_copy_bits Abs_uint32 Rep_uint32 signed_drop_bit_uint32 by standard (fact bit_uint32.rep_eq not_uint32.rep_eq and_uint32.rep_eq or_uint32.rep_eq xor_uint32.rep_eq mask_uint32.rep_eq push_bit_uint32.rep_eq drop_bit_uint32.rep_eq signed_drop_bit_uint32.rep_eq take_bit_uint32.rep_eq set_bit_uint32.rep_eq unset_bit_uint32.rep_eq flip_bit_uint32.rep_eq)+ instance by (fact uint32.of_class_ring_bit_operations) end lift_definition uint32_of_nat :: \nat \ uint32\ is word_of_nat . lift_definition nat_of_uint32 :: \uint32 \ nat\ is unat . lift_definition uint32_of_int :: \int \ uint32\ is word_of_int . lift_definition int_of_uint32 :: \uint32 \ int\ is uint . context includes integer.lifting begin lift_definition Uint32 :: \integer \ uint32\ is word_of_int . lift_definition integer_of_uint32 :: \uint32 \ integer\ is uint . end global_interpretation uint32: word_type_copy_more Abs_uint32 Rep_uint32 signed_drop_bit_uint32 uint32_of_nat nat_of_uint32 uint32_of_int int_of_uint32 Uint32 integer_of_uint32 apply standard apply (simp_all add: uint32_of_nat.rep_eq nat_of_uint32.rep_eq uint32_of_int.rep_eq int_of_uint32.rep_eq Uint32.rep_eq integer_of_uint32.rep_eq integer_eq_iff) done instantiation uint32 :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint32 :: \uint32 \ nat\ is size . lift_definition msb_uint32 :: \uint32 \ bool\ is msb . lift_definition lsb_uint32 :: \uint32 \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint32 :: \uint32 \ nat \ bool \ uint32\ where set_bit_uint32_eq: \set_bit_uint32 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint32_transfer [transfer_rule]: \(cr_uint32 ===> (=) ===> (\) ===> cr_uint32) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint32_eq [abs_def]) transfer_prover end lift_definition set_bits_uint32 :: \(nat \ bool) \ uint32\ is set_bits . lift_definition set_bits_aux_uint32 :: \(nat \ bool) \ nat \ uint32 \ uint32\ is set_bits_aux . global_interpretation uint32: word_type_copy_misc Abs_uint32 Rep_uint32 signed_drop_bit_uint32 uint32_of_nat nat_of_uint32 uint32_of_int int_of_uint32 Uint32 integer_of_uint32 32 set_bits_aux_uint32 by (standard; transfer) simp_all instance using uint32.of_class_bit_comprehension uint32.of_class_set_bit uint32.of_class_lsb by simp_all standard end section \Code setup\ code_printing code_module Uint32 \ (SML) \(* Test that words can handle numbers between 0 and 31 *) val _ = if 5 <= Word.wordSize then () else raise (Fail ("wordSize less than 5")); structure Uint32 : sig val set_bit : Word32.word -> IntInf.int -> bool -> Word32.word val shiftl : Word32.word -> IntInf.int -> Word32.word val shiftr : Word32.word -> IntInf.int -> Word32.word val shiftr_signed : Word32.word -> IntInf.int -> Word32.word val test_bit : Word32.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word32.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word32.orb (x, mask) else Word32.andb (x, Word32.notb mask) end fun shiftl x n = Word32.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word32.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word32.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word32.andb (x, Word32.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word32.fromInt 0 end; (* struct Uint32 *)\ code_reserved SML Uint32 code_printing code_module Uint32 \ (Haskell) \module Uint32(Int32, Word32) where import Data.Int(Int32) import Data.Word(Word32)\ code_reserved Haskell Uint32 text \ OCaml and Scala provide only signed 32bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module "Uint32" \ (OCaml) \module Uint32 : sig val less : int32 -> int32 -> bool val less_eq : int32 -> int32 -> bool val set_bit : int32 -> Z.t -> bool -> int32 val shiftl : int32 -> Z.t -> int32 val shiftr : int32 -> Z.t -> int32 val shiftr_signed : int32 -> Z.t -> int32 val test_bit : int32 -> Z.t -> bool end = struct (* negative numbers have their highest bit set, so they are greater than positive ones *) let less x y = if Int32.compare x Int32.zero < 0 then Int32.compare y Int32.zero < 0 && Int32.compare x y < 0 else Int32.compare y Int32.zero < 0 || Int32.compare x y < 0;; let less_eq x y = if Int32.compare x Int32.zero < 0 then Int32.compare y Int32.zero < 0 && Int32.compare x y <= 0 else Int32.compare y Int32.zero < 0 || Int32.compare x y <= 0;; let set_bit x n b = let mask = Int32.shift_left Int32.one (Z.to_int n) in if b then Int32.logor x mask else Int32.logand x (Int32.lognot mask);; let shiftl x n = Int32.shift_left x (Z.to_int n);; let shiftr x n = Int32.shift_right_logical x (Z.to_int n);; let shiftr_signed x n = Int32.shift_right x (Z.to_int n);; let test_bit x n = Int32.compare (Int32.logand x (Int32.shift_left Int32.one (Z.to_int n))) Int32.zero <> 0;; end;; (*struct Uint32*)\ code_reserved OCaml Uint32 code_printing code_module Uint32 \ (Scala) \object Uint32 { def less(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Int, n: BigInt, b: Boolean) : Int = if (b) x | (1 << n.intValue) else x & (1 << n.intValue).unary_~ def shiftl(x: Int, n: BigInt) : Int = x << n.intValue def shiftr(x: Int, n: BigInt) : Int = x >>> n.intValue def shiftr_signed(x: Int, n: BigInt) : Int = x >> n.intValue def test_bit(x: Int, n: BigInt) : Boolean = (x & (1 << n.intValue)) != 0 } /* object Uint32 */\ code_reserved Scala Uint32 text \ OCaml's conversion from Big\_int to int32 demands that the value fits int a signed 32-bit integer. The following justifies the implementation. \ +context + includes bit_operations_syntax +begin + definition Uint32_signed :: "integer \ uint32" where "Uint32_signed i = (if i < -(0x80000000) \ i \ 0x80000000 then undefined Uint32 i else Uint32 i)" lemma Uint32_code [code]: "Uint32 i = (let i' = i AND 0xFFFFFFFF in if bit i' 31 then Uint32_signed (i' - 0x100000000) else Uint32_signed i')" including undefined_transfer integer.lifting unfolding Uint32_signed_def apply transfer apply (subst word_of_int_via_signed) apply (auto simp add: push_bit_of_1 mask_eq_exp_minus_1 word_of_int_via_signed cong del: if_cong) done lemma Uint32_signed_code [code]: "Rep_uint32 (Uint32_signed i) = (if i < -(0x80000000) \ i \ 0x80000000 then Rep_uint32 (undefined Uint32 i) else word_of_int (int_of_integer_symbolic i))" unfolding Uint32_signed_def Uint32_def int_of_integer_symbolic_def word_of_integer_def by(simp add: Abs_uint32_inverse) +end + text \ Avoid @{term Abs_uint32} in generated code, use @{term Rep_uint32'} instead. The symbolic implementations for code\_simp use @{term Rep_uint32}. The new destructor @{term Rep_uint32'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint32} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint32} ([code abstract] equations for @{typ uint32} may use @{term Rep_uint32} because these instances will be folded away.) To convert @{typ "32 word"} values into @{typ uint32}, use @{term "Abs_uint32'"}. \ definition Rep_uint32' where [simp]: "Rep_uint32' = Rep_uint32" lemma Rep_uint32'_transfer [transfer_rule]: "rel_fun cr_uint32 (=) (\x. x) Rep_uint32'" unfolding Rep_uint32'_def by(rule uint32.rep_transfer) lemma Rep_uint32'_code [code]: "Rep_uint32' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint32' :: "32 word \ uint32" is "\x :: 32 word. x" . lemma Abs_uint32'_code [code]: "Abs_uint32' x = Uint32 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint32 \ _"]] lemma term_of_uint32_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint32.uint32.Abs_uint32'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]]]], TR (STR ''Uint32.uint32'') []])) (term_of_class.term_of (Rep_uint32' x))" by(simp add: term_of_anything) code_printing type_constructor uint32 \ (SML) "Word32.word" and (Haskell) "Uint32.Word32" and (OCaml) "int32" and (Scala) "Int" and (Eval) "Word32.word" | constant Uint32 \ (SML) "Word32.fromLargeInt (IntInf.toLarge _)" and (Haskell) "(Prelude.fromInteger _ :: Uint32.Word32)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint32.Word32)" and (Scala) "_.intValue" | constant Uint32_signed \ (OCaml) "Z.to'_int32" | constant "0 :: uint32" \ (SML) "(Word32.fromInt 0)" and (Haskell) "(0 :: Uint32.Word32)" and (OCaml) "Int32.zero" and (Scala) "0" | constant "1 :: uint32" \ (SML) "(Word32.fromInt 1)" and (Haskell) "(1 :: Uint32.Word32)" and (OCaml) "Int32.one" and (Scala) "1" | constant "plus :: uint32 \ _ " \ (SML) "Word32.+ ((_), (_))" and (Haskell) infixl 6 "+" and (OCaml) "Int32.add" and (Scala) infixl 7 "+" | constant "uminus :: uint32 \ _" \ (SML) "Word32.~" and (Haskell) "negate" and (OCaml) "Int32.neg" and (Scala) "!(- _)" | constant "minus :: uint32 \ _" \ (SML) "Word32.- ((_), (_))" and (Haskell) infixl 6 "-" and (OCaml) "Int32.sub" and (Scala) infixl 7 "-" | constant "times :: uint32 \ _ \ _" \ (SML) "Word32.* ((_), (_))" and (Haskell) infixl 7 "*" and (OCaml) "Int32.mul" and (Scala) infixl 8 "*" | constant "HOL.equal :: uint32 \ _ \ bool" \ (SML) "!((_ : Word32.word) = _)" and (Haskell) infix 4 "==" and (OCaml) "(Int32.compare _ _ = 0)" and (Scala) infixl 5 "==" | class_instance uint32 :: equal \ (Haskell) - | constant "less_eq :: uint32 \ _ \ bool" \ (SML) "Word32.<= ((_), (_))" and (Haskell) infix 4 "<=" and (OCaml) "Uint32.less'_eq" and (Scala) "Uint32.less'_eq" | constant "less :: uint32 \ _ \ bool" \ (SML) "Word32.< ((_), (_))" and (Haskell) infix 4 "<" and (OCaml) "Uint32.less" and (Scala) "Uint32.less" | constant "NOT :: uint32 \ _" \ (SML) "Word32.notb" and (Haskell) "Data'_Bits.complement" and (OCaml) "Int32.lognot" and (Scala) "_.unary'_~" -| constant "(AND) :: uint32 \ _" \ +| constant "Bit_Operations.and :: uint32 \ _" \ (SML) "Word32.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (OCaml) "Int32.logand" and (Scala) infixl 3 "&" -| constant "(OR) :: uint32 \ _" \ +| constant "Bit_Operations.or :: uint32 \ _" \ (SML) "Word32.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (OCaml) "Int32.logor" and (Scala) infixl 1 "|" -| constant "(XOR) :: uint32 \ _" \ +| constant "Bit_Operations.xor :: uint32 \ _" \ (SML) "Word32.xorb ((_),/ (_))" and (Haskell) "Data'_Bits.xor" and (OCaml) "Int32.logxor" and (Scala) infixl 2 "^" definition uint32_divmod :: "uint32 \ uint32 \ uint32 \ uint32" where "uint32_divmod x y = (if y = 0 then (undefined ((div) :: uint32 \ _) x (0 :: uint32), undefined ((mod) :: uint32 \ _) x (0 :: uint32)) else (x div y, x mod y))" definition uint32_div :: "uint32 \ uint32 \ uint32" where "uint32_div x y = fst (uint32_divmod x y)" definition uint32_mod :: "uint32 \ uint32 \ uint32" where "uint32_mod x y = snd (uint32_divmod x y)" lemma div_uint32_code [code]: "x div y = (if y = 0 then 0 else uint32_div x y)" including undefined_transfer unfolding uint32_divmod_def uint32_div_def by transfer (simp add: word_div_def) lemma mod_uint32_code [code]: "x mod y = (if y = 0 then x else uint32_mod x y)" including undefined_transfer unfolding uint32_mod_def uint32_divmod_def by transfer (simp add: word_mod_def) definition uint32_sdiv :: "uint32 \ uint32 \ uint32" where [code del]: "uint32_sdiv x y = (if y = 0 then undefined ((div) :: uint32 \ _) x (0 :: uint32) else Abs_uint32 (Rep_uint32 x sdiv Rep_uint32 y))" definition div0_uint32 :: "uint32 \ uint32" where [code del]: "div0_uint32 x = undefined ((div) :: uint32 \ _) x (0 :: uint32)" declare [[code abort: div0_uint32]] definition mod0_uint32 :: "uint32 \ uint32" where [code del]: "mod0_uint32 x = undefined ((mod) :: uint32 \ _) x (0 :: uint32)" declare [[code abort: mod0_uint32]] lemma uint32_divmod_code [code]: "uint32_divmod x y = (if 0x80000000 \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint32 x, mod0_uint32 x) else let q = push_bit 1 (uint32_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" including undefined_transfer unfolding uint32_divmod_def uint32_sdiv_def div0_uint32_def mod0_uint32_def less_eq_uint32.rep_eq apply transfer apply (simp add: divmod_via_sdivmod push_bit_eq_mult) done lemma uint32_sdiv_code [code]: "Rep_uint32 (uint32_sdiv x y) = (if y = 0 then Rep_uint32 (undefined ((div) :: uint32 \ _) x (0 :: uint32)) else Rep_uint32 x sdiv Rep_uint32 y)" unfolding uint32_sdiv_def by(simp add: Abs_uint32_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint32_divmod_code} computes both with division only. \ code_printing constant uint32_div \ (SML) "Word32.div ((_), (_))" and (Haskell) "Prelude.div" | constant uint32_mod \ (SML) "Word32.mod ((_), (_))" and (Haskell) "Prelude.mod" | constant uint32_divmod \ (Haskell) "divmod" | constant uint32_sdiv \ (OCaml) "Int32.div" and (Scala) "_ '/ _" definition uint32_test_bit :: "uint32 \ integer \ bool" where [code del]: "uint32_test_bit x n = (if n < 0 \ 31 < n then undefined (bit :: uint32 \ _) x n else bit x (nat_of_integer n))" lemma test_bit_uint32_code [code]: "bit x n \ n < 32 \ uint32_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint32_test_bit_def by (transfer, simp, transfer, simp) lemma uint32_test_bit_code [code]: "uint32_test_bit w n = (if n < 0 \ 31 < n then undefined (bit :: uint32 \ _) w n else bit (Rep_uint32 w) (nat_of_integer n))" unfolding uint32_test_bit_def by(simp add: bit_uint32.rep_eq) code_printing constant uint32_test_bit \ (SML) "Uint32.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (OCaml) "Uint32.test'_bit" and (Scala) "Uint32.test'_bit" and (Eval) "(fn w => fn n => if n < 0 orelse 32 <= n then raise (Fail \"argument to uint32'_test'_bit out of bounds\") else Uint32.test'_bit w n)" definition uint32_set_bit :: "uint32 \ integer \ bool \ uint32" where [code del]: "uint32_set_bit x n b = (if n < 0 \ 31 < n then undefined (set_bit :: uint32 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint32_code [code]: "set_bit x n b = (if n < 32 then uint32_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint32_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint32_set_bit_code [code]: "Rep_uint32 (uint32_set_bit w n b) = (if n < 0 \ 31 < n then Rep_uint32 (undefined (set_bit :: uint32 \ _) w n b) else set_bit (Rep_uint32 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint32_set_bit_def by transfer simp code_printing constant uint32_set_bit \ (SML) "Uint32.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (OCaml) "Uint32.set'_bit" and (Scala) "Uint32.set'_bit" and (Eval) "(fn w => fn n => fn b => if n < 0 orelse 32 <= n then raise (Fail \"argument to uint32'_set'_bit out of bounds\") else Uint32.set'_bit x n b)" definition uint32_shiftl :: "uint32 \ integer \ uint32" where [code del]: "uint32_shiftl x n = (if n < 0 \ 32 \ n then undefined (push_bit :: nat \ uint32 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint32_code [code]: "push_bit n x = (if n < 32 then uint32_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint32_shiftl_def by transfer simp lemma uint32_shiftl_code [code]: "Rep_uint32 (uint32_shiftl w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined (push_bit :: nat \ uint32 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint32 w))" including undefined_transfer unfolding uint32_shiftl_def by transfer simp code_printing constant uint32_shiftl \ (SML) "Uint32.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (OCaml) "Uint32.shiftl" and (Scala) "Uint32.shiftl" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 32 then raise Fail \"argument to uint32'_shiftl out of bounds\" else Uint32.shiftl x i)" definition uint32_shiftr :: "uint32 \ integer \ uint32" where [code del]: "uint32_shiftr x n = (if n < 0 \ 32 \ n then undefined (drop_bit :: nat \ uint32 \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint32_code [code]: "drop_bit n x = (if n < 32 then uint32_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint32_shiftr_def by transfer simp lemma uint32_shiftr_code [code]: "Rep_uint32 (uint32_shiftr w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined (drop_bit :: nat \ uint32 \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint32 w))" including undefined_transfer unfolding uint32_shiftr_def by transfer simp code_printing constant uint32_shiftr \ (SML) "Uint32.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (OCaml) "Uint32.shiftr" and (Scala) "Uint32.shiftr" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 32 then raise Fail \"argument to uint32'_shiftr out of bounds\" else Uint32.shiftr x i)" definition uint32_sshiftr :: "uint32 \ integer \ uint32" where [code del]: "uint32_sshiftr x n = (if n < 0 \ 32 \ n then undefined signed_drop_bit_uint32 n x else signed_drop_bit_uint32 (nat_of_integer n) x)" lemma sshiftr_uint32_code [code]: "signed_drop_bit_uint32 n x = (if n < 32 then uint32_sshiftr x (integer_of_nat n) else if bit x 31 then - 1 else 0)" including undefined_transfer integer.lifting unfolding uint32_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond) lemma uint32_sshiftr_code [code]: "Rep_uint32 (uint32_sshiftr w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined signed_drop_bit_uint32 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint32 w))" including undefined_transfer unfolding uint32_sshiftr_def by transfer simp code_printing constant uint32_sshiftr \ (SML) "Uint32.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint32.Int32) _)) :: Uint32.Word32)" and (OCaml) "Uint32.shiftr'_signed" and (Scala) "Uint32.shiftr'_signed" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 32 then raise Fail \"argument to uint32'_shiftr'_signed out of bounds\" else Uint32.shiftr'_signed x i)" +context + includes bit_operations_syntax +begin + lemma uint32_msb_test_bit: "msb x \ bit (x :: uint32) 31" by transfer (simp add: msb_word_iff_bit) lemma msb_uint32_code [code]: "msb x \ uint32_test_bit x 31" by (simp add: uint32_test_bit_def uint32_msb_test_bit) lemma uint32_of_int_code [code]: "uint32_of_int i = Uint32 (integer_of_int i)" including integer.lifting by transfer simp lemma int_of_uint32_code [code]: "int_of_uint32 x = int_of_integer (integer_of_uint32 x)" including integer.lifting by transfer simp lemma nat_of_uint32_code [code]: "nat_of_uint32 x = nat_of_integer (integer_of_uint32 x)" unfolding integer_of_uint32_def including integer.lifting by transfer simp - definition integer_of_uint32_signed :: "uint32 \ integer" where "integer_of_uint32_signed n = (if bit n 31 then undefined integer_of_uint32 n else integer_of_uint32 n)" lemma integer_of_uint32_signed_code [code]: "integer_of_uint32_signed n = (if bit n 31 then undefined integer_of_uint32 n else integer_of_int (uint (Rep_uint32' n)))" by (simp add: integer_of_uint32_signed_def integer_of_uint32_def) lemma integer_of_uint32_code [code]: "integer_of_uint32 n = (if bit n 31 then integer_of_uint32_signed (n AND 0x7FFFFFFF) OR 0x80000000 else integer_of_uint32_signed n)" proof - have \integer_of_uint32_signed (n AND 0x7FFFFFFF) OR 0x80000000 = Bit_Operations.set_bit 31 (integer_of_uint32_signed (take_bit 31 n))\ by (simp add: take_bit_eq_mask set_bit_eq_or push_bit_eq_mult mask_eq_exp_minus_1) moreover have \integer_of_uint32 n = Bit_Operations.set_bit 31 (integer_of_uint32 (take_bit 31 n))\ if \bit n 31\ proof (rule bit_eqI) fix m from that show \bit (integer_of_uint32 n) m = bit (Bit_Operations.set_bit 31 (integer_of_uint32 (take_bit 31 n))) m\ for m including integer.lifting by transfer (auto simp add: bit_simps dest: bit_imp_le_length) qed ultimately show ?thesis by simp (simp add: integer_of_uint32_signed_def bit_simps) qed +end + code_printing constant "integer_of_uint32" \ (SML) "IntInf.fromLarge (Word32.toLargeInt _) : IntInf.int" and (Haskell) "Prelude.toInteger" | constant "integer_of_uint32_signed" \ (OCaml) "Z.of'_int32" and (Scala) "BigInt" section \Quickcheck setup\ definition uint32_of_natural :: "natural \ uint32" where "uint32_of_natural x \ Uint32 (integer_of_natural x)" instantiation uint32 :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint32 \ qc_random_cnv uint32_of_natural" definition "exhaustive_uint32 \ qc_exhaustive_cnv uint32_of_natural" definition "full_exhaustive_uint32 \ qc_full_exhaustive_cnv uint32_of_natural" instance .. end instantiation uint32 :: narrowing begin interpretation quickcheck_narrowing_samples "\i. let x = Uint32 i in (x, 0xFFFFFFFF - x)" "0" "Typerep.Typerep (STR ''Uint32.uint32'') []" . definition "narrowing_uint32 d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint32 itself \ _"]] lemmas partial_term_of_uint32 [code] = partial_term_of_code instance .. end end diff --git a/thys/Native_Word/Uint64.thy b/thys/Native_Word/Uint64.thy --- a/thys/Native_Word/Uint64.thy +++ b/thys/Native_Word/Uint64.thy @@ -1,871 +1,883 @@ (* Title: Uint64.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 64 bits\ theory Uint64 imports Code_Target_Word_Base Word_Type_Copies begin text \ PolyML (in version 5.7) provides a Word64 structure only when run in 64-bit mode. Therefore, we by default provide an implementation of 64-bit words using \verb$IntInf.int$ and masking. The code target \texttt{SML\_word} replaces this implementation and maps the operations directly to the \verb$Word64$ structure provided by the Standard ML implementations. The \verb$Eval$ target used by @{command value} and @{method eval} dynamically tests at runtime for the version of PolyML and uses PolyML's Word64 structure if it detects a 64-bit version which does not suffer from a division bug found in PolyML 5.6. \ section \Type definition and primitive operations\ typedef uint64 = \UNIV :: 64 word set\ .. global_interpretation uint64: word_type_copy Abs_uint64 Rep_uint64 using type_definition_uint64 by (rule word_type_copy.intro) setup_lifting type_definition_uint64 declare uint64.of_word_of [code abstype] declare Quotient_uint64 [transfer_rule] instantiation uint64 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint64 :: uint64 is 0 . lift_definition one_uint64 :: uint64 is 1 . lift_definition plus_uint64 :: \uint64 \ uint64 \ uint64\ is \(+)\ . lift_definition uminus_uint64 :: \uint64 \ uint64\ is uminus . lift_definition minus_uint64 :: \uint64 \ uint64 \ uint64\ is \(-)\ . lift_definition times_uint64 :: \uint64 \ uint64 \ uint64\ is \(*)\ . lift_definition divide_uint64 :: \uint64 \ uint64 \ uint64\ is \(div)\ . lift_definition modulo_uint64 :: \uint64 \ uint64 \ uint64\ is \(mod)\ . lift_definition equal_uint64 :: \uint64 \ uint64 \ bool\ is \HOL.equal\ . lift_definition less_eq_uint64 :: \uint64 \ uint64 \ bool\ is \(\)\ . lift_definition less_uint64 :: \uint64 \ uint64 \ bool\ is \(<)\ . global_interpretation uint64: word_type_copy_ring Abs_uint64 Rep_uint64 by standard (fact zero_uint64.rep_eq one_uint64.rep_eq plus_uint64.rep_eq uminus_uint64.rep_eq minus_uint64.rep_eq times_uint64.rep_eq divide_uint64.rep_eq modulo_uint64.rep_eq equal_uint64.rep_eq less_eq_uint64.rep_eq less_uint64.rep_eq)+ instance proof - show \OFCLASS(uint64, comm_ring_1_class)\ by (rule uint64.of_class_comm_ring_1) show \OFCLASS(uint64, semiring_modulo_class)\ by (fact uint64.of_class_semiring_modulo) show \OFCLASS(uint64, equal_class)\ by (fact uint64.of_class_equal) show \OFCLASS(uint64, linorder_class)\ by (fact uint64.of_class_linorder) qed end instantiation uint64 :: ring_bit_operations begin lift_definition bit_uint64 :: \uint64 \ nat \ bool\ is bit . lift_definition not_uint64 :: \uint64 \ uint64\ is NOT . -lift_definition and_uint64 :: \uint64 \ uint64 \ uint64\ is \(AND)\ . -lift_definition or_uint64 :: \uint64 \ uint64 \ uint64\ is \(OR)\ . -lift_definition xor_uint64 :: \uint64 \ uint64 \ uint64\ is \(XOR)\ . +lift_definition and_uint64 :: \uint64 \ uint64 \ uint64\ is \Bit_Operations.and\ . +lift_definition or_uint64 :: \uint64 \ uint64 \ uint64\ is \Bit_Operations.or\ . +lift_definition xor_uint64 :: \uint64 \ uint64 \ uint64\ is \Bit_Operations.xor\ . lift_definition mask_uint64 :: \nat \ uint64\ is mask . lift_definition push_bit_uint64 :: \nat \ uint64 \ uint64\ is push_bit . lift_definition drop_bit_uint64 :: \nat \ uint64 \ uint64\ is drop_bit . lift_definition signed_drop_bit_uint64 :: \nat \ uint64 \ uint64\ is signed_drop_bit . lift_definition take_bit_uint64 :: \nat \ uint64 \ uint64\ is take_bit . lift_definition set_bit_uint64 :: \nat \ uint64 \ uint64\ is Bit_Operations.set_bit . lift_definition unset_bit_uint64 :: \nat \ uint64 \ uint64\ is unset_bit . lift_definition flip_bit_uint64 :: \nat \ uint64 \ uint64\ is flip_bit . global_interpretation uint64: word_type_copy_bits Abs_uint64 Rep_uint64 signed_drop_bit_uint64 by standard (fact bit_uint64.rep_eq not_uint64.rep_eq and_uint64.rep_eq or_uint64.rep_eq xor_uint64.rep_eq mask_uint64.rep_eq push_bit_uint64.rep_eq drop_bit_uint64.rep_eq signed_drop_bit_uint64.rep_eq take_bit_uint64.rep_eq set_bit_uint64.rep_eq unset_bit_uint64.rep_eq flip_bit_uint64.rep_eq)+ instance by (fact uint64.of_class_ring_bit_operations) end lift_definition uint64_of_nat :: \nat \ uint64\ is word_of_nat . lift_definition nat_of_uint64 :: \uint64 \ nat\ is unat . lift_definition uint64_of_int :: \int \ uint64\ is word_of_int . lift_definition int_of_uint64 :: \uint64 \ int\ is uint . context includes integer.lifting begin lift_definition Uint64 :: \integer \ uint64\ is word_of_int . lift_definition integer_of_uint64 :: \uint64 \ integer\ is uint . end global_interpretation uint64: word_type_copy_more Abs_uint64 Rep_uint64 signed_drop_bit_uint64 uint64_of_nat nat_of_uint64 uint64_of_int int_of_uint64 Uint64 integer_of_uint64 apply standard apply (simp_all add: uint64_of_nat.rep_eq nat_of_uint64.rep_eq uint64_of_int.rep_eq int_of_uint64.rep_eq Uint64.rep_eq integer_of_uint64.rep_eq integer_eq_iff) done instantiation uint64 :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint64 :: \uint64 \ nat\ is size . lift_definition msb_uint64 :: \uint64 \ bool\ is msb . lift_definition lsb_uint64 :: \uint64 \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint64 :: \uint64 \ nat \ bool \ uint64\ where set_bit_uint64_eq: \set_bit_uint64 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint64_transfer [transfer_rule]: \(cr_uint64 ===> (=) ===> (\) ===> cr_uint64) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint64_eq [abs_def]) transfer_prover end lift_definition set_bits_uint64 :: \(nat \ bool) \ uint64\ is set_bits . lift_definition set_bits_aux_uint64 :: \(nat \ bool) \ nat \ uint64 \ uint64\ is set_bits_aux . global_interpretation uint64: word_type_copy_misc Abs_uint64 Rep_uint64 signed_drop_bit_uint64 uint64_of_nat nat_of_uint64 uint64_of_int int_of_uint64 Uint64 integer_of_uint64 64 set_bits_aux_uint64 by (standard; transfer) simp_all instance using uint64.of_class_bit_comprehension uint64.of_class_set_bit uint64.of_class_lsb by simp_all standard end section \Code setup\ text \ For SML, we generate an implementation of unsigned 64-bit words using \verb$IntInf.int$. If @{ML "LargeWord.wordSize > 63"} of the Isabelle/ML runtime environment holds, then we assume that there is also a \Word64\ structure available and accordingly replace the implementation for the target \verb$Eval$. \ code_printing code_module "Uint64" \ (SML) \(* Test that words can handle numbers between 0 and 63 *) val _ = if 6 <= Word.wordSize then () else raise (Fail ("wordSize less than 6")); structure Uint64 : sig eqtype uint64; val zero : uint64; val one : uint64; val fromInt : IntInf.int -> uint64; val toInt : uint64 -> IntInf.int; val toLarge : uint64 -> LargeWord.word; val fromLarge : LargeWord.word -> uint64 val plus : uint64 -> uint64 -> uint64; val minus : uint64 -> uint64 -> uint64; val times : uint64 -> uint64 -> uint64; val divide : uint64 -> uint64 -> uint64; val modulus : uint64 -> uint64 -> uint64; val negate : uint64 -> uint64; val less_eq : uint64 -> uint64 -> bool; val less : uint64 -> uint64 -> bool; val notb : uint64 -> uint64; val andb : uint64 -> uint64 -> uint64; val orb : uint64 -> uint64 -> uint64; val xorb : uint64 -> uint64 -> uint64; val shiftl : uint64 -> IntInf.int -> uint64; val shiftr : uint64 -> IntInf.int -> uint64; val shiftr_signed : uint64 -> IntInf.int -> uint64; val set_bit : uint64 -> IntInf.int -> bool -> uint64; val test_bit : uint64 -> IntInf.int -> bool; end = struct type uint64 = IntInf.int; val mask = 0xFFFFFFFFFFFFFFFF : IntInf.int; val zero = 0 : IntInf.int; val one = 1 : IntInf.int; fun fromInt x = IntInf.andb(x, mask); fun toInt x = x fun toLarge x = LargeWord.fromLargeInt (IntInf.toLarge x); fun fromLarge x = IntInf.fromLarge (LargeWord.toLargeInt x); fun plus x y = IntInf.andb(IntInf.+(x, y), mask); fun minus x y = IntInf.andb(IntInf.-(x, y), mask); fun negate x = IntInf.andb(IntInf.~(x), mask); fun times x y = IntInf.andb(IntInf.*(x, y), mask); fun divide x y = IntInf.div(x, y); fun modulus x y = IntInf.mod(x, y); fun less_eq x y = IntInf.<=(x, y); fun less x y = IntInf.<(x, y); fun notb x = IntInf.andb(IntInf.notb(x), mask); fun orb x y = IntInf.orb(x, y); fun andb x y = IntInf.andb(x, y); fun xorb x y = IntInf.xorb(x, y); val maxWord = IntInf.pow (2, Word.wordSize); fun shiftl x n = if n < maxWord then IntInf.andb(IntInf.<< (x, Word.fromLargeInt (IntInf.toLarge n)), mask) else 0; fun shiftr x n = if n < maxWord then IntInf.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) else 0; val msb_mask = 0x8000000000000000 : IntInf.int; fun shiftr_signed x i = if IntInf.andb(x, msb_mask) = 0 then shiftr x i else if i >= 64 then 0xFFFFFFFFFFFFFFFF else let val x' = shiftr x i val m' = IntInf.andb(IntInf.<<(mask, Word.max(0w64 - Word.fromLargeInt (IntInf.toLarge i), 0w0)), mask) in IntInf.orb(x', m') end; fun test_bit x n = if n < maxWord then IntInf.andb (x, IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n))) <> 0 else false; fun set_bit x n b = if n < 64 then if b then IntInf.orb (x, IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n))) else IntInf.andb (x, IntInf.notb (IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n)))) else x; end \ code_reserved SML Uint64 setup \ let val polyml64 = LargeWord.wordSize > 63; (* PolyML 5.6 has bugs in its Word64 implementation. We test for one such bug and refrain from using Word64 in that case. Testing is done with dynamic code evaluation such that the compiler does not choke on the Word64 structure, which need not be present in a 32bit environment. *) val error_msg = "Buggy Word64 structure"; val test_code = "val _ = if Word64.div (0w18446744073709551611 : Word64.word, 0w3) = 0w6148914691236517203 then ()\n" ^ "else raise (Fail \"" ^ error_msg ^ "\");"; val f = Exn.interruptible_capture (fn () => ML_Compiler.eval ML_Compiler.flags Position.none (ML_Lex.tokenize test_code)) val use_Word64 = polyml64 andalso (case f () of Exn.Res _ => true | Exn.Exn (e as ERROR m) => if String.isSuffix error_msg m then false else Exn.reraise e | Exn.Exn e => Exn.reraise e) ; val newline = "\n"; val content = "structure Uint64 : sig" ^ newline ^ " eqtype uint64;" ^ newline ^ " val zero : uint64;" ^ newline ^ " val one : uint64;" ^ newline ^ " val fromInt : IntInf.int -> uint64;" ^ newline ^ " val toInt : uint64 -> IntInf.int;" ^ newline ^ " val toLarge : uint64 -> LargeWord.word;" ^ newline ^ " val fromLarge : LargeWord.word -> uint64" ^ newline ^ " val plus : uint64 -> uint64 -> uint64;" ^ newline ^ " val minus : uint64 -> uint64 -> uint64;" ^ newline ^ " val times : uint64 -> uint64 -> uint64;" ^ newline ^ " val divide : uint64 -> uint64 -> uint64;" ^ newline ^ " val modulus : uint64 -> uint64 -> uint64;" ^ newline ^ " val negate : uint64 -> uint64;" ^ newline ^ " val less_eq : uint64 -> uint64 -> bool;" ^ newline ^ " val less : uint64 -> uint64 -> bool;" ^ newline ^ " val notb : uint64 -> uint64;" ^ newline ^ " val andb : uint64 -> uint64 -> uint64;" ^ newline ^ " val orb : uint64 -> uint64 -> uint64;" ^ newline ^ " val xorb : uint64 -> uint64 -> uint64;" ^ newline ^ " val shiftl : uint64 -> IntInf.int -> uint64;" ^ newline ^ " val shiftr : uint64 -> IntInf.int -> uint64;" ^ newline ^ " val shiftr_signed : uint64 -> IntInf.int -> uint64;" ^ newline ^ " val set_bit : uint64 -> IntInf.int -> bool -> uint64;" ^ newline ^ " val test_bit : uint64 -> IntInf.int -> bool;" ^ newline ^ "end = struct" ^ newline ^ "" ^ newline ^ "type uint64 = Word64.word;" ^ newline ^ "" ^ newline ^ "val zero = (0wx0 : uint64);" ^ newline ^ "" ^ newline ^ "val one = (0wx1 : uint64);" ^ newline ^ "" ^ newline ^ "fun fromInt x = Word64.fromLargeInt (IntInf.toLarge x);" ^ newline ^ "" ^ newline ^ "fun toInt x = IntInf.fromLarge (Word64.toLargeInt x);" ^ newline ^ "" ^ newline ^ "fun fromLarge x = Word64.fromLarge x;" ^ newline ^ "" ^ newline ^ "fun toLarge x = Word64.toLarge x;" ^ newline ^ "" ^ newline ^ "fun plus x y = Word64.+(x, y);" ^ newline ^ "" ^ newline ^ "fun minus x y = Word64.-(x, y);" ^ newline ^ "" ^ newline ^ "fun negate x = Word64.~(x);" ^ newline ^ "" ^ newline ^ "fun times x y = Word64.*(x, y);" ^ newline ^ "" ^ newline ^ "fun divide x y = Word64.div(x, y);" ^ newline ^ "" ^ newline ^ "fun modulus x y = Word64.mod(x, y);" ^ newline ^ "" ^ newline ^ "fun less_eq x y = Word64.<=(x, y);" ^ newline ^ "" ^ newline ^ "fun less x y = Word64.<(x, y);" ^ newline ^ "" ^ newline ^ "fun set_bit x n b =" ^ newline ^ " let val mask = Word64.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ " in if b then Word64.orb (x, mask)" ^ newline ^ " else Word64.andb (x, Word64.notb mask)" ^ newline ^ " end" ^ newline ^ "" ^ newline ^ "fun shiftl x n =" ^ newline ^ " Word64.<< (x, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ "" ^ newline ^ "fun shiftr x n =" ^ newline ^ " Word64.>> (x, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ "" ^ newline ^ "fun shiftr_signed x n =" ^ newline ^ " Word64.~>> (x, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ "" ^ newline ^ "fun test_bit x n =" ^ newline ^ " Word64.andb (x, Word64.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word64.fromInt 0" ^ newline ^ "" ^ newline ^ "val notb = Word64.notb" ^ newline ^ "" ^ newline ^ "fun andb x y = Word64.andb(x, y);" ^ newline ^ "" ^ newline ^ "fun orb x y = Word64.orb(x, y);" ^ newline ^ "" ^ newline ^ "fun xorb x y = Word64.xorb(x, y);" ^ newline ^ "" ^ newline ^ "end (*struct Uint64*)" val target_SML64 = "SML_word"; in (if use_Word64 then Code_Target.set_printings (Code_Symbol.Module ("Uint64", [(Code_Runtime.target, SOME (content, []))])) else I) #> Code_Target.set_printings (Code_Symbol.Module ("Uint64", [(target_SML64, SOME (content, []))])) end \ code_printing code_module Uint64 \ (Haskell) \module Uint64(Int64, Word64) where import Data.Int(Int64) import Data.Word(Word64)\ code_reserved Haskell Uint64 text \ OCaml and Scala provide only signed 64bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module "Uint64" \ (OCaml) \module Uint64 : sig val less : int64 -> int64 -> bool val less_eq : int64 -> int64 -> bool val set_bit : int64 -> Z.t -> bool -> int64 val shiftl : int64 -> Z.t -> int64 val shiftr : int64 -> Z.t -> int64 val shiftr_signed : int64 -> Z.t -> int64 val test_bit : int64 -> Z.t -> bool end = struct (* negative numbers have their highest bit set, so they are greater than positive ones *) let less x y = if Int64.compare x Int64.zero < 0 then Int64.compare y Int64.zero < 0 && Int64.compare x y < 0 else Int64.compare y Int64.zero < 0 || Int64.compare x y < 0;; let less_eq x y = if Int64.compare x Int64.zero < 0 then Int64.compare y Int64.zero < 0 && Int64.compare x y <= 0 else Int64.compare y Int64.zero < 0 || Int64.compare x y <= 0;; let set_bit x n b = let mask = Int64.shift_left Int64.one (Z.to_int n) in if b then Int64.logor x mask else Int64.logand x (Int64.lognot mask);; let shiftl x n = Int64.shift_left x (Z.to_int n);; let shiftr x n = Int64.shift_right_logical x (Z.to_int n);; let shiftr_signed x n = Int64.shift_right x (Z.to_int n);; let test_bit x n = Int64.compare (Int64.logand x (Int64.shift_left Int64.one (Z.to_int n))) Int64.zero <> 0;; end;; (*struct Uint64*)\ code_reserved OCaml Uint64 code_printing code_module Uint64 \ (Scala) \object Uint64 { def less(x: Long, y: Long) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Long, y: Long) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Long, n: BigInt, b: Boolean) : Long = if (b) x | (1L << n.intValue) else x & (1L << n.intValue).unary_~ def shiftl(x: Long, n: BigInt) : Long = x << n.intValue def shiftr(x: Long, n: BigInt) : Long = x >>> n.intValue def shiftr_signed(x: Long, n: BigInt) : Long = x >> n.intValue def test_bit(x: Long, n: BigInt) : Boolean = (x & (1L << n.intValue)) != 0 } /* object Uint64 */\ code_reserved Scala Uint64 text \ OCaml's conversion from Big\_int to int64 demands that the value fits int a signed 64-bit integer. The following justifies the implementation. \ +context + includes bit_operations_syntax +begin + definition Uint64_signed :: "integer \ uint64" where "Uint64_signed i = (if i < -(0x8000000000000000) \ i \ 0x8000000000000000 then undefined Uint64 i else Uint64 i)" lemma Uint64_code [code]: "Uint64 i = (let i' = i AND 0xFFFFFFFFFFFFFFFF in if bit i' 63 then Uint64_signed (i' - 0x10000000000000000) else Uint64_signed i')" including undefined_transfer integer.lifting unfolding Uint64_signed_def apply transfer apply (subst word_of_int_via_signed) apply (auto simp add: push_bit_of_1 mask_eq_exp_minus_1 word_of_int_via_signed cong del: if_cong) done lemma Uint64_signed_code [code]: "Rep_uint64 (Uint64_signed i) = (if i < -(0x8000000000000000) \ i \ 0x8000000000000000 then Rep_uint64 (undefined Uint64 i) else word_of_int (int_of_integer_symbolic i))" unfolding Uint64_signed_def Uint64_def int_of_integer_symbolic_def word_of_integer_def by(simp add: Abs_uint64_inverse) +end + text \ Avoid @{term Abs_uint64} in generated code, use @{term Rep_uint64'} instead. The symbolic implementations for code\_simp use @{term Rep_uint64}. The new destructor @{term Rep_uint64'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint64} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint64} ([code abstract] equations for @{typ uint64} may use @{term Rep_uint64} because these instances will be folded away.) To convert @{typ "64 word"} values into @{typ uint64}, use @{term "Abs_uint64'"}. \ definition Rep_uint64' where [simp]: "Rep_uint64' = Rep_uint64" lemma Rep_uint64'_transfer [transfer_rule]: "rel_fun cr_uint64 (=) (\x. x) Rep_uint64'" unfolding Rep_uint64'_def by(rule uint64.rep_transfer) lemma Rep_uint64'_code [code]: "Rep_uint64' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint64' :: "64 word \ uint64" is "\x :: 64 word. x" . lemma Abs_uint64'_code [code]: "Abs_uint64' x = Uint64 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint64 \ _"]] lemma term_of_uint64_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint64.uint64.Abs_uint64'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]]]]], TR (STR ''Uint64.uint64'') []])) (term_of_class.term_of (Rep_uint64' x))" by(simp add: term_of_anything) code_printing type_constructor uint64 \ (SML) "Uint64.uint64" and (Haskell) "Uint64.Word64" and (OCaml) "int64" and (Scala) "Long" | constant Uint64 \ (SML) "Uint64.fromInt" and (Haskell) "(Prelude.fromInteger _ :: Uint64.Word64)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint64.Word64)" and (Scala) "_.longValue" | constant Uint64_signed \ (OCaml) "Z.to'_int64" | constant "0 :: uint64" \ (SML) "Uint64.zero" and (Haskell) "(0 :: Uint64.Word64)" and (OCaml) "Int64.zero" and (Scala) "0" | constant "1 :: uint64" \ (SML) "Uint64.one" and (Haskell) "(1 :: Uint64.Word64)" and (OCaml) "Int64.one" and (Scala) "1" | constant "plus :: uint64 \ _ " \ (SML) "Uint64.plus" and (Haskell) infixl 6 "+" and (OCaml) "Int64.add" and (Scala) infixl 7 "+" | constant "uminus :: uint64 \ _" \ (SML) "Uint64.negate" and (Haskell) "negate" and (OCaml) "Int64.neg" and (Scala) "!(- _)" | constant "minus :: uint64 \ _" \ (SML) "Uint64.minus" and (Haskell) infixl 6 "-" and (OCaml) "Int64.sub" and (Scala) infixl 7 "-" | constant "times :: uint64 \ _ \ _" \ (SML) "Uint64.times" and (Haskell) infixl 7 "*" and (OCaml) "Int64.mul" and (Scala) infixl 8 "*" | constant "HOL.equal :: uint64 \ _ \ bool" \ (SML) "!((_ : Uint64.uint64) = _)" and (Haskell) infix 4 "==" and (OCaml) "(Int64.compare _ _ = 0)" and (Scala) infixl 5 "==" | class_instance uint64 :: equal \ (Haskell) - | constant "less_eq :: uint64 \ _ \ bool" \ (SML) "Uint64.less'_eq" and (Haskell) infix 4 "<=" and (OCaml) "Uint64.less'_eq" and (Scala) "Uint64.less'_eq" | constant "less :: uint64 \ _ \ bool" \ (SML) "Uint64.less" and (Haskell) infix 4 "<" and (OCaml) "Uint64.less" and (Scala) "Uint64.less" | constant "NOT :: uint64 \ _" \ (SML) "Uint64.notb" and (Haskell) "Data'_Bits.complement" and (OCaml) "Int64.lognot" and (Scala) "_.unary'_~" -| constant "(AND) :: uint64 \ _" \ +| constant "Bit_Operations.and :: uint64 \ _" \ (SML) "Uint64.andb" and (Haskell) infixl 7 "Data_Bits..&." and (OCaml) "Int64.logand" and (Scala) infixl 3 "&" -| constant "(OR) :: uint64 \ _" \ +| constant "Bit_Operations.or :: uint64 \ _" \ (SML) "Uint64.orb" and (Haskell) infixl 5 "Data_Bits..|." and (OCaml) "Int64.logor" and (Scala) infixl 1 "|" -| constant "(XOR) :: uint64 \ _" \ +| constant "Bit_Operations.xor :: uint64 \ _" \ (SML) "Uint64.xorb" and (Haskell) "Data'_Bits.xor" and (OCaml) "Int64.logxor" and (Scala) infixl 2 "^" definition uint64_divmod :: "uint64 \ uint64 \ uint64 \ uint64" where "uint64_divmod x y = (if y = 0 then (undefined ((div) :: uint64 \ _) x (0 :: uint64), undefined ((mod) :: uint64 \ _) x (0 :: uint64)) else (x div y, x mod y))" definition uint64_div :: "uint64 \ uint64 \ uint64" where "uint64_div x y = fst (uint64_divmod x y)" definition uint64_mod :: "uint64 \ uint64 \ uint64" where "uint64_mod x y = snd (uint64_divmod x y)" lemma div_uint64_code [code]: "x div y = (if y = 0 then 0 else uint64_div x y)" including undefined_transfer unfolding uint64_divmod_def uint64_div_def by transfer (simp add: word_div_def) lemma mod_uint64_code [code]: "x mod y = (if y = 0 then x else uint64_mod x y)" including undefined_transfer unfolding uint64_mod_def uint64_divmod_def by transfer (simp add: word_mod_def) definition uint64_sdiv :: "uint64 \ uint64 \ uint64" where [code del]: "uint64_sdiv x y = (if y = 0 then undefined ((div) :: uint64 \ _) x (0 :: uint64) else Abs_uint64 (Rep_uint64 x sdiv Rep_uint64 y))" definition div0_uint64 :: "uint64 \ uint64" where [code del]: "div0_uint64 x = undefined ((div) :: uint64 \ _) x (0 :: uint64)" declare [[code abort: div0_uint64]] definition mod0_uint64 :: "uint64 \ uint64" where [code del]: "mod0_uint64 x = undefined ((mod) :: uint64 \ _) x (0 :: uint64)" declare [[code abort: mod0_uint64]] lemma uint64_divmod_code [code]: "uint64_divmod x y = (if 0x8000000000000000 \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint64 x, mod0_uint64 x) else let q = push_bit 1 (uint64_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" including undefined_transfer unfolding uint64_divmod_def uint64_sdiv_def div0_uint64_def mod0_uint64_def less_eq_uint64.rep_eq apply transfer apply (simp add: divmod_via_sdivmod push_bit_eq_mult) done lemma uint64_sdiv_code [code]: "Rep_uint64 (uint64_sdiv x y) = (if y = 0 then Rep_uint64 (undefined ((div) :: uint64 \ _) x (0 :: uint64)) else Rep_uint64 x sdiv Rep_uint64 y)" unfolding uint64_sdiv_def by(simp add: Abs_uint64_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint64_divmod_code} computes both with division only. \ code_printing constant uint64_div \ (SML) "Uint64.divide" and (Haskell) "Prelude.div" | constant uint64_mod \ (SML) "Uint64.modulus" and (Haskell) "Prelude.mod" | constant uint64_divmod \ (Haskell) "divmod" | constant uint64_sdiv \ (OCaml) "Int64.div" and (Scala) "_ '/ _" definition uint64_test_bit :: "uint64 \ integer \ bool" where [code del]: "uint64_test_bit x n = (if n < 0 \ 63 < n then undefined (bit :: uint64 \ _) x n else bit x (nat_of_integer n))" lemma bit_uint64_code [code]: "bit x n \ n < 64 \ uint64_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint64_test_bit_def by transfer (auto dest: bit_imp_le_length) lemma uint64_test_bit_code [code]: "uint64_test_bit w n = (if n < 0 \ 63 < n then undefined (bit :: uint64 \ _) w n else bit (Rep_uint64 w) (nat_of_integer n))" unfolding uint64_test_bit_def by(simp add: bit_uint64.rep_eq) code_printing constant uint64_test_bit \ (SML) "Uint64.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (OCaml) "Uint64.test'_bit" and (Scala) "Uint64.test'_bit" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_test'_bit out of bounds\") else Uint64.test'_bit x i)" definition uint64_set_bit :: "uint64 \ integer \ bool \ uint64" where [code del]: "uint64_set_bit x n b = (if n < 0 \ 63 < n then undefined (set_bit :: uint64 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint64_code [code]: "set_bit x n b = (if n < 64 then uint64_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint64_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint64_set_bit_code [code]: "Rep_uint64 (uint64_set_bit w n b) = (if n < 0 \ 63 < n then Rep_uint64 (undefined (set_bit :: uint64 \ _) w n b) else set_bit (Rep_uint64 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint64_set_bit_def by transfer simp code_printing constant uint64_set_bit \ (SML) "Uint64.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (OCaml) "Uint64.set'_bit" and (Scala) "Uint64.set'_bit" and (Eval) "(fn x => fn i => fn b => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_set'_bit out of bounds\") else Uint64.set'_bit x i b)" definition uint64_shiftl :: "uint64 \ integer \ uint64" where [code del]: "uint64_shiftl x n = (if n < 0 \ 64 \ n then undefined (push_bit :: nat \ uint64 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint64_code [code]: "push_bit n x = (if n < 64 then uint64_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint64_shiftl_def by transfer simp lemma uint64_shiftl_code [code]: "Rep_uint64 (uint64_shiftl w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined (push_bit :: nat \ uint64 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint64 w))" including undefined_transfer unfolding uint64_shiftl_def by transfer simp code_printing constant uint64_shiftl \ (SML) "Uint64.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (OCaml) "Uint64.shiftl" and (Scala) "Uint64.shiftl" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_shiftl out of bounds\") else Uint64.shiftl x i)" definition uint64_shiftr :: "uint64 \ integer \ uint64" where [code del]: "uint64_shiftr x n = (if n < 0 \ 64 \ n then undefined (drop_bit :: nat \ uint64 \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint64_code [code]: "drop_bit n x = (if n < 64 then uint64_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint64_shiftr_def by transfer simp lemma uint64_shiftr_code [code]: "Rep_uint64 (uint64_shiftr w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined (drop_bit :: nat \ uint64 \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint64 w))" including undefined_transfer unfolding uint64_shiftr_def by transfer simp code_printing constant uint64_shiftr \ (SML) "Uint64.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (OCaml) "Uint64.shiftr" and (Scala) "Uint64.shiftr" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_shiftr out of bounds\") else Uint64.shiftr x i)" definition uint64_sshiftr :: "uint64 \ integer \ uint64" where [code del]: "uint64_sshiftr x n = (if n < 0 \ 64 \ n then undefined signed_drop_bit_uint64 n x else signed_drop_bit_uint64 (nat_of_integer n) x)" lemma sshiftr_uint64_code [code]: "signed_drop_bit_uint64 n x = (if n < 64 then uint64_sshiftr x (integer_of_nat n) else if bit x 63 then - 1 else 0)" including undefined_transfer integer.lifting unfolding uint64_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond) lemma uint64_sshiftr_code [code]: "Rep_uint64 (uint64_sshiftr w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined signed_drop_bit_uint64 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint64 w))" including undefined_transfer unfolding uint64_sshiftr_def by transfer simp code_printing constant uint64_sshiftr \ (SML) "Uint64.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint64.Int64) _)) :: Uint64.Word64)" and (OCaml) "Uint64.shiftr'_signed" and (Scala) "Uint64.shiftr'_signed" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_shiftr'_signed out of bounds\") else Uint64.shiftr'_signed x i)" +context + includes bit_operations_syntax +begin + lemma uint64_msb_test_bit: "msb x \ bit (x :: uint64) 63" by transfer (simp add: msb_word_iff_bit) lemma msb_uint64_code [code]: "msb x \ uint64_test_bit x 63" by (simp add: uint64_test_bit_def uint64_msb_test_bit) lemma uint64_of_int_code [code]: "uint64_of_int i = Uint64 (integer_of_int i)" including integer.lifting by transfer simp lemma int_of_uint64_code [code]: "int_of_uint64 x = int_of_integer (integer_of_uint64 x)" including integer.lifting by transfer simp lemma nat_of_uint64_code [code]: "nat_of_uint64 x = nat_of_integer (integer_of_uint64 x)" unfolding integer_of_uint64_def including integer.lifting by transfer simp definition integer_of_uint64_signed :: "uint64 \ integer" where "integer_of_uint64_signed n = (if bit n 63 then undefined integer_of_uint64 n else integer_of_uint64 n)" lemma integer_of_uint64_signed_code [code]: "integer_of_uint64_signed n = (if bit n 63 then undefined integer_of_uint64 n else integer_of_int (uint (Rep_uint64' n)))" by (simp add: integer_of_uint64_signed_def integer_of_uint64_def) lemma integer_of_uint64_code [code]: "integer_of_uint64 n = (if bit n 63 then integer_of_uint64_signed (n AND 0x7FFFFFFFFFFFFFFF) OR 0x8000000000000000 else integer_of_uint64_signed n)" proof - have \integer_of_uint64_signed (n AND 0x7FFFFFFFFFFFFFFF) OR 0x8000000000000000 = Bit_Operations.set_bit 63 (integer_of_uint64_signed (take_bit 63 n))\ by (simp add: take_bit_eq_mask set_bit_eq_or push_bit_eq_mult mask_eq_exp_minus_1) moreover have \integer_of_uint64 n = Bit_Operations.set_bit 63 (integer_of_uint64 (take_bit 63 n))\ if \bit n 63\ proof (rule bit_eqI) fix m from that show \bit (integer_of_uint64 n) m = bit (Bit_Operations.set_bit 63 (integer_of_uint64 (take_bit 63 n))) m\ for m including integer.lifting by transfer (auto simp add: bit_simps dest: bit_imp_le_length) qed ultimately show ?thesis by simp (simp add: integer_of_uint64_signed_def bit_simps) qed +end + code_printing constant "integer_of_uint64" \ (SML) "Uint64.toInt" and (Haskell) "Prelude.toInteger" | constant "integer_of_uint64_signed" \ (OCaml) "Z.of'_int64" and (Scala) "BigInt" section \Quickcheck setup\ definition uint64_of_natural :: "natural \ uint64" where "uint64_of_natural x \ Uint64 (integer_of_natural x)" instantiation uint64 :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint64 \ qc_random_cnv uint64_of_natural" definition "exhaustive_uint64 \ qc_exhaustive_cnv uint64_of_natural" definition "full_exhaustive_uint64 \ qc_full_exhaustive_cnv uint64_of_natural" instance .. end instantiation uint64 :: narrowing begin interpretation quickcheck_narrowing_samples "\i. let x = Uint64 i in (x, 0xFFFFFFFFFFFFFFFF - x)" "0" "Typerep.Typerep (STR ''Uint64.uint64'') []" . definition "narrowing_uint64 d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint64 itself \ _"]] lemmas partial_term_of_uint64 [code] = partial_term_of_code instance .. end end diff --git a/thys/Native_Word/Uint8.thy b/thys/Native_Word/Uint8.thy --- a/thys/Native_Word/Uint8.thy +++ b/thys/Native_Word/Uint8.thy @@ -1,599 +1,605 @@ (* Title: Uint8.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 8 bits\ theory Uint8 imports Code_Target_Word_Base Word_Type_Copies begin text \ Restriction for OCaml code generation: OCaml does not provide an int8 type, so no special code generation for this type is set up. If the theory \Code_Target_Bits_Int\ is imported, the type \uint8\ is emulated via \<^typ>\8 word\. \ section \Type definition and primitive operations\ typedef uint8 = \UNIV :: 8 word set\ .. global_interpretation uint8: word_type_copy Abs_uint8 Rep_uint8 using type_definition_uint8 by (rule word_type_copy.intro) setup_lifting type_definition_uint8 declare uint8.of_word_of [code abstype] declare Quotient_uint8 [transfer_rule] instantiation uint8 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint8 :: uint8 is 0 . lift_definition one_uint8 :: uint8 is 1 . lift_definition plus_uint8 :: \uint8 \ uint8 \ uint8\ is \(+)\ . lift_definition uminus_uint8 :: \uint8 \ uint8\ is uminus . lift_definition minus_uint8 :: \uint8 \ uint8 \ uint8\ is \(-)\ . lift_definition times_uint8 :: \uint8 \ uint8 \ uint8\ is \(*)\ . lift_definition divide_uint8 :: \uint8 \ uint8 \ uint8\ is \(div)\ . lift_definition modulo_uint8 :: \uint8 \ uint8 \ uint8\ is \(mod)\ . lift_definition equal_uint8 :: \uint8 \ uint8 \ bool\ is \HOL.equal\ . lift_definition less_eq_uint8 :: \uint8 \ uint8 \ bool\ is \(\)\ . lift_definition less_uint8 :: \uint8 \ uint8 \ bool\ is \(<)\ . global_interpretation uint8: word_type_copy_ring Abs_uint8 Rep_uint8 by standard (fact zero_uint8.rep_eq one_uint8.rep_eq plus_uint8.rep_eq uminus_uint8.rep_eq minus_uint8.rep_eq times_uint8.rep_eq divide_uint8.rep_eq modulo_uint8.rep_eq equal_uint8.rep_eq less_eq_uint8.rep_eq less_uint8.rep_eq)+ instance proof - show \OFCLASS(uint8, comm_ring_1_class)\ by (rule uint8.of_class_comm_ring_1) show \OFCLASS(uint8, semiring_modulo_class)\ by (fact uint8.of_class_semiring_modulo) show \OFCLASS(uint8, equal_class)\ by (fact uint8.of_class_equal) show \OFCLASS(uint8, linorder_class)\ by (fact uint8.of_class_linorder) qed end instantiation uint8 :: ring_bit_operations begin lift_definition bit_uint8 :: \uint8 \ nat \ bool\ is bit . lift_definition not_uint8 :: \uint8 \ uint8\ is NOT . -lift_definition and_uint8 :: \uint8 \ uint8 \ uint8\ is \(AND)\ . -lift_definition or_uint8 :: \uint8 \ uint8 \ uint8\ is \(OR)\ . -lift_definition xor_uint8 :: \uint8 \ uint8 \ uint8\ is \(XOR)\ . +lift_definition and_uint8 :: \uint8 \ uint8 \ uint8\ is \Bit_Operations.and\ . +lift_definition or_uint8 :: \uint8 \ uint8 \ uint8\ is \Bit_Operations.or\ . +lift_definition xor_uint8 :: \uint8 \ uint8 \ uint8\ is \Bit_Operations.xor\ . lift_definition mask_uint8 :: \nat \ uint8\ is mask . lift_definition push_bit_uint8 :: \nat \ uint8 \ uint8\ is push_bit . lift_definition drop_bit_uint8 :: \nat \ uint8 \ uint8\ is drop_bit . lift_definition signed_drop_bit_uint8 :: \nat \ uint8 \ uint8\ is signed_drop_bit . lift_definition take_bit_uint8 :: \nat \ uint8 \ uint8\ is take_bit . lift_definition set_bit_uint8 :: \nat \ uint8 \ uint8\ is Bit_Operations.set_bit . lift_definition unset_bit_uint8 :: \nat \ uint8 \ uint8\ is unset_bit . lift_definition flip_bit_uint8 :: \nat \ uint8 \ uint8\ is flip_bit . global_interpretation uint8: word_type_copy_bits Abs_uint8 Rep_uint8 signed_drop_bit_uint8 by standard (fact bit_uint8.rep_eq not_uint8.rep_eq and_uint8.rep_eq or_uint8.rep_eq xor_uint8.rep_eq mask_uint8.rep_eq push_bit_uint8.rep_eq drop_bit_uint8.rep_eq signed_drop_bit_uint8.rep_eq take_bit_uint8.rep_eq set_bit_uint8.rep_eq unset_bit_uint8.rep_eq flip_bit_uint8.rep_eq)+ instance by (fact uint8.of_class_ring_bit_operations) end lift_definition uint8_of_nat :: \nat \ uint8\ is word_of_nat . lift_definition nat_of_uint8 :: \uint8 \ nat\ is unat . lift_definition uint8_of_int :: \int \ uint8\ is word_of_int . lift_definition int_of_uint8 :: \uint8 \ int\ is uint . context includes integer.lifting begin lift_definition Uint8 :: \integer \ uint8\ is word_of_int . lift_definition integer_of_uint8 :: \uint8 \ integer\ is uint . end global_interpretation uint8: word_type_copy_more Abs_uint8 Rep_uint8 signed_drop_bit_uint8 uint8_of_nat nat_of_uint8 uint8_of_int int_of_uint8 Uint8 integer_of_uint8 apply standard apply (simp_all add: uint8_of_nat.rep_eq nat_of_uint8.rep_eq uint8_of_int.rep_eq int_of_uint8.rep_eq Uint8.rep_eq integer_of_uint8.rep_eq integer_eq_iff) done instantiation uint8 :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint8 :: \uint8 \ nat\ is size . lift_definition msb_uint8 :: \uint8 \ bool\ is msb . lift_definition lsb_uint8 :: \uint8 \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint8 :: \uint8 \ nat \ bool \ uint8\ where set_bit_uint8_eq: \set_bit_uint8 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint8_transfer [transfer_rule]: \(cr_uint8 ===> (=) ===> (\) ===> cr_uint8) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint8_eq [abs_def]) transfer_prover end lift_definition set_bits_uint8 :: \(nat \ bool) \ uint8\ is set_bits . lift_definition set_bits_aux_uint8 :: \(nat \ bool) \ nat \ uint8 \ uint8\ is set_bits_aux . global_interpretation uint8: word_type_copy_misc Abs_uint8 Rep_uint8 signed_drop_bit_uint8 uint8_of_nat nat_of_uint8 uint8_of_int int_of_uint8 Uint8 integer_of_uint8 8 set_bits_aux_uint8 by (standard; transfer) simp_all instance using uint8.of_class_bit_comprehension uint8.of_class_set_bit uint8.of_class_lsb by simp_all standard end section \Code setup\ code_printing code_module Uint8 \ (SML) \(* Test that words can handle numbers between 0 and 3 *) val _ = if 3 <= Word.wordSize then () else raise (Fail ("wordSize less than 3")); structure Uint8 : sig val set_bit : Word8.word -> IntInf.int -> bool -> Word8.word val shiftl : Word8.word -> IntInf.int -> Word8.word val shiftr : Word8.word -> IntInf.int -> Word8.word val shiftr_signed : Word8.word -> IntInf.int -> Word8.word val test_bit : Word8.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word8.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word8.orb (x, mask) else Word8.andb (x, Word8.notb mask) end fun shiftl x n = Word8.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word8.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word8.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word8.andb (x, Word8.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word8.fromInt 0 end; (* struct Uint8 *)\ code_reserved SML Uint8 code_printing code_module Uint8 \ (Haskell) \module Uint8(Int8, Word8) where import Data.Int(Int8) import Data.Word(Word8)\ code_reserved Haskell Uint8 text \ Scala provides only signed 8bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module Uint8 \ (Scala) \object Uint8 { def less(x: Byte, y: Byte) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Byte, y: Byte) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Byte, n: BigInt, b: Boolean) : Byte = if (b) (x | (1 << n.intValue)).toByte else (x & (1 << n.intValue).unary_~).toByte def shiftl(x: Byte, n: BigInt) : Byte = (x << n.intValue).toByte def shiftr(x: Byte, n: BigInt) : Byte = ((x & 255) >>> n.intValue).toByte def shiftr_signed(x: Byte, n: BigInt) : Byte = (x >> n.intValue).toByte def test_bit(x: Byte, n: BigInt) : Boolean = (x & (1 << n.intValue)) != 0 } /* object Uint8 */\ code_reserved Scala Uint8 text \ Avoid @{term Abs_uint8} in generated code, use @{term Rep_uint8'} instead. The symbolic implementations for code\_simp use @{term Rep_uint8}. The new destructor @{term Rep_uint8'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint8} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint8} ([code abstract] equations for @{typ uint8} may use @{term Rep_uint8} because these instances will be folded away.) To convert @{typ "8 word"} values into @{typ uint8}, use @{term "Abs_uint8'"}. \ definition Rep_uint8' where [simp]: "Rep_uint8' = Rep_uint8" lemma Rep_uint8'_transfer [transfer_rule]: "rel_fun cr_uint8 (=) (\x. x) Rep_uint8'" unfolding Rep_uint8'_def by(rule uint8.rep_transfer) lemma Rep_uint8'_code [code]: "Rep_uint8' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint8' :: "8 word \ uint8" is "\x :: 8 word. x" . lemma Abs_uint8'_code [code]: "Abs_uint8' x = Uint8 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint8 \ _"]] lemma term_of_uint8_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint8.uint8.Abs_uint8'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]], TR (STR ''Uint8.uint8'') []])) (term_of_class.term_of (Rep_uint8' x))" by(simp add: term_of_anything) lemma Uin8_code [code]: "Rep_uint8 (Uint8 i) = word_of_int (int_of_integer_symbolic i)" unfolding Uint8_def int_of_integer_symbolic_def by(simp add: Abs_uint8_inverse) code_printing type_constructor uint8 \ (SML) "Word8.word" and (Haskell) "Uint8.Word8" and (Scala) "Byte" | constant Uint8 \ (SML) "Word8.fromLargeInt (IntInf.toLarge _)" and (Haskell) "(Prelude.fromInteger _ :: Uint8.Word8)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint8.Word8)" and (Scala) "_.byteValue" | constant "0 :: uint8" \ (SML) "(Word8.fromInt 0)" and (Haskell) "(0 :: Uint8.Word8)" and (Scala) "0.toByte" | constant "1 :: uint8" \ (SML) "(Word8.fromInt 1)" and (Haskell) "(1 :: Uint8.Word8)" and (Scala) "1.toByte" | constant "plus :: uint8 \ _ \ _" \ (SML) "Word8.+ ((_), (_))" and (Haskell) infixl 6 "+" and (Scala) "(_ +/ _).toByte" | constant "uminus :: uint8 \ _" \ (SML) "Word8.~" and (Haskell) "negate" and (Scala) "(- _).toByte" | constant "minus :: uint8 \ _" \ (SML) "Word8.- ((_), (_))" and (Haskell) infixl 6 "-" and (Scala) "(_ -/ _).toByte" | constant "times :: uint8 \ _ \ _" \ (SML) "Word8.* ((_), (_))" and (Haskell) infixl 7 "*" and (Scala) "(_ */ _).toByte" | constant "HOL.equal :: uint8 \ _ \ bool" \ (SML) "!((_ : Word8.word) = _)" and (Haskell) infix 4 "==" and (Scala) infixl 5 "==" | class_instance uint8 :: equal \ (Haskell) - | constant "less_eq :: uint8 \ _ \ bool" \ (SML) "Word8.<= ((_), (_))" and (Haskell) infix 4 "<=" and (Scala) "Uint8.less'_eq" | constant "less :: uint8 \ _ \ bool" \ (SML) "Word8.< ((_), (_))" and (Haskell) infix 4 "<" and (Scala) "Uint8.less" | constant "NOT :: uint8 \ _" \ (SML) "Word8.notb" and (Haskell) "Data'_Bits.complement" and (Scala) "_.unary'_~.toByte" -| constant "(AND) :: uint8 \ _" \ +| constant "Bit_Operations.and :: uint8 \ _" \ (SML) "Word8.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (Scala) "(_ & _).toByte" -| constant "(OR) :: uint8 \ _" \ +| constant "Bit_Operations.or :: uint8 \ _" \ (SML) "Word8.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (Scala) "(_ | _).toByte" -| constant "(XOR) :: uint8 \ _" \ +| constant "Bit_Operations.xor :: uint8 \ _" \ (SML) "Word8.xorb ((_),/ (_))" and (Haskell) "Data'_Bits.xor" and (Scala) "(_ ^ _).toByte" definition uint8_divmod :: "uint8 \ uint8 \ uint8 \ uint8" where "uint8_divmod x y = (if y = 0 then (undefined ((div) :: uint8 \ _) x (0 :: uint8), undefined ((mod) :: uint8 \ _) x (0 :: uint8)) else (x div y, x mod y))" definition uint8_div :: "uint8 \ uint8 \ uint8" where "uint8_div x y = fst (uint8_divmod x y)" definition uint8_mod :: "uint8 \ uint8 \ uint8" where "uint8_mod x y = snd (uint8_divmod x y)" lemma div_uint8_code [code]: "x div y = (if y = 0 then 0 else uint8_div x y)" including undefined_transfer unfolding uint8_divmod_def uint8_div_def by transfer (simp add: word_div_def) lemma mod_uint8_code [code]: "x mod y = (if y = 0 then x else uint8_mod x y)" including undefined_transfer unfolding uint8_mod_def uint8_divmod_def by transfer (simp add: word_mod_def) definition uint8_sdiv :: "uint8 \ uint8 \ uint8" where "uint8_sdiv x y = (if y = 0 then undefined ((div) :: uint8 \ _) x (0 :: uint8) else Abs_uint8 (Rep_uint8 x sdiv Rep_uint8 y))" definition div0_uint8 :: "uint8 \ uint8" where [code del]: "div0_uint8 x = undefined ((div) :: uint8 \ _) x (0 :: uint8)" declare [[code abort: div0_uint8]] definition mod0_uint8 :: "uint8 \ uint8" where [code del]: "mod0_uint8 x = undefined ((mod) :: uint8 \ _) x (0 :: uint8)" declare [[code abort: mod0_uint8]] lemma uint8_divmod_code [code]: "uint8_divmod x y = (if 0x80 \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint8 x, mod0_uint8 x) else let q = push_bit 1 (uint8_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" including undefined_transfer unfolding uint8_divmod_def uint8_sdiv_def div0_uint8_def mod0_uint8_def less_eq_uint8.rep_eq apply transfer apply (simp add: divmod_via_sdivmod push_bit_eq_mult) done lemma uint8_sdiv_code [code]: "Rep_uint8 (uint8_sdiv x y) = (if y = 0 then Rep_uint8 (undefined ((div) :: uint8 \ _) x (0 :: uint8)) else Rep_uint8 x sdiv Rep_uint8 y)" unfolding uint8_sdiv_def by(simp add: Abs_uint8_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint8_divmod_code} computes both with division only. \ code_printing constant uint8_div \ (SML) "Word8.div ((_), (_))" and (Haskell) "Prelude.div" | constant uint8_mod \ (SML) "Word8.mod ((_), (_))" and (Haskell) "Prelude.mod" | constant uint8_divmod \ (Haskell) "divmod" | constant uint8_sdiv \ (Scala) "(_ '/ _).toByte" definition uint8_test_bit :: "uint8 \ integer \ bool" where [code del]: "uint8_test_bit x n = (if n < 0 \ 7 < n then undefined (bit :: uint8 \ _) x n else bit x (nat_of_integer n))" lemma bit_uint8_code [code]: "bit x n \ n < 8 \ uint8_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint8_test_bit_def by (transfer, simp, transfer, simp) lemma uint8_test_bit_code [code]: "uint8_test_bit w n = (if n < 0 \ 7 < n then undefined (bit :: uint8 \ _) w n else bit (Rep_uint8 w) (nat_of_integer n))" unfolding uint8_test_bit_def by (simp add: bit_uint8.rep_eq) code_printing constant uint8_test_bit \ (SML) "Uint8.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (Scala) "Uint8.test'_bit" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_test'_bit out of bounds\") else Uint8.test'_bit x i)" definition uint8_set_bit :: "uint8 \ integer \ bool \ uint8" where [code del]: "uint8_set_bit x n b = (if n < 0 \ 7 < n then undefined (set_bit :: uint8 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint8_code [code]: "set_bit x n b = (if n < 8 then uint8_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint8_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint8_set_bit_code [code]: "Rep_uint8 (uint8_set_bit w n b) = (if n < 0 \ 7 < n then Rep_uint8 (undefined (set_bit :: uint8 \ _) w n b) else set_bit (Rep_uint8 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint8_set_bit_def by transfer simp code_printing constant uint8_set_bit \ (SML) "Uint8.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (Scala) "Uint8.set'_bit" and (Eval) "(fn x => fn i => fn b => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_set'_bit out of bounds\") else Uint8.set'_bit x i b)" definition uint8_shiftl :: "uint8 \ integer \ uint8" where [code del]: "uint8_shiftl x n = (if n < 0 \ 8 \ n then undefined (push_bit :: nat \ uint8 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint8_code [code]: "push_bit n x = (if n < 8 then uint8_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint8_shiftl_def by transfer simp lemma uint8_shiftl_code [code]: "Rep_uint8 (uint8_shiftl w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined (push_bit :: nat \ uint8 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint8 w))" including undefined_transfer unfolding uint8_shiftl_def by transfer simp code_printing constant uint8_shiftl \ (SML) "Uint8.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (Scala) "Uint8.shiftl" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_shiftl out of bounds\") else Uint8.shiftl x i)" definition uint8_shiftr :: "uint8 \ integer \ uint8" where [code del]: "uint8_shiftr x n = (if n < 0 \ 8 \ n then undefined (drop_bit :: _ \ _ \ uint8) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint8_code [code]: "drop_bit n x = (if n < 8 then uint8_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint8_shiftr_def by transfer simp lemma uint8_shiftr_code [code]: "Rep_uint8 (uint8_shiftr w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined (drop_bit :: _ \ _ \ uint8) w n) else drop_bit (nat_of_integer n) (Rep_uint8 w))" including undefined_transfer unfolding uint8_shiftr_def by transfer simp code_printing constant uint8_shiftr \ (SML) "Uint8.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (Scala) "Uint8.shiftr" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_shiftr out of bounds\") else Uint8.shiftr x i)" definition uint8_sshiftr :: "uint8 \ integer \ uint8" where [code del]: "uint8_sshiftr x n = (if n < 0 \ 8 \ n then undefined signed_drop_bit_uint8 n x else signed_drop_bit_uint8 (nat_of_integer n) x)" lemma sshiftr_uint8_code [code]: "signed_drop_bit_uint8 n x = (if n < 8 then uint8_sshiftr x (integer_of_nat n) else if bit x 7 then -1 else 0)" including undefined_transfer integer.lifting unfolding uint8_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond word_size) lemma uint8_sshiftr_code [code]: "Rep_uint8 (uint8_sshiftr w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined signed_drop_bit_uint8 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint8 w))" including undefined_transfer unfolding uint8_sshiftr_def by transfer simp code_printing constant uint8_sshiftr \ (SML) "Uint8.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint8.Int8) _)) :: Uint8.Word8)" and (Scala) "Uint8.shiftr'_signed" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_sshiftr out of bounds\") else Uint8.shiftr'_signed x i)" +context + includes bit_operations_syntax +begin + lemma uint8_msb_test_bit: "msb x \ bit (x :: uint8) 7" by transfer (simp add: msb_word_iff_bit) lemma msb_uint16_code [code]: "msb x \ uint8_test_bit x 7" by (simp add: uint8_test_bit_def uint8_msb_test_bit) lemma uint8_of_int_code [code]: "uint8_of_int i = Uint8 (integer_of_int i)" including integer.lifting by transfer simp lemma int_of_uint8_code [code]: "int_of_uint8 x = int_of_integer (integer_of_uint8 x)" by (simp add: int_of_uint8.rep_eq integer_of_uint8_def) lemma nat_of_uint8_code [code]: "nat_of_uint8 x = nat_of_integer (integer_of_uint8 x)" unfolding integer_of_uint8_def including integer.lifting by transfer simp definition integer_of_uint8_signed :: "uint8 \ integer" where "integer_of_uint8_signed n = (if bit n 7 then undefined integer_of_uint8 n else integer_of_uint8 n)" lemma integer_of_uint8_signed_code [code]: "integer_of_uint8_signed n = (if bit n 7 then undefined integer_of_uint8 n else integer_of_int (uint (Rep_uint8' n)))" by (simp add: integer_of_uint8_signed_def integer_of_uint8_def) lemma integer_of_uint8_code [code]: "integer_of_uint8 n = (if bit n 7 then integer_of_uint8_signed (n AND 0x7F) OR 0x80 else integer_of_uint8_signed n)" proof - have \integer_of_uint8_signed (n AND 0x7F) OR 0x80 = Bit_Operations.set_bit 7 (integer_of_uint8_signed (take_bit 7 n))\ by (simp add: take_bit_eq_mask set_bit_eq_or push_bit_eq_mult mask_eq_exp_minus_1) moreover have \integer_of_uint8 n = Bit_Operations.set_bit 7 (integer_of_uint8 (take_bit 7 n))\ if \bit n 7\ proof (rule bit_eqI) fix m from that show \bit (integer_of_uint8 n) m = bit (Bit_Operations.set_bit 7 (integer_of_uint8 (take_bit 7 n))) m\ for m including integer.lifting by transfer (auto simp add: bit_simps dest: bit_imp_le_length) qed ultimately show ?thesis by simp (simp add: integer_of_uint8_signed_def bit_simps) qed +end + code_printing constant "integer_of_uint8" \ (SML) "IntInf.fromLarge (Word8.toLargeInt _)" and (Haskell) "Prelude.toInteger" | constant "integer_of_uint8_signed" \ (Scala) "BigInt" section \Quickcheck setup\ definition uint8_of_natural :: "natural \ uint8" where "uint8_of_natural x \ Uint8 (integer_of_natural x)" instantiation uint8 :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint8 \ qc_random_cnv uint8_of_natural" definition "exhaustive_uint8 \ qc_exhaustive_cnv uint8_of_natural" definition "full_exhaustive_uint8 \ qc_full_exhaustive_cnv uint8_of_natural" instance .. end instantiation uint8 :: narrowing begin interpretation quickcheck_narrowing_samples "\i. let x = Uint8 i in (x, 0xFF - x)" "0" "Typerep.Typerep (STR ''Uint8.uint8'') []" . definition "narrowing_uint8 d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint8 itself \ _"]] lemmas partial_term_of_uint8 [code] = partial_term_of_code instance .. end end diff --git a/thys/Native_Word/Word_Type_Copies.thy b/thys/Native_Word/Word_Type_Copies.thy --- a/thys/Native_Word/Word_Type_Copies.thy +++ b/thys/Native_Word/Word_Type_Copies.thy @@ -1,344 +1,344 @@ (* Title: Word_Type_Copies.thy Author: Florian Haftmann, TU Muenchen *) chapter \Systematic approach towards type copies of word type\ theory Word_Type_Copies imports "HOL-Library.Word" "Word_Lib.Most_significant_bit" "Word_Lib.Least_significant_bit" "Word_Lib.Generic_set_bit" Code_Target_Word_Base begin lemma int_of_integer_unsigned_eq [simp]: \int_of_integer (unsigned w) = uint w\ by transfer simp lemma int_of_integer_signed_eq [simp]: \int_of_integer (signed w) = sint w\ by transfer simp lemma word_of_int_of_integer_eq [simp]: \word_of_int (int_of_integer k) = word_of_integer k\ by (simp add: word_of_integer.rep_eq) text \The lifting machinery is not localized, hence the abstract proofs are carried out using morphisms.\ locale word_type_copy = fixes of_word :: \'b::len word \ 'a\ and word_of :: \'a \ 'b word\ assumes type_definition: \type_definition word_of of_word UNIV\ begin lemma word_of_word: \word_of (of_word w) = w\ using type_definition by (simp add: type_definition_def) lemma of_word_of [code abstype]: \of_word (word_of p) = p\ \ \Use an abstract type for code generation to disable pattern matching on \<^term>\of_word\.\ using type_definition by (simp add: type_definition_def) lemma word_of_eqI: \p = q\ if \word_of p = word_of q\ proof - from that have \of_word (word_of p) = of_word (word_of q)\ by simp then show ?thesis by (simp add: of_word_of) qed lemma eq_iff_word_of: \p = q \ word_of p = word_of q\ by (auto intro: word_of_eqI) end bundle constraintless begin declaration \ let val cs = map (rpair NONE o fst o dest_Const) [\<^term>\0\, \<^term>\(+)\, \<^term>\uminus\, \<^term>\(-)\, \<^term>\1\, \<^term>\(*)\, \<^term>\(div)\, \<^term>\(mod)\, \<^term>\HOL.equal\, \<^term>\(\)\, \<^term>\(<)\, \<^term>\(dvd)\, \<^term>\of_bool\, \<^term>\numeral\, \<^term>\of_nat\, \<^term>\bit\, - \<^term>\NOT\, \<^term>\(AND)\, \<^term>\(OR)\, \<^term>\(XOR)\, \<^term>\mask\, + \<^term>\NOT\, \<^term>\Bit_Operations.and\, \<^term>\Bit_Operations.or\, \<^term>\Bit_Operations.xor\, \<^term>\mask\, \<^term>\push_bit\, \<^term>\drop_bit\, \<^term>\take_bit\, \<^term>\Bit_Operations.set_bit\, \<^term>\unset_bit\, \<^term>\flip_bit\, \<^term>\msb\, \<^term>\lsb\, \<^term>\size\, \<^term>\Generic_set_bit.set_bit\, \<^term>\set_bits\] in K (Context.mapping I (fold Proof_Context.add_const_constraint cs)) end \ end locale word_type_copy_ring = word_type_copy opening constraintless + constrains word_of :: \'a \ 'b::len word\ assumes word_of_0 [code]: \word_of 0 = 0\ and word_of_1 [code]: \word_of 1 = 1\ and word_of_add [code]: \word_of (p + q) = word_of p + word_of q\ and word_of_minus [code]: \word_of (- p) = - (word_of p)\ and word_of_diff [code]: \word_of (p - q) = word_of p - word_of q\ and word_of_mult [code]: \word_of (p * q) = word_of p * word_of q\ and word_of_div [code]: \word_of (p div q) = word_of p div word_of q\ and word_of_mod [code]: \word_of (p mod q) = word_of p mod word_of q\ and equal_iff_word_of [code]: \HOL.equal p q \ HOL.equal (word_of p) (word_of q)\ and less_eq_iff_word_of [code]: \p \ q \ word_of p \ word_of q\ and less_iff_word_of [code]: \p < q \ word_of p < word_of q\ begin lemma of_class_comm_ring_1: \OFCLASS('a, comm_ring_1_class)\ by standard (simp_all add: eq_iff_word_of word_of_0 word_of_1 word_of_add word_of_minus word_of_diff word_of_mult algebra_simps) lemma of_class_semiring_modulo: \OFCLASS('a, semiring_modulo_class)\ by standard (simp_all add: eq_iff_word_of word_of_0 word_of_1 word_of_add word_of_minus word_of_diff word_of_mult word_of_mod word_of_div algebra_simps mod_mult_div_eq) lemma of_class_equal: \OFCLASS('a, equal_class)\ by standard (simp add: eq_iff_word_of equal_iff_word_of equal) lemma of_class_linorder: \OFCLASS('a, linorder_class)\ by standard (auto simp add: eq_iff_word_of less_eq_iff_word_of less_iff_word_of) end locale word_type_copy_bits = word_type_copy_ring - opening constraintless + + opening constraintless bit_operations_syntax + constrains word_of :: \'a::{comm_ring_1, semiring_modulo, equal, linorder} \ 'b::len word\ fixes signed_drop_bit :: \nat \ 'a \ 'a\ assumes bit_eq_word_of [code]: \bit p = bit (word_of p)\ and word_of_not [code]: \word_of (NOT p) = NOT (word_of p)\ and word_of_and [code]: \word_of (p AND q) = word_of p AND word_of q\ and word_of_or [code]: \word_of (p OR q) = word_of p OR word_of q\ and word_of_xor [code]: \word_of (p XOR q) = word_of p XOR word_of q\ and word_of_mask [code]: \word_of (mask n) = mask n\ and word_of_push_bit [code]: \word_of (push_bit n p) = push_bit n (word_of p)\ and word_of_drop_bit [code]: \word_of (drop_bit n p) = drop_bit n (word_of p)\ and word_of_signed_drop_bit [code]: \word_of (signed_drop_bit n p) = Word.signed_drop_bit n (word_of p)\ and word_of_take_bit [code]: \word_of (take_bit n p) = take_bit n (word_of p)\ and word_of_set_bit [code]: \word_of (Bit_Operations.set_bit n p) = Bit_Operations.set_bit n (word_of p)\ and word_of_unset_bit [code]: \word_of (unset_bit n p) = unset_bit n (word_of p)\ and word_of_flip_bit [code]: \word_of (flip_bit n p) = flip_bit n (word_of p)\ begin lemma word_of_bool: \word_of (of_bool n) = of_bool n\ by (simp add: word_of_0 word_of_1) lemma word_of_nat: \word_of (of_nat n) = of_nat n\ by (induction n) (simp_all add: word_of_0 word_of_1 word_of_add) lemma word_of_numeral [simp]: \word_of (numeral n) = numeral n\ proof - have \word_of (of_nat (numeral n)) = of_nat (numeral n)\ by (simp only: word_of_nat) then show ?thesis by simp qed lemma word_of_power: \word_of (p ^ n) = word_of p ^ n\ by (induction n) (simp_all add: word_of_1 word_of_mult word_of_numeral) lemma even_iff_word_of: \2 dvd p \ 2 dvd word_of p\ (is \?P \ ?Q\) proof assume ?P then obtain q where \p = 2 * q\ .. then show ?Q by (simp add: word_of_mult word_of_numeral) next assume ?Q then obtain w where \word_of p = 2 * w\ .. then have \of_word (word_of p) = of_word (2 * w)\ by simp then have \p = 2 * of_word w\ by (simp add: eq_iff_word_of word_of_word word_of_mult word_of_numeral) then show ?P by simp qed lemma of_class_ring_bit_operations: \OFCLASS('a, ring_bit_operations_class)\ proof - have induct: \P p\ if stable: \(\p. p div 2 = p \ P p)\ and rec: \(\p b. P p \ (of_bool b + 2 * p) div 2 = p \ P (of_bool b + 2 * p))\ for p :: 'a and P proof - from stable have stable': \(\p. word_of p div 2 = word_of p \ P p)\ by (simp add: eq_iff_word_of word_of_div word_of_numeral) from rec have rec': \(\p b. P p \ (of_bool b + 2 * word_of p) div 2 = word_of p \ P (of_bool b + 2 * p))\ by (simp add: eq_iff_word_of word_of_add word_of_bool word_of_mult word_of_div word_of_numeral) define w where \w = word_of p\ then have \p = of_word w\ by (simp add: of_word_of) also have \P (of_word w)\ proof (induction w rule: bits_induct) case (stable w) show ?case by (rule stable') (simp add: word_of_word stable) next case (rec w b) have \P (of_bool b + 2 * of_word w)\ by (rule rec') (simp_all add: word_of_word rec) also have \of_bool b + 2 * of_word w = of_word (of_bool b + 2 * w)\ by (simp add: eq_iff_word_of word_of_word word_of_add word_of_1 word_of_mult word_of_numeral word_of_0) finally show ?case . qed finally show \P p\ . qed have \class.semiring_parity_axioms (+) (0::'a) (*) 1 (mod)\ by standard (simp_all add: eq_iff_word_of word_of_0 word_of_1 word_of_numeral even_iff_word_of word_of_mod even_iff_mod_2_eq_zero) with of_class_semiring_modulo have \OFCLASS('a, semiring_parity_class)\ by (rule semiring_parity_class.intro) moreover have \class.semiring_bits_axioms (+) (-) (0::'a) (*) 1 (div) (mod) bit\ apply (standard, fact induct) apply (simp_all only: eq_iff_word_of word_of_0 word_of_1 word_of_bool word_of_numeral word_of_add word_of_diff word_of_mult word_of_div word_of_mod word_of_power even_iff_word_of bit_eq_word_of push_bit_take_bit drop_bit_take_bit even_drop_bit_iff_not_bit flip: push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod mask_eq_exp_minus_1) apply (auto simp add: ac_simps bit_simps drop_bit_exp_eq push_bit_of_1) done ultimately have \OFCLASS('a, semiring_bits_class)\ by (rule semiring_bits_class.intro) moreover have \class.semiring_bit_shifts_axioms (+) (*) (1::'a) (div) (mod) push_bit drop_bit take_bit\ by standard (simp_all add: eq_iff_word_of word_of_push_bit word_of_mult word_of_power word_of_numeral word_of_drop_bit word_of_div word_of_take_bit word_of_mod flip: push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod) ultimately have \OFCLASS('a, semiring_bit_shifts_class)\ by (rule semiring_bit_shifts_class.intro) moreover have \class.semiring_bit_operations_axioms (+) (-) (*) (1::'a) bit push_bit (AND) (OR) (XOR) mask Bit_Operations.set_bit unset_bit flip_bit\ by standard (simp_all add: eq_iff_word_of word_of_push_bit word_of_power word_of_numeral bit_eq_word_of word_of_and word_of_or word_of_xor word_of_mask word_of_diff word_of_1 bit_simps word_of_set_bit set_bit_eq_or word_of_unset_bit word_of_flip_bit flip_bit_eq_xor flip: mask_eq_exp_minus_1) ultimately have \OFCLASS('a, semiring_bit_operations_class)\ by (rule semiring_bit_operations_class.intro) moreover have \OFCLASS('a, ring_parity_class)\ using \OFCLASS('a, semiring_parity_class)\ by (rule ring_parity_class.intro) standard moreover have \class.ring_bit_operations_axioms (+) (-) (0::'a) (*) 1 bit uminus NOT\ by standard (simp_all add: eq_iff_word_of word_of_power word_of_numeral bit_eq_word_of word_of_diff word_of_1 bit_simps word_of_not word_of_0 word_of_minus minus_eq_not_minus_1) ultimately show \OFCLASS('a, ring_bit_operations_class)\ by (rule ring_bit_operations_class.intro) qed lemma [code]: \take_bit n a = a AND mask n\ for a :: 'a by (simp add: eq_iff_word_of word_of_take_bit word_of_and word_of_mask take_bit_eq_mask) lemma [code]: \mask (Suc n) = push_bit n (1 :: 'a) OR mask n\ \mask 0 = (0 :: 'a)\ by (simp_all add: eq_iff_word_of word_of_mask word_of_or word_of_push_bit word_of_0 word_of_1 mask_Suc_exp push_bit_of_1) lemma [code]: \Bit_Operations.set_bit n w = w OR push_bit n 1\ for w :: 'a by (simp add: eq_iff_word_of word_of_set_bit word_of_or word_of_push_bit word_of_1 set_bit_eq_or) lemma [code]: \unset_bit n w = w AND NOT (push_bit n 1)\ for w :: 'a by (simp add: eq_iff_word_of word_of_unset_bit word_of_and word_of_not word_of_push_bit word_of_1 unset_bit_eq_and_not) lemma [code]: \flip_bit n w = w XOR push_bit n 1\ for w :: 'a by (simp add: eq_iff_word_of word_of_flip_bit word_of_xor word_of_push_bit word_of_1 flip_bit_eq_xor) end locale word_type_copy_more = word_type_copy_bits + constrains word_of :: \'a::{ring_bit_operations, equal, linorder} \ 'b::len word\ fixes of_nat :: \nat \ 'a\ and nat_of :: \'a \ nat\ and of_int :: \int \ 'a\ and int_of :: \'a \ int\ and of_integer :: \integer \ 'a\ and integer_of :: \'a \ integer\ assumes word_of_nat_eq: \word_of (of_nat n) = word_of_nat n\ and nat_of_eq_word_of: \nat_of p = unat (word_of p)\ and word_of_int_eq: \word_of (of_int k) = word_of_int k\ and int_of_eq_word_of: \int_of p = uint (word_of p)\ and word_of_integer_eq: \word_of (of_integer l) = word_of_integer l\ and integer_of_eq_word_of: \integer_of p = unsigned (word_of p)\ begin lemma of_word_numeral [code_post]: \of_word (numeral n) = numeral n\ by (simp add: eq_iff_word_of word_of_word word_of_numeral) lemma of_word_0 [code_post]: \of_word 0 = 0\ by (simp add: eq_iff_word_of word_of_0 word_of_word) lemma of_word_1 [code_post]: \of_word 1 = 1\ by (simp add: eq_iff_word_of word_of_1 word_of_word) text \Use pretty numerals from integer for pretty printing\ lemma numeral_eq_integer [code_unfold]: \numeral n = of_integer (numeral n)\ by (simp add: eq_iff_word_of word_of_integer_eq word_of_numeral word_of_integer.rep_eq) lemma neg_numeral_eq_integer [code_unfold]: \- numeral n = of_integer (- numeral n)\ by (simp add: eq_iff_word_of word_of_integer_eq word_of_minus word_of_numeral word_of_integer.rep_eq) end locale word_type_copy_misc = word_type_copy_more - opening constraintless + + opening constraintless bit_operations_syntax + constrains word_of :: \'a::{ring_bit_operations, equal, linorder} \ 'b::len word\ fixes size :: nat and set_bits_aux :: \(nat \ bool) \ nat \ 'a \ 'a\ assumes size_eq_length: \size = LENGTH('b::len)\ and msb_iff_word_of [code]: \msb p \ msb (word_of p)\ and lsb_iff_word_of [code]: \lsb p \ lsb (word_of p)\ and size_eq_word_of: \Nat.size (p :: 'a) = Nat.size (word_of p)\ and word_of_generic_set_bit [code]: \word_of (Generic_set_bit.set_bit p n b) = Generic_set_bit.set_bit (word_of p) n b\ and word_of_set_bits: \word_of (set_bits P) = set_bits P\ and word_of_set_bits_aux: \word_of (set_bits_aux P n p) = Code_Target_Word_Base.set_bits_aux P n (word_of p)\ begin lemma size_eq [code]: \Nat.size p = size\ for p :: 'a by (simp add: size_eq_length size_eq_word_of word_size) lemma set_bits_aux_code [code]: \set_bits_aux f n w = (if n = 0 then w else let n' = n - 1 in set_bits_aux f n' (push_bit 1 w OR (if f n' then 1 else 0)))\ by (simp add: eq_iff_word_of word_of_set_bits_aux Let_def word_of_mult word_of_or word_of_0 word_of_1 set_bits_aux_rec [of f n]) lemma set_bits_code [code]: \set_bits P = set_bits_aux P size 0\ by (simp add: fun_eq_iff eq_iff_word_of word_of_set_bits word_of_set_bits_aux word_of_0 size_eq_length set_bits_conv_set_bits_aux) lemma of_class_lsb: \OFCLASS('a, lsb_class)\ by standard (simp add: fun_eq_iff lsb_iff_word_of even_iff_word_of lsb_odd) lemma of_class_set_bit: \OFCLASS('a, set_bit_class)\ by standard (simp add: eq_iff_word_of word_of_generic_set_bit bit_eq_word_of word_of_power word_of_numeral word_of_0 bit_simps) lemma of_class_bit_comprehension: \OFCLASS('a, bit_comprehension_class)\ by standard (simp add: eq_iff_word_of word_of_set_bits bit_eq_word_of set_bits_bit_eq) end 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,382 +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) 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/Refine_Imperative_HOL/benchmarks/Heapmap/isabelle/Heapmap_Bench.thy b/thys/Refine_Imperative_HOL/benchmarks/Heapmap/isabelle/Heapmap_Bench.thy --- a/thys/Refine_Imperative_HOL/benchmarks/Heapmap/isabelle/Heapmap_Bench.thy +++ b/thys/Refine_Imperative_HOL/benchmarks/Heapmap/isabelle/Heapmap_Bench.thy @@ -1,102 +1,108 @@ theory Heapmap_Bench imports "../../../IICF/Impl/Heaps/IICF_Impl_Heapmap" "../../../Sepref_ICF_Bindings" begin +context + includes bit_operations_syntax +begin + definition rrand :: "uint32 \ uint32" where "rrand s \ (s * 1103515245 + 12345) AND 0x7FFFFFFF" +end + definition rand :: "uint32 \ nat \ (uint32 * nat)" where "rand s m \ let s = rrand s; r = nat_of_uint32 s; r = (r * m) div 0x80000000 in (s,r)" partial_function (heap) rep where "rep i N f s = ( if i f s i; rep (i+1) N f s } else return s )" declare rep.simps[code] term hm_insert_op_impl definition "testsuite N \ do { let s=0; let N2=efficient_nat_div2 N; hm \ hm_empty_op_impl N; (hm,s) \ rep 0 N (\(hm,s) i. do { let (s,v) = rand s N2; hm \ hm_insert_op_impl N id i v hm; return (hm,s) }) (hm,s); (hm,s) \ rep 0 N (\(hm,s) i. do { let (s,v) = rand s N2; hm \ hm_change_key_op_impl id i v hm; return (hm,s) }) (hm,s); hm \ rep 0 N (\hm i. do { (_,hm) \ hm_pop_min_op_impl id hm; return hm }) hm; return () }" export_code rep in SML_imp partial_function (tailrec) drep where "drep i N f s = ( if i do { let s=0; let N2=efficient_nat_div2 N; let hm= aluprioi.empty (); let (hm,s) = drep 0 N (\(hm,s) i. do { let (s,v) = rand s N2; let hm = aluprioi.insert hm i v; (hm,s) }) (hm,s); let (hm,s) = drep 0 N (\(hm,s) i. do { let (s,v) = rand s N2; let hm = aluprioi.insert hm i v; (hm,s) }) (hm,s); let hm = drep 0 N (\hm i. do { let (_,_,hm) = aluprioi.pop hm; hm }) hm; () }" export_code testsuite ftestsuite nat_of_integer integer_of_nat in SML_imp module_name Heapmap file \heapmap_export.sml\ end diff --git a/thys/SPARCv8/SparcModel_MMU/MMU.thy b/thys/SPARCv8/SparcModel_MMU/MMU.thy --- a/thys/SPARCv8/SparcModel_MMU/MMU.thy +++ b/thys/SPARCv8/SparcModel_MMU/MMU.thy @@ -1,312 +1,318 @@ (* Title: Memory.thy Author: David Sanán, Trinity College Dublin, 2012 Zhe Hou, NTU, 2016. *) section \Memory Management Unit (MMU)\ theory MMU imports Main RegistersOps Sparc_Types begin section \MMU Sizing\ text\ We need some citation here for documentation about the MMU. \ text\The MMU uses the Address Space Identifiers (ASI) to control memory access. ASI = 8, 10 are for user; ASI = 9, 11 are for supervisor.\ subsection "MMU Types" type_synonym word_PTE_flags = word8 type_synonym word_length_PTE_flags = word_length8 subsection "MMU length values" text\Definitions for the length of the virtua address, page size, virtual translation tables indexes, virtual address offset and Page protection flags\ definition length_entry_type :: "nat" where "length_entry_type \ LENGTH(word_length_entry_type)" definition length_phys_address:: "nat" where "length_phys_address \ LENGTH(word_length_phys_address)" definition length_virtua_address:: "nat" where "length_virtua_address \ LENGTH(word_length_virtua_address)" definition length_page:: "nat" where "length_page \ LENGTH(word_length_page)" definition length_t1:: "nat" where "length_t1 \ LENGTH(word_length_t1)" definition length_t2:: "nat" where "length_t2 \ LENGTH(word_length_t2)" definition length_t3:: "nat" where "length_t3 \ LENGTH(word_length_t3)" definition length_offset:: "nat" where "length_offset \ LENGTH(word_length_offset)" definition length_PTE_flags :: "nat" where "length_PTE_flags \ LENGTH(word_length_PTE_flags)" subsection "MMU index values" definition va_t1_index :: "nat" where "va_t1_index \ length_virtua_address - length_t1" definition va_t2_index :: "nat" where "va_t2_index \ va_t1_index - length_t2" definition va_t3_index :: "nat" where "va_t3_index \ va_t2_index - length_t3" definition va_offset_index :: "nat" where "va_offset_index \ va_t3_index - length_offset" definition pa_page_index :: "nat" where "pa_page_index \ length_phys_address - length_page" definition pa_offset_index :: "nat" where "pa_offset_index \ pa_page_index -length_page" section \MMU Definition\ record MMU_state = registers :: "MMU_context" (* contexts:: context_table*) text \The following functions access MMU registers via addresses. See UT699LEON3FT manual page 35.\ definition mmu_reg_val:: "MMU_state \ virtua_address \ machine_word option" where "mmu_reg_val mmu_state addr \ if addr = 0x000 then \ \MMU control register\ Some ((registers mmu_state) CR) else if addr = 0x100 then \ \Context pointer register\ Some ((registers mmu_state) CTP) else if addr = 0x200 then \ \Context register\ Some ((registers mmu_state) CNR) else if addr = 0x300 then \ \Fault status register\ Some ((registers mmu_state) FTSR) else if addr = 0x400 then \ \Fault address register\ Some ((registers mmu_state) FAR) else None" definition mmu_reg_mod:: "MMU_state \ virtua_address \ machine_word \ MMU_state option" where "mmu_reg_mod mmu_state addr w \ if addr = 0x000 then \ \MMU control register\ Some (mmu_state\registers := (registers mmu_state)(CR := w)\) else if addr = 0x100 then \ \Context pointer register\ Some (mmu_state\registers := (registers mmu_state)(CTP := w)\) else if addr = 0x200 then \ \Context register\ Some (mmu_state\registers := (registers mmu_state)(CNR := w)\) else if addr = 0x300 then \ \Fault status register\ Some (mmu_state\registers := (registers mmu_state)(FTSR := w)\) else if addr = 0x400 then \ \Fault address register\ Some (mmu_state\registers := (registers mmu_state)(FAR := w)\) else None" section \Virtual Memory\ subsection \MMU Auxiliary Definitions\ definition getCTPVal:: "MMU_state \ machine_word" where "getCTPVal mmu \ (registers mmu) CTP" definition getCNRVal::"MMU_state \ machine_word" where "getCNRVal mmu \ (registers mmu) CNR" text\ The physical context table address is got from the ConText Pointer register (CTP) and the Context Register (CNR) MMU registers. The CTP is shifted to align it with the physical address (36 bits) and we add the table index given on CNR. CTP is right shifted 2 bits, cast to phys address and left shifted 6 bytes to be aligned with the context register. CNR is 2 bits left shifted for alignment with the context table. \ definition compose_context_table_addr :: "machine_word \machine_word \ phys_address" where "compose_context_table_addr ctp cnr \ ((ucast (ctp >> 2)) << 6) + (ucast cnr << 2)" subsection \Virtual Address Translation\ text\Get the context table phys address from the MMU registers\ definition get_context_table_addr :: "MMU_state \ phys_address" where "get_context_table_addr mmu \ compose_context_table_addr (getCTPVal mmu) (getCNRVal mmu)" definition va_list_index :: "nat list" where "va_list_index \ [va_t1_index,va_t2_index,va_t3_index,0]" definition offset_index :: "nat list" where "offset_index \ [ length_machine_word , length_machine_word-length_t1 , length_machine_word-length_t1-length_t2 , length_machine_word-length_t1-length_t2-length_t3 ]" definition index_len_table :: "nat list" where "index_len_table \ [8,6,6,0]" definition n_context_tables :: "nat" where "n_context_tables \ 3" text \The following are basic physical memory read functions. At this level we don't need the write memory yet.\ definition mem_context_val:: "asi_type \ phys_address \ mem_context \ mem_val_type option" where "mem_context_val asi add m \ let asi8 = word_of_int 8; r1 = m asi add in if r1 = None then m asi8 add else r1 " +context + includes bit_operations_syntax +begin + text \Given an ASI (word8), an address (word32) addr, read the 32bit value from the memory addresses starting from address addr' where addr' = addr exception that the last two bits are 0's. That is, read the data from addr', addr'+1, addr'+2, addr'+3.\ definition mem_context_val_w32 :: "asi_type \ phys_address \ mem_context \ word32 option" where "mem_context_val_w32 asi addr m \ let addr' = (AND) addr 0b111111111111111111111111111111111100; addr0 = (OR) addr' 0b000000000000000000000000000000000000; addr1 = (OR) addr' 0b000000000000000000000000000000000001; addr2 = (OR) addr' 0b000000000000000000000000000000000010; addr3 = (OR) addr' 0b000000000000000000000000000000000011; r0 = mem_context_val asi addr0 m; r1 = mem_context_val asi addr1 m; r2 = mem_context_val asi addr2 m; r3 = mem_context_val asi addr3 m in if r0 = None \ r1 = None \ r2 = None \ r3 = None then None else let byte0 = case r0 of Some v \ v; byte1 = case r1 of Some v \ v; byte2 = case r2 of Some v \ v; byte3 = case r3 of Some v \ v in Some ((OR) ((OR) ((OR) ((ucast(byte0)) << 24) ((ucast(byte1)) << 16)) ((ucast(byte2)) << 8)) (ucast(byte3))) " text \ @{term "get_addr_from_table"} browses the page description tables until it finds a PTE (bits==suc (suc 0). If it is a PTE it aligns the 24 most significant bits of the entry with the most significant bits of the phys address and or-ed with the offset, which will vary depending on the entry level. In the case we are looking at the last table level (level 3), the offset is aligned to 0 otherwise it will be 2. If the table entry is a PTD (bits== Suc 0), the index is obtained from the virtual address depending on the current level and or-ed with the PTD. \ function ptd_lookup:: "virtua_address \ virtua_address \ mem_context \ nat \ (phys_address \ PTE_flags) option" where "ptd_lookup va pt m lvl = ( if lvl > 3 then None else let thislvl_offset = ( if lvl = 1 then (ucast ((ucast (va >> 24))::word8))::word32 else if lvl = 2 then (ucast ((ucast (va >> 18))::word6))::word32 else (ucast ((ucast (va >> 12))::word6))::word32); thislvl_addr = (OR) pt thislvl_offset; thislvl_data = mem_context_val_w32 (word_of_int 9) (ucast thislvl_addr) m in case thislvl_data of Some v \ ( let et_val = (AND) v 0b00000000000000000000000000000011 in if et_val = 0 then \ \Invalid\ None else if et_val = 1 then \ \Page Table Descriptor\ let ptp = (AND) v 0b11111111111111111111111111111100 in ptd_lookup va ptp m (lvl+1) else if et_val = 2 then \ \Page Table Entry\ let ppn = (ucast (v >> 8))::word24; va_offset = (ucast ((ucast va)::word12))::word36 in Some (((OR) (((ucast ppn)::word36) << 12) va_offset), ((ucast v)::word8)) else \ \\et_val = 3\, reserved.\ None ) |None \ None) " by pat_completeness auto termination by (relation "measure (\ (va, (pt, (m, lvl))). 4 - lvl)") auto definition get_acc_flag:: "PTE_flags \ word3" where "get_acc_flag w8 \ (ucast (w8 >> 2))::word3" definition mmu_readable:: "word3 \ asi_type \ bool" where "mmu_readable f asi \ if uint asi \ {8, 10} then if uint f \ {0,1,2,3,5} then True else False else if uint asi \ {9, 11} then if uint f \ {0,1,2,3,5,6,7} then True else False else False " definition mmu_writable:: "word3 \ asi_type \ bool" where "mmu_writable f asi \ if uint asi \ {8, 10} then if uint f \ {1,3} then True else False else if uint asi \ {9, 11} then if uint f \ {1,3,5,7} then True else False else False " definition virt_to_phys :: "virtua_address \ MMU_state \ mem_context \ (phys_address \ PTE_flags) option" where "virt_to_phys va mmu m \ let ctp_val = mmu_reg_val mmu (0x100); cnr_val = mmu_reg_val mmu (0x200); mmu_cr_val = (registers mmu) CR in if (AND) mmu_cr_val 1 \ 0 then \ \MMU enabled.\ case (ctp_val,cnr_val) of (Some v1, Some v2) \ let context_table_entry = (OR) ((v1 >> 11) << 11) (((AND) v2 0b00000000000000000000000111111111) << 2); context_table_data = mem_context_val_w32 (word_of_int 9) (ucast context_table_entry) m in ( case context_table_data of Some lvl1_page_table \ ptd_lookup va lvl1_page_table m 1 |None \ None) |_ \ None else Some ((ucast va), ((0b11101111)::word8)) " text \ \newpage The below function gives the initial values of MMU registers. In particular, the MMU context register CR is 0 because: We don't know the bits for IMPL, VER, and SC; the bits for PSO are 0s because we use TSO; the reserved bits are 0s; we assume NF bits are 0s; and most importantly, the E bit is 0 because when the machine starts up, MMU is disabled. An initial boot procedure (bootloader or something like that) should configure the MMU and then enable it if the OS uses MMU.\ definition MMU_registers_init :: "MMU_context" where "MMU_registers_init r \ 0" definition mmu_setup :: "MMU_state" where "mmu_setup \ \registers=MMU_registers_init\" end + +end diff --git a/thys/SPARCv8/SparcModel_MMU/RegistersOps.thy b/thys/SPARCv8/SparcModel_MMU/RegistersOps.thy --- a/thys/SPARCv8/SparcModel_MMU/RegistersOps.thy +++ b/thys/SPARCv8/SparcModel_MMU/RegistersOps.thy @@ -1,79 +1,85 @@ section\Register Operations\ theory RegistersOps imports Main "../lib/WordDecl" "Word_Lib.Bit_Shifts_Infix_Syntax" begin +context + includes bit_operations_syntax +begin + text\ This theory provides operations to get, set and clear bits in registers \ section "Getting Fields" text\ Get a field of type @{typ "'b::len word"} starting at @{term "index"} from @{term "addr"} of type @{typ "'a::len word"} \ definition get_field_from_word_a_b:: "'a::len word \ nat \ 'b::len word" where "get_field_from_word_a_b addr index \ let off = (size addr - LENGTH('b)) in ucast ((addr << (off-index)) >> off)" text\ Obtain, from addr of type @{typ "'a::len word"}, another @{typ "'a::len word"} containing the field of length \len\ starting at \index\ in \addr\. \ definition get_field_from_word_a_a:: "'a::len word \ nat \ nat \ 'a::len word" where "get_field_from_word_a_a addr index len \ (addr << (size addr - (index+len)) >> (size addr - len))" section "Setting Fields" text\ Set the field of type @{typ "'b::len word"} at \index\ from \record\ of type @{typ "'a::len word"}. \ definition set_field :: "'a::len word \ 'b::len word \ nat \ 'a::len word" where "set_field record field index \ let mask:: ('a::len word) = (mask (size field)) << index in (record AND (NOT mask)) OR ((ucast field) << index)" section "Clearing Fields" text\ Zero the \n\ initial bits of \addr\. \ definition clear_n_bits:: "'a::len word \ nat \ 'a::len word" where "clear_n_bits addr n \ addr AND (NOT (mask n))" text\ Gets the natural value of a 32 bit mask \ definition get_nat_from_mask::"word32 \ nat \ nat \ (word32 \ nat)" where " get_nat_from_mask w m v \ if (w AND (mask m) =0) then (w>>m, v+m) else (w,m) " definition get_nat_from_mask32::"word32\ nat" where "get_nat_from_mask32 w \ if (w=0) then len_of TYPE (word_length32) else let (w,res) = get_nat_from_mask w 16 0 in let (w,res)= get_nat_from_mask w 8 res in let (w,res) = get_nat_from_mask w 4 res in let (w,res) = get_nat_from_mask w 2 res in let (w,res) = get_nat_from_mask w 1 res in res " +end + end diff --git a/thys/SPARCv8/SparcModel_MMU/Sparc_Execution.thy b/thys/SPARCv8/SparcModel_MMU/Sparc_Execution.thy --- a/thys/SPARCv8/SparcModel_MMU/Sparc_Execution.thy +++ b/thys/SPARCv8/SparcModel_MMU/Sparc_Execution.thy @@ -1,431 +1,437 @@ (* * 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, David Sanan. *) theory Sparc_Execution imports Main Sparc_Instruction Sparc_State Sparc_Types "HOL-Eisbach.Eisbach_Tools" begin primrec sum :: "nat \ nat" where "sum 0 = 0" | "sum (Suc n) = Suc n + sum n" definition select_trap :: "unit \ ('a,unit) sparc_state_monad" where "select_trap _ \ do traps \ gets (\s. (get_trap_set s)); rt_val \ gets (\s. (reset_trap_val s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); et_val \ gets (\s. (get_ET psr_val)); modify (\s. (emp_trap_set s)); if rt_val = True then \ \ignore \ET\, and leave \tt\ unchaged\ return () else if et_val = 0 then \ \go to error mode, machine needs reset\ do set_err_mode True; set_exe_mode False; fail () od \ \By the SPARCv8 manual only 1 of the following traps could be in traps.\ else if data_store_error \ traps then do write_cpu_tt (0b00101011::word8); return () od else if instruction_access_error \ traps then do write_cpu_tt (0b00100001::word8); return () od else if r_register_access_error \ traps then do write_cpu_tt (0b00100000::word8); return () od else if instruction_access_exception \ traps then do write_cpu_tt (0b00000001::word8); return () od else if privileged_instruction \ traps then do write_cpu_tt (0b00000011::word8); return () od else if illegal_instruction \ traps then do write_cpu_tt (0b00000010::word8); return () od else if fp_disabled \ traps then do write_cpu_tt (0b00000100::word8); return () od else if cp_disabled \ traps then do write_cpu_tt (0b00100100::word8); return () od else if unimplemented_FLUSH \ traps then do write_cpu_tt (0b00100101::word8); return () od else if window_overflow \ traps then do write_cpu_tt (0b00000101::word8); return () od else if window_underflow \ traps then do write_cpu_tt (0b00000110::word8); return () od else if mem_address_not_aligned \ traps then do write_cpu_tt (0b00000111::word8); return () od else if fp_exception \ traps then do write_cpu_tt (0b00001000::word8); return () od else if cp_exception \ traps then do write_cpu_tt (0b00101000::word8); return () od else if data_access_error \ traps then do write_cpu_tt (0b00101001::word8); return () od else if data_access_exception \ traps then do write_cpu_tt (0b00001001::word8); return () od else if tag_overflow \ traps then do write_cpu_tt (0b00001010::word8); return () od else if division_by_zero \ traps then do write_cpu_tt (0b00101010::word8); return () od else if trap_instruction \ traps then do ticc_trap_type \ gets (\s. (ticc_trap_type_val s)); write_cpu_tt (word_cat (1::word1) ticc_trap_type); return () od \<^cancel>\else if interrupt_level > 0 then\ \ \We don't consider \interrupt_level\\ else return () od" definition exe_trap_st_pc :: "unit \ ('a::len,unit) sparc_state_monad" where "exe_trap_st_pc _ \ do annul \ gets (\s. (annul_val s)); pc_val \ gets (\s. (cpu_reg_val PC s)); npc_val \ gets (\s. (cpu_reg_val nPC s)); curr_win \ get_curr_win(); if annul = False then do write_reg pc_val curr_win (word_of_int 17); write_reg npc_val curr_win (word_of_int 18); return () od else \ \\annul = True\\ do write_reg npc_val curr_win (word_of_int 17); write_reg (npc_val + 4) curr_win (word_of_int 18); set_annul False; return () od od" definition exe_trap_wr_pc :: "unit \ ('a::len,unit) sparc_state_monad" where "exe_trap_wr_pc _ \ do psr_val \ gets (\s. (cpu_reg_val PSR s)); new_psr_val \ gets (\s. (update_S (1::word1) psr_val)); write_cpu new_psr_val PSR; reset_trap \ gets (\s. (reset_trap_val s)); tbr_val \ gets (\s. (cpu_reg_val TBR s)); if reset_trap = False then do write_cpu tbr_val PC; write_cpu (tbr_val + 4) nPC; return () od else \ \\reset_trap = True\\ do write_cpu 0 PC; write_cpu 4 nPC; set_reset_trap False; return () od od" definition execute_trap :: "unit \ ('a::len,unit) sparc_state_monad" where "execute_trap _ \ do select_trap(); err_mode \ gets (\s. (err_mode_val s)); if err_mode = True then \ \The SparcV8 manual doesn't say what to do.\ return () else do psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. ((ucast (get_S psr_val))::word1)); curr_win \ get_curr_win(); new_cwp \ gets (\s. ((word_of_int (((uint curr_win) - 1) mod NWINDOWS)))::word5); new_psr_val \ gets (\s. (update_PSR_exe_trap new_cwp (0::word1) s_val psr_val)); write_cpu new_psr_val PSR; exe_trap_st_pc(); exe_trap_wr_pc(); return () od od" definition dispatch_instruction :: "instruction \ ('a::len,unit) sparc_state_monad" where "dispatch_instruction instr \ let instr_name = fst instr in do traps \ gets (\s. (get_trap_set s)); if traps = {} then if instr_name \ {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} then load_instr instr else if instr_name \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD} then store_instr instr else if instr_name \ {sethi_type SETHI} then sethi_instr instr else if instr_name \ {nop_type NOP} then nop_instr instr else if instr_name \ {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} then logical_instr instr else if instr_name \ {shift_type SLL,shift_type SRL,shift_type SRA} then shift_instr instr else if instr_name \ {arith_type ADD,arith_type ADDcc,arith_type ADDX} then add_instr instr else if instr_name \ {arith_type SUB,arith_type SUBcc,arith_type SUBX} then sub_instr instr else if instr_name \ {arith_type UMUL,arith_type SMUL,arith_type SMULcc} then mul_instr instr else if instr_name \ {arith_type UDIV,arith_type UDIVcc,arith_type SDIV} then div_instr instr else if instr_name \ {ctrl_type SAVE,ctrl_type RESTORE} then save_restore_instr instr else if instr_name \ {call_type CALL} then call_instr instr else if instr_name \ {ctrl_type JMPL} then jmpl_instr instr else if instr_name \ {ctrl_type RETT} then rett_instr instr else if instr_name \ {sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM, sreg_type RDTBR} then read_state_reg_instr instr else if instr_name \ {sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM, sreg_type WRTBR} then write_state_reg_instr instr else if instr_name \ {load_store_type FLUSH} then flush_instr instr else if instr_name \ {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} then branch_instr instr else fail () else return () od" definition supported_instruction :: "sparc_operation \ bool" where "supported_instruction instr \ if 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, 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} then True else False " definition execute_instr_sub1 :: "instruction \ ('a::len,unit) sparc_state_monad" where "execute_instr_sub1 instr \ do instr_name \ gets (\s. (fst instr)); traps2 \ gets (\s. (get_trap_set s)); if traps2 = {} \ instr_name \ {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} then do npc_val \ gets (\s. (cpu_reg_val nPC s)); write_cpu npc_val PC; write_cpu (npc_val + 4) nPC; return () od else return () od" definition execute_instruction :: "unit \ ('a::len,unit) sparc_state_monad" where "execute_instruction _ \ do traps \ gets (\s. (get_trap_set s)); if traps = {} then do exe_mode \ gets (\s. (exe_mode_val s)); if exe_mode = True then do modify (\s. (delayed_pool_write s)); fetch_result \ gets (\s. (fetch_instruction s)); case fetch_result of Inl e1 \ (do \ \Memory address in PC is not aligned.\ \ \Actually, SparcV8 manual doens't check alignment here.\ raise_trap instruction_access_exception; return () od) | Inr v1 \ (do dec \ gets (\s. (decode_instruction v1)); case dec of Inl e2 \ (\ \Instruction is ill-formatted.\ fail () ) | Inr v2 \ (do instr \ gets (\s. (v2)); annul \ gets (\s. (annul_val s)); if annul = False then do dispatch_instruction instr; execute_instr_sub1 instr; return () od else \ \\annul \ False\\ do set_annul False; npc_val \ gets (\s. (cpu_reg_val nPC s)); write_cpu npc_val PC; write_cpu (npc_val + 4) nPC; return () od od) od) od else return () \ \Not in \execute_mode\.\ od else \ \traps is not empty, which means \trap = 1\.\ do execute_trap(); return () od od" definition NEXT :: "('a::len)sparc_state \ ('a)sparc_state option" where "NEXT s \ case execute_instruction () s of (_,True) \ None | (s',False) \ Some (snd s')" +context + includes bit_operations_syntax +begin + definition good_context :: "('a::len) sparc_state \ bool" where "good_context s \ let traps = get_trap_set s; psr_val = cpu_reg_val PSR s; et_val = get_ET psr_val; rt_val = reset_trap_val s in if traps \ {} \ rt_val = False \ et_val = 0 then False \ \enter \error_mode\ in \select_traps\.\ else let s' = delayed_pool_write s in case fetch_instruction s' of \ \\instruction_access_exception\ is handled in the next state.\ Inl _ \ True |Inr v \ ( case decode_instruction v of Inl _ \ False |Inr instr \ ( let annul = annul_val s' in if annul = True then True else \ \\annul = False\\ if supported_instruction (fst instr) then \ \The only instruction that could fail is \RETT\.\ if (fst instr) = ctrl_type RETT then let curr_win_r = (get_CWP (cpu_reg_val PSR s')); new_cwp_int_r = (((uint curr_win_r) + 1) mod NWINDOWS); wim_val_r = cpu_reg_val WIM s'; psr_val_r = cpu_reg_val PSR s'; et_val_r = get_ET psr_val_r; s_val_r = (ucast (get_S psr_val_r))::word1; op_list_r = snd instr; addr_r = get_addr (snd instr) s' in if et_val_r = 1 then True else if s_val_r = 0 then False else if (get_WIM_bit (nat new_cwp_int_r) wim_val_r) \ 0 then False else if ((AND) addr_r (0b00000000000000000000000000000011::word32)) \ 0 then False else True else True else False \ \Unsupported instruction.\ ) ) " +end + function (sequential) seq_exec:: "nat \ ('a::len,unit) sparc_state_monad" where "seq_exec 0 = return ()" | "seq_exec n = (do execute_instruction(); (seq_exec (n-1)) od) " by pat_completeness auto termination by lexicographic_order type_synonym leon3_state = "(word_length5) sparc_state" type_synonym ('e) leon3_state_monad = "(leon3_state, 'e) det_monad" definition execute_leon3_instruction:: "unit \ (unit) leon3_state_monad" where "execute_leon3_instruction \ execute_instruction" definition seq_exec_leon3:: "nat \ (unit) leon3_state_monad" where "seq_exec_leon3 \ seq_exec" end diff --git a/thys/SPARCv8/SparcModel_MMU/Sparc_Instruction.thy b/thys/SPARCv8/SparcModel_MMU/Sparc_Instruction.thy --- a/thys/SPARCv8/SparcModel_MMU/Sparc_Instruction.thy +++ b/thys/SPARCv8/SparcModel_MMU/Sparc_Instruction.thy @@ -1,2788 +1,2794 @@ (* * 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, David Sanan. *) section \SPARC instruction model\ theory Sparc_Instruction imports Main Sparc_Types Sparc_State "HOL-Eisbach.Eisbach_Tools" begin text\ This theory provides a formal model for assembly instruction to be executed in the model. An instruction is defined as a tuple composed of a @{term sparc_operation} element, defining the operation the instruction carries out, and a list of operands @{term inst_operand}. @{term inst_operand} can be a user register @{term user_reg} or a memory address @{term mem_add_type}. \ datatype inst_operand = W5 word5 |W30 word30 |W22 word22 |Cond word4 |Flag word1 |Asi asi_type |Simm13 word13 |Opf word9 |Imm7 word7 primrec get_operand_w5::"inst_operand \ word5" where "get_operand_w5 (W5 r) = r" primrec get_operand_w30::"inst_operand \ word30" where "get_operand_w30 (W30 r) = r" primrec get_operand_w22::"inst_operand \ word22" where "get_operand_w22 (W22 r) = r" primrec get_operand_cond::"inst_operand \ word4" where "get_operand_cond (Cond r) = r" primrec get_operand_flag::"inst_operand \ word1" where "get_operand_flag (Flag r) = r" primrec get_operand_asi::"inst_operand \ asi_type" where "get_operand_asi (Asi r) = r" primrec get_operand_simm13::"inst_operand \ word13" where "get_operand_simm13 (Simm13 r) = r" primrec get_operand_opf::"inst_operand \ word9" where "get_operand_opf (Opf r) = r" primrec get_operand_imm7:: "inst_operand \ word7" where "get_operand_imm7 (Imm7 r) = r" +context + includes bit_operations_syntax +begin + type_synonym instruction = "(sparc_operation \ inst_operand list)" definition get_op::"word32 \ int" where "get_op w \ uint (w >> 30)" definition get_op2::"word32 \ int" where "get_op2 w \ let mask_op2 = 0b00000001110000000000000000000000 in uint (((AND) mask_op2 w) >> 22)" definition get_op3::"word32 \ int" where "get_op3 w \ let mask_op3 = 0b00000001111110000000000000000000 in uint (((AND) mask_op3 w) >> 19)" definition get_disp30::"word32 \ int" where "get_disp30 w \ let mask_disp30 = 0b00111111111111111111111111111111 in uint ((AND) mask_disp30 w)" definition get_a::"word32 \ int" where "get_a w \ let mask_a = 0b00100000000000000000000000000000 in uint (((AND) mask_a w) >> 29)" definition get_cond::"word32 \ int" where "get_cond w \ let mask_cond = 0b00011110000000000000000000000000 in uint (((AND) mask_cond w) >> 25)" definition get_disp_imm22::"word32 \ int" where "get_disp_imm22 w \ let mask_disp_imm22 = 0b00000000001111111111111111111111 in uint ((AND) mask_disp_imm22 w)" definition get_rd::"word32 \ int" where "get_rd w \ let mask_rd = 0b00111110000000000000000000000000 in uint (((AND) mask_rd w) >> 25)" definition get_rs1::"word32 \ int" where "get_rs1 w \ let mask_rs1 = 0b00000000000001111100000000000000 in uint (((AND) mask_rs1 w) >> 14)" definition get_i::"word32 \ int" where "get_i w \ let mask_i = 0b00000000000000000010000000000000 in uint (((AND) mask_i w) >> 13)" definition get_opf::"word32 \ int" where "get_opf w \ let mask_opf = 0b00000000000000000011111111100000 in uint (((AND) mask_opf w) >> 5)" definition get_rs2::"word32 \ int" where "get_rs2 w \ let mask_rs2 = 0b00000000000000000000000000011111 in uint ((AND) mask_rs2 w)" definition get_simm13::"word32 \ int" where "get_simm13 w \ let mask_simm13 = 0b00000000000000000001111111111111 in uint ((AND) mask_simm13 w)" definition get_asi::"word32 \ int" where "get_asi w \ let mask_asi = 0b00000000000000000001111111100000 in uint (((AND) mask_asi w) >> 5)" definition get_trap_cond:: "word32 \ int" where "get_trap_cond w \ let mask_cond = 0b00011110000000000000000000000000 in uint (((AND) mask_cond w) >> 25)" definition get_trap_imm7:: "word32 \ int" where "get_trap_imm7 w \ let mask_imm7 = 0b00000000000000000000000001111111 in uint ((AND) mask_imm7 w)" definition parse_instr_f1::"word32 \ (Exception list + instruction)" where \ \\CALL\, with a single operand \disp30+"00"\\ "parse_instr_f1 w \ Inr (call_type CALL,[W30 (word_of_int (get_disp30 w))])" definition parse_instr_f2::"word32 \ (Exception list + instruction)" where "parse_instr_f2 w \ let op2 = get_op2 w in if op2 = uint(0b100::word3) then \ \\SETHI\ or \NOP\\ let rd = get_rd w in let imm22 = get_disp_imm22 w in if rd = 0 \ imm22 = 0 then \ \\NOP\\ Inr (nop_type NOP,[]) else \ \\SETHI\, with operands \[imm22,rd]\\ Inr (sethi_type SETHI,[(W22 (word_of_int imm22)), (W5 (word_of_int rd))]) else if op2 = uint(0b010::word3) then \ \\Bicc\, with operands \[a,disp22]\\ let cond = get_cond w in let flaga = Flag (word_of_int (get_a w)) in let disp22 = W22 (word_of_int (get_disp_imm22 w)) in if cond = uint(0b0001::word4) then \ \\BE\\ Inr (bicc_type BE,[flaga,disp22]) else if cond = uint(0b1001::word4) then \ \\BNE\\ Inr (bicc_type BNE,[flaga,disp22]) else if cond = uint(0b1100::word4) then \ \\BGU\\ Inr (bicc_type BGU,[flaga,disp22]) else if cond = uint(0b0010::word4) then \ \\BLE\\ Inr (bicc_type BLE,[flaga,disp22]) else if cond = uint(0b0011::word4) then \ \\BL\\ Inr (bicc_type BL,[flaga,disp22]) else if cond = uint(0b1011::word4) then \ \\BGE\\ Inr (bicc_type BGE,[flaga,disp22]) else if cond = uint(0b0110::word4) then \ \\BNEG\\ Inr (bicc_type BNEG,[flaga,disp22]) else if cond = uint(0b1010::word4) then \ \\BG\\ Inr (bicc_type BG,[flaga,disp22]) else if cond = uint(0b0101::word4) then \ \\BCS\\ Inr (bicc_type BCS,[flaga,disp22]) else if cond = uint(0b0100::word4) then \ \\BLEU\\ Inr (bicc_type BLEU,[flaga,disp22]) else if cond = uint(0b1101::word4) then \ \\BCC\\ Inr (bicc_type BCC,[flaga,disp22]) else if cond = uint(0b1000::word4) then \ \\BA\\ Inr (bicc_type BA,[flaga,disp22]) else if cond = uint(0b0000::word4) then \ \\BN\\ Inr (bicc_type BN,[flaga,disp22]) else if cond = uint(0b1110::word4) then \ \\BPOS\\ Inr (bicc_type BPOS,[flaga,disp22]) else if cond = uint(0b1111::word4) then \ \\BVC\\ Inr (bicc_type BVC,[flaga,disp22]) else if cond = uint(0b0111::word4) then \ \\BVS\\ Inr (bicc_type BVS,[flaga,disp22]) else Inl [invalid_cond_f2] else Inl [invalid_op2_f2] " text \We don't consider floating-point operations, so we don't consider the third type of format 3.\ definition parse_instr_f3::"word32 \ (Exception list + instruction)" where "parse_instr_f3 w \ let this_op = get_op w in let rd = get_rd w in let op3 = get_op3 w in let rs1 = get_rs1 w in let flagi = get_i w in let asi = get_asi w in let rs2 = get_rs2 w in let simm13 = get_simm13 w in if this_op = uint(0b11::word2) then \ \Load and Store\ \ \If an instruction accesses alternative space but \flagi = 1\,\ \ \may need to throw a trap.\ if op3 = uint(0b001001::word6) then \ \\LDSB\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LDSB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LDSB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b011001::word6) then \ \\LDSBA\\ Inr (load_store_type LDSBA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b001010::word6) then \ \\LDSH\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LDSH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LDSH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b011010::word6) then \ \\LDSHA\\ Inr (load_store_type LDSHA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000001::word6) then \ \\LDUB\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LDUB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LDUB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010001::word6) then \ \\LDUBA\\ Inr (load_store_type LDUBA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000010::word6) then \ \\LDUH\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LDUH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LDUH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010010::word6) then \ \\LDUHA\\ Inr (load_store_type LDUHA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000000::word6) then \ \\LD\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010000::word6) then \ \\LDA\\ Inr (load_store_type LDA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000011::word6) then \ \\LDD\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LDD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LDD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010011::word6) then \ \\LDDA\\ Inr (load_store_type LDDA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b001101::word6) then \ \\LDSTUB\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type LDSTUB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type LDSTUB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b011101::word6) then \ \\LDSTUBA\\ Inr (load_store_type LDSTUBA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000101::word6) then \ \\STB\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type STB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type STB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010101::word6) then \ \\STBA\\ Inr (load_store_type STBA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000110::word6) then \ \\STH\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type STH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type STH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010110::word6) then \ \\STHA\\ Inr (load_store_type STHA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000100::word6) then \ \\ST\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type ST,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type ST,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010100::word6) then \ \\STA\\ Inr (load_store_type STA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b000111::word6) then \ \\STD\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type STD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type STD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b010111::word6) then \ \\STDA\\ Inr (load_store_type STDA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else if op3 = uint(0b001111::word6) then \ \\SWAP\\ if flagi = 1 then \ \Operant list is \[i,rs1,simm13,rd]\\ Inr (load_store_type SWAP,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else \ \Operant list is \[i,rs1,rs2,rd]\\ Inr (load_store_type SWAP,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else if op3 = uint(0b011111::word6) then \ \\SWAPA\\ Inr (load_store_type SWAPA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (Asi (word_of_int asi)), (W5 (word_of_int rd))]) else Inl [invalid_op3_f3_op11] else if this_op = uint(0b10::word2) then \ \Others\ if op3 = uint(0b111000::word6) then \ \\JMPL\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (ctrl_type JMPL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (ctrl_type JMPL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b111001::word6) then \ \\RETT\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ctrl_type RETT,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,simm13]\\ Inr (ctrl_type RETT,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13))]) \ \The following are Read and Write instructions,\ \ \only return \[rs1,rd]\ as operand.\ else if op3 = uint(0b101000::word6) \ rs1 \ 0 then \ \\RDASR\\ if rs1 = uint(0b01111::word6) \ rd = 0 then \ \\STBAR\ is a special case of \RDASR\\ Inr (load_store_type STBAR,[]) else Inr (sreg_type RDASR,[(W5 (word_of_int rs1)), (W5 (word_of_int rd))]) else if op3 = uint(0b101000::word6) \ rs1 = 0 then \ \\RDY\\ Inr (sreg_type RDY,[(W5 (word_of_int rs1)), (W5 (word_of_int rd))]) else if op3 = uint(0b101001::word6) then \ \\RDPSR\\ Inr (sreg_type RDPSR,[(W5 (word_of_int rs1)), (W5 (word_of_int rd))]) else if op3 = uint(0b101010::word6) then \ \\RDWIM\\ Inr (sreg_type RDWIM,[(W5 (word_of_int rs1)), (W5 (word_of_int rd))]) else if op3 = uint(0b101011::word6) then \ \\RDTBR\\ Inr (sreg_type RDTBR,[(W5 (word_of_int rs1)), (W5 (word_of_int rd))]) else if op3 = uint(0b110000::word6) \ rd \ 0 then \ \\WRASR\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (sreg_type WRASR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (sreg_type WRASR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b110000::word6) \ rd = 0 then \ \\WRY\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (sreg_type WRY,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (sreg_type WRY,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b110001::word6) then \ \\WRPSR\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (sreg_type WRPSR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (sreg_type WRPSR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b110010::word6) then \ \\WRWIM\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (sreg_type WRWIM,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (sreg_type WRWIM,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b110011::word6) then \ \\WRTBR\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (sreg_type WRTBR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (sreg_type WRTBR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) \ \\FLUSH\ instruction\ else if op3 = uint(0b111011::word6) then \ \\FLUSH\\ if flagi = 0 then \ \return \[1,rs1,rs2]\\ Inr (load_store_type FLUSH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,simm13]\\ Inr (load_store_type FLUSH,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13))]) \ \The following are arithmetic instructions.\ else if op3 = uint(0b000001::word6) then \ \\AND\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ANDs,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ANDs,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010001::word6) then \ \\ANDcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ANDcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ANDcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b000101::word6) then \ \\ANDN\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ANDN,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ANDN,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010101::word6) then \ \\ANDNcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ANDNcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ANDNcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b000010::word6) then \ \\OR\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ORs,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ORs,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010010::word6) then \ \\ORcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ORcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ORcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b000110::word6) then \ \\ORN\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ORN,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ORN,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010110::word6) then \ \\ORNcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type ORNcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type ORNcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b000011::word6) then \ \\XORs\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type XORs,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type XORs,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010011::word6) then \ \\XORcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type XORcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type XORcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b000111::word6) then \ \\XNOR\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type XNOR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type XNOR,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010111::word6) then \ \\XNORcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (logic_type XNORcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (logic_type XNORcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b100101::word6) then \ \\SLL\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (shift_type SLL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,shcnt,rd]\\ let shcnt = rs2 in Inr (shift_type SLL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int shcnt)), (W5 (word_of_int rd))]) else if op3 = uint (0b100110::word6) then \ \\SRL\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (shift_type SRL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,shcnt,rd]\\ let shcnt = rs2 in Inr (shift_type SRL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int shcnt)), (W5 (word_of_int rd))]) else if op3 = uint(0b100111::word6) then \ \\SRA\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (shift_type SRA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,shcnt,rd]\\ let shcnt = rs2 in Inr (shift_type SRA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int shcnt)), (W5 (word_of_int rd))]) else if op3 = uint(0b000000::word6) then \ \\ADD\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type ADD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type ADD,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010000::word6) then \ \\ADDcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type ADDcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type ADDcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b001000::word6) then \ \\ADDX\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type ADDX,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type ADDX,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b011000::word6) then \ \\ADDXcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type ADDXcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type ADDXcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b100000::word6) then \ \\TADDcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type TADDcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type TADDcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b100010::word6) then \ \\TADDccTV\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type TADDccTV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type TADDccTV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b000100::word6) then \ \\SUB\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SUB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SUB,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b010100::word6) then \ \\SUBcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SUBcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SUBcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b001100::word6) then \ \\SUBX\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SUBX,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SUBX,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b011100::word6) then \ \\SUBXcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SUBXcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SUBXcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b100001::word6) then \ \\TSUBcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type TSUBcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type TSUBcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b100011::word6) then \ \\TSUBccTV\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type TSUBccTV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type TSUBccTV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b100100::word6) then \ \\MULScc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type MULScc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type MULScc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b001010::word6) then \ \\UMUL\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type UMUL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type UMUL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b011010::word6) then \ \\UMULcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type UMULcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type UMULcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b001011::word6) then \ \\SMUL\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SMUL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SMUL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b011011::word6) then \ \\SMULcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SMULcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SMULcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b001110::word6) then \ \\UDIV\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type UDIV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type UDIV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b011110::word6) then \ \\UDIVcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type UDIVcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type UDIVcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b001111::word6) then \ \\SDIV\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SDIV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SDIV,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b011111::word6) then \ \\SDIVcc\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (arith_type SDIVcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (arith_type SDIVcc,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b111100::word6) then \ \\SAVE\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (ctrl_type SAVE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (ctrl_type SAVE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b111101::word6) then \ \\RESTORE\\ if flagi = 0 then \ \return \[i,rs1,rs2,rd]\\ Inr (ctrl_type RESTORE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2)), (W5 (word_of_int rd))]) else \ \return \[i,rs1,simm13,rd]\\ Inr (ctrl_type RESTORE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Simm13 (word_of_int simm13)), (W5 (word_of_int rd))]) else if op3 = uint(0b111010::word6) then \ \\Ticc\\ let trap_cond = get_trap_cond w in let trap_imm7 = get_trap_imm7 w in if trap_cond = uint(0b1000::word4) then \ \\TA\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TA,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0000::word4) then \ \\TN\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TN,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TN,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1001::word4) then \ \\TNE\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TNE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TNE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0001::word4) then \ \\TE\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1010::word4) then \ \\TG\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TG,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TG,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0010::word4) then \ \\TLE\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TLE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TLE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1011::word4) then \ \\TGE\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TGE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TGE,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0011::word4) then \ \\TL\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TL,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1100::word4) then \ \\TGU\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TGU,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TGU,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0100::word4) then \ \\TLEU\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TLEU,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TLEU,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1101::word4) then \ \\TCC\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TCC,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TCC,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0101::word4) then \ \\TCS\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TCS,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TCS,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1110::word4) then \ \\TPOS\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TPOS,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TPOS,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0110::word4) then \ \\TNEG\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TNEG,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TNEG,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b1111::word4) then \ \\TVC\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TVC,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TVC,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else if trap_cond = uint(0b0111::word4) then \ \\TVS\\ if flagi = 0 then \ \return \[i,rs1,rs2]\\ Inr (ticc_type TVS,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (W5 (word_of_int rs2))]) else \ \return \[i,rs1,trap_imm7]\\ Inr (ticc_type TVS,[(Flag (word_of_int flagi)), (W5 (word_of_int rs1)), (Imm7 (word_of_int trap_imm7))]) else Inl [invalid_trap_cond] else Inl [invalid_op3_f3_op10] else Inl [invalid_op_f3] " text \Read the word32 value from the Program Counter in the current state. Find the instruction in the memory address of the word32 value. Return a word32 value of the insturction.\ definition fetch_instruction::"('a) sparc_state \ (Exception list + word32)" where "fetch_instruction s \ \ \\pc_val\ is the 32-bit memory address of the instruction.\ let pc_val = cpu_reg_val PC s; psr_val = cpu_reg_val PSR s; s_val = get_S psr_val; asi = if s_val = 0 then word_of_int 8 else word_of_int 9 in \ \Check if \pc_val\ is aligned to 4-byte (32-bit) boundary.\ \ \That is, check if the least significant two bits of\ \ \\pc_val\ are 0s.\ if uint((AND) (0b00000000000000000000000000000011) pc_val) = 0 then \ \Get the 32-bit value from the address of \pc_val\\ \ \to the address of \pc_val+3\\ let (mem_result,n_s) = memory_read asi pc_val s in case mem_result of None \ Inl [fetch_instruction_error] |Some v \ Inr v else Inl [fetch_instruction_error] " text \Decode the word32 value of an instruction into the name of the instruction and its operands.\ definition decode_instruction::"word32 \ Exception list + instruction" where "decode_instruction w \ let this_op = get_op w in if this_op = uint(0b01::word2) then \ \Instruction format 1\ parse_instr_f1 w else if this_op = uint(0b00::word2) then \ \Instruction format 2\ parse_instr_f2 w else \ \\op = 11 0r 10\, instruction format 3\ parse_instr_f3 w " text \Get the current window from the PSR\ definition get_curr_win::"unit \ ('a,('a::len window_size)) sparc_state_monad" where "get_curr_win _ \ do curr_win \ gets (\s. (ucast (get_CWP (cpu_reg_val PSR s)))); return curr_win od" text \Operational semantics for CALL\ definition call_instr::"instruction \ ('a::len,unit) sparc_state_monad" where "call_instr instr \ let op_list = snd instr; mem_addr = ((ucast (get_operand_w30 (op_list!0)))::word32) << 2 in do curr_win \ get_curr_win(); pc_val \ gets (\s. (cpu_reg_val PC s)); npc_val \ gets (\s. (cpu_reg_val nPC s)); write_reg pc_val curr_win (word_of_int 15); write_cpu npc_val PC; write_cpu (pc_val + mem_addr) nPC; return () od" text\Evaluate icc based on the bits N, Z, V, C in PSR and the type of branching instruction. See Sparcv8 manual Page 178.\ definition eval_icc::"sparc_operation \ word1 \ word1 \ word1 \ word1 \ int" where "eval_icc instr_name n_val z_val v_val c_val \ if instr_name = bicc_type BNE then if z_val = 0 then 1 else 0 else if instr_name = bicc_type BE then if z_val = 1 then 1 else 0 else if instr_name = bicc_type BG then if ((OR) z_val (n_val XOR v_val)) = 0 then 1 else 0 else if instr_name = bicc_type BLE then if ((OR) z_val (n_val XOR v_val)) = 1 then 1 else 0 else if instr_name = bicc_type BGE then if (n_val XOR v_val) = 0 then 1 else 0 else if instr_name = bicc_type BL then if (n_val XOR v_val) = 1 then 1 else 0 else if instr_name = bicc_type BGU then if (c_val = 0 \ z_val = 0) then 1 else 0 else if instr_name = bicc_type BLEU then if (c_val = 1 \ z_val = 1) then 1 else 0 else if instr_name = bicc_type BCC then if c_val = 0 then 1 else 0 else if instr_name = bicc_type BCS then if c_val = 1 then 1 else 0 else if instr_name = bicc_type BNEG then if n_val = 1 then 1 else 0 else if instr_name = bicc_type BA then 1 else if instr_name = bicc_type BN then 0 else if instr_name = bicc_type BPOS then if n_val = 0 then 1 else 0 else if instr_name = bicc_type BVC then if v_val = 0 then 1 else 0 else if instr_name = bicc_type BVS then if v_val = 1 then 1 else 0 else -1 " definition branch_instr_sub1:: "sparc_operation \ ('a) sparc_state \ int" where "branch_instr_sub1 instr_name s \ let n_val = get_icc_N ((cpu_reg s) PSR); z_val = get_icc_Z ((cpu_reg s) PSR); v_val = get_icc_V ((cpu_reg s) PSR); c_val = get_icc_C ((cpu_reg s) PSR) in eval_icc instr_name n_val z_val v_val c_val" text \Operational semantics for Branching insturctions. Return exception or a bool value for annulment. If the bool value is 1, then the delay instruciton is not executed, otherwise the delay instruction is executed.\ definition branch_instr::"instruction \ ('a,unit) sparc_state_monad" where "branch_instr instr \ let instr_name = fst instr; op_list = snd instr; disp22 = get_operand_w22 (op_list!1); flaga = get_operand_flag (op_list!0) in do icc_val \ gets( \s. (branch_instr_sub1 instr_name s)); npc_val \ gets (\s. (cpu_reg_val nPC s)); pc_val \ gets (\s. (cpu_reg_val PC s)); write_cpu npc_val PC; if icc_val = 1 then do write_cpu (pc_val + (sign_ext24 (((ucast(disp22))::word24) << 2))) nPC; if (instr_name = bicc_type BA) \ (flaga = 1) then do set_annul True; return () od else return () od else \ \\icc_val = 0\\ do write_cpu (npc_val + 4) nPC; if flaga = 1 then do set_annul True; return () od else return () od od" text \Operational semantics for NOP\ definition nop_instr::"instruction \ ('a,unit) sparc_state_monad" where "nop_instr instr \ return ()" text \Operational semantics for SETHI\ definition sethi_instr::"instruction \ ('a::len,unit) sparc_state_monad" where "sethi_instr instr \ let op_list = snd instr; imm22 = get_operand_w22 (op_list!0); rd = get_operand_w5 (op_list!1) in if rd \ 0 then do curr_win \ get_curr_win(); write_reg (((ucast(imm22))::word32) << 10) curr_win rd; return () od else return () " text \ Get \operand2\ based on the flag \i\, \rs1\, \rs2\, and \simm13\. If \i = 0\ then \operand2 = r[rs2]\, else \operand2 = sign_ext13(simm13)\. \op_list\ should be \[i,rs1,rs2,\]\ or \[i,rs1,simm13,\]\. \ definition get_operand2::"inst_operand list \ ('a::len) sparc_state \ virtua_address" where "get_operand2 op_list s \ let flagi = get_operand_flag (op_list!0); curr_win = ucast (get_CWP (cpu_reg_val PSR s)) in if flagi = 0 then let rs2 = get_operand_w5 (op_list!2); rs2_val = user_reg_val curr_win rs2 s in rs2_val else let ext_simm13 = sign_ext13 (get_operand_simm13 (op_list!2)) in ext_simm13 " text \ Get \operand2_val\ based on the flag \i\, \rs1\, \rs2\, and \simm13\. If \i = 0\ then \operand2_val = uint r[rs2]\, else \operand2_val = sint sign_ext13(simm13)\. \op_list\ should be \[i,rs1,rs2,\]\ or \[i,rs1,simm13,\]\. \ definition get_operand2_val::"inst_operand list \ ('a::len) sparc_state \ int" where "get_operand2_val op_list s \ let flagi = get_operand_flag (op_list!0); curr_win = ucast (get_CWP (cpu_reg_val PSR s)) in if flagi = 0 then let rs2 = get_operand_w5 (op_list!2); rs2_val = user_reg_val curr_win rs2 s in sint rs2_val else let ext_simm13 = sign_ext13 (get_operand_simm13 (op_list!2)) in sint ext_simm13 " text \ Get the address based on the flag \i\, \rs1\, \rs2\, and \simm13\. If \i = 0\ then \addr = r[rs1] + r[rs2]\, else \addr = r[rs1] + sign_ext13(simm13)\. \op_list\ should be \[i,rs1,rs2,\]\ or \[i,rs1,simm13,\]\. \ definition get_addr::"inst_operand list \ ('a::len) sparc_state \ virtua_address" where "get_addr op_list s \ let rs1 = get_operand_w5 (op_list!1); curr_win = ucast (get_CWP (cpu_reg_val PSR s)); rs1_val = user_reg_val curr_win rs1 s; op2 = get_operand2 op_list s in (rs1_val + op2) " text \Operational semantics for JMPL\ definition jmpl_instr::"instruction \ ('a::len,unit) sparc_state_monad" where "jmpl_instr instr \ let op_list = snd instr; rd = get_operand_w5 (op_list!3) in do curr_win \ get_curr_win(); jmp_addr \ gets (\s. (get_addr op_list s)); if ((AND) jmp_addr 0b00000000000000000000000000000011) \ 0 then do raise_trap mem_address_not_aligned; return () od else do rd_next_val \ gets (\s. (if rd \ 0 then (cpu_reg_val PC s) else user_reg_val curr_win rd s)); write_reg rd_next_val curr_win rd; npc_val \ gets (\s. (cpu_reg_val nPC s)); write_cpu npc_val PC; write_cpu jmp_addr nPC; return () od od" text \Operational semantics for RETT\ definition rett_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "rett_instr instr \ let op_list = snd instr in do psr_val \ gets (\s. (cpu_reg_val PSR s)); curr_win \ gets (\s. (get_CWP (cpu_reg_val PSR s))); new_cwp \ gets (\s. (word_of_int (((uint curr_win) + 1) mod NWINDOWS))); new_cwp_int \ gets (\s. ((uint curr_win) + 1) mod NWINDOWS); addr \ gets (\s. (get_addr op_list s)); et_val \ gets (\s. ((ucast (get_ET psr_val))::word1)); s_val \ gets (\s. ((ucast (get_S psr_val))::word1)); ps_val \ gets (\s. ((ucast (get_PS psr_val))::word1)); wim_val \ gets (\s. (cpu_reg_val WIM s)); npc_val \ gets (\s. (cpu_reg_val nPC s)); if et_val = 1 then if s_val = 0 then do raise_trap privileged_instruction; return () od else do raise_trap illegal_instruction; return () od else if s_val = 0 then do write_cpu_tt (0b00000011::word8); set_exe_mode False; set_err_mode True; raise_trap privileged_instruction; fail () od else if (get_WIM_bit (nat new_cwp_int) wim_val) \ 0 then do write_cpu_tt (0b00000110::word8); set_exe_mode False; set_err_mode True; raise_trap window_underflow; fail () od else if ((AND) addr (0b00000000000000000000000000000011::word32)) \ 0 then do write_cpu_tt (0b00000111::word8); set_exe_mode False; set_err_mode True; raise_trap mem_address_not_aligned; fail () od else do write_cpu npc_val PC; write_cpu addr nPC; new_psr_val \ gets (\s. (update_PSR_rett new_cwp 1 ps_val psr_val)); write_cpu new_psr_val PSR; return () od od" definition save_retore_sub1 :: "word32 \ word5 \ word5 \ ('a::len,unit) sparc_state_monad" where "save_retore_sub1 result new_cwp rd \ do psr_val \ gets (\s. (cpu_reg_val PSR s)); new_psr_val \ gets (\s. (update_CWP new_cwp psr_val)); write_cpu new_psr_val PSR; \ \Change \CWP\ to the new window value.\ write_reg result (ucast new_cwp) rd; \ \Write result in \rd\ of the new window.\ return () od" text \Operational semantics for SAVE and RESTORE.\ definition save_restore_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "save_restore_instr instr \ let instr_name = fst instr; op_list = snd instr; rd = get_operand_w5 (op_list!3) in do psr_val \ gets (\s. (cpu_reg_val PSR s)); curr_win \ get_curr_win(); wim_val \ gets (\s. (cpu_reg_val WIM s)); if instr_name = ctrl_type SAVE then do new_cwp \ gets (\s. ((word_of_int (((uint curr_win) - 1) mod NWINDOWS)))::word5); if (get_WIM_bit (unat new_cwp) wim_val) \ 0 then do raise_trap window_overflow; return () od else do result \ gets (\s. (get_addr op_list s)); \ \operands are from the old window.\ save_retore_sub1 result new_cwp rd od od else \ \\instr_name = RESTORE\\ do new_cwp \ gets (\s. ((word_of_int (((uint curr_win) + 1) mod NWINDOWS)))::word5); if (get_WIM_bit (unat new_cwp) wim_val) \ 0 then do raise_trap window_underflow; return () od else do result \ gets (\s. (get_addr op_list s)); \ \operands are from the old window.\ save_retore_sub1 result new_cwp rd od od od" definition flush_cache_line :: "word32 \ ('a,unit) sparc_state_monad" where "flush_cache_line \ undefined" definition flush_Ibuf_and_pipeline :: "word32 \ ('a,unit) sparc_state_monad" where "flush_Ibuf_and_pipeline \ undefined" text \Operational semantics for FLUSH. Flush the all the caches.\ definition flush_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "flush_instr instr \ let op_list = snd instr in do addr \ gets (\s. (get_addr op_list s)); modify (\s. (flush_cache_all s)); \<^cancel>\flush_cache_line(addr);\ \<^cancel>\flush_Ibuf_and_pipeline(addr);\ return () od" text \Operational semantics for read state register instructions. We do not consider RDASR here.\ definition read_state_reg_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "read_state_reg_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!0); rd = get_operand_w5 (op_list!1) in do curr_win \ get_curr_win(); psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. (get_S psr_val)); if (instr_name \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (instr_name = sreg_type RDASR \ privileged_ASR rs1)) \ ((ucast s_val)::word1) = 0 then do raise_trap privileged_instruction; return () od else if illegal_instruction_ASR rs1 then do raise_trap illegal_instruction; return () od else if rd \ 0 then if instr_name = sreg_type RDY then do y_val \ gets (\s. (cpu_reg_val Y s)); write_reg y_val curr_win rd; return () od else if instr_name = sreg_type RDASR then do asr_val \ gets (\s. (cpu_reg_val (ASR rs1) s)); write_reg asr_val curr_win rd; return () od else if instr_name = sreg_type RDPSR then do write_reg psr_val curr_win rd; return () od else if instr_name = sreg_type RDWIM then do wim_val \ gets (\s. (cpu_reg_val WIM s)); write_reg wim_val curr_win rd; return () od else \ \Must be \RDTBR\.\ do tbr_val \ gets (\s. (cpu_reg_val TBR s)); write_reg tbr_val curr_win rd; return () od else return () od" text \Operational semantics for write state register instructions. We do not consider WRASR here.\ definition write_state_reg_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "write_state_reg_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do curr_win \ get_curr_win(); psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. (get_S psr_val)); op2 \ gets (\s. (get_operand2 op_list s)); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); result \ gets (\s. ((XOR) rs1_val op2)); if instr_name = sreg_type WRY then do modify (\s. (delayed_pool_add (DELAYNUM, result, Y) s)); return () od else if instr_name = sreg_type WRASR then if privileged_ASR rd \ s_val = 0 then do raise_trap privileged_instruction; return () od else if illegal_instruction_ASR rd then do raise_trap illegal_instruction; return () od else do modify (\s. (delayed_pool_add (DELAYNUM, result, (ASR rd)) s)); return () od else if instr_name = sreg_type WRPSR then if s_val = 0 then do raise_trap privileged_instruction; return () od else if (uint ((ucast result)::word5)) \ NWINDOWS then do raise_trap illegal_instruction; return () od else do \ \\ET\ and \PIL\ appear to be written IMMEDIATELY w.r.t. interrupts.\ pil_val \ gets (\s. (get_PIL result)); et_val \ gets (\s. (get_ET result)); new_psr_val \ gets (\s. (update_PSR_et_pil et_val pil_val psr_val)); write_cpu new_psr_val PSR; modify (\s. (delayed_pool_add (DELAYNUM, result, PSR) s)); return () od else if instr_name = sreg_type WRWIM then if s_val = 0 then do raise_trap privileged_instruction; return () od else do \ \Don't write bits corresponding to non-existent windows.\ result_f \ gets (\s. ((result << nat (32 - NWINDOWS)) >> nat (32 - NWINDOWS))); modify (\s. (delayed_pool_add (DELAYNUM, result_f, WIM) s)); return () od else \ \Must be \WRTBR\\ if s_val = 0 then do raise_trap privileged_instruction; return () od else do \ \Only write the bits \<31:12>\ of the result to \TBR\.\ tbr_val \ gets (\s. (cpu_reg_val TBR s)); tbr_val_11_0 \ gets (\s. ((AND) tbr_val 0b00000000000000000000111111111111)); result_tmp \ gets (\s. ((AND) result 0b11111111111111111111000000000000)); result_f \ gets (\s. ((OR) tbr_val_11_0 result_tmp)); modify (\s. (delayed_pool_add (DELAYNUM, result_f, TBR) s)); return () od od" definition logical_result :: "sparc_operation \ word32 \ word32 \ word32" where "logical_result instr_name rs1_val operand2 \ if (instr_name = logic_type ANDs) \ (instr_name = logic_type ANDcc) then (AND) rs1_val operand2 else if (instr_name = logic_type ANDN) \ (instr_name = logic_type ANDNcc) then (AND) rs1_val (NOT operand2) else if (instr_name = logic_type ORs) \ (instr_name = logic_type ORcc) then (OR) rs1_val operand2 else if instr_name \ {logic_type ORN,logic_type ORNcc} then (OR) rs1_val (NOT operand2) else if instr_name \ {logic_type XORs,logic_type XORcc} then (XOR) rs1_val operand2 else \ \Must be \XNOR\ or \XNORcc\\ (XOR) rs1_val (NOT operand2) " definition logical_new_psr_val :: "word32 \ ('a) sparc_state \ word32" where "logical_new_psr_val result s \ let psr_val = cpu_reg_val PSR s; n_val = (ucast (result >> 31))::word1; z_val = if (result = 0) then 1 else 0; v_val = 0; c_val = 0 in update_PSR_icc n_val z_val v_val c_val psr_val " definition logical_instr_sub1 :: "sparc_operation \ word32 \ ('a::len,unit) sparc_state_monad" where "logical_instr_sub1 instr_name result \ if instr_name \ {logic_type ANDcc,logic_type ANDNcc,logic_type ORcc, logic_type ORNcc,logic_type XORcc,logic_type XNORcc} then do new_psr_val \ gets (\s. (logical_new_psr_val result s)); write_cpu new_psr_val PSR; return () od else return () " text \Operational semantics for logical instructions.\ definition logical_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "logical_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); rd_val \ gets (\s. (user_reg_val curr_win rd s)); result \ gets (\s. (logical_result instr_name rs1_val operand2)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; logical_instr_sub1 instr_name result od" text \Operational semantics for shift instructions.\ definition shift_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "shift_instr instr \ let instr_name = fst instr; op_list = snd instr; flagi = get_operand_flag (op_list!0); rs1 = get_operand_w5 (op_list!1); rs2_shcnt = get_operand_w5 (op_list!2); rd = get_operand_w5 (op_list!3) in do curr_win \ get_curr_win(); shift_count \ gets (\s. (if flagi = 0 then ucast (user_reg_val curr_win rs2_shcnt s) else rs2_shcnt)); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); if (instr_name = shift_type SLL) \ (rd \ 0) then do rd_val \ gets (\s. (rs1_val << (unat shift_count))); write_reg rd_val curr_win rd; return () od else if (instr_name = shift_type SRL) \ (rd \ 0) then do rd_val \ gets (\s. (rs1_val >> (unat shift_count))); write_reg rd_val curr_win rd; return () od else if (instr_name = shift_type SRA) \ (rd \ 0) then do rd_val \ gets (\s. (rs1_val >>> (unat shift_count))); write_reg rd_val curr_win rd; return () od else return () od" definition add_instr_sub1 :: "sparc_operation \ word32 \ word32 \ word32 \ ('a::len,unit) sparc_state_monad" where "add_instr_sub1 instr_name result rs1_val operand2 \ if instr_name \ {arith_type ADDcc,arith_type ADDXcc} then do psr_val \ gets (\s. (cpu_reg_val PSR s)); result_31 \ gets (\s. ((ucast (result >> 31))::word1)); rs1_val_31 \ gets (\s. ((ucast (rs1_val >> 31))::word1)); operand2_31 \ gets (\s. ((ucast (operand2 >> 31))::word1)); new_n_val \ gets (\s. (result_31)); new_z_val \ gets (\s. (if result = 0 then 1::word1 else 0::word1)); new_v_val \ gets (\s. ((OR) ((AND) rs1_val_31 ((AND) operand2_31 (NOT result_31))) ((AND) (NOT rs1_val_31) ((AND) (NOT operand2_31) result_31)))); new_c_val \ gets (\s. ((OR) ((AND) rs1_val_31 operand2_31) ((AND) (NOT result_31) ((OR) rs1_val_31 operand2_31)))); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; return () od else return () " text \Operational semantics for add instructions. These include ADD, ADDcc, ADDX.\ definition add_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "add_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); c_val \ gets (\s. (get_icc_C psr_val)); result \ gets (\s. (if (instr_name = arith_type ADD) \ (instr_name = arith_type ADDcc) then rs1_val + operand2 else \ \Must be \ADDX\ or \ADDXcc\\ rs1_val + operand2 + (ucast c_val))); rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; add_instr_sub1 instr_name result rs1_val operand2 od" definition sub_instr_sub1 :: "sparc_operation \ word32 \ word32 \ word32 \ ('a::len,unit) sparc_state_monad" where "sub_instr_sub1 instr_name result rs1_val operand2 \ if instr_name \ {arith_type SUBcc,arith_type SUBXcc} then do psr_val \ gets (\s. (cpu_reg_val PSR s)); result_31 \ gets (\s. ((ucast (result >> 31))::word1)); rs1_val_31 \ gets (\s. ((ucast (rs1_val >> 31))::word1)); operand2_31 \ gets (\s. ((ucast (operand2 >> 31))::word1)); new_n_val \ gets (\s. (result_31)); new_z_val \ gets (\s. (if result = 0 then 1::word1 else 0::word1)); new_v_val \ gets (\s. ((OR) ((AND) rs1_val_31 ((AND) (NOT operand2_31) (NOT result_31))) ((AND) (NOT rs1_val_31) ((AND) operand2_31 result_31)))); new_c_val \ gets (\s. ((OR) ((AND) (NOT rs1_val_31) operand2_31) ((AND) result_31 ((OR) (NOT rs1_val_31) operand2_31)))); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; return () od else return () " text \Operational semantics for subtract instructions. These include SUB, SUBcc, SUBX.\ definition sub_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "sub_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); c_val \ gets (\s. (get_icc_C psr_val)); result \ gets (\s. (if (instr_name = arith_type SUB) \ (instr_name = arith_type SUBcc) then rs1_val - operand2 else \ \Must be \SUBX\ or \SUBXcc\\ rs1_val - operand2 - (ucast c_val))); rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; sub_instr_sub1 instr_name result rs1_val operand2 od" definition mul_instr_sub1 :: "sparc_operation \ word32 \ ('a::len,unit) sparc_state_monad" where "mul_instr_sub1 instr_name result \ if instr_name \ {arith_type SMULcc,arith_type UMULcc} then do psr_val \ gets (\s. (cpu_reg_val PSR s)); new_n_val \ gets (\s. ((ucast (result >> 31))::word1)); new_z_val \ gets (\s. (if result = 0 then 1 else 0)); new_v_val \ gets (\s. 0); new_c_val \ gets (\s. 0); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; return () od else return () " text \Operational semantics for multiply instructions.\ definition mul_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "mul_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); result0 \ gets (\s. (if instr_name \ {arith_type UMUL,arith_type UMULcc} then (word_of_int ((uint rs1_val) * (uint operand2)))::word64 else \ \Must be \SMUL\ or \SMULcc\\ (word_of_int ((sint rs1_val) * (sint operand2)))::word64)); \ \whether to use \ucast\ or \scast\ does not matter below.\ y_val \ gets (\s. ((ucast (result0 >> 32))::word32)); write_cpu y_val Y; result \ gets (\s. ((ucast result0)::word32)); rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; mul_instr_sub1 instr_name result od" definition div_comp_temp_64bit :: "instruction \ word64 \ virtua_address \ word64" where "div_comp_temp_64bit i y_rs1 operand2 \ if ((fst i) = arith_type UDIV) \ ((fst i) = arith_type UDIVcc) then (word_of_int ((uint y_rs1) div (uint operand2)))::word64 else \ \Must be \SDIV\ or \SDIVcc\.\ \ \Due to Isabelle's rounding method is not nearest to zero,\ \ \we have to implement division in a different way.\ let sop1 = sint y_rs1; sop2 = sint operand2; pop1 = abs sop1; pop2 = abs sop2 in if sop1 > 0 \ sop2 > 0 then (word_of_int (sop1 div sop2)) else if sop1 > 0 \ sop2 < 0 then (word_of_int (- (sop1 div pop2))) else if sop1 < 0 \ sop2 > 0 then (word_of_int (- (pop1 div sop2))) else \ \\sop1 < 0 \ sop2 < 0\\ (word_of_int (pop1 div pop2))" definition div_comp_temp_V :: "instruction \ word32 \ word33 \ word1" where "div_comp_temp_V i w32 w33 \ if ((fst i) = arith_type UDIV) \ ((fst i) = arith_type UDIVcc) then if w32 = 0 then 0 else 1 else \ \Must be \SDIV\ or \SDIVcc\.\ if (w33 = 0) \ (w33 = (0b111111111111111111111111111111111::word33)) then 0 else 1" definition div_comp_result :: "instruction \ word1 \ word64 \ word32" where "div_comp_result i temp_V temp_64bit \ if temp_V = 1 then if ((fst i) = arith_type UDIV) \ ((fst i) = arith_type UDIVcc) then (0b11111111111111111111111111111111::word32) else if (fst i) \ {arith_type SDIV,arith_type SDIVcc} then if temp_64bit > 0 then (0b01111111111111111111111111111111::word32) else ((word_of_int (0 - (uint (0b10000000000000000000000000000000::word32))))::word32) else ((ucast temp_64bit)::word32) else ((ucast temp_64bit)::word32)" definition div_write_new_val :: "instruction \ word32 \ word1 \ ('a::len,unit) sparc_state_monad" where "div_write_new_val i result temp_V \ if (fst i) \ {arith_type UDIVcc,arith_type SDIVcc} then do psr_val \ gets (\s. (cpu_reg_val PSR s)); new_n_val \ gets (\s. ((ucast (result >> 31))::word1)); new_z_val \ gets (\s. (if result = 0 then 1 else 0)); new_v_val \ gets (\s. temp_V); new_c_val \ gets (\s. 0); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; return () od else return ()" definition div_comp :: "instruction \ word5 \ word5 \ virtua_address \ ('a::len,unit) sparc_state_monad" where "div_comp instr rs1 rd operand2 \ do curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); y_val \ gets (\s. (cpu_reg_val Y s)); y_rs1 \ gets (\s. ((word_cat y_val rs1_val)::word64)); temp_64bit \ gets (\s. (div_comp_temp_64bit instr y_rs1 operand2)); \<^cancel>\result \ gets (\s. ((ucast temp_64bit)::word32));\ temp_high32 \ gets (\s. ((ucast (temp_64bit >> 32))::word32)); temp_high33 \ gets (\s. ((ucast (temp_64bit >> 31))::word33)); temp_V \ gets (\s. (div_comp_temp_V instr temp_high32 temp_high33)); result \ gets (\s. (div_comp_result instr temp_V temp_64bit)); rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; div_write_new_val instr result temp_V od" text \Operational semantics for divide instructions.\ definition div_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "div_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); if (uint operand2) = 0 then do raise_trap division_by_zero; return () od else div_comp instr rs1 rd operand2 od" definition ld_word0 :: "instruction \ word32 \ virtua_address \ word32" where "ld_word0 instr data_word address \ if (fst instr) \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDSBA} then let byte = if (uint ((ucast address)::word2)) = 0 then (ucast (data_word >> 24))::word8 else if (uint ((ucast address)::word2)) = 1 then (ucast (data_word >> 16))::word8 else if (uint ((ucast address)::word2)) = 2 then (ucast (data_word >> 8))::word8 else \ \Must be 3.\ (ucast data_word)::word8 in if (fst instr) = load_store_type LDSB \ (fst instr) = load_store_type LDSBA then sign_ext8 byte else zero_ext8 byte else if (fst instr) = load_store_type LDUH \ (fst instr) = load_store_type LDSH \ (fst instr) = load_store_type LDSHA \ (fst instr) = load_store_type LDUHA then let halfword = if (uint ((ucast address)::word2)) = 0 then (ucast (data_word >> 16))::word16 else \ \Must be 2.\ (ucast data_word)::word16 in if (fst instr) = load_store_type LDSH \ (fst instr) = load_store_type LDSHA then sign_ext16 halfword else zero_ext16 halfword else \ \Must be LDD\ data_word " definition ld_asi :: "instruction \ word1 \ asi_type" where "ld_asi instr s_val \ if (fst instr) \ {load_store_type LDD,load_store_type LD,load_store_type LDUH, load_store_type LDSB,load_store_type LDUB,load_store_type LDSH} then if s_val = 0 then (word_of_int 10)::asi_type else (word_of_int 11)::asi_type else \ \Must be \LDA\, \LDUBA\, \LDSBA\, \LDSHA\, \LDUHA\, or \LDDA\.\ get_operand_asi ((snd instr)!3) " definition load_sub2 :: "virtua_address \ asi_type \ word5 \ ('a::len) window_size \ word32 \ ('a,unit) sparc_state_monad" where "load_sub2 address asi rd curr_win word0 \ do write_reg word0 curr_win ((AND) rd 0b11110); (result1,new_state1) \ gets (\s. (memory_read asi (address + 4) s)); if result1 = None then do raise_trap data_access_exception; return () od else do word1 \ gets (\s. (case result1 of Some v \ v)); modify (\s. (new_state1)); write_reg word1 curr_win ((OR) rd 1); return () od od" definition load_sub3 :: "instruction \ ('a::len) window_size \ word5 \ asi_type \ virtua_address \ ('a::len,unit) sparc_state_monad" where "load_sub3 instr curr_win rd asi address \ do (result,new_state) \ gets (\s. (memory_read asi address s)); if result = None then do raise_trap data_access_exception; return () od else do data_word \ gets (\s. (case result of Some v \ v)); modify (\s. (new_state)); word0 \ gets (\s. (ld_word0 instr data_word address)); if rd \ 0 \ (fst instr) \ {load_store_type LD,load_store_type LDA, load_store_type LDUH,load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDSH,load_store_type LDSHA, load_store_type LDUHA,load_store_type LDSBA} then do write_reg word0 curr_win rd; return () od else \ \Must be \LDD\ or \LDDA\\ load_sub2 address asi rd curr_win word0 od od" definition load_sub1 :: "instruction \ word5 \ word1 \ ('a::len,unit) sparc_state_monad" where "load_sub1 instr rd s_val \ do curr_win \ get_curr_win(); address \ gets (\s. (get_addr (snd instr) s)); asi \ gets (\s. (ld_asi instr s_val)); if (((fst instr) = load_store_type LDD \ (fst instr) = load_store_type LDDA) \ ((ucast address)::word3) \ 0) \ ((fst instr) \ {load_store_type LD,load_store_type LDA} \ ((ucast address)::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 address)::word1) \ 0) then do raise_trap mem_address_not_aligned; return () od else load_sub3 instr curr_win rd asi address od" text \Operational semantics for Load instructions.\ definition load_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "load_instr instr \ let instr_name = fst instr; op_list = snd instr; flagi = get_operand_flag (op_list!0); rd = if instr_name \ {load_store_type LDUBA,load_store_type LDA, load_store_type LDSBA,load_store_type LDSHA, load_store_type LDSHA,load_store_type LDDA} then \ \\rd\ is member 4\ get_operand_w5 (op_list!4) else \ \\rd\ is member 3\ get_operand_w5 (op_list!3) in do psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. (get_S psr_val)); if instr_name \ {load_store_type LDA,load_store_type LDUBA, load_store_type LDSBA,load_store_type LDSHA, load_store_type LDUHA,load_store_type LDDA} \ s_val = 0 then do raise_trap privileged_instruction; return () od else if instr_name \ {load_store_type LDA,load_store_type LDUBA, load_store_type LDSBA,load_store_type LDSHA,load_store_type LDUHA, load_store_type LDDA} \ flagi = 1 then do raise_trap illegal_instruction; return () od else load_sub1 instr rd s_val od" definition st_asi :: "instruction \ word1 \ asi_type" where "st_asi instr s_val \ if (fst instr) \ {load_store_type STD,load_store_type ST, load_store_type STH,load_store_type STB} then if s_val = 0 then (word_of_int 10)::asi_type else (word_of_int 11)::asi_type else \ \Must be \STA\, \STBA\, \STHA\, \STDA\.\ get_operand_asi ((snd instr)!3) " definition st_byte_mask :: "instruction \ virtua_address \ word4" where "st_byte_mask instr address \ if (fst instr) \ {load_store_type STD,load_store_type ST, load_store_type STA,load_store_type STDA} then (0b1111::word4) else if (fst instr) \ {load_store_type STH,load_store_type STHA} then if ((ucast address)::word2) = 0 then (0b1100::word4) else \ \Must be 2.\ (0b0011::word4) else \ \Must be \STB\ or \STBA\.\ if ((ucast address)::word2) = 0 then (0b1000::word4) else if ((ucast address)::word2) = 1 then (0b0100::word4) else if ((ucast address)::word2) = 2 then (0b0010::word4) else \ \Must be 3.\ (0b0001::word4) " definition st_data0 :: "instruction \ ('a::len) window_size \ word5 \ virtua_address \ ('a) sparc_state \ reg_type" where "st_data0 instr curr_win rd address s \ if (fst instr) \ {load_store_type STD,load_store_type STDA} then user_reg_val curr_win ((AND) rd 0b11110) s else if (fst instr) \ {load_store_type ST,load_store_type STA} then user_reg_val curr_win rd s else if (fst instr) \ {load_store_type STH,load_store_type STHA} then if ((ucast address)::word2) = 0 then (user_reg_val curr_win rd s) << 16 else \ \Must be 2.\ user_reg_val curr_win rd s else \ \Must be \STB\ or \STBA\.\ if ((ucast address)::word2) = 0 then (user_reg_val curr_win rd s) << 24 else if ((ucast address)::word2) = 1 then (user_reg_val curr_win rd s) << 16 else if ((ucast address)::word2) = 2 then (user_reg_val curr_win rd s) << 8 else \ \Must be 3.\ user_reg_val curr_win rd s " definition store_sub2 :: "instruction \ ('a::len) window_size \ word5 \ asi_type \ virtua_address \ ('a::len,unit) sparc_state_monad" where "store_sub2 instr curr_win rd asi address \ do byte_mask \ gets (\s. (st_byte_mask instr address)); data0 \ gets (\s. (st_data0 instr curr_win rd address s)); result0 \ gets (\s. (memory_write asi address byte_mask data0 s)); if result0 = None then do raise_trap data_access_exception; return () od else do new_state \ gets (\s. (case result0 of Some v \ v)); modify (\s. (new_state)); if (fst instr) \ {load_store_type STD,load_store_type STDA} then do data1 \ gets (\s. (user_reg_val curr_win ((OR) rd 0b00001) s)); result1 \ gets (\s. (memory_write asi (address + 4) (0b1111::word4) data1 s)); if result1 = None then do raise_trap data_access_exception; return () od else do new_state1 \ gets (\s. (case result1 of Some v \ v)); modify (\s. (new_state1)); return () od od else return () od od" definition store_sub1 :: "instruction \ word5 \ word1 \ ('a::len,unit) sparc_state_monad" where "store_sub1 instr rd s_val \ do curr_win \ get_curr_win(); address \ gets (\s. (get_addr (snd instr) s)); asi \ gets (\s. (st_asi instr s_val)); \ \The following code is intentionally long to match the definitions in SPARCv8.\ if ((fst instr) = load_store_type STH \ (fst instr) = load_store_type STHA) \ ((ucast address)::word1) \ 0 then do raise_trap mem_address_not_aligned; return () od else if (fst instr) \ {load_store_type ST,load_store_type STA} \ ((ucast address)::word2) \ 0 then do raise_trap mem_address_not_aligned; return () od else if (fst instr) \ {load_store_type STD,load_store_type STDA} \ ((ucast address)::word3) \ 0 then do raise_trap mem_address_not_aligned; return () od else store_sub2 instr curr_win rd asi address od" text \Operational semantics for Store instructions.\ definition store_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "store_instr instr \ let instr_name = fst instr; op_list = snd instr; flagi = get_operand_flag (op_list!0); rd = if instr_name \ {load_store_type STA,load_store_type STBA, load_store_type STHA,load_store_type STDA} then \ \\rd\ is member 4\ get_operand_w5 (op_list!4) else \ \\rd\ is member 3\ get_operand_w5 (op_list!3) in do psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. (get_S psr_val)); if instr_name \ {load_store_type STA,load_store_type STDA, load_store_type STHA,load_store_type STBA} \ s_val = 0 then do raise_trap privileged_instruction; return () od else if instr_name \ {load_store_type STA,load_store_type STDA, load_store_type STHA,load_store_type STBA} \ flagi = 1 then do raise_trap illegal_instruction; return () od else store_sub1 instr rd s_val od" text \The instructions below are not used by Xtratum and they are not tested.\ definition ldst_asi :: "instruction \ word1 \ asi_type" where "ldst_asi instr s_val \ if (fst instr) \ {load_store_type LDSTUB} then if s_val = 0 then (word_of_int 10)::asi_type else (word_of_int 11)::asi_type else \ \Must be \LDSTUBA\.\ get_operand_asi ((snd instr)!3) " definition ldst_word0 :: "instruction \ word32 \ virtua_address \ word32" where "ldst_word0 instr data_word address \ let byte = if (uint ((ucast address)::word2)) = 0 then (ucast (data_word >> 24))::word8 else if (uint ((ucast address)::word2)) = 1 then (ucast (data_word >> 16))::word8 else if (uint ((ucast address)::word2)) = 2 then (ucast (data_word >> 8))::word8 else \ \Must be 3.\ (ucast data_word)::word8 in zero_ext8 byte " definition ldst_byte_mask :: "instruction \ virtua_address \ word4" where "ldst_byte_mask instr address \ if ((ucast address)::word2) = 0 then (0b1000::word4) else if ((ucast address)::word2) = 1 then (0b0100::word4) else if ((ucast address)::word2) = 2 then (0b0010::word4) else \ \Must be 3.\ (0b0001::word4) " definition load_store_sub1 :: "instruction \ word5 \ word1 \ ('a::len,unit) sparc_state_monad" where "load_store_sub1 instr rd s_val \ do curr_win \ get_curr_win(); address \ gets (\s. (get_addr (snd instr) s)); asi \ gets (\s. (ldst_asi instr s_val)); \ \wait for locks to be lifted.\ \ \an implementation actually need only block when another \LDSTUB\ or \SWAP\\ \ \is pending on the same byte in memory as the one addressed by this \LDSTUB\\ \ \Should wait when \block_type = 1 \ block_word = 1\\ \ \until another processes write both to be 0.\ \ \We implement this as setting \pc\ as \npc\ when the instruction\ \ \is blocked. This way, in the next iteration, we will still execution\ \ \the current instruction.\ block_byte \ gets (\s. (pb_block_ldst_byte_val address s)); block_word \ gets (\s. (pb_block_ldst_word_val address s)); if block_byte \ block_word then do pc_val \ gets (\s. (cpu_reg_val PC s)); write_cpu pc_val nPC; return () od else do modify (\s. (pb_block_ldst_byte_mod address True s)); (result,new_state) \ gets (\s. (memory_read asi address s)); if result = None then do raise_trap data_access_exception; return () od else do data_word \ gets (\s. (case result of Some v \ v)); modify (\s. (new_state)); byte_mask \ gets (\s. (ldst_byte_mask instr address)); data0 \ gets (\s. (0b11111111111111111111111111111111::word32)); result0 \ gets (\s. (memory_write asi address byte_mask data0 s)); modify (\s. (pb_block_ldst_byte_mod address False s)); if result0 = None then do raise_trap data_access_exception; return () od else do new_state1 \ gets (\s. (case result0 of Some v \ v)); modify (\s. (new_state1)); word0 \ gets (\s. (ldst_word0 instr data_word address)); if rd \ 0 then do write_reg word0 curr_win rd; return () od else return () od od od od" text \Operational semantics for atomic load-store.\ definition load_store_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "load_store_instr instr \ let instr_name = fst instr; op_list = snd instr; flagi = get_operand_flag (op_list!0); rd = if instr_name \ {load_store_type LDSTUBA} then \ \\rd\ is member 4\ get_operand_w5 (op_list!4) else \ \\rd\ is member 3\ get_operand_w5 (op_list!3) in do psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. (get_S psr_val)); if instr_name \ {load_store_type LDSTUBA} \ s_val = 0 then do raise_trap privileged_instruction; return () od else if instr_name \ {load_store_type LDSTUBA} \ flagi = 1 then do raise_trap illegal_instruction; return () od else load_store_sub1 instr rd s_val od" definition swap_sub1 :: "instruction \ word5 \ word1 \ ('a::len,unit) sparc_state_monad" where "swap_sub1 instr rd s_val \ do curr_win \ get_curr_win(); address \ gets (\s. (get_addr (snd instr) s)); asi \ gets (\s. (ldst_asi instr s_val)); temp \ gets (\s. (user_reg_val curr_win rd s)); \ \wait for locks to be lifted.\ \ \an implementation actually need only block when another \LDSTUB\ or \SWAP\\ \ \is pending on the same byte in memory as the one addressed by this \LDSTUB\\ \ \Should wait when \block_type = 1 \ block_word = 1\\ \ \until another processes write both to be 0.\ \ \We implement this as setting \pc\ as \npc\ when the instruction\ \ \is blocked. This way, in the next iteration, we will still execution\ \ \the current instruction.\ block_byte \ gets (\s. (pb_block_ldst_byte_val address s)); block_word \ gets (\s. (pb_block_ldst_word_val address s)); if block_byte \ block_word then do pc_val \ gets (\s. (cpu_reg_val PC s)); write_cpu pc_val nPC; return () od else do modify (\s. (pb_block_ldst_word_mod address True s)); (result,new_state) \ gets (\s. (memory_read asi address s)); if result = None then do raise_trap data_access_exception; return () od else do word \ gets (\s. (case result of Some v \ v)); modify (\s. (new_state)); byte_mask \ gets (\s. (0b1111::word4)); result0 \ gets (\s. (memory_write asi address byte_mask temp s)); modify (\s. (pb_block_ldst_word_mod address False s)); if result0 = None then do raise_trap data_access_exception; return () od else do new_state1 \ gets (\s. (case result0 of Some v \ v)); modify (\s. (new_state1)); if rd \ 0 then do write_reg word curr_win rd; return () od else return () od od od od" text \Operational semantics for swap.\ definition swap_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "swap_instr instr \ let instr_name = fst instr; op_list = snd instr; flagi = get_operand_flag (op_list!0); rd = if instr_name \ {load_store_type SWAPA} then \ \\rd\ is member 4\ get_operand_w5 (op_list!4) else \ \\rd\ is member 3\ get_operand_w5 (op_list!3) in do psr_val \ gets (\s. (cpu_reg_val PSR s)); s_val \ gets (\s. (get_S psr_val)); if instr_name \ {load_store_type SWAPA} \ s_val = 0 then do raise_trap privileged_instruction; return () od else if instr_name \ {load_store_type SWAPA} \ flagi = 1 then do raise_trap illegal_instruction; return () od else swap_sub1 instr rd s_val od" definition bit2_zero :: "word2 \ word1" where "bit2_zero w2 \ if w2 \ 0 then 1 else 0" text \Operational semantics for tagged add instructions.\ definition tadd_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "tadd_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); c_val \ gets (\s. (get_icc_C psr_val)); result \ gets (\s. (rs1_val + operand2)); result_31 \ gets (\s. ((ucast (result >> 31))::word1)); rs1_val_31 \ gets (\s. ((ucast (rs1_val >> 31))::word1)); operand2_31 \ gets (\s. ((ucast (operand2 >> 31))::word1)); rs1_val_2 \ gets (\s. (bit2_zero ((ucast rs1_val)::word2))); operand2_2 \ gets (\s. (bit2_zero ((ucast operand2)::word2))); temp_V \ gets (\s. ((OR) ((OR) ((AND) rs1_val_31 ((AND) operand2_31 (NOT result_31))) ((AND) (NOT rs1_val_31) ((AND) (NOT operand2_31) result_31))) ((OR) rs1_val_2 operand2_2))); if instr_name = arith_type TADDccTV \ temp_V = 1 then do raise_trap tag_overflow; return () od else do rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; new_n_val \ gets (\s. (result_31)); new_z_val \ gets (\s. (if result = 0 then 1::word1 else 0::word1)); new_v_val \ gets (\s. temp_V); new_c_val \ gets (\s. ((OR) ((AND) rs1_val_31 operand2_31) ((AND) (NOT result_31) ((OR) rs1_val_31 operand2_31)))); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; return () od od" text \Operational semantics for tagged add instructions.\ definition tsub_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "tsub_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do operand2 \ gets (\s. (get_operand2 op_list s)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); c_val \ gets (\s. (get_icc_C psr_val)); result \ gets (\s. (rs1_val - operand2)); result_31 \ gets (\s. ((ucast (result >> 31))::word1)); rs1_val_31 \ gets (\s. ((ucast (rs1_val >> 31))::word1)); operand2_31 \ gets (\s. ((ucast (operand2 >> 31))::word1)); rs1_val_2 \ gets (\s. (bit2_zero ((ucast rs1_val)::word2))); operand2_2 \ gets (\s. (bit2_zero ((ucast operand2)::word2))); temp_V \ gets (\s. ((OR) ((OR) ((AND) rs1_val_31 ((AND) operand2_31 (NOT result_31))) ((AND) (NOT rs1_val_31) ((AND) (NOT operand2_31) result_31))) ((OR) rs1_val_2 operand2_2))); if instr_name = arith_type TSUBccTV \ temp_V = 1 then do raise_trap tag_overflow; return () od else do rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; new_n_val \ gets (\s. (result_31)); new_z_val \ gets (\s. (if result = 0 then 1::word1 else 0::word1)); new_v_val \ gets (\s. temp_V); new_c_val \ gets (\s. ((OR) ((AND) rs1_val_31 operand2_31) ((AND) (NOT result_31) ((OR) rs1_val_31 operand2_31)))); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; return () od od" definition muls_op2 :: "inst_operand list \ ('a::len) sparc_state \ word32" where "muls_op2 op_list s \ let y_val = cpu_reg_val Y s in if ((ucast y_val)::word1) = 0 then 0 else get_operand2 op_list s " text \Operational semantics for multiply step instruction.\ definition muls_instr :: "instruction \ ('a::len,unit) sparc_state_monad" where "muls_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1); rd = get_operand_w5 (op_list!3) in do curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); psr_val \ gets (\s. (cpu_reg_val PSR s)); n_val \ gets (\s. (get_icc_N psr_val)); v_val \ gets (\s. (get_icc_V psr_val)); c_val \ gets (\s. (get_icc_C psr_val)); y_val \ gets (\s. (cpu_reg_val Y s)); operand1 \ gets (\s. (word_cat ((XOR) n_val v_val) ((ucast (rs1_val >> 1))::word31))); operand2 \ gets (\s. (muls_op2 op_list s)); result \ gets (\s. (operand1 + operand2)); new_y_val \ gets (\s. (word_cat ((ucast rs1_val)::word1) ((ucast (y_val >> 1))::word31))); write_cpu new_y_val Y; rd_val \ gets (\s. (user_reg_val curr_win rd s)); new_rd_val \ gets (\s. (if rd \ 0 then result else rd_val)); write_reg new_rd_val curr_win rd; result_31 \ gets (\s. ((ucast (result >> 31))::word1)); operand1_31 \ gets (\s. ((ucast (operand1 >> 31))::word1)); operand2_31 \ gets (\s. ((ucast (operand2 >> 31))::word1)); new_n_val \ gets (\s. (result_31)); new_z_val \ gets (\s. (if result = 0 then 1::word1 else 0::word1)); new_v_val \ gets (\s. ((OR) ((AND) operand1_31 ((AND) operand2_31 (NOT result_31))) ((AND) (NOT operand1_31) ((AND) (NOT operand2_31) result_31)))); new_c_val \ gets (\s. ((OR) ((AND) operand1_31 operand2_31) ((AND) (NOT result_31) ((OR) operand1_31 operand2_31)))); new_psr_val \ gets (\s. (update_PSR_icc new_n_val new_z_val new_v_val new_c_val psr_val)); write_cpu new_psr_val PSR; return () od" text\Evaluate icc based on the bits N, Z, V, C in PSR and the type of ticc instruction. See Sparcv8 manual Page 182.\ definition trap_eval_icc::"sparc_operation \ word1 \ word1 \ word1 \ word1 \ int" where "trap_eval_icc instr_name n_val z_val v_val c_val \ if instr_name = ticc_type TNE then if z_val = 0 then 1 else 0 else if instr_name = ticc_type TE then if z_val = 1 then 1 else 0 else if instr_name = ticc_type TG then if ((OR) z_val (n_val XOR v_val)) = 0 then 1 else 0 else if instr_name = ticc_type TLE then if ((OR) z_val (n_val XOR v_val)) = 1 then 1 else 0 else if instr_name = ticc_type TGE then if (n_val XOR v_val) = 0 then 1 else 0 else if instr_name = ticc_type TL then if (n_val XOR v_val) = 1 then 1 else 0 else if instr_name = ticc_type TGU then if (c_val = 0 \ z_val = 0) then 1 else 0 else if instr_name = ticc_type TLEU then if (c_val = 1 \ z_val = 1) then 1 else 0 else if instr_name = ticc_type TCC then if c_val = 0 then 1 else 0 else if instr_name = ticc_type TCS then if c_val = 1 then 1 else 0 else if instr_name = ticc_type TPOS then if n_val = 0 then 1 else 0 else if instr_name = ticc_type TNEG then if n_val = 1 then 1 else 0 else if instr_name = ticc_type TVC then if v_val = 0 then 1 else 0 else if instr_name = ticc_type TVS then if v_val = 1 then 1 else 0 else if instr_name = ticc_type TA then 1 else if instr_name = ticc_type TN then 0 else -1 " text \ Get \operand2\ for \ticc\ based on the flag \i\, \rs1\, \rs2\, and \trap_imm7\. If \i = 0\ then \operand2 = r[rs2]\, else \operand2 = sign_ext7(trap_imm7)\. \op_list\ should be \[i,rs1,rs2]\ or \[i,rs1,trap_imm7]\. \ definition get_trap_op2::"inst_operand list \ ('a::len) sparc_state \ virtua_address" where "get_trap_op2 op_list s \ let flagi = get_operand_flag (op_list!0); curr_win = ucast (get_CWP (cpu_reg_val PSR s)) in if flagi = 0 then let rs2 = get_operand_w5 (op_list!2); rs2_val = user_reg_val curr_win rs2 s in rs2_val else let ext_simm7 = sign_ext7 (get_operand_imm7 (op_list!2)) in ext_simm7 " text \Operational semantics for Ticc insturctions.\ definition ticc_instr::"instruction \ ('a::len,unit) sparc_state_monad" where "ticc_instr instr \ let instr_name = fst instr; op_list = snd instr; rs1 = get_operand_w5 (op_list!1) in do n_val \ gets (\s. get_icc_N ((cpu_reg s) PSR)); z_val \ gets (\s. get_icc_Z ((cpu_reg s) PSR)); v_val \ gets (\s. get_icc_V ((cpu_reg s) PSR)); c_val \ gets (\s. get_icc_C ((cpu_reg s) PSR)); icc_val \ gets(\s. (trap_eval_icc instr_name n_val z_val v_val c_val)); curr_win \ get_curr_win(); rs1_val \ gets (\s. (user_reg_val curr_win rs1 s)); trap_number \ gets (\s. (rs1_val + (get_trap_op2 op_list s))); npc_val \ gets (\s. (cpu_reg_val nPC s)); pc_val \ gets (\s. (cpu_reg_val PC s)); if icc_val = 1 then do raise_trap trap_instruction; trap_number7 \ gets (\s. ((ucast trap_number)::word7)); modify (\s. (ticc_trap_type_mod trap_number7 s)); return () od else \ \\icc_val = 0\\ do write_cpu npc_val PC; write_cpu (npc_val + 4) nPC; return () od od" text \Operational semantics for store barrier.\ definition store_barrier_instr::"instruction \ ('a::len,unit) sparc_state_monad" where "store_barrier_instr instr \ do modify (\s. (store_barrier_pending_mod True s)); return () od" end + +end 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,8582 +1,8588 @@ (* * 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: 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: 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 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/SPARCv8/SparcModel_MMU/Sparc_State.thy b/thys/SPARCv8/SparcModel_MMU/Sparc_State.thy --- a/thys/SPARCv8/SparcModel_MMU/Sparc_State.thy +++ b/thys/SPARCv8/SparcModel_MMU/Sparc_State.thy @@ -1,1007 +1,1013 @@ (* * 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, David Sanan. *) section \SPARC V8 state model\ theory Sparc_State imports Main Sparc_Types "../lib/wp/DetMonadLemmas" MMU begin section \state as a function\ record cpu_cache = dcache:: cache_context icache:: cache_context text\ The state @{term sparc_state} is defined as a tuple @{term cpu_context}, @{term user_context}, @{term mem_context}, defining the state of the CPU registers, user registers, memory, cache, and delayed write pool respectively. Additionally, a boolean indicates whether the state is undefined or not. \ record (overloaded) ('a) sparc_state = cpu_reg:: cpu_context user_reg:: "('a) user_context" sys_reg:: sys_context mem:: mem_context mmu:: MMU_state cache:: cpu_cache dwrite:: delayed_write_pool state_var:: sparc_state_var traps:: "Trap set" undef:: bool section\functions for state member access\ definition cpu_reg_val:: "CPU_register \ ('a) sparc_state \ reg_type" where "cpu_reg_val reg state \ (cpu_reg state) reg" definition cpu_reg_mod :: "word32 \ CPU_register \ ('a) sparc_state \ ('a) sparc_state" where "cpu_reg_mod data_w32 cpu state \ state\cpu_reg := ((cpu_reg state)(cpu := data_w32))\" text \r[0] = 0. Otherwise read the actual value.\ definition user_reg_val:: "('a) window_size \ user_reg_type \ ('a) sparc_state \ reg_type" where "user_reg_val window ur state \ if ur = 0 then 0 else (user_reg state) window ur" text \Write a global register. win should be initialised as NWINDOWS.\ fun (sequential) global_reg_mod :: "word32 \ nat \ user_reg_type \ ('a::len) sparc_state \ ('a) sparc_state" where "global_reg_mod data_w32 0 ur state = state" | "global_reg_mod data_w32 win ur state = ( let win_word = word_of_int (int (win-1)); ns = state\user_reg := (user_reg state)(win_word := ((user_reg state) win_word)(ur := data_w32))\ in global_reg_mod data_w32 (win-1) ur ns )" text \Compute the next window.\ definition next_window :: "('a::len) window_size \ ('a) window_size" where "next_window win \ if (uint win) < (NWINDOWS - 1) then (win + 1) else 0 " text \Compute the previous window.\ definition pre_window :: "('a::len) window_size \ ('a::len) window_size" where "pre_window win \ if (uint win) > 0 then (win - 1) else (word_of_int (NWINDOWS - 1)) " text \write an output register. Also write ur+16 of the previous window.\ definition out_reg_mod :: "word32 \ ('a::len) window_size \ user_reg_type \ ('a) sparc_state \ ('a) sparc_state" where "out_reg_mod data_w32 win ur state \ let state' = state\user_reg := (user_reg state)(win := ((user_reg state) win)(ur := data_w32))\; win' = pre_window win; ur' = ur + 16 in state'\user_reg := (user_reg state')(win' := ((user_reg state') win')(ur' := data_w32))\ " text \Write a input register. Also write ur-16 of the next window.\ definition in_reg_mod :: "word32 \ ('a::len) window_size \ user_reg_type \ ('a) sparc_state \ ('a) sparc_state" where "in_reg_mod data_w32 win ur state \ let state' = state\user_reg := (user_reg state)(win := ((user_reg state) win)(ur := data_w32))\; win' = next_window win; ur' = ur - 16 in state'\user_reg := (user_reg state')(win' := ((user_reg state') win')(ur' := data_w32))\ " text \Do not modify r[0].\ definition user_reg_mod :: "word32 \ ('a::len) window_size \ user_reg_type \ ('a) sparc_state \ ('a) sparc_state" where "user_reg_mod data_w32 win ur state \ if ur = 0 then state else if 0 < ur \ ur < 8 then global_reg_mod data_w32 (nat NWINDOWS) ur state else if 7 < ur \ ur < 16 then out_reg_mod data_w32 win ur state else if 15 < ur \ ur < 24 then state\user_reg := (user_reg state)(win := ((user_reg state) win)(ur := data_w32))\ else \<^cancel>\if 23 < ur \ ur < 32 then\ in_reg_mod data_w32 win ur state \<^cancel>\else state\ " definition sys_reg_val :: "sys_reg \ ('a) sparc_state \ reg_type" where "sys_reg_val reg state \ (sys_reg state) reg" definition sys_reg_mod :: "word32 \ sys_reg \ ('a) sparc_state \ ('a) sparc_state" where "sys_reg_mod data_w32 sys state \ state\sys_reg := (sys_reg state)(sys := data_w32)\" text \The following fucntions deal with physical memory. N.B. Physical memory address in SPARCv8 is 36-bit.\ text \LEON3 doesn't distinguish ASI 8 and 9; 10 and 11 for read access for both user and supervisor. We recently discovered that the compiled machine code by the sparc-elf compiler often reads asi = 10 (user data) when the actual content is store in asi = 8 (user instruction). For testing purposes, we don't distinguish asi = 8,9,10,11 for reading access.\ definition mem_val:: "asi_type \ phys_address \ ('a) sparc_state \ mem_val_type option" where "mem_val asi add state \ let asi8 = word_of_int 8; asi9 = word_of_int 9; asi10 = word_of_int 10; asi11 = word_of_int 11; r1 = (mem state) asi8 add in if r1 = None then let r2 = (mem state) asi9 add in if r2 = None then let r3 = (mem state) asi10 add in if r3 = None then (mem state) asi11 add else r3 else r2 else r1 " text \An alternative way to read values from memory. Some implementations may use this definition.\ definition mem_val_alt:: "asi_type \ phys_address \ ('a) sparc_state \ mem_val_type option" where "mem_val_alt asi add state \ let r1 = (mem state) asi add; asi8 = word_of_int 8; asi9 = word_of_int 9; asi10 = word_of_int 10; asi11 = word_of_int 11 in if r1 = None \ (uint asi) = 8 then let r2 = (mem state) asi9 add in r2 else if r1 = None \ (uint asi) = 9 then let r2 = (mem state) asi8 add in r2 else if r1 = None \ (uint asi) = 10 then let r2 = (mem state) asi11 add in if r2 = None then let r3 = (mem state) asi8 add in if r3 = None then (mem state) asi9 add else r3 else r2 else if r1 = None \ (uint asi) = 11 then let r2 = (mem state) asi10 add in if r2 = None then let r3 = (mem state) asi8 add in if r3 = None then (mem state) asi9 add else r3 else r2 else r1" definition mem_mod :: "asi_type \ phys_address \ mem_val_type \ ('a) sparc_state \ ('a) sparc_state" where "mem_mod asi addr val state \ let state1 = state\mem := (mem state) (asi := ((mem state) asi)(addr := Some val))\ in \ \Only allow one of \asi\ 8 and 9 (10 and 11) to have value.\ if (uint asi) = 8 \ (uint asi) = 10 then let asi2 = word_of_int ((uint asi) + 1) in state1\mem := (mem state1) (asi2 := ((mem state1) asi2)(addr := None))\ else if (uint asi) = 9 \ (uint asi) = 11 then let asi2 = word_of_int ((uint asi) - 1) in state1\mem := (mem state1)(asi2 := ((mem state1) asi2)(addr := None))\ else state1 " text \An alternative way to write memory. This method insists that for each address, it can only hold a value in one of ASI = 8,9,10,11.\ definition mem_mod_alt :: "asi_type \ phys_address \ mem_val_type \ ('a) sparc_state \ ('a) sparc_state" where "mem_mod_alt asi addr val state \ let state1 = state\mem := (mem state) (asi := ((mem state) asi)(addr := Some val))\; asi8 = word_of_int 8; asi9 = word_of_int 9; asi10 = word_of_int 10; asi11 = word_of_int 11 in \ \Only allow one of \asi\ 8, 9, 10, 11 to have value.\ if (uint asi) = 8 then let state2 = state1\mem := (mem state1) (asi9 := ((mem state1) asi9)(addr := None))\; state3 = state2\mem := (mem state2) (asi10 := ((mem state2) asi10)(addr := None))\; state4 = state3\mem := (mem state3) (asi11 := ((mem state3) asi11)(addr := None))\ in state4 else if (uint asi) = 9 then let state2 = state1\mem := (mem state1) (asi8 := ((mem state1) asi8)(addr := None))\; state3 = state2\mem := (mem state2) (asi10 := ((mem state2) asi10)(addr := None))\; state4 = state3\mem := (mem state3) (asi11 := ((mem state3) asi11)(addr := None))\ in state4 else if (uint asi) = 10 then let state2 = state1\mem := (mem state1) (asi9 := ((mem state1) asi9)(addr := None))\; state3 = state2\mem := (mem state2) (asi8 := ((mem state2) asi8)(addr := None))\; state4 = state3\mem := (mem state3) (asi11 := ((mem state3) asi11)(addr := None))\ in state4 else if (uint asi) = 11 then let state2 = state1\mem := (mem state1) (asi9 := ((mem state1) asi9)(addr := None))\; state3 = state2\mem := (mem state2) (asi10 := ((mem state2) asi10)(addr := None))\; state4 = state3\mem := (mem state3) (asi8 := ((mem state3) asi8)(addr := None))\ in state4 else state1 " +context + includes bit_operations_syntax +begin + text \Given an ASI (word8), an address (word32) addr, read the 32bit value from the memory addresses starting from address addr' where addr' = addr exception that the last two bits are 0's. That is, read the data from addr', addr'+1, addr'+2, addr'+3.\ definition mem_val_w32 :: "asi_type \ phys_address \ ('a) sparc_state \ word32 option" where "mem_val_w32 asi addr state \ let addr' = (AND) addr 0b111111111111111111111111111111111100; addr0 = addr'; addr1 = addr' + 1; addr2 = addr' + 2; addr3 = addr' + 3; r0 = mem_val_alt asi addr0 state; r1 = mem_val_alt asi addr1 state; r2 = mem_val_alt asi addr2 state; r3 = mem_val_alt asi addr3 state in if r0 = None \ r1 = None \ r2 = None \ r3 = None then None else let byte0 = case r0 of Some v \ v; byte1 = case r1 of Some v \ v; byte2 = case r2 of Some v \ v; byte3 = case r3 of Some v \ v in Some ((OR) ((OR) ((OR) ((ucast(byte0)) << 24) ((ucast(byte1)) << 16)) ((ucast(byte2)) << 8)) (ucast(byte3))) " text \ Let \addr'\ be \addr\ with last two bits set to 0's. Write the 32bit data in the memory address \addr'\ (and the following 3 addresses). \byte_mask\ decides which byte of the 32bits are written. \ definition mem_mod_w32 :: "asi_type \ phys_address \ word4 \ word32 \ ('a) sparc_state \ ('a) sparc_state" where "mem_mod_w32 asi addr byte_mask data_w32 state \ let addr' = (AND) addr 0b111111111111111111111111111111111100; addr0 = (OR) addr' 0b000000000000000000000000000000000000; addr1 = (OR) addr' 0b000000000000000000000000000000000001; addr2 = (OR) addr' 0b000000000000000000000000000000000010; addr3 = (OR) addr' 0b000000000000000000000000000000000011; byte0 = (ucast (data_w32 >> 24))::mem_val_type; byte1 = (ucast (data_w32 >> 16))::mem_val_type; byte2 = (ucast (data_w32 >> 8))::mem_val_type; byte3 = (ucast data_w32)::mem_val_type; s0 = if (((AND) byte_mask (0b1000::word4)) >> 3) = 1 then mem_mod asi addr0 byte0 state else state; s1 = if (((AND) byte_mask (0b0100::word4)) >> 2) = 1 then mem_mod asi addr1 byte1 s0 else s0; s2 = if (((AND) byte_mask (0b0010::word4)) >> 1) = 1 then mem_mod asi addr2 byte2 s1 else s1; s3 = if ((AND) byte_mask (0b0001::word4)) = 1 then mem_mod asi addr3 byte3 s2 else s2 in s3 " text \The following functions deal with virtual addresses. These are based on functions written by David Sanan.\ definition load_word_mem :: "('a) sparc_state \ virtua_address \ asi_type \ machine_word option" where "load_word_mem state va asi \ let pair = (virt_to_phys va (mmu state) (mem state)) in case pair of Some pair \ ( if mmu_readable (get_acc_flag (snd pair)) asi then (mem_val_w32 asi (fst pair) state) else None) | None \ None" definition store_word_mem ::"('a) sparc_state \ virtua_address \ machine_word \ word4 \ asi_type \ ('a) sparc_state option" where "store_word_mem state va wd byte_mask asi \ let pair = (virt_to_phys va (mmu state) (mem state)) in case pair of Some pair \ ( if mmu_writable (get_acc_flag (snd pair)) asi then Some (mem_mod_w32 asi (fst pair) byte_mask wd state) else None) | None \ None" definition icache_val:: "cache_type \ ('a) sparc_state \ mem_val_type option" where "icache_val c state \ icache (cache state) c" definition dcache_val:: "cache_type \ ('a) sparc_state \ mem_val_type option" where "dcache_val c state \ dcache (cache state) c" definition icache_mod :: "cache_type \ mem_val_type \ ('a) sparc_state \ ('a) sparc_state" where "icache_mod c val state \ state\cache := ((cache state) \icache := (icache (cache state))(c := Some val)\)\ " definition dcache_mod :: "cache_type \ mem_val_type \ ('a) sparc_state \ ('a) sparc_state" where "dcache_mod c val state \ state\cache := ((cache state) \dcache := (dcache (cache state))(c := Some val)\)\ " text \Check if the memory address is in the cache or not.\ definition icache_miss :: "virtua_address \ ('a) sparc_state \ bool" where "icache_miss addr state \ let line_len = 12; tag = (ucast (addr >> line_len))::cache_tag; line = (ucast (0b0::word1))::cache_line_size in if (icache_val (tag,line) state) = None then True else False " text \Check if the memory address is in the cache or not.\ definition dcache_miss :: "virtua_address \ ('a) sparc_state \ bool" where "dcache_miss addr state \ let line_len = 12; tag = (ucast (addr >> line_len))::cache_tag; line = (ucast (0b0::word1))::cache_line_size in if (dcache_val (tag,line) state) = None then True else False " definition read_data_cache:: "('a) sparc_state \ virtua_address \ machine_word option" where "read_data_cache state va \ let tag = (ucast (va >> 12))::word20; offset0 = (AND) ((ucast va)::word12) 0b111111111100; offset1 = (OR) offset0 0b000000000001; offset2 = (OR) offset0 0b000000000010; offset3 = (OR) offset0 0b000000000011; r0 = dcache_val (tag,offset0) state; r1 = dcache_val (tag,offset1) state; r2 = dcache_val (tag,offset2) state; r3 = dcache_val (tag,offset3) state in if r0 = None \ r1 = None \ r2 = None \ r3 = None then None else let byte0 = case r0 of Some v \ v; byte1 = case r1 of Some v \ v; byte2 = case r2 of Some v \ v; byte3 = case r3 of Some v \ v in Some ((OR) ((OR) ((OR) ((ucast(byte0)) << 24) ((ucast(byte1)) << 16)) ((ucast(byte2)) << 8)) (ucast(byte3))) " definition read_instr_cache:: "('a) sparc_state \ virtua_address \ machine_word option" where "read_instr_cache state va \ let tag = (ucast (va >> 12))::word20; offset0 = (AND) ((ucast va)::word12) 0b111111111100; offset1 = (OR) offset0 0b000000000001; offset2 = (OR) offset0 0b000000000010; offset3 = (OR) offset0 0b000000000011; r0 = icache_val (tag,offset0) state; r1 = icache_val (tag,offset1) state; r2 = icache_val (tag,offset2) state; r3 = icache_val (tag,offset3) state in if r0 = None \ r1 = None \ r2 = None \ r3 = None then None else let byte0 = case r0 of Some v \ v; byte1 = case r1 of Some v \ v; byte2 = case r2 of Some v \ v; byte3 = case r3 of Some v \ v in Some ((OR) ((OR) ((OR) ((ucast(byte0)) << 24) ((ucast(byte1)) << 16)) ((ucast(byte2)) << 8)) (ucast(byte3))) " definition add_data_cache :: "('a) sparc_state \ virtua_address \ machine_word \ word4 \ ('a) sparc_state" where "add_data_cache state va word byte_mask \ let tag = (ucast (va >> 12))::word20; offset0 = (AND) ((ucast va)::word12) 0b111111111100; offset1 = (OR) offset0 0b000000000001; offset2 = (OR) offset0 0b000000000010; offset3 = (OR) offset0 0b000000000011; byte0 = (ucast (word >> 24))::mem_val_type; byte1 = (ucast (word >> 16))::mem_val_type; byte2 = (ucast (word >> 8))::mem_val_type; byte3 = (ucast word)::mem_val_type; s0 = if (((AND) byte_mask (0b1000::word4)) >> 3) = 1 then dcache_mod (tag,offset0) byte0 state else state; s1 = if (((AND) byte_mask (0b0100::word4)) >> 2) = 1 then dcache_mod (tag,offset1) byte1 s0 else s0; s2 = if (((AND) byte_mask (0b0010::word4)) >> 1) = 1 then dcache_mod (tag,offset2) byte2 s1 else s1; s3 = if ((AND) byte_mask (0b0001::word4)) = 1 then dcache_mod (tag,offset3) byte3 s2 else s2 in s3 " definition add_instr_cache :: "('a) sparc_state \ virtua_address \ machine_word \ word4 \ ('a) sparc_state" where "add_instr_cache state va word byte_mask \ let tag = (ucast (va >> 12))::word20; offset0 = (AND) ((ucast va)::word12) 0b111111111100; offset1 = (OR) offset0 0b000000000001; offset2 = (OR) offset0 0b000000000010; offset3 = (OR) offset0 0b000000000011; byte0 = (ucast (word >> 24))::mem_val_type; byte1 = (ucast (word >> 16))::mem_val_type; byte2 = (ucast (word >> 8))::mem_val_type; byte3 = (ucast word)::mem_val_type; s0 = if (((AND) byte_mask (0b1000::word4)) >> 3) = 1 then icache_mod (tag,offset0) byte0 state else state; s1 = if (((AND) byte_mask (0b0100::word4)) >> 2) = 1 then icache_mod (tag,offset1) byte1 s0 else s0; s2 = if (((AND) byte_mask (0b0010::word4)) >> 1) = 1 then icache_mod (tag,offset2) byte2 s1 else s1; s3 = if ((AND) byte_mask (0b0001::word4)) = 1 then icache_mod (tag,offset3) byte3 s2 else s2 in s3 " definition empty_cache ::"cache_context" where "empty_cache c \ None" definition flush_data_cache:: "('a) sparc_state \ ('a) sparc_state" where "flush_data_cache state \ state\cache := ((cache state)\dcache := empty_cache\)\" definition flush_instr_cache:: "('a) sparc_state \ ('a) sparc_state" where "flush_instr_cache state \ state\cache := ((cache state)\icache := empty_cache\)\" definition flush_cache_all:: "('a) sparc_state \ ('a) sparc_state" where "flush_cache_all state \ state\cache := ((cache state)\ icache := empty_cache, dcache := empty_cache\)\" text \Check if the FI or FD bit of CCR is 1. If FI is 1 then flush instruction cache. If FD is 1 then flush data cache.\ definition ccr_flush :: "('a) sparc_state \ ('a) sparc_state" where "ccr_flush state \ let ccr_val = sys_reg_val CCR state; \ \\FI\ is bit 21 of \CCR\\ fi_val = ((AND) ccr_val (0b00000000001000000000000000000000)) >> 21; fd_val = ((AND) ccr_val (0b00000000010000000000000000000000)) >> 22; state1 = (if fi_val = 1 then flush_instr_cache state else state) in if fd_val = 1 then flush_data_cache state1 else state1" definition get_delayed_pool :: "('a) sparc_state \ delayed_write_pool" where "get_delayed_pool state \ dwrite state" definition exe_pool :: "(int \ reg_type \ CPU_register) \ (int \ reg_type \ CPU_register)" where "exe_pool w \ case w of (n,v,c) \ ((n-1),v,c)" text \Minus 1 to the delayed count for all the members in the set. Assuming all members have delay > 0.\ primrec delayed_pool_minus :: "delayed_write_pool \ delayed_write_pool" where "delayed_pool_minus [] = []" | "delayed_pool_minus (x#xs) = (exe_pool x)#(delayed_pool_minus xs)" text \Add a delayed-write to the pool.\ definition delayed_pool_add :: "(int \ reg_type \ CPU_register) \ ('a) sparc_state \ ('a) sparc_state" where "delayed_pool_add dw s \ let (i,v,cr) = dw in if i = 0 then \ \Write the value to the register immediately.\ cpu_reg_mod v cr s else \ \Add to delayed write pool.\ let curr_pool = get_delayed_pool s in s\dwrite := curr_pool@[dw]\" text \Remove a delayed-write from the pool. Assume that the delayed-write to be removed has delay 0. i.e., it has been executed.\ definition delayed_pool_rm :: "(int \ reg_type \ CPU_register) \ ('a) sparc_state \ ('a) sparc_state" where "delayed_pool_rm dw s \ let curr_pool = get_delayed_pool s in case dw of (n,v,cr) \ (if n = 0 then s\dwrite := List.remove1 dw curr_pool\ else s) " text \Remove all the entries with delay = 0, i.e., those that are written.\ primrec delayed_pool_rm_written :: "delayed_write_pool \ delayed_write_pool" where "delayed_pool_rm_written [] = []" | "delayed_pool_rm_written (x#xs) = (if fst x = 0 then delayed_pool_rm_written xs else x#(delayed_pool_rm_written xs)) " definition annul_val :: "('a) sparc_state \ bool" where "annul_val state \ get_annul (state_var state)" definition annul_mod :: "bool \ ('a) sparc_state \ ('a) sparc_state" where "annul_mod b s \ s\state_var := write_annul b (state_var s)\" definition reset_trap_val :: "('a) sparc_state \ bool" where "reset_trap_val state \ get_reset_trap (state_var state)" definition reset_trap_mod :: "bool \ ('a) sparc_state \ ('a) sparc_state" where "reset_trap_mod b s \ s\state_var := write_reset_trap b (state_var s)\" definition exe_mode_val :: "('a) sparc_state \ bool" where "exe_mode_val state \ get_exe_mode (state_var state)" definition exe_mode_mod :: "bool \ ('a) sparc_state \ ('a) sparc_state" where "exe_mode_mod b s \ s\state_var := write_exe_mode b (state_var s)\" definition reset_mode_val :: "('a) sparc_state \ bool" where "reset_mode_val state \ get_reset_mode (state_var state)" definition reset_mode_mod :: "bool \ ('a) sparc_state \ ('a) sparc_state" where "reset_mode_mod b s \ s\state_var := write_reset_mode b (state_var s)\" definition err_mode_val :: "('a) sparc_state \ bool" where "err_mode_val state \ get_err_mode (state_var state)" definition err_mode_mod :: "bool \ ('a) sparc_state \ ('a) sparc_state" where "err_mode_mod b s \ s\state_var := write_err_mode b (state_var s)\" definition ticc_trap_type_val :: "('a) sparc_state \ word7" where "ticc_trap_type_val state \ get_ticc_trap_type (state_var state)" definition ticc_trap_type_mod :: "word7 \ ('a) sparc_state \ ('a) sparc_state" where "ticc_trap_type_mod w s \ s\state_var := write_ticc_trap_type w (state_var s)\" definition interrupt_level_val :: "('a) sparc_state \ word3" where "interrupt_level_val state \ get_interrupt_level (state_var state)" definition interrupt_level_mod :: "word3 \ ('a) sparc_state \ ('a) sparc_state" where "interrupt_level_mod w s \ s\state_var := write_interrupt_level w (state_var s)\" definition store_barrier_pending_val :: "('a) sparc_state \ bool" where "store_barrier_pending_val state \ get_store_barrier_pending (state_var state)" definition store_barrier_pending_mod :: "bool \ ('a) sparc_state \ ('a) sparc_state" where "store_barrier_pending_mod w s \ s\state_var := write_store_barrier_pending w (state_var s)\" definition pb_block_ldst_byte_val :: "virtua_address \ ('a) sparc_state \ bool" where "pb_block_ldst_byte_val add state \ (atm_ldst_byte (state_var state)) add" definition pb_block_ldst_byte_mod :: "virtua_address \ bool \ ('a) sparc_state \ ('a) sparc_state" where "pb_block_ldst_byte_mod add b s \ s\state_var := ((state_var s) \atm_ldst_byte := (atm_ldst_byte (state_var s))(add := b)\)\" text \We only read the address such that add mod 4 = 0. add mod 4 represents the current word.\ definition pb_block_ldst_word_val :: "virtua_address \ ('a) sparc_state \ bool" where "pb_block_ldst_word_val add state \ let add0 = ((AND) add (0b11111111111111111111111111111100::word32)) in (atm_ldst_word (state_var state)) add0" text \We only write the address such that add mod 4 = 0. add mod 4 represents the current word.\ definition pb_block_ldst_word_mod :: "virtua_address \ bool \ ('a) sparc_state \ ('a) sparc_state" where "pb_block_ldst_word_mod add b s \ let add0 = ((AND) add (0b11111111111111111111111111111100::word32)) in s\state_var := ((state_var s) \atm_ldst_word := (atm_ldst_word (state_var s))(add0 := b)\)\" definition get_trap_set :: "('a) sparc_state \ Trap set" where "get_trap_set state \ (traps state)" definition add_trap_set :: "Trap \ ('a) sparc_state \ ('a) sparc_state" where "add_trap_set t s \ s\traps := (traps s) \ {t}\" definition emp_trap_set :: "('a) sparc_state \ ('a) sparc_state" where "emp_trap_set s \ s\traps := {}\" definition state_undef:: "('a) sparc_state \ bool" where "state_undef state \ (undef state)" text \The \memory_read\ interface that conforms with the SPARCv8 manual.\ definition memory_read :: "asi_type \ virtua_address \ ('a) sparc_state \ ((word32 option) \ ('a) sparc_state)" where "memory_read asi addr state \ let asi_int = uint asi in \ \See Page 25 and 35 for ASI usage in LEON 3FT.\ if asi_int = 1 then \ \Forced cache miss.\ \ \Directly read from memory.\ let r1 = load_word_mem state addr (word_of_int 8) in if r1 = None then let r2 = load_word_mem state addr (word_of_int 10) in if r2 = None then (None,state) else (r2,state) else (r1,state) else if asi_int = 2 then \ \System registers.\ \ \See Table 19, Page 34 for System Register address map in LEON 3FT.\ if uint addr = 0 then \ \Cache control register.\ ((Some (sys_reg_val CCR state)), state) else if uint addr = 8 then \ \Instruction cache configuration register.\ ((Some (sys_reg_val ICCR state)), state) else if uint addr = 12 then \ \Data cache configuration register.\ ((Some (sys_reg_val DCCR state)), state) else \ \Invalid address.\ (None, state) else if asi_int \ {8,9} then \ \Access instruction memory.\ let ccr_val = (sys_reg state) CCR in if ccr_val AND 1 \ 0 then \ \Cache is enabled. Update cache.\ \ \We don't go through the tradition, i.e., read from cache first,\ \ \if the address is not cached, then read from memory,\ \ \because performance is not an issue here.\ \ \Thus we directly read from memory and update the cache.\ let data = load_word_mem state addr asi in case data of Some w \ (Some w,(add_instr_cache state addr w (0b1111::word4))) |None \ (None, state) else \ \Cache is disabled. Just read from memory.\ ((load_word_mem state addr asi),state) else if asi_int \ {10,11} then \ \Access data memory.\ let ccr_val = (sys_reg state) CCR in if ccr_val AND 1 \ 0 then \ \Cache is enabled. Update cache.\ \ \We don't go through the tradition, i.e., read from cache first,\ \ \if the address is not cached, then read from memory,\ \ \because performance is not an issue here.\ \ \Thus we directly read from memory and update the cache.\ let data = load_word_mem state addr asi in case data of Some w \ (Some w,(add_data_cache state addr w (0b1111::word4))) |None \ (None, state) else \ \Cache is disabled. Just read from memory.\ ((load_word_mem state addr asi),state) \ \We don't access instruction cache tag. i.e., \asi = 12\.\ else if asi_int = 13 then \ \Read instruction cache data.\ let cache_result = read_instr_cache state addr in case cache_result of Some w \ (Some w, state) |None \ (None, state) \ \We don't access data cache tag. i.e., \asi = 14\.\ else if asi_int = 15 then \ \Read data cache data.\ let cache_result = read_data_cache state addr in case cache_result of Some w \ (Some w, state) |None \ (None, state) else if asi_int \ {16,17} then \ \Flush entire instruction/data cache.\ (None, state) \ \Has no effect for memory read.\ else if asi_int \ {20,21} then \ \MMU diagnostic cache access.\ (None, state) \ \Not considered in this model.\ else if asi_int = 24 then \ \Flush cache and TLB in LEON3.\ \ \But is not used for memory read.\ (None, state) else if asi_int = 25 then \ \MMU registers.\ \ \Treat MMU registers as memory addresses that are not in the main memory.\ ((mmu_reg_val (mmu state) addr), state) else if asi_int = 28 then \ \MMU bypass.\ \ \Directly use addr as a physical address.\ \ \Append 0000 in the front of addr.\ \ \In this case, (ucast addr) suffices.\ ((mem_val_w32 asi (ucast addr) state), state) else if asi_int = 29 then \ \MMU diagnostic access.\ (None, state) \ \Not considered in this model.\ else \ \Not considered in this model.\ (None, state) " text \Get the value of a memory address and an ASI.\ definition mem_val_asi:: "asi_type \ phys_address \ ('a) sparc_state \ mem_val_type option" where "mem_val_asi asi add state \ (mem state) asi add" text \Check if an address is used in ASI 9 or 11.\ definition sup_addr :: "phys_address \ ('a) sparc_state \ bool" where "sup_addr addr state \ let addr' = (AND) addr 0b111111111111111111111111111111111100; addr0 = (OR) addr' 0b000000000000000000000000000000000000; addr1 = (OR) addr' 0b000000000000000000000000000000000001; addr2 = (OR) addr' 0b000000000000000000000000000000000010; addr3 = (OR) addr' 0b000000000000000000000000000000000011; r0 = mem_val_asi 9 addr0 state; r1 = mem_val_asi 9 addr1 state; r2 = mem_val_asi 9 addr2 state; r3 = mem_val_asi 9 addr3 state; r4 = mem_val_asi 11 addr0 state; r5 = mem_val_asi 11 addr1 state; r6 = mem_val_asi 11 addr2 state; r7 = mem_val_asi 11 addr3 state in if r0 = None \ r1 = None \ r2 = None \ r3 = None \ r4 = None \ r5 = None \ r6 = None \ r7 = None then False else True " text \The \memory_write\ interface that conforms with SPARCv8 manual.\ text \LEON3 forbids user to write an address in ASI 9 and 11.\ definition memory_write_asi :: "asi_type \ virtua_address \ word4 \ word32 \ ('a) sparc_state \ ('a) sparc_state option" where "memory_write_asi asi addr byte_mask data_w32 state \ let asi_int = uint asi; \ \See Page 25 and 35 for ASI usage in LEON 3FT.\ psr_val = cpu_reg_val PSR state; s_val = get_S psr_val in if asi_int = 1 then \ \Forced cache miss.\ \ \Directly write to memory.\ \ \Assuming writing into \asi = 10\.\ store_word_mem state addr data_w32 byte_mask (word_of_int 10) else if asi_int = 2 then \ \System registers.\ \ \See Table 19, Page 34 for System Register address map in LEON 3FT.\ if uint addr = 0 then \ \Cache control register.\ let s1 = (sys_reg_mod data_w32 CCR state) in \ \Flush the instruction cache if FI of CCR is 1;\ \ \flush the data cache if FD of CCR is 1.\ Some (ccr_flush s1) else if uint addr = 8 then \ \Instruction cache configuration register.\ Some (sys_reg_mod data_w32 ICCR state) else if uint addr = 12 then \ \Data cache configuration register.\ Some (sys_reg_mod data_w32 DCCR state) else \ \Invalid address.\ None else if asi_int \ {8,9} then \ \Access instruction memory.\ \ \Write to memory. LEON3 does write-through. Both cache and the memory are updated.\ let ns = add_instr_cache state addr data_w32 byte_mask in store_word_mem ns addr data_w32 byte_mask asi else if asi_int \ {10,11} then \ \Access data memory.\ \ \Write to memory. LEON3 does write-through. Both cache and the memory are updated.\ let ns = add_data_cache state addr data_w32 byte_mask in store_word_mem ns addr data_w32 byte_mask asi \ \We don't access instruction cache tag. i.e., \asi = 12\.\ else if asi_int = 13 then \ \Write instruction cache data.\ Some (add_instr_cache state addr data_w32 (0b1111::word4)) \ \We don't access data cache tag. i.e., asi = 14.\ else if asi_int = 15 then \ \Write data cache data.\ Some (add_data_cache state addr data_w32 (0b1111::word4)) else if asi_int = 16 then \ \Flush instruction cache.\ Some (flush_instr_cache state) else if asi_int = 17 then \ \Flush data cache.\ Some (flush_data_cache state) else if asi_int \ {20,21} then \ \MMU diagnostic cache access.\ None \ \Not considered in this model.\ else if asi_int = 24 then \ \Flush TLB and cache in LEON3.\ \ \We don't consider TLB here.\ Some (flush_cache_all state) else if asi_int = 25 then \ \MMU registers.\ \ \Treat MMU registers as memory addresses that are not in the main memory.\ let mmu_state' = mmu_reg_mod (mmu state) addr data_w32 in case mmu_state' of Some mmus \ Some (state\mmu := mmus\) |None \ None else if asi_int = 28 then \ \MMU bypass.\ \ \Write to virtual address as physical address.\ \ \Append 0000 in front of addr.\ Some (mem_mod_w32 asi (ucast addr) byte_mask data_w32 state) else if asi_int = 29 then \ \MMU diagnostic access.\ None \ \Not considered in this model.\ else \ \Not considered in this model.\ None " definition memory_write :: "asi_type \ virtua_address \ word4 \ word32 \ ('a) sparc_state \ ('a) sparc_state option" where "memory_write asi addr byte_mask data_w32 state \ let result = memory_write_asi asi addr byte_mask data_w32 state in case result of None \ None | Some s1 \ Some (store_barrier_pending_mod False s1)" text \monad for sequential operations over the register representation\ type_synonym ('a,'e) sparc_state_monad = "(('a) sparc_state,'e) det_monad" text \Given a word32 value, a cpu register, write the value in the cpu register.\ definition write_cpu :: "word32 \ CPU_register \ ('a,unit) sparc_state_monad" where "write_cpu w cr \ do modify (\s. (cpu_reg_mod w cr s)); return () od" definition write_cpu_tt :: "word8 \ ('a,unit) sparc_state_monad" where "write_cpu_tt w \ do tbr_val \ gets (\s. (cpu_reg_val TBR s)); new_tbr_val \ gets (\s. (write_tt w tbr_val)); write_cpu new_tbr_val TBR; return () od" text \Given a word32 value, a word4 window, a user register, write the value in the user register. N.B. CWP is a 5 bit value, but we only use the last 4 bits, since there are only 16 windows.\ definition write_reg :: "word32 \ ('a::len) word \ user_reg_type \ ('a,unit) sparc_state_monad" where "write_reg w win ur \ do modify (\s.(user_reg_mod w win ur s)); return () od" definition set_annul :: "bool \ ('a,unit) sparc_state_monad" where "set_annul b \ do modify (\s. (annul_mod b s)); return () od" definition set_reset_trap :: "bool \ ('a,unit) sparc_state_monad" where "set_reset_trap b \ do modify (\s. (reset_trap_mod b s)); return () od" definition set_exe_mode :: "bool \ ('a,unit) sparc_state_monad" where "set_exe_mode b \ do modify (\s. (exe_mode_mod b s)); return () od" definition set_reset_mode :: "bool \ ('a,unit) sparc_state_monad" where "set_reset_mode b \ do modify (\s. (reset_mode_mod b s)); return () od" definition set_err_mode :: "bool \ ('a,unit) sparc_state_monad" where "set_err_mode b \ do modify (\s. (err_mode_mod b s)); return () od" fun get_delayed_0 :: "(int \ reg_type \ CPU_register) list \ (int \ reg_type \ CPU_register) list" where "get_delayed_0 [] = []" | "get_delayed_0 (x # xs) = (if fst x = 0 then x # (get_delayed_0 xs) else get_delayed_0 xs)" text \Get a list of delayed-writes with delay 0.\ definition get_delayed_write :: "delayed_write_pool \ (int \ reg_type \ CPU_register) list" where "get_delayed_write dwp \ get_delayed_0 dwp" definition delayed_write :: "(int \ reg_type \ CPU_register) \ ('a) sparc_state \ ('a) sparc_state" where "delayed_write dw s \ let (n,v,r) = dw in if n = 0 then cpu_reg_mod v r s else s" primrec delayed_write_all :: "(int \ reg_type \ CPU_register) list \ ('a) sparc_state \ ('a) sparc_state" where "delayed_write_all [] s = s" |"delayed_write_all (x # xs) s = delayed_write_all xs (delayed_write x s)" primrec delayed_pool_rm_list :: "(int \ reg_type \ CPU_register) list\ ('a) sparc_state \ ('a) sparc_state" where "delayed_pool_rm_list [] s = s" |"delayed_pool_rm_list (x # xs) s = delayed_pool_rm_list xs (delayed_pool_rm x s)" definition delayed_pool_write :: "('a) sparc_state \ ('a) sparc_state" where "delayed_pool_write s \ let dwp0 = get_delayed_pool s; dwp1 = delayed_pool_minus dwp0; wl = get_delayed_write dwp1; s1 = delayed_write_all wl s; s2 = delayed_pool_rm_list wl s1 in s2" definition raise_trap :: "Trap \ ('a,unit) sparc_state_monad" where "raise_trap t \ do modify (\s. (add_trap_set t s)); return () od" end +end + diff --git a/thys/SPARCv8/SparcModel_MMU/Sparc_Types.thy b/thys/SPARCv8/SparcModel_MMU/Sparc_Types.thy --- a/thys/SPARCv8/SparcModel_MMU/Sparc_Types.thy +++ b/thys/SPARCv8/SparcModel_MMU/Sparc_Types.thy @@ -1,791 +1,803 @@ (* * 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, David Sanan. *) section \SPARC V8 architecture CPU model\ theory Sparc_Types imports Main "../lib/WordDecl" "Word_Lib.Bit_Shifts_Infix_Syntax" begin text \The following type definitions are taken from David Sanan's definitions for SPARC machines.\ type_synonym machine_word = word32 type_synonym byte = word8 type_synonym phys_address = word36 type_synonym virtua_address = word32 type_synonym page_address = word24 type_synonym offset = word12 type_synonym table_entry = word8 definition page_size :: "word32" where "page_size \ 4096" type_synonym virtua_page_address = word20 type_synonym context_type = word8 type_synonym word_length_t1 = word_length8 type_synonym word_length_t2 = word_length6 type_synonym word_length_t3 = word_length6 type_synonym word_length_offset = word_length12 type_synonym word_length_page = word_length24 type_synonym word_length_phys_address = word_length36 type_synonym word_length_virtua_address = word_length32 type_synonym word_length_entry_type = word_length2 type_synonym word_length_machine_word = word_length32 definition length_machine_word :: "nat" where "length_machine_word \ LENGTH(word_length_machine_word)" text_raw \\newpage\ section \CPU Register Definitions\ text\ The definitions below come from the SPARC Architecture Manual, Version 8. The LEON3 processor has been certified SPARC V8 conformant (2005). \ definition leon3khz ::"word32" where "leon3khz \ 33000" text \The following type definitions for MMU is taken from David Sanan's definitions for MMU.\ text\ The definitions below come from the UT699 LEON 3FT/SPARC V8 Microprocessor Functional Manual, Aeroflex, June 20, 2012, p35. \ datatype MMU_register = CR \ \Control Register\ | CTP \ \ConText Pointer register\ | CNR \ \Context Register\ | FTSR \ \Fault Status Register\ | FAR \ \Fault Address Register\ lemma MMU_register_induct: "P CR \ P CTP \ P CNR \ P FTSR \ P FAR \ P x" by (cases x) auto lemma UNIV_MMU_register [no_atp]: "UNIV = {CR, CTP, CNR, FTSR, FAR}" apply (safe) apply (case_tac x) apply (auto intro:MMU_register_induct) done instantiation MMU_register :: enum begin definition "enum_MMU_register = [ CR, CTP, CNR, FTSR, FAR ]" definition "enum_all_MMU_register P \ P CR \ P CTP \ P CNR \ P FTSR \ P FAR " definition "enum_ex_MMU_register P \ P CR \ P CTP \ P CNR \ P FTSR \ P FAR" instance proof qed (simp_all only: enum_MMU_register_def enum_all_MMU_register_def enum_ex_MMU_register_def UNIV_MMU_register, simp_all) end type_synonym MMU_context = "MMU_register \ machine_word" text \\PTE_flags\ is the last 8 bits of a PTE. See page 242 of SPARCv8 manual. \<^item> C - bit 7 \<^item> M - bit 6, \<^item> R - bit 5 \<^item> ACC - bit 4~2 \<^item> ET - bit 1~0.\ type_synonym PTE_flags = word8 text \ @{term CPU_register} datatype is an enumeration with the CPU registers defined in the SPARC V8 architecture. \ datatype CPU_register = PSR \ \Processor State Register\ | WIM \ \Window Invalid Mask\ | TBR \ \Trap Base Register\ | Y \ \Multiply/Divide Register\ | PC \ \Program Counter\ | nPC \ \next Program Counter\ | DTQ \ \Deferred-Trap Queue\ | FSR \ \Floating-Point State Register\ | FQ \ \Floating-Point Deferred-Trap Queue\ | CSR \ \Coprocessor State Register\ | CQ \ \Coprocessor Deferred-Trap Queue\ (*| CCR -- "Cache Control Register"*) | ASR "word5" \ \Ancillary State Register\ text \The following two functions are dummies since we will not use ASRs. Future formalisation may add more details to this.\ +context + includes bit_operations_syntax +begin + definition privileged_ASR :: "word5 \ bool" where "privileged_ASR r \ False " definition illegal_instruction_ASR :: "word5 \ bool" where "illegal_instruction_ASR r \ False " definition get_tt :: "word32 \ word8" where "get_tt tbr \ ucast (((AND) tbr 0b00000000000000000000111111110000) >> 4) " text \Write the tt field of the TBR register. Return the new value of TBR.\ definition write_tt :: "word8 \ word32 \ word32" where "write_tt new_tt_val tbr_val \ let tmp = (AND) tbr_val 0b111111111111111111111000000001111 in (OR) tmp (((ucast new_tt_val)::word32) << 4) " text \Get the nth bit of WIM. This equals ((AND) WIM $2^n$). N.B. the first bit of WIM is the 0th bit.\ definition get_WIM_bit :: "nat \ word32 \ word1" where "get_WIM_bit n wim \ let mask = ((ucast (0b1::word1))::word32) << n in ucast (((AND) mask wim) >> n) " definition get_CWP :: "word32 \ word5" where "get_CWP psr \ ucast ((AND) psr 0b00000000000000000000000000011111) " definition get_ET :: "word32 \ word1" where "get_ET psr \ ucast (((AND) psr 0b00000000000000000000000000100000) >> 5) " definition get_PIL :: "word32 \ word4" where "get_PIL psr \ ucast (((AND) psr 0b00000000000000000000111100000000) >> 8) " definition get_PS :: "word32 \ word1" where "get_PS psr \ ucast (((AND) psr 0b00000000000000000000000001000000) >> 6) " definition get_S :: "word32 \ word1" where "get_S psr \ \<^cancel>\ucast (((AND) psr 0b00000000000000000000000010000000) >> 7)\ if ((AND) psr (0b00000000000000000000000010000000::word32)) = 0 then 0 else 1 " definition get_icc_N :: "word32 \ word1" where "get_icc_N psr \ ucast (((AND) psr 0b00000000100000000000000000000000) >> 23) " definition get_icc_Z :: "word32 \ word1" where "get_icc_Z psr \ ucast (((AND) psr 0b00000000010000000000000000000000) >> 22) " definition get_icc_V :: "word32 \ word1" where "get_icc_V psr \ ucast (((AND) psr 0b00000000001000000000000000000000) >> 21) " definition get_icc_C :: "word32 \ word1" where "get_icc_C psr \ ucast (((AND) psr 0b00000000000100000000000000000000) >> 20) " definition update_S :: "word1 \ word32 \ word32" where "update_S s_val psr_val \ let tmp0 = (AND) psr_val 0b11111111111111111111111101111111 in (OR) tmp0 (((ucast s_val)::word32) << 7) " text \Update the CWP field of PSR. Return the new value of PSR.\ definition update_CWP :: "word5 \ word32 \ word32" where "update_CWP cwp_val psr_val \ let tmp0 = (AND) psr_val (0b11111111111111111111111111100000::word32); s_val = ((ucast (get_S psr_val))::word1) in if s_val = 0 then (AND) ((OR) tmp0 ((ucast cwp_val)::word32)) (0b11111111111111111111111101111111::word32) else (OR) ((OR) tmp0 ((ucast cwp_val)::word32)) (0b00000000000000000000000010000000::word32) " text \Update the the ET, CWP, and S fields of PSR. Return the new value of PSR.\ definition update_PSR_rett :: "word5 \ word1 \ word1 \ word32 \ word32" where "update_PSR_rett cwp_val et_val s_val psr_val \ let tmp0 = (AND) psr_val 0b11111111111111111111111101000000; tmp1 = (OR) tmp0 ((ucast cwp_val)::word32); tmp2 = (OR) tmp1 (((ucast et_val)::word32) << 5); tmp3 = (OR) tmp2 (((ucast s_val)::word32) << 7) in tmp3 " definition update_PSR_exe_trap :: "word5 \ word1 \ word1 \ word32 \ word32" where "update_PSR_exe_trap cwp_val et_val ps_val psr_val \ let tmp0 = (AND) psr_val 0b11111111111111111111111110000000; tmp1 = (OR) tmp0 ((ucast cwp_val)::word32); tmp2 = (OR) tmp1 (((ucast et_val)::word32) << 5); tmp3 = (OR) tmp2 (((ucast ps_val)::word32) << 6) in tmp3 " text \Update the N, Z, V, C fields of PSR. Return the new value of PSR.\ definition update_PSR_icc :: "word1 \ word1 \ word1 \ word1 \ word32 \ word32" where "update_PSR_icc n_val z_val v_val c_val psr_val \ let n_val_32 = if n_val = 0 then 0 else (0b00000000100000000000000000000000::word32); z_val_32 = if z_val = 0 then 0 else (0b00000000010000000000000000000000::word32); v_val_32 = if v_val = 0 then 0 else (0b00000000001000000000000000000000::word32); c_val_32 = if c_val = 0 then 0 else (0b00000000000100000000000000000000::word32); tmp0 = (AND) psr_val (0b11111111000011111111111111111111::word32); tmp1 = (OR) tmp0 n_val_32; tmp2 = (OR) tmp1 z_val_32; tmp3 = (OR) tmp2 v_val_32; tmp4 = (OR) tmp3 c_val_32 in tmp4 " text \Update the ET, PIL fields of PSR. Return the new value of PSR.\ definition update_PSR_et_pil :: "word1 \ word4 \ word32 \ word32" where "update_PSR_et_pil et pil psr_val \ let tmp0 = (AND) psr_val 0b111111111111111111111000011011111; tmp1 = (OR) tmp0 (((ucast et)::word32) << 5); tmp2 = (OR) tmp1 (((ucast pil)::word32) << 8) in tmp2 " +end + text \ SPARC V8 architecture is organized in windows of 32 user registers. The data stored in a register is defined as a 32 bits word @{term reg_type}: \ type_synonym reg_type = "word32" text \ The access to the value of a CPU register of type @{term CPU_register} is defined by a total function @{term cpu_context} \ type_synonym cpu_context = "CPU_register \ reg_type" text \ User registers are defined with the type @{term user_reg} represented by a 5 bits word. \ type_synonym user_reg_type = "word5" definition PSR_S ::"reg_type" where "PSR_S \ 6" text \ Each window context is defined by a total function @{term window_context} from @{term user_register} to @{term reg_type} (32 bits word storing the actual value of the register). \ type_synonym window_context = "user_reg_type \ reg_type" text \ The number of windows is implementation dependent. The LEON architecture is composed of 16 different windows (a 4 bits word). \ definition NWINDOWS :: "int" where "NWINDOWS \ 8" text \Maximum number of windows is 32 in SPARCv8.\ type_synonym ('a) window_size = "'a word" text \ Finally the user context is defined by another total function @{term user_context} from @{term window_size} to @{term window_context}. That is, the user context is a function taking as argument a register set window and a register within that window, and it returns the value stored in that user register. \ type_synonym ('a) user_context = "('a) window_size \ window_context" datatype sys_reg = CCR \ \Cache control register\ |ICCR \ \Instruction cache configuration register\ |DCCR \ \Data cache configuration register\ type_synonym sys_context = "sys_reg \ reg_type" text\ The memory model is defined by a total function from 32 bits words to 8 bits words \ type_synonym asi_type = "word8" text \ The memory is defined as a function from page address to page, which is also defined as a function from physical address to @{term "machine_word"} \ type_synonym mem_val_type = "word8" type_synonym mem_context = "asi_type \ phys_address \ mem_val_type option" type_synonym cache_tag = "word20" type_synonym cache_line_size = "word12" type_synonym cache_type = "(cache_tag \ cache_line_size)" type_synonym cache_context = "cache_type \ mem_val_type option" text \The delayed-write pool generated from write state register instructions.\ type_synonym delayed_write_pool = "(int \ reg_type \ CPU_register) list" definition DELAYNUM :: "int" where "DELAYNUM \ 0" text \Convert a set to a list.\ definition list_of_set :: "'a set \ 'a list" where "list_of_set s = (SOME l. set l = s)" lemma set_list_of_set: "finite s \ set (list_of_set s) = s" unfolding list_of_set_def by (metis (mono_tags) finite_list some_eq_ex) type_synonym ANNUL = "bool" type_synonym RESET_TRAP = "bool" type_synonym EXECUTE_MODE = "bool" type_synonym RESET_MODE = "bool" type_synonym ERROR_MODE = "bool" type_synonym TICC_TRAP_TYPE = "word7" type_synonym INTERRUPT_LEVEL = "word3" type_synonym STORE_BARRIER_PENDING = "bool" text \The processor asserts this signal to ensure that the memory system will not process another SWAP or LDSTUB operation to the same memory byte.\ type_synonym pb_block_ldst_byte = "virtua_address \ bool" text\The processor asserts this signal to ensure that the memory system will not process another SWAP or LDSTUB operation to the same memory word.\ type_synonym pb_block_ldst_word = "virtua_address \ bool" record sparc_state_var = annul:: ANNUL resett:: RESET_TRAP exe:: EXECUTE_MODE reset:: RESET_MODE err:: ERROR_MODE ticc:: TICC_TRAP_TYPE itrpt_lvl:: INTERRUPT_LEVEL st_bar:: STORE_BARRIER_PENDING atm_ldst_byte:: pb_block_ldst_byte atm_ldst_word:: pb_block_ldst_word definition get_annul :: "sparc_state_var \ bool" where "get_annul v \ annul v" definition get_reset_trap :: "sparc_state_var \ bool" where "get_reset_trap v \ resett v" definition get_exe_mode :: "sparc_state_var \ bool" where "get_exe_mode v \ exe v" definition get_reset_mode :: "sparc_state_var \ bool" where "get_reset_mode v \ reset v" definition get_err_mode :: "sparc_state_var \ bool" where "get_err_mode v \ err v" definition get_ticc_trap_type :: "sparc_state_var \ word7" where "get_ticc_trap_type v \ ticc v" definition get_interrupt_level :: "sparc_state_var \ word3" where "get_interrupt_level v \ itrpt_lvl v" definition get_store_barrier_pending :: "sparc_state_var \ bool" where "get_store_barrier_pending v \ st_bar v" definition write_annul :: "bool \ sparc_state_var \ sparc_state_var" where "write_annul b v \ v\annul := b\" definition write_reset_trap :: "bool \ sparc_state_var \ sparc_state_var" where "write_reset_trap b v \ v\resett := b\" definition write_exe_mode :: "bool \ sparc_state_var \ sparc_state_var" where "write_exe_mode b v \ v\exe := b\" definition write_reset_mode :: "bool \ sparc_state_var \ sparc_state_var" where "write_reset_mode b v \ v\reset := b\" definition write_err_mode :: "bool \ sparc_state_var \ sparc_state_var" where "write_err_mode b v \ v\err := b\" definition write_ticc_trap_type :: "word7 \ sparc_state_var \ sparc_state_var" where "write_ticc_trap_type w v \ v\ticc := w\" definition write_interrupt_level :: "word3 \ sparc_state_var \ sparc_state_var" where "write_interrupt_level w v \ v\itrpt_lvl := w\" definition write_store_barrier_pending :: "bool \ sparc_state_var \ sparc_state_var" where "write_store_barrier_pending b v \ v\st_bar := b\" +context + includes bit_operations_syntax +begin + text \Given a word7 value, find the highest bit, and fill the left bits to be the highest bit.\ definition sign_ext7::"word7 \ word32" where "sign_ext7 w \ let highest_bit = ((AND) w 0b1000000) >> 6 in if highest_bit = 0 then (ucast w)::word32 else (OR) ((ucast w)::word32) 0b11111111111111111111111110000000 " definition zero_ext8 :: "word8 \ word32" where "zero_ext8 w \ (ucast w)::word32 " text \Given a word8 value, find the highest bit, and fill the left bits to be the highest bit.\ definition sign_ext8::"word8 \ word32" where "sign_ext8 w \ let highest_bit = ((AND) w 0b10000000) >> 7 in if highest_bit = 0 then (ucast w)::word32 else (OR) ((ucast w)::word32) 0b11111111111111111111111100000000 " text \Given a word13 value, find the highest bit, and fill the left bits to be the highest bit.\ definition sign_ext13::"word13 \ word32" where "sign_ext13 w \ let highest_bit = ((AND) w 0b1000000000000) >> 12 in if highest_bit = 0 then (ucast w)::word32 else (OR) ((ucast w)::word32) 0b11111111111111111110000000000000 " definition zero_ext16 :: "word16 \ word32" where "zero_ext16 w \ (ucast w)::word32 " text \Given a word16 value, find the highest bit, and fill the left bits to be the highest bit.\ definition sign_ext16::"word16 \ word32" where "sign_ext16 w \ let highest_bit = ((AND) w 0b1000000000000000) >> 15 in if highest_bit = 0 then (ucast w)::word32 else (OR) ((ucast w)::word32) 0b11111111111111110000000000000000 " text \Given a word22 value, find the highest bit, and fill the left bits to tbe the highest bit.\ definition sign_ext22::"word22 \ word32" where "sign_ext22 w \ let highest_bit = ((AND) w 0b1000000000000000000000) >> 21 in if highest_bit = 0 then (ucast w)::word32 else (OR) ((ucast w)::word32) 0b11111111110000000000000000000000 " text \Given a word24 value, find the highest bit, and fill the left bits to tbe the highest bit.\ definition sign_ext24::"word24 \ word32" where "sign_ext24 w \ let highest_bit = ((AND) w 0b100000000000000000000000) >> 23 in if highest_bit = 0 then (ucast w)::word32 else (OR) ((ucast w)::word32) 0b11111111000000000000000000000000 " text\ Operations to be defined. The SPARC V8 architecture is composed of the following set of instructions: \<^item> Load Integer Instructions \<^item> Load Floating-point Instructions \<^item> Load Coprocessor Instructions \<^item> Store Integer Instructions \<^item> Store Floating-point Instructions \<^item> Store Coprocessor Instructions \<^item> Atomic Load-Store Unsigned Byte Instructions \<^item> SWAP Register With Memory Instruction \<^item> SETHI Instructions \<^item> NOP Instruction \<^item> Logical Instructions \<^item> Shift Instructions \<^item> Add Instructions \<^item> Tagged Add Instructions \<^item> Subtract Instructions \<^item> Tagged Subtract Instructions \<^item> Multiply Step Instruction \<^item> Multiply Instructions \<^item> Divide Instructions \<^item> SAVE and RESTORE Instructions \<^item> Branch on Integer Condition Codes Instructions \<^item> Branch on Floating-point Condition Codes Instructions \<^item> Branch on Coprocessor Condition Codes Instructions \<^item> Call and Link Instruction \<^item> Jump and Link Instruction \<^item> Return from Trap Instruction \<^item> Trap on Integer Condition Codes Instructions \<^item> Read State Register Instructions \<^item> Write State Register Instructions \<^item> STBAR Instruction \<^item> Unimplemented Instruction \<^item> Flush Instruction Memory \<^item> Floating-point Operate (FPop) Instructions \<^item> Convert Integer to Floating point Instructions \<^item> Convert Floating point to Integer Instructions \<^item> Convert Between Floating-point Formats Instructions \<^item> Floating-point Move Instructions \<^item> Floating-point Square Root Instructions \<^item> Floating-point Add and Subtract Instructions \<^item> Floating-point Multiply and Divide Instructions \<^item> Floating-point Compare Instructions \<^item> Coprocessor Operate Instructions \ text \The CALL instruction.\ datatype call_type = CALL \ \Call and Link\ text \The SETHI instruction.\ datatype sethi_type = SETHI \ \Set High 22 bits of r Register\ text \The NOP instruction.\ datatype nop_type = NOP \ \No Operation\ text \The Branch on integer condition codes instructions.\ datatype bicc_type = BE \ \Branch on Equal\ | BNE \ \Branch on Not Equal\ | BGU \ \Branch on Greater Unsigned\ | BLE \ \Branch on Less or Equal\ | BL \ \Branch on Less\ | BGE \ \Branch on Greater or Equal\ | BNEG \ \Branch on Negative\ | BG \ \Branch on Greater\ | BCS \ \Branch on Carry Set (Less than, Unsigned)\ | BLEU \ \Branch on Less or Equal Unsigned\ | BCC \ \Branch on Carry Clear (Greater than or Equal, Unsigned)\ | BA \ \Branch Always\ | BN \ \Branch Never\ \ \Added for unconditional branches\ | BPOS \ \Branch on Positive\ | BVC \ \Branch on Overflow Clear\ | BVS \ \Branch on Overflow Set\ text \Memory instructions. That is, load and store.\ datatype load_store_type = LDSB \ \Load Signed Byte\ | LDUB \ \Load Unsigned Byte\ | LDUBA \ \Load Unsigned Byte from Alternate space\ | LDUH \ \Load Unsigned Halfword\ | LD \ \Load Word\ | LDA \ \Load Word from Alternate space\ | LDD \ \Load Doubleword\ | STB \ \Store Byte\ | STH \ \Store Halfword\ | ST \ \Store Word\ | STA \ \Store Word into Alternate space\ | STD \ \Store Doubleword\ | LDSBA \ \Load Signed Byte from Alternate space\ | LDSH \ \Load Signed Halfword\ | LDSHA \ \Load Signed Halfword from Alternate space\ | LDUHA \ \Load Unsigned Halfword from Alternate space\ | LDDA \ \Load Doubleword from Alternate space\ | STBA \ \Store Byte into Alternate space\ | STHA \ \Store Halfword into Alternate space\ | STDA \ \Store Doubleword into Alternate space\ | LDSTUB \ \Atomic Load Store Unsigned Byte\ | LDSTUBA \ \Atomic Load Store Unsinged Byte in Alternate space\ | SWAP \ \Swap r Register with Mmemory\ | SWAPA \ \Swap r Register with Mmemory in Alternate space\ | FLUSH \ \Flush Instruction Memory\ | STBAR \ \Store Barrier\ text \Arithmetic instructions.\ datatype arith_type = ADD \ \Add\ | ADDcc \ \Add and modify icc\ | ADDX \ \Add with Carry\ | SUB \ \Subtract\ | SUBcc \ \Subtract and modify icc\ | SUBX \ \Subtract with Carry\ | UMUL \ \Unsigned Integer Multiply\ | SMUL \ \Signed Integer Multiply\ | SMULcc \ \Signed Integer Multiply and modify icc\ | UDIV \ \Unsigned Integer Divide\ | UDIVcc \ \Unsigned Integer Divide and modify icc\ | SDIV \ \Signed Integer Divide\ | ADDXcc \ \Add with Carry and modify icc\ | TADDcc \ \Tagged Add and modify icc\ | TADDccTV \ \Tagged Add and modify icc and Trap on overflow\ | SUBXcc \ \Subtract with Carry and modify icc\ | TSUBcc \ \Tagged Subtract and modify icc\ | TSUBccTV \ \Tagged Subtract and modify icc and Trap on overflow\ | MULScc \ \Multiply Step and modify icc\ | UMULcc \ \Unsigned Integer Multiply and modify icc\ | SDIVcc \ \Signed Integer Divide and modify icc\ text \Logical instructions.\ datatype logic_type = ANDs \ \And\ | ANDcc \ \And and modify icc\ | ANDN \ \And Not\ | ANDNcc \ \And Not and modify icc\ | ORs \ \Inclusive-Or\ | ORcc \ \Inclusive-Or and modify icc\ | ORN \ \Inclusive Or Not\ | XORs \ \Exclusive-Or\ | XNOR \ \Exclusive-Nor\ | ORNcc \ \Inclusive-Or Not and modify icc\ | XORcc \ \Exclusive-Or and modify icc\ | XNORcc \ \Exclusive-Nor and modify icc\ text \Shift instructions.\ datatype shift_type = SLL \ \Shift Left Logical\ | SRL \ \Shift Right Logical\ | SRA \ \Shift Right Arithmetic\ text \Other Control-transfer instructions.\ datatype ctrl_type = JMPL \ \Jump and Link\ | RETT \ \Return from Trap\ | SAVE \ \Save caller's window\ | RESTORE \ \Restore caller's window\ text \Access state registers instructions.\ datatype sreg_type = RDASR \ \Read Ancillary State Register\ | RDY \ \Read Y Register\ | RDPSR \ \Read Processor State Register\ | RDWIM \ \Read Window Invalid Mask Register\ | RDTBR \ \Read Trap Base Regiser\ | WRASR \ \Write Ancillary State Register\ | WRY \ \Write Y Register\ | WRPSR \ \Write Processor State Register\ | WRWIM \ \Write Window Invalid Mask Register\ | WRTBR \ \Write Trap Base Register\ text \Unimplemented instruction.\ datatype uimp_type = UNIMP \ \Unimplemented\ text \Trap on integer condition code instructions.\ datatype ticc_type = TA \ \Trap Always\ | TN \ \Trap Never\ | TNE \ \Trap on Not Equal\ | TE \ \Trap on Equal\ | TG \ \Trap on Greater\ | TLE \ \Trap on Less or Equal\ | TGE \ \Trap on Greater or Equal\ | TL \ \Trap on Less\ | TGU \ \Trap on Greater Unsigned\ | TLEU \ \Trap on Less or Equal Unsigned\ | TCC \ \Trap on Carry Clear (Greater than or Equal, Unsigned)\ | TCS \ \Trap on Carry Set (Less Than, Unsigned)\ | TPOS \ \Trap on Postive\ | TNEG \ \Trap on Negative\ | TVC \ \Trap on Overflow Clear\ | TVS \ \Trap on Overflow Set\ datatype sparc_operation = call_type call_type | sethi_type sethi_type | nop_type nop_type | bicc_type bicc_type | load_store_type load_store_type | arith_type arith_type | logic_type logic_type | shift_type shift_type | ctrl_type ctrl_type | sreg_type sreg_type | uimp_type uimp_type | ticc_type ticc_type datatype Trap = reset |data_store_error |instruction_access_MMU_miss |instruction_access_error |r_register_access_error |instruction_access_exception |privileged_instruction |illegal_instruction |unimplemented_FLUSH |watchpoint_detected |fp_disabled |cp_disabled |window_overflow |window_underflow |mem_address_not_aligned |fp_exception |cp_exception |data_access_error |data_access_MMU_miss |data_access_exception |tag_overflow |division_by_zero |trap_instruction |interrupt_level_n datatype Exception = \ \The following are processor states that are not in the instruction model,\ \ \but we MAY want to deal with these from hardware perspective.\ \<^cancel>\|execute_mode\ \<^cancel>\|reset_mode\ \<^cancel>\|error_mode\ \ \The following are self-defined exceptions.\ invalid_cond_f2 |invalid_op2_f2 |illegal_instruction2 \ \when \i = 0\ for load/store not from alternate space\ |invalid_op3_f3_op11 |case_impossible |invalid_op3_f3_op10 |invalid_op_f3 |unsupported_instruction |fetch_instruction_error |invalid_trap_cond 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,1328 +1,1334 @@ (* * 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 (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) apply (simp_all add: is_aligned_iff_take_bit_eq_0) apply (simp add: bit_eq_iff) apply (auto simp add: bit_simps) subgoal for m apply (cases \m < n\) apply (auto simp add: not_less) apply (metis bit_take_bit_iff less_mask_eq take_bit_eq_mask) 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 (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 (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/Ancient_Numeral.thy b/thys/Word_Lib/Ancient_Numeral.thy --- a/thys/Word_Lib/Ancient_Numeral.thy +++ b/thys/Word_Lib/Ancient_Numeral.thy @@ -1,231 +1,237 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) theory Ancient_Numeral imports Main Reversed_Bit_Lists Legacy_Aliases begin definition Bit :: "int \ bool \ int" (infixl "BIT" 90) where "k BIT b = (if b then 1 else 0) + k + k" lemma Bit_B0: "k BIT False = k + k" by (simp add: Bit_def) lemma Bit_B1: "k BIT True = k + k + 1" by (simp add: Bit_def) lemma Bit_B0_2t: "k BIT False = 2 * k" by (rule trans, rule Bit_B0) simp lemma Bit_B1_2t: "k BIT True = 2 * k + 1" by (rule trans, rule Bit_B1) simp lemma uminus_Bit_eq: "- k BIT b = (- k - of_bool b) BIT b" by (cases b) (simp_all add: Bit_def) lemma power_BIT: "2 ^ Suc n - 1 = (2 ^ n - 1) BIT True" by (simp add: Bit_B1) lemma bin_rl_simp [simp]: "bin_rest w BIT bin_last w = w" by (simp add: Bit_def) lemma bin_rest_BIT [simp]: "bin_rest (x BIT b) = x" by (simp add: Bit_def) lemma even_BIT [simp]: "even (x BIT b) \ \ b" by (simp add: Bit_def) lemma bin_last_BIT [simp]: "bin_last (x BIT b) = b" by simp lemma BIT_eq_iff [iff]: "u BIT b = v BIT c \ u = v \ b = c" by (auto simp: Bit_def) arith+ lemma BIT_bin_simps [simp]: "numeral k BIT False = numeral (Num.Bit0 k)" "numeral k BIT True = numeral (Num.Bit1 k)" "(- numeral k) BIT False = - numeral (Num.Bit0 k)" "(- numeral k) BIT True = - numeral (Num.BitM k)" by (simp_all only: Bit_B0 Bit_B1 numeral.simps numeral_BitM) lemma BIT_special_simps [simp]: shows "0 BIT False = 0" and "0 BIT True = 1" and "1 BIT False = 2" and "1 BIT True = 3" and "(- 1) BIT False = - 2" and "(- 1) BIT True = - 1" by (simp_all add: Bit_def) lemma Bit_eq_0_iff: "w BIT b = 0 \ w = 0 \ \ b" by (auto simp: Bit_def) arith lemma Bit_eq_m1_iff: "w BIT b = -1 \ w = -1 \ b" by (auto simp: Bit_def) arith lemma expand_BIT: "numeral (Num.Bit0 w) = numeral w BIT False" "numeral (Num.Bit1 w) = numeral w BIT True" "- numeral (Num.Bit0 w) = (- numeral w) BIT False" "- numeral (Num.Bit1 w) = (- numeral (w + Num.One)) BIT True" by (simp_all add: BitM_inc_eq add_One) lemma less_Bits: "v BIT b < w BIT c \ v < w \ v \ w \ \ b \ c" by (auto simp: Bit_def) lemma le_Bits: "v BIT b \ w BIT c \ v < w \ v \ w \ (\ b \ c)" by (auto simp: Bit_def) lemma pred_BIT_simps [simp]: "x BIT False - 1 = (x - 1) BIT True" "x BIT True - 1 = x BIT False" by (simp_all add: Bit_B0_2t Bit_B1_2t) lemma succ_BIT_simps [simp]: "x BIT False + 1 = x BIT True" "x BIT True + 1 = (x + 1) BIT False" by (simp_all add: Bit_B0_2t Bit_B1_2t) lemma add_BIT_simps [simp]: "x BIT False + y BIT False = (x + y) BIT False" "x BIT False + y BIT True = (x + y) BIT True" "x BIT True + y BIT False = (x + y) BIT True" "x BIT True + y BIT True = (x + y + 1) BIT False" by (simp_all add: Bit_B0_2t Bit_B1_2t) lemma mult_BIT_simps [simp]: "x BIT False * y = (x * y) BIT False" "x * y BIT False = (x * y) BIT False" "x BIT True * y = (x * y) BIT False + y" by (simp_all add: Bit_B0_2t Bit_B1_2t algebra_simps) lemma B_mod_2': "X = 2 \ (w BIT True) mod X = 1 \ (w BIT False) mod X = 0" by (simp add: Bit_B0 Bit_B1) lemma bin_ex_rl: "\w b. w BIT b = bin" by (metis bin_rl_simp) lemma bin_exhaust: "(\x b. bin = x BIT b \ Q) \ Q" by (metis bin_ex_rl) lemma bin_abs_lem: "bin = (w BIT b) \ bin \ -1 \ bin \ 0 \ nat \w\ < nat \bin\" apply clarsimp apply (unfold Bit_def) apply (cases b) apply (clarsimp, arith) apply (clarsimp, arith) done lemma bin_induct: assumes PPls: "P 0" and PMin: "P (- 1)" and PBit: "\bin bit. P bin \ P (bin BIT bit)" shows "P bin" apply (rule_tac P=P and a=bin and f1="nat \ abs" in wf_measure [THEN wf_induct]) apply (simp add: measure_def inv_image_def) apply (case_tac x rule: bin_exhaust) apply (frule bin_abs_lem) apply (auto simp add : PPls PMin PBit) done lemma Bit_div2: "(w BIT b) div 2 = w" by (fact bin_rest_BIT) lemma twice_conv_BIT: "2 * x = x BIT False" by (simp add: Bit_def) lemma BIT_lt0 [simp]: "x BIT b < 0 \ x < 0" by(cases b)(auto simp add: Bit_def) lemma BIT_ge0 [simp]: "x BIT b \ 0 \ x \ 0" by(cases b)(auto simp add: Bit_def) lemma bin_to_bl_aux_Bit_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (w BIT b) bl = bin_to_bl_aux (n - 1) w (b # bl)" by (cases n) auto lemma bl_to_bin_BIT: "bl_to_bin bs BIT b = bl_to_bin (bs @ [b])" by (simp add: bl_to_bin_append Bit_def) lemma bin_nth_0_BIT: "bin_nth (w BIT b) 0 \ b" by simp lemma bin_nth_Suc_BIT: "bin_nth (w BIT b) (Suc n) = bin_nth w n" by (simp add: bit_Suc) lemma bin_nth_minus [simp]: "0 < n \ bin_nth (w BIT b) n = bin_nth w (n - 1)" by (cases n) (simp_all add: bit_Suc) lemma bin_sign_simps [simp]: "bin_sign (w BIT b) = bin_sign w" by (simp add: bin_sign_def Bit_def) lemma bin_nth_Bit: "bin_nth (w BIT b) n \ n = 0 \ b \ (\m. n = Suc m \ bin_nth w m)" by (cases n) auto lemmas sbintrunc_Suc_BIT [simp] = signed_take_bit_Suc [where a="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b lemmas sbintrunc_0_BIT_B0 [simp] = signed_take_bit_0 [where a="w BIT False", simplified bin_last_numeral_simps bin_rest_numeral_simps] for w lemmas sbintrunc_0_BIT_B1 [simp] = signed_take_bit_0 [where a="w BIT True", simplified bin_last_BIT bin_rest_numeral_simps] for w lemma sbintrunc_Suc_minus_Is: \0 < n \ sbintrunc (n - 1) w = y \ sbintrunc n (w BIT b) = y BIT b\ by (cases n) (simp_all add: Bit_def signed_take_bit_Suc) lemma bin_cat_Suc_Bit: "bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b" by (auto simp add: Bit_def concat_bit_Suc) lemma int_not_BIT [simp]: "NOT (w BIT b) = (NOT w) BIT (\ b)" by (simp add: not_int_def Bit_def) +context + includes bit_operations_syntax +begin + lemma int_and_Bits [simp]: "(x BIT b) AND (y BIT c) = (x AND y) BIT (b \ c)" using and_int_rec [of \x BIT b\ \y BIT c\] by (auto simp add: Bit_B0_2t Bit_B1_2t) lemma int_or_Bits [simp]: "(x BIT b) OR (y BIT c) = (x OR y) BIT (b \ c)" using or_int_rec [of \x BIT b\ \y BIT c\] by (auto simp add: Bit_B0_2t Bit_B1_2t) lemma int_xor_Bits [simp]: "(x BIT b) XOR (y BIT c) = (x XOR y) BIT ((b \ c) \ \ (b \ c))" using xor_int_rec [of \x BIT b\ \y BIT c\] by (auto simp add: Bit_B0_2t Bit_B1_2t) +end + lemma mod_BIT: "bin BIT bit mod 2 ^ Suc n = (bin mod 2 ^ n) BIT bit" for bit proof - have "2 * (bin mod 2 ^ n) + 1 = (2 * bin mod 2 ^ Suc n) + 1" by (simp add: mod_mult_mult1) also have "\ = ((2 * bin mod 2 ^ Suc n) + 1) mod 2 ^ Suc n" by (simp add: ac_simps pos_zmod_mult_2) also have "\ = (2 * bin + 1) mod 2 ^ Suc n" by (simp only: mod_simps) finally show ?thesis by (auto simp add: Bit_def) qed lemma minus_BIT_0: fixes x y :: int shows "x BIT b - y BIT False = (x - y) BIT b" by(simp add: Bit_def) lemma int_lsb_BIT [simp]: fixes x :: int shows "lsb (x BIT b) \ b" by(simp add: lsb_int_def) lemma int_shiftr_BIT [simp]: fixes x :: int shows int_shiftr0: "drop_bit 0 x = x" and int_shiftr_Suc: "drop_bit (Suc n) (x BIT b) = drop_bit n x" by (simp_all add: drop_bit_Suc) lemma msb_BIT [simp]: "msb (x BIT b) = msb x" by(simp add: msb_int_def) end \ No newline at end of file diff --git a/thys/Word_Lib/Bits_Int.thy b/thys/Word_Lib/Bits_Int.thy --- a/thys/Word_Lib/Bits_Int.thy +++ b/thys/Word_Lib/Bits_Int.thy @@ -1,1561 +1,1573 @@ (* * Copyright Brian Huffman, PSU; Jeremy Dawson and Gerwin Klein, NICTA * * SPDX-License-Identifier: BSD-2-Clause *) section \Bitwise Operations on integers\ theory Bits_Int imports "Word_Lib.Most_significant_bit" "Word_Lib.Least_significant_bit" "Word_Lib.Generic_set_bit" "Word_Lib.Bit_Comprehension" begin subsection \Implicit bit representation of \<^typ>\int\\ lemma bin_last_def: "(odd :: int \ bool) w \ w mod 2 = 1" by (fact odd_iff_mod_2_eq_one) lemma bin_last_numeral_simps [simp]: "\ odd (0 :: int)" "odd (1 :: int)" "odd (- 1 :: int)" "odd (Numeral1 :: int)" "\ odd (numeral (Num.Bit0 w) :: int)" "odd (numeral (Num.Bit1 w) :: int)" "\ odd (- numeral (Num.Bit0 w) :: int)" "odd (- numeral (Num.Bit1 w) :: int)" by simp_all lemma bin_rest_numeral_simps [simp]: "(\k::int. k div 2) 0 = 0" "(\k::int. k div 2) 1 = 0" "(\k::int. k div 2) (- 1) = - 1" "(\k::int. k div 2) Numeral1 = 0" "(\k::int. k div 2) (numeral (Num.Bit0 w)) = numeral w" "(\k::int. k div 2) (numeral (Num.Bit1 w)) = numeral w" "(\k::int. k div 2) (- numeral (Num.Bit0 w)) = - numeral w" "(\k::int. k div 2) (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)" by simp_all lemma bin_rl_eqI: "\(\k::int. k div 2) x = (\k::int. k div 2) y; odd x = odd y\ \ x = y" by (auto elim: oddE) lemma [simp]: shows bin_rest_lt0: "(\k::int. k div 2) i < 0 \ i < 0" and bin_rest_ge_0: "(\k::int. k div 2) i \ 0 \ i \ 0" by auto lemma bin_rest_gt_0 [simp]: "(\k::int. k div 2) x > 0 \ x > 1" by auto subsection \Bit projection\ lemma bin_nth_eq_iff: "(bit :: int \ nat \ bool) x = (bit :: int \ nat \ bool) y \ x = y" by (simp add: bit_eq_iff fun_eq_iff) lemma bin_eqI: "x = y" if "\n. (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" using that by (rule bit_eqI) lemma bin_eq_iff: "x = y \ (\n. (bit :: int \ nat \ bool) x n = (bit :: int \ nat \ bool) y n)" by (fact bit_eq_iff) lemma bin_nth_zero [simp]: "\ (bit :: int \ nat \ bool) 0 n" by simp lemma bin_nth_1 [simp]: "(bit :: int \ nat \ bool) 1 n \ n = 0" by (cases n) (simp_all add: bit_Suc) lemma bin_nth_minus1 [simp]: "(bit :: int \ nat \ bool) (- 1) n" by simp lemma bin_nth_numeral: "(\k::int. k div 2) x = y \ (bit :: int \ nat \ bool) x (numeral n) = (bit :: int \ nat \ bool) y (pred_numeral n)" by (simp add: numeral_eq_Suc bit_Suc) lemmas bin_nth_numeral_simps [simp] = bin_nth_numeral [OF bin_rest_numeral_simps(8)] lemmas bin_nth_simps = bit_0 bit_Suc bin_nth_zero bin_nth_minus1 bin_nth_numeral_simps lemma nth_2p_bin: "(bit :: int \ nat \ bool) (2 ^ n) m = (m = n)" \ \for use when simplifying with \bin_nth_Bit\\ by (auto simp add: bit_exp_iff) lemma nth_rest_power_bin: "(bit :: int \ nat \ bool) (((\k::int. k div 2) ^^ k) w) n = (bit :: int \ nat \ bool) w (n + k)" apply (induct k arbitrary: n) apply clarsimp apply clarsimp apply (simp only: bit_Suc [symmetric] add_Suc) done lemma bin_nth_numeral_unfold: "(bit :: int \ nat \ bool) (numeral (num.Bit0 x)) n \ n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1)" "(bit :: int \ nat \ bool) (numeral (num.Bit1 x)) n \ (n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1))" by (cases n; simp)+ subsection \Truncating\ definition bin_sign :: "int \ int" where "bin_sign k = (if k \ 0 then 0 else - 1)" lemma bin_sign_simps [simp]: "bin_sign 0 = 0" "bin_sign 1 = 0" "bin_sign (- 1) = - 1" "bin_sign (numeral k) = 0" "bin_sign (- numeral k) = -1" by (simp_all add: bin_sign_def) lemma bin_sign_rest [simp]: "bin_sign ((\k::int. k div 2) w) = bin_sign w" by (simp add: bin_sign_def) lemma bintrunc_mod2p: "(take_bit :: nat \ int \ int) n w = w mod 2 ^ n" by (fact take_bit_eq_mod) lemma sbintrunc_mod2p: "(signed_take_bit :: nat \ int \ int) n w = (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n" by (simp add: bintrunc_mod2p signed_take_bit_eq_take_bit_shift) lemma sbintrunc_eq_take_bit: \(signed_take_bit :: nat \ int \ int) n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\ by (fact signed_take_bit_eq_take_bit_shift) lemma sign_bintr: "bin_sign ((take_bit :: nat \ int \ int) n w) = 0" by (simp add: bin_sign_def) lemma bintrunc_n_0: "(take_bit :: nat \ int \ int) n 0 = 0" by (fact take_bit_of_0) lemma sbintrunc_n_0: "(signed_take_bit :: nat \ int \ int) n 0 = 0" by (fact signed_take_bit_of_0) lemma sbintrunc_n_minus1: "(signed_take_bit :: nat \ int \ int) n (- 1) = -1" by (fact signed_take_bit_of_minus_1) lemma bintrunc_Suc_numeral: "(take_bit :: nat \ int \ int) (Suc n) 1 = 1" "(take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * (take_bit :: nat \ int \ int) n (- 1)" "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (- numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (- numeral (w + Num.One))" by (simp_all add: take_bit_Suc) lemma sbintrunc_0_numeral [simp]: "(signed_take_bit :: nat \ int \ int) 0 1 = -1" "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit0 w)) = 0" "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit1 w)) = -1" "(signed_take_bit :: nat \ int \ int) 0 (- numeral (Num.Bit0 w)) = 0" "(signed_take_bit :: nat \ int \ int) 0 (- numeral (Num.Bit1 w)) = -1" by simp_all lemma sbintrunc_Suc_numeral: "(signed_take_bit :: nat \ int \ int) (Suc n) 1 = 1" "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (- numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) n (- numeral (w + Num.One))" by (simp_all add: signed_take_bit_Suc) lemma bin_sign_lem: "(bin_sign ((signed_take_bit :: nat \ int \ int) n bin) = -1) = bit bin n" by (simp add: bin_sign_def) lemma nth_bintr: "(bit :: int \ nat \ bool) ((take_bit :: nat \ int \ int) m w) n \ n < m \ (bit :: int \ nat \ bool) w n" by (fact bit_take_bit_iff) lemma nth_sbintr: "(bit :: int \ nat \ bool) ((signed_take_bit :: nat \ int \ int) m w) n = (if n < m then (bit :: int \ nat \ bool) w n else (bit :: int \ nat \ bool) w m)" by (simp add: bit_signed_take_bit_iff min_def) lemma bin_nth_Bit0: "(bit :: int \ nat \ bool) (numeral (Num.Bit0 w)) n \ (\m. n = Suc m \ (bit :: int \ nat \ bool) (numeral w) m)" using bit_double_iff [of \numeral w :: int\ n] by (auto intro: exI [of _ \n - 1\]) lemma bin_nth_Bit1: "(bit :: int \ nat \ bool) (numeral (Num.Bit1 w)) n \ n = 0 \ (\m. n = Suc m \ (bit :: int \ nat \ bool) (numeral w) m)" using even_bit_succ_iff [of \2 * numeral w :: int\ n] bit_double_iff [of \numeral w :: int\ n] by auto lemma bintrunc_bintrunc_l: "n \ m \ (take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) n w" by (simp add: min.absorb2) lemma sbintrunc_sbintrunc_l: "n \ m \ (signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) n w" by (simp add: min.absorb2) lemma bintrunc_bintrunc_ge: "n \ m \ (take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (take_bit :: nat \ int \ int) n w" by (rule bin_eqI) (auto simp: nth_bintr) lemma bintrunc_bintrunc_min [simp]: "(take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) (min m n) w" by (rule take_bit_take_bit) lemma sbintrunc_sbintrunc_min [simp]: "(signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (min m n) w" by (rule signed_take_bit_signed_take_bit) lemmas sbintrunc_Suc_Pls = signed_take_bit_Suc [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Suc_Min = signed_take_bit_Suc [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_numeral lemmas sbintrunc_Pls = signed_take_bit_0 [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Min = signed_take_bit_0 [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_0_simps = sbintrunc_Pls sbintrunc_Min lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs lemma bintrunc_minus: "0 < n \ (take_bit :: nat \ int \ int) (Suc (n - 1)) w = (take_bit :: nat \ int \ int) n w" by auto lemma sbintrunc_minus: "0 < n \ (signed_take_bit :: nat \ int \ int) (Suc (n - 1)) w = (signed_take_bit :: nat \ int \ int) n w" by auto lemmas sbintrunc_minus_simps = sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] lemma sbintrunc_BIT_I: \0 < n \ (signed_take_bit :: nat \ int \ int) (n - 1) 0 = y \ (signed_take_bit :: nat \ int \ int) n 0 = 2 * y\ by simp lemma sbintrunc_Suc_Is: \(signed_take_bit :: nat \ int \ int) n (- 1) = y \ (signed_take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * y\ by auto lemma sbintrunc_Suc_lem: "(signed_take_bit :: nat \ int \ int) (Suc n) x = y \ m = Suc n \ (signed_take_bit :: nat \ int \ int) m x = y" by (rule ssubst) lemmas sbintrunc_Suc_Ialts = sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem] lemma sbintrunc_bintrunc_lt: "m > n \ (signed_take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (signed_take_bit :: nat \ int \ int) n w" by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) lemma bintrunc_sbintrunc_le: "m \ Suc n \ (take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) m w" by (rule take_bit_signed_take_bit) lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] lemma bintrunc_sbintrunc' [simp]: "0 < n \ (take_bit :: nat \ int \ int) n ((signed_take_bit :: nat \ int \ int) (n - 1) w) = (take_bit :: nat \ int \ int) n w" by (cases n) simp_all lemma sbintrunc_bintrunc' [simp]: "0 < n \ (signed_take_bit :: nat \ int \ int) (n - 1) ((take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (n - 1) w" by (cases n) simp_all lemma bin_sbin_eq_iff: "(take_bit :: nat \ int \ int) (Suc n) x = (take_bit :: nat \ int \ int) (Suc n) y \ (signed_take_bit :: nat \ int \ int) n x = (signed_take_bit :: nat \ int \ int) n y" apply (rule iffI) apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) apply simp apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) apply simp done lemma bin_sbin_eq_iff': "0 < n \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y \ (signed_take_bit :: nat \ int \ int) (n - 1) x = (signed_take_bit :: nat \ int \ int) (n - 1) y" by (cases n) (simp_all add: bin_sbin_eq_iff) lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] (* although bintrunc_minus_simps, if added to default simpset, tends to get applied where it's not wanted in developing the theories, we get a version for when the word length is given literally *) lemmas nat_non0_gr = trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl] lemma bintrunc_numeral: "(take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc take_bit_Suc mod_2_eq_odd) lemma sbintrunc_numeral: "(signed_take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc signed_take_bit_Suc mod2_eq_if) lemma bintrunc_numeral_simps [simp]: "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" "(take_bit :: nat \ int \ int) (numeral k) 1 = 1" by (simp_all add: bintrunc_numeral) lemma sbintrunc_numeral_simps [simp]: "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" "(signed_take_bit :: nat \ int \ int) (numeral k) 1 = 1" by (simp_all add: sbintrunc_numeral) lemma no_bintr_alt1: "(take_bit :: nat \ int \ int) n = (\w. w mod 2 ^ n :: int)" by (rule ext) (rule bintrunc_mod2p) lemma range_bintrunc: "range ((take_bit :: nat \ int \ int) n) = {i. 0 \ i \ i < 2 ^ n}" by (auto simp add: take_bit_eq_mod image_iff) (metis mod_pos_pos_trivial) lemma no_sbintr_alt2: "(signed_take_bit :: nat \ int \ int) n = (\w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" by (rule ext) (simp add : sbintrunc_mod2p) lemma range_sbintrunc: "range ((signed_take_bit :: nat \ int \ int) n) = {i. - (2 ^ n) \ i \ i < 2 ^ n}" proof - have \surj (\k::int. k + 2 ^ n)\ by (rule surjI [of _ \(\k. k - 2 ^ n)\]) simp moreover have \(signed_take_bit :: nat \ int \ int) n = ((\k. k - 2 ^ n) \ take_bit (Suc n) \ (\k. k + 2 ^ n))\ by (simp add: sbintrunc_eq_take_bit fun_eq_iff) ultimately show ?thesis apply (simp only: fun.set_map range_bintrunc) apply (auto simp add: image_iff) apply presburger done qed lemma sbintrunc_inc: \k + 2 ^ Suc n \ (signed_take_bit :: nat \ int \ int) n k\ if \k < - (2 ^ n)\ using that by (fact signed_take_bit_int_greater_eq) lemma sbintrunc_dec: \(signed_take_bit :: nat \ int \ int) n k \ k - 2 ^ (Suc n)\ if \k \ 2 ^ n\ using that by (fact signed_take_bit_int_less_eq) lemma bintr_ge0: "0 \ (take_bit :: nat \ int \ int) n w" by (simp add: bintrunc_mod2p) lemma bintr_lt2p: "(take_bit :: nat \ int \ int) n w < 2 ^ n" by (simp add: bintrunc_mod2p) lemma bintr_Min: "(take_bit :: nat \ int \ int) n (- 1) = 2 ^ n - 1" by (simp add: stable_imp_take_bit_eq) lemma sbintr_ge: "- (2 ^ n) \ (signed_take_bit :: nat \ int \ int) n w" by (fact signed_take_bit_int_greater_eq_minus_exp) lemma sbintr_lt: "(signed_take_bit :: nat \ int \ int) n w < 2 ^ n" by (fact signed_take_bit_int_less_exp) lemma sign_Pls_ge_0: "bin_sign bin = 0 \ bin \ 0" for bin :: int by (simp add: bin_sign_def) lemma sign_Min_lt_0: "bin_sign bin = -1 \ bin < 0" for bin :: int by (simp add: bin_sign_def) lemma bin_rest_trunc: "(\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - 1) ((\k::int. k div 2) bin)" by (simp add: take_bit_rec [of n bin]) lemma bin_rest_power_trunc: "((\k::int. k div 2) ^^ k) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - k) (((\k::int. k div 2) ^^ k) bin)" by (induct k) (auto simp: bin_rest_trunc) lemma bin_rest_trunc_i: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) (Suc n) bin)" by (auto simp add: take_bit_Suc) lemma bin_rest_strunc: "(\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) (Suc n) bin) = (signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin)" by (simp add: signed_take_bit_Suc) lemma bintrunc_rest [simp]: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin)" by (induct n arbitrary: bin) (simp_all add: take_bit_Suc) lemma sbintrunc_rest [simp]: "(signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) n bin)" by (induct n arbitrary: bin) (simp_all add: signed_take_bit_Suc mod2_eq_if) lemma bintrunc_rest': "(take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n" by (rule ext) auto lemma sbintrunc_rest': "(signed_take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n" by (rule ext) auto lemma rco_lem: "f \ g \ f = g \ f \ f \ (g \ f) ^^ n = g ^^ n \ f" apply (rule ext) apply (induct_tac n) apply (simp_all (no_asm)) apply (drule fun_cong) apply (unfold o_def) apply (erule trans) apply simp done lemmas rco_bintr = bintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] lemmas rco_sbintr = sbintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] subsection \Splitting and concatenation\ definition bin_split :: \nat \ int \ int \ int\ where [simp]: \bin_split n k = (drop_bit n k, take_bit n k)\ lemma [code]: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (w div 2) in (w1, of_bool (odd w) + 2 * w2))" "bin_split 0 w = (w, 0)" by (simp_all add: drop_bit_Suc take_bit_Suc mod_2_eq_odd) lemma bin_cat_eq_push_bit_add_take_bit: \concat_bit n l k = push_bit n k + take_bit n l\ by (simp add: concat_bit_eq) lemma bin_sign_cat: "bin_sign ((\k n l. concat_bit n l k) x n y) = bin_sign x" proof - have \0 \ x\ if \0 \ x * 2 ^ n + y mod 2 ^ n\ proof - have \y mod 2 ^ n < 2 ^ n\ using pos_mod_bound [of \2 ^ n\ y] by simp then have \\ y mod 2 ^ n \ 2 ^ n\ by (simp add: less_le) with that have \x \ - 1\ by auto have *: \- 1 \ (- (y mod 2 ^ n)) div 2 ^ n\ by (simp add: zdiv_zminus1_eq_if) from that have \- (y mod 2 ^ n) \ x * 2 ^ n\ by simp then have \(- (y mod 2 ^ n)) div 2 ^ n \ (x * 2 ^ n) div 2 ^ n\ using zdiv_mono1 zero_less_numeral zero_less_power by blast with * have \- 1 \ x * 2 ^ n div 2 ^ n\ by simp with \x \ - 1\ show ?thesis by simp qed then show ?thesis by (simp add: bin_sign_def not_le not_less bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult take_bit_eq_mod) qed lemma bin_cat_assoc: "(\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) x m y) n z = (\k n l. concat_bit n l k) x (m + n) ((\k n l. concat_bit n l k) y n z)" by (fact concat_bit_assoc) lemma bin_cat_assoc_sym: "(\k n l. concat_bit n l k) x m ((\k n l. concat_bit n l k) y n z) = (\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) x (m - n) y) (min m n) z" by (fact concat_bit_assoc_sym) definition bin_rcat :: \nat \ int list \ int\ where \bin_rcat n = horner_sum (take_bit n) (2 ^ n) \ rev\ lemma bin_rcat_eq_foldl: \bin_rcat n = foldl (\u v. (\k n l. concat_bit n l k) u n v) 0\ proof fix ks :: \int list\ show \bin_rcat n ks = foldl (\u v. (\k n l. concat_bit n l k) u n v) 0 ks\ by (induction ks rule: rev_induct) (simp_all add: bin_rcat_def concat_bit_eq push_bit_eq_mult) qed fun bin_rsplit_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit_aux n (m - n) a (b # bs))" definition bin_rsplit :: "nat \ nat \ int \ int list" where "bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" fun bin_rsplitl_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplitl_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split (min m n) c in bin_rsplitl_aux n (m - n) a (b # bs))" definition bin_rsplitl :: "nat \ nat \ int \ int list" where "bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" declare bin_rsplit_aux.simps [simp del] declare bin_rsplitl_aux.simps [simp del] lemma bin_nth_cat: "(bit :: int \ nat \ bool) ((\k n l. concat_bit n l k) x k y) n = (if n < k then (bit :: int \ nat \ bool) y n else (bit :: int \ nat \ bool) x (n - k))" by (simp add: bit_concat_bit_iff) lemma bin_nth_drop_bit_iff: \(bit :: int \ nat \ bool) (drop_bit n c) k \ (bit :: int \ nat \ bool) c (n + k)\ by (simp add: bit_drop_bit_eq) lemma bin_nth_take_bit_iff: \(bit :: int \ nat \ bool) (take_bit n c) k \ k < n \ (bit :: int \ nat \ bool) c k\ by (fact bit_take_bit_iff) lemma bin_nth_split: "bin_split n c = (a, b) \ (\k. (bit :: int \ nat \ bool) a k = (bit :: int \ nat \ bool) c (n + k)) \ (\k. (bit :: int \ nat \ bool) b k = (k < n \ (bit :: int \ nat \ bool) c k))" by (auto simp add: bin_nth_drop_bit_iff bin_nth_take_bit_iff) lemma bin_cat_zero [simp]: "(\k n l. concat_bit n l k) 0 n w = (take_bit :: nat \ int \ int) n w" by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bintr_cat1: "(take_bit :: nat \ int \ int) (k + n) ((\k n l. concat_bit n l k) a n b) = (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) k a) n b" by (metis bin_cat_assoc bin_cat_zero) lemma bintr_cat: "(take_bit :: nat \ int \ int) m ((\k n l. concat_bit n l k) a n b) = (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) (m - n) a) n ((take_bit :: nat \ int \ int) (min m n) b)" by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr) lemma bintr_cat_same [simp]: "(take_bit :: nat \ int \ int) n ((\k n l. concat_bit n l k) a n b) = (take_bit :: nat \ int \ int) n b" by (auto simp add : bintr_cat) lemma cat_bintr [simp]: "(\k n l. concat_bit n l k) a n ((take_bit :: nat \ int \ int) n b) = (\k n l. concat_bit n l k) a n b" by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma split_bintrunc: "bin_split n c = (a, b) \ b = (take_bit :: nat \ int \ int) n c" by simp lemma bin_cat_split: "bin_split n w = (u, v) \ w = (\k n l. concat_bit n l k) u n v" by (auto simp add: bin_cat_eq_push_bit_add_take_bit bits_ident) lemma drop_bit_bin_cat_eq: \drop_bit n ((\k n l. concat_bit n l k) v n w) = v\ by (rule bit_eqI) (simp add: bit_drop_bit_eq bit_concat_bit_iff) lemma take_bit_bin_cat_eq: \take_bit n ((\k n l. concat_bit n l k) v n w) = take_bit n w\ by (rule bit_eqI) (simp add: bit_concat_bit_iff) lemma bin_split_cat: "bin_split n ((\k n l. concat_bit n l k) v n w) = (v, (take_bit :: nat \ int \ int) n w)" by (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) lemma bin_split_zero [simp]: "bin_split n 0 = (0, 0)" by simp lemma bin_split_minus1 [simp]: "bin_split n (- 1) = (- 1, (take_bit :: nat \ int \ int) n (- 1))" by simp lemma bin_split_trunc: "bin_split (min m n) c = (a, b) \ bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (m - n) a, b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done lemma bin_split_trunc1: "bin_split n c = (a, b) \ bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (m - n) a, (take_bit :: nat \ int \ int) m b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done lemma bin_cat_num: "(\k n l. concat_bit n l k) a n b = a * 2 ^ n + (take_bit :: nat \ int \ int) n b" by (simp add: bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult) lemma bin_split_num: "bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" by (simp add: drop_bit_eq_div take_bit_eq_mod) lemmas bin_rsplit_aux_simps = bin_rsplit_aux.simps bin_rsplitl_aux.simps lemmas rsplit_aux_simps = bin_rsplit_aux_simps lemmas th_if_simp1 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct1, THEN mp] for l lemmas th_if_simp2 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct2, THEN mp] for l lemmas rsplit_aux_simp1s = rsplit_aux_simps [THEN th_if_simp1] lemmas rsplit_aux_simp2ls = rsplit_aux_simps [THEN th_if_simp2] \ \these safe to \[simp add]\ as require calculating \m - n\\ lemmas bin_rsplit_aux_simp2s [simp] = rsplit_aux_simp2ls [unfolded Let_def] lemmas rbscl = bin_rsplit_aux_simp2s (2) lemmas rsplit_aux_0_simps [simp] = rsplit_aux_simp1s [OF disjI1] rsplit_aux_simp1s [OF disjI2] lemma bin_rsplit_aux_append: "bin_rsplit_aux n m c (bs @ cs) = bin_rsplit_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp split: prod.split) done lemma bin_rsplitl_aux_append: "bin_rsplitl_aux n m c (bs @ cs) = bin_rsplitl_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplitl_aux.induct) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplitl_aux.simps) apply (clarsimp split: prod.split) done lemmas rsplit_aux_apps [where bs = "[]"] = bin_rsplit_aux_append bin_rsplitl_aux_append lemmas rsplit_def_auxs = bin_rsplit_def bin_rsplitl_def lemmas rsplit_aux_alts = rsplit_aux_apps [unfolded append_Nil rsplit_def_auxs [symmetric]] lemma bin_split_minus: "0 < n \ bin_split (Suc (n - 1)) w = bin_split n w" by auto lemma bin_split_pred_simp [simp]: "(0::nat) < numeral bin \ bin_split (numeral bin) w = (let (w1, w2) = bin_split (numeral bin - 1) ((\k::int. k div 2) w) in (w1, of_bool (odd w) + 2 * w2))" by (simp add: take_bit_rec drop_bit_rec mod_2_eq_odd) lemma bin_rsplit_aux_simp_alt: "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit n (m - n, a) @ b # bs)" apply (simp add: bin_rsplit_aux.simps [of n m c bs]) apply (subst rsplit_aux_alts) apply (simp add: bin_rsplit_def) done lemmas bin_rsplit_simp_alt = trans [OF bin_rsplit_def bin_rsplit_aux_simp_alt] lemmas bthrs = bin_rsplit_simp_alt [THEN [2] trans] lemma bin_rsplit_size_sign' [rule_format]: "n > 0 \ rev sw = bin_rsplit n (nw, w) \ \v\set sw. (take_bit :: nat \ int \ int) n v = v" apply (induct sw arbitrary: nw w) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply clarify apply simp done lemmas bin_rsplit_size_sign = bin_rsplit_size_sign' [OF asm_rl rev_rev_ident [THEN trans] set_rev [THEN equalityD2 [THEN subsetD]]] lemma bin_nth_rsplit [rule_format] : "n > 0 \ m < n \ \w k nw. rev sw = bin_rsplit n (nw, w) \ k < size sw \ (bit :: int \ nat \ bool) (sw ! k) m = (bit :: int \ nat \ bool) w (k * n + m)" apply (induct sw) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply (erule allE, erule impE, erule exI) apply (case_tac k) apply clarsimp prefer 2 apply clarsimp apply (erule allE) apply (erule (1) impE) apply (simp add: bit_drop_bit_eq ac_simps) apply (simp add: bit_take_bit_iff ac_simps) done lemma bin_rsplit_all: "0 < nw \ nw \ n \ bin_rsplit n (nw, w) = [(take_bit :: nat \ int \ int) n w]" by (auto simp: bin_rsplit_def rsplit_aux_simp2ls split: prod.split dest!: split_bintrunc) lemma bin_rsplit_l [rule_format]: "\bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, (take_bit :: nat \ int \ int) m bin)" apply (rule_tac a = "m" in wf_less_than [THEN wf_induct]) apply (simp (no_asm) add: bin_rsplitl_def bin_rsplit_def) apply (rule allI) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (simp add: ac_simps) apply (subst rsplit_aux_alts(1)) apply (subst rsplit_aux_alts(2)) apply clarsimp unfolding bin_rsplit_def bin_rsplitl_def apply (simp add: drop_bit_take_bit) apply (case_tac \x < n\) apply (simp_all add: not_less min_def) done lemma bin_rsplit_rcat [rule_format]: "n > 0 \ bin_rsplit n (n * size ws, bin_rcat n ws) = map ((take_bit :: nat \ int \ int) n) ws" apply (unfold bin_rsplit_def bin_rcat_eq_foldl) apply (rule_tac xs = ws in rev_induct) apply clarsimp apply clarsimp apply (subst rsplit_aux_alts) apply (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) done lemma bin_rsplit_aux_len_le [rule_format] : "\ws m. n \ 0 \ ws = bin_rsplit_aux n nw w bs \ length ws \ m \ nw + length bs * n \ m * n" proof - have *: R if d: "i \ j \ m < j'" and R1: "i * k \ j * k \ R" and R2: "Suc m * k' \ j' * k' \ R" for i j j' k k' m :: nat and R using d apply safe apply (rule R1, erule mult_le_mono1) apply (rule R2, erule Suc_le_eq [THEN iffD2 [THEN mult_le_mono1]]) done have **: "0 < sc \ sc - n + (n + lb * n) \ m * n \ sc + lb * n \ m * n" for sc m n lb :: nat apply safe apply arith apply (case_tac "sc \ n") apply arith apply (insert linorder_le_less_linear [of m lb]) apply (erule_tac k=n and k'=n in *) apply arith apply simp done show ?thesis apply (induct n nw w bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (simp add: ** Let_def split: prod.split) done qed lemma bin_rsplit_len_le: "n \ 0 \ ws = bin_rsplit n (nw, w) \ length ws \ m \ nw \ m * n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len_le) lemma bin_rsplit_aux_len: "n \ 0 \ length (bin_rsplit_aux n nw w cs) = (nw + n - 1) div n + length cs" apply (induct n nw w cs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (erule thin_rl) apply (case_tac m) apply simp apply (case_tac "m \ n") apply (auto simp add: div_add_self2) done lemma bin_rsplit_len: "n \ 0 \ length (bin_rsplit n (nw, w)) = (nw + n - 1) div n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len) lemma bin_rsplit_aux_len_indep: "n \ 0 \ length bs = length cs \ length (bin_rsplit_aux n nw v bs) = length (bin_rsplit_aux n nw w cs)" proof (induct n nw w cs arbitrary: v bs rule: bin_rsplit_aux.induct) case (1 n m w cs v bs) show ?case proof (cases "m = 0") case True with \length bs = length cs\ show ?thesis by simp next case False from "1.hyps" [of \bin_split n w\ \drop_bit n w\ \take_bit n w\] \m \ 0\ \n \ 0\ have hyp: "\v bs. length bs = Suc (length cs) \ length (bin_rsplit_aux n (m - n) v bs) = length (bin_rsplit_aux n (m - n) (drop_bit n w) (take_bit n w # cs))" using bin_rsplit_aux_len by fastforce from \length bs = length cs\ \n \ 0\ show ?thesis by (auto simp add: bin_rsplit_aux_simp_alt Let_def bin_rsplit_len split: prod.split) qed qed lemma bin_rsplit_len_indep: "n \ 0 \ length (bin_rsplit n (nw, v)) = length (bin_rsplit n (nw, w))" apply (unfold bin_rsplit_def) apply (simp (no_asm)) apply (erule bin_rsplit_aux_len_indep) apply (rule refl) done subsection \Logical operations\ abbreviation (input) bin_sc :: \nat \ bool \ int \ int\ where \bin_sc n b i \ set_bit i n b\ lemma bin_sc_0 [simp]: "bin_sc 0 b w = of_bool b + 2 * (\k::int. k div 2) w" by (simp add: set_bit_int_def) lemma bin_sc_Suc [simp]: "bin_sc (Suc n) b w = of_bool (odd w) + 2 * bin_sc n b (w div 2)" by (simp add: set_bit_int_def set_bit_Suc unset_bit_Suc bin_last_def) lemma bin_nth_sc [bit_simps]: "bit (bin_sc n b w) n \ b" by (simp add: bit_simps) lemma bin_sc_sc_same [simp]: "bin_sc n c (bin_sc n b w) = bin_sc n c w" by (induction n arbitrary: w) (simp_all add: bit_Suc) lemma bin_sc_sc_diff: "m \ n \ bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" apply (induct n arbitrary: w m) apply (case_tac [!] m) apply auto done lemma bin_nth_sc_gen: "(bit :: int \ nat \ bool) (bin_sc n b w) m = (if m = n then b else (bit :: int \ nat \ bool) w m)" apply (induct n arbitrary: w m) apply (case_tac m; simp add: bit_Suc) apply (case_tac m; simp add: bit_Suc) done lemma bin_sc_eq: \bin_sc n False = unset_bit n\ \bin_sc n True = Bit_Operations.set_bit n\ apply (simp_all add: fun_eq_iff bit_eq_iff) apply (simp_all add: bit_simps bin_nth_sc_gen) done lemma bin_sc_nth [simp]: "bin_sc n ((bit :: int \ nat \ bool) w n) w = w" by (rule bit_eqI) (simp add: bin_nth_sc_gen) lemma bin_sign_sc [simp]: "bin_sign (bin_sc n b w) = bin_sign w" proof (induction n arbitrary: w) case 0 then show ?case by (auto simp add: bin_sign_def) (use bin_rest_ge_0 in fastforce) next case (Suc n) from Suc [of \w div 2\] show ?case by (auto simp add: bin_sign_def split: if_splits) qed lemma bin_sc_bintr [simp]: "(take_bit :: nat \ int \ int) m (bin_sc n x ((take_bit :: nat \ int \ int) m w)) = (take_bit :: nat \ int \ int) m (bin_sc n x w)" apply (rule bit_eqI) apply (cases x) apply (auto simp add: bit_simps bin_sc_eq) done lemma bin_clr_le: "bin_sc n False w \ w" by (simp add: set_bit_int_def unset_bit_less_eq) lemma bin_set_ge: "bin_sc n True w \ w" by (simp add: set_bit_int_def set_bit_greater_eq) lemma bintr_bin_clr_le: "(take_bit :: nat \ int \ int) n (bin_sc m False w) \ (take_bit :: nat \ int \ int) n w" by (simp add: set_bit_int_def take_bit_unset_bit_eq unset_bit_less_eq) lemma bintr_bin_set_ge: "(take_bit :: nat \ int \ int) n (bin_sc m True w) \ (take_bit :: nat \ int \ int) n w" by (simp add: set_bit_int_def take_bit_set_bit_eq set_bit_greater_eq) lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0" by (induct n) auto lemma bin_sc_TM [simp]: "bin_sc n True (- 1) = - 1" by (induct n) auto lemmas bin_sc_simps = bin_sc_0 bin_sc_Suc bin_sc_TM bin_sc_FP lemma bin_sc_minus: "0 < n \ bin_sc (Suc (n - 1)) b w = bin_sc n b w" by auto lemmas bin_sc_Suc_minus = trans [OF bin_sc_minus [symmetric] bin_sc_Suc] lemma bin_sc_numeral [simp]: "bin_sc (numeral k) b w = of_bool (odd w) + 2 * bin_sc (pred_numeral k) b (w div 2)" by (simp add: numeral_eq_Suc) lemmas bin_sc_minus_simps = bin_sc_simps (2,3,4) [THEN [2] trans, OF bin_sc_minus [THEN sym]] lemma int_set_bit_0 [simp]: fixes x :: int shows "set_bit x 0 b = of_bool b + 2 * (x div 2)" by (fact bin_sc_0) lemma int_set_bit_Suc: fixes x :: int shows "set_bit x (Suc n) b = of_bool (odd x) + 2 * set_bit (x div 2) n b" by (fact bin_sc_Suc) lemma bin_last_set_bit: "odd (set_bit x n b :: int) = (if n > 0 then odd x else b)" by (cases n) (simp_all add: int_set_bit_Suc) lemma bin_rest_set_bit: "(set_bit x n b :: int) div 2 = (if n > 0 then set_bit (x div 2) (n - 1) b else x div 2)" by (cases n) (simp_all add: int_set_bit_Suc) lemma int_set_bit_numeral: fixes x :: int shows "set_bit x (numeral w) b = of_bool (odd x) + 2 * set_bit (x div 2) (pred_numeral w) b" by (fact bin_sc_numeral) lemmas int_set_bit_numerals [simp] = int_set_bit_numeral[where x="numeral w'"] int_set_bit_numeral[where x="- numeral w'"] int_set_bit_numeral[where x="Numeral1"] int_set_bit_numeral[where x="1"] int_set_bit_numeral[where x="0"] int_set_bit_Suc[where x="numeral w'"] int_set_bit_Suc[where x="- numeral w'"] int_set_bit_Suc[where x="Numeral1"] int_set_bit_Suc[where x="1"] int_set_bit_Suc[where x="0"] for w' lemma msb_set_bit [simp]: "msb (set_bit (x :: int) n b) \ msb x" by (simp add: msb_int_def set_bit_int_def) lemma word_set_bit_def: \set_bit a n x = word_of_int (bin_sc n x (uint a))\ apply (rule bit_word_eqI) apply (cases x) apply (simp_all add: bit_simps bin_sc_eq) done lemma set_bit_word_of_int: "set_bit (word_of_int x) n b = word_of_int (bin_sc n b x)" unfolding word_set_bit_def by (rule word_eqI) (simp add: word_size bin_nth_sc_gen nth_bintr bit_simps) lemma word_set_numeral [simp]: "set_bit (numeral bin::'a::len word) n b = word_of_int (bin_sc n b (numeral bin))" unfolding word_numeral_alt by (rule set_bit_word_of_int) lemma word_set_neg_numeral [simp]: "set_bit (- numeral bin::'a::len word) n b = word_of_int (bin_sc n b (- numeral bin))" unfolding word_neg_numeral_alt by (rule set_bit_word_of_int) lemma word_set_bit_0 [simp]: "set_bit 0 n b = word_of_int (bin_sc n b 0)" unfolding word_0_wi by (rule set_bit_word_of_int) lemma word_set_bit_1 [simp]: "set_bit 1 n b = word_of_int (bin_sc n b 1)" unfolding word_1_wi by (rule set_bit_word_of_int) lemma shiftl_int_def: "push_bit n x = x * 2 ^ n" for x :: int by (fact push_bit_eq_mult) lemma shiftr_int_def: "drop_bit n x = x div 2 ^ n" for x :: int by (fact drop_bit_eq_div) subsubsection \Basic simplification rules\ +context + includes bit_operations_syntax +begin + lemmas int_not_def = not_int_def lemma int_not_simps: "NOT (0::int) = -1" "NOT (1::int) = -2" "NOT (- 1::int) = 0" "NOT (numeral w::int) = - numeral (w + Num.One)" "NOT (- numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)" "NOT (- numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)" by (simp_all add: not_int_def) lemma int_not_not: "NOT (NOT x) = x" for x :: int by (fact bit.double_compl) lemma int_and_0 [simp]: "0 AND x = 0" for x :: int by (fact bit.conj_zero_left) lemma int_and_m1 [simp]: "-1 AND x = x" for x :: int by (fact bit.conj_one_left) lemma int_or_zero [simp]: "0 OR x = x" for x :: int by (fact bit.disj_zero_left) lemma int_or_minus1 [simp]: "-1 OR x = -1" for x :: int by (fact bit.disj_one_left) lemma int_xor_zero [simp]: "0 XOR x = x" for x :: int by (fact bit.xor_zero_left) subsubsection \Binary destructors\ lemma bin_rest_NOT [simp]: "(\k::int. k div 2) (NOT x) = NOT ((\k::int. k div 2) x)" by (fact not_int_div_2) lemma bin_last_NOT [simp]: "(odd :: int \ bool) (NOT x) \ \ (odd :: int \ bool) x" by simp lemma bin_rest_AND [simp]: "(\k::int. k div 2) (x AND y) = (\k::int. k div 2) x AND (\k::int. k div 2) y" by (subst and_int_rec) auto lemma bin_last_AND [simp]: "(odd :: int \ bool) (x AND y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst and_int_rec) auto lemma bin_rest_OR [simp]: "(\k::int. k div 2) (x OR y) = (\k::int. k div 2) x OR (\k::int. k div 2) y" by (subst or_int_rec) auto lemma bin_last_OR [simp]: "(odd :: int \ bool) (x OR y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst or_int_rec) auto lemma bin_rest_XOR [simp]: "(\k::int. k div 2) (x XOR y) = (\k::int. k div 2) x XOR (\k::int. k div 2) y" by (subst xor_int_rec) auto lemma bin_last_XOR [simp]: "(odd :: int \ bool) (x XOR y) \ ((odd :: int \ bool) x \ (odd :: int \ bool) y) \ \ ((odd :: int \ bool) x \ (odd :: int \ bool) y)" by (subst xor_int_rec) auto lemma bin_nth_ops: "\x y. (bit :: int \ nat \ bool) (x AND y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x y. (bit :: int \ nat \ bool) (x OR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x y. (bit :: int \ nat \ bool) (x XOR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x. (bit :: int \ nat \ bool) (NOT x) n \ \ (bit :: int \ nat \ bool) x n" by (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff) subsubsection \Derived properties\ lemma int_xor_minus1 [simp]: "-1 XOR x = NOT x" for x :: int by (fact bit.xor_one_left) lemma int_xor_extra_simps [simp]: "w XOR 0 = w" "w XOR -1 = NOT w" for w :: int by simp_all lemma int_or_extra_simps [simp]: "w OR 0 = w" "w OR -1 = -1" for w :: int by simp_all lemma int_and_extra_simps [simp]: "w AND 0 = 0" "w AND -1 = w" for w :: int by simp_all text \Commutativity of the above.\ lemma bin_ops_comm: fixes x y :: int shows int_and_comm: "x AND y = y AND x" and int_or_comm: "x OR y = y OR x" and int_xor_comm: "x XOR y = y XOR x" by (simp_all add: ac_simps) lemma bin_ops_same [simp]: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: int by simp_all lemmas bin_log_esimps = int_and_extra_simps int_or_extra_simps int_xor_extra_simps int_and_0 int_and_m1 int_or_zero int_or_minus1 int_xor_zero int_xor_minus1 subsubsection \Basic properties of logical (bit-wise) operations\ lemma bbw_ao_absorb: "x AND (y OR x) = x \ x OR (y AND x) = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_absorbs_other: "x AND (x OR y) = x \ (y AND x) OR x = x" "(y OR x) AND x = x \ x OR (x AND y) = x" "(x OR y) AND x = x \ (x AND y) OR x = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other lemma int_xor_not: "(NOT x) XOR y = NOT (x XOR y) \ x XOR (NOT y) = NOT (x XOR y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_and_assoc: "(x AND y) AND z = x AND (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_or_assoc: "(x OR y) OR z = x OR (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_xor_assoc: "(x XOR y) XOR z = x XOR (y XOR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc (* BH: Why are these declared as simp rules??? *) lemma bbw_lcs [simp]: "y AND (x AND z) = x AND (y AND z)" "y OR (x OR z) = x OR (y OR z)" "y XOR (x XOR z) = x XOR (y XOR z)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_not_dist: "NOT (x OR y) = (NOT x) AND (NOT y)" "NOT (x AND y) = (NOT x) OR (NOT y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_oa_dist: "(x AND y) OR z = (x OR z) AND (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_dist: "(x OR y) AND z = (x AND z) OR (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) subsubsection \Simplification with numerals\ text \Cases for \0\ and \-1\ are already covered by other simp rules.\ lemma bin_rest_neg_numeral_BitM [simp]: "(\k::int. k div 2) (- numeral (Num.BitM w)) = - numeral w" by simp lemma bin_last_neg_numeral_BitM [simp]: "(odd :: int \ bool) (- numeral (Num.BitM w))" by simp subsubsection \Interactions with arithmetic\ lemma le_int_or: "bin_sign y = 0 \ x \ x OR y" for x y :: int by (simp add: bin_sign_def or_greater_eq split: if_splits) lemmas int_and_le = xtrans(3) [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] text \Interaction between bit-wise and arithmetic: good example of \bin_induction\.\ lemma bin_add_not: "x + NOT x = (-1::int)" by (simp add: not_int_def) lemma AND_mod: "x AND (2 ^ n - 1) = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask mask_eq_exp_minus_1) subsubsection \Truncating results of bit-wise operations\ lemma bin_trunc_ao: "(take_bit :: nat \ int \ int) n x AND (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) n (x AND y)" "(take_bit :: nat \ int \ int) n x OR (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) n (x OR y)" by simp_all lemma bin_trunc_xor: "(take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) n x XOR (take_bit :: nat \ int \ int) n y) = (take_bit :: nat \ int \ int) n (x XOR y)" by simp lemma bin_trunc_not: "(take_bit :: nat \ int \ int) n (NOT ((take_bit :: nat \ int \ int) n x)) = (take_bit :: nat \ int \ int) n (NOT x)" by (fact take_bit_not_take_bit) text \Want theorems of the form of \bin_trunc_xor\.\ lemma bintr_bintr_i: "x = (take_bit :: nat \ int \ int) n y \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y" by auto lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i] subsubsection \More lemmas\ lemma not_int_cmp_0 [simp]: fixes i :: int shows "0 < NOT i \ i < -1" "0 \ NOT i \ i < 0" "NOT i < 0 \ i \ 0" "NOT i \ 0 \ i \ -1" by(simp_all add: int_not_def) arith+ lemma bbw_ao_dist2: "(x :: int) AND (y OR z) = x AND y OR x AND z" by (fact bit.conj_disj_distrib) lemmas int_and_ac = bbw_lcs(1) int_and_comm int_and_assoc lemma int_nand_same [simp]: fixes x :: int shows "x AND NOT x = 0" by simp lemma int_nand_same_middle: fixes x :: int shows "x AND y AND NOT x = 0" by (simp add: bit_eq_iff bit_and_iff bit_not_iff) lemma and_xor_dist: fixes x :: int shows "x AND (y XOR z) = (x AND y) XOR (x AND z)" by (fact bit.conj_xor_distrib) lemma int_and_lt0 [simp]: \x AND y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact and_negative_int_iff) lemma int_and_ge0 [simp]: \x AND y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact and_nonnegative_int_iff) lemma int_and_1: fixes x :: int shows "x AND 1 = x mod 2" by (fact and_one_eq) lemma int_1_and: fixes x :: int shows "1 AND x = x mod 2" by (fact one_and_eq) lemma int_or_lt0 [simp]: \x OR y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact or_negative_int_iff) lemma int_or_ge0 [simp]: \x OR y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact or_nonnegative_int_iff) lemma int_xor_lt0 [simp]: \x XOR y < 0 \ (x < 0) \ (y < 0)\ for x y :: int by (fact xor_negative_int_iff) lemma int_xor_ge0 [simp]: \x XOR y \ 0 \ (x \ 0 \ y \ 0)\ for x y :: int by (fact xor_nonnegative_int_iff) lemma even_conv_AND: \even i \ i AND 1 = 0\ for i :: int by (simp add: and_one_eq mod2_eq_if) lemma bin_last_conv_AND: "(odd :: int \ bool) i \ i AND 1 \ 0" by (simp add: and_one_eq mod2_eq_if) lemma bitval_bin_last: "of_bool ((odd :: int \ bool) i) = i AND 1" by (simp add: and_one_eq mod2_eq_if) lemma bin_sign_and: "bin_sign (i AND j) = - (bin_sign i * bin_sign j)" by(simp add: bin_sign_def) lemma int_not_neg_numeral: "NOT (- numeral n) = (Num.sub n num.One :: int)" by(simp add: int_not_def) lemma int_neg_numeral_pOne_conv_not: "- numeral (n + num.One) = (NOT (numeral n) :: int)" by(simp add: int_not_def) subsection \Setting and clearing bits\ lemma int_shiftl_BIT: fixes x :: int shows int_shiftl0: "push_bit 0 x = x" and int_shiftl_Suc: "push_bit (Suc n) x = 2 * push_bit n x" by (auto simp add: shiftl_int_def) lemma int_0_shiftl: "push_bit n 0 = (0 :: int)" by (fact push_bit_of_0) lemma bin_last_shiftl: "odd (push_bit n x) \ n = 0 \ (odd :: int \ bool) x" by simp lemma bin_rest_shiftl: "(\k::int. k div 2) (push_bit n x) = (if n > 0 then push_bit (n - 1) x else (\k::int. k div 2) x)" by (cases n) (simp_all add: push_bit_eq_mult) lemma bin_nth_shiftl: "(bit :: int \ nat \ bool) (push_bit n x) m \ n \ m \ (bit :: int \ nat \ bool) x (m - n)" by (fact bit_push_bit_iff_int) lemma bin_last_shiftr: "odd (drop_bit n x) \ bit x n" for x :: int by (simp add: bit_iff_odd_drop_bit) lemma bin_rest_shiftr: "(\k::int. k div 2) (drop_bit n x) = drop_bit (Suc n) x" by (simp add: drop_bit_Suc drop_bit_half) lemma bin_nth_shiftr: "(bit :: int \ nat \ bool) (drop_bit n x) m = (bit :: int \ nat \ bool) x (n + m)" by (simp add: bit_simps) lemma bin_nth_conv_AND: fixes x :: int shows "(bit :: int \ nat \ bool) x n \ x AND (push_bit n 1) \ 0" by (fact bit_iff_and_push_bit_not_eq_0) lemma int_shiftl_numeral [simp]: "push_bit (numeral w') (numeral w :: int) = push_bit (pred_numeral w') (numeral (num.Bit0 w))" "push_bit (numeral w') (- numeral w :: int) = push_bit (pred_numeral w') (- numeral (num.Bit0 w))" by(simp_all add: numeral_eq_Suc shiftl_int_def) (metis add_One mult_inc semiring_norm(11) semiring_norm(13) semiring_norm(2) semiring_norm(6) semiring_norm(87))+ lemma int_shiftl_One_numeral [simp]: "push_bit (numeral w) (1::int) = push_bit (pred_numeral w) 2" using int_shiftl_numeral [of Num.One w] by (simp add: numeral_eq_Suc) lemma shiftl_ge_0: fixes i :: int shows "push_bit n i \ 0 \ i \ 0" by (fact push_bit_nonnegative_int_iff) lemma shiftl_lt_0: fixes i :: int shows "push_bit n i < 0 \ i < 0" by (fact push_bit_negative_int_iff) lemma int_shiftl_test_bit: "bit (push_bit i n :: int) m \ m \ i \ bit n (m - i)" by (fact bit_push_bit_iff_int) lemma int_0shiftr: "drop_bit x (0 :: int) = 0" by (fact drop_bit_of_0) lemma int_minus1_shiftr: "drop_bit x (-1 :: int) = -1" by (fact drop_bit_minus_one) lemma int_shiftr_ge_0: fixes i :: int shows "drop_bit n i \ 0 \ i \ 0" by (fact drop_bit_nonnegative_int_iff) lemma int_shiftr_lt_0 [simp]: fixes i :: int shows "drop_bit n i < 0 \ i < 0" by (fact drop_bit_negative_int_iff) lemma int_shiftr_numeral [simp]: "drop_bit (numeral w') (1 :: int) = 0" "drop_bit (numeral w') (numeral num.One :: int) = 0" "drop_bit (numeral w') (numeral (num.Bit0 w) :: int) = drop_bit (pred_numeral w') (numeral w)" "drop_bit (numeral w') (numeral (num.Bit1 w) :: int) = drop_bit (pred_numeral w') (numeral w)" "drop_bit (numeral w') (- numeral (num.Bit0 w) :: int) = drop_bit (pred_numeral w') (- numeral w)" "drop_bit (numeral w') (- numeral (num.Bit1 w) :: int) = drop_bit (pred_numeral w') (- numeral (Num.inc w))" by (simp_all add: numeral_eq_Suc add_One drop_bit_Suc) lemma int_shiftr_numeral_Suc0 [simp]: "drop_bit (Suc 0) (1 :: int) = 0" "drop_bit (Suc 0) (numeral num.One :: int) = 0" "drop_bit (Suc 0) (numeral (num.Bit0 w) :: int) = numeral w" "drop_bit (Suc 0) (numeral (num.Bit1 w) :: int) = numeral w" "drop_bit (Suc 0) (- numeral (num.Bit0 w) :: int) = - numeral w" "drop_bit (Suc 0) (- numeral (num.Bit1 w) :: int) = - numeral (Num.inc w)" by (simp_all add: drop_bit_Suc add_One) lemma bin_nth_minus_p2: assumes sign: "bin_sign x = 0" and y: "y = push_bit n 1" and m: "m < n" and x: "x < y" shows "bit (x - y) m = bit x m" proof - from sign y x have \x \ 0\ and \y = 2 ^ n\ and \x < 2 ^ n\ by (simp_all add: bin_sign_def push_bit_eq_mult split: if_splits) from \0 \ x\ \x < 2 ^ n\ \m < n\ have \bit x m \ bit (x - 2 ^ n) m\ proof (induction m arbitrary: x n) case 0 then show ?case by simp next case (Suc m) moreover define q where \q = n - 1\ ultimately have n: \n = Suc q\ by simp have \(x - 2 ^ Suc q) div 2 = x div 2 - 2 ^ q\ by simp moreover from Suc.IH [of \x div 2\ q] Suc.prems have \bit (x div 2) m \ bit (x div 2 - 2 ^ q) m\ by (simp add: n) ultimately show ?case by (simp add: bit_Suc n) qed with \y = 2 ^ n\ show ?thesis by simp qed lemma bin_clr_conv_NAND: "bin_sc n False i = i AND NOT (push_bit n 1)" by (rule bit_eqI) (auto simp add: bin_sc_eq bit_simps) lemma bin_set_conv_OR: "bin_sc n True i = i OR (push_bit n 1)" by (rule bit_eqI) (auto simp add: bin_sc_eq bit_simps) +end + subsection \More lemmas on words\ lemma msb_conv_bin_sign: "msb x \ bin_sign x = -1" by (simp add: bin_sign_def not_le msb_int_def) lemma msb_bin_sc [simp]: "msb (bin_sc n b x) \ msb x" by (simp add: msb_conv_bin_sign) lemma msb_word_def: \msb a \ bin_sign (signed_take_bit (LENGTH('a) - 1) (uint a)) = - 1\ for a :: \'a::len word\ by (simp add: bin_sign_def bit_simps msb_word_iff_bit) lemma word_msb_def: "msb a \ bin_sign (sint a) = - 1" by (simp add: msb_word_def sint_uint) lemma word_rcat_eq: \word_rcat ws = word_of_int (bin_rcat (LENGTH('a::len)) (map uint ws))\ for ws :: \'a::len word list\ apply (simp add: word_rcat_def bin_rcat_def rev_map) apply transfer apply (simp add: horner_sum_foldr foldr_map comp_def) done lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" by (simp add: sign_Pls_ge_0) lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or \ \following definitions require both arithmetic and bit-wise word operations\ \ \to get \word_no_log_defs\ from \word_log_defs\, using \bin_log_bintrs\\ lemmas wils1 = bin_log_bintrs [THEN word_of_int_eq_iff [THEN iffD2], folded uint_word_of_int_eq, THEN eq_reflection] \ \the binary operations only\ (* BH: why is this needed? *) lemmas word_log_binary_defs = word_and_def word_or_def word_xor_def lemma setBit_no: "Bit_Operations.set_bit n (numeral bin) = word_of_int (bin_sc n True (numeral bin))" by (rule bit_word_eqI) (simp add: bit_simps) lemma clearBit_no: "unset_bit n (numeral bin) = word_of_int (bin_sc n False (numeral bin))" by (rule bit_word_eqI) (simp add: bit_simps) lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" for b n :: int by auto (metis pos_mod_conj)+ lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ a = take_bit (LENGTH('a) - n) a \ b = take_bit (LENGTH('a)) b" for w :: "'a::len word" by transfer (simp add: drop_bit_take_bit ac_simps) \ \limited hom result\ lemma word_cat_hom: "LENGTH('a::len) \ LENGTH('b::len) + LENGTH('c::len) \ (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = word_of_int ((\k n l. concat_bit n l k) w (size b) (uint b))" by transfer (simp add: take_bit_concat_bit_eq) lemma bintrunc_shiftl: "take_bit n (push_bit i m) = push_bit i (take_bit (n - i) m)" for m :: int by (fact take_bit_push_bit) lemma uint_shiftl: "uint (push_bit i n) = take_bit (size n) (push_bit i (uint n))" by (simp add: unsigned_push_bit_eq word_size) lemma bin_mask_conv_pow2: "mask n = 2 ^ n - (1 :: int)" by (fact mask_eq_exp_minus_1) lemma bin_mask_ge0: "mask n \ (0 :: int)" by (fact mask_nonnegative_int) +context + includes bit_operations_syntax +begin + lemma and_bin_mask_conv_mod: "x AND mask n = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask) +end + lemma bin_mask_numeral: "mask (numeral n) = (1 :: int) + 2 * mask (pred_numeral n)" by (fact mask_numeral) lemma bin_nth_mask: "bit (mask n :: int) i \ i < n" by (simp add: bit_mask_iff) lemma bin_sign_mask [simp]: "bin_sign (mask n) = 0" by (simp add: bin_sign_def bin_mask_conv_pow2) lemma bin_mask_p1_conv_shift: "mask n + 1 = push_bit n (1 :: int)" by (simp add: bin_mask_conv_pow2 shiftl_int_def) lemma sbintrunc_eq_in_range: "((signed_take_bit :: nat \ int \ int) n x = x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" "(x = (signed_take_bit :: nat \ int \ int) n x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) \ (signed_take_bit :: nat \ int \ int) n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) else if x \ 2 ^ n then x - 2 * (2 ^ n) else x)" apply (simp add: no_sbintr_alt2, safe) apply (simp add: mod_pos_geq) apply (subst mod_add_self1[symmetric], simp) done lemma sint_range': \- (2 ^ (LENGTH('a) - Suc 0)) \ sint x \ sint x < 2 ^ (LENGTH('a) - Suc 0)\ for x :: \'a::len word\ apply transfer using sbintr_ge sbintr_lt apply auto done lemma signed_arith_eq_checks_to_ord: "(sint a + sint b = sint (a + b )) = ((a <=s a + b) = (0 <=s b))" "(sint a - sint b = sint (a - b )) = ((0 <=s a - b) = (b <=s a))" "(- sint a = sint (- a)) = (0 <=s (- a) = (a <=s 0))" using sint_range'[where x=a] sint_range'[where x=b] by (simp_all add: sint_word_ariths word_sle_eq word_sless_alt sbintrunc_If) lemma signed_mult_eq_checks_double_size: assumes mult_le: "(2 ^ (len_of TYPE ('a) - 1) + 1) ^ 2 \ (2 :: int) ^ (len_of TYPE ('b) - 1)" and le: "2 ^ (LENGTH('a) - 1) \ (2 :: int) ^ (len_of TYPE ('b) - 1)" shows "(sint (a :: 'a :: len word) * sint b = sint (a * b)) = (scast a * scast b = (scast (a * b) :: 'b :: len word))" proof - have P: "(signed_take_bit :: nat \ int \ int) (size a - 1) (sint a * sint b) \ range ((signed_take_bit :: nat \ int \ int) (size a - 1))" by simp have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1" apply (cut_tac x=x in sint_range') apply (simp add: abs_le_iff word_size) done have abs_ab: "abs (sint a * sint b) < 2 ^ (LENGTH('b) - 1)" using abs_mult_less[OF abs[where x=a] abs[where x=b]] mult_le by (simp add: abs_mult power2_eq_square word_size) define r s where \r = LENGTH('a) - 1\ \s = LENGTH('b) - 1\ then have \LENGTH('a) = Suc r\ \LENGTH('b) = Suc s\ \size a = Suc r\ \size b = Suc r\ by (simp_all add: word_size) then show ?thesis using P[unfolded range_sbintrunc] abs_ab le apply clarsimp apply (transfer fixing: r s) apply (auto simp add: signed_take_bit_int_eq_self min.absorb2 simp flip: signed_take_bit_eq_iff_take_bit_eq) done qed lemma bintrunc_id: "\m \ int n; 0 < m\ \ take_bit n m = m" by (simp add: take_bit_int_eq_self_iff le_less_trans less_exp) lemma bin_cat_cong: "concat_bit n b a = concat_bit m d c" if "n = m" "a = c" "take_bit m b = take_bit m d" using that(3) unfolding that(1,2) by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bin_cat_eqD1: "concat_bit n b a = concat_bit n d c \ a = c" by (metis drop_bit_bin_cat_eq) lemma bin_cat_eqD2: "concat_bit n b a = concat_bit n d c \ take_bit n b = take_bit n d" by (metis take_bit_bin_cat_eq) lemma bin_cat_inj: "(concat_bit n b a) = concat_bit n d c \ a = c \ take_bit n b = take_bit n d" by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) code_identifier code_module Bits_Int \ (SML) Bit_Operations and (OCaml) Bit_Operations and (Haskell) Bit_Operations and (Scala) Bit_Operations end diff --git a/thys/Word_Lib/Examples.thy b/thys/Word_Lib/Examples.thy --- a/thys/Word_Lib/Examples.thy +++ b/thys/Word_Lib/Examples.thy @@ -1,226 +1,232 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) theory Examples imports Bitwise Next_and_Prev Generic_set_bit Word_Syntax Signed_Division_Word begin +context + includes bit_operations_syntax +begin + text "modulus" lemma "(27 :: 4 word) = -5" by simp lemma "(27 :: 4 word) = 11" by simp lemma "27 \ (11 :: 6 word)" by simp text "signed" lemma "(127 :: 6 word) = -1" by simp text "number ring simps" lemma "27 + 11 = (38::'a::len word)" "27 + 11 = (6::5 word)" "7 * 3 = (21::'a::len word)" "11 - 27 = (-16::'a::len word)" "- (- 11) = (11::'a::len word)" "-40 + 1 = (-39::'a::len word)" by simp_all lemma "word_pred 2 = 1" by simp lemma "word_succ (- 3) = -2" by simp lemma "23 < (27::8 word)" by simp lemma "23 \ (27::8 word)" by simp lemma "\ 23 < (27::2 word)" by simp lemma "0 < (4::3 word)" by simp lemma "1 < (4::3 word)" by simp lemma "0 < (1::3 word)" by simp text "ring operations" lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by simp text "casting" lemma "uint (234567 :: 10 word) = 71" by simp lemma "uint (-234567 :: 10 word) = 953" by simp lemma "sint (234567 :: 10 word) = 71" by simp lemma "sint (-234567 :: 10 word) = -71" by simp lemma "uint (1 :: 10 word) = 1" by simp lemma "unat (-234567 :: 10 word) = 953" by simp lemma "unat (1 :: 10 word) = 1" by simp lemma "ucast (0b1010 :: 4 word) = (0b10 :: 2 word)" by simp lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by simp lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by simp lemma "ucast (1 :: 4 word) = (1 :: 2 word)" by simp text "reducing goals to nat or int and arith:" lemma "i < x \ i < i + 1" for i x :: "'a::len word" by unat_arith lemma "i < x \ i < i + 1" for i x :: "'a::len word" by unat_arith text "bool lists" lemma "of_bl [True, False, True, True] = (0b1011::'a::len word)" by simp lemma "to_bl (0b110::4 word) = [False, True, True, False]" by simp lemma "of_bl (replicate 32 True) = (0xFFFFFFFF::32 word)" by (simp add: numeral_eq_Suc) text "bit operations" lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by simp lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by simp lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)" by simp lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by simp lemma "0 AND 5 = (0 :: 8 word)" by simp lemma "1 AND 1 = (1 :: 8 word)" by simp lemma "1 AND 0 = (0 :: 8 word)" by simp lemma "1 AND 5 = (1 :: 8 word)" by simp lemma "1 OR 6 = (7 :: 8 word)" by simp lemma "1 OR 1 = (1 :: 8 word)" by simp lemma "1 XOR 7 = (6 :: 8 word)" by simp lemma "1 XOR 1 = (0 :: 8 word)" by simp lemma "NOT 1 = (254 :: 8 word)" by simp lemma "NOT 0 = (255 :: 8 word)" by simp lemma "(-1 :: 32 word) = 0xFFFFFFFF" by simp lemma "bit (0b0010 :: 4 word) 1" by simp lemma "\ bit (0b0010 :: 4 word) 0" by simp lemma "\ bit (0b1000 :: 3 word) 4" by simp lemma "\ bit (1 :: 3 word) 2" by simp lemma "bit (0b11000 :: 10 word) n = (n = 4 \ n = 3)" by (auto simp add: bin_nth_Bit0 bin_nth_Bit1) lemma "set_bit 55 7 True = (183::'a::len word)" by simp lemma "set_bit 0b0010 7 True = (0b10000010::'a::len word)" by simp lemma "set_bit 0b0010 1 False = (0::'a::len word)" by simp lemma "set_bit 1 3 True = (0b1001::'a::len word)" by simp lemma "set_bit 1 0 False = (0::'a::len word)" by simp lemma "set_bit 0 3 True = (0b1000::'a::len word)" by simp lemma "set_bit 0 3 False = (0::'a::len word)" by simp lemma "odd (0b0101::'a::len word)" by simp lemma "even (0b1000::'a::len word)" by simp lemma "odd (1::'a::len word)" by simp lemma "even (0::'a::len word)" by simp lemma "\ msb (0b0101::4 word)" by simp lemma "msb (0b1000::4 word)" by simp lemma "\ msb (1::4 word)" by simp lemma "\ msb (0::4 word)" by simp lemma "word_cat (27::4 word) (27::8 word) = (2843::'a::len word)" by simp lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" by simp lemma "0b1011 << 2 = (0b101100::'a::len word)" by simp lemma "0b1011 >> 2 = (0b10::8 word)" by simp lemma "0b1011 >>> 2 = (0b10::8 word)" by simp lemma "1 << 2 = (0b100::'a::len word)" apply simp? oops lemma "slice 3 (0b101111::6 word) = (0b101::3 word)" by simp lemma "slice 3 (1::6 word) = (0::3 word)" apply simp? oops lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by simp lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by simp lemma "word_roti 2 0b1110 = (0b1011::4 word)" by simp lemma "word_roti (- 2) 0b0110 = (0b1001::4 word)" by simp lemma "word_rotr 2 0 = (0::4 word)" by simp lemma "word_rotr 2 1 = (0b0100::4 word)" apply simp? oops lemma "word_rotl 2 1 = (0b0100::4 word)" apply simp? oops lemma "word_roti (- 2) 1 = (0b0100::4 word)" apply simp? oops lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" proof - have "(x AND 0xff00) OR (x AND 0x00ff) = x AND (0xff00 OR 0x00ff)" by (simp only: word_ao_dist2) also have "0xff00 OR 0x00ff = (-1::16 word)" by simp also have "x AND -1 = x" by simp finally show ?thesis . qed lemma "word_next (2:: 8 word) = 3" by eval lemma "word_next (255:: 8 word) = 255" by eval lemma "word_prev (2:: 8 word) = 1" by eval lemma "word_prev (0:: 8 word) = 0" by eval text "proofs using bitwise expansion" lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" by word_bitwise lemma "(x AND NOT 3) >> 4 << 2 = ((x >> 2) AND NOT 3)" for x :: "10 word" by word_bitwise lemma "((x AND -8) >> 3) AND 7 = (x AND 56) >> 3" for x :: "12 word" by word_bitwise text "some problems require further reasoning after bit expansion" lemma "x \ 42 \ x \ 89" for x :: "8 word" apply word_bitwise apply blast done lemma "(x AND 1023) = 0 \ x \ -1024" for x :: \32 word\ apply word_bitwise apply clarsimp done text "operations like shifts by non-numerals will expose some internal list representations but may still be easy to solve" lemma shiftr_overflow: "32 \ a \ b >> a = 0" for b :: \32 word\ apply word_bitwise apply simp done (* testing for presence of word_bitwise *) lemma "((x :: 32 word) >> 3) AND 7 = (x AND 56) >> 3" by word_bitwise (* Tests *) lemma "( 4 :: 32 word) sdiv 4 = 1" "(-4 :: 32 word) sdiv 4 = -1" "(-3 :: 32 word) sdiv 4 = 0" "( 3 :: 32 word) sdiv -4 = 0" "(-3 :: 32 word) sdiv -4 = 0" "(-5 :: 32 word) sdiv -4 = 1" "( 5 :: 32 word) sdiv -4 = -1" by (simp_all add: sdiv_word_def signed_divide_int_def) lemma "( 4 :: 32 word) smod 4 = 0" "( 3 :: 32 word) smod 4 = 3" "(-3 :: 32 word) smod 4 = -3" "( 3 :: 32 word) smod -4 = 3" "(-3 :: 32 word) smod -4 = -3" "(-5 :: 32 word) smod -4 = -1" "( 5 :: 32 word) smod -4 = 1" by (simp_all add: smod_word_def signed_modulo_int_def signed_divide_int_def) lemma "1 < (1024::32 word) \ 1 \ (1024::32 word)" by simp end + +end diff --git a/thys/Word_Lib/Guide.thy b/thys/Word_Lib/Guide.thy --- a/thys/Word_Lib/Guide.thy +++ b/thys/Word_Lib/Guide.thy @@ -1,411 +1,414 @@ (* * Copyright Florian Haftmann * * SPDX-License-Identifier: BSD-2-Clause *) (*<*) theory Guide imports Word_Lib_Sumo Word_64 Ancient_Numeral begin notation (output) push_bit (\push'_bit\) notation (output) drop_bit (\drop'_bit\) notation (output) signed_drop_bit (\signed'_drop'_bit\) notation (output) Generic_set_bit.set_bit (\Generic'_set'_bit.set'_bit\) hide_const (open) Generic_set_bit.set_bit push_bit drop_bit signed_drop_bit no_notation bit (infixl \!!\ 100) abbreviation \push_bit n a \ a << n\ abbreviation \drop_bit n a \ a >> n\ abbreviation \signed_drop_bit n a \ a >>> n\ (*>*) section \A short overview over bit operations and word types\ subsection \Basic theories and key ideas\ text \ When formalizing bit operations, it is tempting to represent bit values as explicit lists over a binary type. This however is a bad idea, mainly due to the inherent ambiguities in representation concerning repeating leading bits. Hence this approach avoids such explicit lists altogether following an algebraic path: \<^item> Bit values are represented by numeric types: idealized unbounded bit values can be represented by type \<^typ>\int\, bounded bit values by quotient types over \<^typ>\int\, aka \<^typ>\'a word\. \<^item> (A special case are idealized unbounded bit values ending in @{term [source] 0} which can be represented by type \<^typ>\nat\ but only support a restricted set of operations). The most fundamental ideas are developed in theory \<^theory>\HOL.Parity\ (which is part of \<^theory>\Main\): \<^item> Multiplication by \<^term>\2 :: int\ is a bit shift to the left and \<^item> Division by \<^term>\2 :: int\ is a bit shift to the right. \<^item> Concerning bounded bit values, iterated shifts to the left may result in eliminating all bits by shifting them all beyond the boundary. The property \<^prop>\(2 :: int) ^ n \ 0\ represents that \<^term>\n\ is \<^emph>\not\ beyond that boundary. \<^item> The projection on a single bit is then @{thm [mode=iff] bit_iff_odd [where ?'a = int, no_vars]}. \<^item> This leads to the most fundamental properties of bit values: \<^item> Equality rule: @{thm [display, mode=iff] bit_eq_iff [where ?'a = int, no_vars]} \<^item> Induction rule: @{thm [display, mode=iff] bits_induct [where ?'a = int, no_vars]} \<^item> Characteristic properties @{prop [source] \bit (f x) n \ P x n\} are available in fact collection \<^text>\bit_simps\. On top of this, the following generic operations are provided after import of theory \<^theory>\HOL-Library.Bit_Operations\: \<^item> Singleton \<^term>\n\th bit: \<^term>\(2 :: int) ^ n\ \<^item> Bit mask upto bit \<^term>\n\: @{thm mask_eq_exp_minus_1 [where ?'a = int, no_vars]} \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]} \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]} \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]} \<^item> Bitwise negation: @{thm [mode=iff] bit_not_iff [where ?'a = int, no_vars]} \<^item> Bitwise conjunction: @{thm [mode=iff] bit_and_iff [where ?'a = int, no_vars]} \<^item> Bitwise disjunction: @{thm [mode=iff] bit_or_iff [where ?'a = int, no_vars]} \<^item> Bitwise exclusive disjunction: @{thm [mode=iff] bit_xor_iff [where ?'a = int, no_vars]} \<^item> Setting a single bit: @{thm set_bit_def [where ?'a = int, no_vars]} \<^item> Unsetting a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]} \<^item> Flipping a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]} \<^item> Signed truncation, or modulus centered around \<^term>\0::int\: @{thm [display] signed_take_bit_def [where ?'a = int, no_vars]} \<^item> (Bounded) conversion from and to a list of bits: @{thm [display] horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} Bit concatenation on \<^typ>\int\ as given by @{thm [display] concat_bit_def [no_vars]} appears quite technical but is the logical foundation for the quite natural bit concatenation on \<^typ>\'a word\ (see below). Proper word types are introduced in theory \<^theory>\HOL-Library.Word\, with the following specific operations: \<^item> Standard arithmetic: @{term \(+) :: 'a::len word \ 'a word \ 'a word\}, @{term \uminus :: 'a::len word \ 'a word\}, @{term \(-) :: 'a::len word \ 'a word \ 'a word\}, @{term \(*) :: 'a::len word \ 'a word \ 'a word\}, @{term \0 :: 'a::len word\}, @{term \1 :: 'a::len word\}, numerals etc. \<^item> Standard bit operations: see above. \<^item> Conversion with unsigned interpretation of words: \<^item> @{term [source] \unsigned :: 'a::len word \ 'b::semiring_1\} \<^item> Important special cases as abbreviations: \<^item> @{term [source] \unat :: 'a::len word \ nat\} \<^item> @{term [source] \uint :: 'a::len word \ int\} \<^item> @{term [source] \ucast :: 'a::len word \ 'b::len word\} \<^item> Conversion with signed interpretation of words: \<^item> @{term [source] \signed :: 'a::len word \ 'b::ring_1\} \<^item> Important special cases as abbreviations: \<^item> @{term [source] \sint :: 'a::len word \ int\} \<^item> @{term [source] \scast :: 'a::len word \ 'b::len word\} \<^item> Operations with unsigned interpretation of words: \<^item> @{thm [mode=iff] word_le_nat_alt [no_vars]} \<^item> @{thm [mode=iff] word_less_nat_alt [no_vars]} \<^item> @{thm unat_div_distrib [no_vars]} \<^item> @{thm unat_drop_bit_eq [no_vars]} \<^item> @{thm unat_mod_distrib [no_vars]} \<^item> @{thm [mode=iff] udvd_iff_dvd [no_vars]} \<^item> Operations with signed interpretation of words: \<^item> @{thm [mode=iff] word_sle_eq [no_vars]} \<^item> @{thm [mode=iff] word_sless_alt [no_vars]} \<^item> @{thm sint_signed_drop_bit_eq [no_vars]} \<^item> Rotation and reversal: \<^item> @{term [source] \word_rotl :: nat \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_rotr :: nat \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_roti :: int \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_reverse :: 'a::len word \ 'a word\} \<^item> Concatenation: @{term [source, display] \word_cat :: 'a::len word \ 'b::len word \ 'c::len word\} For proofs about words the following default strategies are applicable: \<^item> Using bit extensionality (facts \<^text>\bit_eq_iff\, \<^text>\bit_eqI\; fact collection \<^text>\bit_simps\). \<^item> Using the @{method transfer} method. \ subsection \More library theories\ text \ Note: currently, the theories listed here are hardly separate entities since they import each other in various ways. Always inspect them to understand what you pull in if you want to import one. \<^descr>[Syntax] \<^descr>[\<^theory>\Word_Lib.Syntax_Bundles\] Bundles to provide alternative syntax for various bit operations. \<^descr>[\<^theory>\Word_Lib.Hex_Words\] Printing word numerals as hexadecimal numerals. \<^descr>[\<^theory>\Word_Lib.Type_Syntax\] Pretty type-sensitive syntax for cast operations. \<^descr>[\<^theory>\Word_Lib.Word_Syntax\] Specific ASCII syntax for prominent bit operations on word. \<^descr>[Proof tools] \<^descr>[\<^theory>\Word_Lib.Norm_Words\] Rewriting word numerals to normal forms. \<^descr>[\<^theory>\Word_Lib.Bitwise\] Method @{method word_bitwise} decomposes word equalities and inequalities into bit propositions. \<^descr>[\<^theory>\Word_Lib.Word_EqI\] Method @{method word_eqI_solve} decomposes word equalities and inequalities into bit propositions. \<^descr>[Operations] \<^descr>[\<^theory>\Word_Lib.Signed_Division_Word\] Signed division on word: \<^item> @{term [source] \(sdiv) :: 'a::len word \ 'a word \ 'a word\} \<^item> @{term [source] \(smod) :: 'a::len word \ 'a word \ 'a word\} \<^descr>[\<^theory>\Word_Lib.Aligned\] \ \<^item> @{thm [mode=iff] is_aligned_iff_udvd [no_vars]} \<^descr>[\<^theory>\Word_Lib.Least_significant_bit\] The least significant bit as an alias: @{thm [mode=iff] lsb_odd [where ?'a = int, no_vars]} \<^descr>[\<^theory>\Word_Lib.Most_significant_bit\] The most significant bit: \<^item> @{thm [mode=iff] msb_int_def [of k]} \<^item> @{thm [mode=iff] word_msb_sint [no_vars]} \<^item> @{thm [mode=iff] msb_word_iff_sless_0 [no_vars]} \<^item> @{thm [mode=iff] msb_word_iff_bit [no_vars]} \<^descr>[\<^theory>\Word_Lib.Bit_Shifts_Infix_Syntax\] Abbreviations for bit shifts decorated with traditional infix syntax: \<^item> @{abbrev shiftl} \<^item> @{abbrev shiftr} \<^item> @{abbrev sshiftr} \<^descr>[\<^theory>\Word_Lib.Next_and_Prev\] \ \<^item> @{thm word_next_unfold [no_vars]} \<^item> @{thm word_prev_unfold [no_vars]} \<^descr>[\<^theory>\Word_Lib.Enumeration_Word\] More on explicit enumeration of word types. \<^descr>[\<^theory>\Word_Lib.More_Word_Operations\] Even more operations on word. \<^descr>[Types] \<^descr>[\<^theory>\Word_Lib.Signed_Words\] Formal tagging of word types with a \<^text>\signed\ marker. \<^descr>[Lemmas] \<^descr>[\<^theory>\Word_Lib.More_Word\] More lemmas on words. \<^descr>[\<^theory>\Word_Lib.Word_Lemmas\] More lemmas on words, covering many other theories mentioned here. \<^descr>[Words of popular lengths]. \<^descr>[\<^theory>\Word_Lib.Word_8\] for 8-bit words. \<^descr>[\<^theory>\Word_Lib.Word_16\] for 16-bit words. \<^descr>[\<^theory>\Word_Lib.Word_32\] for 32-bit words. \<^descr>[\<^theory>\Word_Lib.Word_64\] for 64-bit words. This theory is not part of \<^text>\Word_Lib_Sumo\, because it shadows names from \<^theory>\Word_Lib.Word_32\. They can be used together, but then require to use qualified names in applications. \ subsection \More library sessions\ text \ \<^descr>[\<^text>\Native_Word\] Makes machine words and machine arithmetic available for code generation. It provides a common abstraction that hides the differences between the different target languages. The code generator maps these operations to the APIs of the target languages. \ subsection \Legacy theories\ text \ The following theories contain material which has been factored out since it is not recommended to use it in new applications, mostly because matters can be expressed succinctly using already existing operations. This section gives some indication how to migrate away from those theories. However theorem coverage may still be terse in some cases. \<^descr>[\<^theory>\Word_Lib.Word_Lib_Sumo\] An entry point importing any relevant theory in that session. Intended for backward compatibility: start importing this theory when migrating applications to Isabelle2021, and later sort out what you really need. You may need to include \<^theory>\Word_Lib.Word_64\ separately. \<^descr>[\<^theory>\Word_Lib.Generic_set_bit\] Kind of an alias: @{thm set_bit_eq [no_vars]} \<^descr>[\<^theory>\Word_Lib.Typedef_Morphisms\] A low-level extension to HOL typedef providing conversions along type morphisms. The @{method transfer} method seems to be sufficient for most applications though. \<^descr>[\<^theory>\Word_Lib.Bit_Comprehension\] Comprehension syntax for bit values over predicates \<^typ>\nat \ bool\. For \<^typ>\'a::len word\, straightforward alternatives exist; difficult to handle for \<^typ>\int\. \<^descr>[\<^theory>\Word_Lib.Reversed_Bit_Lists\] Representation of bit values as explicit list in \<^emph>\reversed\ order. This should rarely be necessary: the \<^const>\bit\ projection should be sufficient in most cases. In case explicit lists are needed, existing operations can be used: @{thm [display] horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} \<^descr>[\<^theory>\Word_Lib.Many_More\] Collection of operations and theorems which are kept for backward compatibility and not used in other theories in session \<^text>\Word_Lib\. They are used in applications of \<^text>\Word_Lib\, but should be migrated to there. \ section \Changelog\ text \ \<^descr>[Changes since AFP 2021] ~ \<^item> Theory \<^theory>\Word_Lib.Ancient_Numeral\ is no part of \<^theory>\Word_Lib.Word_Lib_Sumo\ any longer. + \<^item> Infix syntax for \<^term>\(AND)\, \<^term>\(OR)\, \<^term>\(XOR)\ organized in + syntax bundle \<^text>\bit_operations_syntax\. + \<^item> Abbreviation \<^abbrev>\max_word\ moved from distribution into theory \<^theory>\Word_Lib.Legacy_Aliases\. \<^item> Operation \<^const>\test_bit\ replaced by input abbreviation \<^abbrev>\test_bit\. \<^item> Operation \<^const>\shiftl\ replaced by input abbreviation \<^abbrev>\shiftl\. \<^item> Operation \<^const>\shiftr\ replaced by input abbreviation \<^abbrev>\shiftr\. \<^item> Operation \<^const>\sshiftr\ replaced by input abbreviation \<^abbrev>\sshiftr\. \<^item> Abbreviations \<^abbrev>\bin_nth\, \<^abbrev>\bin_last\, \<^abbrev>\bin_rest\, \<^abbrev>\bintrunc\, \<^abbrev>\sbintrunc\, \<^abbrev>\norm_sint\, \<^abbrev>\bin_cat\ moved into theory \<^theory>\Word_Lib.Legacy_Aliases\. \<^item> Operations \<^abbrev>\shiftl1\, \<^abbrev>\shiftr1\, \<^abbrev>\sshiftr1\, \<^abbrev>\bshiftr1\, \<^abbrev>\setBit\, \<^abbrev>\clearBit\ moved from distribution into theory \<^theory>\Word_Lib.Legacy_Aliases\ and replaced by input abbreviations. \<^item> Operation \<^const>\complement\ replaced by input abbreviation \<^abbrev>\complement\. \ (*<*) end (*>*) diff --git a/thys/Word_Lib/Legacy_Aliases.thy b/thys/Word_Lib/Legacy_Aliases.thy --- a/thys/Word_Lib/Legacy_Aliases.thy +++ b/thys/Word_Lib/Legacy_Aliases.thy @@ -1,206 +1,212 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section \Legacy aliases\ theory Legacy_Aliases imports "HOL-Library.Word" begin abbreviation (input) test_bit :: \'a::semiring_bits \ nat \ bool\ where \test_bit \ bit\ abbreviation (input) bin_nth :: \int \ nat \ bool\ where \bin_nth \ bit\ abbreviation (input) bin_last :: \int \ bool\ where \bin_last \ odd\ abbreviation (input) bin_rest :: \int \ int\ where \bin_rest w \ w div 2\ abbreviation (input) bintrunc :: \nat \ int \ int\ where \bintrunc \ take_bit\ abbreviation (input) sbintrunc :: \nat \ int \ int\ where \sbintrunc \ signed_take_bit\ abbreviation (input) bin_cat :: \int \ nat \ int \ int\ where \bin_cat k n l \ concat_bit n l k\ abbreviation (input) norm_sint :: \nat \ int \ int\ where \norm_sint n \ signed_take_bit (n - 1)\ abbreviation (input) max_word :: \'a::len word\ where \max_word \ - 1\ abbreviation (input) complement :: \'a::len word \ 'a word\ where \complement \ not\ lemma complement_mask: "complement (2 ^ n - 1) = NOT (mask n)" unfolding mask_eq_decr_exp by simp abbreviation (input) shiftl1 :: \'a::len word \ 'a word\ where \shiftl1 \ (*) 2\ abbreviation (input) shiftr1 :: \'a::len word \ 'a word\ where \shiftr1 w \ w div 2\ abbreviation (input) sshiftr1 :: \'a::len word \ 'a word\ where \sshiftr1 \ signed_drop_bit (Suc 0)\ +context + includes bit_operations_syntax +begin + abbreviation (input) bshiftr1 :: \bool \ 'a::len word \ 'a word\ where \bshiftr1 b w \ w div 2 OR push_bit (LENGTH('a) - Suc 0) (of_bool b) \ +end + lemma shiftr1_1: "shiftr1 (1::'a::len word) = 0" by (fact bits_1_div_2) lemma sshiftr1_eq: \sshiftr1 w = word_of_int (sint w div 2)\ by (rule bit_word_eqI) (auto simp add: bit_simps min_def simp flip: bit_Suc elim: le_SucE) lemma shiftl1_eq: \shiftl1 w = word_of_int (2 * uint w)\ by (rule bit_word_eqI) (simp add: bit_simps) lemma bit_shiftl1_iff: \bit (shiftl1 w) n \ 0 < n \ n < LENGTH('a) \ bit w (n - 1)\ for w :: \'a::len word\ by (auto simp add: bit_simps) lemma bit_shiftr1_iff: \bit (shiftr1 w) n \ bit w (Suc n)\ by (simp add: bit_Suc) lemma shiftr1_eq: \shiftr1 w = word_of_int (uint w div 2)\ by (rule bit_word_eqI) (simp add: bit_simps flip: bit_Suc) lemma shiftl1_rev: "shiftl1 w = word_reverse (shiftr1 (word_reverse w))" by (rule bit_word_eqI) (auto simp add: bit_simps Suc_diff_Suc simp flip: bit_Suc) lemma shiftl1_p: "shiftl1 w = w + w" for w :: "'a::len word" by (fact mult_2) lemma shiftr1_bintr: "(shiftr1 (numeral w) :: 'a::len word) = word_of_int (take_bit LENGTH('a) (numeral w) div 2)" by (rule bit_word_eqI) (simp add: bit_simps flip: bit_Suc) lemma sshiftr1_sbintr: "(sshiftr1 (numeral w) :: 'a::len word) = word_of_int (signed_take_bit (LENGTH('a) - 1) (numeral w) div 2)" apply (cases \LENGTH('a)\) apply simp_all apply (rule bit_word_eqI) apply (auto simp add: bit_simps min_def simp flip: bit_Suc elim: le_SucE) done lemma shiftl1_wi: "shiftl1 (word_of_int w) = word_of_int (2 * w)" by transfer simp lemma shiftl1_numeral: "shiftl1 (numeral w) = numeral (Num.Bit0 w)" unfolding word_numeral_alt shiftl1_wi by simp lemma shiftl1_neg_numeral: "shiftl1 (- numeral w) = - numeral (Num.Bit0 w)" unfolding word_neg_numeral_alt shiftl1_wi by simp lemma shiftl1_0: "shiftl1 0 = 0" by (fact mult_zero_right) lemma shiftl1_def_u: "shiftl1 w = word_of_int (2 * uint w)" by (fact shiftl1_eq) lemma shiftl1_def_s: "shiftl1 w = word_of_int (2 * sint w)" by simp lemma shiftr1_0: "shiftr1 0 = 0" by (fact bits_div_0) lemma sshiftr1_0: "sshiftr1 0 = 0" by (fact signed_drop_bit_of_0) lemma sshiftr1_n1: "sshiftr1 (- 1) = - 1" by (fact signed_drop_bit_of_minus_1) lemma uint_shiftr1: "uint (shiftr1 w) = uint w div 2" by (rule bit_eqI) (simp add: bit_simps flip: bit_Suc) lemma shiftr1_div_2: "uint (shiftr1 w) = uint w div 2" by (fact uint_shiftr1) lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2" by (rule bit_eqI) (auto simp add: bit_simps ac_simps min_def simp flip: bit_Suc elim: le_SucE) lemma bit_sshiftr1_iff: \bit (sshiftr1 w) n \ bit w (if n = LENGTH('a) - 1 then LENGTH('a) - 1 else Suc n)\ for w :: \'a::len word\ by (auto simp add: bit_simps) lemma bit_bshiftr1_iff: \bit (bshiftr1 b w) n \ b \ n = LENGTH('a) - 1 \ bit w (Suc n)\ for w :: \'a::len word\ by (auto simp add: bit_simps simp flip: bit_Suc) lemma nth_shiftl1: "bit (shiftl1 w) n \ n < size w \ n > 0 \ bit w (n - 1)" by (auto simp add: word_size bit_simps) lemma nth_shiftr1: "bit (shiftr1 w) n = bit w (Suc n)" by (simp add: bit_Suc) lemma nth_sshiftr1: "bit (sshiftr1 w) n = (if n = size w - 1 then bit w n else bit w (Suc n))" by (auto simp add: word_size bit_simps) lemma shiftl_power: "(shiftl1 ^^ x) (y::'a::len word) = 2 ^ x * y" by (induction x) simp_all lemma le_shiftr1: \shiftr1 u \ shiftr1 v\ if \u \ v\ using that by (simp add: word_le_nat_alt unat_div div_le_mono) lemma le_shiftr1': "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" by (meson dual_order.antisym le_cases le_shiftr1) abbreviation (input) setBit :: \'a::len word \ nat \ 'a word\ where \setBit w n \ set_bit n w\ abbreviation (input) clearBit :: \'a::len word \ nat \ 'a word\ where \clearBit w n \ unset_bit n w\ lemma bit_setBit_iff: \bit (setBit w m) n \ (m = n \ n < LENGTH('a) \ bit w n)\ for w :: \'a::len word\ by (auto simp add: bit_simps) lemma bit_clearBit_iff: \bit (clearBit w m) n \ m \ n \ bit w n\ for w :: \'a::len word\ by (auto simp add: bit_simps) lemmas less_def = less_eq [symmetric] lemmas le_def = not_less [symmetric, where ?'a = nat] end diff --git a/thys/Word_Lib/Many_More.thy b/thys/Word_Lib/Many_More.thy --- a/thys/Word_Lib/Many_More.thy +++ b/thys/Word_Lib/Many_More.thy @@ -1,684 +1,690 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) theory Many_More imports Main "HOL-Library.Word" More_Word Even_More_List begin lemma nat_less_mult_monoish: "\ a < b; c < (d :: nat) \ \ (a + 1) * (c + 1) <= b * d" apply (drule Suc_leI)+ apply (drule(1) mult_le_mono) apply simp done +context + includes bit_operations_syntax +begin + lemma if_and_helper: "(If x v v') AND v'' = If x (v AND v'') (v' AND v'')" by (rule if_distrib) +end + lemma eq_eqI: "a = b \ (a = x) = (b = x)" by simp lemma map2_Cons_2_3: "(map2 f xs (y # ys) = (z # zs)) = (\x xs'. xs = x # xs' \ f x y = z \ map2 f xs' ys = zs)" by (case_tac xs, simp_all) lemma map2_xor_replicate_False: "map2 (\x y. x \ \ y) xs (replicate n False) = take n xs" apply (induct xs arbitrary: n, simp) apply (case_tac n; simp) done lemma plus_Collect_helper: "(+) x ` {xa. P (xa :: 'a :: len word)} = {xa. P (xa - x)}" by (fastforce simp add: image_def) lemma plus_Collect_helper2: "(+) (- x) ` {xa. P (xa :: 'a :: len word)} = {xa. P (x + xa)}" using plus_Collect_helper [of "- x" P] by (simp add: ac_simps) lemma range_subset_eq2: "{a :: 'a :: len word .. b} \ {} \ ({a .. b} \ {c .. d}) = (c \ a \ b \ d)" by simp lemma nat_mod_power_lem: fixes a :: nat shows "1 < a \ a ^ n mod a ^ m = (if m \ n then 0 else a ^ n)" apply (clarsimp) apply (clarsimp simp add: le_iff_add power_add) done lemma i_hate_words_helper: "i \ (j - k :: nat) \ i \ j" by simp lemma i_hate_words: "unat (a :: 'a word) \ unat (b :: 'a :: len word) - Suc 0 \ a \ -1" apply (frule i_hate_words_helper) apply (subst(asm) word_le_nat_alt[symmetric]) apply (clarsimp simp only: word_minus_one_le) apply (simp only: linorder_not_less[symmetric]) apply (erule notE) apply (rule diff_Suc_less) apply (subst neq0_conv[symmetric]) apply (subst unat_eq_0) apply (rule notI, drule arg_cong[where f="(+) 1"]) apply simp done lemma If_eq_obvious: "x \ z \ ((if P then x else y) = z) = (\ P \ y = z)" by simp lemma Some_to_the: "v = Some x \ x = the v" by simp lemma dom_if_Some: "dom (\x. if P x then Some (f x) else g x) = {x. P x} \ dom g" by fastforce lemma dom_insert_absorb: "x \ dom f \ insert x (dom f) = dom f" by (fact insert_absorb) lemma emptyE2: "\ S = {}; x \ S \ \ P" by simp lemma ptr_add_image_multI: "\ \x y. (x * val = y * val') = (x * val'' = y); x * val'' \ S \ \ ptr_add ptr (x * val) \ (\p. ptr_add ptr (p * val')) ` S" by (auto simp add: image_iff) metis lemmas map_prod_split_imageI' = map_prod_imageI[where f="case_prod f" and g="case_prod g" and a="(a, b)" and b="(c, d)" for a b c d f g] lemmas map_prod_split_imageI = map_prod_split_imageI'[simplified] lemma dom_if: "dom (\a. if a \ addrs then Some (f a) else g a) = addrs \ dom g" by (auto simp: dom_def split: if_split) lemmas arg_cong_Not = arg_cong [where f=Not] lemma drop_append_miracle: "n = length xs \ drop n (xs @ ys) = ys" by simp lemma foldr_does_nothing_to_xf: "\ \x s. x \ set xs \ xf (f x s) = xf s \ \ xf (foldr f xs s) = xf s" by (induct xs, simp_all) lemma mod_mod_power_int: fixes k :: int shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" by (fact mod_exp_eq) lemma le_step_down_nat:"\(i::nat) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma le_step_down_int:"\(i::int) \ n; i = n \ P; i \ n - 1 \ P\ \ P" by arith lemma replicate_numeral [simp]: "replicate (numeral k) x = x # replicate (pred_numeral k) x" by (simp add: numeral_eq_Suc) lemma list_exhaust_size_gt0: assumes "\a list. y = a # list \ P" shows "0 < length y \ P" apply (cases y) apply simp apply (rule assms) apply fastforce done lemma list_exhaust_size_eq0: assumes "y = [] \ P" shows "length y = 0 \ P" apply (cases y) apply (rule assms) apply simp apply simp done lemma size_Cons_lem_eq: "y = xa # list \ size y = Suc k \ size list = k" by auto lemma takeWhile_take_has_property: "n \ length (takeWhile P xs) \ \x \ set (take n xs). P x" by (induct xs arbitrary: n; simp split: if_split_asm) (case_tac n, simp_all) lemma takeWhile_take_has_property_nth: "\ n < length (takeWhile P xs) \ \ P (xs ! n)" by (induct xs arbitrary: n; simp split: if_split_asm) (case_tac n, simp_all) lemma takeWhile_replicate: "takeWhile f (replicate len x) = (if f x then replicate len x else [])" by (induct_tac len) auto lemma takeWhile_replicate_empty: "\ f x \ takeWhile f (replicate len x) = []" by (simp add: takeWhile_replicate) lemma takeWhile_replicate_id: "f x \ takeWhile f (replicate len x) = replicate len x" by (simp add: takeWhile_replicate) lemma nth_rev: "n < length xs \ rev xs ! n = xs ! (length xs - 1 - n)" using rev_nth by simp lemma nth_rev_alt: "n < length ys \ ys ! n = rev ys ! (length ys - Suc n)" by (simp add: nth_rev) lemma hd_butlast: "length xs > 1 \ hd (butlast xs) = hd xs" by (cases xs) auto lemma split_upt_on_n: "n < m \ [0 ..< m] = [0 ..< n] @ [n] @ [Suc n ..< m]" by (metis append_Cons append_Nil less_Suc_eq_le less_imp_le_nat upt_add_eq_append' upt_rec zero_less_Suc) lemma drop_eq_mono: assumes le: "m \ n" assumes drop: "drop m xs = drop m ys" shows "drop n xs = drop n ys" proof - have ex: "\p. n = p + m" by (rule exI[of _ "n - m"]) (simp add: le) then obtain p where p: "n = p + m" by blast show ?thesis unfolding p drop_drop[symmetric] drop by simp qed lemma drop_Suc_nth: "n < length xs \ drop n xs = xs!n # drop (Suc n) xs" by (simp add: Cons_nth_drop_Suc) lemma and_len: "xs = ys \ xs = ys \ length xs = length ys" by auto lemma tl_if: "tl (if p then xs else ys) = (if p then tl xs else tl ys)" by auto lemma hd_if: "hd (if p then xs else ys) = (if p then hd xs else hd ys)" by auto lemma if_single: "(if xc then [xab] else [an]) = [if xc then xab else an]" by auto \ \note -- \if_Cons\ can cause blowup in the size, if \p\ is complex, so make a simproc\ lemma if_Cons: "(if p then x # xs else y # ys) = If p x y # If p xs ys" by auto lemma list_of_false: "True \ set xs \ xs = replicate (length xs) False" by (induct xs, simp_all) lemma list_all2_induct [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q xs ys" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys\ \ P (x # xs) (y # ys)" shows "P xs ys" using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 then show ?case by auto fact+ next case (2 x xs y ys) show ?case proof (rule consr) from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all then show "P xs ys" by (intro "2.hyps") qed qed lemma replicate_minus: "k < n \ replicate n False = replicate (n - k) False @ replicate k False" by (subst replicate_add [symmetric]) simp lemma cart_singleton_empty: "(S \ {e} = {}) = (S = {})" by blast lemma MinI: assumes fa: "finite A" and ne: "A \ {}" and xv: "m \ A" and min: "\y \ A. m \ y" shows "Min A = m" using fa ne xv min proof (induct A arbitrary: m rule: finite_ne_induct) case singleton then show ?case by simp next case (insert y F) from insert.prems have yx: "m \ y" and fx: "\y \ F. m \ y" by auto have "m \ insert y F" by fact then show ?case proof assume mv: "m = y" have mlt: "m \ Min F" by (rule iffD2 [OF Min_ge_iff [OF insert.hyps(1) insert.hyps(2)] fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mv [symmetric]) apply (auto simp: min_def mlt) done next assume "m \ F" then have mf: "Min F = m" by (rule insert.hyps(4) [OF _ fx]) show ?case apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)]) apply (subst mf) apply (rule iffD2 [OF _ yx]) apply (auto simp: min_def) done qed qed lemma power_numeral: "a ^ numeral k = a * a ^ (pred_numeral k)" by (simp add: numeral_eq_Suc) lemma funpow_numeral [simp]: "f ^^ numeral k = f \ f ^^ (pred_numeral k)" by (simp add: numeral_eq_Suc) lemma funpow_minus_simp: "0 < n \ f ^^ n = f \ f ^^ (n - 1)" by (auto dest: gr0_implies_Suc) lemma rco_alt: "(f \ g) ^^ n \ f = f \ (g \ f) ^^ n" apply (rule ext) apply (induct n) apply (simp_all add: o_def) done lemma union_sub: "\B \ A; C \ B\ \ (A - B) \ (B - C) = (A - C)" by fastforce lemma insert_sub: "x \ xs \ (insert x (xs - ys)) = (xs - (ys - {x}))" by blast lemma ran_upd: "\ inj_on f (dom f); f y = Some z \ \ ran (\x. if x = y then None else f x) = ran f - {z}" unfolding ran_def apply (rule set_eqI) apply simp by (metis domI inj_on_eq_iff option.sel) lemma if_apply_def2: "(if P then F else G) = (\x. (P \ F x) \ (\ P \ G x))" by simp lemma case_bool_If: "case_bool P Q b = (if b then P else Q)" by simp lemma if_f: "(if a then f b else f c) = f (if a then b else c)" by simp lemma size_if: "size (if p then xs else ys) = (if p then size xs else size ys)" by (fact if_distrib) lemma if_Not_x: "(if p then \ x else x) = (p = (\ x))" by auto lemma if_x_Not: "(if p then x else \ x) = (p = x)" by auto lemma if_same_and: "(If p x y \ If p u v) = (if p then x \ u else y \ v)" by auto lemma if_same_eq: "(If p x y = (If p u v)) = (if p then x = u else y = v)" by auto lemma if_same_eq_not: "(If p x y = (\ If p u v)) = (if p then x = (\ u) else y = (\ v))" by auto lemma the_elemI: "y = {x} \ the_elem y = x" by simp lemma nonemptyE: "S \ {} \ (\x. x \ S \ R) \ R" by auto lemmas xtr1 = xtrans(1) lemmas xtr2 = xtrans(2) lemmas xtr3 = xtrans(3) lemmas xtr4 = xtrans(4) lemmas xtr5 = xtrans(5) lemmas xtr6 = xtrans(6) lemmas xtr7 = xtrans(7) lemmas xtr8 = xtrans(8) lemmas if_fun_split = if_apply_def2 lemma not_empty_eq: "(S \ {}) = (\x. x \ S)" by auto lemma range_subset_lower: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ c \ a" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_upper: fixes c :: "'a ::linorder" shows "\ {a..b} \ {c..d}; x \ {a..b} \ \ b \ d" apply (frule (1) subsetD) apply (rule classical) apply clarsimp done lemma range_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d)" apply (insert non_empty) apply (rule iffI) apply (frule range_subset_lower [where x=a], simp) apply (drule range_subset_upper [where x=a], simp) apply simp apply auto done lemma range_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} = {c..d}) = (a = c \ b = d)" by (metis atLeastatMost_subset_iff eq_iff non_empty) lemma range_strict_subset_eq: fixes a::"'a::linorder" assumes non_empty: "a \ b" shows "({a..b} \ {c..d}) = (c \ a \ b \ d \ (a = c \ b \ d))" apply (insert non_empty) apply (subst psubset_eq) apply (subst range_subset_eq, assumption+) apply (subst range_eq, assumption+) apply simp done lemma range_subsetI: fixes x :: "'a :: order" assumes xX: "X \ x" and yY: "y \ Y" shows "{x .. y} \ {X .. Y}" using xX yY by auto lemma set_False [simp]: "(set bs \ {False}) = (True \ set bs)" by auto lemma int_not_emptyD: "A \ B \ {} \ \x. x \ A \ x \ B" by (erule contrapos_np, clarsimp simp: disjoint_iff_not_equal) definition sum_map :: "('a \ 'b) \ ('c \ 'd) \ 'a + 'c \ 'b + 'd" where "sum_map f g x \ case x of Inl v \ Inl (f v) | Inr v' \ Inr (g v')" lemma sum_map_simps[simp]: "sum_map f g (Inl v) = Inl (f v)" "sum_map f g (Inr w) = Inr (g w)" by (simp add: sum_map_def)+ lemma if_Some_None_eq_None: "((if P then Some v else None) = None) = (\ P)" by simp lemma CollectPairFalse [iff]: "{(a,b). False} = {}" by (simp add: split_def) lemma if_conj_dist: "((if b then w else x) \ (if b then y else z) \ X) = ((if b then w \ y else x \ z) \ X)" by simp lemma if_P_True1: "Q \ (if P then True else Q)" by simp lemma if_P_True2: "Q \ (if P then Q else True)" by simp lemmas nat_simps = diff_add_inverse2 diff_add_inverse lemmas nat_iffs = le_add1 le_add2 lemma nat_min_simps: "(a::nat) \ b \ min b a = a" "a \ b \ min a b = a" by simp_all lemmas zadd_diff_inverse = trans [OF diff_add_cancel [symmetric] add.commute] lemmas add_diff_cancel2 = add.commute [THEN diff_eq_eq [THEN iffD2]] lemmas mcl = mult_cancel_left [THEN iffD1, THEN make_pos_rule] lemma pl_pl_rels: "a + b = c + d \ a \ c \ b \ d \ a \ c \ b \ d" for a b c d :: nat by arith lemmas pl_pl_rels' = add.commute [THEN [2] trans, THEN pl_pl_rels] lemma iszero_minus: \iszero (- z) \ iszero z\ by (simp add: iszero_def) lemma diff_le_eq': "a - b \ c \ a \ b + c" for a b c :: int by arith lemma zless2: "0 < (2 :: int)" by (fact zero_less_numeral) lemma zless2p: "0 < (2 ^ n :: int)" by arith lemma zle2p: "0 \ (2 ^ n :: int)" by arith lemma ex_eq_or: "(\m. n = Suc m \ (m = k \ P m)) \ n = Suc k \ (\m. n = Suc m \ P m)" by auto lemma power_minus_simp: "0 < n \ a ^ n = a * a ^ (n - 1)" by (auto dest: gr0_implies_Suc) lemma n2s_ths: \2 + n = Suc (Suc n)\ \n + 2 = Suc (Suc n)\ by (fact add_2_eq_Suc add_2_eq_Suc')+ lemma s2n_ths: \Suc (Suc n) = 2 + n\ \Suc (Suc n) = n + 2\ by simp_all lemma gt_or_eq_0: "0 < y \ 0 = y" for y :: nat by arith lemma sum_imp_diff: "j = k + i \ j - i = k" for k :: nat by simp lemma le_diff_eq': "a \ c - b \ b + a \ c" for a b c :: int by arith lemma less_diff_eq': "a < c - b \ b + a < c" for a b c :: int by arith lemma diff_less_eq': "a - b < c \ a < b + c" for a b c :: int by arith lemma axxbyy: "a + m + m = b + n + n \ a = 0 \ a = 1 \ b = 0 \ b = 1 \ a = b \ m = n" for a b m n :: int by arith lemma minus_eq: "m - k = m \ k = 0 \ m = 0" for k m :: nat by arith lemma pl_pl_mm: "a + b = c + d \ a - c = d - b" for a b c d :: nat by arith lemmas pl_pl_mm' = add.commute [THEN [2] trans, THEN pl_pl_mm] lemma less_le_mult': "w * c < b * c \ 0 \ c \ (w + 1) * c \ b * c" for b c w :: int apply (rule mult_right_mono) apply (rule zless_imp_add1_zle) apply (erule (1) mult_right_less_imp_less) apply assumption done lemma less_le_mult: "w * c < b * c \ 0 \ c \ w * c + c \ b * c" for b c w :: int using less_le_mult' [of w c b] by (simp add: algebra_simps) lemmas less_le_mult_minus = iffD2 [OF le_diff_eq less_le_mult, simplified left_diff_distrib] lemma gen_minus: "0 < n \ f n = f (Suc (n - 1))" by auto lemma mpl_lem: "j \ i \ k < j \ i - j + k < i" for i j k :: nat by arith lemmas dme = div_mult_mod_eq lemmas dtle = div_times_less_eq_dividend lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] div_times_less_eq_dividend] lemmas sdl = div_nat_eqI lemma given_quot: "f > 0 \ (f * l + (f - 1)) div f = l" for f l :: nat by (rule div_nat_eqI) (simp_all) lemma given_quot_alt: "f > 0 \ (l * f + f - Suc 0) div f = l" for f l :: nat apply (frule given_quot) apply (rule trans) prefer 2 apply (erule asm_rl) apply (rule_tac f="\n. n div f" in arg_cong) apply (simp add : ac_simps) done lemma x_power_minus_1: fixes x :: "'a :: {ab_group_add, power, numeral, one}" shows "x + (2::'a) ^ n - (1::'a) = x + (2 ^ n - 1)" by simp lemma nat_diff_add: fixes i :: nat shows "\ i + j = k \ \ i = k - j" by arith lemma pow_2_gt: "n \ 2 \ (2::int) < 2 ^ n" by (induct n) auto lemma sum_to_zero: "(a :: 'a :: ring) + b = 0 \ a = (- b)" by (drule arg_cong[where f="\ x. x - a"], simp) lemma arith_is_1: "\ x \ Suc 0; x > 0 \ \ x = 1" by arith lemma suc_le_pow_2: "1 < (n::nat) \ Suc n < 2 ^ n" by (induct n; clarsimp) (case_tac "n = 1"; clarsimp) lemma nat_le_Suc_less_imp: "x < y \ x \ y - Suc 0" by arith lemma power_sub_int: "\ m \ n; 0 < b \ \ b ^ n div b ^ m = (b ^ (n - m) :: int)" apply (subgoal_tac "\n'. n = m + n'") apply (clarsimp simp: power_add) apply (rule exI[where x="n - m"]) apply simp done lemma nat_Suc_less_le_imp: "(k::nat) < Suc n \ k \ n" by auto lemma nat_add_less_by_max: "\ (x::nat) \ xmax ; y < k - xmax \ \ x + y < k" by simp lemma nat_le_Suc_less: "0 < y \ (x \ y - Suc 0) = (x < y)" by arith lemma nat_power_minus_less: "a < 2 ^ (x - n) \ (a :: nat) < 2 ^ x" by (erule order_less_le_trans) simp lemma less_le_mult_nat': "w * c < b * c ==> 0 \ c ==> Suc w * c \ b * (c::nat)" apply (rule mult_right_mono) apply (rule Suc_leI) apply (erule (1) mult_right_less_imp_less) apply assumption done lemma less_le_mult_nat: \0 < c \ w < b \ c + w * c \ b * c\ for b c w :: nat using less_le_mult_nat' [of w c b] by simp lemma p_assoc_help: fixes p :: "'a::{ring,power,numeral,one}" shows "p + 2^sz - 1 = p + (2^sz - 1)" by simp lemma pow_mono_leq_imp_lt: "x \ y \ x < 2 ^ y" by (simp add: le_less_trans) lemma small_powers_of_2: "x \ 3 \ x < 2 ^ (x - 1)" by (induct x; simp add: suc_le_pow_2) lemma nat_less_power_trans2: fixes n :: nat shows "\n < 2 ^ (m - k); k \ m\ \ n * 2 ^ k < 2 ^ m" by (subst mult.commute, erule (1) nat_less_power_trans) lemma nat_move_sub_le: "(a::nat) + b \ c \ a \ c - b" by arith lemma plus_minus_one_rewrite: "v + (- 1 :: ('a :: {ring, one, uminus})) \ v - 1" by (simp add: field_simps) lemma Suc_0_lt_2p_len_of: "Suc 0 < 2 ^ LENGTH('a :: len)" by (metis One_nat_def len_gt_0 lessI numeral_2_eq_2 one_less_power) 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,2544 +1,2550 @@ (* * 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\ proof (rule bit_eqI) fix n assume \2 ^ n \ (0 :: 'a word)\ then have \n < LENGTH('a)\ by simp then show \bit (x AND Bit_Operations.mask LENGTH('a)) n \ bit x n\ by (simp add: bit_and_iff bit_mask_iff) qed 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) (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) 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) 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) 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) 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) 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) 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) 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) 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) 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 (clarsimp simp add: mask_eq_decr_exp Parity.bit_eq_iff bit_and_iff bit_not_iff bit_mask_iff) 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 (simp add: bit_eq_iff 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 simp add: bit_eq_iff) 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 simp add: Parity.bit_eq_iff bit_xor_iff max_def) 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 (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,1026 +1,1032 @@ (* * 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_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 lemma unat_of_nat_ctz_smw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len signed word) = word_ctz w" by simp 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 (simp add: bit_eq_iff) 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) 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 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 simp apply (drule_tac x="x AND mask n >> m" in spec) apply (auto simp add: and_mask_less_size wsst_TYs multiple_mask_trivia neg_mask_twice word_bw_assocs max_absorb2) apply (erule notE) apply (simp add: shiftr_mask2) apply (rule and_mask_less') apply simp apply (subst (asm) disjunctive_add) apply (simp add: bit_simps) apply auto apply (erule notE) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) 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 (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 (rule not_msb_from_less) apply (simp add: word_less_nat_alt) 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 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/Most_significant_bit.thy b/thys/Word_Lib/Most_significant_bit.thy --- a/thys/Word_Lib/Most_significant_bit.thy +++ b/thys/Word_Lib/Most_significant_bit.thy @@ -1,188 +1,193 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson, NICTA *) section \Dedicated operation for the most significant bit\ theory Most_significant_bit imports "HOL-Library.Word" More_Word More_Arithmetic begin class msb = fixes msb :: \'a \ bool\ instantiation int :: msb begin definition \msb x \ x < 0\ for x :: int instance .. end lemma msb_bin_rest [simp]: "msb (x div 2) = msb x" for x :: int by (simp add: msb_int_def) +context + includes bit_operations_syntax +begin + lemma int_msb_and [simp]: "msb ((x :: int) AND y) \ msb x \ msb y" by(simp add: msb_int_def) lemma int_msb_or [simp]: "msb ((x :: int) OR y) \ msb x \ msb y" by(simp add: msb_int_def) lemma int_msb_xor [simp]: "msb ((x :: int) XOR y) \ msb x \ msb y" by(simp add: msb_int_def) +end + lemma int_msb_not [simp]: "msb (NOT (x :: int)) \ \ msb x" by(simp add: msb_int_def not_less) lemma msb_shiftl [simp]: "msb (push_bit n (x :: int)) \ msb x" by(simp add: msb_int_def) lemma msb_shiftr [simp]: "msb (drop_bit r (x :: int)) \ msb x" by(simp add: msb_int_def) lemma msb_0 [simp]: "msb (0 :: int) = False" by(simp add: msb_int_def) lemma msb_1 [simp]: "msb (1 :: int) = False" by(simp add: msb_int_def) lemma msb_numeral [simp]: "msb (numeral n :: int) = False" "msb (- numeral n :: int) = True" by(simp_all add: msb_int_def) instantiation word :: (len) msb begin definition msb_word :: \'a word \ bool\ where msb_word_iff_bit: \msb w \ bit w (LENGTH('a) - Suc 0)\ for w :: \'a::len word\ instance .. end lemma msb_word_eq: \msb w \ bit w (LENGTH('a) - 1)\ for w :: \'a::len word\ by (simp add: msb_word_iff_bit) lemma word_msb_sint: "msb w \ sint w < 0" by (simp add: msb_word_eq bit_last_iff) lemma msb_word_iff_sless_0: \msb w \ w by (simp add: word_msb_sint word_sless_alt) lemma msb_word_of_int: "msb (word_of_int x::'a::len word) = bit x (LENGTH('a) - 1)" by (simp add: msb_word_iff_bit bit_simps) lemma word_msb_numeral [simp]: "msb (numeral w::'a::len word) = bit (numeral w :: int) (LENGTH('a) - 1)" unfolding word_numeral_alt by (rule msb_word_of_int) lemma word_msb_neg_numeral [simp]: "msb (- numeral w::'a::len word) = bit (- numeral w :: int) (LENGTH('a) - 1)" unfolding word_neg_numeral_alt by (rule msb_word_of_int) lemma word_msb_0 [simp]: "\ msb (0::'a::len word)" by (simp add: msb_word_iff_bit) lemma word_msb_1 [simp]: "msb (1::'a::len word) \ LENGTH('a) = 1" by (simp add: msb_word_iff_bit le_Suc_eq) lemma word_msb_nth: "msb w = bit (uint w) (LENGTH('a) - 1)" for w :: "'a::len word" by (simp add: msb_word_iff_bit bit_simps) lemma msb_nth: "msb w = bit w (LENGTH('a) - 1)" for w :: "'a::len word" by (fact msb_word_eq) lemma word_msb_n1 [simp]: "msb (-1::'a::len word)" by (simp add: msb_word_eq not_le) lemma msb_shift: "msb w \ drop_bit (LENGTH('a) - 1) w \ 0" for w :: "'a::len word" by (simp add: msb_word_eq bit_iff_odd_drop_bit drop_bit_eq_zero_iff_not_bit_last) lemmas word_ops_msb = msb1 [unfolded msb_nth [symmetric, unfolded One_nat_def]] lemma word_sint_msb_eq: "sint x = uint x - (if msb x then 2 ^ size x else 0)" apply (cases \LENGTH('a)\) apply (simp_all add: msb_word_iff_bit word_size) apply transfer apply (simp add: signed_take_bit_eq_take_bit_minus) done lemma word_sle_msb_le: "x <=s y \ (msb y \ msb x) \ ((msb x \ \ msb y) \ x \ y)" apply (simp add: word_sle_eq word_sint_msb_eq word_size word_le_def) apply safe apply (rule order_trans[OF _ uint_ge_0]) apply (simp add: order_less_imp_le) apply (erule notE[OF leD]) apply (rule order_less_le_trans[OF _ uint_ge_0]) apply simp done lemma word_sless_msb_less: "x (msb y \ msb x) \ ((msb x \ \ msb y) \ x < y)" by (auto simp add: word_sless_eq word_sle_msb_le) lemma not_msb_from_less: "(v :: 'a word) < 2 ^ (LENGTH('a :: len) - 1) \ \ msb v" apply (clarsimp simp add: msb_nth) apply (drule less_mask_eq) apply (drule word_eqD, drule(1) iffD2) apply (simp add: bit_simps) done lemma sint_eq_uint: "\ msb x \ sint x = uint x" apply (cases \LENGTH('a)\) apply (simp_all add: msb_word_iff_bit) apply transfer apply (simp add: signed_take_bit_eq_take_bit_minus) done lemma scast_eq_ucast: "\ msb x \ scast x = ucast x" apply (cases \LENGTH('a)\) apply simp apply (rule bit_word_eqI) apply (auto simp add: bit_signed_iff bit_unsigned_iff min_def msb_word_eq) apply (erule notE) apply (metis le_less_Suc_eq test_bit_bin) done lemma msb_ucast_eq: "LENGTH('a) = LENGTH('b) \ msb (ucast x :: ('a::len) word) = msb (x :: ('b::len) word)" by (simp add: msb_word_eq bit_simps) lemma msb_big: \msb a \ 2 ^ (LENGTH('a) - Suc 0) \ a\ for a :: \'a::len word\ using bang_is_le [of a \LENGTH('a) - Suc 0\] apply (auto simp add: msb_nth word_le_not_less) apply (rule ccontr) apply (erule notE) apply (rule ccontr) - apply (clarsimp simp: not_less) - apply (subgoal_tac "a = a AND mask (LENGTH('a) - Suc 0)") + apply (clarsimp simp: not_less) + apply (subgoal_tac "a = take_bit (LENGTH('a) - Suc 0) a") apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"]) - apply clarsimp - apply simp + apply auto apply (simp flip: take_bit_eq_mask) apply (rule sym) apply (simp add: take_bit_eq_self_iff_drop_bit_eq_0 drop_bit_eq_zero_iff_not_bit_last) done end diff --git a/thys/Word_Lib/Reversed_Bit_Lists.thy b/thys/Word_Lib/Reversed_Bit_Lists.thy --- a/thys/Word_Lib/Reversed_Bit_Lists.thy +++ b/thys/Word_Lib/Reversed_Bit_Lists.thy @@ -1,2228 +1,2240 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson, NICTA *) section \Bit values as reversed lists of bools\ theory Reversed_Bit_Lists imports "HOL-Library.Word" Typedef_Morphisms Least_significant_bit Most_significant_bit Even_More_List "HOL-Library.Sublist" Aligned Legacy_Aliases begin +context + includes bit_operations_syntax +begin + lemma horner_sum_of_bool_2_concat: \horner_sum of_bool 2 (concat (map (\x. map (bit x) [0.. for ws :: \'a::len word list\ proof (induction ws) case Nil then show ?case by simp next case (Cons w ws) moreover have \horner_sum of_bool 2 (map (bit w) [0.. proof transfer fix k :: int have \map (\n. n < LENGTH('a) \ bit k n) [0.. by simp then show \horner_sum of_bool 2 (map (\n. n < LENGTH('a) \ bit k n) [0.. by (simp only: horner_sum_bit_eq_take_bit) qed ultimately show ?case by (simp add: horner_sum_append) qed subsection \Implicit augmentation of list prefixes\ primrec takefill :: "'a \ nat \ 'a list \ 'a list" where Z: "takefill fill 0 xs = []" | Suc: "takefill fill (Suc n) xs = (case xs of [] \ fill # takefill fill n xs | y # ys \ y # takefill fill n ys)" lemma nth_takefill: "m < n \ takefill fill n l ! m = (if m < length l then l ! m else fill)" apply (induct n arbitrary: m l) apply clarsimp apply clarsimp apply (case_tac m) apply (simp split: list.split) apply (simp split: list.split) done lemma takefill_alt: "takefill fill n l = take n l @ replicate (n - length l) fill" by (induct n arbitrary: l) (auto split: list.split) lemma takefill_replicate [simp]: "takefill fill n (replicate m fill) = replicate n fill" by (simp add: takefill_alt replicate_add [symmetric]) lemma takefill_le': "n = m + k \ takefill x m (takefill x n l) = takefill x m l" by (induct m arbitrary: l n) (auto split: list.split) lemma length_takefill [simp]: "length (takefill fill n l) = n" by (simp add: takefill_alt) lemma take_takefill': "n = k + m \ take k (takefill fill n w) = takefill fill k w" by (induct k arbitrary: w n) (auto split: list.split) lemma drop_takefill: "drop k (takefill fill (m + k) w) = takefill fill m (drop k w)" by (induct k arbitrary: w) (auto split: list.split) lemma takefill_le [simp]: "m \ n \ takefill x m (takefill x n l) = takefill x m l" by (auto simp: le_iff_add takefill_le') lemma take_takefill [simp]: "m \ n \ take m (takefill fill n w) = takefill fill m w" by (auto simp: le_iff_add take_takefill') lemma takefill_append: "takefill fill (m + length xs) (xs @ w) = xs @ (takefill fill m w)" by (induct xs) auto lemma takefill_same': "l = length xs \ takefill fill l xs = xs" by (induct xs arbitrary: l) auto lemmas takefill_same [simp] = takefill_same' [OF refl] lemma tf_rev: "n + k = m + length bl \ takefill x m (rev (takefill y n bl)) = rev (takefill y m (rev (takefill x k (rev bl))))" apply (rule nth_equalityI) apply (auto simp add: nth_takefill rev_nth) apply (rule_tac f = "\n. bl ! n" in arg_cong) apply arith done lemma takefill_minus: "0 < n \ takefill fill (Suc (n - 1)) w = takefill fill n w" by auto lemmas takefill_Suc_cases = list.cases [THEN takefill.Suc [THEN trans]] lemmas takefill_Suc_Nil = takefill_Suc_cases (1) lemmas takefill_Suc_Cons = takefill_Suc_cases (2) lemmas takefill_minus_simps = takefill_Suc_cases [THEN [2] takefill_minus [symmetric, THEN trans]] lemma takefill_numeral_Nil [simp]: "takefill fill (numeral k) [] = fill # takefill fill (pred_numeral k) []" by (simp add: numeral_eq_Suc) lemma takefill_numeral_Cons [simp]: "takefill fill (numeral k) (x # xs) = x # takefill fill (pred_numeral k) xs" by (simp add: numeral_eq_Suc) subsection \Range projection\ definition bl_of_nth :: "nat \ (nat \ 'a) \ 'a list" where "bl_of_nth n f = map f (rev [0.. rev (bl_of_nth n f) ! m = f m" by (simp add: bl_of_nth_def rev_map) lemma bl_of_nth_inj: "(\k. k < n \ f k = g k) \ bl_of_nth n f = bl_of_nth n g" by (simp add: bl_of_nth_def) lemma bl_of_nth_nth_le: "n \ length xs \ bl_of_nth n (nth (rev xs)) = drop (length xs - n) xs" apply (induct n arbitrary: xs) apply clarsimp apply clarsimp apply (rule trans [OF _ hd_Cons_tl]) apply (frule Suc_le_lessD) apply (simp add: rev_nth trans [OF drop_Suc drop_tl, symmetric]) apply (subst hd_drop_conv_nth) apply force apply simp_all apply (rule_tac f = "\n. drop n xs" in arg_cong) apply simp done lemma bl_of_nth_nth [simp]: "bl_of_nth (length xs) ((!) (rev xs)) = xs" by (simp add: bl_of_nth_nth_le) subsection \More\ definition rotater1 :: "'a list \ 'a list" where "rotater1 ys = (case ys of [] \ [] | x # xs \ last ys # butlast ys)" definition rotater :: "nat \ 'a list \ 'a list" where "rotater n = rotater1 ^^ n" lemmas rotater_0' [simp] = rotater_def [where n = "0", simplified] lemma rotate1_rl': "rotater1 (l @ [a]) = a # l" by (cases l) (auto simp: rotater1_def) lemma rotate1_rl [simp] : "rotater1 (rotate1 l) = l" apply (unfold rotater1_def) apply (cases "l") apply (case_tac [2] "list") apply auto done lemma rotate1_lr [simp] : "rotate1 (rotater1 l) = l" by (cases l) (auto simp: rotater1_def) lemma rotater1_rev': "rotater1 (rev xs) = rev (rotate1 xs)" by (cases "xs") (simp add: rotater1_def, simp add: rotate1_rl') lemma rotater_rev': "rotater n (rev xs) = rev (rotate n xs)" by (induct n) (auto simp: rotater_def intro: rotater1_rev') lemma rotater_rev: "rotater n ys = rev (rotate n (rev ys))" using rotater_rev' [where xs = "rev ys"] by simp lemma rotater_drop_take: "rotater n xs = drop (length xs - n mod length xs) xs @ take (length xs - n mod length xs) xs" by (auto simp: rotater_rev rotate_drop_take rev_take rev_drop) lemma rotater_Suc [simp]: "rotater (Suc n) xs = rotater1 (rotater n xs)" unfolding rotater_def by auto lemma nth_rotater: \rotater m xs ! n = xs ! ((n + (length xs - m mod length xs)) mod length xs)\ if \n < length xs\ using that by (simp add: rotater_drop_take nth_append not_less less_diff_conv ac_simps le_mod_geq) lemma nth_rotater1: \rotater1 xs ! n = xs ! ((n + (length xs - 1)) mod length xs)\ if \n < length xs\ using that nth_rotater [of n xs 1] by simp lemma rotate_inv_plus [rule_format]: "\k. k = m + n \ rotater k (rotate n xs) = rotater m xs \ rotate k (rotater n xs) = rotate m xs \ rotater n (rotate k xs) = rotate m xs \ rotate n (rotater k xs) = rotater m xs" by (induct n) (auto simp: rotater_def rotate_def intro: funpow_swap1 [THEN trans]) lemmas rotate_inv_rel = le_add_diff_inverse2 [symmetric, THEN rotate_inv_plus] lemmas rotate_inv_eq = order_refl [THEN rotate_inv_rel, simplified] lemmas rotate_lr [simp] = rotate_inv_eq [THEN conjunct1] lemmas rotate_rl [simp] = rotate_inv_eq [THEN conjunct2, THEN conjunct1] lemma rotate_gal: "rotater n xs = ys \ rotate n ys = xs" by auto lemma rotate_gal': "ys = rotater n xs \ xs = rotate n ys" by auto lemma length_rotater [simp]: "length (rotater n xs) = length xs" by (simp add : rotater_rev) lemma rotate_eq_mod: "m mod length xs = n mod length xs \ rotate m xs = rotate n xs" apply (rule box_equals) defer apply (rule rotate_conv_mod [symmetric])+ apply simp done lemma restrict_to_left: "x = y \ x = z \ y = z" by simp lemmas rotate_eqs = trans [OF rotate0 [THEN fun_cong] id_apply] rotate_rotate [symmetric] rotate_id rotate_conv_mod rotate_eq_mod lemmas rrs0 = rotate_eqs [THEN restrict_to_left, simplified rotate_gal [symmetric] rotate_gal' [symmetric]] lemmas rrs1 = rrs0 [THEN refl [THEN rev_iffD1]] lemmas rotater_eqs = rrs1 [simplified length_rotater] lemmas rotater_0 = rotater_eqs (1) lemmas rotater_add = rotater_eqs (2) lemma butlast_map: "xs \ [] \ butlast (map f xs) = map f (butlast xs)" by (induct xs) auto lemma rotater1_map: "rotater1 (map f xs) = map f (rotater1 xs)" by (cases xs) (auto simp: rotater1_def last_map butlast_map) lemma rotater_map: "rotater n (map f xs) = map f (rotater n xs)" by (induct n) (auto simp: rotater_def rotater1_map) lemma but_last_zip [rule_format] : "\ys. length xs = length ys \ xs \ [] \ last (zip xs ys) = (last xs, last ys) \ butlast (zip xs ys) = zip (butlast xs) (butlast ys)" apply (induct xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ done lemma but_last_map2 [rule_format] : "\ys. length xs = length ys \ xs \ [] \ last (map2 f xs ys) = f (last xs) (last ys) \ butlast (map2 f xs ys) = map2 f (butlast xs) (butlast ys)" apply (induct xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ done lemma rotater1_zip: "length xs = length ys \ rotater1 (zip xs ys) = zip (rotater1 xs) (rotater1 ys)" apply (unfold rotater1_def) apply (cases xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv but_last_zip)[1])+ done lemma rotater1_map2: "length xs = length ys \ rotater1 (map2 f xs ys) = map2 f (rotater1 xs) (rotater1 ys)" by (simp add: rotater1_map rotater1_zip) lemmas lrth = box_equals [OF asm_rl length_rotater [symmetric] length_rotater [symmetric], THEN rotater1_map2] lemma rotater_map2: "length xs = length ys \ rotater n (map2 f xs ys) = map2 f (rotater n xs) (rotater n ys)" by (induct n) (auto intro!: lrth) lemma rotate1_map2: "length xs = length ys \ rotate1 (map2 f xs ys) = map2 f (rotate1 xs) (rotate1 ys)" by (cases xs; cases ys) auto lemmas lth = box_equals [OF asm_rl length_rotate [symmetric] length_rotate [symmetric], THEN rotate1_map2] lemma rotate_map2: "length xs = length ys \ rotate n (map2 f xs ys) = map2 f (rotate n xs) (rotate n ys)" by (induct n) (auto intro!: lth) subsection \Explicit bit representation of \<^typ>\int\\ primrec bl_to_bin_aux :: "bool list \ int \ int" where Nil: "bl_to_bin_aux [] w = w" | Cons: "bl_to_bin_aux (b # bs) w = bl_to_bin_aux bs (of_bool b + 2 * w)" definition bl_to_bin :: "bool list \ int" where "bl_to_bin bs = bl_to_bin_aux bs 0" primrec bin_to_bl_aux :: "nat \ int \ bool list \ bool list" where Z: "bin_to_bl_aux 0 w bl = bl" | Suc: "bin_to_bl_aux (Suc n) w bl = bin_to_bl_aux n (w div 2) (odd w # bl)" definition bin_to_bl :: "nat \ int \ bool list" where "bin_to_bl n w = bin_to_bl_aux n w []" lemma bin_to_bl_aux_zero_minus_simp [simp]: "0 < n \ bin_to_bl_aux n 0 bl = bin_to_bl_aux (n - 1) 0 (False # bl)" by (cases n) auto lemma bin_to_bl_aux_minus1_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (- 1) bl = bin_to_bl_aux (n - 1) (- 1) (True # bl)" by (cases n) auto lemma bin_to_bl_aux_one_minus_simp [simp]: "0 < n \ bin_to_bl_aux n 1 bl = bin_to_bl_aux (n - 1) 0 (True # bl)" by (cases n) auto lemma bin_to_bl_aux_Bit0_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (numeral (Num.Bit0 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (False # bl)" by (cases n) simp_all lemma bin_to_bl_aux_Bit1_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (numeral (Num.Bit1 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (True # bl)" by (cases n) simp_all lemma bl_to_bin_aux_append: "bl_to_bin_aux (bs @ cs) w = bl_to_bin_aux cs (bl_to_bin_aux bs w)" by (induct bs arbitrary: w) auto lemma bin_to_bl_aux_append: "bin_to_bl_aux n w bs @ cs = bin_to_bl_aux n w (bs @ cs)" by (induct n arbitrary: w bs) auto lemma bl_to_bin_append: "bl_to_bin (bs @ cs) = bl_to_bin_aux cs (bl_to_bin bs)" unfolding bl_to_bin_def by (rule bl_to_bin_aux_append) lemma bin_to_bl_aux_alt: "bin_to_bl_aux n w bs = bin_to_bl n w @ bs" by (simp add: bin_to_bl_def bin_to_bl_aux_append) lemma bin_to_bl_0 [simp]: "bin_to_bl 0 bs = []" by (auto simp: bin_to_bl_def) lemma size_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" by (induct n arbitrary: w bs) auto lemma size_bin_to_bl [simp]: "length (bin_to_bl n w) = n" by (simp add: bin_to_bl_def size_bin_to_bl_aux) lemma bl_bin_bl': "bin_to_bl (n + length bs) (bl_to_bin_aux bs w) = bin_to_bl_aux n w bs" apply (induct bs arbitrary: w n) apply auto apply (simp_all only: add_Suc [symmetric]) apply (auto simp add: bin_to_bl_def) done lemma bl_bin_bl [simp]: "bin_to_bl (length bs) (bl_to_bin bs) = bs" unfolding bl_to_bin_def apply (rule box_equals) apply (rule bl_bin_bl') prefer 2 apply (rule bin_to_bl_aux.Z) apply simp done lemma bl_to_bin_inj: "bl_to_bin bs = bl_to_bin cs \ length bs = length cs \ bs = cs" apply (rule_tac box_equals) defer apply (rule bl_bin_bl) apply (rule bl_bin_bl) apply simp done lemma bl_to_bin_False [simp]: "bl_to_bin (False # bl) = bl_to_bin bl" by (auto simp: bl_to_bin_def) lemma bl_to_bin_Nil [simp]: "bl_to_bin [] = 0" by (auto simp: bl_to_bin_def) lemma bin_to_bl_zero_aux: "bin_to_bl_aux n 0 bl = replicate n False @ bl" by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) lemma bin_to_bl_zero: "bin_to_bl n 0 = replicate n False" by (simp add: bin_to_bl_def bin_to_bl_zero_aux) lemma bin_to_bl_minus1_aux: "bin_to_bl_aux n (- 1) bl = replicate n True @ bl" by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) lemma bin_to_bl_minus1: "bin_to_bl n (- 1) = replicate n True" by (simp add: bin_to_bl_def bin_to_bl_minus1_aux) subsection \Semantic interpretation of \<^typ>\bool list\ as \<^typ>\int\\ lemma bin_bl_bin': "bl_to_bin (bin_to_bl_aux n w bs) = bl_to_bin_aux bs (take_bit n w)" by (induct n arbitrary: w bs) (auto simp: bl_to_bin_def take_bit_Suc ac_simps mod_2_eq_odd) lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = take_bit n w" by (auto simp: bin_to_bl_def bin_bl_bin') lemma bl_to_bin_rep_F: "bl_to_bin (replicate n False @ bl) = bl_to_bin bl" by (simp add: bin_to_bl_zero_aux [symmetric] bin_bl_bin') (simp add: bl_to_bin_def) lemma bin_to_bl_trunc [simp]: "n \ m \ bin_to_bl n (take_bit m w) = bin_to_bl n w" by (auto intro: bl_to_bin_inj) lemma bin_to_bl_aux_bintr: "bin_to_bl_aux n (take_bit m bin) bl = replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl" apply (induct n arbitrary: m bin bl) apply clarsimp apply clarsimp apply (case_tac "m") apply (clarsimp simp: bin_to_bl_zero_aux) apply (erule thin_rl) apply (induct_tac n) apply (auto simp add: take_bit_Suc) done lemma bin_to_bl_bintr: "bin_to_bl n (take_bit m bin) = replicate (n - m) False @ bin_to_bl (min n m) bin" unfolding bin_to_bl_def by (rule bin_to_bl_aux_bintr) lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = 0" by (induct n) auto lemma len_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" by (fact size_bin_to_bl_aux) lemma len_bin_to_bl: "length (bin_to_bl n w) = n" by (fact size_bin_to_bl) (* FIXME: duplicate *) lemma sign_bl_bin': "bin_sign (bl_to_bin_aux bs w) = bin_sign w" by (induction bs arbitrary: w) (simp_all add: bin_sign_def) lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = 0" by (simp add: bl_to_bin_def sign_bl_bin') lemma bl_sbin_sign_aux: "hd (bin_to_bl_aux (Suc n) w bs) = (bin_sign (signed_take_bit n w) = -1)" by (induction n arbitrary: w bs) (auto simp add: bin_sign_def even_iff_mod_2_eq_zero bit_Suc) lemma bl_sbin_sign: "hd (bin_to_bl (Suc n) w) = (bin_sign (signed_take_bit n w) = -1)" unfolding bin_to_bl_def by (rule bl_sbin_sign_aux) lemma bin_nth_of_bl_aux: "bit (bl_to_bin_aux bl w) n = (n < size bl \ rev bl ! n \ n \ length bl \ bit w (n - size bl))" apply (induction bl arbitrary: w) apply simp_all apply safe apply (simp_all add: not_le nth_append bit_double_iff even_bit_succ_iff split: if_splits) done lemma bin_nth_of_bl: "bit (bl_to_bin bl) n = (n < length bl \ rev bl ! n)" by (simp add: bl_to_bin_def bin_nth_of_bl_aux) lemma bin_nth_bl: "n < m \ bit w n = nth (rev (bin_to_bl m w)) n" apply (induct n arbitrary: m w) apply clarsimp apply (case_tac m, clarsimp) apply (clarsimp simp: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt) apply (case_tac m, clarsimp) apply (clarsimp simp: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt bit_Suc) done lemma nth_bin_to_bl_aux: "n < m + length bl \ (bin_to_bl_aux m w bl) ! n = (if n < m then bit w (m - 1 - n) else bl ! (n - m))" apply (induction bl arbitrary: w) apply simp_all apply (simp add: bin_nth_bl [of \m - Suc n\ m] rev_nth flip: bin_to_bl_def) apply (metis One_nat_def Suc_pred add_diff_cancel_left' add_diff_cancel_right' bin_to_bl_aux_alt bin_to_bl_def diff_Suc_Suc diff_is_0_eq diff_zero less_Suc_eq_0_disj less_antisym less_imp_Suc_add list.size(3) nat_less_le nth_append size_bin_to_bl_aux) done lemma nth_bin_to_bl: "n < m \ (bin_to_bl m w) ! n = bit w (m - Suc n)" by (simp add: bin_to_bl_def nth_bin_to_bl_aux) lemma takefill_bintrunc: "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))" apply (rule nth_equalityI) apply simp apply (clarsimp simp: nth_takefill rev_nth nth_bin_to_bl bin_nth_of_bl) done lemma bl_bin_bl_rtf: "bin_to_bl n (bl_to_bin bl) = rev (takefill False n (rev bl))" by (simp add: takefill_bintrunc) lemma bl_to_bin_lt2p_aux: "bl_to_bin_aux bs w < (w + 1) * (2 ^ length bs)" proof (induction bs arbitrary: w) case Nil then show ?case by simp next case (Cons b bs) from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] show ?case apply (auto simp add: algebra_simps) apply (subst mult_2 [of \2 ^ length bs\]) apply (simp only: add.assoc) apply (rule pos_add_strict) apply simp_all done qed lemma bl_to_bin_lt2p_drop: "bl_to_bin bs < 2 ^ length (dropWhile Not bs)" proof (induct bs) case Nil then show ?case by simp next case (Cons b bs) with bl_to_bin_lt2p_aux[where w=1] show ?case by (simp add: bl_to_bin_def) qed lemma bl_to_bin_lt2p: "bl_to_bin bs < 2 ^ length bs" by (metis bin_bl_bin bintr_lt2p bl_bin_bl) lemma bl_to_bin_ge2p_aux: "bl_to_bin_aux bs w \ w * (2 ^ length bs)" proof (induction bs arbitrary: w) case Nil then show ?case by simp next case (Cons b bs) from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] show ?case apply (auto simp add: algebra_simps) apply (rule add_le_imp_le_left [of \2 ^ length bs\]) apply (rule add_increasing) apply simp_all done qed lemma bl_to_bin_ge0: "bl_to_bin bs \ 0" apply (unfold bl_to_bin_def) apply (rule xtrans(4)) apply (rule bl_to_bin_ge2p_aux) apply simp done lemma butlast_rest_bin: "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (w div 2)" apply (unfold bin_to_bl_def) apply (cases n, clarsimp) apply clarsimp apply (auto simp add: bin_to_bl_aux_alt) done lemma butlast_bin_rest: "butlast bl = bin_to_bl (length bl - Suc 0) (bl_to_bin bl div 2)" using butlast_rest_bin [where w="bl_to_bin bl" and n="length bl"] by simp lemma butlast_rest_bl2bin_aux: "bl \ [] \ bl_to_bin_aux (butlast bl) w = bl_to_bin_aux bl w div 2" by (induct bl arbitrary: w) auto lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bl_to_bin bl div 2" by (cases bl) (auto simp: bl_to_bin_def butlast_rest_bl2bin_aux) lemma trunc_bl2bin_aux: "take_bit m (bl_to_bin_aux bl w) = bl_to_bin_aux (drop (length bl - m) bl) (take_bit (m - length bl) w)" proof (induct bl arbitrary: w) case Nil show ?case by simp next case (Cons b bl) show ?case proof (cases "m - length bl") case 0 then have "Suc (length bl) - m = Suc (length bl - m)" by simp with Cons show ?thesis by simp next case (Suc n) then have "m - Suc (length bl) = n" by simp with Cons Suc show ?thesis by (simp add: take_bit_Suc ac_simps) qed qed lemma trunc_bl2bin: "take_bit m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" by (simp add: bl_to_bin_def trunc_bl2bin_aux) lemma trunc_bl2bin_len [simp]: "take_bit (length bl) (bl_to_bin bl) = bl_to_bin bl" by (simp add: trunc_bl2bin) lemma bl2bin_drop: "bl_to_bin (drop k bl) = take_bit (length bl - k) (bl_to_bin bl)" apply (rule trans) prefer 2 apply (rule trunc_bl2bin [symmetric]) apply (cases "k \ length bl") apply auto done lemma take_rest_power_bin: "m \ n \ take m (bin_to_bl n w) = bin_to_bl m (((\w. w div 2) ^^ (n - m)) w)" apply (rule nth_equalityI) apply simp apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin) done lemma last_bin_last': "size xs > 0 \ last xs \ odd (bl_to_bin_aux xs w)" by (induct xs arbitrary: w) auto lemma last_bin_last: "size xs > 0 \ last xs \ odd (bl_to_bin xs)" unfolding bl_to_bin_def by (erule last_bin_last') lemma bin_last_last: "odd w \ last (bin_to_bl (Suc n) w)" by (simp add: bin_to_bl_def) (auto simp: bin_to_bl_aux_alt) lemma drop_bin2bl_aux: "drop m (bin_to_bl_aux n bin bs) = bin_to_bl_aux (n - m) bin (drop (m - n) bs)" apply (induction n arbitrary: m bin bs) apply auto apply (case_tac "m \ n") apply (auto simp add: not_le Suc_diff_le) apply (case_tac "m - n") apply auto apply (use Suc_diff_Suc in fastforce) done lemma drop_bin2bl: "drop m (bin_to_bl n bin) = bin_to_bl (n - m) bin" by (simp add: bin_to_bl_def drop_bin2bl_aux) lemma take_bin2bl_lem1: "take m (bin_to_bl_aux m w bs) = bin_to_bl m w" apply (induct m arbitrary: w bs) apply clarsimp apply clarsimp apply (simp add: bin_to_bl_aux_alt) apply (simp add: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt) done lemma take_bin2bl_lem: "take m (bin_to_bl_aux (m + n) w bs) = take m (bin_to_bl (m + n) w)" by (induct n arbitrary: w bs) (simp_all (no_asm) add: bin_to_bl_def take_bin2bl_lem1, simp) lemma bin_split_take: "bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl (m + n) c)" apply (induct n arbitrary: b c) apply clarsimp apply (clarsimp simp: Let_def split: prod.split_asm) apply (simp add: bin_to_bl_def) apply (simp add: take_bin2bl_lem drop_bit_Suc) done lemma bin_to_bl_drop_bit: "k = m + n \ bin_to_bl m (drop_bit n c) = take m (bin_to_bl k c)" using bin_split_take by simp lemma bin_split_take1: "k = m + n \ bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl k c)" using bin_split_take by simp lemma bl_bin_bl_rep_drop: "bin_to_bl n (bl_to_bin bl) = replicate (n - length bl) False @ drop (length bl - n) bl" by (simp add: bl_to_bin_inj bl_to_bin_rep_F trunc_bl2bin) lemma bl_to_bin_aux_cat: "bl_to_bin_aux bs (concat_bit nv v w) = concat_bit (nv + length bs) (bl_to_bin_aux bs v) w" by (rule bit_eqI) (auto simp add: bin_nth_of_bl_aux bin_nth_cat algebra_simps) lemma bin_to_bl_aux_cat: "bin_to_bl_aux (nv + nw) (concat_bit nw w v) bs = bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)" by (induction nw arbitrary: w bs) (simp_all add: concat_bit_Suc) lemma bl_to_bin_aux_alt: "bl_to_bin_aux bs w = concat_bit (length bs) (bl_to_bin bs) w" using bl_to_bin_aux_cat [where nv = "0" and v = "0"] by (simp add: bl_to_bin_def [symmetric]) lemma bin_to_bl_cat: "bin_to_bl (nv + nw) (concat_bit nw w v) = bin_to_bl_aux nv v (bin_to_bl nw w)" by (simp add: bin_to_bl_def bin_to_bl_aux_cat) lemmas bl_to_bin_aux_app_cat = trans [OF bl_to_bin_aux_append bl_to_bin_aux_alt] lemmas bin_to_bl_aux_cat_app = trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt] lemma bl_to_bin_app_cat: "bl_to_bin (bsa @ bs) = concat_bit (length bs) (bl_to_bin bs) (bl_to_bin bsa)" by (simp only: bl_to_bin_aux_app_cat bl_to_bin_def) lemma bin_to_bl_cat_app: "bin_to_bl (n + nw) (concat_bit nw wa w) = bin_to_bl n w @ bin_to_bl nw wa" by (simp only: bin_to_bl_def bin_to_bl_aux_cat_app) text \\bl_to_bin_app_cat_alt\ and \bl_to_bin_app_cat\ are easily interderivable.\ lemma bl_to_bin_app_cat_alt: "concat_bit n w (bl_to_bin cs) = bl_to_bin (cs @ bin_to_bl n w)" by (simp add: bl_to_bin_app_cat) lemma mask_lem: "(bl_to_bin (True # replicate n False)) = bl_to_bin (replicate n True) + 1" apply (unfold bl_to_bin_def) apply (induct n) apply simp apply (simp only: Suc_eq_plus1 replicate_add append_Cons [symmetric] bl_to_bin_aux_append) apply simp done lemma bin_exhaust: "(\x b. bin = of_bool b + 2 * x \ Q) \ Q" for bin :: int apply (cases \even bin\) apply (auto elim!: evenE oddE) apply fastforce apply fastforce done primrec rbl_succ :: "bool list \ bool list" where Nil: "rbl_succ Nil = Nil" | Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)" primrec rbl_pred :: "bool list \ bool list" where Nil: "rbl_pred Nil = Nil" | Cons: "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)" primrec rbl_add :: "bool list \ bool list \ bool list" where \ \result is length of first arg, second arg may be longer\ Nil: "rbl_add Nil x = Nil" | Cons: "rbl_add (y # ys) x = (let ws = rbl_add ys (tl x) in (y \ hd x) # (if hd x \ y then rbl_succ ws else ws))" primrec rbl_mult :: "bool list \ bool list \ bool list" where \ \result is length of first arg, second arg may be longer\ Nil: "rbl_mult Nil x = Nil" | Cons: "rbl_mult (y # ys) x = (let ws = False # rbl_mult ys x in if y then rbl_add ws x else ws)" lemma size_rbl_pred: "length (rbl_pred bl) = length bl" by (induct bl) auto lemma size_rbl_succ: "length (rbl_succ bl) = length bl" by (induct bl) auto lemma size_rbl_add: "length (rbl_add bl cl) = length bl" by (induct bl arbitrary: cl) (auto simp: Let_def size_rbl_succ) lemma size_rbl_mult: "length (rbl_mult bl cl) = length bl" by (induct bl arbitrary: cl) (auto simp add: Let_def size_rbl_add) lemmas rbl_sizes [simp] = size_rbl_pred size_rbl_succ size_rbl_add size_rbl_mult lemmas rbl_Nils = rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil lemma rbl_add_app2: "length blb \ length bla \ rbl_add bla (blb @ blc) = rbl_add bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def) done lemma rbl_add_take2: "length blb \ length bla \ rbl_add bla (take (length bla) blb) = rbl_add bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def) done lemma rbl_mult_app2: "length blb \ length bla \ rbl_mult bla (blb @ blc) = rbl_mult bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def rbl_add_app2) done lemma rbl_mult_take2: "length blb \ length bla \ rbl_mult bla (take (length bla) blb) = rbl_mult bla blb" apply (rule trans) apply (rule rbl_mult_app2 [symmetric]) apply simp apply (rule_tac f = "rbl_mult bla" in arg_cong) apply (rule append_take_drop_id) done lemma rbl_add_split: "P (rbl_add (y # ys) (x # xs)) = (\ws. length ws = length ys \ ws = rbl_add ys xs \ (y \ ((x \ P (False # rbl_succ ws)) \ (\ x \ P (True # ws)))) \ (\ y \ P (x # ws)))" by (cases y) (auto simp: Let_def) lemma rbl_mult_split: "P (rbl_mult (y # ys) xs) = (\ws. length ws = Suc (length ys) \ ws = False # rbl_mult ys xs \ (y \ P (rbl_add ws xs)) \ (\ y \ P ws))" by (auto simp: Let_def) lemma rbl_pred: "rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin - 1))" proof (unfold bin_to_bl_def, induction n arbitrary: bin) case 0 then show ?case by simp next case (Suc n) obtain b k where \bin = of_bool b + 2 * k\ using bin_exhaust by blast moreover have \(2 * k - 1) div 2 = k - 1\ using even_succ_div_2 [of \2 * (k - 1)\] by simp ultimately show ?case using Suc [of \bin div 2\] by simp (auto simp add: bin_to_bl_aux_alt) qed lemma rbl_succ: "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))" apply (unfold bin_to_bl_def) apply (induction n arbitrary: bin) apply simp_all apply (case_tac bin rule: bin_exhaust) apply (simp_all add: bin_to_bl_aux_alt ac_simps) done lemma rbl_add: "\bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = rev (bin_to_bl n (bina + binb))" apply (unfold bin_to_bl_def) apply (induct n) apply simp apply clarsimp apply (case_tac bina rule: bin_exhaust) apply (case_tac binb rule: bin_exhaust) apply (case_tac b) apply (case_tac [!] "ba") apply (auto simp: rbl_succ bin_to_bl_aux_alt Let_def ac_simps) done lemma rbl_add_long: "m \ n \ rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = rev (bin_to_bl n (bina + binb))" apply (rule box_equals [OF _ rbl_add_take2 rbl_add]) apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong) apply (rule rev_swap [THEN iffD1]) apply (simp add: rev_take drop_bin2bl) apply simp done lemma rbl_mult_gt1: "m \ length bl \ rbl_mult bl (rev (bin_to_bl m binb)) = rbl_mult bl (rev (bin_to_bl (length bl) binb))" apply (rule trans) apply (rule rbl_mult_take2 [symmetric]) apply simp_all apply (rule_tac f = "rbl_mult bl" in arg_cong) apply (rule rev_swap [THEN iffD1]) apply (simp add: rev_take drop_bin2bl) done lemma rbl_mult_gt: "m > n \ rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))" by (auto intro: trans [OF rbl_mult_gt1]) lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt] lemma rbbl_Cons: "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (of_bool b + 2 * x))" by (simp add: bin_to_bl_def) (simp add: bin_to_bl_aux_alt) lemma rbl_mult: "rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = rev (bin_to_bl n (bina * binb))" apply (induct n arbitrary: bina binb) apply simp_all apply (unfold bin_to_bl_def) apply clarsimp apply (case_tac bina rule: bin_exhaust) apply (case_tac binb rule: bin_exhaust) apply (simp_all add: bin_to_bl_aux_alt) apply (simp_all add: rbbl_Cons rbl_mult_Suc rbl_add algebra_simps) done lemma sclem: "size (concat (map (bin_to_bl n) xs)) = length xs * n" by (simp add: length_concat comp_def sum_list_triv) lemma bin_cat_foldl_lem: "foldl (\u k. concat_bit n k u) x xs = concat_bit (size xs * n) (foldl (\u k. concat_bit n k u) y xs) x" apply (induct xs arbitrary: x) apply simp apply (simp (no_asm)) apply (frule asm_rl) apply (drule meta_spec) apply (erule trans) apply (drule_tac x = "concat_bit n a y" in meta_spec) apply (simp add: bin_cat_assoc_sym) done lemma bin_rcat_bl: "bin_rcat n wl = bl_to_bin (concat (map (bin_to_bl n) wl))" apply (unfold bin_rcat_eq_foldl) apply (rule sym) apply (induct wl) apply (auto simp add: bl_to_bin_append) apply (simp add: bl_to_bin_aux_alt sclem) apply (simp add: bin_cat_foldl_lem [symmetric]) done lemma bin_last_bl_to_bin: "odd (bl_to_bin bs) \ bs \ [] \ last bs" by(cases "bs = []")(auto simp add: bl_to_bin_def last_bin_last'[where w=0]) lemma bin_rest_bl_to_bin: "bl_to_bin bs div 2 = bl_to_bin (butlast bs)" by(cases "bs = []")(simp_all add: bl_to_bin_def butlast_rest_bl2bin_aux) lemma bl_xor_aux_bin: "map2 (\x y. x \ y) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v XOR w) (map2 (\x y. x \ y) bs cs)" apply (induction n arbitrary: v w bs cs) apply auto apply (case_tac v rule: bin_exhaust) apply (case_tac w rule: bin_exhaust) apply clarsimp done lemma bl_or_aux_bin: "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v OR w) (map2 (\) bs cs)" by (induct n arbitrary: v w bs cs) simp_all lemma bl_and_aux_bin: "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v AND w) (map2 (\) bs cs)" by (induction n arbitrary: v w bs cs) simp_all lemma bl_not_aux_bin: "map Not (bin_to_bl_aux n w cs) = bin_to_bl_aux n (NOT w) (map Not cs)" by (induct n arbitrary: w cs) auto lemma bl_not_bin: "map Not (bin_to_bl n w) = bin_to_bl n (NOT w)" by (simp add: bin_to_bl_def bl_not_aux_bin) lemma bl_and_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v AND w)" by (simp add: bin_to_bl_def bl_and_aux_bin) lemma bl_or_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v OR w)" by (simp add: bin_to_bl_def bl_or_aux_bin) lemma bl_xor_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v XOR w)" using bl_xor_aux_bin by (simp add: bin_to_bl_def) subsection \Type \<^typ>\'a word\\ lift_definition of_bl :: \bool list \ 'a::len word\ is bl_to_bin . lift_definition to_bl :: \'a::len word \ bool list\ is \bin_to_bl LENGTH('a)\ by (simp add: bl_to_bin_inj) lemma to_bl_eq: \to_bl w = bin_to_bl (LENGTH('a)) (uint w)\ for w :: \'a::len word\ by transfer simp lemma bit_of_bl_iff [bit_simps]: \bit (of_bl bs :: 'a word) n \ rev bs ! n \ n < LENGTH('a::len) \ n < length bs\ by transfer (simp add: bin_nth_of_bl ac_simps) lemma rev_to_bl_eq: \rev (to_bl w) = map (bit w) [0.. for w :: \'a::len word\ apply (rule nth_equalityI) apply (simp add: to_bl.rep_eq) apply (simp add: bin_nth_bl bit_word.rep_eq to_bl.rep_eq) done lemma to_bl_eq_rev: \to_bl w = map (bit w) (rev [0.. for w :: \'a::len word\ using rev_to_bl_eq [of w] apply (subst rev_is_rev_conv [symmetric]) apply (simp add: rev_map) done lemma of_bl_rev_eq: \of_bl (rev bs) = horner_sum of_bool 2 bs\ apply (rule bit_word_eqI) apply (simp add: bit_of_bl_iff) apply transfer apply (simp add: bit_horner_sum_bit_iff ac_simps) done lemma of_bl_eq: \of_bl bs = horner_sum of_bool 2 (rev bs)\ using of_bl_rev_eq [of \rev bs\] by simp lemma bshiftr1_eq: \bshiftr1 b w = of_bl (b # butlast (to_bl w))\ apply (rule bit_word_eqI) apply (auto simp add: bit_simps to_bl_eq_rev nth_append rev_nth nth_butlast not_less simp flip: bit_Suc) apply (metis Suc_pred len_gt_0 less_eq_decr_length_iff not_bit_length verit_la_disequality) done lemma length_to_bl_eq: \length (to_bl w) = LENGTH('a)\ for w :: \'a::len word\ by transfer simp lemma word_rotr_eq: \word_rotr n w = of_bl (rotater n (to_bl w))\ apply (rule bit_word_eqI) subgoal for n apply (cases \n < LENGTH('a)\) apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff rotater_rev length_to_bl_eq nth_rotate rev_to_bl_eq ac_simps) done done lemma word_rotl_eq: \word_rotl n w = of_bl (rotate n (to_bl w))\ proof - have \rotate n (to_bl w) = rev (rotater n (rev (to_bl w)))\ by (simp add: rotater_rev') then show ?thesis apply (simp add: word_rotl_eq_word_rotr bit_of_bl_iff length_to_bl_eq rev_to_bl_eq) apply (rule bit_word_eqI) subgoal for n apply (cases \n < LENGTH('a)\) apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff nth_rotater) done done qed lemma to_bl_def': "(to_bl :: 'a::len word \ bool list) = bin_to_bl (LENGTH('a)) \ uint" by transfer (simp add: fun_eq_iff) \ \type definitions theorem for in terms of equivalent bool list\ lemma td_bl: "type_definition (to_bl :: 'a::len word \ bool list) of_bl {bl. length bl = LENGTH('a)}" apply (standard; transfer) apply (auto dest: sym) done -interpretation word_bl: +global_interpretation word_bl: type_definition "to_bl :: 'a::len word \ bool list" of_bl "{bl. length bl = LENGTH('a::len)}" by (fact td_bl) lemmas word_bl_Rep' = word_bl.Rep [unfolded mem_Collect_eq, iff] lemma word_size_bl: "size w = size (to_bl w)" by (auto simp: word_size) lemma to_bl_use_of_bl: "to_bl w = bl \ w = of_bl bl \ length bl = length (to_bl w)" by (fastforce elim!: word_bl.Abs_inverse [unfolded mem_Collect_eq]) lemma length_bl_gt_0 [iff]: "0 < length (to_bl x)" for x :: "'a::len word" unfolding word_bl_Rep' by (rule len_gt_0) lemma bl_not_Nil [iff]: "to_bl x \ []" for x :: "'a::len word" by (fact length_bl_gt_0 [unfolded length_greater_0_conv]) lemma length_bl_neq_0 [iff]: "length (to_bl x) \ 0" for x :: "'a::len word" by (fact length_bl_gt_0 [THEN gr_implies_not0]) lemma hd_bl_sign_sint: "hd (to_bl w) = (bin_sign (sint w) = -1)" apply transfer apply (auto simp add: bin_sign_def) using bin_sign_lem bl_sbin_sign apply fastforce using bin_sign_lem bl_sbin_sign apply force done lemma of_bl_drop': "lend = length bl - LENGTH('a::len) \ of_bl (drop lend bl) = (of_bl bl :: 'a word)" by transfer (simp flip: trunc_bl2bin) lemma test_bit_of_bl: "bit (of_bl bl::'a::len word) n = (rev bl ! n \ n < LENGTH('a) \ n < length bl)" by transfer (simp add: bin_nth_of_bl ac_simps) lemma no_of_bl: "(numeral bin ::'a::len word) = of_bl (bin_to_bl (LENGTH('a)) (numeral bin))" by transfer simp lemma uint_bl: "to_bl w = bin_to_bl (size w) (uint w)" by transfer simp lemma to_bl_bin: "bl_to_bin (to_bl w) = uint w" by (simp add: uint_bl word_size) lemma to_bl_of_bin: "to_bl (word_of_int bin::'a::len word) = bin_to_bl (LENGTH('a)) bin" by (auto simp: uint_bl word_ubin.eq_norm word_size) lemma to_bl_numeral [simp]: "to_bl (numeral bin::'a::len word) = bin_to_bl (LENGTH('a)) (numeral bin)" unfolding word_numeral_alt by (rule to_bl_of_bin) lemma to_bl_neg_numeral [simp]: "to_bl (- numeral bin::'a::len word) = bin_to_bl (LENGTH('a)) (- numeral bin)" unfolding word_neg_numeral_alt by (rule to_bl_of_bin) lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w" by (simp add: uint_bl word_size) lemma uint_bl_bin: "bl_to_bin (bin_to_bl (LENGTH('a)) (uint x)) = uint x" for x :: "'a::len word" by (rule trans [OF bin_bl_bin word_ubin.norm_Rep]) lemma ucast_bl: "ucast w = of_bl (to_bl w)" by transfer simp lemma ucast_down_bl: \(ucast :: 'a::len word \ 'b::len word) (of_bl bl) = of_bl bl\ if \is_down (ucast :: 'a::len word \ 'b::len word)\ using that by transfer simp lemma of_bl_append_same: "of_bl (X @ to_bl w) = w" by transfer (simp add: bl_to_bin_app_cat) lemma ucast_of_bl_up: \ucast (of_bl bl :: 'a::len word) = of_bl bl\ if \size bl \ size (of_bl bl :: 'a::len word)\ using that apply transfer apply (rule bit_eqI) apply (auto simp add: bit_take_bit_iff) apply (subst (asm) trunc_bl2bin_len [symmetric]) apply (auto simp only: bit_take_bit_iff) done lemma word_rev_tf: "to_bl (of_bl bl::'a::len word) = rev (takefill False (LENGTH('a)) (rev bl))" by transfer (simp add: bl_bin_bl_rtf) lemma word_rep_drop: "to_bl (of_bl bl::'a::len word) = replicate (LENGTH('a) - length bl) False @ drop (length bl - LENGTH('a)) bl" by (simp add: word_rev_tf takefill_alt rev_take) lemma to_bl_ucast: "to_bl (ucast (w::'b::len word) ::'a::len word) = replicate (LENGTH('a) - LENGTH('b)) False @ drop (LENGTH('b) - LENGTH('a)) (to_bl w)" apply (unfold ucast_bl) apply (rule trans) apply (rule word_rep_drop) apply simp done lemma ucast_up_app: \to_bl (ucast w :: 'b::len word) = replicate n False @ (to_bl w)\ if \source_size (ucast :: 'a word \ 'b word) + n = target_size (ucast :: 'a word \ 'b word)\ for w :: \'a::len word\ using that by (auto simp add : source_size target_size to_bl_ucast) lemma ucast_down_drop [OF refl]: "uc = ucast \ source_size uc = target_size uc + n \ to_bl (uc w) = drop n (to_bl w)" by (auto simp add : source_size target_size to_bl_ucast) lemma scast_down_drop [OF refl]: "sc = scast \ source_size sc = target_size sc + n \ to_bl (sc w) = drop n (to_bl w)" apply (subgoal_tac "sc = ucast") apply safe apply simp apply (erule ucast_down_drop) apply (rule down_cast_same [symmetric]) apply (simp add : source_size target_size is_down) done lemma word_0_bl [simp]: "of_bl [] = 0" by transfer simp lemma word_1_bl: "of_bl [True] = 1" by transfer (simp add: bl_to_bin_def) lemma of_bl_0 [simp]: "of_bl (replicate n False) = 0" by transfer (simp add: bl_to_bin_rep_False) lemma to_bl_0 [simp]: "to_bl (0::'a::len word) = replicate (LENGTH('a)) False" by (simp add: uint_bl word_size bin_to_bl_zero) \ \links with \rbl\ operations\ lemma word_succ_rbl: "to_bl w = bl \ to_bl (word_succ w) = rev (rbl_succ (rev bl))" by transfer (simp add: rbl_succ) lemma word_pred_rbl: "to_bl w = bl \ to_bl (word_pred w) = rev (rbl_pred (rev bl))" by transfer (simp add: rbl_pred) lemma word_add_rbl: "to_bl v = vbl \ to_bl w = wbl \ to_bl (v + w) = rev (rbl_add (rev vbl) (rev wbl))" apply transfer apply (drule sym) apply (drule sym) apply (simp add: rbl_add) done lemma word_mult_rbl: "to_bl v = vbl \ to_bl w = wbl \ to_bl (v * w) = rev (rbl_mult (rev vbl) (rev wbl))" apply transfer apply (drule sym) apply (drule sym) apply (simp add: rbl_mult) done lemma rtb_rbl_ariths: "rev (to_bl w) = ys \ rev (to_bl (word_succ w)) = rbl_succ ys" "rev (to_bl w) = ys \ rev (to_bl (word_pred w)) = rbl_pred ys" "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v * w)) = rbl_mult ys xs" "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v + w)) = rbl_add ys xs" by (auto simp: rev_swap [symmetric] word_succ_rbl word_pred_rbl word_mult_rbl word_add_rbl) lemma of_bl_length_less: \(of_bl x :: 'a::len word) < 2 ^ k\ if \length x = k\ \k < LENGTH('a)\ proof - from that have \length x < LENGTH('a)\ by simp then have \(of_bl x :: 'a::len word) < 2 ^ length x\ apply (simp add: of_bl_eq) apply transfer apply (simp add: take_bit_horner_sum_bit_eq) apply (subst length_rev [symmetric]) apply (simp only: horner_sum_of_bool_2_less) done with that show ?thesis by simp qed lemma word_eq_rbl_eq: "x = y \ rev (to_bl x) = rev (to_bl y)" by simp lemma bl_word_not: "to_bl (NOT w) = map Not (to_bl w)" by transfer (simp add: bl_not_bin) lemma bl_word_xor: "to_bl (v XOR w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_xor_bin) lemma bl_word_or: "to_bl (v OR w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_or_bin) lemma bl_word_and: "to_bl (v AND w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_and_bin) lemma bin_nth_uint': "bit (uint w) n \ rev (bin_to_bl (size w) (uint w)) ! n \ n < size w" apply (unfold word_size) apply (safe elim!: bin_nth_uint_imp) apply (frule bin_nth_uint_imp) apply (fast dest!: bin_nth_bl)+ done lemmas bin_nth_uint = bin_nth_uint' [unfolded word_size] lemma test_bit_bl: "bit w n \ rev (to_bl w) ! n \ n < size w" by transfer (auto simp add: bin_nth_bl) lemma to_bl_nth: "n < size w \ to_bl w ! n = bit w (size w - Suc n)" by (simp add: word_size rev_nth test_bit_bl) lemma map_bit_interval_eq: \map (bit w) [0.. for w :: \'a::len word\ proof (rule nth_equalityI) show \length (map (bit w) [0.. by simp fix m assume \m < length (map (bit w) [0.. then have \m < n\ by simp then have \bit w m \ takefill False n (rev (to_bl w)) ! m\ by (auto simp add: nth_takefill not_less rev_nth to_bl_nth word_size dest: bit_imp_le_length) with \m < n \show \map (bit w) [0.. takefill False n (rev (to_bl w)) ! m\ by simp qed lemma to_bl_unfold: \to_bl w = rev (map (bit w) [0.. for w :: \'a::len word\ by (simp add: map_bit_interval_eq takefill_bintrunc to_bl_def flip: bin_to_bl_def) lemma nth_rev_to_bl: \rev (to_bl w) ! n \ bit w n\ if \n < LENGTH('a)\ for w :: \'a::len word\ using that by (simp add: to_bl_unfold) lemma nth_to_bl: \to_bl w ! n \ bit w (LENGTH('a) - Suc n)\ if \n < LENGTH('a)\ for w :: \'a::len word\ using that by (simp add: to_bl_unfold rev_nth) lemma of_bl_rep_False: "of_bl (replicate n False @ bs) = of_bl bs" by (auto simp: of_bl_def bl_to_bin_rep_F) lemma [code abstract]: \Word.the_int (of_bl bs :: 'a word) = horner_sum of_bool 2 (take LENGTH('a::len) (rev bs))\ apply (simp add: of_bl_eq flip: take_bit_horner_sum_bit_eq) apply transfer apply simp done lemma [code]: \to_bl w = map (bit w) (rev [0.. for w :: \'a::len word\ by (fact to_bl_eq_rev) lemma word_reverse_eq_of_bl_rev_to_bl: \word_reverse w = of_bl (rev (to_bl w))\ by (rule bit_word_eqI) (auto simp add: bit_word_reverse_iff bit_of_bl_iff nth_to_bl) lemmas word_reverse_no_def [simp] = word_reverse_eq_of_bl_rev_to_bl [of "numeral w"] for w lemma to_bl_word_rev: "to_bl (word_reverse w) = rev (to_bl w)" by (rule nth_equalityI) (simp_all add: nth_rev_to_bl word_reverse_def word_rep_drop flip: of_bl_eq) lemma to_bl_n1 [simp]: "to_bl (-1::'a::len word) = replicate (LENGTH('a)) True" apply (rule word_bl.Abs_inverse') apply simp apply (rule word_eqI) apply (clarsimp simp add: word_size) apply (auto simp add: word_bl.Abs_inverse test_bit_bl word_size) done lemma rbl_word_or: "rev (to_bl (x OR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_or rev_map) lemma rbl_word_and: "rev (to_bl (x AND y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_and rev_map) lemma rbl_word_xor: "rev (to_bl (x XOR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_xor rev_map) lemma rbl_word_not: "rev (to_bl (NOT x)) = map Not (rev (to_bl x))" by (simp add: bl_word_not rev_map) lemma bshiftr1_numeral [simp]: \bshiftr1 b (numeral w :: 'a word) = of_bl (b # butlast (bin_to_bl LENGTH('a::len) (numeral w)))\ by (rule bit_word_eqI) (auto simp add: bit_simps rev_nth nth_append nth_butlast nth_bin_to_bl simp flip: bit_Suc) lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)" unfolding bshiftr1_eq by (rule word_bl.Abs_inverse) simp lemma shiftl1_of_bl: "shiftl1 (of_bl bl) = of_bl (bl @ [False])" by transfer (simp add: bl_to_bin_append) lemma shiftl1_bl: "shiftl1 w = of_bl (to_bl w @ [False])" for w :: "'a::len word" proof - have "shiftl1 w = shiftl1 (of_bl (to_bl w))" by simp also have "\ = of_bl (to_bl w @ [False])" by (rule shiftl1_of_bl) finally show ?thesis . qed lemma bl_shiftl1: "to_bl (shiftl1 w) = tl (to_bl w) @ [False]" for w :: "'a::len word" by (simp add: shiftl1_bl word_rep_drop drop_Suc drop_Cons') (fast intro!: Suc_leI) \ \Generalized version of \bl_shiftl1\. Maybe this one should replace it?\ lemma bl_shiftl1': "to_bl (shiftl1 w) = tl (to_bl w @ [False])" by (simp add: shiftl1_bl word_rep_drop drop_Suc del: drop_append) lemma shiftr1_bl: \shiftr1 w = of_bl (butlast (to_bl w))\ proof (rule bit_word_eqI) fix n assume \n < LENGTH('a)\ show \bit (shiftr1 w) n \ bit (of_bl (butlast (to_bl w)) :: 'a word) n\ proof (cases \n = LENGTH('a) - 1\) case True then show ?thesis by (simp add: bit_shiftr1_iff bit_of_bl_iff) next case False with \n < LENGTH('a)\ have \n < LENGTH('a) - 1\ by simp with \n < LENGTH('a)\ show ?thesis by (simp add: bit_shiftr1_iff bit_of_bl_iff rev_nth nth_butlast word_size to_bl_nth) qed qed lemma bl_shiftr1: "to_bl (shiftr1 w) = False # butlast (to_bl w)" for w :: "'a::len word" by (simp add: shiftr1_bl word_rep_drop len_gt_0 [THEN Suc_leI]) \ \Generalized version of \bl_shiftr1\. Maybe this one should replace it?\ lemma bl_shiftr1': "to_bl (shiftr1 w) = butlast (False # to_bl w)" apply (rule word_bl.Abs_inverse') apply (simp del: butlast.simps) apply (simp add: shiftr1_bl of_bl_def) done lemma bl_sshiftr1: "to_bl (sshiftr1 w) = hd (to_bl w) # butlast (to_bl w)" for w :: "'a::len word" proof (rule nth_equalityI) fix n assume \n < length (to_bl (sshiftr1 w))\ then have \n < LENGTH('a)\ by simp then show \to_bl (sshiftr1 w) ! n \ (hd (to_bl w) # butlast (to_bl w)) ! n\ apply (cases n) apply (simp_all add: to_bl_nth word_size hd_conv_nth bit_sshiftr1_iff nth_butlast Suc_diff_Suc nth_to_bl) done qed simp lemma drop_shiftr: "drop n (to_bl (drop_bit n w)) = take (size w - n) (to_bl w)" for w :: "'a::len word" apply (rule nth_equalityI) apply (simp_all add: word_size to_bl_nth bit_simps) done lemma drop_sshiftr: "drop n (to_bl (signed_drop_bit n w)) = take (size w - n) (to_bl w)" for w :: "'a::len word" apply (simp_all add: word_size) apply (rule nth_equalityI) apply (simp_all add: word_size nth_to_bl bit_signed_drop_bit_iff) done lemma take_shiftr: "n \ size w \ take n (to_bl (drop_bit n w)) = replicate n False" apply (rule nth_equalityI) apply (simp_all add: word_size to_bl_nth bit_simps) using bit_imp_le_length by fastforce lemma take_sshiftr': "n \ size w \ hd (to_bl (signed_drop_bit n w)) = hd (to_bl w) \ take n (to_bl (signed_drop_bit n w)) = replicate n (hd (to_bl w))" for w :: "'a::len word" apply (auto simp add: hd_bl_sign_sint bin_sign_def not_le word_size sint_signed_drop_bit_eq) apply (rule nth_equalityI) apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff) apply (rule nth_equalityI) apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff) done lemmas hd_sshiftr = take_sshiftr' [THEN conjunct1] lemmas take_sshiftr = take_sshiftr' [THEN conjunct2] lemma atd_lem: "take n xs = t \ drop n xs = d \ xs = t @ d" by (auto intro: append_take_drop_id [symmetric]) lemmas bl_shiftr = atd_lem [OF take_shiftr drop_shiftr] lemmas bl_sshiftr = atd_lem [OF take_sshiftr drop_sshiftr] lemma shiftl_of_bl: "push_bit n (of_bl bl) = of_bl (bl @ replicate n False)" apply (rule bit_word_eqI) apply (auto simp add: bit_simps nth_append) done lemma shiftl_bl: "push_bit n w = of_bl (to_bl w @ replicate n False)" for w :: "'a::len word" proof - have "push_bit n w = push_bit n (of_bl (to_bl w))" by simp also have "\ = of_bl (to_bl w @ replicate n False)" by (rule shiftl_of_bl) finally show ?thesis . qed lemma bl_shiftl: "to_bl (push_bit n w) = drop n (to_bl w) @ replicate (min (size w) n) False" by (simp add: shiftl_bl word_rep_drop word_size) lemma shiftr1_bl_of: "length bl \ LENGTH('a) \ shiftr1 (of_bl bl::'a::len word) = of_bl (butlast bl)" apply (cases bl rule: rev_cases) apply simp_all apply (rule bit_word_eqI) apply (auto simp add: bit_simps simp flip: bit_Suc) done lemma shiftr_bl_of: "length bl \ LENGTH('a) \ drop_bit n (of_bl bl::'a::len word) = of_bl (take (length bl - n) bl)" by (rule bit_word_eqI) (auto simp add: bit_simps rev_nth) lemma shiftr_bl: "drop_bit n x \ of_bl (take (LENGTH('a) - n) (to_bl x))" for x :: "'a::len word" using shiftr_bl_of [where 'a='a, of "to_bl x"] by simp lemma aligned_bl_add_size [OF refl]: "size x - n = m \ n \ size x \ drop m (to_bl x) = replicate n False \ take m (to_bl y) = replicate m False \ to_bl (x + y) = take m (to_bl x) @ drop m (to_bl y)" for x :: \'a::len word\ apply (subgoal_tac "x AND y = 0") prefer 2 apply (rule word_bl.Rep_eqD) apply (simp add: bl_word_and) apply (rule align_lem_and [THEN trans]) apply (simp_all add: word_size)[5] apply simp apply (subst word_plus_and_or [symmetric]) apply (simp add : bl_word_or) apply (rule align_lem_or) apply (simp_all add: word_size) done lemma mask_bl: "mask n = of_bl (replicate n True)" by (auto simp add: bit_simps intro!: word_eqI) lemma bl_and_mask': "to_bl (w AND mask n :: 'a::len word) = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl w)" apply (rule nth_equalityI) apply simp apply (clarsimp simp add: to_bl_nth word_size bit_simps) apply (auto simp add: word_size test_bit_bl nth_append rev_nth) done lemma slice1_eq_of_bl: \(slice1 n w :: 'b::len word) = of_bl (takefill False n (to_bl w))\ for w :: \'a::len word\ proof (rule bit_word_eqI) fix m assume \m < LENGTH('b)\ show \bit (slice1 n w :: 'b::len word) m \ bit (of_bl (takefill False n (to_bl w)) :: 'b word) m\ by (cases \m \ n\; cases \LENGTH('a) \ n\) (auto simp add: bit_slice1_iff bit_of_bl_iff not_less rev_nth not_le nth_takefill nth_to_bl algebra_simps) qed lemma slice1_no_bin [simp]: "slice1 n (numeral w :: 'b word) = of_bl (takefill False n (bin_to_bl (LENGTH('b::len)) (numeral w)))" by (simp add: slice1_eq_of_bl) (* TODO: neg_numeral *) lemma slice_no_bin [simp]: "slice n (numeral w :: 'b word) = of_bl (takefill False (LENGTH('b::len) - n) (bin_to_bl (LENGTH('b::len)) (numeral w)))" by (simp add: slice_def) (* TODO: neg_numeral *) lemma slice_take': "slice n w = of_bl (take (size w - n) (to_bl w))" by (simp add: slice_def word_size slice1_eq_of_bl takefill_alt) lemmas slice_take = slice_take' [unfolded word_size] \ \shiftr to a word of the same size is just slice, slice is just shiftr then ucast\ lemmas shiftr_slice = trans [OF shiftr_bl [THEN meta_eq_to_obj_eq] slice_take [symmetric]] lemma slice1_down_alt': "sl = slice1 n w \ fs = size sl \ fs + k = n \ to_bl sl = takefill False fs (drop k (to_bl w))" apply (simp add: slice1_eq_of_bl) apply transfer apply (simp add: bl_bin_bl_rep_drop) using drop_takefill apply force done lemma slice1_up_alt': "sl = slice1 n w \ fs = size sl \ fs = n + k \ to_bl sl = takefill False fs (replicate k False @ (to_bl w))" apply (simp add: slice1_eq_of_bl) apply transfer apply (simp add: bl_bin_bl_rep_drop flip: takefill_append) apply (metis diff_add_inverse) done lemmas sd1 = slice1_down_alt' [OF refl refl, unfolded word_size] lemmas su1 = slice1_up_alt' [OF refl refl, unfolded word_size] lemmas slice1_down_alt = le_add_diff_inverse [THEN sd1] lemmas slice1_up_alts = le_add_diff_inverse [symmetric, THEN su1] le_add_diff_inverse2 [symmetric, THEN su1] lemma slice1_tf_tf': "to_bl (slice1 n w :: 'a::len word) = rev (takefill False (LENGTH('a)) (rev (takefill False n (to_bl w))))" unfolding slice1_eq_of_bl by (rule word_rev_tf) lemmas slice1_tf_tf = slice1_tf_tf' [THEN word_bl.Rep_inverse', symmetric] lemma revcast_eq_of_bl: \(revcast w :: 'b::len word) = of_bl (takefill False (LENGTH('b)) (to_bl w))\ for w :: \'a::len word\ by (simp add: revcast_def slice1_eq_of_bl) lemmas revcast_no_def [simp] = revcast_eq_of_bl [where w="numeral w", unfolded word_size] for w lemma to_bl_revcast: "to_bl (revcast w :: 'a::len word) = takefill False (LENGTH('a)) (to_bl w)" apply (rule nth_equalityI) apply simp apply (cases \LENGTH('a) \ LENGTH('b)\) apply (auto simp add: nth_to_bl nth_takefill bit_revcast_iff) done lemma word_cat_bl: "word_cat a b = of_bl (to_bl a @ to_bl b)" apply (rule bit_word_eqI) apply (simp add: bit_word_cat_iff bit_of_bl_iff nth_append not_less nth_rev_to_bl) apply (meson bit_word.rep_eq less_diff_conv2 nth_rev_to_bl) done lemma of_bl_append: "(of_bl (xs @ ys) :: 'a::len word) = of_bl xs * 2^(length ys) + of_bl ys" apply transfer apply (simp add: bl_to_bin_app_cat bin_cat_num) done lemma of_bl_False [simp]: "of_bl (False#xs) = of_bl xs" by (rule word_eqI) (auto simp: test_bit_of_bl nth_append) lemma of_bl_True [simp]: "(of_bl (True # xs) :: 'a::len word) = 2^length xs + of_bl xs" by (subst of_bl_append [where xs="[True]", simplified]) (simp add: word_1_bl) lemma of_bl_Cons: "of_bl (x#xs) = of_bool x * 2^length xs + of_bl xs" by (cases x) simp_all lemma word_split_bl': "std = size c - size b \ (word_split c = (a, b)) \ (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c)))" apply (simp add: word_split_def) apply transfer apply (cases \LENGTH('b) \ LENGTH('a)\) apply (auto simp add: drop_bit_take_bit drop_bin2bl bin_to_bl_drop_bit [symmetric, of \LENGTH('a)\ \LENGTH('a) - LENGTH('b)\ \LENGTH('b)\] min_absorb2) done lemma word_split_bl: "std = size c - size b \ (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c))) \ word_split c = (a, b)" apply (rule iffI) defer apply (erule (1) word_split_bl') apply (case_tac "word_split c") apply (auto simp add: word_size) apply (frule word_split_bl' [rotated]) apply (auto simp add: word_size) done lemma word_split_bl_eq: "(word_split c :: ('c::len word \ 'd::len word)) = (of_bl (take (LENGTH('a::len) - LENGTH('d::len)) (to_bl c)), of_bl (drop (LENGTH('a) - LENGTH('d)) (to_bl c)))" for c :: "'a::len word" apply (rule word_split_bl [THEN iffD1]) apply (unfold word_size) apply (rule refl conjI)+ done lemma word_rcat_bl: \word_rcat wl = of_bl (concat (map to_bl wl))\ proof - define ws where \ws = rev wl\ moreover have \word_rcat (rev ws) = of_bl (concat (map to_bl (rev ws)))\ apply (simp add: word_rcat_def of_bl_eq rev_concat rev_map comp_def rev_to_bl_eq flip: horner_sum_of_bool_2_concat) apply transfer apply simp done ultimately show ?thesis by simp qed lemma size_rcat_lem': "size (concat (map to_bl wl)) = length wl * size (hd wl)" by (induct wl) (auto simp: word_size) lemmas size_rcat_lem = size_rcat_lem' [unfolded word_size] lemma nth_rcat_lem: "n < length (wl::'a word list) * LENGTH('a::len) \ rev (concat (map to_bl wl)) ! n = rev (to_bl (rev wl ! (n div LENGTH('a)))) ! (n mod LENGTH('a))" apply (induct wl) apply clarsimp apply (clarsimp simp add : nth_append size_rcat_lem) apply (simp flip: mult_Suc minus_div_mult_eq_mod add: less_Suc_eq_le not_less) apply (metis (no_types, lifting) diff_is_0_eq div_le_mono len_not_eq_0 less_Suc_eq less_mult_imp_div_less nonzero_mult_div_cancel_right not_le nth_Cons_0) done lemma foldl_eq_foldr: "foldl (+) x xs = foldr (+) (x # xs) 0" for x :: "'a::comm_monoid_add" by (induct xs arbitrary: x) (auto simp: add.assoc) lemmas word_cat_bl_no_bin [simp] = word_cat_bl [where a="numeral a" and b="numeral b", unfolded to_bl_numeral] for a b (* FIXME: negative numerals, 0 and 1 *) lemmas word_split_bl_no_bin [simp] = word_split_bl_eq [where c="numeral c", unfolded to_bl_numeral] for c lemmas word_rot_defs = word_roti_eq_word_rotr_word_rotl word_rotr_eq word_rotl_eq lemma to_bl_rotl: "to_bl (word_rotl n w) = rotate n (to_bl w)" by (simp add: word_rotl_eq to_bl_use_of_bl) lemmas blrs0 = rotate_eqs [THEN to_bl_rotl [THEN trans]] lemmas word_rotl_eqs = blrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotl [symmetric]] lemma to_bl_rotr: "to_bl (word_rotr n w) = rotater n (to_bl w)" by (simp add: word_rotr_eq to_bl_use_of_bl) lemmas brrs0 = rotater_eqs [THEN to_bl_rotr [THEN trans]] lemmas word_rotr_eqs = brrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotr [symmetric]] declare word_rotr_eqs (1) [simp] declare word_rotl_eqs (1) [simp] lemmas abl_cong = arg_cong [where f = "of_bl"] +end + locale word_rotate begin lemmas word_rot_defs' = to_bl_rotl to_bl_rotr lemmas blwl_syms [symmetric] = bl_word_not bl_word_and bl_word_or bl_word_xor lemmas lbl_lbl = trans [OF word_bl_Rep' word_bl_Rep' [symmetric]] lemmas ths_map2 [OF lbl_lbl] = rotate_map2 rotater_map2 lemmas ths_map [where xs = "to_bl v"] = rotate_map rotater_map for v lemmas th1s [simplified word_rot_defs' [symmetric]] = ths_map2 ths_map end lemmas bl_word_rotl_dt = trans [OF to_bl_rotl rotate_drop_take, simplified word_bl_Rep'] lemmas bl_word_rotr_dt = trans [OF to_bl_rotr rotater_drop_take, simplified word_bl_Rep'] lemma bl_word_roti_dt': "n = nat ((- i) mod int (size (w :: 'a::len word))) \ to_bl (word_roti i w) = drop n (to_bl w) @ take n (to_bl w)" apply (unfold word_roti_eq_word_rotr_word_rotl) apply (simp add: bl_word_rotl_dt bl_word_rotr_dt word_size) apply safe apply (simp add: zmod_zminus1_eq_if) apply safe apply (simp add: nat_mult_distrib) apply (simp add: nat_diff_distrib [OF pos_mod_sign pos_mod_conj [THEN conjunct2, THEN order_less_imp_le]] nat_mod_distrib) apply (simp add: nat_mod_distrib) done lemmas bl_word_roti_dt = bl_word_roti_dt' [unfolded word_size] lemmas word_rotl_dt = bl_word_rotl_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_rotr_dt = bl_word_rotr_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_roti_dt = bl_word_roti_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_rotr_dt_no_bin' [simp] = word_rotr_dt [where w="numeral w", unfolded to_bl_numeral] for w (* FIXME: negative numerals, 0 and 1 *) lemmas word_rotl_dt_no_bin' [simp] = word_rotl_dt [where w="numeral w", unfolded to_bl_numeral] for w (* FIXME: negative numerals, 0 and 1 *) lemma max_word_bl: "to_bl (- 1::'a::len word) = replicate LENGTH('a) True" by (fact to_bl_n1) lemma to_bl_mask: "to_bl (mask n :: 'a::len word) = replicate (LENGTH('a) - n) False @ replicate (min (LENGTH('a)) n) True" by (simp add: mask_bl word_rep_drop min_def) lemma map_replicate_True: "n = length xs \ map (\(x,y). x \ y) (zip xs (replicate n True)) = xs" by (induct xs arbitrary: n) auto lemma map_replicate_False: "n = length xs \ map (\(x,y). x \ y) (zip xs (replicate n False)) = replicate n False" by (induct xs arbitrary: n) auto +context + includes bit_operations_syntax +begin + lemma bl_and_mask: fixes w :: "'a::len word" and n :: nat defines "n' \ LENGTH('a) - n" shows "to_bl (w AND mask n) = replicate n' False @ drop n' (to_bl w)" proof - note [simp] = map_replicate_True map_replicate_False have "to_bl (w AND mask n) = map2 (\) (to_bl w) (to_bl (mask n::'a::len word))" by (simp add: bl_word_and) also have "to_bl w = take n' (to_bl w) @ drop n' (to_bl w)" by simp also have "map2 (\) \ (to_bl (mask n::'a::len word)) = replicate n' False @ drop n' (to_bl w)" unfolding to_bl_mask n'_def by (subst zip_append) auto finally show ?thesis . qed lemma drop_rev_takefill: "length xs \ n \ drop (n - length xs) (rev (takefill False n (rev xs))) = xs" by (simp add: takefill_alt rev_take) declare bin_to_bl_def [simp] lemmas of_bl_reasoning = to_bl_use_of_bl of_bl_append lemma uint_of_bl_is_bl_to_bin_drop: "length (dropWhile Not l) \ LENGTH('a) \ uint (of_bl l :: 'a::len word) = bl_to_bin l" apply transfer apply (simp add: take_bit_eq_mod) apply (rule Divides.mod_less) apply (rule bl_to_bin_ge0) using bl_to_bin_lt2p_drop apply (rule order.strict_trans2) apply simp done corollary uint_of_bl_is_bl_to_bin: "length l\LENGTH('a) \ uint ((of_bl::bool list\ ('a :: len) word) l) = bl_to_bin l" apply(rule uint_of_bl_is_bl_to_bin_drop) using le_trans length_dropWhile_le by blast lemma bin_to_bl_or: "bin_to_bl n (a OR b) = map2 (\) (bin_to_bl n a) (bin_to_bl n b)" using bl_or_aux_bin[where n=n and v=a and w=b and bs="[]" and cs="[]"] by simp lemma word_and_1_bl: fixes x::"'a::len word" shows "(x AND 1) = of_bl [bit x 0]" by (simp add: mod_2_eq_odd and_one_eq) lemma word_1_and_bl: fixes x::"'a::len word" shows "(1 AND x) = of_bl [bit x 0]" by (simp add: mod_2_eq_odd one_and_eq) lemma of_bl_drop: "of_bl (drop n xs) = (of_bl xs AND mask (length xs - n))" apply (rule bit_word_eqI) apply (auto simp: rev_nth bit_simps cong: rev_conj_cong) done lemma to_bl_1: "to_bl (1::'a::len word) = replicate (LENGTH('a) - 1) False @ [True]" by (rule nth_equalityI) (auto simp add: to_bl_unfold nth_append rev_nth bit_1_iff not_less not_le) lemma eq_zero_set_bl: "(w = 0) = (True \ set (to_bl w))" apply (auto simp add: to_bl_unfold) apply (rule bit_word_eqI) apply auto done lemma of_drop_to_bl: "of_bl (drop n (to_bl x)) = (x AND mask (size x - n))" by (simp add: of_bl_drop word_size_bl) lemma unat_of_bl_length: "unat (of_bl xs :: 'a::len word) < 2 ^ (length xs)" proof (cases "length xs < LENGTH('a)") case True then have "(of_bl xs::'a::len word) < 2 ^ length xs" by (simp add: of_bl_length_less) with True show ?thesis by (simp add: word_less_nat_alt unat_of_nat) next case False have "unat (of_bl xs::'a::len word) < 2 ^ LENGTH('a)" by (simp split: unat_split) also from False have "LENGTH('a) \ length xs" by simp then have "2 ^ LENGTH('a) \ (2::nat) ^ length xs" by (rule power_increasing) simp finally show ?thesis . qed lemma word_msb_alt: "msb w \ hd (to_bl w)" for w :: "'a::len word" apply (simp add: msb_word_eq) apply (subst hd_conv_nth) apply simp apply (subst nth_to_bl) apply simp apply simp done lemma word_lsb_last: \lsb w \ last (to_bl w)\ for w :: \'a::len word\ using nth_to_bl [of \LENGTH('a) - Suc 0\ w] by (simp add: lsb_odd last_conv_nth) lemma is_aligned_to_bl: "is_aligned (w :: 'a :: len word) n = (True \ set (drop (size w - n) (to_bl w)))" by (simp add: is_aligned_mask eq_zero_set_bl bl_and_mask word_size) lemma is_aligned_replicate: fixes w::"'a::len word" assumes aligned: "is_aligned w n" and nv: "n \ LENGTH('a)" shows "to_bl w = (take (LENGTH('a) - n) (to_bl w)) @ replicate n False" apply (rule nth_equalityI) using assms apply (simp_all add: nth_append not_less word_size to_bl_nth is_aligned_imp_not_bit) done lemma is_aligned_drop: fixes w::"'a::len word" assumes "is_aligned w n" "n \ LENGTH('a)" shows "drop (LENGTH('a) - n) (to_bl w) = replicate n False" proof - have "to_bl w = take (LENGTH('a) - n) (to_bl w) @ replicate n False" by (rule is_aligned_replicate) fact+ then have "drop (LENGTH('a) - n) (to_bl w) = drop (LENGTH('a) - n) \" by simp also have "\ = replicate n False" by simp finally show ?thesis . qed lemma less_is_drop_replicate: fixes x::"'a::len word" assumes lt: "x < 2 ^ n" shows "to_bl x = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl x)" by (metis assms bl_and_mask' less_mask_eq) lemma is_aligned_add_conv: fixes off::"'a::len word" assumes aligned: "is_aligned w n" and offv: "off < 2 ^ n" shows "to_bl (w + off) = (take (LENGTH('a) - n) (to_bl w)) @ (drop (LENGTH('a) - n) (to_bl off))" proof cases assume nv: "n \ LENGTH('a)" show ?thesis proof (subst aligned_bl_add_size, simp_all only: word_size) show "drop (LENGTH('a) - n) (to_bl w) = replicate n False" by (subst is_aligned_replicate [OF aligned nv]) (simp add: word_size) from offv show "take (LENGTH('a) - n) (to_bl off) = replicate (LENGTH('a) - n) False" by (subst less_is_drop_replicate, assumption) simp qed fact next assume "\ n \ LENGTH('a)" with offv show ?thesis by (simp add: power_overflow) qed lemma is_aligned_replicateI: "to_bl p = addr @ replicate n False \ is_aligned (p::'a::len word) n" apply (simp add: is_aligned_to_bl word_size) apply (subgoal_tac "length addr = LENGTH('a) - n") apply (simp add: replicate_not_True) apply (drule arg_cong [where f=length]) apply simp done lemma to_bl_2p: "n < LENGTH('a) \ to_bl ((2::'a::len word) ^ n) = replicate (LENGTH('a) - Suc n) False @ True # replicate n False" apply (rule nth_equalityI) apply (auto simp add: nth_append to_bl_nth word_size bit_simps not_less nth_Cons le_diff_conv) subgoal for i apply (cases \Suc (i + n) - LENGTH('a)\) apply simp_all done done lemma xor_2p_to_bl: fixes x::"'a::len word" shows "to_bl (x XOR 2^n) = (if n < LENGTH('a) then take (LENGTH('a)-Suc n) (to_bl x) @ (\rev (to_bl x)!n) # drop (LENGTH('a)-n) (to_bl x) else to_bl x)" apply (auto simp add: to_bl_eq_rev take_map drop_map take_rev drop_rev bit_simps) apply (rule nth_equalityI) apply (auto simp add: bit_simps rev_nth nth_append Suc_diff_Suc) done lemma is_aligned_replicateD: "\ is_aligned (w::'a::len word) n; n \ LENGTH('a) \ \ \xs. to_bl w = xs @ replicate n False \ length xs = size w - n" apply (subst is_aligned_replicate, assumption+) apply (rule exI, rule conjI, rule refl) apply (simp add: word_size) done text \right-padding a word to a certain length\ definition "bl_pad_to bl sz \ bl @ (replicate (sz - length bl) False)" lemma bl_pad_to_length: assumes lbl: "length bl \ sz" shows "length (bl_pad_to bl sz) = sz" using lbl by (simp add: bl_pad_to_def) lemma bl_pad_to_prefix: "prefix bl (bl_pad_to bl sz)" by (simp add: bl_pad_to_def) lemma of_bl_length: "length xs < LENGTH('a) \ of_bl xs < (2 :: 'a::len word) ^ length xs" by (simp add: of_bl_length_less) lemma of_bl_mult_and_not_mask_eq: "\is_aligned (a :: 'a::len word) n; length b + m \ n\ \ a + of_bl b * (2^m) AND NOT(mask n) = a" apply (simp flip: push_bit_eq_mult subtract_mask(1) take_bit_eq_mask) apply (subst disjunctive_add) apply (auto simp add: bit_simps not_le not_less) apply (meson is_aligned_imp_not_bit is_aligned_weaken less_diff_conv2) apply (erule is_alignedE') apply (simp add: take_bit_push_bit) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma bin_to_bl_of_bl_eq: "\is_aligned (a::'a::len word) n; length b + c \ n; length b + c < LENGTH('a)\ \ bin_to_bl (length b) (uint (drop_bit c (a + of_bl b * 2^c))) = b" apply (simp flip: push_bit_eq_mult take_bit_eq_mask) apply (subst disjunctive_add) apply (auto simp add: bit_simps not_le not_less unsigned_or_eq unsigned_drop_bit_eq unsigned_push_bit_eq bin_to_bl_or simp flip: bin_to_bl_def) apply (meson is_aligned_imp_not_bit is_aligned_weaken less_diff_conv2) apply (erule is_alignedE') apply (rule nth_equalityI) apply (auto simp add: nth_bin_to_bl bit_simps rev_nth simp flip: bin_to_bl_def) done (* FIXME: move to Word distribution *) lemma bin_nth_minus_Bit0[simp]: "0 < n \ bit (numeral (num.Bit0 w) :: int) n = bit (numeral w :: int) (n - 1)" by (cases n; simp) lemma bin_nth_minus_Bit1[simp]: "0 < n \ bit (numeral (num.Bit1 w) :: int) n = bit (numeral w :: int) (n - 1)" by (cases n; simp) (* casting a long word to a shorter word and casting back to the long word is equal to the original long word -- if the word is small enough. 'l is the longer word. 's is the shorter word. *) lemma bl_cast_long_short_long_ingoreLeadingZero_generic: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (of_bl :: _ \ 'l::len word) (to_bl ((of_bl::_ \ 's::len word) (to_bl w))) = w" by (rule word_uint_eqI) (simp add: uint_of_bl_is_bl_to_bin uint_of_bl_is_bl_to_bin_drop) (* Casting between longer and shorter word. 'l is the longer word. 's is the shorter word. For example: 'l::len word is 128 word (full ipv6 address) 's::len word is 16 word (address piece of ipv6 address in colon-text-representation) *) corollary ucast_short_ucast_long_ingoreLeadingZero: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (ucast:: 's::len word \ 'l::len word) ((ucast:: 'l::len word \ 's::len word) w) = w" apply (subst ucast_bl)+ apply (rule bl_cast_long_short_long_ingoreLeadingZero_generic; simp) done lemma length_drop_mask: fixes w::"'a::len word" shows "length (dropWhile Not (to_bl (w AND mask n))) \ n" proof - have "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)" for ls n by(subst takeWhile_append2) simp+ then show ?thesis unfolding bl_and_mask by (simp add: dropWhile_eq_drop) qed lemma map_bits_rev_to_bl: "map (bit x) [0.. of_bl xs * 2^c < (2::'a::len word) ^ (length xs + c)" by (simp add: of_bl_length word_less_power_trans2) lemma of_bl_max: "(of_bl xs :: 'a::len word) \ mask (length xs)" proof - define ys where \ys = rev xs\ have \take_bit (length ys) (horner_sum of_bool 2 ys :: 'a word) = horner_sum of_bool 2 ys\ by transfer (simp add: take_bit_horner_sum_bit_eq min_def) then have \(of_bl (rev ys) :: 'a word) \ mask (length ys)\ by (simp only: of_bl_rev_eq less_eq_mask_iff_take_bit_eq_self) with ys_def show ?thesis by simp qed text\Some auxiliaries for sign-shifting by the entire word length or more\ lemma sshiftr_clamp_pos: assumes "LENGTH('a) \ n" "0 \ sint x" shows "signed_drop_bit n (x::'a::len word) = 0" apply (rule word_sint.Rep_eqD) apply (simp add: sint_signed_drop_bit_eq) using assms by (metis Word.sint_0 bit_last_iff not_less signed_drop_bit_beyond sint_signed_drop_bit_eq) lemma sshiftr_clamp_neg: assumes "LENGTH('a) \ n" "sint x < 0" shows "signed_drop_bit n (x::'a::len word) = -1" apply (rule word_sint.Rep_eqD) apply (simp add: sint_signed_drop_bit_eq) using assms by (metis bit_last_iff signed_drop_bit_beyond sint_n1 sint_signed_drop_bit_eq) lemma sshiftr_clamp: assumes "LENGTH('a) \ n" shows "signed_drop_bit n (x::'a::len word) = signed_drop_bit LENGTH('a) x" apply (cases "0 \ sint x") subgoal apply (subst sshiftr_clamp_pos[OF assms]) defer apply (subst sshiftr_clamp_pos) by auto apply (subst sshiftr_clamp_neg[OF assms]) defer apply (subst sshiftr_clamp_neg) by auto text\ Like @{thm shiftr1_bl_of}, but the precondition is stronger because we need to pick the msb out of the list. \ lemma sshiftr1_bl_of: "length bl = LENGTH('a) \ sshiftr1 (of_bl bl::'a::len word) = of_bl (hd bl # butlast bl)" apply (rule word_bl.Rep_eqD) apply (subst bl_sshiftr1[of "of_bl bl :: 'a word"]) by (simp add: word_bl.Abs_inverse) text\ Like @{thm sshiftr1_bl_of}, with a weaker precondition. We still get a direct equation for @{term \sshiftr1 (of_bl bl)\}, it's just uglier. \ lemma sshiftr1_bl_of': "LENGTH('a) \ length bl \ sshiftr1 (of_bl bl::'a::len word) = of_bl (hd (drop (length bl - LENGTH('a)) bl) # butlast (drop (length bl - LENGTH('a)) bl))" apply (subst of_bl_drop'[symmetric, of "length bl - LENGTH('a)"]) using sshiftr1_bl_of[of "drop (length bl - LENGTH('a)) bl"] by auto text\ Like @{thm shiftr_bl_of}. \ lemma sshiftr_bl_of: assumes "length bl = LENGTH('a)" shows "signed_drop_bit n (of_bl bl::'a::len word) = of_bl (replicate n (hd bl) @ take (length bl - n) bl)" proof - from assms obtain b bs where \bl = b # bs\ by (cases bl) simp_all then have *: \bl ! 0 \ b\ \hd bl \ b\ by simp_all show ?thesis apply (rule bit_word_eqI) using assms * by (auto simp add: bit_simps nth_append rev_nth not_less) qed text\Like @{thm shiftr_bl}\ lemma sshiftr_bl: "signed_drop_bit n x \ of_bl (replicate n (msb x) @ take (LENGTH('a) - n) (to_bl x))" for x :: "'a::len word" unfolding word_msb_alt by (smt (z3) length_to_bl_eq sshiftr_bl_of word_bl.Rep_inverse) end + +end diff --git a/thys/Word_Lib/Word_16.thy b/thys/Word_Lib/Word_16.thy --- a/thys/Word_Lib/Word_16.thy +++ b/thys/Word_Lib/Word_16.thy @@ -1,22 +1,28 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Words of Length 16" theory Word_16 imports More_Word Signed_Words begin lemma len16: "len_of (x :: 16 itself) = 16" by simp +context + includes bit_operations_syntax +begin + lemma word16_and_max_simp: \x AND 0xFFFF = x\ for x :: \16 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_32.thy b/thys/Word_Lib/Word_32.thy --- a/thys/Word_Lib/Word_32.thy +++ b/thys/Word_Lib/Word_32.thy @@ -1,342 +1,348 @@ (* * 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) 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 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,306 +1,312 @@ (* * 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) 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 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,108 +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) 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,1934 +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 (clarsimp simp add: mask_eq_decr_exp Parity.bit_eq_iff bit_and_iff bit_not_iff bit_push_bit_iff) (* 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) 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" by (metis (no_types, lifting) and_not_mask bit.conj_ac(1) bit.conj_ac(2) mask_eq_0_eq_x shiftl_mask_is_0 shiftl_over_and_dist) 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 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 diff --git a/thys/Word_Lib/Word_Lib_Sumo.thy b/thys/Word_Lib/Word_Lib_Sumo.thy --- a/thys/Word_Lib/Word_Lib_Sumo.thy +++ b/thys/Word_Lib/Word_Lib_Sumo.thy @@ -1,134 +1,135 @@ (* * Copyright Florian Haftmann * * SPDX-License-Identifier: BSD-2-Clause *) section \Ancient comprehensive Word Library\ theory Word_Lib_Sumo imports "HOL-Library.Word" Aligned Bit_Comprehension Bit_Shifts_Infix_Syntax Bits_Int Bitwise_Signed Bitwise Enumeration_Word Generic_set_bit Hex_Words Least_significant_bit More_Arithmetic More_Divides More_Sublist Even_More_List More_Misc Strict_part_mono Legacy_Aliases Most_significant_bit Next_and_Prev Norm_Words Reversed_Bit_Lists Rsplit Signed_Words Syntax_Bundles Typedef_Morphisms Type_Syntax Word_EqI Word_Lemmas Word_8 Word_16 Word_32 Word_Syntax Signed_Division_Word More_Word_Operations Many_More begin +unbundle bit_operations_syntax unbundle bit_projection_infix_syntax declare word_induct2[induct type] declare word_nat_cases[cases type] declare signed_take_bit_Suc [simp] (* these generate take_bit terms, which we often don't want for concrete lengths *) lemmas of_int_and_nat = unsigned_of_nat unsigned_of_int signed_of_int signed_of_nat bundle no_take_bit begin declare of_int_and_nat[simp del] end lemmas bshiftr1_def = bshiftr1_eq lemmas is_down_def = is_down_eq lemmas is_up_def = is_up_eq lemmas mask_def = mask_eq lemmas scast_def = scast_eq lemmas shiftl1_def = shiftl1_eq lemmas shiftr1_def = shiftr1_eq lemmas sshiftr1_def = sshiftr1_eq lemmas sshiftr_def = sshiftr_eq_funpow_sshiftr1 lemmas to_bl_def = to_bl_eq lemmas ucast_def = ucast_eq lemmas unat_def = unat_eq_nat_uint lemmas word_cat_def = word_cat_eq lemmas word_reverse_def = word_reverse_eq_of_bl_rev_to_bl lemmas word_roti_def = word_roti_eq_word_rotr_word_rotl lemmas word_rotl_def = word_rotl_eq lemmas word_rotr_def = word_rotr_eq lemmas word_sle_def = word_sle_eq lemmas word_sless_def = word_sless_eq lemmas uint_0 = uint_nonnegative lemmas uint_lt = uint_bounded lemmas uint_mod_same = uint_idem lemmas of_nth_def = word_set_bits_def lemmas of_nat_word_eq_iff = word_of_nat_eq_iff lemmas of_nat_word_eq_0_iff = word_of_nat_eq_0_iff lemmas of_int_word_eq_iff = word_of_int_eq_iff lemmas of_int_word_eq_0_iff = word_of_int_eq_0_iff lemmas word_next_def = word_next_unfold lemmas word_prev_def = word_prev_unfold lemmas is_aligned_def = is_aligned_iff_dvd_nat lemmas word_and_max_simps = word8_and_max_simp word16_and_max_simp word32_and_max_simp lemma distinct_lemma: "f x \ f y \ x \ y" by auto lemmas and_bang = word_and_nth lemmas sdiv_int_def = signed_divide_int_def lemmas smod_int_def = signed_modulo_int_def (* shortcut for some specific lengths *) lemma word_fixed_sint_1[simp]: "sint (1::8 word) = 1" "sint (1::16 word) = 1" "sint (1::32 word) = 1" "sint (1::64 word) = 1" by (auto simp: sint_word_ariths) declare of_nat_diff [simp] (* Haskellish names/syntax *) notation (input) bit ("testBit") lemmas cast_simps = cast_simps ucast_down_bl (* shadows the slightly weaker Word.nth_ucast *) lemma nth_ucast: "(ucast (w::'a::len word)::'b::len word) !! n = (w !! n \ n < min LENGTH('a) LENGTH('b))" by (auto simp add: bit_simps not_le dest: bit_imp_le_length) end diff --git a/thys/Word_Lib/Word_Syntax.thy b/thys/Word_Lib/Word_Syntax.thy --- a/thys/Word_Lib/Word_Syntax.thy +++ b/thys/Word_Lib/Word_Syntax.thy @@ -1,36 +1,42 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Additional Syntax for Word Bit Operations" theory Word_Syntax imports "HOL-Library.Word" begin text \Additional bit and type syntax that forces word types.\ +context + includes bit_operations_syntax +begin + abbreviation wordNOT :: "'a::len word \ 'a word" ("~~ _" [70] 71) where "~~ x == NOT x" abbreviation wordAND :: "'a::len word \ 'a word \ 'a word" (infixr "&&" 64) where "a && b == a AND b" abbreviation wordOR :: "'a::len word \ 'a word \ 'a word" (infixr "||" 59) where "a || b == a OR b" abbreviation wordXOR :: "'a::len word \ 'a word \ 'a word" (infixr "xor" 59) where "a xor b == a XOR b" end + +end