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,1613 +1,1635 @@ (* 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 (shiftr n k) = (int_of_uint32 n) div (2 ^ k)" by (transfer, rule shiftr_div_2n) 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: int_word_uint) + by transfer (simp add: take_bit_int_eq_self) 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) (shiftr n 1) 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 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]: "shiftr x 1 = (uint32_shiftr x 1)" unfolding shiftr_uint32_code using integer_of_nat_1 by auto (* ******************************************************************************** *) 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, auto) 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]) 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 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'" unfolding r' using y unfolding urel32_def using small - unfolding ppp by transfer (auto simp add: uint_nat unat_mod and_one_eq) + have "urel32 (y AND 1) r'" + unfolding r' + using y + unfolding urel32_def + using small + apply (simp add: ppp and_one_eq) + apply transfer + apply transfer + apply (auto simp add: zmod_int take_bit_int_eq_self) + apply (rule le_less_trans) + apply (rule zmod_le_nonneg_dividend) + apply simp_all + done from urel32_eq[OF this urel32_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel32 (shiftr y 1) (int d')" unfolding d' using y unfolding urel32_def using small unfolding ppp by (transfer, auto simp: shiftr_div_2n) 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 (transfer, auto simp: uint_word_ariths) 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 (shiftr n k) = (int_of_uint64 n) div (2 ^ k)" by (transfer, rule shiftr_div_2n) 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: int_word_uint) + by transfer (simp add: take_bit_int_eq_self) 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) (shiftr n 1) 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 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]: "shiftr x 1 = (uint64_shiftr x 1)" unfolding shiftr_uint64_code using integer_of_nat_1 by auto 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, auto) 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]) 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 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'" unfolding r' using y unfolding urel64_def using small - unfolding ppp by transfer (auto simp add: uint_nat unat_mod and_one_eq) + have "urel64 (y AND 1) r'" + unfolding r' + using y + unfolding urel64_def + using small + apply (simp add: ppp and_one_eq) + apply transfer apply transfer + apply (auto simp add: int_eq_iff nat_take_bit_eq nat_mod_distrib zmod_int) + apply (auto simp add: zmod_int mod_2_eq_odd) + apply (metis (full_types) even_take_bit_eq le_less_trans odd_iff_mod_2_eq_one take_bit_nonnegative zero_neq_numeral zmod_le_nonneg_dividend) + apply (auto simp add: less_le) + apply (simp add: le_less) + done from urel64_eq[OF this urel64_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel64 (shiftr y 1) (int d')" unfolding d' using y unfolding urel64_def using small unfolding ppp by (transfer, auto simp: shiftr_div_2n) 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 (transfer, auto simp: uint_word_ariths) 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 (shiftr n k) = (int_of_integer n) div (2 ^ k)" by (simp add: shiftr_int_def shiftr_integer.rep_eq) 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) (shiftr n 1) 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 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]: "shiftr x 1 = (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]) 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 by (smt aux bitAND_int_code int_and_1 mod2_gr_0 of_nat_0_le_iff of_nat_0_less_iff of_nat_1 one_integer.rep_eq p2 ppp) from urel_integer_eq[OF this urel_integer_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel_integer (shiftr y 1) (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 no_notation shiftr (infixl ">>" 55) (* to avoid conflict with bind *) end diff --git a/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_State.thy b/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_State.thy --- a/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_State.thy +++ b/thys/CAVA_LTL_Modelchecker/SM/RefPoint/SM_State.thy @@ -1,72 +1,70 @@ section \State for SM\ theory SM_State imports SM_Syntax "HOL-Word.Word" "HOL-Library.Multiset" begin section \Values\ text \The primitive values are fixed-size signed integers\ type_synonym word_size = 32 \ \Word size\ type_synonym signed = "word_size Word.word" \ \Signed integer\ text \Currently, we only have signed integer values. This may change if we extend the language, and allow, i.e., channel pointers, pids or process references\ type_synonym val = signed \ \Value type\ section \Configurations\ type_synonym valuation = "ident \ val" record local_state = variables :: valuation record global_state = variables :: valuation text \The effect of actions is on focused states\ type_synonym focused_state = "local_state \ global_state" section \Utilities\ abbreviation "word_len \ LENGTH(word_size)" abbreviation "signeds \ sints (LENGTH(word_size))" definition min_signed :: int where "min_signed \ -(2^(word_len - 1))" definition max_signed :: int where "max_signed \ 2^(word_len - 1) - 1" definition signed_of_int :: "int \ signed" where "signed_of_int i \ word_of_int i" definition int_of_signed :: "signed \ int" where "int_of_signed i == sint i" lemma si_inv[simp]: "signed_of_int (int_of_signed i) = i" unfolding signed_of_int_def int_of_signed_def by simp lemma int_of_signed_in_range[simp]: "int_of_signed i \ min_signed" "int_of_signed i \ max_signed" unfolding int_of_signed_def min_signed_def max_signed_def apply - apply (rule sint_ge) using sint_lt[of i] by simp lemma is_inv[simp]: "\ i\min_signed; i\max_signed \ \ (int_of_signed (signed_of_int i)) = i" - unfolding int_of_signed_def signed_of_int_def min_signed_def max_signed_def - by (simp add: sints_num word_sint.Abs_inverse) + by (simp add: signed_take_bit_int_eq_self min_signed_def max_signed_def int_of_signed_def signed_of_int_def) - primrec val_of_bool :: "bool \ val" where "val_of_bool False = 0" | "val_of_bool True = 1" definition bool_of_val :: "val \ bool" where "bool_of_val v \ v\0" lemma bool_of_val_invs[simp]: "bool_of_val (val_of_bool b) = b" "val_of_bool (bool_of_val v) = (if v=0 then 0 else 1)" unfolding bool_of_val_def by (cases b) auto end diff --git a/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy b/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy --- a/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy +++ b/thys/IEEE_Floating_Point/Conversion_IEEE_Float.thy @@ -1,348 +1,349 @@ (* Author: Fabian Hellauer Fabian Immler *) theory Conversion_IEEE_Float imports "HOL-Library.Float" IEEE_Properties "HOL-Library.Code_Target_Numeral" begin definition "of_finite (x::('e, 'f)float) = (if is_normal x then (Float (normal_mantissa x) (normal_exponent x)) else if is_denormal x then (Float (denormal_mantissa x) (denormal_exponent TYPE(('e, 'f)float))) else 0)" lemma float_val_of_finite: "is_finite x \ of_finite x = valof x" by (induction x) (auto simp: normal_imp_not_denormal of_finite_def) definition is_normal_Float::"('e, 'f)float itself \ Float.float \ bool" where "is_normal_Float x f \ mantissa f \ 0 \ bitlen \mantissa f\ \ fracwidth x + 1 \ - int (bias x) - bitlen \mantissa f\ + 1 < Float.exponent f \ Float.exponent f < 2^(LENGTH('e)) - bitlen \mantissa f\ - bias x" definition is_denormal_Float::"('e, 'f)float itself \ Float.float \ bool" where "is_denormal_Float x f \ mantissa f \ 0 \ bitlen \mantissa f\ \ 1 - Float.exponent f - int (bias x) \ 1 - 2^(LENGTH('e) - 1) - int LENGTH('f) < Float.exponent f" lemmas is_denormal_FloatD = is_denormal_Float_def[THEN iffD1, THEN conjunct1] is_denormal_Float_def[THEN iffD1, THEN conjunct2] definition is_finite_Float::"('e, 'f)float itself \ Float.float \ bool" where "is_finite_Float x f \ is_normal_Float x f \ is_denormal_Float x f \ f = 0" lemma is_finite_Float_eq: "is_finite_Float TYPE(('e, 'f)float) f \ (let e = Float.exponent f; bm = bitlen (abs (mantissa f)) in bm \ Suc LENGTH('f) \ bm \ 2 ^ (LENGTH('e) - 1) - e \ 1 - 2 ^ (LENGTH('e) - 1) - int LENGTH('f) < e)" proof - have *: "(2::int) ^ (LENGTH('e) - Suc 0) - 1 < 2 ^ LENGTH('e)" by (metis Suc_1 diff_le_self lessI linorder_not_less one_less_numeral_iff power_strict_increasing_iff zle_diff1_eq) have **: "1 - 2 ^ (LENGTH('e) - Suc 0) < int LENGTH('f)" by (smt len_gt_0 of_nat_0_less_iff zero_less_power) have ***: "2 ^ (LENGTH('e) - 1) + 1 = 2 ^ LENGTH('e) - int (bias TYPE(('e, 'f) IEEE.float))" by (simp add: bias_def power_Suc[symmetric]) have rewr: "x \ 2 ^ n - e \ x + e < 2 ^ n + 1" for x::int and n e by auto show ?thesis unfolding *** rewr using * ** unfolding is_finite_Float_def is_normal_Float_def is_denormal_Float_def by (auto simp: Let_def bias_def mantissa_eq_zero_iff intro: le_less_trans[OF add_right_mono]) qed lift_definition normal_of_Float :: "Float.float \ ('e, 'f)float" is "\x. let m = mantissa x; e = Float.exponent x in (if m > 0 then 0 else 1, word_of_int (e + int (bias TYPE(('e, 'f)float)) + bitlen \m\ - 1), word_of_int (\m\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \m\)) - 2 ^ (LENGTH('f))))" . lemma sign_normal_of_Float:"sign (normal_of_Float x) = (if x > 0 then 0 else 1)" by transfer (auto simp: Let_def mantissa_pos_iff) lemma uints_bitlen_eq: "uints n = {i. 0 \ i \ bitlen i \ n}" by (auto simp: uints_num bitlen_le_iff_power) lemma uint_word_of_int_bitlen_eq: "uint (word_of_int x::'a::len word) = x" if "bitlen x \ LENGTH('a)" "x \ 0" by (subst word_uint.Abs_inverse) (simp_all add: uints_bitlen_eq that) lemma fraction_normal_of_Float:"fraction (normal_of_Float x::('e, 'f)float) = (nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f))" if "is_normal_Float TYPE(('e, 'f)float) x" proof - from that have bmp: "bitlen \mantissa x\ > 0" by (metis abs_of_nonneg bitlen_bounds bitlen_def is_normal_Float_def nat_code(2) of_nat_0_le_iff power.simps(1) zabs_less_one_iff zero_less_abs_iff) have mless: "\mantissa x\ < 2 ^ nat (bitlen \mantissa x\)" using bitlen_bounds by force have lem: "2 ^ nat (bitlen \mantissa x\ - 1) \ \mantissa x\" using bitlen_bounds is_normal_Float_def that zero_less_abs_iff by blast from that have nble: "nat (bitlen \mantissa x\) \ Suc LENGTH('f)" using bitlen_bounds by (auto simp: is_normal_Float_def) have nn: "0 \ \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f)" apply (rule add_le_imp_le_diff) apply (rule order_trans[rotated]) apply (rule mult_right_mono) apply (rule lem, force) unfolding power_add[symmetric] using nble bmp by (auto) have "\mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) < 2 * 2 ^ LENGTH('f)" apply (rule less_le_trans) apply (rule mult_strict_right_mono) apply (rule mless) apply force unfolding power_add[symmetric] power_Suc[symmetric] apply (rule power_increasing) using nble by auto then have "bitlen (\mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f)) \ int LENGTH('f)" unfolding bitlen_le_iff_power by simp then show ?thesis apply (transfer fixing: x) unfolding Let_def split_beta' fst_conv snd_conv uint_nat [symmetric] nat_uint_eq [symmetric] using nn apply (subst uint_word_of_int_bitlen_eq) apply (auto simp: nat_mult_distrib nat_diff_distrib nat_power_eq) done qed lemma exponent_normal_of_Float:"exponent (normal_of_Float x::('e, 'f)float) = nat (Float.exponent x + (bias TYPE(('e, 'f)float)) + bitlen \mantissa x\ - 1)" if "is_normal_Float TYPE(('e, 'f)float) x" using that apply (transfer fixing: x) apply (simp flip: uint_nat nat_uint_eq add: Let_def) - apply (subst uint_word_of_int_bitlen_eq) - apply (auto simp: is_normal_Float_def bitlen_le_iff_power uint_word_of_int_bitlen_eq Let_def) + apply (auto simp: is_normal_Float_def bitlen_le_iff_power uint_word_of_int_bitlen_eq Let_def) + apply transfer + apply (simp add: nat_take_bit_eq take_bit_int_eq_self) done lift_definition denormal_of_Float :: "Float.float \ ('e, 'f)float" is "\x. let m = mantissa x; e = Float.exponent x in (if m \ 0 then 0 else 1, 0, word_of_int (\m\ * 2 ^ nat (e + bias TYPE(('e, 'f)float) + fracwidth TYPE(('e, 'f)float) - 1)))" . lemma sign_denormal_of_Float:"sign (denormal_of_Float x) = (if x \ 0 then 0 else 1)" by transfer (auto simp: Let_def mantissa_nonneg_iff) lemma exponent_denormal_of_Float:"exponent (denormal_of_Float x::('e, 'f)float) = 0" by (transfer fixing: x) (auto simp: Let_def) lemma fraction_denormal_of_Float:"fraction (denormal_of_Float x::('e, 'f)float) = (nat \mantissa x\ * 2 ^ nat (Float.exponent x + bias TYPE(('e, 'f)float) + LENGTH('f) - 1))" if "is_denormal_Float TYPE(('e, 'f)float) x" proof - have mless: "\mantissa x\ < 2 ^ nat (bitlen \mantissa x\)" using bitlen_bounds by force have *: "nat (bitlen \mantissa x\) + nat (Float.exponent x + (2 ^ (LENGTH('e) - Suc 0) + int LENGTH('f)) - 2) \ LENGTH('f)" using that by (auto simp: is_denormal_Float_def nat_diff_distrib' le_diff_conv bitlen_nonneg nat_le_iff bias_def nat_add_distrib[symmetric]) have "\mantissa x\ * 2 ^ nat (Float.exponent x + int (bias TYPE(('e, 'f)float)) + LENGTH('f) - 1) < 2 ^ LENGTH('f)" apply (rule less_le_trans) apply (rule mult_strict_right_mono) apply (rule mless, force) unfolding power_add[symmetric] power_Suc[symmetric] apply (rule power_increasing) apply (auto simp: bias_def) using that * by (auto simp: is_denormal_Float_def algebra_simps) then show ?thesis apply (transfer fixing: x) apply transfer apply (simp add: Let_def nat_eq_iff take_bit_eq_mod) done qed definition of_finite_Float :: "Float.float \ ('e, 'f) float" where "of_finite_Float x = (if is_normal_Float TYPE(('e, 'f)float) x then normal_of_Float x else if is_denormal_Float TYPE(('e, 'f)float) x then denormal_of_Float x else 0)" lemma valof_normal_of_Float: "valof (normal_of_Float x::('e, 'f)float) = x" if "is_normal_Float TYPE(('e, 'f)float) x" proof - have "valof (normal_of_Float x::('e, 'f)float) = (- 1) ^ sign (normal_of_Float x::('e, 'f)float) * ((1 + real (nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\)) - 2 ^ LENGTH('f)) / 2 ^ LENGTH('f)) * 2 powr (bitlen \mantissa x\ - 1)) * 2 powr Float.exponent x" (is "_ = ?s * ?m * ?e") using that by (auto simp: is_normal_Float_def valof_eq fraction_normal_of_Float powr_realpow[symmetric] exponent_normal_of_Float powr_diff powr_add) also have "\mantissa x\ > 0" using that by (auto simp: is_normal_Float_def) have bound: "2 ^ LENGTH('f) \ nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\))" proof - have "(2::nat) ^ LENGTH('f) \ 2 ^ nat (bitlen \mantissa x\ - 1) * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\))" by (simp add: power_add[symmetric]) also have "\ \ nat \mantissa x\ * 2 ^ (Suc LENGTH('f) - nat (bitlen \mantissa x\))" using bitlen_bounds[of "\mantissa x\"] that by (auto simp: is_normal_Float_def) finally show ?thesis . qed have "?m = abs (mantissa x)" apply (subst of_nat_diff) subgoal using bound by auto subgoal using that by (auto simp: powr_realpow[symmetric] powr_add[symmetric] is_normal_Float_def bitlen_nonneg of_nat_diff divide_simps) done finally show ?thesis by (auto simp: mantissa_exponent sign_normal_of_Float abs_real_def zero_less_mult_iff) qed lemma valof_denormal_of_Float: "valof (denormal_of_Float x::('e, 'f)float) = x" if "is_denormal_Float TYPE(('e, 'f)float) x" proof - have less: "0 < Float.exponent x + (int (bias TYPE(('e, 'f) IEEE.float)) + int LENGTH('f))" using that by (auto simp: is_denormal_Float_def bias_def) have "valof (denormal_of_Float x::('e, 'f)float) = ((- 1) ^ sign (denormal_of_Float x::('e, 'f)float) * \real_of_int (mantissa x)\) * (2 powr real (nat (Float.exponent x + int (bias TYPE(('e, 'f) IEEE.float)) + int LENGTH('f) - 1)) / (2 powr real (bias TYPE(('e, 'f) IEEE.float)) * 2 powr LENGTH('f)) * 2)" (is "_ = ?m * ?e") by (auto simp: valof_eq exponent_denormal_of_Float fraction_denormal_of_Float that mantissa_exponent powr_realpow[symmetric]) also have "?m = mantissa x" by (auto simp: sign_denormal_of_Float abs_real_def mantissa_neg_iff) also have "?e = 2 powr Float.exponent x" by (auto simp: powr_add[symmetric] divide_simps powr_mult_base less ac_simps) finally show ?thesis by (simp add: mantissa_exponent) qed lemma valof_of_finite_Float: "is_finite_Float (TYPE(('e, 'f) IEEE.float)) x \ valof (of_finite_Float x::('e, 'f)float) = x" by (auto simp: of_finite_Float_def is_finite_Float_def valof_denormal_of_Float valof_normal_of_Float) lemma is_normal_normal_of_Float: "is_normal (normal_of_Float x::('e, 'f)float)" if "is_normal_Float TYPE(('e, 'f)float) x" using that by (auto simp: is_normal_def exponent_normal_of_Float that is_normal_Float_def emax_eq nat_less_iff) lemma is_denormal_denormal_of_Float: "is_denormal (denormal_of_Float x::('e, 'f)float)" if "is_denormal_Float TYPE(('e, 'f)float) x" using that by (auto simp: is_denormal_def exponent_denormal_of_Float that is_denormal_Float_def emax_eq fraction_denormal_of_Float le_nat_iff bias_def) lemma is_finite_of_finite_Float: "is_finite (of_finite_Float x)" by (auto simp: is_finite_def of_finite_Float_def is_normal_normal_of_Float is_denormal_denormal_of_Float) lemma Float_eq_zero_iff: "Float m e = 0 \ m = 0" by (metis Float.compute_is_float_zero Float_0_eq_0) lemma bitlen_mantissa_Float: shows "bitlen \mantissa (Float m e)\ = (if m = 0 then 0 else bitlen \m\ + e) - Float.exponent (Float m e)" using bitlen_Float[of m e] by auto lemma exponent_Float: shows "Float.exponent (Float m e) = (if m = 0 then 0 else bitlen \m\ + e) - bitlen \mantissa (Float m e)\ " using bitlen_Float[of m e] by auto lemma is_normal_Float_normal: "is_normal_Float TYPE(('e, 'f)float) (Float (normal_mantissa x) (normal_exponent x))" if "is_normal x" for x::"('e, 'f)float" proof - define f where "f = Float (normal_mantissa x) (normal_exponent x)" from that have "f \ 0" by (auto simp: f_def is_normal_def zero_float_def[symmetric] Float_eq_zero_iff normal_mantissa_def add_nonneg_eq_0_iff) from denormalize_shift[OF f_def this] obtain i where i: "normal_mantissa x = mantissa f * 2 ^ i" "normal_exponent x = Float.exponent f - int i" by auto have "mantissa f \ 0" by (auto simp: \f \ 0\ i mantissa_eq_zero_iff Float_eq_zero_iff) moreover have "normal_exponent x \ Float.exponent f" unfolding i by simp then have " bitlen \mantissa f\ \ 1 + int LENGTH('f)" unfolding bitlen_mantissa_Float bitlen_normal_mantissa f_def by auto moreover have "- int (bias TYPE(('e, 'f)float)) - bitlen \mantissa f\ + 1 < Float.exponent f" unfolding bitlen_mantissa_Float bitlen_normal_mantissa f_def using that by (auto simp: mantissa_eq_zero_iff abs_mult bias_def normal_mantissa_def normal_exponent_def is_normal_def emax_eq less_diff_conv add_nonneg_eq_0_iff) moreover have "2 ^ (LENGTH('e) - Suc 0) + - (1::int) * 2 ^ LENGTH('e) \ 0" by simp then have "(2::int) ^ (LENGTH('e) - Suc 0) < 1 + 2 ^ LENGTH('e)" by arith then have "Float.exponent f < 2 ^ LENGTH('e) - bitlen \mantissa f\ - int (bias TYPE(('e, 'f)float))" using normal_exponent_bounds_int[OF that] unfolding bitlen_mantissa_Float bitlen_normal_mantissa f_def by (auto simp: bias_def algebra_simps power_Suc[symmetric] intro: le_less_trans[OF add_right_mono] normal_exponent_bounds_int[OF that]) ultimately show ?thesis by (auto simp: is_normal_Float_def f_def) qed lemma is_denormal_Float_denormal: "is_denormal_Float TYPE(('e, 'f)float) (Float (denormal_mantissa x) (denormal_exponent TYPE(('e, 'f)float)))" if "is_denormal x" for x::"('e, 'f)float" proof - define f where "f = Float (denormal_mantissa x) (denormal_exponent TYPE(('e, 'f)float))" from that have "f \ 0" by (auto simp: f_def is_denormal_def zero_float_def[symmetric] Float_eq_zero_iff denormal_mantissa_def add_nonneg_eq_0_iff) from denormalize_shift[OF f_def this] obtain i where i: "denormal_mantissa x = mantissa f * 2 ^ i" "denormal_exponent TYPE(('e, 'f)float) = Float.exponent f - int i" by auto have "mantissa f \ 0" by (auto simp: \f \ 0\ i mantissa_eq_zero_iff Float_eq_zero_iff) moreover have "bitlen \mantissa f\ \ 1 - Float.exponent f - int (bias TYPE(('e, 'f) IEEE.float))" using \mantissa f \ 0\ unfolding f_def bitlen_mantissa_Float using bitlen_denormal_mantissa[of x] by (auto simp: denormal_exponent_def) moreover have "2 - 2 ^ (LENGTH('e) - Suc 0) - int LENGTH('f) \ Float.exponent f" (is "?l \ _") proof - have "?l \ denormal_exponent TYPE(('e, 'f)float) + i" using that by (auto simp: is_denormal_def bias_def denormal_exponent_def) also have "\ = Float.exponent f" unfolding i by auto finally show ?thesis . qed ultimately show ?thesis unfolding is_denormal_Float_def exponent_Float f_def[symmetric] by auto qed lemma is_finite_Float_of_finite: "is_finite_Float TYPE(('e, 'f)float) (of_finite x)" for x::"('e, 'f)float" by (auto simp: is_finite_Float_def of_finite_def is_normal_Float_normal is_denormal_Float_denormal) end diff --git a/thys/IP_Addresses/CIDR_Split.thy b/thys/IP_Addresses/CIDR_Split.thy --- a/thys/IP_Addresses/CIDR_Split.thy +++ b/thys/IP_Addresses/CIDR_Split.thy @@ -1,381 +1,382 @@ (* Title: CIDR_Split.thy Authors: Julius Michaelis, Cornelius Diekmann *) theory CIDR_Split imports IP_Address Prefix_Match Hs_Compat begin section\CIDR Split Motivation (Example for IPv4)\ text\When talking about ranges of IP addresses, we can make the ranges explicit by listing their elements.\ context begin private lemma "map (of_nat \ nat) [1 .. 4] = ([1, 2, 3, 4]:: 32 word list)" by eval private definition ipv4addr_upto :: "32 word \ 32 word \ 32 word list" where "ipv4addr_upto i j \ map (of_nat \ nat) [int (unat i) .. int (unat j)]" private lemma ipv4addr_upto: "set (ipv4addr_upto i j) = {i .. j}" proof - have int_interval_eq_image: "{int m..int n} = int ` {m..n}" for m n by (auto intro!: image_eqI [of _ int "nat k" for k]) have helpX:"\f (i::nat) (j::nat). (f \ nat) ` {int i..int j} = f ` {i .. j}" by (auto simp add: image_comp int_interval_eq_image) - { fix xa :: int - assume a1: "int (unat i) \ xa \ xa \ int (unat j)" - then have f2: "int (nat xa) = xa" - by force - then have "unat (of_int xa::32 word) = nat xa" - using a1 by (metis (full_types) le_unat_uoi nat_int nat_mono of_int_of_nat_eq) - then have "i \ of_int xa" and "of_int xa \ j" - using f2 a1 by (metis (no_types) uint_nat word_le_def)+ - } note hlp=this + have hlp: \i \ word_of_nat (nat xa)\ \word_of_nat (nat xa) \ j\ + if \uint i \ xa\ \xa \ uint j\ for xa :: int + proof - + from uint_nonnegative [of i] \uint i \ xa\ + have \0 \ xa\ by (rule order_trans) + moreover from \xa \ uint j\ uint_bounded [of j] + have \xa < 2 ^ 32\ by simp + ultimately have xa: \take_bit 32 xa = xa\ + by (simp add: take_bit_int_eq_self) + from xa \uint i \ xa\ show \i \ word_of_nat (nat xa)\ + by transfer simp + from xa \xa \ uint j\ show \word_of_nat (nat xa) \ j\ + by transfer simp + qed show ?thesis unfolding ipv4addr_upto_def - apply(intro set_eqI) - apply(simp) - apply(safe) - apply(simp_all) - using hlp apply blast - using hlp apply blast - apply (metis atLeastAtMost_iff imageI image_int_atLeastAtMost of_int_of_nat_eq word_le_nat_alt word_unat.Rep_inverse) - done + apply (rule set_eqI) + apply (auto simp add: hlp) + apply (metis (mono_tags) atLeastAtMost_iff image_iff unat_eq_nat_uint word_less_eq_iff_unsigned word_unat.Rep_inverse) + done qed text\The function @{const ipv4addr_upto} gives back a list of all the ips in the list. This list can be pretty huge! In the following, we will use CIDR notation (e.g. 192.168.0.0/24) to describe the list more compactly.\ end section\CIDR Split\ context begin private lemma find_SomeD: "find f x = Some y \ f y \ y \ set x" by(induction x; simp split: if_splits) (*pfxes needs a dummy parameter. The first parameter is a dummy that we have the 'a::len0 type and can refer to its length.*) private definition pfxes :: "'a::len0 itself \ nat list" where "pfxes _ = map nat [0..int(len_of TYPE ('a))]" private lemma "pfxes TYPE(32) = map nat [0 .. 32]" by eval private definition "largest_contained_prefix (a::('a :: len) word) r = ( let cs = (map (\s. PrefixMatch a s) (pfxes TYPE('a))); \ \anything that is a subset should also be a valid prefix. but try proving that.\ cfs = find (\s. valid_prefix s \ wordinterval_subset (prefix_to_wordinterval s) r) cs in cfs) " (* The joke is that it is always Some, given that a \ r. *) text\Split off one prefix:\ private definition wordinterval_CIDR_split1 :: "'a::len wordinterval \ 'a prefix_match option \ 'a wordinterval" where "wordinterval_CIDR_split1 r \ ( let ma = wordinterval_lowest_element r in case ma of None \ (None, r) | Some a \ (case largest_contained_prefix a r of None \ (None, r) | Some m \ (Some m, wordinterval_setminus r (prefix_to_wordinterval m))))" private lemma wordinterval_CIDR_split1_innard_helper: fixes a::"'a::len word" shows "wordinterval_lowest_element r = Some a \ largest_contained_prefix a r \ None" proof - assume a: "wordinterval_lowest_element r = Some a" have b: "(a,len_of(TYPE('a))) \ set (map (Pair a) (pfxes TYPE('a)))" unfolding pfxes_def set_map set_upto using Set.image_iff atLeastAtMost_iff int_eq_iff order_refl by metis (*400ms*) have c: "valid_prefix (PrefixMatch a (len_of(TYPE('a))))" by(simp add: valid_prefix_def pfxm_mask_def) have "wordinterval_to_set (prefix_to_wordinterval (PrefixMatch a (len_of(TYPE('a))))) = {a}" unfolding prefix_to_wordinterval_def pfxm_mask_def by simp moreover have "a \ wordinterval_to_set r" using a wordinterval_lowest_element_set_eq wordinterval_lowest_none_empty by (metis is_lowest_element_def option.distinct(1)) ultimately have d: "wordinterval_to_set (prefix_to_wordinterval (PrefixMatch a (LENGTH('a)))) \ wordinterval_to_set r" by simp show ?thesis unfolding largest_contained_prefix_def Let_def using b c d by(auto simp add: find_None_iff) qed private lemma r_split1_not_none: fixes r:: "'a::len wordinterval" shows "\ wordinterval_empty r \ fst (wordinterval_CIDR_split1 r) \ None" unfolding wordinterval_CIDR_split1_def Let_def by(cases "wordinterval_lowest_element r") (auto simp add: wordinterval_lowest_none_empty dest: wordinterval_CIDR_split1_innard_helper) private lemma largest_contained_prefix_subset: "largest_contained_prefix a r = Some p \ wordinterval_to_set (prefix_to_wordinterval p) \ wordinterval_to_set r" unfolding largest_contained_prefix_def Let_def by(drule find_SomeD) simp private lemma wordinterval_CIDR_split1_snd: "wordinterval_CIDR_split1 r = (Some s, u) \ u = wordinterval_setminus r (prefix_to_wordinterval s)" unfolding wordinterval_CIDR_split1_def Let_def by(clarsimp split: option.splits) private lemma largest_contained_prefix_subset_s1D: "wordinterval_CIDR_split1 r = (Some s, u) \ wordinterval_to_set (prefix_to_wordinterval s) \ wordinterval_to_set r" by(intro largest_contained_prefix_subset[where a = "the (wordinterval_lowest_element r)"]) (simp add: wordinterval_CIDR_split1_def split: option.splits) private theorem wordinterval_CIDR_split1_preserve: fixes r:: "'a::len wordinterval" shows "wordinterval_CIDR_split1 r = (Some s, u) \ wordinterval_eq (wordinterval_union (prefix_to_wordinterval s) u) r" proof(unfold wordinterval_eq_set_eq) assume as: "wordinterval_CIDR_split1 r = (Some s, u)" have ud: "u = wordinterval_setminus r (prefix_to_wordinterval s)" using as[THEN wordinterval_CIDR_split1_snd] . with largest_contained_prefix_subset_s1D[OF as] show "wordinterval_to_set (wordinterval_union (prefix_to_wordinterval s) u) = wordinterval_to_set r" unfolding ud by auto qed private lemma wordinterval_CIDR_split1_some_r_ne: "wordinterval_CIDR_split1 r = (Some s, u) \ \ wordinterval_empty r" proof(rule ccontr, goal_cases) case 1 have "wordinterval_lowest_element r = None" unfolding wordinterval_lowest_none_empty using 1(2) unfolding not_not . then have "wordinterval_CIDR_split1 r = (None, r)" unfolding wordinterval_CIDR_split1_def Let_def by simp then show False using 1(1) by simp qed private lemma wordinterval_CIDR_split1_distinct: fixes r:: "'a::len wordinterval" shows "wordinterval_CIDR_split1 r = (Some s, u) \ wordinterval_empty (wordinterval_intersection (prefix_to_wordinterval s) u)" proof(goal_cases) case 1 have nn: "wordinterval_lowest_element r \ None" using wordinterval_CIDR_split1_some_r_ne 1 wordinterval_lowest_none_empty by metis from 1 have "u = wordinterval_setminus r (prefix_to_wordinterval s)" by(elim wordinterval_CIDR_split1_snd) then show ?thesis by simp qed private lemma wordinterval_CIDR_split1_distinct2: fixes r:: "'a::len wordinterval" shows "wordinterval_CIDR_split1 r = (Some s, u) \ wordinterval_empty (wordinterval_intersection (prefix_to_wordinterval s) u)" by(rule wordinterval_CIDR_split1_distinct[where r = r]) simp function wordinterval_CIDR_split_prefixmatch :: "'a::len wordinterval \ 'a prefix_match list" where "wordinterval_CIDR_split_prefixmatch rs = ( if \ wordinterval_empty rs then case wordinterval_CIDR_split1 rs of (Some s, u) \ s # wordinterval_CIDR_split_prefixmatch u | _ \ [] else [] )" by pat_completeness simp termination wordinterval_CIDR_split_prefixmatch proof(relation "measure (card \ wordinterval_to_set)", rule wf_measure, unfold in_measure comp_def, goal_cases) note vernichter = wordinterval_empty_set_eq wordinterval_intersection_set_eq wordinterval_union_set_eq wordinterval_eq_set_eq case (1 rs x y x2) note some = 1(2)[unfolded 1(3), symmetric] from prefix_never_empty have "wordinterval_to_set (prefix_to_wordinterval x2) \ {}" unfolding vernichter . thus ?case unfolding wordinterval_CIDR_split1_preserve[OF some, unfolded vernichter, symmetric] unfolding card_Un_disjoint[OF finite finite wordinterval_CIDR_split1_distinct[OF some, unfolded vernichter]] by auto qed private lemma unfold_rsplit_case: assumes su: "(Some s, u) = wordinterval_CIDR_split1 rs" shows "(case wordinterval_CIDR_split1 rs of (None, u) \ [] | (Some s, u) \ s # wordinterval_CIDR_split_prefixmatch u) = s # wordinterval_CIDR_split_prefixmatch u" using su by (metis option.simps(5) split_conv) lemma "wordinterval_CIDR_split_prefixmatch (RangeUnion (WordInterval (0x40000000) 0x5FEFBBCC) (WordInterval 0x5FEEBB1C 0x7FFFFFFF)) = [PrefixMatch (0x40000000::32 word) 2]" by eval lemma "length (wordinterval_CIDR_split_prefixmatch (WordInterval 0 (0xFFFFFFFE::32 word))) = 32" by eval declare wordinterval_CIDR_split_prefixmatch.simps[simp del] theorem wordinterval_CIDR_split_prefixmatch: "wordinterval_to_set r = (\x\set (wordinterval_CIDR_split_prefixmatch r). prefix_to_wordset x)" proof(induction r rule: wordinterval_CIDR_split_prefixmatch.induct) case (1 rs) show ?case proof(cases "wordinterval_empty rs") case True thus ?thesis by(simp add: wordinterval_CIDR_split_prefixmatch.simps) next case False obtain x y where s1: "wordinterval_CIDR_split1 rs = (Some x, y)" using r_split1_not_none[OF False] by(auto simp add: fst_def split: prod.splits) have mIH: "wordinterval_to_set y = (\x\set (wordinterval_CIDR_split_prefixmatch y). prefix_to_wordset x)" using 1[OF False s1[symmetric] refl] . have *: "wordinterval_to_set rs = prefix_to_wordset x \ (\x\set (wordinterval_CIDR_split_prefixmatch y). prefix_to_wordset x)" unfolding mIH[symmetric] proof - have ud: "y = wordinterval_setminus rs (prefix_to_wordinterval x)" using wordinterval_CIDR_split1_snd[OF s1] . have ss: "prefix_to_wordset x \ wordinterval_to_set rs" using largest_contained_prefix_subset_s1D[OF s1] by simp show "wordinterval_to_set rs = prefix_to_wordset x \ wordinterval_to_set y" unfolding ud using ss by simp blast qed show ?thesis apply(subst wordinterval_CIDR_split_prefixmatch.simps) apply(unfold if_P[OF False] s1 prod.simps option.simps *) (* WOOOOO simplifier bug (* try making this a simp add: *) *) apply(simp) done qed qed lemma wordinterval_CIDR_split_prefixmatch_all_valid_Ball: fixes r:: "'a::len wordinterval" shows "\e\set (wordinterval_CIDR_split_prefixmatch r). valid_prefix e \ pfxm_length e \ LENGTH('a)" (* The induction is somewhat verbose, so it is less annoying to write the two down at once *) proof(induction r rule: wordinterval_CIDR_split_prefixmatch.induct) case 1 case (1 rs) show ?case proof(cases "wordinterval_empty rs") case False obtain x y where s1: "wordinterval_CIDR_split1 rs = (Some x, y)" using r_split1_not_none[OF False] by(auto simp add: fst_def split: prod.splits) hence i1: "valid_prefix x" unfolding wordinterval_CIDR_split1_def Let_def largest_contained_prefix_def by(auto dest: find_SomeD split: option.splits) have i2: "pfxm_length x \ LENGTH('a)" using s1 unfolding wordinterval_CIDR_split1_def Let_def largest_contained_prefix_def pfxes_def by(force split: option.splits dest: find_SomeD simp: nat_le_iff) have mIH: "\a\set (wordinterval_CIDR_split_prefixmatch y). valid_prefix a \ pfxm_length a \ LENGTH('a)" using 1[OF False s1[symmetric] refl] . with i1 i2 show ?thesis apply(subst wordinterval_CIDR_split_prefixmatch.simps) apply(unfold if_P[OF False] s1 prod.simps option.simps) apply(simp) done qed (simp add: wordinterval_CIDR_split_prefixmatch.simps) qed private lemma wordinterval_CIDR_split_prefixmatch_all_valid_less_Ball_hlp: "x \ set [s\map (PrefixMatch x2) (pfxes TYPE('a::len0)) . valid_prefix s \ wordinterval_to_set (prefix_to_wordinterval s) \ wordinterval_to_set rs] \ pfxm_length x \ LENGTH('a)" by(clarsimp simp: pfxes_def) presburger text\Since @{const wordinterval_CIDR_split_prefixmatch} only returns valid prefixes, we can safely convert it to CIDR lists\ (* actually, just valid_prefix doesn't mean that the prefix length is sane. Fortunately, wordinterval_CIDR_split_prefixmatch_all_valid_Ball does entail that *) lemma "valid_prefix (PrefixMatch (0::16 word) 20)" by(simp add: valid_prefix_def) lemma wordinterval_CIDR_split_disjunct: "a \ set (wordinterval_CIDR_split_prefixmatch i) \ b \ set (wordinterval_CIDR_split_prefixmatch i) \ a \ b \ prefix_to_wordset a \ prefix_to_wordset b = {}" proof(induction i rule: wordinterval_CIDR_split_prefixmatch.induct) case (1 rs) note IH = 1(1) have prema: "a \ set (wordinterval_CIDR_split_prefixmatch rs)" (is "a \ ?os") using 1 by simp have premb: "b \ ?os" using 1 by simp show ?case proof(cases "wordinterval_empty rs") case False obtain x y where s1: "wordinterval_CIDR_split1 rs = (Some x, y)" using r_split1_not_none[OF False] by(auto simp add: fst_def split: prod.splits) have mi: "k \ set (wordinterval_CIDR_split_prefixmatch y)" (is "k \ ?rs") if p: "k \ x" "k \ ?os" for k using p s1 by(subst (asm) wordinterval_CIDR_split_prefixmatch.simps) (simp only: if_P[OF False] split: prod.splits option.splits; simp) have a: "k \ ?rs \ prefix_to_wordset k \ wordinterval_to_set y" for k (* this is actually a quite general statement, might make a lemma out of it *) unfolding wordinterval_CIDR_split_prefixmatch by blast have b: "prefix_to_wordset x \ wordinterval_to_set y = {}" using wordinterval_CIDR_split1_snd[OF s1] by simp show ?thesis proof(cases "a = x"; cases "b = x") assume as: "a = x" "b \ x" with a[OF mi[OF as(2) premb]] b show ?thesis by blast next assume as: "a \ x" "b = x" with a[OF mi[OF as(1) prema]] b show ?thesis by blast next assume as: "a \ x" "b \ x" (* Nothing to do case *) have i: "a \ ?rs" "b \ ?rs" using as mi prema premb by blast+ show "prefix_to_wordset a \ prefix_to_wordset b = {}" by(rule IH[OF False s1[symmetric] refl i]) (fact 1) next assume as: "a = x" "b = x" (* impossible case *) with 1 have False by simp thus ?thesis .. qed next case True hence "wordinterval_CIDR_split_prefixmatch rs = []" by(simp add: wordinterval_CIDR_split_prefixmatch.simps) thus ?thesis using prema by simp qed qed lemma wordinterval_CIDR_split_distinct: "distinct (wordinterval_CIDR_split_prefixmatch i)" (* wish this was a corollary *) proof(induction i rule: wordinterval_CIDR_split_prefixmatch.induct) case (1 rs) show ?case proof(cases "wordinterval_empty rs") case False obtain x y where s1: "wordinterval_CIDR_split1 rs = (Some x, y)" using r_split1_not_none[OF False] by(auto simp add: fst_def split: prod.splits) have mIH: "distinct (wordinterval_CIDR_split_prefixmatch y)" using 1[OF False s1[symmetric] refl] . have "prefix_to_wordset x \ wordinterval_to_set y = {}" using wordinterval_CIDR_split1_snd[OF s1] by simp hence i1: "x \ set (wordinterval_CIDR_split_prefixmatch y)" unfolding wordinterval_CIDR_split_prefixmatch using prefix_never_empty[of x, simplified] by blast show ?thesis using s1 by(subst wordinterval_CIDR_split_prefixmatch.simps) (simp add: if_P[OF False] mIH i1 split: option.splits prod.splits) qed (simp add: wordinterval_CIDR_split_prefixmatch.simps) qed lemma wordinterval_CIDR_split_existential: "x \ wordinterval_to_set w \ \s. s \ set (wordinterval_CIDR_split_prefixmatch w) \ x \ prefix_to_wordset s" using wordinterval_CIDR_split_prefixmatch[symmetric] by fastforce subsection\Versions for @{const ipset_from_cidr}\ definition cidr_split :: "'i::len wordinterval \ ('i word \ nat) list" where "cidr_split rs \ map prefix_match_to_CIDR (wordinterval_CIDR_split_prefixmatch rs)" corollary cidr_split_prefix: fixes r :: "'i::len wordinterval" shows "(\x\set (cidr_split r). uncurry ipset_from_cidr x) = wordinterval_to_set r" unfolding wordinterval_CIDR_split_prefixmatch[symmetric] cidr_split_def apply(simp add: prefix_match_to_CIDR_def2 wordinterval_CIDR_split_prefixmatch) using prefix_to_wordset_ipset_from_cidr wordinterval_CIDR_split_prefixmatch_all_valid_Ball by blast corollary cidr_split_prefix_single: fixes start :: "'i::len word" shows "(\x\set (cidr_split (iprange_interval (start, end))). uncurry ipset_from_cidr x) = {start..end}" unfolding wordinterval_to_set.simps[symmetric] using cidr_split_prefix iprange_interval.simps by metis private lemma interval_in_splitD: "xa \ foo \ prefix_to_wordset xa \ \(prefix_to_wordset ` foo)" by auto lemma cidrsplit_no_overlaps: "\ x \ set (wordinterval_CIDR_split_prefixmatch wi); xa \ set (wordinterval_CIDR_split_prefixmatch wi); pt && ~~ (pfxm_mask x) = pfxm_prefix x; pt && ~~ (pfxm_mask xa) = pfxm_prefix xa \ \ x = xa" proof(rule ccontr, goal_cases) case 1 hence "prefix_match_semantics x pt" "prefix_match_semantics xa pt" unfolding prefix_match_semantics_def by (simp_all add: word_bw_comms(1)) moreover have "valid_prefix x" "valid_prefix xa" using 1(1-2) wordinterval_CIDR_split_prefixmatch_all_valid_Ball by blast+ ultimately have "pt \ prefix_to_wordset x" "pt \ prefix_to_wordset xa" using prefix_match_semantics_wordset by blast+ with wordinterval_CIDR_split_disjunct[OF 1(1,2) 1(5)] show False by blast qed end end diff --git a/thys/IP_Addresses/IP_Address_Parser.thy b/thys/IP_Addresses/IP_Address_Parser.thy --- a/thys/IP_Addresses/IP_Address_Parser.thy +++ b/thys/IP_Addresses/IP_Address_Parser.thy @@ -1,130 +1,130 @@ theory IP_Address_Parser imports IP_Address IPv4 IPv6 "HOL-Library.Code_Target_Nat" (*!!*) begin section\Parsing IP Addresses\ subsection\IPv4 Parser\ ML\ local fun extract_int ss = case ss |> implode |> Int.fromString of SOME i => i | NONE => raise Fail "unparsable int"; fun mk_nat maxval i = if i < 0 orelse i > maxval then raise Fail("nat ("^Int.toString i^") must be between 0 and "^Int.toString maxval) else (HOLogic.mk_number HOLogic.natT i); val mk_nat255 = mk_nat 255; fun mk_quadrupel (((a,b),c),d) = HOLogic.mk_prod (mk_nat255 a, HOLogic.mk_prod (mk_nat255 b, HOLogic.mk_prod (mk_nat255 c, mk_nat255 d))); in fun mk_ipv4addr ip = @{const ipv4addr_of_dotdecimal} $ mk_quadrupel ip; val parser_ipv4 = (Scan.many1 Symbol.is_ascii_digit >> extract_int) --| ($$ ".") -- (Scan.many1 Symbol.is_ascii_digit >> extract_int) --| ($$ ".") -- (Scan.many1 Symbol.is_ascii_digit >> extract_int) --| ($$ ".") -- (Scan.many1 Symbol.is_ascii_digit >> extract_int); end; local val (ip_term, rest) = "10.8.0.255" |> raw_explode |> Scan.finite Symbol.stopper (parser_ipv4 >> mk_ipv4addr); in val _ = if rest <> [] then raise Fail "did not parse everything" else writeln "parsed"; val _ = if Code_Evaluation.dynamic_value_strict @{context} ip_term <> @{term "168296703::ipv4addr"} then raise Fail "parser failed" else writeln "test passed"; end; \ subsection\IPv6 Parser\ definition mk_ipv6addr :: "16 word option list \ ipv6addr_syntax option" where "mk_ipv6addr partslist = ( let \ \remove empty lists to the beginning and end if omission occurs at start/end\ \ \to join over \<^verbatim>\:\ properly\ fix_start = (\ps. case ps of None#None#_ \ tl ps | _ \ ps); fix_end = (\ps. case rev ps of None#None#_ \ butlast ps | _ \ ps); ps = (fix_end \ fix_start) partslist in if length (filter (\p. p = None) ps) = 1 then ipv6_unparsed_compressed_to_preferred ps else case ps of [Some a,Some b,Some c,Some d,Some e,Some f,Some g,Some h] \ Some (IPv6AddrPreferred a b c d e f g h) | _ \ None )" ML\ local val fromHexString = StringCvt.scanString (Int.scan StringCvt.HEX); fun extract_int ss = case ss of "" => NONE | xs => case xs |> fromHexString of SOME i => SOME i | NONE => raise Fail "unparsable int"; in val mk_ipv6addr = map (fn p => case p of NONE => @{const None ("16 word")} | SOME i => @{const Some ("16 word")} $ - (@{const word_of_int (16)} $ + (@{const of_int ("16 word")} $ HOLogic.mk_number HOLogic.intT i) ) #> HOLogic.mk_list @{typ "16 word option"} (*TODO: never use THE! is there some option_dest?*) #> (fn x => @{const ipv6preferred_to_int} $ (@{const the ("ipv6addr_syntax")} $ (@{const mk_ipv6addr} $ x))); val parser_ipv6 = Scan.many1 (fn x => Symbol.is_ascii_hex x orelse x = ":") >> (implode #> space_explode ":" #> map extract_int) (* a different implementation which returns a list of exploded strings: Scan.repeat ((Scan.many Symbol.is_ascii_hex >> extract_int) --| ($$ ":")) @@@ (Scan.many Symbol.is_ascii_hex >> extract_int >> (fn p => [p]))*) end; local val parse_ipv6 = raw_explode #> Scan.finite Symbol.stopper (parser_ipv6 >> mk_ipv6addr); fun unit_test (ip_string, ip_result) = let val (ip_term, rest) = ip_string |> parse_ipv6; val _ = if rest <> [] then raise Fail "did not parse everything" else (); val _ = Code_Evaluation.dynamic_value_strict @{context} ip_term |> Syntax.pretty_term @{context} |> Pretty.writeln; val _ = if Code_Evaluation.dynamic_value_strict @{context} ip_term <> ip_result then raise Fail "parser failed" else writeln ("test passed for "^ip_string); in () end; in val _ = map unit_test [("10:ab:FF:0::FF:4:255", @{term "83090298060623265259947972050027093::ipv6addr"}) ,("2001:db8::8:800:200c:417a", @{term "42540766411282592856906245548098208122::ipv6addr"}) ,("ff01::101", @{term "338958331222012082418099330867817087233::ipv6addr"}) ,("::8:800:200c:417a", @{term "2260596444381562::ipv6addr"}) ,("2001:db8::", @{term "42540766411282592856903984951653826560::ipv6addr"}) ,("ff00::", @{term "338953138925153547590470800371487866880::ipv6addr"}) ,("fe80::", @{term "338288524927261089654018896841347694592::ipv6addr"}) ,("1::", @{term "5192296858534827628530496329220096::ipv6addr"}) ,("1::", @{term "5192296858534827628530496329220096::ipv6addr"}) ,("::", @{term "0::ipv6addr"}) ,("::1", @{term "1::ipv6addr"}) ,("2001:db8:0:1:1:1:1:1", @{term "42540766411282592875351010504635121665::ipv6addr"}) ,("ffff:ffff:ffff:ffff:ffff:ffff:ffff:ffff", @{term "340282366920938463463374607431768211455::ipv6addr"}) ]; 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,257 +1,257 @@ (* 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 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 = max_word" 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: ipv4addr_of_nat_def le_unat_uoi nat_of_ipv4addr_def) - 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) + 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)\ 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 word_of_nat nat_of_ipv4addr_def) + 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) + 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) 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: Word_Lemmas.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) 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) apply(simp add: and_mask_bl_take) 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; fail) apply(simp add: ipv4addr_of_dotdecimal_bit dotdecimal_of_ipv4addr.simps) apply(simp add: ipv4addr_of_nat_nat_of_ipv4addr) apply(simp add: 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) 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 mask_eq) 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_dotdecimal.simps ipv4addr_of_nat_def ipset_from_cidr_def ipset_from_netmask_def mask_eq) 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 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,945 +1,945 @@ (* 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)" - by(simp add: ipv6addr_of_nat_def word_of_nat) +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 = max_word" 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_def ipv6addr_of_nat_def le_unat_uoi) - 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) + 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\ 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\ 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 add: mask_eq)+ 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" by((subst word128_masks_ipv6pieces)+, subst ucast_ipv6_piece, simp_all)+ 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) 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) apply(rule ipv6addr_16word_pieces_compose_or) 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) 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) (*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 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] )" by(simp add: ipv6_preferred_to_compressed.simps max_zero_streak_def List_explode_goup_by_zeros) 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/Interval_Arithmetic_Word32/Interval_Word32.thy b/thys/Interval_Arithmetic_Word32/Interval_Word32.thy --- a/thys/Interval_Arithmetic_Word32/Interval_Word32.thy +++ b/thys/Interval_Arithmetic_Word32/Interval_Word32.thy @@ -1,3886 +1,4146 @@ (* Author: Brandon Bohrer *) theory Interval_Word32 imports Complex_Main "HOL-Word.More_Word" Word_Lib.Word_Lemmas Word_Lib.Word_Lib Word_Lib.Word_Syntax Word_Lib.Bitwise begin +abbreviation signed_real_of_word :: \'a::len word \ real\ + where \signed_real_of_word \ signed\ + +lemma (in linordered_idom) signed_less_numeral_iff: + \signed w < numeral n \ sint w < numeral n\ (is \?P \ ?Q\) +proof - + have \?Q \ of_int (sint w) < of_int (numeral n)\ + by (simp only: of_int_less_iff) + also have \\ \ ?P\ + by (transfer fixing: less less_eq n) simp + finally show ?thesis .. +qed + +lemma (in linordered_idom) signed_less_neg_numeral_iff: + \signed w < - numeral n \ sint w < - numeral n\ (is \?P \ ?Q\) +proof - + have \?Q \ of_int (sint w) < of_int (- numeral n)\ + by (simp only: of_int_less_iff) + also have \\ \ ?P\ + by (transfer fixing: less less_eq uminus n) simp + finally show ?thesis .. +qed + +lemma (in linordered_idom) numeral_less_signed_iff: + \numeral n < signed w \ numeral n < sint w\ (is \?P \ ?Q\) +proof - + have \?Q \ of_int (numeral n) < of_int (sint w)\ + by (simp only: of_int_less_iff) + also have \\ \ ?P\ + by (transfer fixing: less less_eq n) simp + finally show ?thesis .. +qed + +lemma (in linordered_idom) neg_numeral_less_signed_iff: + \- numeral n < signed w \ - numeral n < sint w\ (is \?P \ ?Q\) +proof - + have \?Q \ of_int (- numeral n) < of_int (sint w)\ + by (simp only: of_int_less_iff) + also have \\ \ ?P\ + by (transfer fixing: less less_eq uminus n) simp + finally show ?thesis .. +qed + +lemma (in linordered_idom) signed_nonnegative_iff: + \0 \ signed w \ 0 \ sint w\ (is \?P \ ?Q\) +proof - + have \?Q \ of_int 0 \ of_int (sint w)\ + by (simp only: of_int_le_iff) + also have \\ \ ?P\ + by (transfer fixing: less_eq) simp + finally show ?thesis .. +qed + +lemma signed_real_of_word_plus_numeral_eq_signed_real_of_word_iff: + \signed_real_of_word v + numeral n = signed_real_of_word w + \ sint v + numeral n = sint w\ (is \?P \ ?Q\) +proof - + have \?Q \ real_of_int (sint v + numeral n) = real_of_int (sint w)\ + by (simp only: of_int_eq_iff) + also have \\ \ ?P\ + by simp + finally show ?thesis .. +qed + +lemma signed_real_of_word_sum_less_eq_numeral_iff: + \signed_real_of_word v + signed_real_of_word w \ numeral n + \ sint v + sint w \ numeral n\ (is \?P \ ?Q\) +proof - + have \?Q \ real_of_int (sint v + sint w) \ real_of_int (numeral n)\ + by (simp only: of_int_le_iff) + also have \\ \ ?P\ + by simp + finally show ?thesis .. +qed + +lemma signed_real_of_word_sum_less_eq_neg_numeral_iff: + \signed_real_of_word v + signed_real_of_word w \ - numeral n + \ sint v + sint w \ - numeral n\ (is \?P \ ?Q\) +proof - + have \?Q \ real_of_int (sint v + sint w) \ real_of_int (- numeral n)\ + by (simp only: of_int_le_iff) + also have \\ \ ?P\ + by simp + finally show ?thesis .. +qed + +lemma signed_real_of_word_sum_less_numeral_iff: + \signed_real_of_word v + signed_real_of_word w < numeral n + \ sint v + sint w < numeral n\ (is \?P \ ?Q\) +proof - + have \?Q \ real_of_int (sint v + sint w) < real_of_int (numeral n)\ + by (simp only: of_int_less_iff) + also have \\ \ ?P\ + by simp + finally show ?thesis .. +qed + +lemma signed_real_of_word_sum_less_neg_numeral_iff: + \signed_real_of_word v + signed_real_of_word w < - numeral n + \ sint v + sint w < - numeral n\ (is \?P \ ?Q\) +proof - + have \?Q \ real_of_int (sint v + sint w) < real_of_int (- numeral n)\ + by (simp only: of_int_less_iff) + also have \\ \ ?P\ + by simp + finally show ?thesis .. +qed + +lemma numeral_less_eq_signed_real_of_word_sum: + \numeral n \ signed_real_of_word v + signed_real_of_word w + \ numeral n \ sint v + sint w\ (is \?P \ ?Q\) +proof - + have \?Q \ real_of_int (numeral n) \ real_of_int (sint v + sint w)\ + by (simp only: of_int_le_iff) + also have \\ \ ?P\ + by simp + finally show ?thesis .. +qed + +lemma neg_numeral_less_eq_signed_real_of_word_sum: + \- numeral n \ signed_real_of_word v + signed_real_of_word w + \ - numeral n \ sint v + sint w\ (is \?P \ ?Q\) +proof - + have \?Q \ real_of_int (- numeral n) \ real_of_int (sint v + sint w)\ + by (simp only: of_int_le_iff) + also have \\ \ ?P\ + by simp + finally show ?thesis .. +qed + +lemma numeral_less_signed_real_of_word_sum: + \numeral n < signed_real_of_word v + signed_real_of_word w + \ numeral n < sint v + sint w\ (is \?P \ ?Q\) +proof - + have \?Q \ real_of_int (numeral n) < real_of_int (sint v + sint w)\ + by (simp only: of_int_less_iff) + also have \\ \ ?P\ + by simp + finally show ?thesis .. +qed + +lemma neg_numeral_less_signed_real_of_word_sum: + \- numeral n < signed_real_of_word v + signed_real_of_word w + \ - numeral n < sint v + sint w\ (is \?P \ ?Q\) +proof - + have \?Q \ real_of_int (- numeral n) < real_of_int (sint v + sint w)\ + by (simp only: of_int_less_iff) + also have \\ \ ?P\ + by simp + finally show ?thesis .. +qed + +lemmas real_of_word_simps [simp] = signed_less_numeral_iff [where ?'a = real] + numeral_less_signed_iff [where ?'a = real] + signed_less_neg_numeral_iff [where ?'a = real] + neg_numeral_less_signed_iff [where ?'a = real] + signed_nonnegative_iff [where ?'a = real] + +lemmas more_real_of_word_simps = + signed_real_of_word_plus_numeral_eq_signed_real_of_word_iff + signed_real_of_word_sum_less_eq_numeral_iff + signed_real_of_word_sum_less_eq_neg_numeral_iff + signed_real_of_word_sum_less_numeral_iff + signed_real_of_word_sum_less_neg_numeral_iff + numeral_less_eq_signed_real_of_word_sum + neg_numeral_less_eq_signed_real_of_word_sum + numeral_less_signed_real_of_word_sum + neg_numeral_less_signed_real_of_word_sum + +declare more_real_of_word_simps [simp] + text\Interval-Word32.thy implements conservative interval arithmetic operators on 32-bit word values, with explicit infinities for values outside the representable bounds. It is suitable for use in interpreters for languages which must have a well-understood low-level behavior (see Interpreter.thy). This work was originally part of the paper by Bohrer \emph{et al.}~\cite{BohrerTMMP18}. It is worth noting that this is not the first formalization of interval arithmetic in Isabelle/HOL. This article is presented regardless because it has unique goals in mind which have led to unique design decisions. Our goal is generate code which can be used to perform conservative arithmetic in implementations extracted from a proof. The Isabelle standard library now features interval arithmetic, for example in Approximation.thy. Ours differs in two ways: 1) We use intervals with explicit positive and negative infinities, and with overflow checking. Such checking is often relevant in implementation-level code with unknown inputs. To promote memory-efficient implementations, we moreover use sentinel values for infinities, rather than datatype constructors. This is especially important in real-time settings where the garbarge collection required for datatypes can be a concern. 2) Our goal is not to use interval arithmetic to discharge Isabelle goals, but to generate useful proven-correct implementation code, see Interpreter.thy. On the other hand, we are not concerned with producing interval-based automation for arithmetic goals in HOL. In practice, much of the work in this theory comes down to sheer case-analysis. Bounds-checking requires many edge cases in arithmetic functions, which come with many cases in proofs. Where possible, we attempt to offload interesting facts about word representations of numbers into reusable lemmas, but even then main results require many subcases, each with a certain amount of arithmetic grunt work. \ section \Interval arithmetic definitions\ subsection \Syntax\ text\Words are 32-bit\ type_synonym word = "32 Word.word" text\Sentinel values for infinities. Note that we leave the maximum value ($2^31$) completed unused, so that negation of $(2^{31})-1$ is not an edge case\ definition NEG_INF::"word" where NEG_INF_def[simp]:"NEG_INF = -((2 ^ 31) -1)" definition NegInf::"real" where NegInf[simp]:"NegInf = real_of_int (sint NEG_INF)" definition POS_INF::"word" where POS_INF_def[simp]:"POS_INF = (2^31) - 1" definition PosInf::"real" where PosInf[simp]:"PosInf = real_of_int (sint POS_INF)" text\Subtype of words who represent a finite value. \ typedef bword = "{n::word. sint n \ sint NEG_INF \ sint n \ sint POS_INF}" apply(rule exI[where x=NEG_INF]) by (auto) text\Numeric literals\ type_synonym lit = bword setup_lifting type_definition_bword lift_definition bword_zero::"bword" is "0::32 Word.word" by auto lift_definition bword_one::"bword" is "1::32 Word.word" by(auto simp add: sint_uint) lift_definition bword_neg_one::"bword" is "-1::32 Word.word" by(auto) definition word::"word \ bool" where word_def[simp]:"word w \ w \ {NEG_INF..POS_INF}" named_theorems rep_simps "Simplifications for representation functions" text\Definitions of interval containment and word representation repe w r iff word w encodes real number r\ inductive repe ::"word \ real \ bool" (infix "\\<^sub>E" 10) where repPOS_INF:"r \ real_of_int (sint POS_INF) \ repe POS_INF r" | repNEG_INF:"r \ real_of_int (sint NEG_INF) \ repe NEG_INF r" | repINT:"(sint w) < real_of_int(sint POS_INF) \ (sint w) > real_of_int(sint NEG_INF) \ repe w (sint w)" inductive_simps repePos_simps[rep_simps]:"repe POS_INF r" and repeNeg_simps[rep_simps]:"repe NEG_INF r" and repeInt_simps[rep_simps]:"repe w (sint w)" text\repU w r if w represents an upper bound of r\ definition repU ::"word \ real \ bool" (infix "\\<^sub>U" 10) where "repU w r \ \ r'. r' \ r \ repe w r'" lemma repU_leq:"repU w r \ r' \ r \ repU w r'" unfolding repU_def using order_trans by auto text\repU w r if w represents a lower bound of r\ definition repL ::"word \ real \ bool" (infix "\\<^sub>L" 10) where "repL w r \ \ r'. r' \ r \ repe w r'" lemma repL_geq:"repL w r \ r' \ r \ repL w r'" unfolding repL_def using order_trans by auto text\repP (l,u) r iff l and u encode lower and upper bounds of r\ definition repP ::"word * word \ real \ bool" (infix "\\<^sub>P" 10) where "repP w r \ let (w1, w2) = w in repL w1 r \ repU w2 r" lemma int_not_posinf: assumes b1:"real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:"real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "ra \ POS_INF" using b1 b2 by auto lemma int_not_neginf: assumes b1:" real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:" real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "ra \ NEG_INF" using b1 b2 by auto lemma int_not_undef: assumes b1:"real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:"real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "ra \ NEG_INF-1" - using b1 b2 by auto + using b1 b2 by auto lemma sint_range: assumes b1:"real_of_int (sint ra) < real_of_int (sint POS_INF)" assumes b2:"real_of_int (sint NEG_INF) < real_of_int (sint ra)" shows "sint ra \ {i. i > -((2^31)-1) \ i < (2^31)-1}" using b1 b2 by auto lemma word_size_neg: fixes w :: "32 Word.word" shows "size (-w) = size w" using Word.word_size[of w] Word.word_size[of "-w"] by auto lemma uint_distinct: fixes w1 w2 shows "w1 \ w2 \ uint w1 \ uint w2" by auto section \Preliminary lemmas\ subsection \Case analysis lemmas\ text\Case analysis principle for pairs of intervals, used in proofs of arithmetic operations\ lemma ivl_zero_case: fixes l1 u1 l2 u2 :: real assumes ivl1:"l1 \ u1" assumes ivl2:"l2 \ u2" shows "(l1 \ 0 \ 0 \ u1 \ l2 \ 0 \ 0 \ u2) \(l1 \ 0 \ 0 \ u1 \ 0 \ l2) \(l1 \ 0 \ 0 \ u1 \ u2 \ 0) \(0 \ l1 \ l2 \ 0 \ 0 \ u2) \(u1 \ 0 \ l2 \ 0 \ 0 \ u2) \(u1 \ 0 \ u2 \ 0) \(u1 \ 0 \ 0 \ l2) \(0 \ l1 \ u2 \ 0) \(0 \ l1 \ 0 \ l2)" using ivl1 ivl2 by (metis le_cases) lemma case_ivl_zero [consumes 2, case_names ZeroZero ZeroPos ZeroNeg PosZero NegZero NegNeg NegPos PosNeg PosPos]: fixes l1 u1 l2 u2 :: real shows "l1 \ u1 \ l2 \ u2 \ ((l1 \ 0 \ 0 \ u1 \ l2 \ 0 \ 0 \ u2) \ P) \ ((l1 \ 0 \ 0 \ u1 \ 0 \ l2) \ P) \ ((l1 \ 0 \ 0 \ u1 \ u2 \ 0) \ P) \ ((0 \ l1 \ l2 \ 0 \ 0 \ u2) \ P) \ ((u1 \ 0 \ l2 \ 0 \ 0 \ u2) \ P) \ ((u1 \ 0 \ u2 \ 0) \ P) \ ((u1 \ 0 \ 0 \ l2) \ P) \ ((0 \ l1 \ u2 \ 0) \ P) \ ((0 \ l1 \ 0 \ l2) \ P) \ P" using ivl_zero_case[of l1 u1 l2 u2] by auto lemma case_inf2[case_names PosPos PosNeg PosNum NegPos NegNeg NegNum NumPos NumNeg NumNum]: shows "\w1 w2 P. (w1 = POS_INF \ w2 = POS_INF \ P w1 w2) \ (w1 = POS_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 = POS_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 = POS_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = POS_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ P w1 w2" by(auto) lemma case_pu_inf[case_names PosAny AnyPos NegNeg NegNum NumNeg NumNum]: shows "\w1 w2 P. (w1 = POS_INF \ P w1 w2) \ (w2 = POS_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 = NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ P w1 w2" by(auto) lemma case_pl_inf[case_names NegAny AnyNeg PosPos PosNum NumPos NumNum]: shows "\w1 w2 P. (w1 = NEG_INF \ P w1 w2) \ (w2 = NEG_INF \ P w1 w2) \ (w1 = POS_INF \ w2 = POS_INF \ P w1 w2) \ (w1 = POS_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 = POS_INF \ P w1 w2) \ (w1 \ POS_INF \ w1 \ NEG_INF \ w2 \ POS_INF \ w2 \ NEG_INF \ P w1 w2) \ P w1 w2" by(auto) lemma word_trichotomy[case_names Less Equal Greater]: fixes w1 w2 :: word shows "(w1 P w1 w2) \ (w1 = w2 \ P w1 w2) \ (w2 P w1 w2) \ P w1 w2" using signed.linorder_cases by auto lemma case_times_inf [case_names PosPos NegPos PosNeg NegNeg PosLo PosHi PosZero NegLo NegHi NegZero LoPos HiPos ZeroPos LoNeg HiNeg ZeroNeg AllFinite]: fixes w1 w2 P assumes pp:"(w1 = POS_INF \ w2 = POS_INF \ P w1 w2)" and np:"(w1 = NEG_INF \ w2 = POS_INF \ P w1 w2)" and pn:"(w1 = POS_INF \ w2 = NEG_INF \ P w1 w2)" and nn:"(w1 = NEG_INF \ w2 = NEG_INF \ P w1 w2)" and pl:"(w1 = POS_INF \ w2 \ NEG_INF \ w2 P w1 w2)" and ph:"(w1 = POS_INF \ w2 \ POS_INF \ 0 P w1 w2)" and pz:"(w1 = POS_INF \ w2 = 0 \ P w1 w2)" and nl:"(w1 = NEG_INF \ w2 \ NEG_INF \ w2 P w1 w2)" and nh:"(w1 = NEG_INF \ w2 \ POS_INF \ 0 P w1 w2)" and nz:"(w1 = NEG_INF \ 0 = w2 \ P w1 w2)" and lp:"(w1 \ NEG_INF \ w1 w2 = POS_INF \ P w1 w2)" and hp:"(w1 \ POS_INF \ 0 w2 = POS_INF \ P w1 w2)" and zp:"(0 = w1 \ w2 = POS_INF \ P w1 w2)" and ln:"(w1 \ NEG_INF \ w1 w2 = NEG_INF \ P w1 w2)" and hn:"(w1 \ POS_INF \ 0 w2 = NEG_INF \ P w1 w2)" and zn:"(0 = w1 \ w2 = NEG_INF \ P w1 w2)" and allFinite:"w1 \ NEG_INF \ w1 \ POS_INF \ w2 \ NEG_INF \ w2 \ POS_INF \ P w1 w2" shows " P w1 w2" proof (cases rule: word_trichotomy[of w1 0]) case Less then have w1l:"w1 Trivial arithmetic lemmas\ lemma max_diff_pos:"0 \ 9223372034707292161 + ((-(2 ^ 31))::real)" by auto lemma max_less:"2 ^ 31 < (9223372039002259455::int)" by auto lemma sints64:"sints 64 = {i. - (2 ^ 63) \ i \ i < 2 ^ 63}" using sints_def[of 64] range_sbintrunc[of 63] by auto lemma sints32:"sints 32 = {i. - (2 ^ 31) \ i \ i < 2 ^ 31}" using sints_def[of 32] range_sbintrunc[of 31] by auto lemma upcast_max:"sint((scast(0x80000001::word))::64 Word.word)=sint((0x80000001::32 Word.word))" by auto lemma upcast_min:"(0xFFFFFFFF80000001::64 Word.word) = ((scast (-0x7FFFFFFF::word))::64 Word.word)" by auto lemma min_extend_neg:"sint ((0xFFFFFFFF80000001)::64 Word.word) < 0" by auto lemma min_extend_val':"sint ((-0x7FFFFFFF)::64 Word.word) = (-0x7FFFFFFF)" by auto lemma min_extend_val:"(-0x7FFFFFFF::64 Word.word) = 0xFFFFFFFF80000001" by auto lemma range2s:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto section \Arithmetic operations\ text\This section defines operations which conservatively compute upper and lower bounds of arithmetic functions given upper and lower bounds on their arguments. Each function comes with a proof that it rounds in the advertised direction. \ subsection \Addition upper bound\ text\Upper bound of w1 + w2\ fun pu :: "word \ word \ word" where "pu w1 w2 = (if w1 = POS_INF then POS_INF else if w2 = POS_INF then POS_INF else if w1 = NEG_INF then (if w2 = NEG_INF then NEG_INF else (let sum::64 Word.word = ((scast w2)::64 Word.word) + ((scast NEG_INF)::64 Word.word) in if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum)) else if w2 = NEG_INF then (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast NEG_INF)::64 Word.word) in if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum) else (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s (sum::64 Word.word) then POS_INF else if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum))" lemma scast_down_range: -fixes w::"'a::len Word.word" -assumes "sint w \ sints (len_of (TYPE('b::len)))" -shows "sint w = sint ((scast w)::'b Word.word)" -unfolding scast_def -by (simp add: assms word_sint.Abs_inverse) + fixes w::"'a::len Word.word" + assumes "sint w \ sints (len_of (TYPE('b::len)))" + shows "sint w = sint ((scast w)::'b Word.word)" + using word_sint.Abs_inverse [OF assms] by simp lemma pu_lemma: fixes w1 w2 fixes r1 r2 :: real assumes up1:"w1 \\<^sub>U (r1::real)" assumes up2:"w2 \\<^sub>U (r2::real)" shows "pu w1 w2 \\<^sub>U (r1 + r2)" proof - have scast_eq1:"sint((scast w1)::64 Word.word) = sint w1" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have scast_eq2:"sint((scast (0x80000001::word))::64 Word.word) = sint ((0x80000001::32 Word.word))" by auto have scast_eq3:"sint((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have w2Geq:"sint ((scast w2)::64 Word.word) \ - (2 ^ 31) " using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 Word.word_size scast_eq1 upcast_max scast_eq3 len32 mem_Collect_eq by auto have "sint ((scast w2)::64 Word.word) \ 2 ^ 31" apply (auto simp add: Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 scast_eq3 len32) using Word.word_sint.Rep[of "(w2)::32 Word.word"] len32[of "TYPE(32)"] sints32 by auto then have w2Less:"sint ((scast w2)::64 Word.word) < 9223372039002259455" by auto have w2Range: "-(2 ^ (size ((scast w2)::64 Word.word) - 1)) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size scast_eq1 upcast_max sints64 max_less) using max_diff_pos max_less w2Less w2Geq by auto have w2case1a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w2Range]) have w1Lower:"sint ((scast w1)::64 Word.word) \ - (2 ^ 31) " using Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 Word.word_size scast_eq1 scast_eq2 scast_eq3 len32 mem_Collect_eq by auto have w1Leq:"sint ((scast w1)::64 Word.word) \ 2 ^ 31" apply (auto simp add: Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 scast_eq1 len32) using Word.word_sint.Rep[of "(w1)::32 Word.word"] len32[of "TYPE(32)"] sints32 by auto then have w1Less:"sint ((scast w1)::64 Word.word) < 9223372039002259455" using max_less by auto have w1MinusBound:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size[of "(scast w1)::64 Word.word"] Word.word_size[of "(scast (-0x7FFFFFFF))::64 Word.word"] scast_eq3 scast_eq2 Word.word_sint.Rep[of "(w1)::32 Word.word"] Word.word_sint.Rep[of "0x80000001::32 Word.word"] Word.word_sint.Rep[of "(scast w1)::64 Word.word"] Word.word_sint.Rep[of "-0x7FFFFFFF::64 Word.word"] sints64 sints32) using w1Lower w1Less by auto have w1case1a:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w1)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by (rule signed_arith_sint(1)[of "(scast w1)::64 Word.word" "(- 0x7FFFFFFF)", OF w1MinusBound]) have w1case1a':"sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) = sint ((scast w1)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word)" using min_extend_val w1case1a by auto have w1Leq':"sint w1 \ 2^31 - 1" using Word.word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have neg64:"(((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) = ((scast w2)::64 Word.word) + (-0x7FFFFFFF)" by auto have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto obtain r'\<^sub>1 and r'\<^sub>2 where geq1:"r'\<^sub>1\r1" and equiv1:"w1 \\<^sub>E r'\<^sub>1" and geq2:"r'\<^sub>2\r2" and equiv2:"w2 \\<^sub>E r'\<^sub>2" using up1 up2 unfolding repU_def by auto show ?thesis proof (cases rule: case_pu_inf[where ?w1.0=w1, where ?w2.0=w2]) case PosAny then show ?thesis apply (auto simp add: repU_def repe.simps) using linear by blast next case AnyPos then show ?thesis apply (auto simp add: repU_def repe.simps) using linear by blast next case NegNeg then show ?thesis using up1 up2 by (auto simp add: repU_def repe.simps) next case NegNum assume neq1:"w2 \ POS_INF" assume eq2:"w1 = NEG_INF" assume neq3:"w2 \ NEG_INF" let ?sum = "(scast w2 + scast NEG_INF)::64 Word.word" have leq1:"r'\<^sub>1 \ (real_of_int (sint NEG_INF))" using equiv1 neq1 eq2 neq3 by (auto simp add: repe.simps) have leq2:"r'\<^sub>2 = (real_of_int (sint w2))" using equiv2 neq1 eq2 neq3 by (auto simp add: repe.simps) have case1:"?sum <=s ((scast NEG_INF)::64 Word.word) \ NEG_INF \\<^sub>U r1 + r2" - using up1 up2 apply (simp add: repU_def repe.simps word_sle_eq) - apply(rule exI[where x= "r1 + r2"]) - apply(auto) - using w2case1a min_extend_neg - apply (auto simp add: neq1 eq2 neq3 repINT repU_def repe.simps repeInt_simps - up2 word_sless_alt) - using repINT repU_def repe.simps repeInt_simps up2 word_sless_alt - add.right_neutral add_mono dual_order.trans of_int_le_0_iff scast_eq3 by fastforce+ + using up1 up2 + apply (simp add: repU_def repe.simps word_sle_eq) + apply (rule exI [where x= "r1 + r2"]) + apply auto + using w2case1a + apply (auto simp add: eq2 scast_eq3) + subgoal for r' + proof - + assume \r1 \ r'\ \r' \ - 2147483647\ \r2 \ signed w2\ \sint w2 \ 0\ + from \sint w2 \ 0\ have \real_of_int (sint w2) \ real_of_int 0\ + by (simp only: of_int_le_iff) + then have \signed w2 \ (0::real)\ + by simp + from \r1 \ r'\ \r' \ - 2147483647\ have \r1 \ - 2147483647\ + by (rule order_trans) + moreover from \r2 \ signed w2\ \signed w2 \ (0::real)\ have \r2 \ 0\ + by (rule order_trans) + ultimately show \r1 + r2 \ - 2147483647\ + by simp + qed + done have case2:"\(?sum <=s scast NEG_INF) \ scast ?sum \\<^sub>U r1 + r2" apply(simp add: repU_def repe.simps word_sle_def up1 up2) apply(rule exI[where x= "r'\<^sub>2 - 0x7FFFFFFF"]) apply(rule conjI) subgoal proof - assume " \ sint (scast w2 + 0xFFFFFFFF80000001) \ - 2147483647" have bound1:"r1 \ - 2147483647" using leq1 geq1 by (auto) have bound2:"r2 \ r'\<^sub>2" using leq2 geq2 by auto show "r1 + r2 \ r'\<^sub>2 - 2147483647" using bound1 bound2 by(linarith) qed apply(rule disjI2) apply(rule disjI2) - apply(auto) + apply(auto simp add: not_le) subgoal proof - - assume a:"\ sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" + assume a:"sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" then have sintw2_bound:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding min_extend_val by auto have case1a:" sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w2Range]) have "- 0x7FFFFFFF < sint w2 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq3 by linarith then have w2bound:"0 < sint w2" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" using case1a scast_eq3 min_extend_val' Word.word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word)) " using scast_down_range[OF rightSize] by auto then have b:"sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word) = sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001)" using min_extend_val by auto have c:"sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) = sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word)" using min_extend_val case1a by auto - show "r'\<^sub>2 - 2147483647 - = (real_of_int (sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word)))" - using a b min_extend_val' scast_eq3 leq2 case1a - by auto + show \r'\<^sub>2 - 2147483647 = signed (SCAST(64 \ 32) (SCAST(32 \ 64) w2 + 0xFFFFFFFF80000001))\ + using a b min_extend_val' scast_eq3 leq2 case1a [symmetric] + apply (simp add: algebra_simps) + apply transfer + apply simp + done qed subgoal proof - have range2a:" - (2 ^ (size ((scast w2)::64 Word.word) - 1)) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((-0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size scast_eq1 upcast_max sints64 sints32 max_less) using max_diff_pos max_less w2Geq w2Less by auto have case2a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF range2a]) have neg64:"(((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) = ((scast w2)::64 Word.word) + (-0x7FFFFFFF)" by auto - assume "\ sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" + assume "sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" then have sintw2_bound:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 by auto have a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((-0x7FFFFFFF)::64 Word.word)" using case2a by auto have b:"sint ((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have d:"sint w2 \ 2^31 - 1" using Word.word_sint.Rep[of "(w2)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have "- 0x7FFFFFFF < sint w2 + (- 0x7FFFFFFF)" using sintw2_bound case2a min_extend_val' scast_eq3 by linarith then have w2bound:"0 < sint w2" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case2a b min_extend_val' using Word.word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto have "sint (scast (((scast w2)::64 Word.word) + (-0x7FFFFFFF))::word) < 2147483647" unfolding downcast a b min_extend_val' using range2s[of "sint w2", OF d] by auto then show "sint (scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001)::word) < 2147483647" by auto qed subgoal proof - - assume notLeq:"\ sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" + assume notLeq:"sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" then have gr:"sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" by auto have case2a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w2Range]) from neg64 have sintw2_bound:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 using notLeq by auto have a:"sint (((scast w2)::64 Word.word) + (-0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((-0x7FFFFFFF)::64 Word.word)" using case2a by auto have c:"sint((-0x7FFFFFFF)::64 Word.word) = -0x7FFFFFFF" by auto have d:"sint w2 \ 2^31 - 1" using Word.word_sint.Rep[of "(w2)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have "- 0x7FFFFFFF < sint w2 + (- 0x7FFFFFFF)" using sintw2_bound case2a c scast_eq3 by linarith then have w2bound:"0 < sint w2" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case2a scast_eq3 using Word.word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto have sintEq:" sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word) = sint (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001) " using downcast by auto - show "-2147483647 - < real_of_int (sint ((scast (((scast w2)::64 Word.word) + 0xFFFFFFFF80000001))::word))" + show "- 2147483647 < sint (SCAST(64 \ 32) (SCAST(32 \ 64) w2 + 0xFFFFFFFF80000001))" unfolding sintEq using gr of_int_less_iff of_int_minus of_int_numeral by linarith qed done have castEquiv:"\(?sum <=s scast NEG_INF) \ (scast ?sum) \\<^sub>U r1 + r2" using up1 up2 case1 case2 by fastforce have letRep:"(let sum = ?sum in if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>U r1 + r2" using case1 case2 by(cases "?sum <=s scast NEG_INF"; auto) show "pu w1 w2 \\<^sub>U r1 + r2" using letRep eq2 neq1 by(auto) next case NumNeg assume neq3:"w1 \ NEG_INF" assume neq1:"w1 \ POS_INF" assume eq2:"w2 = NEG_INF" let ?sum = "(scast w1 + scast NEG_INF)::64 Word.word" have case1:"?sum <=s ((scast NEG_INF)::64 Word.word) \ NEG_INF \\<^sub>U r1 + r2" using up1 up2 apply (simp add: repU_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) apply(auto) using w1case1a min_extend_neg apply (auto simp add: neq1 eq2 neq3 repINT repU_def repe.simps repeInt_simps up2 word_sless_alt) using repINT repU_def repe.simps repeInt_simps up2 word_sless_alt proof - fix r' assume a1:"sint ((scast w1)::64 Word.word) \ 0" - then have h3:"sint w1 \ 0" using scast_eq1 by auto + then have "sint w1 \ 0" using scast_eq1 by auto + then have h3: \signed w1 \ (0::real)\ + by transfer simp assume a2:"r2 \ r'" - assume a3:"r1 \ (real_of_int (sint w1))" + assume a3:"r1 \ signed w1" assume a4:"r' \ (- 2147483647)" from a2 a4 have h1:"r2 \ - 2147483647" by auto from a1 a3 h3 have h2:"r1 \ 0" using dual_order.trans of_int_le_0_iff by blast show "r1 + r2 \ (- 2147483647)" using h1 h2 add.right_neutral add_mono by fastforce qed have leq1:"r'\<^sub>2 \ (real_of_int (sint NEG_INF))" and leq2:"r'\<^sub>1 = (real_of_int (sint w1))" using equiv1 equiv2 neq1 eq2 neq3 unfolding repe.simps by auto have case1a:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF::64 Word.word)) = sint ((scast w1)::64 Word.word) + sint (-0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF w1MinusBound]) have case2:"\(?sum <=s scast NEG_INF) \ scast ?sum \\<^sub>U r1 + r2" apply (simp add: repU_def repe.simps word_sle_def up1 up2) apply(rule exI[where x= "r'\<^sub>1 - 0x7FFFFFFF"]) (*r1 + r2*) apply(rule conjI) subgoal using leq1 leq2 geq1 geq2 by auto apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - have f:"r'\<^sub>1 = (real_of_int (sint w1))" by (simp add: leq1 leq2 ) assume a:"\ sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding min_extend_val by auto have "- 0x7FFFFFFF < sint w1 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq1 by linarith then have w2bound:"0 < sint w1" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding w1case1a using w2bound Word.word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"] scast_eq1) have downcast:"sint ((scast (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto - then have b:"sint ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word) + then have "sint ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word) = sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001)" using min_extend_val by auto - have "(real_of_int (sint ((scast w1)::64 Word.word ) - 2147483647)) - = r'\<^sub>1 - (real_of_int 2147483647)" - by (simp add: scast_eq1 leq2) - then show "r'\<^sub>1 - 2147483647 - = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word)))" - by (metis b w1case1a' min_extend_val' diff_minus_eq_add minus_minus of_int_numeral) - qed + then have \signed (SCAST(64 \ 32) (SCAST(32 \ 64) w1 + 0xFFFFFFFF80000001)) = + (signed (SCAST(32 \ 64) w1 + 0xFFFFFFFF80000001) :: real)\ + by transfer simp + moreover have "r'\<^sub>1 - (real_of_int 2147483647) = + (real_of_int (sint ((scast w1)::64 Word.word ) - 2147483647))" + by (simp add: scast_eq1 leq2) + moreover from w1case1a' + have \signed (SCAST(32 \ 64) w1 + 0xFFFFFFFF80000001) = + signed (SCAST(32 \ 64) w1) + (signed (- 0x7FFFFFFF :: 64 Word.word) :: real)\ + by transfer simp + ultimately show "r'\<^sub>1 - 2147483647 + = (signed ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word))" + by simp + qed subgoal proof - assume "\ sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 by auto have "- 0x7FFFFFFF < sint w1 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq1 by linarith then have w2bound:"0 < sint w1" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a scast_eq1 w1case1a' using Word.word_sint.Rep[of "(w1)::32 Word.word"] w2bound by(auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF))))::word) = sint (((scast w1)::64 Word.word) + ((- 0x7FFFFFFF)::64 Word.word)) " using scast_down_range[OF rightSize] by auto show "sint (scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001)::word) < 2147483647" using downcast min_extend_val' w1case1a' scast_eq1 arith[of "sint w1", OF w1Leq'] by auto qed subgoal proof - assume notLeq:"\ sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) \ - 2147483647" then have gr:"sint (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) > - 2147483647" by auto then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (-0x7FFFFFFF)) > - 2147483647" unfolding neg64 using notLeq by auto have "- 0x7FFFFFFF < sint w1 + (- 0x7FFFFFFF)" using sintw2_bound case1a min_extend_val' scast_eq1 by linarith then have w2bound:"0 < sint w1" using less_add_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + - 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a scast_eq1 w1case1a' using Word.word_sint.Rep[of "(w1)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) show "- 2147483647 - < real_of_int (sint ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word))" - using scast_down_range[OF rightSize] gr of_int_less_iff of_int_minus of_int_numeral - by auto + < sint ((scast (((scast w1)::64 Word.word) + 0xFFFFFFFF80000001))::word)" + using scast_down_range[OF rightSize] gr of_int_less_iff of_int_minus of_int_numeral by simp qed done have letUp:"(let sum=?sum in if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>U r1+r2" using case1 case2 by (auto simp add: Let_def) have puSimp:"pu w1 w2=(let sum = ?sum in if sum <=s scast NEG_INF then NEG_INF else scast sum)" using neq3 neq1 eq2 equiv1 leq2 repeInt_simps by force then show "pu w1 w2 \\<^sub>U r1 + r2" using letUp puSimp by auto next case NumNum assume notinf1:"w1 \ POS_INF" assume notinf2:"w2 \ POS_INF" assume notneginf1:"w1 \ NEG_INF" assume notneginf2:"w2 \ NEG_INF" let ?sum = "((scast w1)::64 Word.word) + ((scast w2):: 64 Word.word)" have inf_case:"scast POS_INF <=s ?sum \ POS_INF \\<^sub>U r1 + r2" using repU_def repePos_simps by (meson dual_order.strict_trans not_less order_refl) have truth:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" using Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast w1)::64 Word.word"] scast_eq1 scast_eq3 Word.word_sint.Rep[of "(w1)::32 Word.word"] Word.word_sint.Rep[of "(w2)::32 Word.word"] Word.word_sint.Rep[of "(scast w1)::64 Word.word"] Word.word_sint.Rep[of "(scast w2)::64 Word.word"] sints64 sints32 by auto have sint_eq:"sint((scast w1 + scast w2)::64 Word.word) = sint w1 + sint w2" using signed_arith_sint(1)[of "(scast w1)::64 Word.word" "(scast w2)::64 Word.word", OF truth] scast_eq1 scast_eq3 by auto have bigOne:"scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word) \ \r'\r1 + r2. r' \ (- 0x7FFFFFFF)" proof - assume "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" - then have sum_leq:"sint w1 + sint w2 \ - 0x7FFFFFFF" - and sum_leq':" (sint w1 + sint w2) \ (- 2147483647)" - using sint_eq unfolding word_sle_eq by auto + then have "sint w1 + sint w2 \ - 0x7FFFFFFF" + using sint_eq unfolding word_sle_eq by auto + then have sum_leq: \real_of_int (sint w1 + sint w2) \ real_of_int (- 0x7FFFFFFF)\ + by (simp only: of_int_le_iff) obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using up1 up2 unfolding repU_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using \scast w1 + scast w2 <=s - 0x7FFFFFFF\ word_sle_def notinf1 notinf2 bound1 bound2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce show "\r'\r1 + r2. r' \ (- 0x7FFFFFFF)" apply(rule exI[where x = "r'\<^sub>1 + r'\<^sub>2"]) - using bound1 bound2 add_mono something sum_leq' order.trans - by auto + using bound1 bound2 add_mono something sum_leq + apply (auto intro: order_trans [of _ \signed_real_of_word w1 + + signed_real_of_word w2\]) + done qed have anImp:"\r'. (r'\r1 + r2 \ r' \ (- 2147483647)) \ (\r. - (2 ^ 31 - 1) = - (2 ^ 31 - 1) \ r' = r \ r \ (real_of_int (sint ((- (2 ^ 31 - 1))::32 Word.word))))" by auto have anEq:"((scast ((- (2 ^ 31 - 1))::32 Word.word))::64 Word.word) = (- 0x7FFFFFFF)" by auto have bigTwo: "\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \(?sum <=s ((scast NEG_INF)::64 Word.word)) \ \r'\r1 + r2. r' = (real_of_int (sint (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word))) \ (r' < 0x7FFFFFFF \ (-0x7FFFFFFF) < r')" proof - assume "\(((scast POS_INF)::64 Word.word) <=s ?sum)" - then have sum_leq:"sint w1 + sint w2 < 0x7FFFFFFF" - unfolding word_sle_eq POS_INF_def using sint_eq by auto - then have sum_leq':" (sint w1 + sint w2) < (2147483647)" - by auto - assume "\(?sum <=s ((scast NEG_INF)::64 Word.word))" - then have sum_geq:"(- 0x7FFFFFFF) < sint w1 + sint w2" - unfolding word_sle_eq NEG_INF_def using sint_eq by auto - then have sum_geq':" (- 2147483647) < (sint w1 + sint w2)" - by auto + and "\(?sum <=s ((scast NEG_INF)::64 Word.word))" + then have interval_int: "sint w1 + sint w2 < 0x7FFFFFFF" + "(- 0x7FFFFFFF) < sint w1 + sint w2" + unfolding word_sle_eq POS_INF_def NEG_INF_def using sint_eq by auto + then have interval: \real_of_int (sint w1 + sint w2) < real_of_int (0x7FFFFFFF)\ + \real_of_int (- 0x7FFFFFFF) < real_of_int (sint w1 + sint w2)\ + by (simp_all only: of_int_less_iff) obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using up1 up2 unfolding repU_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using word_sle_def notinf1 notinf2 bound1 bound2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce have "(w1 \\<^sub>E r'\<^sub>1)" using bound1 by auto then have r1w1:"r'\<^sub>1 = (real_of_int (sint w1))" and w1U:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and w1L:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "(w2 \\<^sub>E r'\<^sub>2)" using bound2 by auto then have r2w1:"r'\<^sub>2 = (real_of_int (sint w2))" and w2U:" (real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and w2L:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "sint (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)) = sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)" apply(rule scast_down_range) - unfolding sint_eq using sints32 sum_geq sum_leq by auto + unfolding sint_eq using sints32 interval_int by auto then have cast_eq:"(sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)) = sint w1 + sint w2" - using scast_down_range sints32 sum_geq sum_leq sint_eq by auto + using scast_down_range sints32 interval_int sint_eq by auto from something and cast_eq have r12_sint_scast:"r'\<^sub>1 + r'\<^sub>2 = (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word))" using r1w1 w1U w1L r2w1 w2U w2L by (simp) show ?thesis - using bound1 bound2 add_mono r12_sint_scast cast_eq sum_leq sum_leq' sum_geq' + using bound1 bound2 add_mono r12_sint_scast cast_eq interval \r'\<^sub>1 + r'\<^sub>2 = (real_of_int (sint (scast (scast w1 + scast w2))))\ - by auto + by simp qed have neg_inf_case:"?sum <=s ((scast ((NEG_INF)::word))::64 Word.word) \ NEG_INF \\<^sub>U r1 + r2" proof (unfold repU_def NEG_INF_def repe.simps) assume "scast w1 + scast w2 <=s ((scast ((- (2 ^ 31 - 1))::word))::64 Word.word)" then have "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" by (metis anEq) then obtain r' where geq:"(r' \ r1 + r2)" and leq:"(r' \ (- 0x7FFFFFFF))" using bigOne by auto show "(\r'\plus r1 r2. (\r. uminus (minus(2 ^ 31) 1) = POS_INF \ r' = r \ (real_of_int (sint POS_INF)) \ r) \ (\r. uminus (minus(2 ^ 31) 1) = uminus (minus(2 ^ 31) 1) \ r' = r \ r \ real_of_int (sint ((uminus (minus(2 ^ 31) 1))::word))) \ (\w. uminus (minus(2 ^ 31) 1) = w \ r' = real_of_int (sint w) \ (real_of_int (sint w)) < (real_of_int (sint POS_INF)) \ less (real_of_int (sint (uminus (minus(2 ^ 31) 1)))) (real_of_int (sint w))))" using leq anImp geq by auto qed have int_case:"\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \ (?sum <=s ((scast NEG_INF)::64 Word.word)) \ ((scast ?sum)::word) \\<^sub>U r1 + r2" proof - assume bound1:"\ ((scast POS_INF)::64 Word.word) <=s scast w1 + scast w2" assume bound2:"\ scast w1 + scast w2 <=s ((scast NEG_INF)::64 Word.word)" obtain r'::real where rDef:"r' = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)))" and r12:"r'\r1 + r2" and boundU:"r' < 0x7FFFFFFF" and boundL:"(-0x7FFFFFFF) < r'" using bigTwo[OF bound1 bound2] by auto obtain w::word where wdef:"w = (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word)" by auto then have wr:"r' = (real_of_int (sint w))" using r12 bound1 bound2 rDef by blast show ?thesis unfolding repU_def repe.simps using r12 wdef rDef boundU boundL wr by auto qed have almost:"(let sum::64 Word.word = scast w1 + scast w2 in if scast POS_INF <=s sum then POS_INF else if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>U r1 + r2" apply(cases "((scast POS_INF)::64 Word.word) <=s ((?sum)::64 Word.word)") subgoal using inf_case Let_def int_case neg_inf_case by auto apply(cases "?sum <=s scast NEG_INF") subgoal using inf_case Let_def int_case neg_inf_case proof - assume "\ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2" then have "\ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF \ \ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2 \ \ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF \ ((let w = scast w1 + scast w2 in if scast POS_INF <=s (w::64 Word.word) then POS_INF else if w <=s scast NEG_INF then NEG_INF else scast w) \\<^sub>U r1 + r2)" using neg_inf_case by presburger then show ?thesis using int_case by force qed subgoal using inf_case Let_def int_case neg_inf_case proof - assume a1: "\ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2" assume "\ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF" have "\ (scast w1::64 Word.word) + scast w2 <=s scast NEG_INF \ \ (scast POS_INF::64 Word.word) <=s scast w1 + scast w2 \ ((let w = scast w1 + scast w2 in if scast POS_INF <=s (w::64 Word.word) then POS_INF else if w <=s scast NEG_INF then NEG_INF else scast w) \\<^sub>U r1 + r2)" using a1 neg_inf_case by presburger then show ?thesis using int_case by force qed done then show ?thesis using notinf1 notinf2 notneginf1 notneginf2 by auto qed qed text\Lower bound of w1 + w2\ fun pl :: "word \ word \ word" where "pl w1 w2 = (if w1 = NEG_INF then NEG_INF else if w2 = NEG_INF then NEG_INF else if w1 = POS_INF then (if w2 = POS_INF then POS_INF else (let sum::64 Word.word = ((scast w2)::64 Word.word) + ((scast POS_INF)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s(sum::64 Word.word) then POS_INF else scast sum)) else if w2 = POS_INF then (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast POS_INF)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s(sum::64 Word.word) then POS_INF else scast sum) else (let sum::64 Word.word = ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word) in if ((scast POS_INF)::64 Word.word) <=s (sum::64 Word.word) then POS_INF else if (sum::64 Word.word) <=s ((scast NEG_INF)::64 Word.word) then NEG_INF else scast sum))" subsection \Addition lower bound\ text\Correctness of lower bound of w1 + w2\ lemma pl_lemma: assumes lo1:"w1 \\<^sub>L (r1::real)" assumes lo2:"w2 \\<^sub>L (r2::real)" shows "pl w1 w2 \\<^sub>L (r1 + r2)" proof - have scast_eq1:"sint((scast w1)::64 Word.word) = sint w1" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have scast_eq2:"sint((scast (0x80000001::word))::64 Word.word)=sint((0x80000001::32 Word.word))" by auto have scast_eq3:"sint((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have sints64:"sints 64 = {i. - (2 ^ 63) \ i \ i < 2 ^ 63}" using sints_def[of 64] range_sbintrunc[of 63] by auto have sints32:"sints 32 = {i. - (2 ^ 31) \ i \ i < 2 ^ 31}" using sints_def[of 32] range_sbintrunc[of 31] by auto have thing1:"0 \ 9223372034707292161 + ((-(2 ^ 31))::real)" by auto have "sint (( w2)) \ (-(2 ^ 31))" using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 mem_Collect_eq Word.word_size[of "(scast w2)::64 Word.word"] scast_eq1 scast_eq2 scast_eq3 len32 by auto then have thing4:"sint ((scast w2)::64 Word.word) \ (-(2 ^ 31))" using scast_down_range sint_ge sints_num using scast_eq3 by linarith have aLeq2:"(-(2 ^ 31)::int) \ -9223372039002259455" by auto then have thing2:" (0::int) \ 9223372039002259455 + sint ((scast w2)::64 Word.word)" using thing4 aLeq2 by (metis ab_group_add_class.ab_left_minus add.commute add_mono neg_le_iff_le) have aLeq:"2 ^ 31 \ (9223372039002259455::int)" by auto have bLeq:"sint ((scast w2)::64 Word.word) \ 2 ^ 31" apply ( auto simp add: Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 scast_eq3 len32) using Word.word_sint.Rep[of "(w2)::32 Word.word"] len32[of "TYPE(32)"] sints32 by auto have thing3:" sint ((scast w2)::64 Word.word) \ 9223372034707292160 " using aLeq bLeq by auto have truth:" - (2 ^ (size ((scast w2)::64 Word.word) - 1)) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" by(auto simp add: Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast (0x7FFFFFFF))::64 Word.word"] scast_eq1 scast_eq2 sints64 sints32 thing2 thing1 thing3) have case1a:" sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have case1b:"sint ((0xFFFFFFFF80000001)::64 Word.word) < 0" by auto have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto have neg64:"(((scast w2)::64 Word.word) + 0x7FFFFFFF) = ((scast w2)::64 Word.word) + (0x7FFFFFFF)" by auto obtain r'\<^sub>1 and r'\<^sub>2 where geq1:"r'\<^sub>1\r1" and equiv1:"w1 \\<^sub>E r'\<^sub>1" and geq2:"r'\<^sub>2\r2" and equiv2:"w2 \\<^sub>E r'\<^sub>2" using lo1 lo2 unfolding repL_def by auto show ?thesis proof (cases rule: case_pl_inf[where ?w1.0=w1, where ?w2.0=w2]) case NegAny then show ?thesis apply (auto simp add: repL_def repe.simps) using lo1 lo2 linear by auto next case AnyNeg then show ?thesis apply (auto simp add: repL_def repe.simps) using linear by auto next case PosPos then show ?thesis using lo1 lo2 by (auto simp add: repL_def repe.simps) next case PosNum assume neq1:"w2 \ POS_INF" assume eq2:"w1 = POS_INF" assume neq3:"w2 \ NEG_INF" let ?sum = "(scast w2 + scast POS_INF)::64 Word.word" have case1:"(((scast POS_INF)::64 Word.word) <=s ?sum) \ POS_INF \\<^sub>L r1 + r2" using lo1 lo2 apply (simp add: repL_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) using case1a case1b apply (auto simp add: neq1 eq2 neq3 repINT repL_def repe.simps repeInt_simps lo2 word_sless_alt) proof - fix r' assume a1:"0 \ sint (((scast w2)::64 Word.word))" from a1 have h3:"2147483647 \ sint w2 + 0x7FFFFFFF " using scast_eq3 by auto assume a2:"r' \ r1" - assume a3:"(real_of_int (sint w2)) \ r2" + assume a3:"signed w2 \ r2" assume a4:"(2147483647) \ r'" from a2 a4 have h1:"2147483647 \ r1" by auto from a1 a3 h3 have h2:"0 \ r2" - using dual_order.trans of_int_le_0_iff le_add_same_cancel2 by fastforce + using of_int_le_0_iff le_add_same_cancel2 + apply simp + apply transfer + apply simp + apply (metis (full_types) of_int_0 of_int_le_iff order_trans signed_take_bit_nonnegative_iff) + done show "(2147483647) \ r1 + r2 " using h1 h2 h3 add.right_neutral add_mono by fastforce qed have leq1:"r'\<^sub>1 \ (real_of_int (sint POS_INF))" using equiv1 neq1 eq2 neq3 unfolding repe.simps by auto have leq2:"r'\<^sub>2 = (real_of_int (sint w2))" using equiv2 neq1 eq2 neq3 unfolding repe.simps by auto have case2:"\(scast POS_INF <=s ?sum) \ scast ?sum \\<^sub>L r1 + r2" apply (simp add: repL_def repe.simps word_sle_def lo1 lo2) apply(rule exI[where x= "r'\<^sub>2 + 0x7FFFFFFF"]) (*r1 + r2*) apply(rule conjI) subgoal proof - assume "\ 2147483647 \ sint (scast w2 + 0x7FFFFFFF)" have bound1:"2147483647 \ r1" using leq1 geq1 by (auto) have bound2:"r'\<^sub>2 \ r2 " using leq2 geq2 by auto show "r'\<^sub>2 + 2147483647 \ r1 + r2" using bound1 bound2 by linarith qed apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - assume a:"\ 2147483647 \ sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"2147483647 > sint (((scast w2)::64 Word.word) + (0x7FFFFFFF))" by auto have case1a:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have a1:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((0x7FFFFFFF)::64 Word.word)" using case1a by auto have c1:"sint((0x7FFFFFFF)::64 Word.word) = 0x7FFFFFFF" by auto have "sint w2 + ( 0x7FFFFFFF) < 0x7FFFFFFF" using sintw2_bound case1a c1 scast_eq3 by linarith then have w2bound:"sint w2 < 0" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a scast_eq3 c1 using Word.word_sint.Rep[of "(w2)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have b:"sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word) = sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" by auto have c:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) = sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word)" using case1a by auto have d:"sint ((0x7FFFFFFF)::64 Word.word) = (0x7FFFFFFF)" by auto have f:"r'\<^sub>2 = (real_of_int (sint w2))" by (simp add: leq2) show "r'\<^sub>2 + 2147483647 - = (real_of_int (sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word)))" - using a b c d scast_eq3 f leq2 of_int_numeral by fastforce + = (signed ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word))" + using a b c d scast_eq3 f leq2 of_int_numeral + by auto qed subgoal proof - have truth2a:"-(2^(size ((scast w2)::64 Word.word)-1)) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ sint ((scast w2)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w2)::64 Word.word) - 1) - 1" apply(auto simp add: Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast (0x7FFFFFFF))::64 Word.word"] scast_eq1 scast_eq2 sints64 sints32 thing2) using thing1 thing2 thing3 by auto have case2a:" sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth2a]) have min_cast:"(0x7FFFFFFF::64 Word.word) =((scast (0x7FFFFFFF::word))::64 Word.word)" by auto assume "\ 2147483647 \ sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"2147483647 > sint (((scast w2)::64 Word.word) + (0x7FFFFFFF))" using neg64 by auto have a:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((0x7FFFFFFF)::64 Word.word)" using case2a by auto have c:"sint((0x7FFFFFFF)::64 Word.word) = 0x7FFFFFFF" by auto have " 0x7FFFFFFF > sint w2 + ( 0x7FFFFFFF)" using sintw2_bound case2a c scast_eq3 by linarith then have w2bound:" sint w2 < 0" by simp have rightSize:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case2a scast_eq3 c apply (auto simp add: sints32 len32[of "TYPE(32)"]) using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 len32[of "TYPE(32)"] w2bound by auto have downcast:"sint ((scast (((scast w2)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then show "sint (scast (((scast w2)::64 Word.word) + 0x7FFFFFFF)::word) < 2147483647" unfolding downcast a scast_eq3 c using w2bound by auto qed subgoal proof - assume notLeq:"\ 2147483647 \ sint (((scast w2)::64 Word.word) + 0x7FFFFFFF)" then have gr:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) < 2147483647" by auto have case2a:" sint (((scast w2)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w2)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have min_cast:"(0x7FFFFFFF::64 Word.word) =((scast (0x7FFFFFFF::word))::64 Word.word)" by auto have neg64:"(((scast w2)::64 Word.word) + 0x7FFFFFFF) = ((scast w2)::64 Word.word) + (0x7FFFFFFF)" by auto then have sintw2_bound:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) < 2147483647" using neg64 using notLeq by auto have a:"sint (((scast w2)::64 Word.word) + (0x7FFFFFFF)) = sint((scast w2)::64 Word.word) + sint((0x7FFFFFFF)::64 Word.word)" using case2a by auto have c:"sint((0x7FFFFFFF)::64 Word.word) = 0x7FFFFFFF" by auto have "- 2147483647 \ w2" using neq3 unfolding NEG_INF_def by auto then have "sint((- 2147483647)::word) \ sint w2" using word_sint.Rep_inject by blast then have n1:"- 2147483647 \ sint w2" by auto have "- 2147483648 \ w2" apply(rule repe.cases[OF equiv2]) by auto then have "sint(- 2147483648::word) \ sint w2" using word_sint.Rep_inject by blast then have n2:"- 2147483648 \ sint w2" by auto then have d:"sint w2 > - 2147483647" using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 len32[of "TYPE(32)"] neq3 n1 n2 by auto have w2bound:"- 2147483647 < sint w2 + 0x7FFFFFFF" using sintw2_bound case2a c scast_eq3 d by linarith have rightSize:"sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" using sints32 len32[of "TYPE(32)"] w2bound Word.word_sint.Rep[of "(w2)::32 Word.word"] c case2a scast_eq3 sintw2_bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have downcast:"sint ((scast (((scast w2)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w2)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto have sintEq:" sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word) = sint (((scast w2)::64 Word.word) + 0x7FFFFFFF) " using downcast by auto - show "- 2147483647 - < real_of_int (sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word))" + show "- 2147483647 + < sint ((scast (((scast w2)::64 Word.word) + 0x7FFFFFFF))::word)" unfolding sintEq - using gr of_int_less_iff of_int_minus of_int_numeral c case2a scast_eq3 w2bound by linarith + using gr of_int_less_iff of_int_minus of_int_numeral c case2a scast_eq3 w2bound + by simp qed done have "(let sum = ?sum in if scast POS_INF <=s sum then POS_INF else scast sum) \\<^sub>L r1 + r2" using case1 case2 by (auto simp add: Let_def) then show ?thesis using lo1 lo2 neq1 eq2 neq3 by (auto) next case NumPos assume neq3:"w1 \ NEG_INF" assume neq1:"w1 \ POS_INF" assume eq2:"w2 = POS_INF" let ?sum = "(scast w1 + scast POS_INF)::64 Word.word" have thing1:"0 \ 9223372034707292161 + ((-(2 ^ 31))::real)" by auto have "sint (( w1)) \ (-(2 ^ 31))" using Word.word_sint.Rep[of "(w1)::32 Word.word"] scast_eq1 scast_eq2 scast_eq3 Word.word_size[of "(scast w1)::64 Word.word"] sints32 len32 mem_Collect_eq by auto then have thing4:"sint ((scast w1)::64 Word.word) \ (-(2 ^ 31))" using scast_down_range sint_ge sints_num scast_eq3 scast_eq1 by linarith have aLeq2:"(-(2 ^ 31)::int) \ -9223372039002259455" by auto then have thing2:" (0::int) \ 9223372039002259455 + sint ((scast w1)::64 Word.word)" using thing4 aLeq2 by (metis ab_group_add_class.ab_left_minus add.commute add_mono neg_le_iff_le) have aLeq:"2 ^ 31 \ (9223372039002259455::int)" by auto have bLeq:"sint ((scast w1)::64 Word.word) \ 2 ^ 31" apply (auto simp add: Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 scast_eq1 len32) using Word.word_sint.Rep[of "(w1)::32 Word.word"] len32[of "TYPE(32)"] sints32 by clarsimp have thing3:" sint ((scast w1)::64 Word.word) \ 9223372034707292160 " using aLeq bLeq by auto have truth:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" by(auto simp add: Word.word_size[of "(scast w1)::64 Word.word"] Word.word_size[of "(scast (0x7FFFFFFF))::64 Word.word"] scast_eq3 scast_eq2 sints64 sints32 thing2 thing1 thing3) have case1a:"sint (((scast w1)::64 Word.word) + (0x7FFFFFFF::64 Word.word)) = sint ((scast w1)::64 Word.word) + sint (0x7FFFFFFF::64 Word.word)" by(rule signed_arith_sint(1)[OF truth]) have case1b:"sint ((0xFFFFFFFF80000001)::64 Word.word) < 0" by auto have g:"(0x7FFFFFFF::64 Word.word) = 0x7FFFFFFF" by auto have c:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) = sint ((scast w1)::64 Word.word) + sint ((0x7FFFFFFF)::64 Word.word)" using g case1a by blast have d:"sint ((0x7FFFFFFF)::64 Word.word) = (0x7FFFFFFF)" by auto have e:"sint ((scast w1)::64 Word.word) = sint w1" using scast_eq1 by blast have d2:"sint w1 \ 2^31 - 1" using Word.word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"]) have case1:"scast POS_INF <=s ?sum \ POS_INF \\<^sub>L r1 + r2" using lo1 lo2 apply (simp add: repL_def repe.simps word_sle_def) apply(rule exI[where x= "r1 + r2"]) apply(auto) using case1a case1b apply (auto simp add: neq1 eq2 neq3 repINT repL_def repe.simps repeInt_simps lo2 word_sless_alt) (* close 4 goals *) proof - fix r' have h3:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) = sint (((scast w1)::64 Word.word)) + sint(0x7FFFFFFF::64 Word.word)" using case1a by auto have h4:"sint(0x7FFFFFFF::64 Word.word) = 2147483647" by auto assume a1:"0 \ sint ((scast w1)::64 Word.word)" then have h3:"sint w1 \ 0" using scast_eq1 h3 h4 by auto assume a2:"r' \ r2" - assume a3:"(real_of_int (sint w1)) \ r1" + assume a3:"signed w1 \ r1" assume a4:"(2147483647) \ r'" from a2 a4 have h1:"r2 \ 2147483647" by auto - from a1 a3 h3 have h2:"r1 \ 0" - using dual_order.trans of_int_0_le_iff by fastforce + from a3 h3 have h2:"r1 \ 0" + by (auto intro: order_trans [of _ \signed_real_of_word w1\]) show " 2147483647 \ r1 + r2" using h1 h2 add.right_neutral add_mono by fastforce qed have leq1:"r'\<^sub>2 \ (real_of_int (sint POS_INF))" and leq2:"r'\<^sub>1 = (real_of_int (sint w1))" using equiv1 equiv2 neq1 eq2 neq3 unfolding repe.simps by auto have neg64:"(((scast w1)::64 Word.word) + 0xFFFFFFFF80000001) = ((scast w1)::64 Word.word) + (-0x7FFFFFFF)" by auto have case2:"\(scast POS_INF <=s ?sum) \ scast ?sum \\<^sub>L r1 + r2" apply (simp add: repL_def repe.simps word_sle_def lo1 lo2) apply(rule exI[where x= "r'\<^sub>1 + 0x7FFFFFFF"]) apply(rule conjI) subgoal proof - assume "\ 2147483647 \ sint (scast w1 + 0x7FFFFFFF)" have bound1:"r2 \ 2147483647" using leq1 geq2 by (auto) have bound2:"r1 \ r'\<^sub>1" using leq2 geq1 by auto show "r'\<^sub>1 + 2147483647 \ r1 + r2" using bound1 bound2 by linarith qed apply(rule disjI2) apply(rule disjI2) apply(auto) subgoal proof - have f:"r'\<^sub>1 = (real_of_int (sint w1))" by (simp add: leq1 leq2 ) assume a:"\ 2147483647 \ sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"2147483647 > sint (((scast w1)::64 Word.word) + (0x7FFFFFFF))" by auto have "0x7FFFFFFF > sint w1 + (0x7FFFFFFF)" using sintw2_bound case1a d scast_eq1 by linarith then have w2bound:"0 > sint w1" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a e using w2bound Word.word_sint.Rep[of "(w1)::32 Word.word"] by (auto simp add: sints32 len32[of "TYPE(32)"] ) have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto have downcast:"sint ((scast (((scast w1)::64 Word.word) + (( 0x7FFFFFFF))))::word) = sint (((scast w1)::64 Word.word) + (( 0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have b:"sint((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word) = sint(((scast w1)::64 Word.word) + 0x7FFFFFFF)" using g by auto show "r'\<^sub>1 + 2147483647 - = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word)))" + = ((signed_real_of_word ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word)))" using a b c d e f proof - have "(real_of_int (sint ((scast w1)::64 Word.word ) + 2147483647)) = r'\<^sub>1 + (real_of_int 2147483647)" using e leq2 by auto - then show ?thesis - using b c d of_int_numeral by metis + from this [symmetric] show ?thesis + using b c d of_int_numeral + by simp qed qed subgoal proof - assume "\ 2147483647 \ sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (0x7FFFFFFF)) < 2147483647" unfolding neg64 by auto have "0x7FFFFFFF > sint w1 + (0x7FFFFFFF)" using sintw2_bound case1a d scast_eq1 by linarith then have w2bound:"0 > sint w1" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a e c using Word.word_sint.Rep[of "(w1)::32 Word.word"] w2bound by (auto simp add: sints32 len32[of "TYPE(32)"]) have arith:"\x::int. x \ 2 ^ 31 - 1 \ x + (- 2147483647) < 2147483647" by auto have downcast:"sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word) = sint (((scast w1)::64 Word.word) + ((0x7FFFFFFF)::64 Word.word))" using scast_down_range[OF rightSize] by auto show "sint (scast (((scast w1)::64 Word.word) + 0x7FFFFFFF)::word) < 2147483647" using downcast d e c arith[of "sint w1", OF d2] sintw2_bound by linarith qed subgoal proof - assume notLeq:"\ 2147483647 \ sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" then have gr:"2147483647 > sint (((scast w1)::64 Word.word) + 0x7FFFFFFF)" by auto then have sintw2_bound:"sint (((scast w1)::64 Word.word) + (0x7FFFFFFF)) < 2147483647" unfolding neg64 using notLeq by auto have "0x7FFFFFFF > sint w1 + ( 0x7FFFFFFF)" using sintw2_bound case1a d scast_eq1 by linarith then have useful:"0 > sint w1" using add_less_same_cancel2 by blast have rightSize:"sint (((scast w1)::64 Word.word) + 0x7FFFFFFF) \ sints (len_of TYPE(32))" unfolding case1a e using Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 len32[of "TYPE(32)"] useful by auto have "- 2147483647 \ w1" using neq3 unfolding NEG_INF_def by auto then have "sint((- 2147483647)::word) \ sint w1" using word_sint.Rep_inject by blast then have n1:"- 2147483647 \ sint w1" by auto have "- 2147483648 \ w1" apply(rule repe.cases[OF equiv1]) using int_not_undef[of w1] by auto then have "sint(- 2147483648::word) \ sint w1" using word_sint.Rep_inject by blast then have n2:"- 2147483648 \ sint w1" by auto then have d:"sint w1 > - 2147483647" using Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 len32[of "TYPE(32)"] n1 n2 neq3 by (simp) have d2:"sint (0x7FFFFFFF::64 Word.word) > 0" by auto from d d2 have d3:"- 2147483647 < sint w1 + sint (0x7FFFFFFF::64 Word.word)" by auto have d4:"sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word) = sint w1 + sint (0x7FFFFFFF::64 Word.word)" using case1a rightSize scast_down_range scast_eq1 by fastforce - then show "-2147483647 - < real_of_int (sint ((scast (((scast w1)::64 Word.word) + 0x7FFFFFFF))::word))" + then show "- 2147483647 < sint (SCAST(64 \ 32) (SCAST(32 \ 64) w1 + 0x7FFFFFFF))" using d3 d4 by auto qed done have "(let sum = ?sum in if scast POS_INF <=s sum then POS_INF else scast sum) \\<^sub>L r1 + r2" using case1 case2 by (auto simp add: Let_def) then show ?thesis using neq1 eq2 neq3 by (auto) next case NumNum assume notinf1:"w1 \ POS_INF" assume notinf2:"w2 \ POS_INF" assume notneginf1:"w1 \ NEG_INF" assume notneginf2:"w2 \ NEG_INF" let ?sum = "((scast w1)::64 Word.word) + ((scast w2):: 64 Word.word)" have truth:" - (2 ^ (size ((scast w1)::64 Word.word) - 1)) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ sint ((scast w1)::64 Word.word) + sint ((scast w2)::64 Word.word) \ 2 ^ (size ((scast w1)::64 Word.word) - 1) - 1" using Word.word_size[of "(scast w2)::64 Word.word"] Word.word_size[of "(scast w1)::64 Word.word"] scast_eq1 scast_eq3 sints64 sints32 Word.word_sint.Rep[of "(w1)::32 Word.word"] Word.word_sint.Rep[of "(w2)::32 Word.word"] by auto have sint_eq:"sint((scast w1 + scast w2)::64 Word.word) = sint w1 + sint w2" using signed_arith_sint(1)[of "(scast w1)::64 Word.word" "(scast w2)::64 Word.word", OF truth] scast_eq1 scast_eq3 by auto have bigOne:"scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word) \ \r'\r1 + r2. r' \ -0x7FFFFFFF" proof - assume "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" then have sum_leq:"sint w1 + sint w2 \ - 0x7FFFFFFF" and sum_leq':" (sint w1 + sint w2) \ (- 2147483647)" using sint_eq unfolding word_sle_eq by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using lo1 lo2 unfolding repL_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using bound1 bound2 \scast w1 + scast w2 <=s -0x7FFFFFFF\ word_sle_def notinf1 notinf2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce show "\r'\r1 + r2. r' \ (-0x7FFFFFFF)" - apply(rule exI[where x = "r'\<^sub>1 + r'\<^sub>2"]) - using bound1 bound2 add_mono something sum_leq' order.trans by auto + apply (rule exI[where x = "r'\<^sub>1 + r'\<^sub>2"]) + using bound1 bound2 add_mono something sum_leq' + apply (auto intro: order_trans [of _ \signed_real_of_word w1 + signed_real_of_word w2\]) + done qed have anImp:"\r'. (r'\r1 + r2 \ r' \ (- 2147483647)) \ (\r. - (2 ^ 31 - 1) = - (2 ^ 31 - 1) \ r' = r \ r \ (real_of_int (sint ((- (2 ^ 31 - 1))::32 Word.word))))" by auto have anEq:"((scast ((- (2 ^ 31 - 1))::32 Word.word))::64 Word.word) = (- 0x7FFFFFFF)" by auto have bigTwo: "\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \(?sum <=s ((scast NEG_INF)::64 Word.word)) \ \r'\r1 + r2. r' = (real_of_int (sint (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word))) \ (r' < 0x7FFFFFFF \ (-0x7FFFFFFF) < r')" proof - assume "\(((scast POS_INF)::64 Word.word) <=s ?sum)" then have sum_leq:"sint w1 + sint w2 < 0x7FFFFFFF" unfolding word_sle_eq using sint_eq by auto then have sum_leq':" (sint w1 + sint w2) < (2147483647)" by auto assume "\(?sum <=s ((scast NEG_INF)::64 Word.word))" then have sum_geq:"(- 0x7FFFFFFF) < sint w1 + sint w2" unfolding word_sle_eq using sint_eq by auto then have sum_geq':" (- 2147483647) < (sint w1 + sint w2)" by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using lo1 lo2 unfolding repL_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using word_sle_def notinf1 notinf2 bound1 bound2 unfolding repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono by fastforce have "(w1 \\<^sub>E r'\<^sub>1)" using bound1 by auto then have r1w1:"r'\<^sub>1 = (real_of_int (sint w1))" and w1U:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and w1L:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "(w2 \\<^sub>E r'\<^sub>2)" using bound2 by auto then have r2w1:"r'\<^sub>2 = (real_of_int (sint w2))" and w2U:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and w2L:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" unfolding repe.simps using notinf1 notinf2 notneginf1 notneginf2 by (auto) have "sint (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)) = sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)" apply(rule scast_down_range) unfolding sint_eq using sints32 sum_geq sum_leq by auto then have cast_eq:"(sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)) = sint w1 + sint w2" using scast_down_range sints32 sum_geq sum_leq sint_eq by auto from something and cast_eq have r12_sint_scast:"r'\<^sub>1 + r'\<^sub>2 = (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word))" using r1w1 w1U w1L r2w1 w2U w2L by (simp) have leq_ref:"\x y ::real. x = y ==> x \ y" by auto show ?thesis apply(rule exI[where x="r'\<^sub>1 + r'\<^sub>2"]) apply(rule conjI) subgoal using r12_sint_scast cast_eq leq_ref r2w1 r1w1 add_mono[of r'\<^sub>1 r1 r'\<^sub>2 r2] bound1 bound2 by auto using bound1 bound2 add_mono r12_sint_scast cast_eq sum_leq sum_leq' sum_geq' sum_geq \r'\<^sub>1 + r'\<^sub>2 = (real_of_int (sint (scast (scast w1 + scast w2))))\ by auto qed have neg_inf_case:"?sum <=s ((scast ((NEG_INF)::word))::64 Word.word) \ NEG_INF \\<^sub>L r1 + r2" proof (unfold repL_def NEG_INF_def repe.simps) assume "scast w1 + scast w2 <=s ((scast ((- (2 ^ 31 - 1))::word))::64 Word.word)" then have "scast w1 + scast w2 <=s ((- 0x7FFFFFFF)::64 Word.word)" by (metis anEq) then obtain r' where geq:"(r' \ r1 + r2)" and leq:"(r' \ (-0x7FFFFFFF))" using bigOne by auto show "(\r'\plus r1 r2. (\r. uminus (minus(2 ^ 31) 1) = POS_INF \ r' = r \ (real_of_int (sint POS_INF)) \ r) \ (\r. uminus (minus(2 ^ 31) 1) = uminus (minus(2 ^ 31) 1) \ r' = r \ r \ (real_of_int (sint ((uminus (minus(2 ^ 31) 1))::word)))) \ (\w. uminus (minus(2 ^ 31) 1) = w \ r' = (real_of_int (sint w)) \ (real_of_int (sint w)) < (real_of_int (sint POS_INF)) \ less ( (real_of_int (sint (uminus (minus(2 ^ 31) 1))))) ((real_of_int (sint w)))))" using leq geq by auto qed have bigThree:"0x7FFFFFFF <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word) \ \r'\r1 + r2. 2147483647 \ r'" proof - assume "0x7FFFFFFF <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" then have sum_leq:"0x7FFFFFFF \ sint w1 + sint w2 " and sum_leq':" 2147483647 \ (sint w1 + sint w2)" using sint_eq unfolding word_sle_eq by auto obtain r'\<^sub>1 r'\<^sub>2 ::real where bound1:"r'\<^sub>1 \ r1 \ (w1 \\<^sub>E r'\<^sub>1)" and bound2:"r'\<^sub>2 \ r2 \ (w2 \\<^sub>E r'\<^sub>2)" using lo1 lo2 unfolding repL_def by auto have somethingA:"r'\<^sub>1 \ sint w1" and somethingB:"r'\<^sub>2 \ sint w2" using \ 0x7FFFFFFF <=s scast w1 + scast w2 \ word_sle_def notinf1 notinf2 bound1 bound2 repe.simps by auto have something:"r'\<^sub>1 + r'\<^sub>2 \ sint w1 + sint w2" using somethingA somethingB add_mono of_int_add by fastforce show "\r'\ r1 + r2. (2147483647) \ r'" apply(rule exI[where x = "r'\<^sub>1 + r'\<^sub>2"]) using bound1 bound2 add_mono something sum_leq' order.trans proof - have f1: " (real_of_int (sint w2)) = r'\<^sub>2" by (metis bound2 notinf2 notneginf2 repe.cases) have " (real_of_int (sint w1)) = r'\<^sub>1" by (metis bound1 notinf1 notneginf1 repe.cases) then have f2: " (real_of_int 2147483647) \ r'\<^sub>2 + r'\<^sub>1" - using f1 sum_leq' by force + using f1 sum_leq' by auto have "r'\<^sub>2 + r'\<^sub>1 \ r2 + r1" by (meson add_left_mono add_right_mono bound1 bound2 order.trans) then show "r'\<^sub>1 + r'\<^sub>2 \ r1 + r2 \ 2147483647 \ r'\<^sub>1 + r'\<^sub>2" using f2 by (simp add: add.commute) qed qed have inf_case:"((scast POS_INF)::64 Word.word) <=s ?sum \ POS_INF \\<^sub>L r1 + r2" proof - assume "((scast POS_INF)::64 Word.word) <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" then have "((scast ((2 ^ 31 - 1)::word))::64 Word.word) <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" unfolding repL_def repe.simps by auto then have "(0x7FFFFFFF::64 Word.word) <=s ((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)" by auto then obtain r' where geq:"(r' \ r1 + r2)" and leq:"(0x7FFFFFFF \ r')" using bigThree by auto show "?thesis" unfolding repL_def repe.simps using leq geq by auto qed have int_case:"\(((scast POS_INF)::64 Word.word) <=s ?sum) \ \ (?sum <=s ((scast NEG_INF)::64 Word.word)) \ ((scast ?sum)::word) \\<^sub>L r1 + r2" proof - assume bound1:"\ ((scast POS_INF)::64 Word.word) <=s scast w1 + scast w2" assume bound2:"\ scast w1 + scast w2 <=s ((scast NEG_INF)::64 Word.word)" obtain r'::real where rDef:"r' = (real_of_int (sint ((scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word)))::word)))" and r12:"r'\r1 + r2" and boundU:"r' < 0x7FFFFFFF" and boundL:" (-0x7FFFFFFF) < r'" using bigTwo[OF bound1 bound2] by auto obtain w::word where wdef:"w = (scast (((scast w1)::64 Word.word) + ((scast w2)::64 Word.word))::word)" by auto then have wr:"r' = (real_of_int (sint w))" using r12 bound1 bound2 rDef by blast show ?thesis unfolding repL_def repe.simps using r12 wdef rDef boundU boundL wr by auto qed have "(let sum = ?sum in if scast POS_INF <=s sum then POS_INF else if sum <=s scast NEG_INF then NEG_INF else scast sum) \\<^sub>L r1 + r2" apply(cases "((scast POS_INF)::64 Word.word) <=s ?sum") apply(cases "?sum <=s scast NEG_INF") using inf_case neg_inf_case int_case by (auto simp add: Let_def) then show ?thesis using notinf1 notinf2 notneginf1 notneginf2 by(auto) qed qed subsection \Max function\ text\Maximum of w1 + w2 in 2s-complement\ fun wmax :: "word \ word \ word" where "wmax w1 w2 = (if w1 Correctness of wmax\ lemma wmax_lemma: assumes eq1:"w1 \\<^sub>E (r1::real)" assumes eq2:"w2 \\<^sub>E (r2::real)" shows "wmax w1 w2 \\<^sub>E (max r1 r2)" proof(cases rule: case_inf2[where ?w1.0=w1, where ?w2.0=w2]) case PosPos from PosPos eq1 eq2 have bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2:"(real_of_int (sint POS_INF)) \ r2" by (auto simp add: repe.simps) have eqInf:"wmax w1 w2 = POS_INF" using PosPos unfolding wmax.simps by auto have pos_eq:"POS_INF \\<^sub>E max r1 r2" apply(rule repPOS_INF) using bound1 bound2 by linarith show ?thesis using pos_eq eqInf by auto next case PosNeg from PosNeg have bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2:"r2 \ (real_of_int (sint NEG_INF))" using eq1 eq2 by (auto simp add: repe.simps) have eqNeg:"wmax w1 w2 = POS_INF" unfolding eq1 eq2 wmax.simps PosNeg word_sless_def word_sle_def by(auto) have neg_eq:"POS_INF \\<^sub>E max r1 r2" apply(rule repPOS_INF) using bound1 bound2 eq1 eq2 by auto show "?thesis" using eqNeg neg_eq by auto next case PosNum from PosNum eq1 eq2 have bound1:" (real_of_int (sint POS_INF)) \ r1" and bound2a:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound2b:" (real_of_int (sint w2)) < (real_of_int (sint POS_INF))" by (auto simp add: repe.simps) have eqNeg:"wmax w1 w2 = POS_INF" using PosNum bound2b unfolding wmax.simps word_sless_def word_sle_def by auto have neg_eq:"POS_INF \\<^sub>E max r1 r2" apply (rule repPOS_INF) using bound1 bound2a bound2b word_sless_alt le_max_iff_disj unfolding eq1 eq2 by blast show "?thesis" using eqNeg neg_eq by auto next case NegPos from NegPos eq1 eq2 have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:" (real_of_int (sint POS_INF)) \ r2" by (auto simp add: repe.simps) have eqNeg:"wmax w1 w2 = POS_INF" unfolding NegPos word_sless_def word_sle_def by(auto) have neg_eq:"POS_INF \\<^sub>E max r1 r2" apply(rule repPOS_INF) using bound1 bound2 by auto show "wmax w1 w2 \\<^sub>E max r1 r2" using eqNeg neg_eq by auto next case NegNeg from NegNeg eq1 eq2 have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" by (auto simp add: repe.simps) have eqNeg:"NEG_INF \\<^sub>E max r1 r2" apply(rule repNEG_INF) using eq1 eq2 bound1 bound2 by(auto) have neg_eq:"wmax w1 w2 = NEG_INF" using NegNeg by auto show "wmax w1 w2 \\<^sub>E max r1 r2" using eqNeg neg_eq by auto next case NegNum from NegNum eq1 eq2 have eq3:"r2 = (real_of_int (sint w2))" and bound2a:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound1:"r1 \ (real_of_int (sint NEG_INF))" by (auto simp add: repe.simps) have eqNeg:"max r1 r2 = (real_of_int (sint w2))" using NegNum assms(2) bound2a eq3 repeInt_simps bound1 bound2a bound2b by (metis less_irrefl max.bounded_iff max_def not_less) then have extra_eq:"(wmax w1 w2) = w2" using assms(2) bound2a eq3 NegNum repeInt_simps by (simp add: word_sless_alt) have neg_eq:"wmax w1 w2 \\<^sub>E (real_of_int (sint (wmax w1 w2)))" apply(rule repINT) using extra_eq eq3 bound2a bound2b by(auto) show "wmax w1 w2 \\<^sub>E max r1 r2" using eqNeg neg_eq extra_eq by auto next case NumPos from NumPos eq1 eq2 have p2:"w2 = POS_INF" and eq1:"r1 = (real_of_int (sint w1))" and eq2:"r2 = r2" and bound1a:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:"(real_of_int (sint POS_INF)) \ r2" by (auto simp add: repe.simps) have res1:"wmax w1 w2 = POS_INF" using NumPos p2 eq1 eq2 assms(1) bound1b p2 repeInt_simps by (simp add: word_sless_alt) have res3:"POS_INF \\<^sub>E max r1 r2" using repPOS_INF NumPos bound2 le_max_iff_disj by blast show "wmax w1 w2 \\<^sub>E max r1 r2" using res1 res3 by auto next case NumNeg from NumNeg eq1 eq2 have n2:"w2 = NEG_INF" and rw1:"r1 = (real_of_int (sint w1))" and bound1a:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" by (auto simp add: repe.simps) have res1:"max r1 r2 = (real_of_int (sint (wmax w1 w2)))" using bound1b bound2 NumNeg less_trans wmax.simps of_int_less_iff word_sless_alt rw1 antisym_conv2 less_imp_le max_def by metis have res2:"wmax w1 w2 \\<^sub>E (real_of_int (sint (wmax w1 w2)))" apply(rule repINT) using bound1a bound1b bound2 NumNeg leD leI less_trans n2 wmax.simps by (auto simp add: word_sless_alt) show "wmax w1 w2 \\<^sub>E max r1 r2" using res1 res2 by auto next case NumNum from NumNum eq1 eq2 have eq1:"r1 = (real_of_int (sint w1))" and eq2:"r2 = (real_of_int (sint w2))" and bound1a:"(real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2a:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" by (auto simp add: repe.simps) have res1:"max r1 r2 = (real_of_int (sint (wmax w1 w2)))" using eq1 eq2 bound1a bound1b bound2a bound2b - by (simp add: max_def word_sless_alt) + apply (auto simp add: max_def word_sless_alt not_less; transfer) + apply simp_all + done have res2:"wmax w1 w2 \\<^sub>E (real_of_int (sint (wmax w1 w2)))" apply (rule repINT) using bound1a bound1b bound2a bound2b by (simp add: \max r1 r2 = (real_of_int (sint (wmax w1 w2)))\ eq2 min_less_iff_disj)+ show "wmax w1 w2 \\<^sub>E max r1 r2" using res1 res2 by auto qed lemma max_repU1: assumes "w1 \\<^sub>U x" assumes "w2 \\<^sub>U y" shows "wmax w1 w2 \\<^sub>U x " using wmax_lemma assms repU_def by (meson le_max_iff_disj) lemma max_repU2: assumes "w1 \\<^sub>U y" assumes "w2 \\<^sub>U x" shows "wmax w1 w2 \\<^sub>U x" using wmax_lemma assms repU_def by (meson le_max_iff_disj) text\Product of w1 * w2 with bounds checking\ fun wtimes :: "word \ word \ word" where "wtimes w1 w2 = (if w1 = POS_INF \ w2 = POS_INF then POS_INF else if w1 = NEG_INF \ w2 = POS_INF then NEG_INF else if w1 = POS_INF \ w2 = NEG_INF then NEG_INF else if w1 = NEG_INF \ w2 = NEG_INF then POS_INF else if w1 = POS_INF \ w2 0 0 = w2 then 0 else if w1 = NEG_INF \ w2 0 0 = w2 then 0 else if w1 w2 = POS_INF then NEG_INF else if 0 w2 = POS_INF then POS_INF else if 0 = w1 \ w2 = POS_INF then 0 else if w1 w2 = NEG_INF then POS_INF else if 0 w2 = NEG_INF then NEG_INF else if 0 = w1 \ w2 = NEG_INF then 0 else (let prod::64 Word.word = (scast w1) * (scast w2) in if prod <=s (scast NEG_INF) then NEG_INF else if (scast POS_INF) <=s prod then POS_INF else (scast prod)))" subsection \Multiplication upper bound\ text\Product of 32-bit numbers fits in 64 bits\ lemma times_upcast_lower: fixes x y::int assumes a1:"x \ -2147483648" assumes a2:"y \ -2147483648" assumes a3:"x \ 2147483648" assumes a4:"y \ 2147483648" shows "- 4611686018427387904 \ x * y" proof - let ?thesis = "- 4611686018427387904 \ x * y" have is_neg:"- 4611686018427387904 < (0::int)" by auto have case1:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"x * y \ 0" using a5 a6 by (simp add: zero_le_mult_iff) then show ?thesis using is_neg by auto qed have case2:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"-2147483648 * (2147483648) \ x * 2147483648" using a1 a2 a3 a4 a5 a6 by linarith have h2:"-2147483648 \ y" using a6 by auto have h3:"x * 2147483648 \ x * y" using a1 a2 a3 a4 a5 a6 h2 using mult_left_mono_neg by blast show ?thesis using h1 h2 h3 by auto qed have case3:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * (-2147483648) \ 2147483648 * y" using a1 a2 a3 a4 a5 a6 by linarith have h2:"-2147483648 \ x" using a5 by auto have h3:"2147483648 * y \ x * y" using a1 a2 a3 a4 a5 a6 h2 using mult_left_mono_neg mult_right_mono_neg by blast show ?thesis using h1 h2 h3 by auto qed have case4:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"x * y \ 0" using a5 a6 by (simp add: zero_le_mult_iff) then show ?thesis using is_neg by auto qed show ?thesis using case1 case2 case3 case4 by linarith qed text\Product of 32-bit numbers fits in 64 bits\ lemma times_upcast_upper: fixes x y ::int assumes a1:"x \ -2147483648" assumes a2:"y \ -2147483648" assumes a3:"x \ 2147483648" assumes a4:"y \ 2147483648" shows "x * y \ 4611686018427387904" proof - let ?thesis = "x * y \ 4611686018427387904" have case1:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * 2147483648 \ x * 2147483648" using a1 a2 a3 a4 a5 a6 by linarith have h2:"x * 2147483648 \ x * y" using a1 a2 a3 a4 a5 a6 by (simp add: mult_mono) show ?thesis using h1 h2 by auto qed have case2:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * 2147483648 \ (0::int)" by auto have h2:"0 \ x * y" using a5 a6 mult_nonneg_nonpos2 by blast show ?thesis using h1 h2 by auto qed have case3:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"2147483648 * 2147483648 \ (0::int)" by auto have h2:"0 \ x * y" using a5 a6 mult_nonneg_nonpos by blast show ?thesis using h1 h2 by auto qed have case4:"x \ 0 \ y \ 0 \ ?thesis" proof - assume a5:"x \ 0" and a6:"y \ 0" have h1:"-2147483648 * -2147483648 \ x * -2147483648" using a1 a2 a3 a4 a5 a6 by linarith have h2:"x * -2147483648 \ x * y" using a1 a2 a3 a4 a5 a6 mult_left_mono_neg by blast show ?thesis using h1 h2 by auto qed show "x * y \ 4611686018427387904" using case1 case2 case3 case4 by linarith qed text\Correctness of 32x32 bit multiplication\ subsection \Exact multiplication\ lemma wtimes_exact: assumes eq1:"w1 \\<^sub>E r1" assumes eq2:"w2 \\<^sub>E r2" shows "wtimes w1 w2 \\<^sub>E r1 * r2" proof - have POS_cast:"sint ((scast POS_INF)::64 Word.word) = sint POS_INF" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have POS_sint:"sint POS_INF = (2^31)-1" by auto have w1_cast:"sint ((scast w1)::64 Word.word) = sint w1" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have w2_cast:"sint ((scast w2)::64 Word.word) = sint w2" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have NEG_cast:"sint ((scast NEG_INF)::64 Word.word) = sint NEG_INF" apply(rule Word.sint_up_scast) unfolding Word.is_up by auto have rangew1:"sint ((scast w1)::64 Word.word) \ {- (2 ^ 31).. (2^31)} " using Word.word_sint.Rep[of "(w1)::32 Word.word"] sints32 len32 mem_Collect_eq POS_cast w1_cast by auto have rangew2:"sint ((scast w2)::64 Word.word) \ {- (2 ^ 31).. (2^31)} " using Word.word_sint.Rep[of "(w2)::32 Word.word"] sints32 len32 mem_Collect_eq POS_cast w2_cast by auto show ?thesis proof (cases rule: case_times_inf[of w1 w2]) case PosPos then have a1: "PosInf \ r1" and a2: "PosInf \ r2" using "PosPos" eq1 eq2 repe.simps by (auto) have f3: "\n e::real. 1 \ max ( (numeral n)) e" by (simp add: le_max_iff_disj) have "\n e::real. 0 \ max ( (numeral n)) e" by (simp add: le_max_iff_disj) then have "r1 \ r1 * r2" using f3 "PosPos" eq1 eq2 repe.simps using eq1 eq2 by (auto simp add: repe.simps) then have "PosInf \ r1 * r2" using a1 by linarith then show ?thesis using "PosPos" by(auto simp add: repe.simps) next case NegPos from "NegPos" have notPos:"w1 \ POS_INF" unfolding POS_INF_def NEG_INF_def by auto have a1: "NegInf \ r1" using eq1 "NegPos" by (auto simp add: repe.simps) have a2: "PosInf \ r2" using eq2 "NegPos" by (auto simp add: repe.simps) have f1: "real_of_int Numeral1 = 1" by simp have f3: "(real_of_int 3) \ - r1" using a1 by (auto) have f4:"0 \ r2" using f1 a2 by(auto) have f5: "r1 \ - 1" using f3 by auto have fact:"r1 * r2 \ - r2" using f5 f4 mult_right_mono by fastforce show ?thesis using a1 a2 fact by (auto simp add: repe.simps "NegPos") next case PosNeg have a1: "PosInf \ r1" using eq1 "PosNeg" by (auto simp add: repe.simps) then have h1:"r1 \ 1" by (auto) have a2: " NegInf \ r2" using eq2 "PosNeg" by (auto simp add: repe.simps) have f1: "\ NegInf * (- 1) \ 1" by (auto) have f2: "\e ea::real. (e * (- 1) \ ea) = (ea * (- 1) \ e)" by force then have f3: "\ 1 * (- 1::real) \ NegInf" using f1 by blast have f4: "r1 * (- 1) \ NegInf" using f2 a1 by(auto) have f5: "\e ea eb. (if (ea::real) \ eb then e \ eb else e \ ea) = (e \ ea \ e \ eb)" by force have " 0 * (- 1::real) \ 1" by simp then have "r1 * (- 1) \ 0" using f5 f4 f3 f2 by meson then have f6: "0 \ r1" by fastforce have "1 * (- 1) \ (- 1::real)" using f2 by force then have fact:"r2 \ (- 1)" using f3 a2 by fastforce have rule:"\c. c > 0 \ r1 \ c \ r2 \ -1 \ r1 * r2 \ -c" apply auto by (metis (no_types, hide_lams) f5 mult_less_cancel_left_pos mult_minus1_right neg_le_iff_le not_less) have "r1 * r2 \ NegInf" using "PosNeg" f6 fact rule[of PosInf] a1 by(auto) then show ?thesis using "PosNeg" by (auto simp add: repe.simps) next case NegNeg have a1: "(-2147483647) \ r1" using eq1 "NegNeg" by (auto simp add: repe.simps) then have h1:"r1 \ -1" using max.bounded_iff max_def one_le_numeral by auto have a2: " (-2147483647) \ r2" using eq2 "NegNeg" by (auto simp add: repe.simps) have f1: "\e ea eb. \ (e::real) \ ea \ \ 0 \ eb \ eb * e \ eb * ea" using mult_left_mono by metis have f2: "- 1 = (- 1::real)" by force have f3: " 0 \ (1::real)" by simp have f4: "\e ea eb. (ea::real) \ e \ \ ea \ eb \ \ eb \ e" by (meson less_le_trans not_le) have f5: " 0 \ (2147483647::real)" by simp have f6: "- (- 2147483647) = (2147483647::real)" by force then have f7: "- ( (- 2147483647) * r1) = (2147483647 * r1)" by (metis mult_minus_left) have f8: "- ( (- 2147483647) * (- 1)) = 2147483647 * (- 1::real)" by simp have " 2147483647 = - 1 * (- 2147483647::real)" by simp then have f9: "r1 \ (- 1) \ 2147483647 \ r1 * (- 2147483647)" using f8 f7 f5 f2 f1 by linarith have f10: "- 2147483647 = (- 2147483647::real)" by fastforce have f11: "- (r2 * 1 * (r1 * (- 1))) = r1 * r2" by (simp add: mult.commute) have f12: "r1 * (- 1) = - (r1 * 1)" by simp have "r1 * 1 * ( (- 2147483647) * 1) = (- 2147483647) * r1" by (simp add: mult.commute) then have f13: "r1 * (- 1) * ( (- 2147483647) * 1) = 2147483647 * r1" using f12 f6 by (metis (no_types) mult_minus_left) have " 1 * r1 \ 1 * (- 2147483647)" using a1 by (auto simp add: a1) then have " 2147483647 \ r1 * (- 1)" by fastforce then have "0 \ r1 * (- 1)" using f5 f4 by (metis) then have "r1 \ (- 1) \ - (r1 * 2147483647) \ - (r2 * 1 * (r1 * (- 1)))" by (metis a2 f11 h1 mult_left_mono_neg minus_mult_right mult_minus1_right neg_0_le_iff_le) then have "r1 \ (- 1) \ r1 * (- 2147483647) \ r2 * r1" using f11 f10 by (metis mult_minus_left mult.commute) then have fact:" 2147483647 \ r2 * r1" using f9 f4 by blast show ?thesis using a1 a2 h1 fact by (auto simp add: repe.simps "NegNeg" mult.commute) next case PosLo from "PosLo" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "PosLo" have upper:" (real_of_int (sint POS_INF)) \ r1 " by (auto simp add: repe.simps) have lower1:"sint w2 < 0" using "PosLo" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower2:"sint w2 \ -1" by auto from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "PosLo" by (auto simp add: repe.simps) have f4: "r1 * (- 1) \ (- 2147483647)" using upper by (auto) have f5: "r2 \ (- 1)" - using lower2 rw2 by simp + using lower2 rw2 by transfer simp have "0 < r1" using upper by (auto) have "\r. r < - 2147483647 \ \ r < r1 * - 1" using f4 less_le_trans by blast then have "r1 * (real_of_int (sint w2)) \ (- 2147483647)" using f5 f4 upper lower2 rw2 mult_left_mono by (metis \0 < r1\ dual_order.order_iff_strict f5 mult_left_mono rw2) then have "r1 * r2 \ real_of_int (sint NEG_INF)" using upper lower2 rw2 by (auto) then show ?thesis using "PosLo" by (auto simp add: repe.simps) next case PosHi from "PosHi" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "PosHi" have upper:"(real_of_int (sint POS_INF)) \ r1 " by (auto simp add: repe.simps) have lower1:"sint w2 > 0" using "PosHi" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower2:"sint w2 \ 1" by auto from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "PosHi" by (auto) have "0 \ r1" using upper by (auto) then have "r1 \ r1 * r2" using rw2 lower2 by (metis (no_types) mult_left_mono mult.right_neutral of_int_1_le_iff) have "PosInf \ r1 * r2" using upper lower2 rw2 apply (auto) - using \0 \ r1\ mult_mono mult_numeral_1_right numeral_One of_int_1_le_iff zero_le_one - by metis + using \0 \ r1\ mult_numeral_1_right numeral_One of_int_1_le_iff zero_le_one + apply simp + using mult_mono [of 2147483647 r1 1 \signed_real_of_word (w2::32 Word.word)\] + apply simp + apply transfer + apply simp + done then show ?thesis using "PosHi" by (auto simp add: repe.simps) next case PosZero from "PosZero" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "PosZero" have upper:" (real_of_int (sint POS_INF)) \ r1 " by (auto simp add: repe.simps) have lower1:"sint w2 = 0" using "PosZero" by (auto simp add: word_sless_def word_sle_def) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "PosZero" by auto have "0 = r1 * r2" using "PosZero" rw2 by auto then show ?thesis using "PosZero" by (auto simp add: repe.simps) next case NegHi have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" using "NegHi" by (auto) from eq1 "NegHi" have upper:"(real_of_int (sint NEG_INF)) \ r1 " by (auto simp add: repe.simps) have low:"sint w2 > 0" using "NegHi" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) - then have lower1:"(real_of_int (sint w2)) > 0" by auto + then have lower1:"(real_of_int (sint w2)) > 0" + by transfer simp then have lower2:"(real_of_int (sint w2)) \ 1" - using low - by (simp add: int_le_real_less) + using low by transfer simp + from eq1 have rw1:"r1 \ (real_of_int (sint w1))" using repe.simps "NegHi" by (simp add: upper) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "NegHi" by (auto) have mylem:"\x y z::real. x \ -1 \ y \ 1 \ z \ -1 \ x \ z \ x * y \ z" proof - fix x y z::real assume a1:"x \ -1" assume a2:"y \ 1" then have h1:"-1 \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) assume a5:"x \ z" then have h6:"-x \ -z" by auto have h3:"-x * -z = x * z" by auto show "x * y \ z" using a1 a2 a3 a5 a4 h1 h2 h3 h6 h5 a5 dual_order.trans leD mult.right_neutral by (metis dual_order.order_iff_strict mult_less_cancel_left2) qed have prereqs:"r1 \ - 1" "1 \ (real_of_int (sint w2))" " (- 2147483647::real) \ - 1 " "r1 \ (-2147483647)" using rw1 rw2 "NegHi" lower2 by (auto simp add: word_sless_def word_sle_def) have "r1 * r2 \ real_of_int (sint NEG_INF)" using upper lower1 lower2 rw1 rw2 apply (auto simp add: word_sless_def word_sle_def) using mylem[of "r1" " (real_of_int (sint w2))" " (- 2147483647)"] prereqs by auto then show ?thesis using "NegHi" by (auto simp add: repe.simps) next case NegLo from "NegLo" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq1 "NegLo" have upper:"(real_of_int (sint NEG_INF)) \ r1" by (auto simp add: repe.simps) have low:"sint w2 < 0" using "NegLo" by (auto simp add: word_sless_def word_sle_def dual_order.order_iff_strict) - then have lower1:"(real_of_int (sint w2)) < 0" by auto - then have lower2:"(real_of_int (sint w2)) \ -1" using low - by (simp add: int_le_real_less) + then have lower1:"(real_of_int (sint w2)) < 0" + by transfer simp from eq1 have rw1:"r1 \ (real_of_int (sint w1))" using repe.simps "NegLo" by (simp add: upper) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps "NegLo" by (auto) have hom:"(- 2147483647) = -(2147483647::real)" by auto have mylem:"\x y z::real. y < 0 \ x \ y \ z \ -1 \ -y \ x * z" proof - fix x y z::real assume a1:"y < 0" assume a2:"x \ y" then have h1:"-x \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) have h4:"-x * -z \ -y" using a1 a2 a3 a4 h1 h2 h5 dual_order.trans mult.right_neutral by (metis mult.commute neg_0_less_iff_less real_mult_le_cancel_iff1) have h3:"-x * -z = x * z" by auto show "- y \ x * z " using a1 a2 a3 a4 h1 h2 h3 h4 h5 by simp qed - have prereqs:"- 2147483647 < (0::real)" " r1 \ - 2147483647" " (real_of_int (sint w2)) \ - 1" - using rw1 rw2 "NegLo" by (auto simp add: word_sless_def word_sle_def lower2) - have "2147483647 \ r1 * r2" - using upper lower1 lower2 rw1 rw2 prereqs + have prereqs:"- 2147483647 < (0::real)" " r1 \ - 2147483647" + using rw1 rw2 "NegLo" by (auto simp add: word_sless_def word_sle_def) + moreover have \sint w2 \ - 1\ + using low by simp + then have \real_of_int (sint w2) \ real_of_int (- 1)\ + by (simp only: of_int_le_iff) + then have \signed_real_of_word w2 \ - 1\ + by simp + ultimately have "2147483647 \ r1 * r2" + using upper lower1 rw1 rw2 mylem[of "-2147483647" "r1" "(real_of_int (sint w2))"] by (auto simp add: word_sless_def word_sle_def) then show ?thesis using "NegLo" by (auto simp add: repe.simps) next case NegZero from "NegZero" have w2NotPinf:"w2 \ POS_INF" and w2NotNinf:"w2 \ NEG_INF" by (auto) from eq2 "NegZero" have "r2 = 0" using repe.simps "NegZero" by (auto) then show ?thesis using "NegZero" by (auto simp add: repe.simps) next case LoPos from "LoPos" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "LoPos" have upper:"(real_of_int (sint POS_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 < 0" using "LoPos" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) then have lower2:"sint w1 \ -1" by auto from eq1 have rw1:"r1 = (real_of_int (sint w1))" using repe.simps "LoPos" by (auto simp add: repe.simps) have f4: "r2 * (- 1) \ (- 2147483647)" using upper by(auto) have f5: "r1 \ (- 1)" - using lower2 rw1 by simp + using lower2 rw1 by transfer simp have "0 < r2" using upper by(auto) then have "r2 * r1 \ r2 * (- 1)" by (metis dual_order.order_iff_strict mult_right_mono f5 mult.commute) then have "r2 * r1 \ (- 2147483647)" by (meson f4 less_le_trans not_le) then have "(real_of_int (sint w1)) * r2 \ (- 2147483647)" using f5 f4 rw1 less_le_trans not_le mult.commute rw1 by (auto simp add: mult.commute) then have "r1 * r2 \ NegInf" using rw1 by (auto) then show ?thesis using "LoPos" by (auto simp: repe.simps) next case HiPos from "HiPos" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "HiPos" have upper:"(real_of_int (sint POS_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 > 0" using "HiPos" by (auto simp add: word_sless_def word_sle_def dual_order.order_iff_strict) then have lower2:"sint w1 \ 1" by auto from eq1 have rw2:"r1 = (real_of_int (sint w1))" using "HiPos" by (auto simp add: repe.simps) have "0 \ r2" using upper by(auto) then have "r2 \ r2 * r1" using lower2 rw2 by (metis (no_types) mult_left_mono mult.right_neutral of_int_1_le_iff) have "2147483647 \ r1 * r2" - using upper lower2 rw2 - by (auto simp add: word_sless_def word_sle_def order_trans) + using upper lower2 rw2 + apply (simp add: word_sless_def word_sle_def) + using mult_mono [of 1 \signed_real_of_word w1\ 2147483647 r2] + apply simp + apply transfer + apply simp + done then show ?thesis using "HiPos" by (auto simp add: repe.simps) next case ZeroPos from "ZeroPos" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "ZeroPos" have upper:" (real_of_int (sint POS_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 = 0" using "ZeroPos" by (auto simp add: word_sless_def word_sle_def) from eq1 have rw2:"r1 = (real_of_int (sint w1))" using repe.simps "ZeroPos" by (auto) have "r1 = 0" using lower1 rw2 by auto then show ?thesis using "ZeroPos" by (auto simp add: repe.simps) next case ZeroNeg from "ZeroNeg" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "ZeroNeg" have upper:"(real_of_int (sint NEG_INF)) \ r2 " by (auto simp add: repe.simps) have lower1:"sint w1 = 0" using "ZeroNeg" by (auto simp add: word_sless_def word_sle_def) from eq1 have rw2:"r1 = (real_of_int (sint w1))" using repe.simps "ZeroNeg" by (auto) have "r1 = 0" using lower1 rw2 by auto then show ?thesis using "ZeroNeg" by (auto simp add: repe.simps) next case LoNeg from "LoNeg" have w2NotPinf:"w1 \ POS_INF" and w2NotNinf:"w1 \ NEG_INF" by (auto) from eq2 "LoNeg" have upper:" (real_of_int (sint NEG_INF)) \ r2 " by (auto simp add: repe.simps) have low:"sint w1 < 0" using "LoNeg" apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) - then have lower1:"(real_of_int (sint w1)) < 0" by auto - then have lower2:"(real_of_int (sint w1)) \ -1" using low - by (simp add: int_le_real_less) + then have lower1:"(real_of_int (sint w1)) < 0" by transfer simp + from low have \sint w1 \ - 1\ + by simp + then have lower2:"(real_of_int (sint w1)) \ -1" + by transfer simp from eq1 have rw1:"r2 \ (real_of_int (sint w2))" using "LoNeg" upper by auto from eq1 have rw2:"r1 = (real_of_int (sint w1))" using "LoNeg" by (auto simp add: upper repe.simps) have hom:"(- 2147483647::real) = -(2147483647)" by auto have mylem:"\x y z::real. y < 0 \ x \ y \ z \ -1 \ -y \ x * z" proof - fix x y z::real assume a1:"y < 0" assume a2:"x \ y" then have h1:"-x \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) have h4:"-x * -z \ -y" using a1 a2 a3 a4 h1 h2 h5 dual_order.trans mult_left_mono mult.right_neutral mult.commute by (metis dual_order.order_iff_strict mult_minus_right mult_zero_right neg_le_iff_le) have h3:"-x * -z = x * z" by auto show "- y \ x * z " using a1 a2 a3 a4 h1 h2 h3 h4 h5 by simp qed have prereqs:"- 2147483647 < (0::real)" " r2 \ - 2147483647" " (real_of_int (sint w1)) \ - 1" using rw1 rw2 "LoNeg" lower2 by (auto simp add: word_sless_def word_sle_def lower2) have "2147483647 \ r1 * r2" using upper lower1 lower2 rw1 rw2 mylem[of "-2147483647" "r2" "(real_of_int (sint w1))"] prereqs by (auto simp add:word_sless_def word_sle_def mult.commute) then show ?thesis using "LoNeg" by (auto simp add: repe.simps) next case HiNeg from HiNeg have w1NotPinf:"w1 \ POS_INF" and w1NotNinf:"w1 \ NEG_INF" by (auto) have upper:" (real_of_int (sint NEG_INF)) \ r2 " using HiNeg eq2 by (auto simp add: repe.simps ) have low:"sint w1 > 0" using HiNeg apply (auto simp add: word_sless_def word_sle_def) by (simp add: dual_order.order_iff_strict) - then have lower1:"(real_of_int (sint w1)) > 0" by auto - then have lower2:"(real_of_int (sint w1)) \ 1" using low - by (simp add: int_le_real_less) + then have lower1:"(real_of_int (sint w1)) > 0" by transfer simp + from low have \sint w1 \ 1\ + by simp + then have lower2:"(real_of_int (sint w1)) \ 1" + by transfer simp from eq2 have rw1:"r2 \ (real_of_int (sint w2))" using repe.simps HiNeg by (simp add: upper) from eq1 have rw2:"r1 = (real_of_int (sint w1))" using repe.simps HiNeg by (auto) have mylem:"\x y z::real. x \ -1 \ y \ 1 \ z \ -1 \ x \ z \ x * y \ z" proof - fix x y z::real assume a1:"x \ -1" assume a2:"y \ 1" then have h1:"-1 \ -y" by auto assume a3:"z \ -1" then have a4:"z < 0" by auto from a4 have h2:"-z > 0" using leD leI by auto from a3 have h5:"-z \ 1" by (simp add: leD leI) assume a5:"x \ z" then have h6:"-x \ -z" by auto have h3:"-x * -z = x * z" by auto show "x * y \ z" using a1 a2 a3 a4 h1 h2 h3 h6 h5 a5 dual_order.trans less_eq_real_def by (metis mult_less_cancel_left1 not_le) qed have prereqs:"r2 \ - 1" "1 \ (real_of_int (sint w1))" " (- 2147483647) \ - (1::real )" "r2 \ (- 2147483647)" using rw1 rw2 HiNeg lower2 by (auto simp add: word_sless_def word_sle_def) have "r1 * r2 \ - 2147483647" using upper lower1 lower2 rw1 rw2 apply (auto simp add: word_sless_def word_sle_def) using mylem[of "r2" "(real_of_int (sint w1))" " (- 2147483647)"] prereqs by (auto simp add: mult.commute) then show ?thesis using HiNeg by(auto simp add: repe.simps) next case AllFinite let ?prod = "(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))" consider (ProdNeg) "?prod <=s ((scast NEG_INF)::64 Word.word)" | (ProdPos) "(((scast POS_INF)::64 Word.word) <=s ?prod)" | (ProdFin) "\(?prod <=s ((scast NEG_INF)::64 Word.word)) \ \((scast POS_INF)::64 Word.word) <=s ?prod" by (auto) then show ?thesis proof (cases) case ProdNeg have bigLeq:"(4611686018427387904::real) \ 9223372036854775807" by auto have set_cast:"\x::int. (x \ {-(2^31)..2^31}) = ( (real_of_int x) \ {-(2^31)..2^31})" by auto have eq3:"sint(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) = sint ((scast w1)::64 Word.word) * sint ((scast w2)::64 Word.word)" apply(rule Word_Lemmas.signed_arith_sint(4)) using rangew1 rangew2 w1_cast w2_cast using Word.word_size[of "((scast w1)::64 Word.word)"] using Word.word_size[of "((scast w2)::64 Word.word)"] using times_upcast_upper[of "sint w1" "sint w2"] using times_upcast_lower[of "sint w1" "sint w2"] by auto assume "?prod <=s ((scast NEG_INF)::64 Word.word)" then have sint_leq:"sint ?prod \ sint ((scast NEG_INF)::64 Word.word)" using word_sle_def by blast have neqs:"w1 \ POS_INF" " w1 \ NEG_INF" "w2 \ POS_INF" "w2 \ NEG_INF" using AllFinite word_sless_def signed.not_less_iff_gr_or_eq by force+ from eq1 have rw1:"r1 = (real_of_int (sint w1))" using neqs by (auto simp add: repe.simps) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using neqs by (auto simp add: repe.simps) show ?thesis using AllFinite ProdNeg w1_cast w2_cast rw1 rw2 sint_leq - apply (auto simp add: repe.simps) - by (metis (no_types, hide_lams) eq3 of_int_le_iff of_int_minus of_int_mult of_int_numeral) + apply (auto simp add: repe.simps eq3) + apply (subst (asm) of_int_le_iff [symmetric, where ?'a = real]) + apply simp + done next case ProdPos have bigLeq:"(4611686018427387904::real) \ 9223372036854775807" by auto have set_cast:"\x::int. (x \ {-(2^31)..2^31}) = ( (real_of_int x) \ {-(2^31)..2^31})" by auto have eq3:"sint(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) = sint ((scast w1)::64 Word.word) * sint ((scast w2)::64 Word.word)" apply(rule Word_Lemmas.signed_arith_sint(4)) using rangew1 rangew2 POS_cast POS_sint w1_cast w2_cast using Word.word_size[of "((scast w1)::64 Word.word)"] using Word.word_size[of "((scast w2)::64 Word.word)"] using times_upcast_upper[of "sint w1" "sint w2"] using times_upcast_lower[of "sint w1" "sint w2"] by auto assume cast:"((scast POS_INF)::64 Word.word) <=s ?prod" then have sint_leq:"sint ((scast POS_INF)::64 Word.word) \ sint ?prod" using word_sle_def by blast have neqs:"w1 \ POS_INF" " w1 \ NEG_INF" "w2 \ POS_INF" "w2 \ NEG_INF" using AllFinite word_sless_def signed.not_less_iff_gr_or_eq by force+ from eq1 have rw1:"r1 = (real_of_int (sint w1))" using repe.simps AllFinite neqs by auto from eq2 have rw2:"r2 = (real_of_int (sint w2))" using repe.simps AllFinite neqs by auto have prodHi:"r1 * r2 \ PosInf" - using w1_cast w2_cast rw1 rw2 sint_leq apply(auto) - by (metis (no_types, hide_lams) eq3 of_int_le_iff of_int_mult of_int_numeral) + using w1_cast w2_cast rw1 rw2 sint_leq apply (auto simp add: eq3) + apply (subst (asm) of_int_le_iff [symmetric, where ?'a = real]) + apply simp + done have infs:"SCAST(32 \ 64) NEG_INF 64) POS_INF" by (auto) have casted:"SCAST(32 \ 64) POS_INF <=s SCAST(32 \ 64) w1 * SCAST(32 \ 64) w2" using cast by auto have almostContra:"SCAST(32 \ 64) NEG_INF 64) w1 * SCAST(32 \ 64) w2" using infs cast signed.order.strict_trans2 by blast have contra:"\(SCAST(32 \ 64) w1 * SCAST(32 \ 64) w2 <=s SCAST(32 \ 64) NEG_INF)" using eq3 almostContra by auto have wtimesCase:"wtimes w1 w2 = POS_INF" using neqs ProdPos almostContra wtimes.simps AllFinite ProdPos by (auto simp add: repe.simps Let_def) show ?thesis using prodHi apply(simp only: repe.simps) apply(rule disjI1) apply(rule exI[where x= "r1*r2"]) apply(rule conjI) apply(rule wtimesCase) using prodHi by auto next case ProdFin have bigLeq:"(4611686018427387904::real) \ 9223372036854775807" by auto have set_cast:"\x::int. (x \ {-(2^31)..2^31}) = ( (real_of_int x) \ {-(2^31)..2^31})" by auto have eq3:"sint(((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) = sint ((scast w1)::64 Word.word) * sint ((scast w2)::64 Word.word)" apply(rule Word_Lemmas.signed_arith_sint(4)) using rangew1 rangew2 POS_cast POS_sint w1_cast w2_cast using Word.word_size[of "((scast w1)::64 Word.word)"] using Word.word_size[of "((scast w2)::64 Word.word)"] using times_upcast_upper[of "sint w1" "sint w2"] using times_upcast_lower[of "sint w1" "sint w2"] by auto from ProdFin have a1:"\(?prod <=s ((scast NEG_INF)::64 Word.word))" by auto then have sintGe:"sint (?prod) > sint (((scast NEG_INF)::64 Word.word))" using word_sle_def dual_order.order_iff_strict signed.linear by fastforce from ProdFin have a2:"\((scast POS_INF)::64 Word.word) <=s ?prod" by auto then have sintLe:"sint (((scast POS_INF)::64 Word.word)) > sint (?prod)" using word_sle_def dual_order.order_iff_strict signed.linear by fastforce have neqs:"w1 \ POS_INF" " w1 \ NEG_INF" "w2 \ POS_INF" "w2 \ NEG_INF" using AllFinite word_sless_def signed.not_less_iff_gr_or_eq by force+ from eq1 have rw1:"r1 = (real_of_int (sint w1))" using neqs by(auto simp add: repe.simps) from eq2 have rw2:"r2 = (real_of_int (sint w2))" using neqs by(auto simp add: repe.simps) from rw1 rw2 have "r1 * r2 = (real_of_int ((sint w1) * (sint w2)))" by simp have rightSize:"sint (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)) \ sints (len_of TYPE(32))" using sintLe sintGe sints32 by (simp) have downcast:"sint ((scast (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word)))::word) = sint (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))" using scast_down_range[OF rightSize] by auto then have res_eq:"r1 * r2 = real_of_int(sint((scast (((scast w1)::64 Word.word)*((scast w2)::64 Word.word)))::word))" using rw1 rw2 eq3 POS_cast POS_sint w1_cast w2_cast downcast \r1 * r2 = (real_of_int (sint w1 * sint w2))\ by (auto) have res_up:"sint (scast (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))::word) < sint POS_INF" using rw1 rw2 eq3 POS_cast POS_sint w1_cast w2_cast downcast \r1 * r2 = (real_of_int (sint w1 * sint w2))\ \sint (scast w1 * scast w2) < sint (scast POS_INF)\ of_int_eq_iff res_eq by presburger have res_lo:"sint NEG_INF < sint (scast (((scast w1)::64 Word.word) * ((scast w2)::64 Word.word))::word)" using rw1 rw2 eq3 POS_cast POS_sint w1_cast w2_cast NEG_cast downcast \r1 * r2 = (real_of_int (sint w1 * sint w2))\ \sint (scast NEG_INF) < sint (scast w1 * scast w2)\ of_int_eq_iff res_eq by presburger have "scast ?prod \\<^sub>E (r1 * r2)" using res_eq res_up res_lo - by (auto simp add: rep_simps) + apply (auto simp add: rep_simps) + using repeInt_simps by auto then show ?thesis using AllFinite ProdFin by(auto) qed qed qed subsection \Multiplication upper bound\ text\Upper bound of multiplication from upper and lower bounds\ fun tu :: "word \ word \ word \ word \ word" where "tu w1l w1u w2l w2u = wmax (wmax (wtimes w1l w2l) (wtimes w1u w2l)) (wmax (wtimes w1l w2u) (wtimes w1u w2u))" lemma tu_lemma: assumes u1:"u\<^sub>1 \\<^sub>U (r1::real)" assumes u2:"u\<^sub>2 \\<^sub>U (r2::real)" assumes l1:"l\<^sub>1 \\<^sub>L (r1::real)" assumes l2:"l\<^sub>2 \\<^sub>L (r2::real)" shows "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U (r1 * r2)" proof - obtain rl1 rl2 ru1 ru2 :: real where gru1:"ru1 \ r1" and gru2:"ru2 \ r2" and grl1:"rl1 \ r1" and grl2:"rl2 \ r2" and eru1:"u\<^sub>1 \\<^sub>E ru1" and eru2:"u\<^sub>2 \\<^sub>E ru2" and erl1:"l\<^sub>1 \\<^sub>E rl1" and erl2:"l\<^sub>2 \\<^sub>E rl2" using u1 u2 l1 l2 unfolding repU_def repL_def by auto have timesuu:"wtimes u\<^sub>1 u\<^sub>2 \\<^sub>E ru1 * ru2" using wtimes_exact[OF eru1 eru2] by auto have timesul:"wtimes u\<^sub>1 l\<^sub>2 \\<^sub>E ru1 * rl2" using wtimes_exact[OF eru1 erl2] by auto have timeslu:"wtimes l\<^sub>1 u\<^sub>2 \\<^sub>E rl1 * ru2" using wtimes_exact[OF erl1 eru2] by auto have timesll:"wtimes l\<^sub>1 l\<^sub>2 \\<^sub>E rl1 * rl2" using wtimes_exact[OF erl1 erl2] by auto have maxt12:"wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>E max (rl1 * rl2) (ru1 * rl2)" by (rule wmax_lemma[OF timesll timesul]) have maxt34:"wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>E max (rl1 * ru2) (ru1 * ru2)" by (rule wmax_lemma[OF timeslu timesuu]) have bigMax:"wmax (wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>E max (max (rl1 * rl2) (ru1 * rl2)) (max (rl1 * ru2) (ru1 * ru2))" by (rule wmax_lemma[OF maxt12 maxt34]) obtain maxt12val :: real where maxU12:"wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>U max (rl1 * rl2) (ru1 * rl2)" using maxt12 unfolding repU_def by blast obtain maxt34val :: real where maxU34:"wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>U max (rl1 * ru2) (ru1 * ru2)" using maxt34 unfolding repU_def by blast obtain bigMaxU:"wmax (wmax (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmax (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>U max (max (rl1 * rl2) (ru1 * rl2)) (max (rl1 * ru2) (ru1 * ru2))" using bigMax unfolding repU_def by blast have ivl1:"rl1 \ ru1" using grl1 gru1 by auto have ivl2:"rl2 \ ru2" using grl2 gru2 by auto let ?thesis = "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" show ?thesis using ivl1 ivl2 proof(cases rule: case_ivl_zero) case ZeroZero assume "rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2" then have geq1:"ru1 \ 0" and geq2:"ru2 \ 0" by auto consider "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" using le_cases by auto then show "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" proof (cases) case 1 have g1:"ru1 * ru2 \ ru1 * r2" using "1" geq1 geq2 grl2 gru2 by (simp add: mult_left_mono) have g2:"ru1 * r2 \ r1 * r2" using "1" geq1 geq2 grl1 grl2 gru1 gru2 by (simp add: mult_right_mono) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up eru1 eru2 erl1 erl2 repU_def timesuu tu.simps max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] by (metis wmax.elims) next case 2 have g1:"ru1 * ru2 \ 0" using "2" geq1 geq2 grl2 gru2 by (simp) have g2:"0 \ r1 * r2" using "2" by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) next case 3 have g1:"ru1 * ru2 \ 0" using "3" geq1 geq2 by simp have g2:"0 \ r1 * r2" using "3" by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] repU_def tu.simps timesuu by (metis max.coboundedI1 max.commute maxt34) next case 4 have g1:"rl1 * rl2 \ rl1 * r2" using "4" geq1 geq2 grl1 grl2 gru1 gru2 using \rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2\ less_eq_real_def by (metis mult_left_mono_neg) have g2:"rl1 * r2 \ r1 * r2" using "4" geq1 geq2 grl1 grl2 gru1 gru2 \rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2\ by (metis mult_left_mono_neg mult.commute) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) qed next case ZeroPos assume bounds:"rl1 \ 0 \ 0 \ ru1 \ 0 \ rl2" have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast consider "r1 \ 0" | "r1 \ 0" using le_cases by (auto) then show ?thesis proof (cases) case 1 assume r1:"r1 \ 0" have g1:"ru1 * ru2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_left_mono by blast have g2:"ru1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by blast from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) next case 2 assume r1:"r1 \ 0" have g1:"ru1 * ru2 \ 0" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_left_mono by (simp add: mult_less_0_iff less_le_trans not_less) have g2:"0 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) qed next case ZeroNeg assume bounds:"rl1 \ 0 \ 0 \ ru1 \ ru2 \ 0" have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have case1:"r1 \ 0 \ ?thesis" proof - assume r1:"r1 \ 0" have g1:"rl1 * rl2 \ 0" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_less_0_iff less_le_trans not_less by metis have g2:"0 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by (simp add: mult_le_0_iff) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU2 max_repU1 repU_def timesll tu.simps) qed have case2:"r1 \ 0 \ ?thesis" proof - assume r1:"r1 \ 0" have g1:"rl1 * rl2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 by (metis mult_left_mono_neg) have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult.commute by (metis mult_left_mono_neg) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) qed show "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" using case1 case2 le_cases by blast next case PosZero assume bounds:"0 \ rl1 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"ru1 * ru2 \ ru1 * r2" using "1" bounds grl1 grl2 gru1 gru2 using mult_left_mono using leD leI less_le_trans by metis have g2:"ru1 * r2 \ r1 * r2" using "1" bounds grl1 grl2 gru1 gru2 using mult_right_mono by blast from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) next case 2 have g1:"ru1 * ru2 \ 0" using r1 bounds grl2 gru2 gru1 leD leI less_le_trans by auto have g2:"0 \ r1 * r2" using r1 "2" by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.coboundedI1 max.commute maxt34 by (metis repU_def tu.simps) qed next case NegZero assume bounds:"ru1 \ 0 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"ru1 * rl2 \ 0" using r1 "1" bounds grl1 grl2 gru1 gru2 mult_less_0_iff not_less by metis have g2:"0 \ r1 * r2" using r1 "1" bounds grl1 grl2 gru1 gru2 by (simp add: mult_le_0_iff) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesul tu.simps) next case 2 have lower:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have g1:"rl1 * rl2 \ rl1 * r2" using r1 "2" bounds grl1 grl2 gru1 gru2 less_eq(1) less_le_trans not_less mult_le_cancel_left by metis have g2:"rl1 * r2 \ r1 * r2" using r1 "2" bounds grl1 grl2 gru1 gru2 mult.commute not_le lower mult_le_cancel_left by metis from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) qed next case NegNeg assume bounds:"ru1 \ 0 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds dual_order.trans grl2 r2 by blast have g1:"rl1 * rl2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 less_eq(1) mult_le_cancel_left less_le_trans not_less by metis have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult.commute not_le lower1 lower2 mult_le_cancel_left by metis from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max_repU2 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] max.commute maxt34 by (metis max_repU1 repU_def timesll tu.simps) next case NegPos assume bounds:"ru1 \ 0 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using bounds by auto have upper2:"ru2 \ 0" using bounds dual_order.trans gru2 r2 by blast have g1:"ru1 * rl2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 not_less upper1 lower2 mult_le_cancel_left by metis have g2:"ru1 * r2 \ r1 * r2" using r1 upper1 r2 mult_right_mono gru1 by metis from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims maxt34 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] by (metis max_repU1 repU_def timesul tu.simps) next case PosNeg assume bounds:"0 \ rl1 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using dual_order.trans grl2 r2 by blast have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using bounds by auto have g1:"rl1 * ru2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 not_less upper2 lower1 mult_le_cancel_left by metis have g2:"rl1 * r2 \ r1 * r2" using r1 lower1 r2 not_less gru2 gru1 grl1 grl2 by (metis mult_le_cancel_left mult.commute) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show "tu l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>U r1 * r2" using up maxU12 maxU34 bigMaxU wmax.elims max.coboundedI1 max.commute maxt34 max_repU2[OF maxU12] max_repU2[OF maxU34] max_repU2[OF bigMaxU] by (metis repU_def tu.simps) next case PosPos assume bounds:"0 \ rl1 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using dual_order.trans gru2 u2 r2 bounds by blast have g1:"ru1 * ru2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_left_mono leD leI less_le_trans by metis have g2:"ru1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_right_mono by metis from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims max.coboundedI1 max.commute maxt34 max_repU2[OF bigMaxU] max_repU2[OF maxU12] max_repU2[OF maxU34] by (metis repU_def tu.simps) qed qed subsection \Minimum function\ text\Minimum of 2s-complement words\ fun wmin :: "word \ word \ word" where "wmin w1 w2 = (if w1 Correctness of wmin\ lemma wmin_lemma: assumes eq1:"w1 \\<^sub>E (r1::real)" assumes eq2:"w2 \\<^sub>E (r2::real)" shows "wmin w1 w2 \\<^sub>E (min r1 r2)" proof(cases rule: case_inf2[where ?w1.0=w1, where ?w2.0=w2]) case PosPos assume p1:"w1 = POS_INF" and p2:"w2 = POS_INF" then have bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2:"(real_of_int (sint POS_INF)) \ r2" using eq1 eq2 by (auto simp add: rep_simps repe.simps) have eqInf:"wmin w1 w2 = POS_INF" using p1 p2 unfolding wmin.simps by auto have pos_eq:"POS_INF \\<^sub>E min r1 r2" apply(rule repPOS_INF) using bound1 bound2 unfolding eq1 eq2 by auto show ?thesis using pos_eq eqInf by auto next case PosNeg assume p1:"w1 = POS_INF" assume n2:"w2 = NEG_INF" obtain r ra :: real where bound1:" (real_of_int (sint POS_INF)) \ r" and bound2:"ra \ (real_of_int (sint NEG_INF))" and eq1:"r1 = r" and eq2:"r2 = ra" using p1 n2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"wmin w1 w2 = NEG_INF" unfolding eq1 eq2 wmin.simps p1 n2 word_sless_def word_sle_def by(auto) have neg_eq:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using bound1 bound2 eq1 eq2 by auto show "?thesis" using eqNeg neg_eq by auto next case PosNum assume p1:"w1 = POS_INF" assume np2:"w2 \ POS_INF" assume nn2:"w2 \ NEG_INF" have eq2:"r2 = (real_of_int (sint w2))" and bound1:"(real_of_int (sint POS_INF)) \ r1" and bound2a:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound2b:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" using p1 np2 nn2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"min r1 r2 = sint w2" using p1 by (metis bound1 bound2b dual_order.trans eq2 min_def not_less) have neg_eq:"wmin w1 w2 \\<^sub>E (real_of_int (sint (wmin w1 w2)))" apply (rule repINT) using bound1 bound2a bound2b bound2b p1 unfolding eq1 eq2 by (auto simp add: word_sless_alt) show "?thesis" using eqNeg neg_eq by (metis bound2b less_eq_real_def not_less of_int_less_iff p1 wmin.simps word_sless_alt) next case NegPos assume n1:"w1 = NEG_INF" assume p2:"w2 = POS_INF" have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:"(real_of_int (sint POS_INF)) \ r2" using n1 p2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"wmin w1 w2 = NEG_INF" unfolding eq1 eq2 wmin.simps n1 p2 word_sless_def word_sle_def by(auto) have neg_eq:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using bound1 bound2 unfolding eq1 eq2 by auto show "wmin w1 w2 \\<^sub>E min r1 r2" using eqNeg neg_eq by auto next case NegNeg assume n1:"w1 = NEG_INF" assume n2:"w2 = NEG_INF" have bound1:"r1 \ (real_of_int (sint NEG_INF))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" using n1 n2 eq1 eq2 by(auto simp add: rep_simps repe.simps) have eqNeg:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using eq1 eq2 bound1 bound2 unfolding NEG_INF_def by (auto) have neg_eq:"wmin w1 w2 = NEG_INF" using n1 n2 unfolding NEG_INF_def wmin.simps by auto show "wmin w1 w2 \\<^sub>E min r1 r2" using eqNeg neg_eq by auto next case NegNum assume n1:"w1 = NEG_INF" and nn2:"w2 \ NEG_INF" and np2:"w2 \ POS_INF" have eq2:"r2 = (real_of_int (sint w2))" and bound2a:"(real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:"(real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" and bound1:"r1 \ (real_of_int (sint NEG_INF))" using n1 nn2 np2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have eqNeg:"wmin w1 w2 = NEG_INF" using n1 assms(2) bound2a eq2 n1 repeInt_simps by (auto simp add: word_sless_alt) have neg_eq:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using bound1 bound2a bound2b eq1 min_le_iff_disj by blast show "wmin w1 w2 \\<^sub>E min r1 r2" using eqNeg neg_eq by auto next case NumPos assume p2:"w2 = POS_INF" and nn1:"w1 \ NEG_INF" and np1:"w1 \ POS_INF" have eq1:"r1 = (real_of_int (sint w1))" and bound1a:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:" (real_of_int (sint POS_INF)) \ r2" using nn1 np1 p2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have res1:"wmin w1 w2 = w1" using p2 eq1 eq2 assms(1) bound1b p2 repeInt_simps by (auto simp add: word_sless_alt) have res2:"min r1 r2 = (real_of_int (sint w1))" using eq1 eq2 bound1a bound1b bound2 - by (auto simp add: less_imp_le less_le_trans min_def) + by transfer (auto simp add: less_imp_le less_le_trans min_def) have res3:"wmin w1 w2 \\<^sub>E (real_of_int (sint (wmin w1 w2)))" apply(rule repINT) using p2 bound1a res1 bound1a bound1b bound2 by auto show "wmin w1 w2 \\<^sub>E min r1 r2" using res1 res2 res3 by auto next case NumNeg assume nn1:"w1 \ NEG_INF" assume np1:"w1 \ POS_INF" assume n2:"w2 = NEG_INF" have eq1:"r1 = (real_of_int (sint w1))" and bound1a:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2:"r2 \ (real_of_int (sint NEG_INF))" using nn1 np1 n2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have res1:"wmin w1 w2 = NEG_INF" using n2 bound1b by (metis min.absorb_iff2 min_def n2 not_less of_int_less_iff wmin.simps word_sless_alt) have res2:"NEG_INF \\<^sub>E min r1 r2" apply(rule repNEG_INF) using eq1 eq2 bound1a bound1b bound2 min_le_iff_disj by blast show "wmin w1 w2 \\<^sub>E min r1 r2" using res1 res2 by auto next case NumNum assume np1:"w1 \ POS_INF" assume nn1:"w1 \ NEG_INF" assume np2:"w2 \ POS_INF" assume nn2:"w2 \ NEG_INF" have eq1:"r1 = (real_of_int (sint w1))" and eq2:"r2 = (real_of_int (sint w2))" and bound1a:" (real_of_int (sint w1)) < (real_of_int (sint POS_INF))" and bound1b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w1))" and bound2a:" (real_of_int (sint w2)) < (real_of_int (sint POS_INF))" and bound2b:" (real_of_int (sint NEG_INF)) < (real_of_int (sint w2))" using nn1 np1 nn2 np2 eq2 eq1 eq2 by (auto simp add: rep_simps repe.simps) have res1:"min r1 r2 = (real_of_int (sint (wmin w1 w2)))" using eq1 eq2 bound1a bound1b bound2a bound2b - by (simp add: min_def word_sless_alt) + apply (simp add: min_def word_sless_alt not_less) + apply transfer + apply simp + done have res2:"wmin w1 w2 \\<^sub>E (real_of_int (sint (wmin w1 w2)))" apply (rule repINT) using bound1a bound1b bound2a bound2b by (simp add: \min r1 r2 = (real_of_int (sint (wmin w1 w2)))\ eq2 min_less_iff_disj)+ show "wmin w1 w2 \\<^sub>E min r1 r2" using res1 res2 by auto qed lemma min_repU1: assumes "w1 \\<^sub>L x" assumes "w2 \\<^sub>L y" shows "wmin w1 w2 \\<^sub>L x " using wmin_lemma assms repL_def by (meson min_le_iff_disj) lemma min_repU2: assumes "w1 \\<^sub>L y" assumes "w2 \\<^sub>L x" shows "wmin w1 w2 \\<^sub>L x" using wmin_lemma assms repL_def by (meson min_le_iff_disj) subsection \Multiplication lower bound\ text\Multiplication lower bound\ fun tl :: "word \ word \ word \ word \ word" where "tl w1l w1u w2l w2u = wmin (wmin (wtimes w1l w2l) (wtimes w1u w2l)) (wmin (wtimes w1l w2u) (wtimes w1u w2u))" text\Correctness of multiplication lower bound\ lemma tl_lemma: assumes u1:"u\<^sub>1 \\<^sub>U (r1::real)" assumes u2:"u\<^sub>2 \\<^sub>U (r2::real)" assumes l1:"l\<^sub>1 \\<^sub>L (r1::real)" assumes l2:"l\<^sub>2 \\<^sub>L (r2::real)" shows "tl l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>L (r1 * r2)" proof - obtain rl1 rl2 ru1 ru2 :: real where gru1:"ru1 \ r1" and gru2:"ru2 \ r2" and grl1:"rl1 \ r1" and grl2:"rl2 \ r2" and eru1:"u\<^sub>1 \\<^sub>E ru1" and eru2:"u\<^sub>2 \\<^sub>E ru2" and erl1:"l\<^sub>1 \\<^sub>E rl1" and erl2:"l\<^sub>2 \\<^sub>E rl2" using u1 u2 l1 l2 unfolding repU_def repL_def by auto have timesuu:"wtimes u\<^sub>1 u\<^sub>2 \\<^sub>E ru1 * ru2" using wtimes_exact[OF eru1 eru2] by auto have timesul:"wtimes u\<^sub>1 l\<^sub>2 \\<^sub>E ru1 * rl2" using wtimes_exact[OF eru1 erl2] by auto have timeslu:"wtimes l\<^sub>1 u\<^sub>2 \\<^sub>E rl1 * ru2" using wtimes_exact[OF erl1 eru2] by auto have timesll:"wtimes l\<^sub>1 l\<^sub>2 \\<^sub>E rl1 * rl2" using wtimes_exact[OF erl1 erl2] by auto have maxt12:"wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>E min (rl1 * rl2) (ru1 * rl2)" by (rule wmin_lemma[OF timesll timesul]) have maxt34:"wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>E min (rl1 * ru2) (ru1 * ru2)" by (rule wmin_lemma[OF timeslu timesuu]) have bigMax:"wmin (wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>E min (min(rl1 * rl2) (ru1 * rl2)) (min (rl1 * ru2) (ru1 * ru2))" by (rule wmin_lemma[OF maxt12 maxt34]) obtain maxt12val :: real where maxU12:"wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2) \\<^sub>L min (rl1 * rl2) (ru1 * rl2)" using maxt12 unfolding repL_def by blast obtain maxt34val :: real where maxU34:"wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2) \\<^sub>L min (rl1 * ru2) (ru1 * ru2)" using maxt34 unfolding repL_def by blast obtain bigMaxU:"wmin (wmin (wtimes l\<^sub>1 l\<^sub>2) (wtimes u\<^sub>1 l\<^sub>2)) (wmin (wtimes l\<^sub>1 u\<^sub>2) (wtimes u\<^sub>1 u\<^sub>2)) \\<^sub>L min (min (rl1 * rl2) (ru1 * rl2)) (min (rl1 * ru2) (ru1 * ru2))" using bigMax unfolding repL_def by blast have ivl1:"rl1 \ ru1" using grl1 gru1 by auto have ivl2:"rl2 \ ru2" using grl2 gru2 by auto let ?thesis = "tl l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>L r1 * r2" show ?thesis using ivl1 ivl2 proof(cases rule: case_ivl_zero) case ZeroZero assume "rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2" then have geq1:"ru1 \ 0" and geq2:"ru2 \ 0" and geq3:"rl1 \ 0" and geq4:"rl2 \ 0" by auto consider "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" | "r1 \ 0 \ r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"rl1 * ru2 \ 0" using "1" geq1 geq2 geq3 geq4 grl2 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using "1" geq1 geq2 grl1 grl2 gru1 gru2 by (simp) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis using up eru1 eru2 erl1 erl2 min_repU1 min_repU2 repL_def repU_def timeslu tl.simps wmin.elims by (metis bigMax min_le_iff_disj) next case 2 have g1:"rl1 * ru2 \ rl1 * r2" using "2" geq1 geq2 grl2 gru2 by (metis mult_le_cancel_left geq3 leD) have g2:"rl1 * r2 \ r1 * r2" using "2" geq1 geq2 grl2 gru2 by (simp add: mult_right_mono grl1) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 min_repU2 repL_def tl.simps min.coboundedI1 maxt34) next case 3 have g1:"ru1 * rl2 \ ru1 * r2" using "3" geq1 geq2 grl2 gru2 by (simp add: mult_left_mono) have g2:"ru1 * r2 \ r1 * r2" using "3" geq1 geq2 grl1 grl2 gru1 gru2 mult_minus_right mult_right_mono by (simp add: mult_right_mono_neg) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 maxt34 timesul by (metis repL_def tl.simps) next case 4 have g1:"ru1 * rl2 \ 0" using "4" geq1 geq2 grl1 grl2 gru1 gru2 \rl1 \ 0 \ 0 \ ru1 \ rl2 \ 0 \ 0 \ ru2\ mult_less_0_iff less_eq_real_def not_less by auto have g2:"0 \ r1 * r2" using "4" geq1 geq2 grl1 grl2 gru1 gru2 by (metis mult_less_0_iff not_less) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 maxU34 wmin.elims min_repU1 min_repU2 repL_def timesul tl.simps) qed next case ZeroPos assume bounds:"rl1 \ 0 \ 0 \ ru1 \ 0 \ rl2" have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast consider "r1 \ 0" | "r1 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"rl1 * rl2 \ 0" using "1" r2 bounds grl1 grl2 gru1 gru2 by (simp add: mult_le_0_iff) have g2:"0 \ r1 * r2" using "1" r2 bounds grl1 grl2 gru1 gru2 by (simp) from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis by (metis repL_def timesll tl.simps up maxU12 maxU34 wmin.elims min_repU2 min_repU1) next case 2 have bound:"ru2 \ 0" using "2" r2 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto then have g1:"rl1 * ru2 \ rl1 * r2" using "2" r2 bounds grl1 grl2 gru1 gru2 mult_le_cancel_left by fastforce have g2:"rl1 * r2 \ r1 * r2" using "2" r2 bounds grl1 grl2 gru1 gru2 mult_le_0_iff mult_le_cancel_right by fastforce from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 wmin.elims min_repU2 min.coboundedI1 maxt34 repL_def tl.simps) qed next case ZeroNeg assume bounds:"rl1 \ 0 \ 0 \ ru1 \ ru2 \ 0" have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast consider (Pos) "r1 \ 0" | (Neg) "r1 \ 0" using le_cases by auto then show ?thesis proof (cases) case Pos have bound:"rl2 \ 0" using Pos r2 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto then have g1:"ru1 * rl2 \ ru1 * r2" using Pos bounds grl1 grl2 gru1 gru2 mult_le_cancel_left by fastforce have p1:"\a::real. (0 \ - a) = (a \ 0)" by(auto) have p2:"\a b::real. (- a \ - b) = (b \ a)" by auto have g2:"ru1 * r2 \ r1 * r2" using Pos r2 bounds grl1 grl2 gru1 gru2 p1 p2 by (simp add: mult_right_mono_neg) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis by (metis up maxU12 maxU34 wmin.elims min_repU2 min_repU1 repL_def timesul tl.simps) next case Neg have g1:"ru1 * ru2 \ 0" using Neg r2 bounds grl1 grl2 gru1 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using Neg r2 zero_le_mult_iff by blast from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt34 by (metis repL_def tl.simps) qed next case PosZero assume bounds:"0 \ rl1 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have bound:"0 \ ru1" using r1 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 have g1:"rl1 * rl2 \ 0" using r1 "1" bounds grl1 grl2 gru1 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using r1 "1" bounds grl1 grl2 gru1 gru2 zero_le_mult_iff by blast from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmax.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt12 maxt34 repL_def timesll tl.simps by metis next case 2 have g1:"ru1 * rl2 \ ru1 * r2" using r1 "2" bounds bound grl1 grl2 gru1 gru2 using mult_left_mono by blast have g2:"ru1 * r2 \ r1 * r2" using r1 "2" bounds bound grl2 gru2 by (metis mult_left_mono_neg gru1 mult.commute) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 maxt34 by (metis repL_def timesul tl.simps) qed next case NegZero assume bounds:"ru1 \ 0 \ rl2 \ 0 \ 0 \ ru2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have bound:"rl1 \ 0" using r1 bounds grl1 grl2 gru1 gru2 dual_order.trans by auto consider "r2 \ 0" | "r2 \ 0" using le_cases by auto then show ?thesis proof (cases) case 1 assume r2:"r2 \ 0" have g1:"rl1 * ru2 \ rl1 * r2" using r1 r2 bounds bound grl1 grl2 gru1 gru2 by (metis mult_le_cancel_left leD) have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_right_mono by (simp add: mult_le_0_iff) from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU min_repU2 min_repU1 min.coboundedI1 maxt34 by (metis min_repU2 repL_def tl.simps) next case 2 assume r2:"r2 \ 0" have lower:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have g1:"ru1 * ru2 \ 0" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_le_0_iff by blast have g2:"0 \ r1 * r2" using r1 r2 by (simp add: zero_le_mult_iff) from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt34 by (metis repL_def tl.simps) qed next case NegNeg assume bounds:"ru1 \ 0 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds dual_order.trans grl2 r2 by blast have g1:"ru1 * ru2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using not_less mult_le_cancel_left by metis have g2:"ru1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_le_cancel_left mult.commute not_le lower1 lower2 by metis from g1 and g2 have up:"ru1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 min.coboundedI1 min.commute maxt34 by (metis repL_def tl.simps) next case NegPos assume bounds:"ru1 \ 0 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans gru1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds dual_order.trans grl1 r1 by blast have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using bounds by auto have upper2:"ru2 \ 0" using bounds dual_order.trans gru2 r2 by blast have g1:"rl1 * ru2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 less_le_trans upper1 lower2 by (metis mult_le_cancel_left not_less) have g2:"rl1 * r2 \ r1 * r2" using r1 upper1 r2 mult_right_mono mult_le_0_iff grl1 by blast from g1 and g2 have up:"rl1 * ru2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 maxt12 maxt34 by (metis repL_def timeslu tl.simps) next case PosNeg assume bounds:"0 \ rl1 \ ru2 \ 0" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans gru2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using dual_order.trans grl2 r2 by blast have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using bounds by auto have g1:"ru1 * rl2 \ ru1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 mult_left_mono less_le_trans not_less by metis have g2:"ru1 * r2 \ r1 * r2" using r1 lower1 r2 not_less gru2 gru1 grl1 grl2 by (metis mult_le_cancel_left mult.commute) from g1 and g2 have up:"ru1 * rl2 \ r1 * r2" by auto show "tl l\<^sub>1 u\<^sub>1 l\<^sub>2 u\<^sub>2 \\<^sub>L r1 * r2" using up maxU12 maxU34 bigMaxU wmin.elims min_repU2 min_repU1 by (metis repL_def timesul tl.simps) next case PosPos assume bounds:"0 \ rl1 \ 0 \ rl2" have r1:"r1 \ 0" using bounds dual_order.trans grl1 by blast have r2:"r2 \ 0" using bounds dual_order.trans grl2 by blast have lower1:"rl1 \ 0" using bounds by auto have lower2:"rl2 \ 0" using bounds by auto have upper1:"ru1 \ 0" using dual_order.trans gru1 u1 r1 by blast have upper2:"ru2 \ 0" using dual_order.trans gru2 u2 r2 bounds by blast have g1:"rl1 * rl2 \ rl1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_left_mono using leD leI less_le_trans by auto have g2:"rl1 * r2 \ r1 * r2" using r1 r2 bounds grl1 grl2 gru1 gru2 using mult_right_mono by blast from g1 and g2 have up:"rl1 * rl2 \ r1 * r2" by auto show ?thesis using up maxU12 maxU34 bigMaxU min_repU2 min_repU1 min.coboundedI1 maxt12 maxt34 by (metis repL_def tl.simps) qed qed text\Most significant bit only changes under successor when all other bits are 1\ lemma msb_succ: fixes w :: "32 Word.word" assumes neq1:"uint w \ 0xFFFFFFFF" assumes neq2:"uint w \ 0x7FFFFFFF" shows "msb (w + 1) = msb w" proof - have "w \ 0xFFFFFFFF" using neq1 by auto then have neqneg1:"w \ -1" by auto have "w \ 0x7FFFFFFF" using neq2 by auto then have neqneg2:"w \ (2^31)-1" by auto show ?thesis using neq1 neq2 unfolding msb_big using Word_Lemmas.word_le_make_less[of "w + 1" "0x80000000"] Word_Lemmas.word_le_make_less[of "w " "0x80000000"] neqneg1 neqneg2 by auto qed text\Negation commutes with msb except at edge cases\ lemma msb_non_min: fixes w :: "32 Word.word" assumes neq1:"uint w \ 0" assumes neq2:"uint w \ ((2^(len_of (TYPE(31)))))" shows "msb (uminus w) = HOL.Not(msb(w))" proof - have fact1:"uminus w = word_succ (~~ w)" by (rule twos_complement) have fact2:"msb (~~w) = HOL.Not(msb w)" using word_ops_msb[of w] by auto have neqneg1:"w \ 0" using neq1 by auto have not_undef:"w \ 0x80000000" using neq2 by auto then have neqneg2:"w \ (2^31)" by auto from \w \ 0\ have \~~ w \ ~~ 0\ by (simp only: bit.compl_eq_compl_iff) simp then have "(~~ w) \ 0xFFFFFFFF" by auto then have uintNeq1:"uint (~~ w) \ 0xFFFFFFFF" using uint_distinct[of "~~w" "0xFFFFFFFF"] by auto from \w \ 2 ^ 31\ have \~~ w \ ~~ 2 ^ 31\ by (simp only: bit.compl_eq_compl_iff) simp then have "(~~ w) \ 0x7FFFFFFF" by auto then have uintNeq2:" uint (~~ w) \ 0x7FFFFFFF" using uint_distinct[of "~~w" "0x7FFFFFFF"] by auto have fact3:"msb ((~~w) + 1) = msb (~~w)" apply(rule msb_succ[of "~~w"]) using neq1 neq2 uintNeq1 uintNeq2 by auto show "msb (uminus w) = HOL.Not(msb(w))" using fact1 fact2 fact3 by (simp add: word_succ_p1) qed text\Only 0x80000000 preserves msb=1 under negation\ lemma msb_min_neg: fixes w::"word" assumes msb1:"msb (- w)" assumes msb2:"msb w" shows "uint w = ((2^(len_of (TYPE(31)))))" proof (rule ccontr) from \msb w\ have \w \ 0\ using word_msb_0 by auto then have \uint w \ 0\ by transfer simp moreover assume \uint w \ 2 ^ LENGTH(31)\ ultimately have \msb (- w) \ \ msb w\ by (rule msb_non_min) with assms show False by simp qed text\Only 0x00000000 preserves msb=0 under negation\ lemma msb_zero: fixes w::"word" assumes msb1:"\ msb (- w)" assumes msb2:"\ msb w" shows "uint w = 0" proof - have neq:"w \ ((2 ^ len_of TYPE(31))::word)" using msb1 msb2 by auto have eq:"uint ((2 ^ len_of TYPE(31))::word) = 2 ^ len_of TYPE(31)" by auto then have neq:"uint w \ uint ((2 ^ len_of TYPE(31))::word)" using uint_distinct[of w "2^len_of TYPE(31)"] neq eq by auto show ?thesis using msb1 msb2 minus_zero msb_non_min[of w] neq by force qed text\Finite numbers alternate msb under negation\ lemma msb_pos: fixes w::"word" assumes msb1:"msb (- w)" assumes msb2:"\ msb w" shows "uint w \ {1 .. (2^((len_of TYPE(32)) - 1))-1}" proof - have main: "w \ {1 .. (2^((len_of TYPE(32)) - 1))-1}" using msb1 msb2 apply(clarsimp) unfolding word_msb_sint apply(rule conjI) apply (metis neg_equal_0_iff_equal not_le word_less_1) proof - have imp:"w \ 0x80000000 \ False" proof - assume geq:"w \ 0x80000000" then have "msb w" using Word_Lemmas.msb_big[of w] by auto then show False using msb2 by auto qed have mylem:"\w1 w2::word. uint w1 \ uint w2 \ w1 \ w2" subgoal for w1 w2 by (simp add: word_le_def) done have mylem2:"\w1 w2::word. w1 > w2 \ uint w1 > uint w2" subgoal for w1 w2 by (simp add: word_less_def) done have gr_to_geq:"w > 0x7FFFFFFF \ w \ 0x80000000" apply(rule mylem) using mylem2[of "0x7FFFFFFF" "w"] by auto have taut:"w \ 0x7FFFFFFF \ w > 0x7FFFFFFF" by auto then show "w \ 0x7FFFFFFF" using imp taut gr_to_geq by auto qed have set_eq:"(uint ` (({1..(minus(2 ^ (minus(len_of TYPE(32)) 1)) 1)})::word set)) = ({1..minus(2 ^ (minus (len_of TYPE(32)) 1)) 1}::int set)" apply(auto simp add: word_le_def) subgoal for xa proof - assume lower:"1 \ xa" and upper:"xa \ 2147483647" then have in_range:"xa \ {0 .. 2^32-1}" by auto then have "xa \ range (uint::word \ int)" unfolding Word.word_uint.Rep_range Word.uints_num by auto then obtain w::word where xaw:"xa = uint w" by auto then have "w \ {1..0x7FFFFFFF} " using lower upper apply(clarsimp, auto) by (auto simp add: word_le_def) then show ?thesis using uint_distinct uint_distinct main image_eqI word_le_def xaw by blast qed done then show "uint w \ {1..2 ^ (len_of TYPE(32) - 1) - 1}" using uint_distinct uint_distinct main image_eqI by blast qed lemma msb_neg: fixes w::"word" assumes msb1:"\ msb (- w)" assumes msb2:"msb w" shows "uint w \ {2^((len_of TYPE(32) - 1))+1 .. 2^((len_of TYPE(32)))-1}" proof - have mylem:"\w1 w2::word. uint w1 \ uint w2 \ w1 \ w2" by (simp add: word_le_def) have mylem2:"\w1 w2::word. w1 > w2 \ uint w1 > uint w2" by (simp add: word_less_def) have gr_to_geq:"w > 0x80000000 \ w \ 0x80000001" apply(rule mylem) using mylem2[of "0x80000000" "w"] by auto have taut:"w \ 0x80000000 \ 0x80000000 < w" by auto have imp:"w \ 0x80000000 \ False" proof - assume geq:"w \ 0x80000000" then have "(msb (-w))" using Word_Lemmas.msb_big[of "-w"] Word_Lemmas.msb_big[of "w"] by (simp add: msb2) then show False using msb1 by auto qed have main: "w \ {2^((len_of TYPE(32)) - 1)+1 .. 2^((len_of TYPE(32)))-1}" using msb1 msb2 apply(clarsimp) unfolding word_msb_sint proof - show "0x80000001 \ w" using imp taut gr_to_geq by auto qed have set_eq:"(uint ` (({2^((len_of TYPE(32) - 1))+1 .. 2^((len_of TYPE(32)))-1})::word set)) = {2^((len_of TYPE(32) - 1))+1 .. 2^((len_of TYPE(32)))-1}" apply(auto) subgoal for xa by (simp add: word_le_def) subgoal for w using uint_lt [of w] by simp subgoal for xa proof - assume lower:"2147483649 \ xa" and upper:"xa \ 4294967295" then have in_range:"xa \ {0x80000000 .. 0xFFFFFFFF}" by auto then have "xa \ range (uint::word \ int)" unfolding Word.word_uint.Rep_range Word.uints_num by auto then obtain w::word where xaw:"xa = uint w" by auto then have the_in:"w \ {0x80000001 .. 0xFFFFFFFF} " using lower upper by (auto simp add: word_le_def) have the_eq:"(0xFFFFFFFF::word) = -1" by auto from the_in the_eq have "w \ {0x80000001 .. -1}" by auto then show ?thesis using uint_distinct uint_distinct main image_eqI word_le_def xaw by blast qed done then show "uint w \ {2^((len_of TYPE(32)) - 1)+1 .. 2^((len_of TYPE(32)))-1}" using uint_distinct uint_distinct main image_eqI by blast qed text\2s-complement commutes with negation except edge cases\ lemma sint_neg_hom: fixes w :: "32 Word.word" shows "uint w \ ((2^(len_of (TYPE(31))))) \ (sint(-w) = -(sint w))" unfolding word_sint_msb_eq apply auto subgoal using msb_min_neg by auto prefer 3 subgoal using msb_zero[of w] by (simp add: msb_zero) proof - assume msb1:"msb (- w)" assume msb2:"\ msb w" have "uint w \ {1 .. (2^((len_of TYPE(32)) - 1))-1}" using msb_pos[OF msb1 msb2] by auto then have bound:"uint w \ {1 .. 0x7FFFFFFF}" by auto have size:"size (w::32 Word.word) = 32" using Word.word_size[of w] by auto have lem:"\x::int. \n::nat. x \ {1..(2^n)-1} \ ((- x) mod (2^n)) - (2^n) = - x" subgoal for x n apply(cases "x mod 2^n = 0") by(auto simp add: Divides.zmod_zminus1_eq_if[of x "2^n"]) done have lem_rule:"uint w \ {1..2 ^ 32 - 1} \ (- uint w mod 4294967296) - 4294967296 = - uint w" using lem[of "uint w" 32] by auto have almost:"- uint w mod 4294967296 - 4294967296 = - uint w" apply(rule lem_rule) using bound by auto show "uint (- w) - 2 ^ size (- w) = - uint w" using bound unfolding Word.uint_word_ariths word_size_neg by (auto simp add: size almost) next assume neq:"uint w \ 0x80000000" assume msb1:"\ msb (- w)" assume msb2:"msb w" have bound:"uint w \ {0x80000001.. 0xFFFFFFFF}" using msb1 msb2 msb_neg by auto have size:"size (w::32 Word.word) = 32" using Word.word_size[of w] by auto have lem:"\x::int. \n::nat. x \ {1..(2^n)-1} \ (-x mod (2^n)) = (2^n) - x" subgoal for x n apply(auto) apply(cases "x mod 2^n = 0") by (simp add: Divides.zmod_zminus1_eq_if[of x "2^n"])+ done from bound have wLeq: "uint w \ 4294967295" and wGeq: "2147483649 \ uint w" by auto from wLeq have wLeq':"uint w \ 4294967296" by fastforce have f3: "(0 \ 4294967296 + - 1 * uint w + - 1 * ((4294967296 + - 1 * uint w) mod 4294967296)) = (uint w + (4294967296 + - 1 * uint w) mod 4294967296 \ 4294967296)" by auto have f4: "(0 \ 4294967296 + - 1 * uint w) = (uint w \ 4294967296)" by auto have f5: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * (i mod ia)" by (simp add: zmod_le_nonneg_dividend) then have f6: "uint w + (4294967296 + - 1 * uint w) mod 4294967296 \ 4294967296" using f4 f3 wLeq' by blast have f7: "4294967296 + - 1 * uint w + - 4294967296 = - 1 * uint w" by auto have f8: "- (1::int) * 4294967296 = - 4294967296" by auto have f9: "(0 \ - 1 * uint w) = (uint w \ 0)" by auto have f10: "(4294967296 + -1 * uint w + -1 * ((4294967296 + -1 * uint w) mod 4294967296) \ 0) = (4294967296 \ uint w + (4294967296 + - 1 * uint w) mod 4294967296)" by auto have f11: "\ 4294967296 \ (0::int)" by auto have f12: "\x0. ((0::int) < x0) = (\ x0 \ 0)" by auto have f13: "\x0 x1. ((x1::int) < x0) = (\ 0 \ x1 + - 1 * x0)" by auto have f14: "\x0 x1. ((x1::int) \ x1 mod x0) = (x1 + - 1 * (x1 mod x0) \ 0)" by auto have "\ uint w \ 0" using wGeq by fastforce then have "4294967296 \ uint w + (4294967296 + - 1 * uint w) mod 4294967296" using f14 f13 f12 f11 f10 f9 f8 f7 by (metis (no_types) int_mod_ge) then show "uint (- w) = 2 ^ size w - uint w" using f6 unfolding Word.uint_word_ariths by (auto simp add: size f4) qed text\2s-complement encoding is injective\ lemma sint_dist: fixes x y ::word assumes "x \ y" shows "sint x \ sint y" by (simp add: assms) subsection\Negation\ fun wneg :: "word \ word" where "wneg w = (if w = NEG_INF then POS_INF else if w = POS_INF then NEG_INF else -w)" text\word negation is correct\ lemma wneg_lemma: assumes eq:"w \\<^sub>E (r::real)" shows "wneg w \\<^sub>E -r" apply(rule repe.cases[OF eq]) apply(auto intro!: repNEG_INF repPOS_INF simp add: repe.simps)[2] subgoal for ra proof - assume eq:"w = ra" assume i:"r = (real_of_int (sint ra))" assume bounda:" (real_of_int (sint ra)) < (real_of_int (sint POS_INF))" assume boundb:" (real_of_int (sint NEG_INF)) < (real_of_int (sint ra))" have raNeq:"ra \ 2147483647" using sint_range[OF bounda boundb] by (auto) have raNeqUndef:"ra \ 2147483648" using int_not_undef[OF bounda boundb] by (auto) have "uint ra \ uint ((2 ^ len_of TYPE(31))::word)" apply (rule uint_distinct) using raNeqUndef by auto then have raNeqUndefUint:"uint ra \ ((2 ^ len_of TYPE(31)))" by auto have res1:"wneg w \\<^sub>E (real_of_int (sint (wneg w)))" apply (rule repINT) using sint_range[OF bounda boundb] sint_neg_hom[of ra, OF raNeqUndefUint] raNeq raNeqUndefUint raNeqUndef eq by(auto) have res2:"- r = (real_of_int (sint (wneg w)))" using eq bounda boundb i sint_neg_hom[of ra, OF raNeqUndefUint] raNeq raNeqUndef eq - by (auto) + apply auto + apply transfer + apply simp + done show ?thesis using res1 res2 by auto qed done subsection\Comparison\ fun wgreater :: "word \ word \ bool" where "wgreater w1 w2 = (sint w1 > sint w2)" lemma neg_less_contra:"\x. Suc x < - (Suc x) \ False" by auto text\Comparison < is correct\ lemma wgreater_lemma:"w1 \\<^sub>L (r1::real) \ w2 \\<^sub>U r2 \ wgreater w1 w2 \ r1 > r2" proof (auto simp add: repU_def repL_def) fix r'\<^sub>1 r'\<^sub>2 assume sint_le:"sint w1 > sint w2" then have sless:"(w2 1 \ r1" assume r2_leq:"r2 \ r'\<^sub>2" assume wr1:"w1 \\<^sub>E r'\<^sub>1" assume wr2:"w2 \\<^sub>E r'\<^sub>2" have greater:"r'\<^sub>1 > r'\<^sub>2" using wr1 wr2 apply(auto simp add: repe.simps) - prefer 4 using sless sint_le by (auto simp add: less_le_trans not_le) + prefer 4 using sless sint_le + apply (auto simp add: less_le_trans not_le) + apply transfer apply simp + apply transfer apply simp + apply transfer apply simp + done show "r1 > r2" using r1_leq r2_leq greater by auto qed text\Comparison $\geq$ of words\ fun wgeq :: "word \ word \ bool" where "wgeq w1 w2 = ((\ ((w2 = NEG_INF \ w1 = NEG_INF) \(w2 = POS_INF \ w1 = POS_INF))) \ (sint w2 \ sint w1))" text\Comparison $\geq$ of words is correct\ lemma wgeq_lemma:"w1 \\<^sub>L r1 \ w2 \\<^sub>U (r2::real) \ wgeq w1 w2 \ r1 \ r2" proof (unfold wgeq.simps) assume assms:"\ (w2 = NEG_INF \ w1 = NEG_INF \ w2 = POS_INF \ w1 = POS_INF) \ sint w2 \ sint w1" assume a1:"w1 \\<^sub>L r1" and a2:"w2 \\<^sub>U (r2::real)" from assms have sint_le:"sint w2 \ sint w1" by auto then have sless:"(w2 <=s w1)" using word_sless_alt word_sle_def by auto obtain r'\<^sub>1 r'\<^sub>2 where r1_leq:"r'\<^sub>1 \ r1" and r2_leq:"r2 \ r'\<^sub>2" and wr1:"w1 \\<^sub>E r'\<^sub>1" and wr2:"w2 \\<^sub>E r'\<^sub>2" using a1 a2 unfolding repU_def repL_def by auto from assms have check1:"\ (w1 = NEG_INF \ w2 = NEG_INF)" by auto from assms have check2:"\ (w1 = POS_INF \ w2 = POS_INF)" by auto have less:"r'\<^sub>2 \ r'\<^sub>1" using sless sint_le check1 check2 repe.simps wr2 wr1 - by(auto simp add: repe.simps) + apply (auto simp add: repe.simps) + apply transfer + apply simp + apply transfer + apply simp + apply transfer + apply simp + apply transfer + apply simp + apply transfer + apply simp + apply transfer + apply simp + apply transfer + apply simp + apply transfer + apply simp + done show "r1 \ r2" using r1_leq r2_leq less by auto qed subsection\Absolute value\ text\Absolute value of word\ fun wabs :: "word \ word" where "wabs l1 = (wmax l1 (wneg l1))" text\Correctness of wmax\ lemma wabs_lemma: assumes eq:"w \\<^sub>E (r::real)" shows "wabs w \\<^sub>E (abs r)" proof - have w:"wmax w (wneg w) \\<^sub>E max r (-r)" by (rule wmax_lemma[OF eq wneg_lemma[OF eq]]) have r:"max r (-r) = abs r" by auto from w r show ?thesis by auto qed -end \ No newline at end of file + +declare more_real_of_word_simps [simp del] + +end diff --git a/thys/Iptables_Semantics/Common/Word_Upto.thy b/thys/Iptables_Semantics/Common/Word_Upto.thy --- a/thys/Iptables_Semantics/Common/Word_Upto.thy +++ b/thys/Iptables_Semantics/Common/Word_Upto.thy @@ -1,199 +1,205 @@ section\Word Upto\ theory Word_Upto imports Main IP_Addresses.Hs_Compat Word_Lib.Word_Lemmas begin text\Enumerate a range of machine words.\ text\enumerate from the back (inefficient)\ function word_upto :: "'a word \ 'a word \ ('a::len) word list" where "word_upto a b = (if a = b then [a] else word_upto a (b - 1) @ [b])" by pat_completeness auto (*by the way: does not terminate practically if b < a; will terminate after it reaches the word wrap-around!*) termination word_upto apply(relation "measure (unat \ uncurry (-) \ prod.swap)") apply(rule wf_measure; fail) apply(simp) apply(subgoal_tac "unat (b - a - 1) < unat (b - a)") apply(simp add: diff_right_commute; fail) apply(rule measure_unat) apply auto done declare word_upto.simps[simp del] text\enumerate from the front (more inefficient)\ function word_upto' :: "'a word \ 'a word \ ('a::len) word list" where "word_upto' a b = (if a = b then [a] else a # word_upto' (a + 1) b)" by pat_completeness auto termination word_upto' apply(relation "measure (\ (a, b). unat (b - a))") apply(rule wf_measure; fail) apply(simp) apply(subgoal_tac "unat (b - a - 1) < unat (b - a)") apply (simp add: diff_diff_add; fail) apply(rule measure_unat) apply auto done declare word_upto'.simps[simp del] lemma word_upto_cons_front[code]: "word_upto a b = word_upto' a b" proof(induction a b rule:word_upto'.induct) case (1 a b) have hlp1: "a \ b \ a # word_upto (a + 1) b = word_upto a b" apply(induction a b rule:word_upto.induct) apply simp apply(subst(1) word_upto.simps) apply(simp) apply safe apply(subst(1) word_upto.simps) apply (simp) apply(subst(1) word_upto.simps) apply (simp; fail) apply(case_tac "a \ b - 1") apply(simp) apply (metis Cons_eq_appendI word_upto.simps) apply(simp) done from 1[symmetric] show ?case apply(cases "a = b") subgoal apply(subst word_upto.simps) apply(subst word_upto'.simps) by(simp) apply(subst word_upto'.simps) by(simp add: hlp1) qed (* Most of the lemmas I show about word_upto hold without a \ b, but I don't need that right now and it's giving me a headache *) lemma word_upto_set_eq: "a \ b \ x \ set (word_upto a b) \ a \ x \ x \ b" proof show "a \ b \ x \ set (word_upto a b) \ a \ x \ x \ b" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst(asm) word_upto.simps) apply(simp; fail) apply(subst(asm) word_upto.simps) apply(simp) apply(erule disjE) apply(simp; fail) proof(goal_cases) case (1 a b) from 1(2-3) have "b \ 0" by force from 1(2,3) have "a \ b - 1" by (metis \b \ 0\ eq_iff le_step_down_nat unat_arith_simps(1) unat_minus_one) from 1(1)[OF this 1(4)] show ?case by (metis dual_order.trans 1(2,3) less_imp_le measure_unat word_le_0_iff word_le_nat_alt) qed next show "a \ x \ x \ b \ x \ set (word_upto a b)" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst word_upto.simps) apply(simp; force) apply(subst word_upto.simps) apply(simp) apply(case_tac "x = b") apply(simp;fail) proof(goal_cases) case (1 a b) from 1(2-4) have "b \ 0" by force - from 1(2,4) have "x \ b - 1" by (metis \b \ 0\ dual_order.antisym le_step_down_nat unat_minus_one word_le_nat_alt) + from 1(2,4) have "x \ b - 1" + using le_step_down_word by auto from 1(1) this show ?case by simp qed qed lemma word_upto_distinct_hlp: "a \ b \ a \ b \ b \ set (word_upto a (b - 1))" apply(rule ccontr, unfold not_not) apply(subgoal_tac "a \ b - 1") apply(drule iffD1[OF word_upto_set_eq[of a "b -1" b]]) - apply(simp add: word_upto.simps; fail) + apply(simp add: word_upto.simps) apply(subgoal_tac "b \ 0") apply(meson leD measure_unat word_le_nat_alt) - apply(blast intro: iffD1[OF word_le_0_iff]) - apply(metis antisym_conv le_step_down_nat unat_minus_one word_le_0_iff word_le_nat_alt) + apply(blast intro: iffD1[OF word_le_0_iff]) + using le_step_down_word apply blast done lemma distinct_word_upto: "a \ b \ distinct (word_upto a b)" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst word_upto.simps) apply(simp; force) apply(subst word_upto.simps) apply(case_tac "a \ b - 1") apply(simp) - apply(rule word_upto_distinct_hlp; simp; fail) + apply(rule word_upto_distinct_hlp; simp) apply(simp) -apply(rule ccontr) -apply(metis le_step_down_nat less_le not_le unat_minus_one word_le_nat_alt word_not_simps(1)) + apply(rule ccontr) + apply (simp add: not_le antisym word_minus_one_le_leq) done lemma word_upto_eq_upto: "s \ e \ e \ unat (max_word :: 'l word) \ word_upto ((of_nat :: nat \ ('l :: len) word) s) (of_nat e) = map of_nat (upt s (Suc e))" proof(induction e) let ?mwon = "of_nat :: nat \ 'l word" let ?mmw = "max_word :: 'l word" case (Suc e) show ?case proof(cases "?mwon s = ?mwon (Suc e)") case True have "s = Suc e" using le_unat_uoi Suc.prems True by metis with True show ?thesis by(subst word_upto.simps) (simp) next case False hence le: "s \ e" using le_SucE Suc.prems by blast have lm: "e \ unat ?mmw" using Suc.prems by simp have sucm: "(of_nat :: nat \ ('l :: len) word) (Suc e) - 1 = of_nat e" using Suc.prems(2) by simp note mIH = Suc.IH[OF le lm] show ?thesis by(subst word_upto.simps) (simp add: False[simplified] Suc.prems mIH sucm) qed qed(simp add: word_upto.simps) lemma word_upto_alt: "(a :: ('l :: len) word) \ b \ word_upto a b = map of_nat (upt (unat a) (Suc (unat b)))" proof - let ?mmw = "max_word :: 'l word" assume le: "a \ b" hence nle: "unat a \ unat b" by(unat_arith) have lem: "unat b \ unat ?mmw" by (simp add: word_unat_less_le) note word_upto_eq_upto[OF nle lem, unfolded word_unat.Rep_inverse] thus "word_upto a b = map of_nat [unat a.. b then map of_nat (upt (unat a) (Suc (unat b))) else word_upto a b)" using word_upto_alt by metis lemma sorted_word_upto: fixes a b :: "('l :: len) word" assumes "a \ b" shows "sorted (word_upto a b)" proof - define m and n where \m = unat a\ and \n = Suc (unat b)\ moreover have \sorted (map of_nat [m.. apply (simp add: sorted_map) apply (rule sorted_wrt_mono_rel [of _ \(\)\]) apply (simp_all flip: sorted_sorted_wrt) apply (simp add: le_unat_uoi less_Suc_eq_le n_def word_of_nat_le) + apply transfer + apply simp + apply (subst take_bit_int_eq_self) + apply (simp_all add: le_less_trans) + apply (metis le_unat_uoi of_int_of_nat_eq of_nat_mono uint_word_of_int_eq unat_eq_nat_uint unsigned_of_int) done ultimately have \sorted (map of_nat [unat a.. by simp with assms show ?thesis by (simp only: word_upto_alt) qed end diff --git a/thys/JinjaThreads/Compiler/Correctness1.thy b/thys/JinjaThreads/Compiler/Correctness1.thy --- a/thys/JinjaThreads/Compiler/Correctness1.thy +++ b/thys/JinjaThreads/Compiler/Correctness1.thy @@ -1,2385 +1,2385 @@ (* Title: JinjaThreads/Compiler/Correctness1.thy Author: Andreas Lochbihler, Tobias Nipkow reminiscent of the Jinja theory Compiler/Correctness1 *) section \Semantic Correctness of Stage 1\ theory Correctness1 imports J0J1Bisim "../J/DefAssPreservation" begin lemma finals_map_Val [simp]: "finals (map Val vs)" by(simp add: finals_iff) context J_heap_base begin lemma \red0r_preserves_defass: assumes wf: "wf_J_prog P" shows "\ \red0r extTA P t h (e, xs) (e', xs'); \ e \dom xs\ \ \ \ e' \dom xs'\" by(induct rule: rtranclp_induct2)(auto dest: red_preserves_defass[OF wf]) lemma \red0t_preserves_defass: assumes wf: "wf_J_prog P" shows "\ \red0t extTA P t h (e, xs) (e', xs'); \ e \dom xs\ \ \ \ e' \dom xs'\" by(rule \red0r_preserves_defass[OF wf])(rule tranclp_into_rtranclp) end lemma LAss_lem: "\x \ set xs; size xs \ size ys \ \ m1 \\<^sub>m m2(xs[\]ys) \ m1(x\y) \\<^sub>m m2(xs[\]ys[index xs x := y])" apply(simp add:map_le_def) apply(simp add:fun_upds_apply index_less_aux eq_sym_conv) done lemma Block_lem: fixes l :: "'a \ 'b" assumes 0: "l \\<^sub>m [Vs [\] ls]" and 1: "l' \\<^sub>m [Vs [\] ls', V\v]" and hidden: "V \ set Vs \ ls ! index Vs V = ls' ! index Vs V" and size: "size ls = size ls'" "size Vs < size ls'" shows "l'(V := l V) \\<^sub>m [Vs [\] ls']" proof - have "l'(V := l V) \\<^sub>m [Vs [\] ls', V\v](V := l V)" using 1 by(rule map_le_upd) also have "\ = [Vs [\] ls'](V := l V)" by simp also have "\ \\<^sub>m [Vs [\] ls']" proof (cases "l V") case None thus ?thesis by simp next case (Some w) hence "[Vs [\] ls] V = Some w" using 0 by(force simp add: map_le_def split:if_splits) hence VinVs: "V \ set Vs" and w: "w = ls ! index Vs V" using size by(auto simp add:fun_upds_apply split:if_splits) hence "w = ls' ! index Vs V" using hidden[OF VinVs] by simp hence "[Vs [\] ls'](V := l V) = [Vs [\] ls']" using Some size VinVs by(simp add:index_less_aux map_upds_upd_conv_index) thus ?thesis by simp qed finally show ?thesis . qed subsection \Correctness proof\ locale J0_J1_heap_base = J?: J_heap_base + J1?: J1_heap_base + constrains addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" begin lemma ta_bisim01_extTA2J0_extTA2J1: assumes wf: "wf_J_prog P" and nt: "\n T C M a h. \ n < length \ta\\<^bsub>t\<^esub>; \ta\\<^bsub>t\<^esub> ! n = NewThread T (C, M, a) h \ \ typeof_addr h a = \Class_type C\ \ (\T meth D. P \ C sees M:[]\T =\meth\ in D)" shows "ta_bisim01 (extTA2J0 P ta) (extTA2J1 (compP1 P) ta)" apply(simp add: ta_bisim_def ta_upd_simps) apply(auto intro!: list_all2_all_nthI) apply(case_tac "\ta\\<^bsub>t\<^esub> ! n") apply(auto simp add: bisim_red0_Red1_def) apply(drule (1) nt) apply(clarify) apply(erule bisim_list_extTA2J0_extTA2J1[OF wf, simplified]) done lemma red_external_ta_bisim01: "\ wf_J_prog P; P,t \ \a\M(vs), h\ -ta\ext \va, h'\ \ \ ta_bisim01 (extTA2J0 P ta) (extTA2J1 (compP1 P) ta)" apply(rule ta_bisim01_extTA2J0_extTA2J1, assumption) apply(drule (1) red_external_new_thread_sees, auto simp add: in_set_conv_nth) apply(drule red_ext_new_thread_heap, auto simp add: in_set_conv_nth) done lemmas \red1t_expr = NewArray_\red1t_xt Cast_\red1t_xt InstanceOf_\red1t_xt BinOp_\red1t_xt1 BinOp_\red1t_xt2 LAss_\red1t AAcc_\red1t_xt1 AAcc_\red1t_xt2 AAss_\red1t_xt1 AAss_\red1t_xt2 AAss_\red1t_xt3 ALength_\red1t_xt FAcc_\red1t_xt FAss_\red1t_xt1 FAss_\red1t_xt2 CAS_\red1t_xt1 CAS_\red1t_xt2 CAS_\red1t_xt3 Call_\red1t_obj Call_\red1t_param Block_None_\red1t_xt Block_\red1t_Some Sync_\red1t_xt InSync_\red1t_xt Seq_\red1t_xt Cond_\red1t_xt Throw_\red1t_xt Try_\red1t_xt lemmas \red1r_expr = NewArray_\red1r_xt Cast_\red1r_xt InstanceOf_\red1r_xt BinOp_\red1r_xt1 BinOp_\red1r_xt2 LAss_\red1r AAcc_\red1r_xt1 AAcc_\red1r_xt2 AAss_\red1r_xt1 AAss_\red1r_xt2 AAss_\red1r_xt3 ALength_\red1r_xt FAcc_\red1r_xt FAss_\red1r_xt1 FAss_\red1r_xt2 CAS_\red1r_xt1 CAS_\red1r_xt2 CAS_\red1r_xt3 Call_\red1r_obj Call_\red1r_param Block_None_\red1r_xt Block_\red1r_Some Sync_\red1r_xt InSync_\red1r_xt Seq_\red1r_xt Cond_\red1r_xt Throw_\red1r_xt Try_\red1r_xt definition sim_move01 :: "'addr J1_prog \ 'thread_id \ ('addr, 'thread_id, 'heap) J0_thread_action \ 'addr expr \ 'addr expr1 \ 'heap \ 'addr locals1 \ ('addr, 'thread_id, 'heap) external_thread_action \ 'addr expr1 \ 'heap \ 'addr locals1 \ bool" where "sim_move01 P t ta0 e0 e h xs ta e' h' xs' \ \ final e0 \ (if \move0 P h e0 then h' = h \ ta0 = \ \ ta = \ \ \red1't P t h (e, xs) (e', xs') else ta_bisim01 ta0 (extTA2J1 P ta) \ (if call e0 = None \ call1 e = None then (\e'' xs''. \red1'r P t h (e, xs) (e'', xs'') \ False,P,t \1 \e'', (h, xs'')\ -ta\ \e', (h', xs')\ \ \ \move1 P h e'') else False,P,t \1 \e, (h, xs)\ -ta\ \e', (h', xs')\ \ \ \move1 P h e))" definition sim_moves01 :: "'addr J1_prog \ 'thread_id \ ('addr, 'thread_id, 'heap) J0_thread_action \ 'addr expr list \ 'addr expr1 list \ 'heap \ 'addr locals1 \ ('addr, 'thread_id, 'heap) external_thread_action \ 'addr expr1 list \ 'heap \ 'addr locals1 \ bool" where "sim_moves01 P t ta0 es0 es h xs ta es' h' xs' \ \ finals es0 \ (if \moves0 P h es0 then h' = h \ ta0 = \ \ ta = \ \ \reds1't P t h (es, xs) (es', xs') else ta_bisim01 ta0 (extTA2J1 P ta) \ (if calls es0 = None \ calls1 es = None then (\es'' xs''. \reds1'r P t h (es, xs) (es'', xs'') \ False,P,t \1 \es'', (h, xs'')\ [-ta\] \es', (h', xs')\ \ \ \moves1 P h es'') else False,P,t \1 \es, (h, xs)\ [-ta\] \es', (h', xs')\ \ \ \moves1 P h es))" declare \red1t_expr [elim!] \red1r_expr[elim!] lemma sim_move01_expr: assumes "sim_move01 P t ta0 e0 e h xs ta e' h' xs'" shows "sim_move01 P t ta0 (newA T\e0\) (newA T\e\) h xs ta (newA T\e'\) h' xs'" "sim_move01 P t ta0 (Cast T e0) (Cast T e) h xs ta (Cast T e') h' xs'" "sim_move01 P t ta0 (e0 instanceof T) (e instanceof T) h xs ta (e' instanceof T) h' xs'" "sim_move01 P t ta0 (e0 \bop\ e2) (e \bop\ e2') h xs ta (e' \bop\ e2') h' xs'" "sim_move01 P t ta0 (Val v \bop\ e0) (Val v \bop\ e) h xs ta (Val v \bop\ e') h' xs'" "sim_move01 P t ta0 (V := e0) (V' := e) h xs ta (V' := e') h' xs'" "sim_move01 P t ta0 (e0\e2\) (e\e2'\) h xs ta (e'\e2'\) h' xs'" "sim_move01 P t ta0 (Val v\e0\) (Val v\e\) h xs ta (Val v\e'\) h' xs'" "sim_move01 P t ta0 (e0\e2\ := e3) (e\e2'\ := e3') h xs ta (e'\e2'\ := e3') h' xs'" "sim_move01 P t ta0 (Val v\e0\ := e3) (Val v\e\ := e3') h xs ta (Val v\e'\ := e3') h' xs'" "sim_move01 P t ta0 (AAss (Val v) (Val v') e0) (AAss (Val v) (Val v') e) h xs ta (AAss (Val v) (Val v') e') h' xs'" "sim_move01 P t ta0 (e0\length) (e\length) h xs ta (e'\length) h' xs'" "sim_move01 P t ta0 (e0\F{D}) (e\F'{D'}) h xs ta (e'\F'{D'}) h' xs'" "sim_move01 P t ta0 (FAss e0 F D e2) (FAss e F' D' e2') h xs ta (FAss e' F' D' e2') h' xs'" "sim_move01 P t ta0 (FAss (Val v) F D e0) (FAss (Val v) F' D' e) h xs ta (FAss (Val v) F' D' e') h' xs'" "sim_move01 P t ta0 (CompareAndSwap e0 D F e2 e3) (CompareAndSwap e D F e2' e3') h xs ta (CompareAndSwap e' D F e2' e3') h' xs'" "sim_move01 P t ta0 (CompareAndSwap (Val v) D F e0 e3) (CompareAndSwap (Val v) D F e e3') h xs ta (CompareAndSwap (Val v) D F e' e3') h' xs'" "sim_move01 P t ta0 (CompareAndSwap (Val v) D F (Val v') e0) (CompareAndSwap (Val v) D F (Val v') e) h xs ta (CompareAndSwap (Val v) D F (Val v') e') h' xs'" "sim_move01 P t ta0 (e0\M(es)) (e\M(es')) h xs ta (e'\M(es')) h' xs'" "sim_move01 P t ta0 ({V:T=vo; e0}) ({V':T=None; e}) h xs ta ({V':T=None; e'}) h' xs'" "sim_move01 P t ta0 (sync(e0) e2) (sync\<^bsub>V'\<^esub>(e) e2') h xs ta (sync\<^bsub>V'\<^esub>(e') e2') h' xs'" "sim_move01 P t ta0 (insync(a) e0) (insync\<^bsub>V'\<^esub>(a') e) h xs ta (insync\<^bsub>V'\<^esub>(a') e') h' xs'" "sim_move01 P t ta0 (e0;;e2) (e;;e2') h xs ta (e';;e2') h' xs'" "sim_move01 P t ta0 (if (e0) e2 else e3) (if (e) e2' else e3') h xs ta (if (e') e2' else e3') h' xs'" "sim_move01 P t ta0 (throw e0) (throw e) h xs ta (throw e') h' xs'" "sim_move01 P t ta0 (try e0 catch(C V) e2) (try e catch(C' V') e2') h xs ta (try e' catch(C' V') e2') h' xs'" using assms apply(simp_all add: sim_move01_def final_iff \red1r_Val \red1t_Val split: if_split_asm split del: if_split) apply(fastforce simp add: final_iff \red1r_Val \red1t_Val split!: if_splits intro: red1_reds1.intros)+ done lemma sim_moves01_expr: "sim_move01 P t ta0 e0 e h xs ta e' h' xs' \ sim_moves01 P t ta0 (e0 # es2) (e # es2') h xs ta (e' # es2') h' xs'" "sim_moves01 P t ta0 es0 es h xs ta es' h' xs' \ sim_moves01 P t ta0 (Val v # es0) (Val v # es) h xs ta (Val v # es') h' xs'" apply(simp_all add: sim_move01_def sim_moves01_def final_iff finals_iff Cons_eq_append_conv \red1t_Val \red1r_Val split: if_split_asm split del: if_split) apply(auto simp add: Cons_eq_append_conv \red1t_Val \red1r_Val split!: if_splits intro: List1Red1 List1Red2 \red1t_inj_\reds1t \red1r_inj_\reds1r \reds1t_cons_\reds1t \reds1r_cons_\reds1r) apply(force elim!: \red1r_inj_\reds1r List1Red1) apply(force elim!: \red1r_inj_\reds1r List1Red1) apply(force elim!: \red1r_inj_\reds1r List1Red1) apply(force elim!: \red1r_inj_\reds1r List1Red1) apply(force elim!: \reds1r_cons_\reds1r intro!: List1Red2) apply(force elim!: \reds1r_cons_\reds1r intro!: List1Red2) done lemma sim_move01_CallParams: "sim_moves01 P t ta0 es0 es h xs ta es' h' xs' \ sim_move01 P t ta0 (Val v\M(es0)) (Val v\M(es)) h xs ta (Val v\M(es')) h' xs'" apply(clarsimp simp add: sim_move01_def sim_moves01_def \reds1r_map_Val \reds1t_map_Val is_vals_conv split: if_split_asm split del: if_split) apply(fastforce simp add: sim_move01_def sim_moves01_def \reds1r_map_Val \reds1t_map_Val intro: Call_\red1r_param Call1Params) apply(rule conjI, fastforce) apply(split if_split) apply(rule conjI) apply(clarsimp simp add: finals_iff) apply(clarify) apply(split if_split) apply(rule conjI) apply(simp del: call.simps calls.simps call1.simps calls1.simps) apply(fastforce simp add: sim_move01_def sim_moves01_def \red1r_Val \red1t_Val \reds1r_map_Val_Throw intro: Call_\red1r_param Call1Params split: if_split_asm) apply(fastforce split: if_split_asm simp add: is_vals_conv \reds1r_map_Val \reds1r_map_Val_Throw) apply(rule conjI, fastforce) apply(fastforce simp add: sim_move01_def sim_moves01_def \red1r_Val \red1t_Val \reds1t_map_Val \reds1r_map_Val is_vals_conv intro: Call_\red1r_param Call1Params split: if_split_asm) done lemma sim_move01_reds: "\ (h', a) \ allocate h (Class_type C); ta0 = \NewHeapElem a (Class_type C)\; ta = \NewHeapElem a (Class_type C)\ \ \ sim_move01 P t ta0 (new C) (new C) h xs ta (addr a) h' xs" "allocate h (Class_type C) = {} \ sim_move01 P t \ (new C) (new C) h xs \ (THROW OutOfMemory) h xs" "\ (h', a) \ allocate h (Array_type T (nat (sint i))); 0 <=s i; ta0 = \NewHeapElem a (Array_type T (nat (sint i)))\; ta = \NewHeapElem a (Array_type T (nat (sint i)))\ \ \ sim_move01 P t ta0 (newA T\Val (Intg i)\) (newA T\Val (Intg i)\) h xs ta (addr a) h' xs" "i sim_move01 P t \ (newA T\Val (Intg i)\) (newA T\Val (Intg i)\) h xs \ (THROW NegativeArraySize) h xs" "\ allocate h (Array_type T (nat (sint i))) = {}; 0 <=s i \ \ sim_move01 P t \ (newA T\Val (Intg i)\) (newA T\Val (Intg i)\) h xs \ (THROW OutOfMemory) h xs" "\ typeof\<^bsub>h\<^esub> v = \U\; P \ U \ T \ \ sim_move01 P t \ (Cast T (Val v)) (Cast T (Val v)) h xs \ (Val v) h xs" "\ typeof\<^bsub>h\<^esub> v = \U\; \ P \ U \ T \ \ sim_move01 P t \ (Cast T (Val v)) (Cast T (Val v)) h xs \ (THROW ClassCast) h xs" "\ typeof\<^bsub>h\<^esub> v = \U\; b \ v \ Null \ P \ U \ T \ \ sim_move01 P t \ ((Val v) instanceof T) ((Val v) instanceof T) h xs \ (Val (Bool b)) h xs" "binop bop v1 v2 = Some (Inl v) \ sim_move01 P t \ ((Val v1) \bop\ (Val v2)) (Val v1 \bop\ Val v2) h xs \ (Val v) h xs" "binop bop v1 v2 = Some (Inr a) \ sim_move01 P t \ ((Val v1) \bop\ (Val v2)) (Val v1 \bop\ Val v2) h xs \ (Throw a) h xs" "\ xs!V = v; V < size xs \ \ sim_move01 P t \ (Var V') (Var V) h xs \ (Val v) h xs" "V < length xs \ sim_move01 P t \ (V' := Val v) (V := Val v) h xs \ unit h (xs[V := v])" "sim_move01 P t \ (null\Val v\) (null\Val v\) h xs \ (THROW NullPointer) h xs" "\ typeof_addr h a = \Array_type T n\; i sint i \ int n \ \ sim_move01 P t \ (addr a\Val (Intg i)\) ((addr a)\Val (Intg i)\) h xs \ (THROW ArrayIndexOutOfBounds) h xs" "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; heap_read h a (ACell (nat (sint i))) v; ta0 = \ReadMem a (ACell (nat (sint i))) v\; ta = \ReadMem a (ACell (nat (sint i))) v\ \ \ sim_move01 P t ta0 (addr a\Val (Intg i)\) ((addr a)\Val (Intg i)\) h xs ta (Val v) h xs" "sim_move01 P t \ (null\Val v\ := Val v') (null\Val v\ := Val v') h xs \ (THROW NullPointer) h xs" "\ typeof_addr h a = \Array_type T n\; i sint i \ int n \ \ sim_move01 P t \ (AAss (addr a) (Val (Intg i)) (Val v)) (AAss (addr a) (Val (Intg i)) (Val v)) h xs \ (THROW ArrayIndexOutOfBounds) h xs" "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; typeof\<^bsub>h\<^esub> v = \U\; \ (P \ U \ T) \ \ sim_move01 P t \ (AAss (addr a) (Val (Intg i)) (Val v)) (AAss (addr a) (Val (Intg i)) (Val v)) h xs \ (THROW ArrayStore) h xs" "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; typeof\<^bsub>h\<^esub> v = Some U; P \ U \ T; heap_write h a (ACell (nat (sint i))) v h'; ta0 = \WriteMem a (ACell (nat (sint i))) v\; ta = \WriteMem a (ACell (nat (sint i))) v \ \ \ sim_move01 P t ta0 (AAss (addr a) (Val (Intg i)) (Val v)) (AAss (addr a) (Val (Intg i)) (Val v)) h xs ta unit h' xs" "typeof_addr h a = \Array_type T n\ \ sim_move01 P t \ (addr a\length) (addr a\length) h xs \ (Val (Intg (word_of_int (int n)))) h xs" "sim_move01 P t \ (null\length) (null\length) h xs \ (THROW NullPointer) h xs" "\ heap_read h a (CField D F) v; ta0 = \ReadMem a (CField D F) v\; ta = \ReadMem a (CField D F) v\ \ \ sim_move01 P t ta0 (addr a\F{D}) (addr a\F{D}) h xs ta (Val v) h xs" "sim_move01 P t \ (null\F{D}) (null\F{D}) h xs \ (THROW NullPointer) h xs" "\ heap_write h a (CField D F) v h'; ta0 = \WriteMem a (CField D F) v\; ta = \WriteMem a (CField D F) v\ \ \ sim_move01 P t ta0 (addr a\F{D} := Val v) (addr a\F{D} := Val v) h xs ta unit h' xs" "sim_move01 P t \ (null\compareAndSwap(D\F, Val v, Val v')) (null\compareAndSwap(D\F, Val v, Val v')) h xs \ (THROW NullPointer) h xs" "\ heap_read h a (CField D F) v''; heap_write h a (CField D F) v' h'; v'' = v; ta0 = \ReadMem a (CField D F) v'', WriteMem a (CField D F) v'\; ta = \ReadMem a (CField D F) v'', WriteMem a (CField D F) v'\ \ \ sim_move01 P t ta0 (addr a\compareAndSwap(D\F, Val v, Val v')) (addr a\compareAndSwap(D\F, Val v, Val v')) h xs ta true h' xs" "\ heap_read h a (CField D F) v''; v'' \ v; ta0 = \ReadMem a (CField D F) v''\; ta = \ReadMem a (CField D F) v''\ \ \ sim_move01 P t ta0 (addr a\compareAndSwap(D\F, Val v, Val v')) (addr a\compareAndSwap(D\F, Val v, Val v')) h xs ta false h xs" "sim_move01 P t \ (null\F{D} := Val v) (null\F{D} := Val v) h xs \ (THROW NullPointer) h xs" "sim_move01 P t \ ({V':T=vo; Val u}) ({V:T=None; Val u}) h xs \ (Val u) h xs" "V < length xs \ sim_move01 P t \ (sync(null) e0) (sync\<^bsub>V\<^esub> (null) e1) h xs \ (THROW NullPointer) h (xs[V := Null])" "sim_move01 P t \ (Val v;;e0) (Val v;; e1) h xs \ e1 h xs" "sim_move01 P t \ (if (true) e0 else e0') (if (true) e1 else e1') h xs \ e1 h xs" "sim_move01 P t \ (if (false) e0 else e0') (if (false) e1 else e1') h xs \ e1' h xs" "sim_move01 P t \ (throw null) (throw null) h xs \ (THROW NullPointer) h xs" "sim_move01 P t \ (try (Val v) catch(C V') e0) (try (Val v) catch(C V) e1) h xs \ (Val v) h xs" "\ typeof_addr h a = \Class_type D\; P \ D \\<^sup>* C; V < length xs \ \ sim_move01 P t \ (try (Throw a) catch(C V') e0) (try (Throw a) catch(C V) e1) h xs \ ({V:Class C=None; e1}) h (xs[V := Addr a])" "sim_move01 P t \ (newA T\Throw a\) (newA T\Throw a\) h xs \ (Throw a) h xs" "sim_move01 P t \ (Cast T (Throw a)) (Cast T (Throw a)) h xs \ (Throw a) h xs" "sim_move01 P t \ ((Throw a) instanceof T) ((Throw a) instanceof T) h xs \ (Throw a) h xs" "sim_move01 P t \ ((Throw a) \bop\ e0) ((Throw a) \bop\ e1) h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v \bop\ (Throw a)) (Val v \bop\ (Throw a)) h xs \ (Throw a) h xs" "sim_move01 P t \ (V' := Throw a) (V := Throw a) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\e0\) (Throw a\e1\) h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v\Throw a\) (Val v\Throw a\) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\e0\ := e0') (Throw a\e1\ := e1') h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v\Throw a\ := e0) (Val v\Throw a\ := e1) h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v\Val v'\ := Throw a) (Val v\Val v'\ := Throw a) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\length) (Throw a\length) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\F{D}) (Throw a\F{D}) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\F{D} := e0) (Throw a\F{D} := e1) h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v\F{D} := Throw a) (Val v\F{D} := Throw a) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\compareAndSwap(D\F, e2, e3)) (Throw a\compareAndSwap(D\F, e2', e3')) h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v\compareAndSwap(D\F, Throw a, e3)) (Val v\compareAndSwap(D\F, Throw a, e3')) h xs \ (Throw a) h xs" "sim_move01 P t \ (Val v\compareAndSwap(D\F, Val v', Throw a)) (Val v\compareAndSwap(D\F, Val v', Throw a)) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a\M(es0)) (Throw a\M(es1)) h xs \ (Throw a) h xs" "sim_move01 P t \ ({V':T=vo; Throw a}) ({V:T=None; Throw a}) h xs \ (Throw a) h xs" "sim_move01 P t \ (sync(Throw a) e0) (sync\<^bsub>V\<^esub>(Throw a) e1) h xs \ (Throw a) h xs" "sim_move01 P t \ (Throw a;;e0) (Throw a;;e1) h xs \ (Throw a) h xs" "sim_move01 P t \ (if (Throw a) e0 else e0') (if (Throw a) e1 else e1') h xs \ (Throw a) h xs" "sim_move01 P t \ (throw (Throw a)) (throw (Throw a)) h xs \ (Throw a) h xs" apply(simp_all add: sim_move01_def ta_bisim_def split: if_split_asm split del: if_split) apply(fastforce intro: red1_reds1.intros)+ done lemma sim_move01_ThrowParams: "sim_move01 P t \ (Val v\M(map Val vs @ Throw a # es0)) (Val v\M(map Val vs @ Throw a # es1)) h xs \ (Throw a) h xs" apply(simp add: sim_move01_def split del: if_split) apply(rule conjI, fastforce) apply(split if_split) apply(rule conjI) apply(fastforce intro: red1_reds1.intros) apply(fastforce simp add: sim_move01_def intro: red1_reds1.intros) done lemma sim_move01_CallNull: "sim_move01 P t \ (null\M(map Val vs)) (null\M(map Val vs)) h xs \ (THROW NullPointer) h xs" by(fastforce simp add: sim_move01_def map_eq_append_conv intro: red1_reds1.intros) lemma sim_move01_SyncLocks: "\ V < length xs; ta0 = \Lock\a, SyncLock a\; ta = \Lock\a, SyncLock a\ \ \ sim_move01 P t ta0 (sync(addr a) e0) (sync\<^bsub>V\<^esub> (addr a) e1) h xs ta (insync\<^bsub>V\<^esub> (a) e1) h (xs[V := Addr a])" "\ xs ! V = Addr a'; V < length xs; ta0 = \Unlock\a', SyncUnlock a'\; ta = \Unlock\a', SyncUnlock a'\ \ \ sim_move01 P t ta0 (insync(a') (Val v)) (insync\<^bsub>V\<^esub> (a) (Val v)) h xs ta (Val v) h xs" "\ xs ! V = Addr a'; V < length xs; ta0 = \Unlock\a', SyncUnlock a'\; ta = \Unlock\a', SyncUnlock a'\ \ \ sim_move01 P t ta0 (insync(a') (Throw a'')) (insync\<^bsub>V\<^esub> (a) (Throw a'')) h xs ta (Throw a'') h xs" by(fastforce simp add: sim_move01_def ta_bisim_def expand_finfun_eq fun_eq_iff finfun_upd_apply ta_upd_simps intro: red1_reds1.intros[simplified] split: if_split_asm)+ lemma sim_move01_TryFail: "\ typeof_addr h a = \Class_type D\; \ P \ D \\<^sup>* C \ \ sim_move01 P t \ (try (Throw a) catch(C V') e0) (try (Throw a) catch(C V) e1) h xs \ (Throw a) h xs" by(auto simp add: sim_move01_def intro!: Red1TryFail) lemma sim_move01_BlockSome: "\ sim_move01 P t ta0 e0 e h (xs[V := v]) ta e' h' xs'; V < length xs \ \ sim_move01 P t ta0 ({V':T=\v\; e0}) ({V:T=\v\; e}) h xs ta ({V:T=None; e'}) h' xs'" "V < length xs \ sim_move01 P t \ ({V':T=\v\; Val u}) ({V:T=\v\; Val u}) h xs \ (Val u) h (xs[V := v])" "V < length xs \ sim_move01 P t \ ({V':T=\v\; Throw a}) ({V:T=\v\; Throw a}) h xs \ (Throw a) h (xs[V := v])" apply(auto simp add: sim_move01_def) apply(split if_split_asm) apply(fastforce intro: intro: converse_rtranclp_into_rtranclp Block1Some Block1Red Block_\red1r_Some) apply(fastforce intro: intro: converse_rtranclp_into_rtranclp Block1Some Block1Red Block_\red1r_Some) apply(fastforce simp add: sim_move01_def intro!: \red1t_2step[OF Block1Some] \red1r_1step[OF Block1Some] Red1Block Block1Throw)+ done lemmas sim_move01_intros = sim_move01_expr sim_move01_reds sim_move01_ThrowParams sim_move01_CallNull sim_move01_TryFail sim_move01_BlockSome sim_move01_CallParams declare sim_move01_intros[intro] lemma sim_move01_preserves_len: "sim_move01 P t ta0 e0 e h xs ta e' h' xs' \ length xs' = length xs" by(fastforce simp add: sim_move01_def split: if_split_asm dest: \red1r_preserves_len \red1t_preserves_len red1_preserves_len) lemma sim_move01_preserves_unmod: "\ sim_move01 P t ta0 e0 e h xs ta e' h' xs'; unmod e i; i < length xs \ \ xs' ! i = xs ! i" apply(auto simp add: sim_move01_def split: if_split_asm dest: \red1t_preserves_unmod) apply(frule (2) \red1'r_preserves_unmod) apply(frule (1) \red1r_unmod_preserved) apply(frule \red1r_preserves_len) apply(auto dest: red1_preserves_unmod) apply(frule (2) \red1'r_preserves_unmod) apply(frule (1) \red1r_unmod_preserved) apply(frule \red1r_preserves_len) apply(auto dest: red1_preserves_unmod) done lemma assumes wf: "wf_J_prog P" shows red1_simulates_red_aux: "\ extTA2J0 P,P,t \ \e1, S\ -TA\ \e1', S'\; bisim vs e1 e2 XS; fv e1 \ set vs; lcl S \\<^sub>m [vs [\] XS]; length vs + max_vars e1 \ length XS; \aMvs. call e1 = \aMvs\ \ synthesized_call P (hp S) aMvs \ \ \ta e2' XS'. sim_move01 (compP1 P) t TA e1 e2 (hp S) XS ta e2' (hp S') XS' \ bisim vs e1' e2' XS' \ lcl S' \\<^sub>m [vs [\] XS']" (is "\ _; _; _; _; _; ?synth e1 S \ \ ?concl e1 e2 S XS e1' S' TA vs") and reds1_simulates_reds_aux: "\ extTA2J0 P,P,t \ \es1, S\ [-TA\] \es1', S'\; bisims vs es1 es2 XS; fvs es1 \ set vs; lcl S \\<^sub>m [vs [\] XS]; length vs + max_varss es1 \ length XS; \aMvs. calls es1 = \aMvs\ \ synthesized_call P (hp S) aMvs \ \ \ta es2' xs'. sim_moves01 (compP1 P) t TA es1 es2 (hp S) XS ta es2' (hp S') xs' \ bisims vs es1' es2' xs' \ lcl S' \\<^sub>m [vs [\] xs']" (is "\ _; _; _; _; _; ?synths es1 S \ \ ?concls es1 es2 S XS es1' S' TA vs") proof(induct arbitrary: vs e2 XS and vs es2 XS rule: red_reds.inducts) case (BinOpRed1 e s ta e' s' bop e2 Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs (e \bop\ e2) E2 xs\ obtain E where "E2 = E \bop\ compE1 Vs e2" and "bisim Vs e E xs" and "\ contains_insync e2" by auto with IH[of Vs E xs] \fv (e \bop\ e2) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val e\ \length Vs + max_vars (e \bop\ e2) \ length xs\ \?synth (e \bop\ e2) s\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ show ?case by(cases "is_val e'")(fastforce elim!: sim_move01_expr)+ next case (BinOpRed2 e s ta e' s' v bop Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \bisim Vs (Val v \bop\ e) E2 xs\ obtain E where "E2 = Val v \bop\ E" and "bisim Vs e E xs" by auto with IH[of Vs E xs] \fv (Val v \bop\ e) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars (Val v \bop\ e) \ length xs\ \?synth (Val v \bop\ e) s\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ show ?case by(fastforce elim!: sim_move01_expr) next case RedVar thus ?case by(fastforce simp add: index_less_aux map_le_def fun_upds_apply intro!: exI dest: bspec) next case RedLAss thus ?case by(fastforce intro: index_less_aux LAss_lem intro!: exI simp del: fun_upd_apply) next case (AAccRed1 a s ta a' s' i Vs E2 xs) note IH = \\vs e2 XS. \bisim vs a e2 XS; fv a \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars a \ length XS; ?synth a s \ \ ?concl a e2 s XS a' s' ta vs\ from \extTA2J0 P,P,t \ \a,s\ -ta\ \a',s'\\ have "\ is_val a" by auto with \bisim Vs (a\i\) E2 xs\ obtain E where "E2 = E\compE1 Vs i\" and "bisim Vs a E xs" and "\ contains_insync i" by auto with IH[of Vs E xs] \fv (a\i\) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val a\ \length Vs + max_vars (a\i\) \ length xs\ \?synth (a\i\) s\ \extTA2J0 P,P,t \ \a,s\ -ta\ \a',s'\\ show ?case by(cases "is_val a'")(fastforce elim!: sim_move01_expr)+ next case (AAccRed2 i s ta i' s' a Vs E2 xs) note IH = \\vs e2 XS. \bisim vs i e2 XS; fv i \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars i \ length XS; ?synth i s \ \ ?concl i e2 s XS i' s' ta vs\ from \bisim Vs (Val a\i\) E2 xs\ obtain E where "E2 = Val a\E\" and "bisim Vs i E xs" by auto with IH[of Vs E xs] \fv (Val a\i\) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars (Val a\i\) \ length xs\ \?synth (Val a\i\) s\ \extTA2J0 P,P,t \ \i,s\ -ta\ \i',s'\\ show ?case by(fastforce elim!: sim_move01_expr) next case RedAAcc thus ?case by(auto simp del: split_paired_Ex) next case (AAssRed1 a s ta a' s' i e Vs E2 xs) note IH = \\vs e2 XS. \bisim vs a e2 XS; fv a \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars a \ length XS; ?synth a s \ \ ?concl a e2 s XS a' s' ta vs\ from \extTA2J0 P,P,t \ \a,s\ -ta\ \a',s'\\ have "\ is_val a" by auto with \bisim Vs (a\i\:=e) E2 xs\ obtain E where E2: "E2 = E\compE1 Vs i\:=compE1 Vs e" and "bisim Vs a E xs" and sync: "\ contains_insync i" "\ contains_insync e" by auto with IH[of Vs E xs] \fv (a\i\:=e) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val a\ \extTA2J0 P,P,t \ \a,s\ -ta\ \a',s'\\ \length Vs + max_vars (a\i\:=e) \ length xs\ \?synth (a\i\:=e) s\ obtain ta' e2' xs' where IH': "sim_move01 (compP1 P) t ta a E (hp s) xs ta' e2' (hp s') xs'" "bisim Vs a' e2' xs'" "lcl s' \\<^sub>m [Vs [\] xs']" by auto show ?case proof(cases "is_val a'") case True from \fv (a\i\:=e) \ set Vs\ sync have "bisim Vs i (compE1 Vs i) xs'" "bisim Vs e (compE1 Vs e) xs'" by auto with IH' E2 True sync \\ is_val a\ \extTA2J0 P,P,t \ \a,s\ -ta\ \a',s'\\ show ?thesis by(cases "is_val i")(fastforce elim!: sim_move01_expr)+ next case False with IH' E2 sync \\ is_val a\ \extTA2J0 P,P,t \ \a,s\ -ta\ \a',s'\\ show ?thesis by(fastforce elim!: sim_move01_expr) qed next case (AAssRed2 i s ta i' s' a e Vs E2 xs) note IH = \\vs e2 XS. \bisim vs i e2 XS; fv i \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars i \ length XS; ?synth i s \ \ ?concl i e2 s XS i' s' ta vs\ from \extTA2J0 P,P,t \ \i,s\ -ta\ \i',s'\\ have "\ is_val i" by auto with \bisim Vs (Val a\i\ := e) E2 xs\ obtain E where "E2 = Val a\E\:=compE1 Vs e" and "bisim Vs i E xs" and "\ contains_insync e" by auto with IH[of Vs E xs] \fv (Val a\i\:=e) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val i\ \extTA2J0 P,P,t \ \i,s\ -ta\ \i',s'\\ \length Vs + max_vars (Val a\i\:=e) \ length xs\ \?synth (Val a\i\:=e) s\ show ?case by(cases "is_val i'")(fastforce elim!: sim_move01_expr)+ next case (AAssRed3 e s ta e' s' a i Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \bisim Vs (Val a\Val i\ := e) E2 xs\ obtain E where "E2 = Val a\Val i\:=E" and "bisim Vs e E xs" by auto with IH[of Vs E xs] \fv (Val a\Val i\:=e) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_vars (Val a\Val i\:=e) \ length xs\ \?synth (Val a\Val i\:=e) s\ show ?case by(fastforce elim!: sim_move01_expr) next case RedAAssStore thus ?case by(auto intro!: exI) next case (FAssRed1 e s ta e' s' F D e2 Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs (e\F{D} := e2) E2 xs\ obtain E where "E2 = E\F{D} := compE1 Vs e2" and "bisim Vs e E xs" and "\ contains_insync e2" by auto with IH[of Vs E xs] \fv (e\F{D} := e2) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val e\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_vars (e\F{D} := e2) \ length xs\ \?synth (e\F{D} := e2) s\ show ?case by(cases "is_val e'")(fastforce elim!: sim_move01_expr)+ next case (FAssRed2 e s ta e' s' v F D Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \bisim Vs (Val v\F{D} := e) E2 xs\ obtain E where "E2 = Val v\F{D} := E" and "bisim Vs e E xs" by auto with IH[of Vs E xs] \fv (Val v\F{D} := e) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_vars (Val v\F{D} := e) \ length xs\ \?synth (Val v\F{D} := e) s\ show ?case by(fastforce elim!: sim_move01_expr) next case (CASRed1 e s ta e' s' D F e2 e3 Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs _ E2 xs\ obtain E where E2: "E2 = E\compareAndSwap(D\F, compE1 Vs e2, compE1 Vs e3)" and "bisim Vs e E xs" and sync: "\ contains_insync e2" "\ contains_insync e3" by(auto) with IH[of Vs E xs] \fv _ \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val e\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_vars _ \ length xs\ \?synth _ s\ obtain ta' e2' xs' where IH': "sim_move01 (compP1 P) t ta e E (hp s) xs ta' e2' (hp s') xs'" "bisim Vs e' e2' xs'" "lcl s' \\<^sub>m [Vs [\] xs']" by auto show ?case proof(cases "is_val e'") case True from \fv _ \ set Vs\ sync have "bisim Vs e2 (compE1 Vs e2) xs'" "bisim Vs e3 (compE1 Vs e3) xs'" by auto with IH' E2 True sync \\ is_val e\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ show ?thesis by(cases "is_val e2")(fastforce elim!: sim_move01_expr)+ next case False with IH' E2 sync \\ is_val e\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ show ?thesis by(fastforce elim!: sim_move01_expr) qed next case (CASRed2 e s ta e' s' v D F e3 Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs _ E2 xs\ obtain E where "E2 = Val v\compareAndSwap(D\F, E, compE1 Vs e3)" and "bisim Vs e E xs" and "\ contains_insync e3" by(auto) with IH[of Vs E xs] \fv _ \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \\ is_val e\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_vars _ \ length xs\ \?synth _ s\ show ?case by(cases "is_val e'")(fastforce elim!: sim_move01_expr)+ next case (CASRed3 e s ta e' s' v D F v' Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \bisim Vs _ E2 xs\ obtain E where "E2 = Val v\compareAndSwap(D\F, Val v', E)" and "bisim Vs e E xs" by auto with IH[of Vs E xs] \fv _ \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_vars _ \ length xs\ \?synth _ s\ show ?case by(fastforce elim!: sim_move01_expr) next case (CallObj e s ta e' s' M es Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs (e\M(es)) E2 xs\ obtain E where "E2 = E\M(compEs1 Vs es)" and "bisim Vs e E xs" and "\ contains_insyncs es" by(auto simp add: compEs1_conv_map) with IH[of Vs E xs] \fv (e\M(es)) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars (e\M(es)) \ length xs\ \?synth (e\M(es)) s\ show ?case by(cases "is_val e'")(fastforce elim!: sim_move01_expr split: if_split_asm)+ next case (CallParams es s ta es' s' v M Vs E2 xs) note IH = \\vs es2 XS. \bisims vs es es2 XS; fvs es \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_varss es \ length XS; ?synths es s \ \ ?concls es es2 s XS es' s' ta vs\ from \bisim Vs (Val v\M(es)) E2 xs\ obtain Es where "E2 = Val v\M(Es)" and "bisims Vs es Es xs" by auto moreover from \extTA2J0 P,P,t \ \es,s\ [-ta\] \es',s'\\ have "\ is_vals es" by auto with \?synth (Val v\M(es)) s\ have "?synths es s" by(auto) moreover note IH[of Vs Es xs] \fv (Val v\M(es)) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars (Val v\M(es)) \ length xs\ ultimately show ?case by(fastforce elim!: sim_move01_CallParams) next case (RedCall s a U M Ts T pns body D vs Vs E2 xs) from \typeof_addr (hp s) a = \U\\ have "call (addr a\M(map Val vs)) = \(a, M, vs)\" by auto with \?synth (addr a\M(map Val vs)) s\ have "synthesized_call P (hp s) (a, M, vs)" by auto with \typeof_addr (hp s) a = \U\\ \P \ class_type_of U sees M: Ts\T = \(pns, body)\ in D\ have False by(auto simp add: synthesized_call_conv dest: sees_method_fun) thus ?case .. next case (RedCallExternal s a T M Ts Tr D vs ta va h' ta' e' s' Vs E2 xs) from \bisim Vs (addr a\M(map Val vs)) E2 xs\ have "E2 = addr a\M(map Val vs)" by auto moreover note \P \ class_type_of T sees M: Ts\Tr = Native in D\ \typeof_addr (hp s) a = \T\\ \ta' = extTA2J0 P ta\ \e' = extRet2J (addr a\M(map Val vs)) va\ \s' = (h', lcl s)\ \P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\\ \lcl s \\<^sub>m [Vs [\] xs]\ moreover from wf \P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\\ have "ta_bisim01 (extTA2J0 P ta) (extTA2J1 (compP1 P) ta)" by(rule red_external_ta_bisim01) moreover from \P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\\ \typeof_addr (hp s) a = \T\\ \P \ class_type_of T sees M: Ts\Tr = Native in D\ have "\external_defs D M \ h' = hp s \ ta = \" by(fastforce dest: \external'_red_external_heap_unchanged \external'_red_external_TA_empty simp add: \external'_def \external_def) ultimately show ?case by(cases va)(fastforce intro!: exI[where x=ta] intro: Red1CallExternal simp add: map_eq_append_conv sim_move01_def dest: sees_method_fun simp del: split_paired_Ex)+ next case (BlockRed e h x V vo ta e' h' x' T Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl (h, x(V := vo)) \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e (h, (x(V := vo)))\ \ ?concl e e2 (h, x(V := vo)) XS e' (h', x') ta vs\ note red = \extTA2J0 P,P,t \ \e,(h, x(V := vo))\ -ta\ \e',(h', x')\\ note len = \length Vs + max_vars {V:T=vo; e} \ length xs\ from \fv {V:T=vo; e} \ set Vs\ have fv: "fv e \ set (Vs@[V])" by auto from \bisim Vs {V:T=vo; e} E2 xs\ show ?case proof(cases rule: bisim_cases(7)[consumes 1, case_names BlockNone BlockSome BlockSomeNone]) case (BlockNone E') with red IH[of "Vs@[V]" E' xs] fv \lcl (h, x) \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars {V:T=vo; e} \ length xs\ \?synth {V:T=vo; e} (h, x)\ obtain TA' e2' xs' where red': "sim_move01 (compP1 P) t ta e E' h xs TA' e2' h' xs'" and bisim': "bisim (Vs @ [V]) e' e2' xs'" "x' \\<^sub>m [Vs @ [V] [\] xs']" by auto from red' \length Vs + max_vars {V:T=vo; e} \ length xs\ have "length (Vs@[V]) + max_vars e \ length xs'" by(fastforce simp add: sim_move01_def dest: red1_preserves_len \red1t_preserves_len \red1r_preserves_len split: if_split_asm) with \x' \\<^sub>m [Vs @ [V] [\] xs']\ have "x' \\<^sub>m [Vs [\] xs', V \ xs' ! length Vs]" by(simp) moreover { assume "V \ set Vs" hence "hidden (Vs @ [V]) (index Vs V)" by(rule hidden_index) with \bisim (Vs @ [V]) e E' xs\ have "unmod E' (index Vs V)" by -(rule hidden_bisim_unmod) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ \V \ set Vs\ have "index Vs V < length xs" by(auto intro: index_less_aux) ultimately have "xs ! index Vs V = xs' ! index Vs V" using sim_move01_preserves_unmod[OF red'] by(simp) } moreover from red' have "length xs = length xs'" by(rule sim_move01_preserves_len[symmetric]) ultimately have rel: "x'(V := x V) \\<^sub>m [Vs [\] xs']" using \lcl (h, x) \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars {V:T=vo; e} \ length xs\ by(auto intro: Block_lem) show ?thesis proof(cases "x' V") case None with red' bisim' BlockNone len show ?thesis by(fastforce simp del: split_paired_Ex fun_upd_apply intro: rel) next case (Some v) moreover with \x' \\<^sub>m [Vs @ [V] [\] xs']\ have "[Vs @ [V] [\] xs'] V = \v\" by(auto simp add: map_le_def dest: bspec) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ have "length Vs < length xs" by auto ultimately have "xs' ! length Vs = v" using \length xs = length xs'\ by(simp) with red' bisim' BlockNone Some len show ?thesis by(fastforce simp del: fun_upd_apply intro: rel) qed next case (BlockSome E' v) with red IH[of "Vs@[V]" E' "xs[length Vs := v]"] fv \lcl (h, x) \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars {V:T=vo; e} \ length xs\ \?synth {V:T=vo; e} (h, x)\ obtain TA' e2' xs' where red': "sim_move01 (compP1 P) t ta e E' h (xs[length Vs := v]) TA' e2' h' xs'" and bisim': "bisim (Vs @ [V]) e' e2' xs'" "x' \\<^sub>m [Vs @ [V] [\] xs']" by auto from red' \length Vs + max_vars {V:T=vo; e} \ length xs\ have "length (Vs@[V]) + max_vars e \ length xs'" by(auto dest: sim_move01_preserves_len) with \x' \\<^sub>m [Vs @ [V] [\] xs']\ have "x' \\<^sub>m [Vs [\] xs', V \ xs' ! length Vs]" by(simp) moreover { assume "V \ set Vs" hence "hidden (Vs @ [V]) (index Vs V)" by(rule hidden_index) with \bisim (Vs @ [V]) e E' (xs[length Vs := v])\ have "unmod E' (index Vs V)" by -(rule hidden_bisim_unmod) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ \V \ set Vs\ have "index Vs V < length xs" by(auto intro: index_less_aux) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ \V \ set Vs\ have "(xs[length Vs := v]) ! index Vs V = xs ! index Vs V" by(simp) ultimately have "xs ! index Vs V = xs' ! index Vs V" using sim_move01_preserves_unmod[OF red', of "index Vs V"] by(simp) } moreover from red' have "length xs = length xs'" by(auto dest: sim_move01_preserves_len) ultimately have rel: "x'(V := x V) \\<^sub>m [Vs [\] xs']" using \lcl (h, x) \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars {V:T=vo; e} \ length xs\ by(auto intro: Block_lem) from BlockSome red obtain v' where Some: "x' V = \v'\" by(auto dest!: red_lcl_incr) with \x' \\<^sub>m [Vs @ [V] [\] xs']\ have "[Vs @ [V] [\] xs'] V = \v'\" by(auto simp add: map_le_def dest: bspec) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ have "length Vs < length xs" by auto ultimately have "xs' ! length Vs = v'" using \length xs = length xs'\ by(simp) with red' bisim' BlockSome Some \length Vs < length xs\ show ?thesis by(fastforce simp del: fun_upd_apply intro: rel) next case (BlockSomeNone E') with red IH[of "Vs@[V]" E' xs] fv \lcl (h, x) \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars {V:T=vo; e} \ length xs\ \?synth {V:T=vo; e} (h, x)\ obtain TA' e2' xs' where red': "sim_move01 (compP1 P) t ta e E' h xs TA' e2' h' xs'" and IH': "bisim (Vs @ [V]) e' e2' xs'" "x' \\<^sub>m [Vs @ [V] [\] xs']" by auto from red' \length Vs + max_vars {V:T=vo; e} \ length xs\ have "length (Vs@[V]) + max_vars e \ length xs'" by(auto dest: sim_move01_preserves_len) with \x' \\<^sub>m [Vs @ [V] [\] xs']\ have "x' \\<^sub>m [Vs [\] xs', V \ xs' ! length Vs]" by(simp) moreover { assume "V \ set Vs" hence "hidden (Vs @ [V]) (index Vs V)" by(rule hidden_index) with \bisim (Vs @ [V]) e E' xs\ have "unmod E' (index Vs V)" by -(rule hidden_bisim_unmod) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ \V \ set Vs\ have "index Vs V < length xs" by(auto intro: index_less_aux) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ \V \ set Vs\ have "(xs[length Vs := v]) ! index Vs V = xs ! index Vs V" by(simp) ultimately have "xs ! index Vs V = xs' ! index Vs V" using sim_move01_preserves_unmod[OF red', of "index Vs V"] by(simp) } moreover from red' have "length xs = length xs'" by(auto dest: sim_move01_preserves_len) ultimately have rel: "x'(V := x V) \\<^sub>m [Vs [\] xs']" using \lcl (h, x) \\<^sub>m [Vs [\] xs]\ \length Vs + max_vars {V:T=vo; e} \ length xs\ by(auto intro: Block_lem) from BlockSomeNone red obtain v' where Some: "x' V = \v'\" by(auto dest!: red_lcl_incr) with \x' \\<^sub>m [Vs @ [V] [\] xs']\ have "[Vs @ [V] [\] xs'] V = \v'\" by(auto simp add: map_le_def dest: bspec) moreover from \length Vs + max_vars {V:T=vo; e} \ length xs\ have "length Vs < length xs" by auto ultimately have "xs' ! length Vs = v'" using \length xs = length xs'\ by(simp) with red' IH' BlockSomeNone Some \length Vs < length xs\ show ?thesis by(fastforce simp del: fun_upd_apply intro: rel) qed next case (RedBlock V T vo u s Vs E2 xs) from \bisim Vs {V:T=vo; Val u} E2 xs\ obtain vo' where [simp]: "E2 = {length Vs:T=vo'; Val u}" by auto from RedBlock show ?case proof(cases vo) case (Some v) with \bisim Vs {V:T=vo; Val u} E2 xs\ have vo': "case vo' of None \ xs ! length Vs = v | Some v' \ v = v'" by auto have "sim_move01 (compP1 P) t \ {V:T=vo; Val u} E2 (hp s) xs \ (Val u) (hp s) (xs[length Vs := v])" proof(cases vo') case None with vo' have "xs[length Vs := v] = xs" by auto thus ?thesis using Some None by auto next case Some from \length Vs + max_vars {V:T=vo; Val u} \ length xs\ have "length Vs < length xs" by simp with vo' Some show ?thesis using \vo = Some v\ by auto qed thus ?thesis using RedBlock by fastforce qed fastforce next case SynchronizedNull thus ?case by fastforce next case (LockSynchronized a e s Vs E2 xs) from \bisim Vs (sync(addr a) e) E2 xs\ have E2: "E2 = sync\<^bsub>length Vs\<^esub> (addr a) (compE1 (Vs@[fresh_var Vs]) e)" and sync: "\ contains_insync e" by auto moreover have "fresh_var Vs \ set Vs" by(rule fresh_var_fresh) with \fv (sync(addr a) e) \ set Vs\ have "fresh_var Vs \ fv e" by auto from E2 \fv (sync(addr a) e) \ set Vs\ sync have "bisim (Vs@[fresh_var Vs]) e (compE1 (Vs@[fresh_var Vs]) e) (xs[length Vs := Addr a])" by(auto intro!: compE1_bisim) hence "bisim Vs (insync(a) e) (insync\<^bsub>length Vs\<^esub> (a) (compE1 (Vs@[fresh_var Vs]) e)) (xs[length Vs := Addr a])" using \fresh_var Vs \ fv e\ \length Vs + max_vars (sync(addr a) e) \ length xs\ by auto moreover from \length Vs + max_vars (sync(addr a) e) \ length xs\ have "False,compP1 P,t \1 \sync\<^bsub>length Vs\<^esub> (addr a) (compE1 (Vs@[fresh_var Vs]) e), (hp s, xs)\ -\Lock\a, SyncLock a\\ \insync\<^bsub>length Vs\<^esub> (a) (compE1 (Vs@[fresh_var Vs]) e), (hp s, xs[length Vs := Addr a])\" by -(rule Lock1Synchronized, auto) hence "sim_move01 (compP1 P) t \Lock\a, SyncLock a\ (sync(addr a) e) E2 (hp s) xs \Lock\a, SyncLock a\ (insync\<^bsub>length Vs\<^esub> (a) (compE1 (Vs@[fresh_var Vs]) e)) (hp s) (xs[length Vs := Addr a])" using E2 by(fastforce simp add: sim_move01_def ta_bisim_def) moreover have "zip Vs (xs[length Vs := Addr a]) = (zip Vs xs)[length Vs := (arbitrary, Addr a)]" by(rule sym)(simp add: update_zip) hence "zip Vs (xs[length Vs := Addr a]) = zip Vs xs" by simp with \lcl s \\<^sub>m [Vs [\] xs]\ have "lcl s \\<^sub>m [Vs [\] xs[length Vs := Addr a]]" by(auto simp add: map_le_def map_upds_def) ultimately show ?case using \lcl s \\<^sub>m [Vs [\] xs]\ by fastforce next case (SynchronizedRed2 e s ta e' s' a Vs E2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s \ \ ?concl e e2 s XS e' s' ta vs\ from \bisim Vs (insync(a) e) E2 xs\ obtain E where E2: "E2 = insync\<^bsub>length Vs\<^esub> (a) E" and bisim: "bisim (Vs@[fresh_var Vs]) e E xs" and xsa: "xs ! length Vs = Addr a" by auto from \fv (insync(a) e) \ set Vs\ fresh_var_fresh[of Vs] have fv: "fresh_var Vs \ fv e" by auto from \length Vs + max_vars (insync(a) e) \ length xs\ have "length Vs < length xs" by simp { assume "lcl s (fresh_var Vs) \ None" then obtain v where "lcl s (fresh_var Vs) = \v\" by auto with \lcl s \\<^sub>m [Vs [\] xs]\ have "[Vs [\] xs] (fresh_var Vs) = \v\" by(auto simp add: map_le_def dest: bspec) hence "fresh_var Vs \ set Vs" by(auto simp add: map_upds_def set_zip dest!: map_of_SomeD ) moreover have "fresh_var Vs \ set Vs" by(rule fresh_var_fresh) ultimately have False by contradiction } hence "lcl s (fresh_var Vs) = None" by(cases "lcl s (fresh_var Vs)", auto) hence "(lcl s)(fresh_var Vs := None) = lcl s" by(auto intro: ext) moreover from \lcl s \\<^sub>m [Vs [\] xs]\ have "(lcl s)(fresh_var Vs := None) \\<^sub>m [Vs [\] xs, fresh_var Vs \ xs ! length Vs]" by(simp) ultimately have "lcl s \\<^sub>m [Vs @ [fresh_var Vs] [\] xs]" using \length Vs < length xs\ by(auto) with IH[of "Vs@[fresh_var Vs]" E xs] \fv (insync(a) e) \ set Vs\ bisim \length Vs + max_vars (insync(a) e) \ length xs\ \?synth (insync(a) e) s\ obtain TA' e2' xs' where IH': "sim_move01 (compP1 P) t ta e E (hp s) xs TA' e2' (hp s') xs'" "bisim (Vs @ [fresh_var Vs]) e' e2' xs'" "lcl s' \\<^sub>m [Vs @ [fresh_var Vs] [\] xs']" by auto from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "dom (lcl s') \ dom (lcl s) \ fv e" by(auto dest: red_dom_lcl) with fv \lcl s (fresh_var Vs) = None\ have "(fresh_var Vs) \ dom (lcl s')" by auto hence "lcl s' (fresh_var Vs) = None" by auto moreover from IH' have "length xs = length xs'" by(auto dest: sim_move01_preserves_len) moreover note \lcl s' \\<^sub>m [Vs @ [fresh_var Vs] [\] xs']\ \length Vs < length xs\ ultimately have "lcl s' \\<^sub>m [Vs [\] xs']" by(auto simp add: map_le_def dest: bspec) moreover from bisim fv have "unmod E (length Vs)" by(auto intro: bisim_fv_unmod) with IH' \length Vs < length xs\ have "xs ! length Vs = xs' ! length Vs" by(auto dest!: sim_move01_preserves_unmod) with xsa have "xs' ! length Vs = Addr a" by simp ultimately show ?case using IH' E2 by(fastforce) next case (UnlockSynchronized a v s Vs E2 xs) from \bisim Vs (insync(a) Val v) E2 xs\ have E2: "E2 = insync\<^bsub>length Vs\<^esub> (a) Val v" and xsa: "xs ! length Vs = Addr a" by auto moreover from \length Vs + max_vars (insync(a) Val v) \ length xs\ xsa have "False,compP1 P,t \1 \insync\<^bsub>length Vs\<^esub> (a) (Val v), (hp s, xs)\ -\Unlock\a, SyncUnlock a\\ \Val v, (hp s, xs)\" by-(rule Unlock1Synchronized, auto) hence "sim_move01 (compP1 P) t \Unlock\a, SyncUnlock a\ (insync(a) Val v) (insync\<^bsub>length Vs\<^esub> (a) Val v) (hp s) xs \Unlock\a, SyncUnlock a\ (Val v) (hp s) xs" by(fastforce simp add: sim_move01_def ta_bisim_def) ultimately show ?case using \lcl s \\<^sub>m [Vs [\] xs]\ by fastforce next case (RedWhile b c s Vs E2 xs) from \bisim Vs (while (b) c) E2 xs\ have "E2 = while (compE1 Vs b) (compE1 Vs c)" and sync: "\ contains_insync b" "\ contains_insync c" by auto moreover have "False,compP1 P,t \1 \while(compE1 Vs b) (compE1 Vs c), (hp s, xs)\ -\\ \if (compE1 Vs b) (compE1 Vs c;;while(compE1 Vs b) (compE1 Vs c)) else unit, (hp s, xs)\" by(rule Red1While) hence "sim_move01 (compP1 P) t \ (while (b) c) (while (compE1 Vs b) (compE1 Vs c)) (hp s) xs \ (if (compE1 Vs b) (compE1 Vs c;;while(compE1 Vs b) (compE1 Vs c)) else unit) (hp s) xs" by(auto simp add: sim_move01_def) moreover from \fv (while (b) c) \ set Vs\ sync have "bisim Vs (if (b) (c;; while (b) c) else unit) (if (compE1 Vs b) (compE1 Vs (c;; while(b) c)) else (compE1 Vs unit)) xs" by -(rule bisimCond, auto) ultimately show ?case using \lcl s \\<^sub>m [Vs [\] xs]\ by(simp)((rule exI)+, erule conjI, auto) next case (RedTryCatch s a D C V e2 Vs E2 xs) thus ?case by(auto intro!: exI)(auto intro!: compE1_bisim) next case RedTryFail thus ?case by(auto intro!: exI) next case (ListRed1 e s ta e' s' es Vs ES2 xs) note IH = \\vs e2 XS. \bisim vs e e2 XS; fv e \ set vs; lcl s \\<^sub>m [vs [\] XS]; length vs + max_vars e \ length XS; ?synth e s\ \ ?concl e e2 s XS e' s' ta vs\ from \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisims Vs (e # es) ES2 xs\ obtain E' where "bisim Vs e E' xs" and ES2: "ES2 = E' # compEs1 Vs es" and sync: "\ contains_insyncs es" by(auto simp add: compEs1_conv_map) with IH[of Vs E' xs] \fvs (e # es) \ set Vs\ \lcl s \\<^sub>m [Vs [\] xs]\ \extTA2J0 P,P,t \ \e,s\ -ta\ \e',s'\\ \length Vs + max_varss (e # es) \ length xs\ \?synths (e # es) s\ \\ is_val e\ show ?case by(cases "is_val e'")(fastforce elim!: sim_moves01_expr split: if_split_asm)+ next case (ListRed2 es s ta es' s' v Vs ES2 xs) thus ?case by(fastforce elim!: bisims_cases elim!: sim_moves01_expr) next case CallThrowParams thus ?case by(fastforce elim!:bisim_cases simp add: bisims_map_Val_Throw) next case (BlockThrow V T vo a s Vs e2 xs) thus ?case by(cases vo)(fastforce elim!: bisim_cases)+ next case (SynchronizedThrow2 a ad s Vs E2 xs) from \bisim Vs (insync(a) Throw ad) E2 xs\ have "xs ! length Vs = Addr a" by auto with \length Vs + max_vars (insync(a) Throw ad) \ length xs\ have "False,compP1 P,t \1 \insync\<^bsub>length Vs\<^esub> (a) Throw ad, (hp s, xs)\ -\Unlock\a, SyncUnlock a\\ \Throw ad, (hp s, xs)\" by-(rule Synchronized1Throw2, auto) hence "sim_move01 (compP1 P) t \Unlock\a, SyncUnlock a\ (insync(a) Throw ad) (insync\<^bsub>length Vs\<^esub> (a) Throw ad) (hp s) xs \Unlock\a, SyncUnlock a\ (Throw ad) (hp s) xs" by(fastforce simp add: sim_move01_def ta_bisim_def fun_eq_iff expand_finfun_eq finfun_upd_apply ta_upd_simps split: if_split_asm) moreover note \lcl s \\<^sub>m [Vs [\] xs]\ \bisim Vs (insync(a) Throw ad) E2 xs\ ultimately show ?case by(fastforce) next case InstanceOfRed thus ?case by(fastforce) next case RedInstanceOf thus ?case by(auto intro!: exI) next case InstanceOfThrow thus ?case by fastforce qed(fastforce simp del: fun_upd_apply split: if_split_asm)+ end declare max_dest [iff del] declare split_paired_Ex [simp del] primrec countInitBlock :: "('a, 'b, 'addr) exp \ nat" and countInitBlocks :: "('a, 'b, 'addr) exp list \ nat" where "countInitBlock (new C) = 0" | "countInitBlock (newA T\e\) = countInitBlock e" | "countInitBlock (Cast T e) = countInitBlock e" | "countInitBlock (e instanceof T) = countInitBlock e" | "countInitBlock (Val v) = 0" | "countInitBlock (Var V) = 0" | "countInitBlock (V := e) = countInitBlock e" | "countInitBlock (e \bop\ e') = countInitBlock e + countInitBlock e'" | "countInitBlock (a\i\) = countInitBlock a + countInitBlock i" | "countInitBlock (AAss a i e) = countInitBlock a + countInitBlock i + countInitBlock e" | "countInitBlock (a\length) = countInitBlock a" | "countInitBlock (e\F{D}) = countInitBlock e" | "countInitBlock (FAss e F D e') = countInitBlock e + countInitBlock e'" | "countInitBlock (e\compareAndSwap(D\F, e', e'')) = countInitBlock e + countInitBlock e' + countInitBlock e''" | "countInitBlock (e\M(es)) = countInitBlock e + countInitBlocks es" | "countInitBlock ({V:T=vo; e}) = (case vo of None \ 0 | Some v \ 1) + countInitBlock e" | "countInitBlock (sync\<^bsub>V'\<^esub> (e) e') = countInitBlock e + countInitBlock e'" | "countInitBlock (insync\<^bsub>V'\<^esub> (ad) e) = countInitBlock e" | "countInitBlock (e;;e') = countInitBlock e + countInitBlock e'" | "countInitBlock (if (e) e1 else e2) = countInitBlock e + countInitBlock e1 + countInitBlock e2" | "countInitBlock (while(b) e) = countInitBlock b + countInitBlock e" | "countInitBlock (throw e) = countInitBlock e" | "countInitBlock (try e catch(C V) e') = countInitBlock e + countInitBlock e'" | "countInitBlocks [] = 0" | "countInitBlocks (e # es) = countInitBlock e + countInitBlocks es" context J0_J1_heap_base begin lemmas \red0r_expr = NewArray_\red0r_xt Cast_\red0r_xt InstanceOf_\red0r_xt BinOp_\red0r_xt1 BinOp_\red0r_xt2 LAss_\red0r AAcc_\red0r_xt1 AAcc_\red0r_xt2 AAss_\red0r_xt1 AAss_\red0r_xt2 AAss_\red0r_xt3 ALength_\red0r_xt FAcc_\red0r_xt FAss_\red0r_xt1 FAss_\red0r_xt2 CAS_\red0r_xt1 CAS_\red0r_xt2 CAS_\red0r_xt3 Call_\red0r_obj Call_\red0r_param Block_\red0r_xt Sync_\red0r_xt InSync_\red0r_xt Seq_\red0r_xt Cond_\red0r_xt Throw_\red0r_xt Try_\red0r_xt lemmas \red0t_expr = NewArray_\red0t_xt Cast_\red0t_xt InstanceOf_\red0t_xt BinOp_\red0t_xt1 BinOp_\red0t_xt2 LAss_\red0t AAcc_\red0t_xt1 AAcc_\red0t_xt2 AAss_\red0t_xt1 AAss_\red0t_xt2 AAss_\red0t_xt3 ALength_\red0t_xt FAcc_\red0t_xt FAss_\red0t_xt1 FAss_\red0t_xt2 CAS_\red0t_xt1 CAS_\red0t_xt2 CAS_\red0t_xt3 Call_\red0t_obj Call_\red0t_param Block_\red0t_xt Sync_\red0t_xt InSync_\red0t_xt Seq_\red0t_xt Cond_\red0t_xt Throw_\red0t_xt Try_\red0t_xt declare \red0r_expr [elim!] declare \red0t_expr [elim!] definition sim_move10 :: "'addr J_prog \ 'thread_id \ ('addr, 'thread_id, 'heap) external_thread_action \ 'addr expr1 \ 'addr expr1 \ 'addr expr \ 'heap \ 'addr locals \ ('addr, 'thread_id, 'heap) J0_thread_action \ 'addr expr \ 'heap \ 'addr locals \ bool" where "sim_move10 P t ta1 e1 e1' e h xs ta e' h' xs' \ \ final e1 \ (if \move1 P h e1 then (\red0t (extTA2J0 P) P t h (e, xs) (e', xs') \ countInitBlock e1' < countInitBlock e1 \ e' = e \ xs' = xs) \ h' = h \ ta1 = \ \ ta = \ else ta_bisim01 ta (extTA2J1 (compP1 P) ta1) \ (if call e = None \ call1 e1 = None then (\e'' xs''. \red0r (extTA2J0 P) P t h (e, xs) (e'', xs'') \ extTA2J0 P,P,t \ \e'', (h, xs'')\ -ta\ \e', (h', xs')\ \ no_call P h e'' \ \ \move0 P h e'') else extTA2J0 P,P,t \ \e, (h, xs)\ -ta\ \e', (h', xs')\ \ no_call P h e \ \ \move0 P h e))" definition sim_moves10 :: "'addr J_prog \ 'thread_id \ ('addr, 'thread_id, 'heap) external_thread_action \ 'addr expr1 list \ 'addr expr1 list \ 'addr expr list \ 'heap \ 'addr locals \ ('addr, 'thread_id, 'heap) J0_thread_action \ 'addr expr list \ 'heap \ 'addr locals \ bool" where "sim_moves10 P t ta1 es1 es1' es h xs ta es' h' xs' \ \ finals es1 \ (if \moves1 P h es1 then (\reds0t (extTA2J0 P) P t h (es, xs) (es', xs') \ countInitBlocks es1' < countInitBlocks es1 \ es' = es \ xs' = xs) \ h' = h \ ta1 = \ \ ta = \ else ta_bisim01 ta (extTA2J1 (compP1 P) ta1) \ (if calls es = None \ calls1 es1 = None then (\es'' xs''. \reds0r (extTA2J0 P) P t h (es, xs) (es'', xs'') \ extTA2J0 P,P,t \ \es'', (h, xs'')\ [-ta\] \es', (h', xs')\ \ no_calls P h es'' \ \ \moves0 P h es'') else extTA2J0 P,P,t \ \es, (h, xs)\ [-ta\] \es', (h', xs')\ \ no_calls P h es \ \ \moves0 P h es))" lemma sim_move10_expr: assumes "sim_move10 P t ta1 e1 e1' e h xs ta e' h' xs'" shows "sim_move10 P t ta1 (newA T\e1\) (newA T\e1'\) (newA T\e\) h xs ta (newA T\e'\) h' xs'" "sim_move10 P t ta1 (Cast T e1) (Cast T e1') (Cast T e) h xs ta (Cast T e') h' xs'" "sim_move10 P t ta1 (e1 instanceof T) (e1' instanceof T) (e instanceof T) h xs ta (e' instanceof T) h' xs'" "sim_move10 P t ta1 (e1 \bop\ e2) (e1' \bop\ e2) (e \bop\ e2') h xs ta (e' \bop\ e2') h' xs'" "sim_move10 P t ta1 (Val v \bop\ e1) (Val v \bop\ e1') (Val v \bop\ e) h xs ta (Val v \bop\ e') h' xs'" "sim_move10 P t ta1 (V := e1) (V := e1') (V' := e) h xs ta (V' := e') h' xs'" "sim_move10 P t ta1 (e1\e2\) (e1'\e2\) (e\e2'\) h xs ta (e'\e2'\) h' xs'" "sim_move10 P t ta1 (Val v\e1\) (Val v\e1'\) (Val v\e\) h xs ta (Val v\e'\) h' xs'" "sim_move10 P t ta1 (e1\e2\ := e3) (e1'\e2\ := e3) (e\e2'\ := e3') h xs ta (e'\e2'\ := e3') h' xs'" "sim_move10 P t ta1 (Val v\e1\ := e3) (Val v\e1'\ := e3) (Val v\e\ := e3') h xs ta (Val v\e'\ := e3') h' xs'" "sim_move10 P t ta1 (AAss (Val v) (Val v') e1) (AAss (Val v) (Val v') e1') (AAss (Val v) (Val v') e) h xs ta (AAss (Val v) (Val v') e') h' xs'" "sim_move10 P t ta1 (e1\length) (e1'\length) (e\length) h xs ta (e'\length) h' xs'" "sim_move10 P t ta1 (e1\F{D}) (e1'\F{D}) (e\F'{D'}) h xs ta (e'\F'{D'}) h' xs'" "sim_move10 P t ta1 (FAss e1 F D e2) (FAss e1' F D e2) (FAss e F' D' e2') h xs ta (FAss e' F' D' e2') h' xs'" "sim_move10 P t ta1 (FAss (Val v) F D e1) (FAss (Val v) F D e1') (FAss (Val v) F' D' e) h xs ta (FAss (Val v) F' D' e') h' xs'" "sim_move10 P t ta1 (CompareAndSwap e1 F D e2 e3) (CompareAndSwap e1' F D e2 e3) (CompareAndSwap e F' D' e2' e3') h xs ta (CompareAndSwap e' F' D' e2' e3') h' xs'" "sim_move10 P t ta1 (CompareAndSwap (Val v) F D e1 e3) (CompareAndSwap (Val v) F D e1' e3) (CompareAndSwap (Val v) F' D' e e3') h xs ta (CompareAndSwap (Val v) F' D' e' e3') h' xs'" "sim_move10 P t ta1 (CompareAndSwap (Val v) F D (Val v') e1) (CompareAndSwap (Val v) F D (Val v') e1') (CompareAndSwap (Val v) F' D' (Val v') e) h xs ta (CompareAndSwap (Val v) F' D' (Val v') e') h' xs'" "sim_move10 P t ta1 (e1\M(es)) (e1'\M(es)) (e\M(es')) h xs ta (e'\M(es')) h' xs'" "sim_move10 P t ta1 (sync\<^bsub>V\<^esub>(e1) e2) (sync\<^bsub>V\<^esub>(e1') e2) (sync(e) e2') h xs ta (sync(e') e2') h' xs'" "sim_move10 P t ta1 (insync\<^bsub>V\<^esub>(a) e1) (insync\<^bsub>V\<^esub>(a) e1') (insync(a') e) h xs ta (insync(a') e') h' xs'" "sim_move10 P t ta1 (e1;;e2) (e1';;e2) (e;;e2') h xs ta (e';;e2') h' xs'" "sim_move10 P t ta1 (if (e1) e2 else e3) (if (e1') e2 else e3) (if (e) e2' else e3') h xs ta (if (e') e2' else e3') h' xs'" "sim_move10 P t ta1 (throw e1) (throw e1') (throw e) h xs ta (throw e') h' xs'" "sim_move10 P t ta1 (try e1 catch(C V) e2) (try e1' catch(C V) e2) (try e catch(C' V') e2') h xs ta (try e' catch(C' V') e2') h' xs'" using assms apply(simp_all add: sim_move10_def final_iff split del: if_split split: if_split_asm) apply(fastforce simp: \red0t_Val \red0r_Val intro: red_reds.intros split!: if_splits)+ done lemma sim_moves10_expr: "sim_move10 P t ta1 e1 e1' e h xs ta e' h' xs' \ sim_moves10 P t ta1 (e1 # es2) (e1' # es2) (e # es2') h xs ta (e' # es2') h' xs'" "sim_moves10 P t ta1 es1 es1' es h xs ta es' h' xs' \ sim_moves10 P t ta1 (Val v # es1) (Val v # es1') (Val v # es) h xs ta (Val v # es') h' xs'" unfolding sim_moves10_def sim_move10_def final_iff finals_iff apply(simp_all add: Cons_eq_append_conv split del: if_split split: if_split_asm) apply(safe intro!: if_split) apply(fastforce simp add: is_vals_conv \reds0t_map_Val \reds0r_map_Val \red0t_Val \red0r_Val intro!: \red0r_inj_\reds0r \reds0r_cons_\reds0r \red0t_inj_\reds0t \reds0t_cons_\reds0t ListRed1 ListRed2 split: if_split_asm)+ done lemma sim_move10_CallParams: "sim_moves10 P t ta1 es1 es1' es h xs ta es' h' xs' \ sim_move10 P t ta1 (Val v\M(es1)) (Val v\M(es1')) (Val v\M(es)) h xs ta (Val v\M(es')) h' xs'" unfolding sim_move10_def sim_moves10_def apply(simp split: if_split_asm split del: if_split add: is_vals_conv) apply(fastforce simp add: \red0t_Val \red0r_Val \reds0t_map_Val \reds0r_map_Val intro: Call_\red0r_param Call_\red0t_param CallParams split: if_split_asm split del: if_split intro!: if_split) apply(rule conjI) apply fastforce apply(rule if_intro) apply fastforce apply(clarsimp split del: if_split) apply(rule if_intro) apply(clarsimp split: if_split_asm simp add: is_vals_conv) apply(rule exI conjI)+ apply(erule Call_\red0r_param) apply(fastforce intro: CallParams) apply(rule exI conjI)+ apply(erule Call_\red0r_param) apply(fastforce intro!: CallParams) apply(clarsimp split del: if_split split: if_split_asm simp add: is_vals_conv \reds0r_map_Val) apply fastforce apply(rule conjI) apply fastforce apply(rule if_intro) apply fastforce apply(rule conjI) apply fastforce apply(rule if_intro) apply(clarsimp split: if_split_asm) apply(clarsimp split: if_split_asm split del: if_split simp add: is_vals_conv) apply(fastforce intro: CallParams) done lemma sim_move10_Block: "sim_move10 P t ta1 e1 e1' e h (xs(V' := vo)) ta e' h' xs' \ sim_move10 P t ta1 ({V:T=None; e1}) ({V:T=None; e1'}) ({V':T=vo; e}) h xs ta ({V':T=xs' V'; e'}) h' (xs'(V' := xs V'))" proof - assume "sim_move10 P t ta1 e1 e1' e h (xs(V' := vo)) ta e' h' xs'" moreover { fix e'' xs'' assume "extTA2J0 P,P,t \ \e'',(h, xs'')\ -ta\ \e',(h', xs')\" hence "extTA2J0 P,P,t \ \e'',(h, xs''(V' := xs V', V' := xs'' V'))\ -ta\ \e',(h', xs')\" by simp from BlockRed[OF this, of T] have "extTA2J0 P,P,t \ \{V':T=xs'' V'; e''},(h, xs''(V' := xs V'))\ -ta\ \{V':T=xs' V'; e'},(h', xs'(V' := xs V'))\" by simp } ultimately show ?thesis by(fastforce simp add: sim_move10_def final_iff split: if_split_asm) qed lemma sim_move10_reds: "\ (h', a) \ allocate h (Class_type C); ta1 = \NewHeapElem a (Class_type C)\; ta = \NewHeapElem a (Class_type C)\ \ \ sim_move10 P t ta1 (new C) e1' (new C) h xs ta (addr a) h' xs" "allocate h (Class_type C) = {} \ sim_move10 P t \ (new C) e1' (new C) h xs \ (THROW OutOfMemory) h xs" "\ (h', a) \ allocate h (Array_type T (nat (sint i))); 0 <=s i; ta1 = \NewHeapElem a (Array_type T (nat (sint i)))\; ta = \NewHeapElem a (Array_type T (nat (sint i)))\ \ \ sim_move10 P t ta1 (newA T\Val (Intg i)\) e1' (newA T\Val (Intg i)\) h xs ta (addr a) h' xs" "i sim_move10 P t \ (newA T\Val (Intg i)\) e1' (newA T\Val (Intg i)\) h xs \ (THROW NegativeArraySize) h xs" "\ allocate h (Array_type T (nat (sint i))) = {}; 0 <=s i \ \ sim_move10 P t \ (newA T\Val (Intg i)\) e1' (newA T\Val (Intg i)\) h xs \ (THROW OutOfMemory) h xs" "\ typeof\<^bsub>h\<^esub> v = \U\; P \ U \ T \ \ sim_move10 P t \ (Cast T (Val v)) e1' (Cast T (Val v)) h xs \ (Val v) h xs" "\ typeof\<^bsub>h\<^esub> v = \U\; \ P \ U \ T \ \ sim_move10 P t \ (Cast T (Val v)) e1' (Cast T (Val v)) h xs \ (THROW ClassCast) h xs" "\ typeof\<^bsub>h\<^esub> v = \U\; b \ v \ Null \ P \ U \ T \ \ sim_move10 P t \ ((Val v) instanceof T) e1' ((Val v) instanceof T) h xs \ (Val (Bool b)) h xs" "binop bop v1 v2 = Some (Inl v) \ sim_move10 P t \ ((Val v1) \bop\ (Val v2)) e1' (Val v1 \bop\ Val v2) h xs \ (Val v) h xs" "binop bop v1 v2 = Some (Inr a) \ sim_move10 P t \ ((Val v1) \bop\ (Val v2)) e1' (Val v1 \bop\ Val v2) h xs \ (Throw a) h xs" "xs V = \v\ \ sim_move10 P t \ (Var V') e1' (Var V) h xs \ (Val v) h xs" "sim_move10 P t \ (V' := Val v) e1' (V := Val v) h xs \ unit h (xs(V \ v))" "sim_move10 P t \ (null\Val v\) e1' (null\Val v\) h xs \ (THROW NullPointer) h xs" "\ typeof_addr h a = \Array_type T n\; i sint i \ int n \ \ sim_move10 P t \ (addr a\Val (Intg i)\) e1' ((addr a)\Val (Intg i)\) h xs \ (THROW ArrayIndexOutOfBounds) h xs" "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; heap_read h a (ACell (nat (sint i))) v; ta1 = \ReadMem a (ACell (nat (sint i))) v\; ta = \ReadMem a (ACell (nat (sint i))) v\ \ \ sim_move10 P t ta1 (addr a\Val (Intg i)\) e1' ((addr a)\Val (Intg i)\) h xs ta (Val v) h xs" "sim_move10 P t \ (null\Val v\ := Val v') e1' (null\Val v\ := Val v') h xs \ (THROW NullPointer) h xs" "\ typeof_addr h a = \Array_type T n\; i sint i \ int n \ \ sim_move10 P t \ (AAss (addr a) (Val (Intg i)) (Val v)) e1' (AAss (addr a) (Val (Intg i)) (Val v)) h xs \ (THROW ArrayIndexOutOfBounds) h xs" "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; typeof\<^bsub>h\<^esub> v = \U\; \ (P \ U \ T) \ \ sim_move10 P t \ (AAss (addr a) (Val (Intg i)) (Val v)) e1' (AAss (addr a) (Val (Intg i)) (Val v)) h xs \ (THROW ArrayStore) h xs" "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; typeof\<^bsub>h\<^esub> v = Some U; P \ U \ T; heap_write h a (ACell (nat (sint i))) v h'; ta1 = \WriteMem a (ACell (nat (sint i))) v\; ta = \WriteMem a (ACell (nat (sint i))) v\ \ \ sim_move10 P t ta1 (AAss (addr a) (Val (Intg i)) (Val v)) e1' (AAss (addr a) (Val (Intg i)) (Val v)) h xs ta unit h' xs" - "typeof_addr h a = \Array_type T n\ \ sim_move10 P t \ (addr a\length) e1' (addr a\length) h xs \ (Val (Intg (word_of_int (int n)))) h xs" + "typeof_addr h a = \Array_type T n\ \ sim_move10 P t \ (addr a\length) e1' (addr a\length) h xs \ (Val (Intg (word_of_nat n))) h xs" "sim_move10 P t \ (null\length) e1' (null\length) h xs \ (THROW NullPointer) h xs" "\ heap_read h a (CField D F) v; ta1 = \ReadMem a (CField D F) v\; ta = \ReadMem a (CField D F) v\ \ \ sim_move10 P t ta1 (addr a\F{D}) e1' (addr a\F{D}) h xs ta (Val v) h xs" "sim_move10 P t \ (null\F{D}) e1' (null\F{D}) h xs \ (THROW NullPointer) h xs" "\ heap_write h a (CField D F) v h'; ta1 = \WriteMem a (CField D F) v\; ta = \WriteMem a (CField D F) v\ \ \ sim_move10 P t ta1 (addr a\F{D} := Val v) e1' (addr a\F{D} := Val v) h xs ta unit h' xs" "sim_move10 P t \ (null\compareAndSwap(D\F, Val v, Val v')) e1' (null\compareAndSwap(D\F, Val v, Val v')) h xs \ (THROW NullPointer) h xs" "\ heap_read h a (CField D F) v''; heap_write h a (CField D F) v' h'; v'' = v; ta1 = \ ReadMem a (CField D F) v'', WriteMem a (CField D F) v' \; ta = \ ReadMem a (CField D F) v'', WriteMem a (CField D F) v' \ \ \ sim_move10 P t ta1 (addr a\compareAndSwap(D\F, Val v, Val v')) e1' (addr a\compareAndSwap(D\F, Val v, Val v')) h xs ta true h' xs" "\ heap_read h a (CField D F) v''; v'' \ v; ta1 = \ ReadMem a (CField D F) v'' \; ta = \ ReadMem a (CField D F) v'' \ \ \ sim_move10 P t ta1 (addr a\compareAndSwap(D\F, Val v, Val v')) e1' (addr a\compareAndSwap(D\F, Val v, Val v')) h xs ta false h xs" "sim_move10 P t \ (null\F{D} := Val v) e1' (null\F{D} := Val v) h xs \ (THROW NullPointer) h xs" "sim_move10 P t \ ({V':T=None; Val u}) e1' ({V:T=vo; Val u}) h xs \ (Val u) h xs" "sim_move10 P t \ ({V':T=\v\; e}) ({V':T=None; e}) ({V:T=vo; e'}) h xs \ ({V:T=vo; e'}) h xs" "sim_move10 P t \ (sync\<^bsub>V'\<^esub>(null) e0) e1' (sync(null) e1) h xs \ (THROW NullPointer) h xs" "sim_move10 P t \ (Val v;;e0) e1' (Val v;; e1) h xs \ e1 h xs" "sim_move10 P t \ (if (true) e0 else e0') e1' (if (true) e1 else e2) h xs \ e1 h xs" "sim_move10 P t \ (if (false) e0 else e0') e1' (if (false) e1 else e2) h xs \ e2 h xs" "sim_move10 P t \ (throw null) e1' (throw null) h xs \ (THROW NullPointer) h xs" "sim_move10 P t \ (try (Val v) catch(C V') e0) e1' (try (Val v) catch(C V) e1) h xs \ (Val v) h xs" "\ typeof_addr h a = \Class_type D\; P \ D \\<^sup>* C \ \ sim_move10 P t \ (try (Throw a) catch(C V') e0) e1' (try (Throw a) catch(C V) e1) h xs \ ({V:Class C=\Addr a\; e1}) h xs" "sim_move10 P t \ (newA T\Throw a\) e1' (newA T\Throw a\) h xs \ (Throw a) h xs" "sim_move10 P t \ (Cast T (Throw a)) e1' (Cast T (Throw a)) h xs \ (Throw a) h xs" "sim_move10 P t \ ((Throw a) instanceof T) e1' ((Throw a) instanceof T) h xs \ (Throw a) h xs" "sim_move10 P t \ ((Throw a) \bop\ e0) e1' ((Throw a) \bop\ e1) h xs \ (Throw a) h xs" "sim_move10 P t \ (Val v \bop\ (Throw a)) e1' (Val v \bop\ (Throw a)) h xs \ (Throw a) h xs" "sim_move10 P t \ (V' := Throw a) e1' (V := Throw a) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a\e0\) e1' (Throw a\e1\) h xs \ (Throw a) h xs" "sim_move10 P t \ (Val v\Throw a\) e1' (Val v\Throw a\) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a\e0\ := e0') e1' (Throw a\e1\ := e2) h xs \ (Throw a) h xs" "sim_move10 P t \ (Val v\Throw a\ := e0) e1' (Val v\Throw a\ := e1) h xs \ (Throw a) h xs" "sim_move10 P t \ (Val v\Val v'\ := Throw a) e1' (Val v\Val v'\ := Throw a) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a\length) e1' (Throw a\length) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a\F{D}) e1' (Throw a\F{D}) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a\F{D} := e0) e1' (Throw a\F{D} := e1) h xs \ (Throw a) h xs" "sim_move10 P t \ (Val v\F{D} := Throw a) e1' (Val v\F{D} := Throw a) h xs \ (Throw a) h xs" "sim_move10 P t \ (CompareAndSwap (Throw a) D F e0 e0') e1' (Throw a\compareAndSwap(D\F, e1'', e1''')) h xs \ (Throw a) h xs" "sim_move10 P t \ (CompareAndSwap (Val v) D F (Throw a) e0') e1' (Val v\compareAndSwap(D\F, Throw a, e1'')) h xs \ (Throw a) h xs" "sim_move10 P t \ (CompareAndSwap (Val v) D F (Val v') (Throw a)) e1' (Val v\compareAndSwap(D\F, Val v', Throw a)) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a\M(es0)) e1' (Throw a\M(es1)) h xs \ (Throw a) h xs" "sim_move10 P t \ (Val v\M(map Val vs @ Throw a # es0)) e1' (Val v\M(map Val vs @ Throw a # es1)) h xs \ (Throw a) h xs" "sim_move10 P t \ ({V':T=None; Throw a}) e1' ({V:T=vo; Throw a}) h xs \ (Throw a) h xs" "sim_move10 P t \ (sync\<^bsub>V'\<^esub>(Throw a) e0) e1' (sync(Throw a) e1) h xs \ (Throw a) h xs" "sim_move10 P t \ (Throw a;;e0) e1' (Throw a;;e1) h xs \ (Throw a) h xs" "sim_move10 P t \ (if (Throw a) e0 else e0') e1' (if (Throw a) e1 else e2) h xs \ (Throw a) h xs" "sim_move10 P t \ (throw (Throw a)) e1' (throw (Throw a)) h xs \ (Throw a) h xs" apply(fastforce simp add: sim_move10_def no_calls_def no_call_def ta_bisim_def intro: red_reds.intros)+ done lemma sim_move10_CallNull: "sim_move10 P t \ (null\M(map Val vs)) e1' (null\M(map Val vs)) h xs \ (THROW NullPointer) h xs" by(fastforce simp add: sim_move10_def map_eq_append_conv intro: RedCallNull) lemma sim_move10_SyncLocks: "\ ta1 = \Lock\a, SyncLock a\; ta = \Lock\a, SyncLock a\ \ \ sim_move10 P t ta1 (sync\<^bsub>V\<^esub>(addr a) e0) e1' (sync(addr a) e1) h xs ta (insync(a) e1) h xs" "\ ta1 = \Unlock\a, SyncUnlock a\; ta = \Unlock\a, SyncUnlock a\ \ \ sim_move10 P t ta1 (insync\<^bsub>V\<^esub>(a') (Val v)) e1' (insync(a) (Val v)) h xs ta (Val v) h xs" "\ ta1 = \Unlock\a, SyncUnlock a\; ta = \Unlock\a, SyncUnlock a\ \ \ sim_move10 P t ta1 (insync\<^bsub>V\<^esub>(a') (Throw a'')) e1' (insync(a) (Throw a'')) h xs ta (Throw a'') h xs" by(fastforce simp add: sim_move10_def ta_bisim_def ta_upd_simps intro: red_reds.intros[simplified])+ lemma sim_move10_TryFail: "\ typeof_addr h a = \Class_type D\; \ P \ D \\<^sup>* C \ \ sim_move10 P t \ (try (Throw a) catch(C V') e0) e1' (try (Throw a) catch(C V) e1) h xs \ (Throw a) h xs" by(auto simp add: sim_move10_def intro!: RedTryFail) lemmas sim_move10_intros = sim_move10_expr sim_move10_reds sim_move10_CallNull sim_move10_TryFail sim_move10_Block sim_move10_CallParams lemma sim_move10_preserves_defass: assumes wf: "wf_J_prog P" shows "\ sim_move10 P t ta1 e1 e1' e h xs ta e' h' xs'; \ e \dom xs\ \ \ \ e' \dom xs'\" by(auto simp add: sim_move10_def split: if_split_asm dest!: \red0r_preserves_defass[OF wf] \red0t_preserves_defass[OF wf] red_preserves_defass[OF wf]) declare sim_move10_intros[intro] lemma assumes wf: "wf_J_prog P" shows red_simulates_red1_aux: "\ False,compP1 P,t \1 \e1, S\ -TA\ \e1', S'\; bisim vs e2 e1 (lcl S); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl S]; length vs + max_vars e1 \ length (lcl S); \ e2 \dom x\ \ \ \ta e2' x'. sim_move10 P t TA e1 e1' e2 (hp S) x ta e2' (hp S') x' \ bisim vs e2' e1' (lcl S') \ x' \\<^sub>m [vs [\] lcl S']" (is "\ _; _; _; _; _; _ \ \ ?concl e1 e1' e2 S x TA S' e1' vs") and reds_simulates_reds1_aux: "\ False,compP1 P,t \1 \es1, S\ [-TA\] \es1', S'\; bisims vs es2 es1 (lcl S); fvs es2 \ set vs; x \\<^sub>m [vs [\] lcl S]; length vs + max_varss es1 \ length (lcl S); \s es2 \dom x\ \ \ \ta es2' x'. sim_moves10 P t TA es1 es1' es2 (hp S) x ta es2' (hp S') x' \ bisims vs es2' es1' (lcl S') \ x' \\<^sub>m [vs [\] lcl S']" (is "\ _; _; _; _; _; _ \ \ ?concls es1 es1' es2 S x TA S' es1' vs") proof(induct arbitrary: vs e2 x and vs es2 x rule: red1_reds1.inducts) case (Bin1OpRed1 e s ta e' s' bop e2 Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs E2 (e \bop\ e2) (lcl s)\ obtain E E2' where E2: "E2 = E \bop\ E2'" "e2 = compE1 Vs E2'" and "bisim Vs E e (lcl s)" and sync: "\ contains_insync E2'" by(auto elim!: bisim_cases) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (e \bop\ e2) \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' e' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(auto) with E2 \fv E2 \ set Vs\ sync show ?case by(cases "is_val e2'")(auto intro: BinOpRed1) next case (Bin1OpRed2 e s ta e' s' v bop Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \bisim Vs E2 (Val v \bop\ e) (lcl s)\ obtain E where E2: "E2 = Val v \bop\ E" and "bisim Vs E e (lcl s)" by(auto) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (Val v \bop\ e) \ length (lcl s)\ \\ E2 \dom X\\ E2 ultimately show ?case by(auto intro: BinOpRed2) next case (Red1Var s V v Vs E2 X) from \bisim Vs E2 (Var V) (lcl s)\ \fv E2 \ set Vs\ obtain V' where "E2 = Var V'" "V' = Vs ! V" "V = index Vs V'" by(clarify, simp) from \E2 = Var V'\ \\ E2 \dom X\\ obtain v' where "X V' = \v'\" by(auto simp add: hyperset_defs) with \X \\<^sub>m [Vs [\] lcl s]\ have "[Vs [\] lcl s] V' = \v'\" by(rule map_le_SomeD) with \length Vs + max_vars (Var V) \ length (lcl s)\ have "lcl s ! (index Vs V') = v'" by(auto intro: map_upds_Some_eq_nth_index) with \V = index Vs V'\ \lcl s ! V = v\ have "v = v'" by simp with \X V' = \v'\\ \E2 = Var V'\ \X \\<^sub>m [Vs [\] lcl s]\ show ?case by(fastforce intro: RedVar) next case (LAss1Red e s ta e' s' V Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \bisim Vs E2 (V:=e) (lcl s)\ obtain E V' where E2: "E2 = (V':=E)" "V = index Vs V'" and "bisim Vs E e (lcl s)" by auto with IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (V:=e) \ length (lcl s)\ \\ E2 \dom X\\ E2 show ?case by(auto intro: LAssRed) next case (Red1LAss V l v h Vs E2 X) from \bisim Vs E2 (V:=Val v) (lcl (h, l))\ obtain V' where "E2 = (V' := Val v)" "V = index Vs V'" by(auto) moreover with \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl (h, l)]\ \length Vs + max_vars (V:=Val v) \ length (lcl (h, l))\ have "X(V' \ v) \\<^sub>m [Vs [\] l[index Vs V' := v]]" by(auto intro: LAss_lem) ultimately show ?case by(auto intro: RedLAss simp del: fun_upd_apply) next case (AAcc1Red1 a s ta a' s' i Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 a (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars a \ length (lcl s); \ e2 \dom x\\ \ ?concl a a' e2 s x ta s' a' vs\ from \False,compP1 P,t \1 \a,s\ -ta\ \a',s'\\ have "\ is_val a" by auto with \bisim Vs E2 (a\i\) (lcl s)\ obtain E E2' where E2: "E2 = E\E2'\" "i = compE1 Vs E2'" and "bisim Vs E a (lcl s)" and sync: "\ contains_insync E2'" by(fastforce) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (a\i\) \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where "sim_move10 P t ta a a' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' a' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(auto) with E2 \fv E2 \ set Vs\ sync show ?case by(cases "is_val e2'")(auto intro: AAccRed1) next case (AAcc1Red2 i s ta i' s' a Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 i (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars i \ length (lcl s); \ e2 \dom x\\ \ ?concl i i' e2 s x ta s' i' vs\ from \bisim Vs E2 (Val a\i\) (lcl s)\ obtain E where E2: "E2 = Val a\E\" and "bisim Vs E i (lcl s)" by(auto) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ E2 \length Vs + max_vars (Val a\i\) \ length (lcl s)\ \\ E2 \dom X\\ ultimately show ?case by(auto intro: AAccRed2) next case Red1AAcc thus ?case by(fastforce intro: RedAAcc simp del: fun_upd_apply) next case (AAss1Red1 a s ta a' s' i e Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 a (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars a \ length (lcl s); \ e2 \dom x\ \ \ ?concl a a' e2 s x ta s' a' vs\ from \False,compP1 P,t \1 \a,s\ -ta\ \a',s'\\ have "\ is_val a" by auto with \bisim Vs E2 (a\i\:=e) (lcl s)\ obtain E E2' E2'' where E2: "E2 = E\E2'\:=E2''" "i = compE1 Vs E2'" "e = compE1 Vs E2''" and "bisim Vs E a (lcl s)" and sync: "\ contains_insync E2'" "\ contains_insync E2''" by(fastforce) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (a\i\:=e) \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where IH': "sim_move10 P t ta a a' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' a' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(auto) show ?case proof(cases "is_val e2'") case True from E2 \fv E2 \ set Vs\ sync have "bisim Vs E2' i (lcl s')" "bisim Vs E2'' e (lcl s')" by auto with IH' E2 True sync show ?thesis by(cases "is_val E2'")(fastforce intro: AAssRed1)+ next case False with IH' E2 sync show ?thesis by(fastforce intro: AAssRed1) qed next case (AAss1Red2 i s ta i' s' a e Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 i (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars i \ length (lcl s); \ e2 \dom x\\ \ ?concl i i' e2 s x ta s' i' vs\ from \False,compP1 P,t \1 \i,s\ -ta\ \i',s'\\ have "\ is_val i" by auto with \bisim Vs E2 (Val a\i\:=e) (lcl s)\ obtain E E2' where E2: "E2 = Val a\E\:=E2'" "e = compE1 Vs E2'" and "bisim Vs E i (lcl s)" and sync: "\ contains_insync E2'" by(fastforce) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (Val a\i\:=e) \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where "sim_move10 P t ta i i' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' i' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(auto) with E2 \fv E2 \ set Vs\ sync show ?case by(cases "is_val e2'")(fastforce intro: AAssRed2)+ next case (AAss1Red3 e s ta e' s' a i Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \bisim Vs E2 (Val a\Val i\:=e) (lcl s)\ obtain E where E2: "E2 = Val a\Val i\:=E" and "bisim Vs E e (lcl s)" by(fastforce) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ E2 \length Vs + max_vars (Val a\Val i\:=e) \ length (lcl s)\ \\ E2 \dom X\\ ultimately show ?case by(fastforce intro: AAssRed3) next case Red1AAssStore thus ?case by(auto)((rule exI conjI)+, auto) next case Red1AAss thus ?case by(fastforce simp del: fun_upd_apply) next case (FAss1Red1 e s ta e' s' F D e2 Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs E2 (e\F{D}:=e2) (lcl s)\ obtain E E2' where E2: "E2 = E\F{D}:=E2'" "e2 = compE1 Vs E2'" and "bisim Vs E e (lcl s)" and sync: "\ contains_insync E2'" by(fastforce) with IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (e\F{D}:=e2) \ length (lcl s)\ \\ E2 \dom X\\ obtain TA' e2' x' where "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' e' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(fastforce) with E2 \fv E2 \ set Vs\ sync show ?case by(cases "is_val e2'")(auto intro: FAssRed1) next case (FAss1Red2 e s ta e' s' v F D Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \bisim Vs E2 (Val v\F{D}:=e) (lcl s)\ obtain E where E2: "E2 = (Val v\F{D}:=E)" and "bisim Vs E e (lcl s)" by(fastforce) with IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (Val v\F{D}:=e) \ length (lcl s)\ \\ E2 \dom X\\ E2 show ?case by(fastforce intro: FAssRed2) next case (CAS1Red1 e s ta e' s' D F e2 e3 Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\ \ \ ?concl e e' e2 s x ta s' e' vs\ from \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs E2 _ (lcl s)\ obtain E E2' E2'' where E2: "E2 = E\compareAndSwap(D\F, E2', E2'')" "e2 = compE1 Vs E2'" "e3 = compE1 Vs E2''" and "bisim Vs E e (lcl s)" and sync: "\ contains_insync E2'" "\ contains_insync E2''" by(fastforce) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars _ \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where IH': "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' e' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(auto) show ?case proof(cases "is_val e2'") case True from E2 \fv E2 \ set Vs\ sync have "bisim Vs E2' e2 (lcl s')" "bisim Vs E2'' e3 (lcl s')" by auto with IH' E2 True sync show ?thesis by(cases "is_val E2'")(fastforce)+ next case False with IH' E2 sync show ?thesis by(fastforce) qed next case (CAS1Red2 e s ta e' s' v D F e3 Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ have "\ is_val e" by auto with \bisim Vs E2 _ (lcl s)\ obtain E E2' where E2: "E2 = (Val v\compareAndSwap(D\F, E, E2'))" "e3 = compE1 Vs E2'" and "bisim Vs E e (lcl s)" and sync: "\ contains_insync E2'" by(auto) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars _ \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' e' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(auto) with E2 \fv E2 \ set Vs\ sync show ?case by(cases "is_val e2'")(fastforce)+ next case (CAS1Red3 e s ta e' s' v D F v' Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ ?concl e e' e2 s x ta s' e' vs\ from \bisim Vs E2 _ (lcl s)\ obtain E where E2: "E2 = (Val v\compareAndSwap(D\F, Val v', E))" and "bisim Vs E e (lcl s)" by(fastforce) moreover note IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ E2 \length Vs + max_vars _ \ length (lcl s)\ \\ E2 \dom X\\ ultimately show ?case by(fastforce) next case (Call1Obj e s ta e' s' M es Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\ \ \ ?concl e e' e2 s x ta s' e' vs\ from \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ \bisim Vs E2 (e\M(es)) (lcl s)\ obtain E es' where E2: "E2 = E\M(es')" "es = compEs1 Vs es'" and "bisim Vs E e (lcl s)" and sync: "\ contains_insyncs es'" by(auto elim!: bisim_cases simp add: compEs1_conv_map) with IH[of Vs E X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (e\M(es)) \ length (lcl s)\ \\ E2 \dom X\\ obtain TA' e2' x' where IH': "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' e' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by(fastforce) with E2 \fv E2 \ set Vs\ \E2 = E\M(es')\ sync show ?case by(cases "is_val e2'")(auto intro: CallObj) next case (Call1Params es s ta es' s' v M Vs E2 X) note IH = \\vs es2 x. \ bisims vs es2 es (lcl s); fvs es2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_varss es \ length (lcl s); \s es2 \dom x\\ \ ?concls es es' es2 s x ta s' es' vs\ from \bisim Vs E2 (Val v\M(es)) (lcl s)\ obtain Es where "E2 = Val v \M(Es)" "bisims Vs Es es (lcl s)" by(auto) with IH[of Vs Es X] \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (Val v\M(es)) \ length (lcl s)\ \\ E2 \dom X\\ \E2 = Val v \M(Es)\ show ?case by(fastforce intro: CallParams) next case (Red1CallExternal s a T M Ts Tr D vs ta va h' e' s' Vs E2 X) from \bisim Vs E2 (addr a\M(map Val vs)) (lcl s)\ have E2: "E2 = addr a\M(map Val vs)" by auto moreover from \compP1 P \ class_type_of T sees M: Ts\Tr = Native in D\ have "P \ class_type_of T sees M: Ts\Tr = Native in D" by simp moreover from \compP1 P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\\ have "P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\" by simp moreover from wf \P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\\ have "ta_bisim01 (extTA2J0 P ta) (extTA2J1 (compP1 P) ta)" by(rule red_external_ta_bisim01) moreover note \typeof_addr (hp s) a = \T\\ \e' = extRet2J1 (addr a\M(map Val vs)) va\ \s' = (h', lcl s)\ moreover from \typeof_addr (hp s) a = \T\\ \P,t \ \a\M(vs),hp s\ -ta\ext \va,h'\\ \P \ class_type_of T sees M: Ts\Tr = Native in D\ have "\external_defs D M \ ta = \ \ h' = hp s" by(fastforce dest: \external'_red_external_heap_unchanged \external'_red_external_TA_empty simp add: \external'_def \external_def) ultimately show ?case using \X \\<^sub>m [Vs [\] lcl s]\ by(fastforce intro!: exI[where x="extTA2J0 P ta"] intro: RedCallExternal simp add: map_eq_append_conv sim_move10_def synthesized_call_def dest: sees_method_fun del: disjCI intro: disjI1 disjI2) next case (Block1Red e h x ta e' h' x' V T Vs E2 X) note IH = \\vs e2 xa. \ bisim vs e2 e (lcl (h, x)); fv e2 \ set vs; xa \\<^sub>m [vs [\] lcl (h, x)]; length vs + max_vars e \ length (lcl (h, x)); \ e2 \dom xa\\ \ ?concl e e' e2 (h, x) xa ta (h', x') e' vs\ from \False,compP1 P,t \1 \e,(h, x)\ -ta\ \e',(h', x')\\ have "length x = length x'" by(auto dest: red1_preserves_len) with \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ have "length Vs < length x'" by simp from \bisim Vs E2 {V:T=None; e} (lcl (h, x))\ show ?case proof(cases rule: bisim_cases(14)[consumes 1, case_names BlockNone BlockSome BlockSomeNone]) case (BlockNone V' E) with \fv E2 \ set Vs\ have "fv E \ set (Vs@[V'])" by auto with IH[of "Vs@[V']" E "X(V' := None)"] BlockNone \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl (h, x)]\ \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ \\ E2 \dom X\\ obtain TA' e2' X' where IH': "sim_move10 P t ta e e' E h (X(V' := None)) TA' e2' h' X'" "bisim (Vs @ [V']) e2' e' x'" "X' \\<^sub>m [Vs @ [V'] [\] x']" by(fastforce simp del: fun_upd_apply) { assume "V' \ set Vs" hence "hidden (Vs @ [V']) (index Vs V')" by(rule hidden_index) with \bisim (Vs @ [V']) E e (lcl (h, x))\ have "unmod e (index Vs V')" by(auto intro: hidden_bisim_unmod) moreover from \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ \V' \ set Vs\ have "index Vs V' < length x" by(auto intro: index_less_aux) ultimately have "x ! index Vs V' = x' ! index Vs V'" using red1_preserves_unmod[OF \False,compP1 P,t \1 \e,(h, x)\ -ta\ \e',(h', x')\\] by(simp) } with \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ \X' \\<^sub>m [Vs @ [V'] [\] x']\ \length x = length x'\ \X \\<^sub>m [Vs [\] lcl (h, x)]\ have rel: "X'(V' := X V') \\<^sub>m [Vs [\] x']" by(auto intro: Block_lem) show ?thesis proof(cases "X' V'") case None with BlockNone IH' rel show ?thesis by(fastforce intro: BlockRed) next case (Some v) with \X' \\<^sub>m [Vs @ [V'] [\] x']\ \length Vs < length x'\ have "x' ! length Vs = v" by(auto dest: map_le_SomeD) with BlockNone IH' rel Some show ?thesis by(fastforce intro: BlockRed) qed next case BlockSome thus ?thesis by simp next case (BlockSomeNone V' E) with \fv E2 \ set Vs\ have "fv E \ set (Vs@[V'])" by auto with IH[of "Vs@[V']" E "X(V' \ x ! length Vs)"] BlockSomeNone \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl (h, x)]\ \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ \\ E2 \dom X\\ obtain TA' e2' X' where IH': "sim_move10 P t ta e e' E h (X(V' \ x ! length Vs)) TA' e2' h' X'" "bisim (Vs @ [V']) e2' e' x'" "X' \\<^sub>m [Vs @ [V'] [\] x']" by(fastforce simp del: fun_upd_apply) { assume "V' \ set Vs" hence "hidden (Vs @ [V']) (index Vs V')" by(rule hidden_index) with \bisim (Vs @ [V']) E e (lcl (h, x))\ have "unmod e (index Vs V')" by(auto intro: hidden_bisim_unmod) moreover from \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ \V' \ set Vs\ have "index Vs V' < length x" by(auto intro: index_less_aux) ultimately have "x ! index Vs V' = x' ! index Vs V'" using red1_preserves_unmod[OF \False,compP1 P,t \1 \e,(h, x)\ -ta\ \e',(h', x')\\] by(simp) } with \length Vs + max_vars {V:T=None; e} \ length (lcl (h, x))\ \X' \\<^sub>m [Vs @ [V'] [\] x']\ \length x = length x'\ \X \\<^sub>m [Vs [\] lcl (h, x)]\ have rel: "X'(V' := X V') \\<^sub>m [Vs [\] x']" by(auto intro: Block_lem) from \sim_move10 P t ta e e' E h (X(V' \ x ! length Vs)) TA' e2' h' X'\ obtain v' where "X' V' = \v'\" by(auto simp: sim_move10_def split: if_split_asm dest!: \red0t_lcl_incr \red0r_lcl_incr red_lcl_incr subsetD) with \X' \\<^sub>m [Vs @ [V'] [\] x']\ \length Vs < length x'\ have "x' ! length Vs = v'" by(auto dest: map_le_SomeD) with BlockSomeNone IH' rel \X' V' = \v'\\ show ?thesis by(fastforce intro: BlockRed) qed next case(Block1Some V xs T v e h) from \bisim vs e2 {V:T=\v\; e} (lcl (h, xs))\ obtain e' V' where "e2 = {V':T=\v\; e'}" and "V = length vs" "bisim (vs @ [V']) e' e (xs[length vs := v])" by(fastforce) moreover have "sim_move10 P t \ {length vs:T=\v\; e} {length vs:T=None; e} {V':T=\v\; e'} h x \ {V':T=\v\; e'} h x" by(auto) moreover from \bisim (vs @ [V']) e' e (xs[length vs := v])\ \length vs + max_vars {V:T=\v\; e} \ length (lcl (h, xs))\ have "bisim vs {V':T=\v\; e'} {length vs:T=None; e} (xs[length vs := v])" by auto moreover from \x \\<^sub>m [vs [\] lcl (h, xs)]\ \length vs + max_vars {V:T=\v\; e} \ length (lcl (h, xs))\ have "x \\<^sub>m [vs [\] xs[length vs := v]]" by auto ultimately show ?case by auto next case (Lock1Synchronized V xs a e h Vs E2 X) note len = \length Vs + max_vars (sync\<^bsub>V\<^esub> (addr a) e) \ length (lcl (h, xs))\ from \bisim Vs E2 (sync\<^bsub>V\<^esub> (addr a) e) (lcl (h, xs))\ obtain E where E2: "E2 = sync(addr a) E" "e = compE1 (Vs@[fresh_var Vs]) E" and sync: "\ contains_insync E" and [simp]: "V = length Vs" by auto moreover have "extTA2J0 P,P,t \ \sync(addr a) E, (h, X)\ -\Lock\a, SyncLock a\\ \insync(a) E, (h, X)\" by(rule LockSynchronized) moreover from \fv E2 \ set Vs\ fresh_var_fresh[of Vs] sync len have "bisim Vs (insync(a) E) (insync\<^bsub>length Vs\<^esub> (a) e) (xs[length Vs := Addr a])" unfolding \e = compE1 (Vs@[fresh_var Vs]) E\ \E2 = sync(addr a) E\ by -(rule bisimInSynchronized,rule compE1_bisim, auto) moreover have "zip Vs (xs[length Vs := Addr a]) = (zip Vs xs)[length Vs := (arbitrary, Addr a)]" by(rule sym)(simp add: update_zip) hence "zip Vs (xs[length Vs := Addr a]) = zip Vs xs" by simp with \X \\<^sub>m [Vs [\] (lcl (h, xs))]\ have "X \\<^sub>m [Vs [\] xs[length Vs := Addr a]]" by(auto simp add: map_le_def map_upds_def) ultimately show ?case by(fastforce intro: sim_move10_SyncLocks) next case (Synchronized1Red2 e s ta e' s' V a Vs E2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\ \ \ ?concl e e' e2 s x ta s' e' vs\ from \bisim Vs E2 (insync\<^bsub>V\<^esub> (a) e) (lcl s)\ obtain E where E2: "E2 = insync(a) E" and bisim: "bisim (Vs@[fresh_var Vs]) E e (lcl s)" and xsa: "lcl s ! length Vs = Addr a" and [simp]: "V = length Vs" by auto with \fv E2 \ set Vs\ fresh_var_fresh[of Vs] have fv: "(fresh_var Vs) \ fv E" by auto from \length Vs + max_vars (insync\<^bsub>V\<^esub> (a) e) \ length (lcl s)\ have "length Vs < length (lcl s)" by simp { assume "X (fresh_var Vs) \ None" then obtain v where "X (fresh_var Vs) = \v\" by auto with \X \\<^sub>m [Vs [\] lcl s]\ have "[Vs [\] lcl s] (fresh_var Vs) = \v\" by(auto simp add: map_le_def dest: bspec) hence "(fresh_var Vs) \ set Vs" by(auto simp add: map_upds_def set_zip dest!: map_of_SomeD ) moreover have "(fresh_var Vs) \ set Vs" by(rule fresh_var_fresh) ultimately have False by contradiction } hence "X (fresh_var Vs) = None" by(cases "X (fresh_var Vs)", auto) hence "X(fresh_var Vs := None) = X" by(auto intro: ext) moreover from \X \\<^sub>m [Vs [\] lcl s]\ have "X(fresh_var Vs := None) \\<^sub>m [Vs [\] lcl s, (fresh_var Vs) \ (lcl s) ! length Vs]" by(simp) ultimately have "X \\<^sub>m [Vs @ [fresh_var Vs] [\] lcl s]" using \length Vs < length (lcl s)\ by(auto) moreover note IH[of "Vs@[fresh_var Vs]" E X] bisim E2 \fv E2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_vars (insync\<^bsub>V\<^esub> (a) e) \ length (lcl s)\ \\ E2 \dom X\\ ultimately obtain TA' e2' x' where IH': "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim (Vs @ [fresh_var Vs]) e2' e' (lcl s')" "x' \\<^sub>m [Vs @ [fresh_var Vs] [\] lcl s']" by auto hence "dom x' \ dom X \ fv E" by(fastforce iff del: domIff simp add: sim_move10_def dest: red_dom_lcl \red0r_dom_lcl[OF wf_prog_wwf_prog[OF wf]] \red0t_dom_lcl[OF wf_prog_wwf_prog[OF wf]] \red0r_fv_subset[OF wf_prog_wwf_prog[OF wf]] split: if_split_asm) with fv \X (fresh_var Vs) = None\ have "(fresh_var Vs) \ dom x'" by auto hence "x' (fresh_var Vs) = None" by auto moreover from \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ have "length (lcl s) = length (lcl s')" by(auto dest: red1_preserves_len) moreover note \x' \\<^sub>m [Vs @ [fresh_var Vs] [\] lcl s']\ \length Vs < length (lcl s)\ ultimately have "x' \\<^sub>m [Vs [\] lcl s']" by(auto simp add: map_le_def dest: bspec) moreover from bisim fv have "unmod e (length Vs)" by(auto intro: bisim_fv_unmod) with \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ \length Vs < length (lcl s)\ have "lcl s ! length Vs = lcl s' ! length Vs" by(auto dest!: red1_preserves_unmod) with xsa have "lcl s' ! length Vs = Addr a" by simp ultimately show ?case using IH' E2 by(auto intro: SynchronizedRed2) next case (Unlock1Synchronized xs V a' a v h Vs E2 X) from \bisim Vs E2 (insync\<^bsub>V\<^esub> (a) Val v) (lcl (h, xs))\ have E2: "E2 = insync(a) Val v" "V = length Vs" "xs ! length Vs = Addr a" by auto moreover with \xs ! V = Addr a'\ have [simp]: "a' = a" by simp have "extTA2J0 P,P,t \ \insync(a) (Val v), (h, X)\ -\Unlock\a, SyncUnlock a\\ \Val v, (h, X)\" by(rule UnlockSynchronized) ultimately show ?case using \X \\<^sub>m [Vs [\] lcl (h, xs)]\ by(fastforce intro: sim_move10_SyncLocks) next case (Unlock1SynchronizedNull xs V a v h Vs E2 X) from \bisim Vs E2 (insync\<^bsub>V\<^esub> (a) Val v) (lcl (h, xs))\ have "V = length Vs" "xs ! length Vs = Addr a" by(auto) with \xs ! V = Null\ have False by simp thus ?case .. next case (Unlock1SynchronizedFail xs V A' a' v h Vs E2 X) from \False\ show ?case .. next case (Red1While b c s Vs E2 X) from \bisim Vs E2 (while (b) c) (lcl s)\ obtain B C where E2: "E2 = while (B) C" "b = compE1 Vs B" "c = compE1 Vs C" and sync: "\ contains_insync B" "\ contains_insync C" by auto moreover have "extTA2J0 P,P,t \ \while (B) C, (hp s, X)\ -\\ \if (B) (C;;while (B) C) else unit, (hp s, X)\" by(rule RedWhile) hence "sim_move10 P t \ (while (compE1 Vs B) (compE1 Vs C)) (if (compE1 Vs B) (compE1 Vs C;;while (compE1 Vs B) (compE1 Vs C)) else unit) (while (B) C) (hp s) X \ (if (B) (C;;while (B) C) else unit) (hp s) X" by(auto simp add: sim_move10_def) moreover from \fv E2 \ set Vs\ E2 sync have "bisim Vs (if (B) (C;; while (B) C) else unit) (if (compE1 Vs B) (compE1 Vs (C;; while(B) C)) else (compE1 Vs unit)) (lcl s)" by -(rule bisimCond, auto) ultimately show ?case using \X \\<^sub>m [Vs [\] lcl s]\ by(simp)(rule exI, rule exI, rule exI, erule conjI, auto) next case (Red1TryCatch h a D C V x e2 Vs E2 X) from \bisim Vs E2 (try Throw a catch(C V) e2) (lcl (h, x))\ obtain E2' V' where "E2 = try Throw a catch(C V') E2'" "V = length Vs" "e2 = compE1 (Vs @ [V']) E2'" and sync: "\ contains_insync E2'" by(auto) with \fv E2 \ set Vs\ have "fv E2' \ set (Vs @[V'])" by auto with \e2 = compE1 (Vs @ [V']) E2'\ sync have "bisim (Vs @[V']) E2' e2 (x[V := Addr a])" by(auto intro!: compE1_bisim) with \V = length Vs\ \length Vs + max_vars (try Throw a catch(C V) e2) \ length (lcl (h, x))\ have "bisim Vs {V':Class C=\Addr a\; E2'} {length Vs:Class C=None; e2} (x[V := Addr a])" by(auto intro: bisimBlockSomeNone) moreover from \length Vs + max_vars (try Throw a catch(C V) e2) \ length (lcl (h, x))\ have "[Vs [\] x[length Vs := Addr a]] = [Vs [\] x]" by simp with \X \\<^sub>m [Vs [\] lcl (h, x)]\ have "X \\<^sub>m [Vs [\] x[length Vs := Addr a]]" by simp moreover note \e2 = compE1 (Vs @ [V']) E2'\ \E2 = try Throw a catch(C V') E2'\ \typeof_addr h a = \Class_type D\\ \compP1 P \ D \\<^sup>* C\ \V = length Vs\ ultimately show ?case by(auto intro!: exI) next case Red1TryFail thus ?case by(auto intro!: exI sim_move10_TryFail) next case (List1Red1 e s ta e' s' es Vs ES2 X) note IH = \\vs e2 x. \ bisim vs e2 e (lcl s); fv e2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_vars e \ length (lcl s); \ e2 \dom x\\ \ \TA' e2' x'. sim_move10 P t ta e e' e2 (hp s) x TA' e2' (hp s') x' \ bisim vs e2' e' (lcl s') \ x' \\<^sub>m [vs [\] lcl s']\ from \bisims Vs ES2 (e # es) (lcl s)\ \False,compP1 P,t \1 \e,s\ -ta\ \e',s'\\ obtain E ES where "ES2 = E # ES" "\ is_val E" "es = compEs1 Vs ES" "bisim Vs E e (lcl s)" and sync: "\ contains_insyncs ES" by(auto elim!: bisims_cases simp add: compEs1_conv_map) with IH[of Vs E X] \fvs ES2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_varss (e # es) \ length (lcl s)\ \\s ES2 \dom X\\ obtain TA' e2' x' where IH': "sim_move10 P t ta e e' E (hp s) X TA' e2' (hp s') x'" "bisim Vs e2' e' (lcl s')" "x' \\<^sub>m [Vs [\] lcl s']" by fastforce show ?case proof(cases "is_val e2'") case False with IH' \ES2 = E # ES\ \es = compEs1 Vs ES\ sync show ?thesis by(auto intro: sim_moves10_expr) next case True from \fvs ES2 \ set Vs\ \ES2 = E # ES\ \es = compEs1 Vs ES\ sync have "bisims Vs ES es (lcl s')" by(auto intro: compEs1_bisims) with IH' True \ES2 = E # ES\ \es = compEs1 Vs ES\ show ?thesis by(auto intro: sim_moves10_expr) qed next case (List1Red2 es s ta es' s' v Vs ES2 X) note IH = \\vs es2 x. \bisims vs es2 es (lcl s); fvs es2 \ set vs; x \\<^sub>m [vs [\] lcl s]; length vs + max_varss es \ length (lcl s); \s es2 \dom x\\ \ \TA' es2' x'. sim_moves10 P t ta es es' es2 (hp s) x TA' es2' (hp s') x' \ bisims vs es2' es' (lcl s') \ x' \\<^sub>m [vs [\] lcl s']\ from \bisims Vs ES2 (Val v # es) (lcl s)\ obtain ES where "ES2 = Val v # ES" "bisims Vs ES es (lcl s)" by(auto elim!: bisims_cases) with IH[of Vs ES X] \fvs ES2 \ set Vs\ \X \\<^sub>m [Vs [\] lcl s]\ \length Vs + max_varss (Val v # es) \ length (lcl s)\ \\s ES2 \dom X\\ \ES2 = Val v # ES\ show ?case by(fastforce intro: sim_moves10_expr) next case Call1ThrowParams thus ?case by(fastforce intro: CallThrowParams elim!: bisim_cases simp add: bisims_map_Val_Throw2) next case (Synchronized1Throw2 xs V a' a ad h Vs E2 X) from \bisim Vs E2 (insync\<^bsub>V\<^esub> (a) Throw ad) (lcl (h, xs))\ have "xs ! length Vs = Addr a" and "V = length Vs" by auto with \xs ! V = Addr a'\ have [simp]: "a' = a" by simp have "extTA2J0 P,P,t \ \insync(a) Throw ad, (h, X)\ -\Unlock\a, SyncUnlock a\\ \Throw ad, (h, X)\" by(rule SynchronizedThrow2) with \X \\<^sub>m [Vs [\] lcl (h, xs)]\ \bisim Vs E2 (insync\<^bsub>V\<^esub> (a) Throw ad) (lcl (h, xs))\ show ?case by(auto intro: sim_move10_SyncLocks intro!: exI) next case (Synchronized1Throw2Null xs V a a' h Vs E2 X) from \bisim Vs E2 (insync\<^bsub>V\<^esub> (a) Throw a') (lcl (h, xs))\ have "V = length Vs" "xs ! length Vs = Addr a" by(auto) with \xs ! V = Null\ have False by simp thus ?case .. next case (Synchronized1Throw2Fail xs V A' a' a h Vs E2 X) from \False\ show ?case .. next case InstanceOf1Red thus ?case by auto(blast) next case Red1InstanceOf thus ?case by hypsubst_thin auto next case InstanceOf1Throw thus ?case by auto next case CAS1Throw thus ?case by fastforce next case CAS1Throw2 thus ?case by fastforce next case CAS1Throw3 thus ?case by fastforce qed(simp_all del: fun_upd_apply, (fastforce intro: red_reds.intros simp del: fun_upd_apply simp add: finfun_upd_apply)+) lemma bisim_call_Some_call1: "\ bisim Vs e e' xs; call e = \aMvs\; length Vs + max_vars e' \ length xs \ \ \e'' xs'. \red1'r P t h (e', xs) (e'', xs') \ call1 e'' = \aMvs\ \ bisim Vs e e'' xs' \ take (length Vs) xs = take (length Vs) xs'" and bisims_calls_Some_calls1: "\ bisims Vs es es' xs; calls es = \aMvs\; length Vs + max_varss es' \ length xs \ \ \es'' xs'. \reds1'r P t h (es', xs) (es'', xs') \ calls1 es'' = \aMvs\ \ bisims Vs es es'' xs' \ take (length Vs) xs = take (length Vs) xs'" proof(induct rule: bisim_bisims.inducts) case bisimCallParams thus ?case by(fastforce simp add: is_vals_conv split: if_split_asm) next case bisimBlockNone thus ?case by(fastforce intro: take_eq_take_le_eq) next case (bisimBlockSome Vs V e e' xs v T) from \length Vs + max_vars {length Vs:T=\v\; e'} \ length xs\ have "\red1'r P t h ({length Vs:T=\v\; e'}, xs) ({length Vs:T=None; e'}, xs[length Vs := v])" by(auto intro!: \red1r_1step Block1Some) also from bisimBlockSome obtain e'' xs' where "\red1'r P t h (e', xs[length Vs := v]) (e'', xs')" and "call1 e'' = \aMvs\" "bisim (Vs @ [V]) e e'' xs'" and "take (length (Vs @ [V])) (xs[length Vs := v]) = take (length (Vs @ [V])) xs'" by auto hence "\red1'r P t h ({length Vs:T=None; e'}, xs[length Vs := v]) ({length Vs:T=None; e''}, xs')" by auto also from \call1 e'' = \aMvs\\ have "call1 {length Vs:T=None; e''} = \aMvs\" by simp moreover from \take (length (Vs @ [V])) (xs[length Vs := v]) = take (length (Vs @ [V])) xs'\ have "take (length Vs) xs = take (length Vs) xs'" by(auto dest: take_eq_take_le_eq[where m="length Vs"] simp add: take_list_update_beyond) moreover { have "xs' ! length Vs = take (length (Vs @ [V])) xs' ! length Vs" by simp also note \take (length (Vs @ [V])) (xs[length Vs := v]) = take (length (Vs @ [V])) xs'\[symmetric] also have "take (length (Vs @ [V])) (xs[length Vs := v]) ! length Vs = v" using \length Vs + max_vars {length Vs:T=\v\; e'} \ length xs\ by simp finally have "bisim Vs {V:T=\v\; e} {length Vs:T=None; e''} xs'" using \bisim (Vs @ [V]) e e'' xs'\ by auto } ultimately show ?case by blast next case (bisimBlockSomeNone Vs V e e' xs v T) then obtain e'' xs' where "\red1'r P t h (e', xs) (e'', xs')" "call1 e'' = \aMvs\" "bisim (Vs @ [V]) e e'' xs'" and "take (length (Vs @ [V])) xs = take (length (Vs @ [V])) xs'" by auto hence "\red1'r P t h ({length Vs:T=None; e'}, xs) ({length Vs:T=None; e''}, xs')" by auto moreover from \call1 e'' = \aMvs\\ have "call1 ({length Vs:T=None; e''}) = \aMvs\" by simp moreover from \take (length (Vs @ [V])) xs = take (length (Vs @ [V])) xs'\ have "take (length Vs) xs = take (length Vs) xs'" by(auto intro: take_eq_take_le_eq) moreover { have "xs' ! length Vs = take (length (Vs @ [V])) xs' ! length Vs" by simp also note \take (length (Vs @ [V])) xs = take (length (Vs @ [V])) xs'\[symmetric] also have "take (length (Vs @ [V])) xs ! length Vs = v" using \xs ! length Vs = v\ by simp finally have "bisim Vs {V:T=\v\; e} {length Vs:T=None; e''} xs'" using \bisim (Vs @ [V]) e e'' xs'\ by auto } ultimately show ?case by blast next case (bisimInSynchronized Vs e e' xs a) then obtain e'' xs' where "\red1'r P t h (e', xs) (e'', xs')" "call1 e'' = \aMvs\" "bisim (Vs @ [fresh_var Vs]) e e'' xs'" and "take (Suc (length Vs)) xs = take (Suc (length Vs)) xs'" by auto hence "\red1'r P t h (insync\<^bsub>length Vs\<^esub> (a) e', xs) (insync\<^bsub>length Vs\<^esub> (a) e'', xs')" by auto moreover from \call1 e'' = \aMvs\\ have "call1 (insync\<^bsub>length Vs\<^esub> (a) e'') = \aMvs\" by simp moreover from \take (Suc (length Vs)) xs = take (Suc (length Vs)) xs'\ have "take (length Vs) xs = take (length Vs) xs'" by(auto intro: take_eq_take_le_eq) moreover { have "xs' ! length Vs = take (Suc (length Vs)) xs' ! length Vs" by simp also note \take (Suc (length Vs)) xs = take (Suc (length Vs)) xs'\[symmetric] also have "take (Suc (length Vs)) xs ! length Vs = Addr a" using \xs ! length Vs = Addr a\ by simp finally have "bisim Vs (insync(a) e) (insync\<^bsub>length Vs\<^esub> (a) e'') xs'" using \bisim (Vs @ [fresh_var Vs]) e e'' xs'\ by auto } ultimately show ?case by blast next case bisimsCons1 thus ?case by(fastforce intro!: \red1r_inj_\reds1r) next case bisimsCons2 thus ?case by(fastforce intro!: \reds1r_cons_\reds1r) qed fastforce+ lemma sim_move01_into_Red1: "sim_move01 P t ta e E' h xs ta' e2' h' xs' \ if \Move0 P h (e, es1) then \Red1't P t h ((E', xs), exs2) ((e2', xs'), exs2) \ ta = \ \ h = h' else \ex2' exs2' ta'. \Red1'r P t h ((E', xs), exs2) (ex2', exs2') \ (call e = None \ call1 E' = None \ ex2' = (E', xs) \ exs2' = exs2) \ False,P,t \1 \ex2'/exs2',h\ -ta'\ \(e2', xs')/exs2,h'\ \ \ \Move1 P h (ex2', exs2') \ ta_bisim01 ta ta'" apply(auto simp add: sim_move01_def intro: \red1t_into_\Red1t \red1r_into_\Red1r red1Red split: if_split_asm) apply(fastforce intro: red1Red intro!: \red1r_into_\Red1r)+ done lemma sim_move01_max_vars_decr: "sim_move01 P t ta e0 e h xs ta' e' h' xs' \ max_vars e' \ max_vars e" by(fastforce simp add: sim_move01_def split: if_split_asm dest: \red1t_max_vars red1_max_vars_decr \red1r_max_vars) lemma Red1_simulates_red0: assumes wf: "wf_J_prog P" and red: "P,t \0 \e1/es1, h\ -ta\ \e1'/es1', h'\" and bisiml: "bisim_list1 (e1, es1) (ex2, exs2)" shows "\ex2'' exs2''. bisim_list1 (e1', es1') (ex2'', exs2'') \ (if \Move0 P h (e1, es1) then \Red1't (compP1 P) t h (ex2, exs2) (ex2'', exs2'') \ ta = \ \ h = h' else \ex2' exs2' ta'. \Red1'r (compP1 P) t h (ex2, exs2) (ex2', exs2') \ (call e1 = None \ call1 (fst ex2) = None \ ex2' = ex2 \ exs2' = exs2) \ False,compP1 P,t \1 \ex2'/exs2', h\ -ta'\ \ex2''/exs2'', h'\ \ \ \Move1 P h (ex2', exs2') \ ta_bisim01 ta ta')" (is "\ex2'' exs2'' . _ \ ?red ex2'' exs2''") using red proof(cases) case (red0Red XS') note [simp] = \es1' = es1\ and red = \extTA2J0 P,P,t \ \e1,(h, Map.empty)\ -ta\ \e1',(h', XS')\\ and notsynth = \\aMvs. call e1 = \aMvs\ \ synthesized_call P h aMvs\ from bisiml obtain E xs where ex2: "ex2 = (E, xs)" and bisim: "bisim [] e1 E xs" and fv: "fv e1 = {}" and length: "max_vars E \ length xs" and bsl: "bisim_list es1 exs2" and D: "\ e1 \{}\" by(auto elim!: bisim_list1_elim) from bisim_max_vars[OF bisim] length red1_simulates_red_aux[OF wf red bisim] fv notsynth obtain ta' e2' xs' where sim: "sim_move01 (compP1 P) t ta e1 E h xs ta' e2' h' xs'" and bisim': "bisim [] e1' e2' xs'" and XS': "XS' \\<^sub>m Map.empty" by auto from sim_move01_into_Red1[OF sim, of es1 exs2] have "?red (e2', xs') exs2" unfolding ex2 by auto moreover { note bsl bisim' moreover from fv red_fv_subset[OF wf_prog_wwf_prog[OF wf] red] have "fv e1' = {}" by simp moreover from red D have "\ e1' \dom XS'\" by(auto dest: red_preserves_defass[OF wf] split: if_split_asm) with red_dom_lcl[OF red] \fv e1 = {}\ have "\ e1' \{}\" by(auto elim!: D_mono' simp add: hyperset_defs) moreover from sim have "length xs = length xs'" "max_vars e2' \ max_vars E" by(auto dest: sim_move01_preserves_len sim_move01_max_vars_decr) with length have length': "max_vars e2' \ length xs'" by(auto) ultimately have "bisim_list1 (e1', es1) ((e2', xs'), exs2)" by(rule bisim_list1I) } ultimately show ?thesis using ex2 by(simp split del: if_split)(rule exI conjI|assumption)+ next case (red0Call a M vs U Ts T pns body D) note [simp] = \ta = \\ \h' = h\ and es1' = \es1' = e1 # es1\ and e1' = \e1' = blocks (this # pns) (Class D # Ts) (Addr a # vs) body\ and call = \call e1 = \(a, M, vs)\\ and ha = \typeof_addr h a = \U\\ and sees = \P \ class_type_of U sees M: Ts\T = \(pns, body)\ in D\ and len = \length vs = length pns\ \length Ts = length pns\ from bisiml obtain E xs where ex2: "ex2 = (E, xs)" and bisim: "bisim [] e1 E xs" and fv: "fv e1 = {}" and length: "max_vars E \ length xs" and bsl: "bisim_list es1 exs2" and D: "\ e1 \{}\" by(auto elim!: bisim_list1_elim) from bisim_call_Some_call1[OF bisim call, of "compP1 P" t h] length obtain e' xs' where red: "\red1'r (compP1 P) t h (E, xs) (e', xs')" and "call1 e' = \(a, M, vs)\" "bisim [] e1 e' xs'" and "take 0 xs = take 0 xs'" by auto let ?e1' = "blocks (this # pns) (Class D # Ts) (Addr a # vs) body" let ?e2' = "blocks1 0 (Class D#Ts) (compE1 (this # pns) body)" let ?xs2' = "Addr a # vs @ replicate (max_vars (compE1 (this # pns) body)) undefined_value" let ?exs2' = "(e', xs') # exs2" have "\Red1'r (compP1 P) t h ((E, xs), exs2) ((e', xs'), exs2)" using red by(rule \red1r_into_\Red1r) moreover { note \call1 e' = \(a, M, vs)\\ moreover note ha moreover have "compP1 P \ class_type_of U sees M:Ts \ T = map_option (\(pns, body). compE1 (this # pns) body) \(pns, body)\ in D" using sees unfolding compP1_def by(rule sees_method_compP) hence sees': "compP1 P \ class_type_of U sees M:Ts \ T = \compE1 (this # pns) body\ in D" by simp moreover from len have "length vs = length Ts" by simp ultimately have "False,compP1 P,t \1 \(e', xs')/exs2,h\ -\\ \(?e2', ?xs2')/?exs2', h\" by(rule red1Call) moreover have "\Move1 P h ((e', xs'), exs2)" using \call1 e' = \(a, M, vs)\\ ha sees by(auto simp add: synthesized_call_def \external'_def dest: sees_method_fun dest!: \move1_not_call1[where P=P and h=h]) ultimately have "\Red1' (compP1 P) t h ((e', xs'), exs2) ((?e2', ?xs2'), ?exs2')" by auto } ultimately have "\Red1't (compP1 P) t h ((E, xs), exs2) ((?e2', ?xs2'), ?exs2')" by(rule rtranclp_into_tranclp1) moreover { from red have "length xs' = length xs" by(rule \red1r_preserves_len) moreover from red have "max_vars e' \ max_vars E" by(rule \red1r_max_vars) ultimately have "max_vars e' \ length xs'" using length by simp } with bsl \bisim [] e1 e' xs'\ fv D have "bisim_list (e1 # es1) ?exs2'" using \call e1 = \(a, M, vs)\\ \call1 e' = \(a, M, vs)\\ by(rule bisim_listCons) hence "bisim_list1 (e1', es1') ((?e2', ?xs2'), ?exs2')" unfolding e1' es1' proof(rule bisim_list1I) from wf_prog_wwf_prog[OF wf] sees have "wf_mdecl wwf_J_mdecl P D (M,Ts,T,\(pns,body)\)" by(rule sees_wf_mdecl) hence fv': "fv body \ set pns \ {this}" by(auto simp add: wf_mdecl_def) moreover from \P \ class_type_of U sees M: Ts\T = \(pns, body)\ in D\ have "\ contains_insync body" by(auto dest!: sees_wf_mdecl[OF wf] WT_expr_locks simp add: wf_mdecl_def contains_insync_conv) ultimately have "bisim ([this] @ pns) body (compE1 ([this] @ pns) body) ?xs2'" by -(rule compE1_bisim, auto) with \length vs = length pns\ \length Ts = length pns\ have "bisim ([] @ [this]) (blocks pns Ts vs body) (blocks1 (Suc 0) Ts (compE1 (this # pns) body)) ?xs2'" by -(drule blocks_bisim,auto simp add: nth_append) from bisimBlockSomeNone[OF this, of "Addr a" "Class D"] show "bisim [] ?e1' ?e2' ?xs2'" by simp from fv' len show "fv ?e1' = {}" by auto from wf sees have "wf_mdecl wf_J_mdecl P D (M,Ts,T,\(pns,body)\)" by(rule sees_wf_mdecl) hence "\ body \set pns \ {this}\" by(auto simp add: wf_mdecl_def) with \length vs = length pns\ \length Ts = length pns\ have "\ (blocks pns Ts vs body) \dom [this \ Addr a]\" by(auto) thus "\ ?e1' \{}\" by auto from len show "max_vars ?e2' \ length ?xs2'" by(auto simp add: blocks1_max_vars) qed moreover have "\Move0 P h (e1, es1)" using call ha sees by(fastforce simp add: synthesized_call_def \external'_def dest: sees_method_fun \move0_callD[where P=P and h=h]) ultimately show ?thesis using ex2 \call e1 = \(a, M, vs)\\ by(simp del: \Move1.simps)(rule exI conjI|assumption)+ next case (red0Return e) note es1 = \es1 = e # es1'\ and e1' = \e1' = inline_call e1 e\ and [simp] = \ta = \\ \h' = h\ and fin = \final e1\ from bisiml es1 obtain E' xs' E xs exs' aMvs where ex2: "ex2 = (E', xs')" and exs2: "exs2 = (E, xs) # exs'" and bisim': "bisim [] e1 E' xs'" and bisim: "bisim [] e E xs" and fv: "fv e = {}" and length: "max_vars E \ length xs" and bisiml: "bisim_list es1' exs'" and D: "\ e \{}\" and call: "call e = \aMvs\" and call1: "call1 E = \aMvs\" by(fastforce elim: bisim_list1_elim) let ?e2' = "inline_call E' E" from \final e1\ bisim' have "final E'" by(auto) hence red': "False,compP1 P,t \1 \ex2/exs2, h\ -\\ \(?e2', xs)/exs', h\" unfolding ex2 exs2 by(rule red1Return) moreover have "\Move0 P h (e1, es1) = \Move1 P h ((E', xs'), exs2)" using \final e1\ \final E'\ by auto moreover { note bisiml moreover from bisim' fin bisim have "bisim [] (inline_call e1 e) (inline_call E' E) xs" using call by(rule bisim_inline_call)(simp add: fv) moreover from fv_inline_call[of e1 e] fv fin have "fv (inline_call e1 e) = {}" by auto moreover from fin have "\ (inline_call e1 e) \{}\" using call D by(rule defass_inline_call) moreover have "max_vars ?e2' \ max_vars E + max_vars E'" by(rule inline_call_max_vars1) with \final E'\ length have "max_vars ?e2' \ length xs" by(auto elim!: final.cases) ultimately have "bisim_list1 (inline_call e1 e, es1') ((?e2', xs), exs')" by(rule bisim_list1I) } ultimately show ?thesis using e1' \final e1\ \final E'\ ex2 apply(simp del: \Move0.simps \Move1.simps) apply(rule exI conjI impI|assumption)+ apply(rule tranclp.r_into_trancl, simp) apply blast done qed lemma sim_move10_into_red0: assumes wwf: "wwf_J_prog P" and sim: "sim_move10 P t ta e2 e2' e h Map.empty ta' e' h' x'" and fv: "fv e = {}" shows "if \move1 P h e2 then (\Red0t P t h (e, es) (e', es) \ countInitBlock e2' < countInitBlock e2 \ e' = e \ x' = Map.empty) \ ta = \ \ h = h' else \e'' ta'. \Red0r P t h (e, es) (e'', es) \ (call1 e2 = None \ call e = None \ e'' = e) \ P,t \0 \e''/es,h\ -ta'\ \e'/es,h'\ \ \ \Move0 P h (e'', es) \ ta_bisim01 ta' (extTA2J1 (compP1 P) ta)" proof(cases "\move1 P h e2") case True with sim have "\ final e2" and red: "\red0t (extTA2J0 P) P t h (e, Map.empty) (e', x') \ countInitBlock e2' < countInitBlock e2 \ e' = e \ x' = Map.empty" and [simp]: "h' = h" "ta = \" "ta' = \" by(simp_all add: sim_move10_def) from red have "\Red0t P t h (e, es) (e', es) \ countInitBlock e2' < countInitBlock e2 \ e' = e \ x' = Map.empty" proof assume red: "\red0t (extTA2J0 P) P t h (e, Map.empty) (e', x')" from \red0t_fv_subset[OF wwf red] \red0t_dom_lcl[OF wwf red] fv have "dom x' \ {}" by(auto split: if_split_asm) hence "x' = Map.empty" by auto with red have "\red0t (extTA2J0 P) P t h (e, Map.empty) (e', Map.empty)" by simp with wwf have "\Red0t P t h (e, es) (e', es)" using fv by(rule \red0t_into_\Red0t) thus ?thesis .. qed simp with True show ?thesis by simp next case False with sim obtain e'' xs'' where "\ final e2" and \red: "\red0r (extTA2J0 P) P t h (e, Map.empty) (e'', xs'')" and red: "extTA2J0 P,P,t \ \e'',(h, xs'')\ -ta'\ \e',(h', x')\" and call: "call1 e2 = None \ call e = None \ e'' = e" and "\ \move0 P h e''" "ta_bisim01 ta' (extTA2J1 (compP1 P) ta)" "no_call P h e''" by(auto simp add: sim_move10_def split: if_split_asm) from \red0r_fv_subset[OF wwf \red] \red0r_dom_lcl[OF wwf \red] fv have "dom xs'' \ {}" by(auto) hence "xs'' = Map.empty" by(auto) with \red have "\red0r (extTA2J0 P) P t h (e, Map.empty) (e'', Map.empty)" by simp with wwf have "\Red0r P t h (e, es) (e'', es)" using fv by(rule \red0r_into_\Red0r) moreover from red \xs'' = Map.empty\ have "extTA2J0 P,P,t \ \e'',(h, Map.empty)\ -ta'\ \e',(h', x')\" by simp from red0Red[OF this] \no_call P h e''\ have "P,t \0 \e''/es,h\ -ta'\ \e'/es,h'\" by(simp add: no_call_def) moreover from \\ \move0 P h e''\ red have "\ \Move0 P h (e'', es)" by auto ultimately show ?thesis using False \ta_bisim01 ta' (extTA2J1 (compP1 P) ta)\ call by(simp del: \Move0.simps) blast qed lemma red0_simulates_Red1: assumes wf: "wf_J_prog P" and red: "False,compP1 P,t \1 \ex2/exs2, h\ -ta\ \ex2'/exs2', h'\" and bisiml: "bisim_list1 (e, es) (ex2, exs2)" shows "\e' es'. bisim_list1 (e', es') (ex2', exs2') \ (if \Move1 P h (ex2, exs2) then (\Red0t P t h (e, es) (e', es') \ countInitBlock (fst ex2') < countInitBlock (fst ex2) \ exs2' = exs2 \ e' = e \ es' = es) \ ta = \ \ h = h' else \e'' es'' ta'. \Red0r P t h (e, es) (e'', es'') \ (call1 (fst ex2) = None \ call e = None \ e'' = e \ es'' = es) \ P,t \0 \e''/es'', h\ -ta'\ \e'/es', h'\ \ \ \Move0 P h (e'', es'') \ ta_bisim01 ta' ta)" (is "\e' es' . _ \ ?red e' es'") using red proof(cases) case (red1Red E xs TA E' xs') note red = \False,compP1 P,t \1 \E,(h, xs)\ -TA\ \E',(h', xs')\\ and ex2 = \ex2 = (E, xs)\ and ex2' = \ex2' = (E', xs')\ and [simp] = \ta = extTA2J1 (compP1 P) TA\ \exs2' = exs2\ from bisiml ex2 have bisim: "bisim [] e E xs" and fv: "fv e = {}" and length: "max_vars E \ length xs" and bsl: "bisim_list es exs2" and D: "\ e \{}\" by(auto elim: bisim_list1_elim) from red_simulates_red1_aux[OF wf red, simplified, OF bisim, of Map.empty] fv length D obtain TA' e2' x' where red': "sim_move10 P t TA E E' e h Map.empty TA' e2' h' x'" and bisim'': "bisim [] e2' E' xs'" and lcl': "x' \\<^sub>m Map.empty" by auto from red have "\ final E" by auto with sim_move10_into_red0[OF wf_prog_wwf_prog[OF wf] red', of es] fv ex2 ex2' have red'': "?red e2' es" by fastforce moreover { note bsl bisim'' moreover from red' fv have "fv e2' = {}" by(fastforce simp add: sim_move10_def split: if_split_asm dest: \red0r_fv_subset[OF wf_prog_wwf_prog[OF wf]] \red0t_fv_subset[OF wf_prog_wwf_prog[OF wf]] red_fv_subset[OF wf_prog_wwf_prog[OF wf]]) moreover from red' have "dom x' \ dom (Map.empty) \ fv e" unfolding sim_move10_def apply(auto split: if_split_asm del: subsetI dest: \red0r_dom_lcl[OF wf_prog_wwf_prog[OF wf]] \red0t_dom_lcl[OF wf_prog_wwf_prog[OF wf]]) apply(frule_tac [1-2] \red0r_fv_subset[OF wf_prog_wwf_prog[OF wf]]) apply(auto dest!: \red0r_dom_lcl[OF wf_prog_wwf_prog[OF wf]] red_dom_lcl del: subsetI, blast+) done with fv have "dom x' \ {}" by(auto) hence "x' = Map.empty" by(auto) with D red' have "\ e2' \{}\" by(auto dest!: sim_move10_preserves_defass[OF wf] split: if_split_asm) moreover from red have "length xs' = length xs" by(auto dest: red1_preserves_len) with red1_max_vars[OF red] length have "max_vars E' \ length xs'" by simp ultimately have "bisim_list1 (e2', es) ((E', xs'), exs2)" by(rule bisim_list1I) } ultimately show ?thesis using ex2' by(clarsimp split: if_split_asm)(rule exI conjI|assumption|simp)+ next case (red1Call E a M vs U Ts T body D xs) note [simp] = \ex2 = (E, xs)\ \h' = h\ \ta = \\ and ex2' = \ex2' = (blocks1 0 (Class D#Ts) body, Addr a # vs @ replicate (max_vars body) undefined_value)\ and exs' = \exs2' = (E, xs) # exs2\ and call = \call1 E = \(a, M, vs)\\ and ha = \typeof_addr h a = \U\\ and sees = \compP1 P \ class_type_of U sees M: Ts\T = \body\ in D\ and len = \length vs = length Ts\ let ?e2' = "blocks1 0 (Class D#Ts) body" let ?xs2' = "Addr a # vs @ replicate (max_vars body) undefined_value" from bisiml have bisim: "bisim [] e E xs" and fv: "fv e = {}" and length: "max_vars E \ length xs" and bsl: "bisim_list es exs2" and D: "\ e \{}\" by(auto elim: bisim_list1_elim) from bisim \call1 E = \(a, M, vs)\\ have "call e = \(a, M, vs)\" by(rule bisim_call1_Some_call) moreover note ha moreover from \compP1 P \ class_type_of U sees M: Ts\T = \body\ in D\ obtain pns Jbody where sees': "P \ class_type_of U sees M: Ts\T = \(pns, Jbody)\ in D" and body: "body = compE1 (this # pns) Jbody" by(auto dest: sees_method_compPD) let ?e' = "blocks (this # pns) (Class D # Ts) (Addr a # vs) Jbody" note sees' moreover from wf sees' have "length Ts = length pns" by(auto dest!: sees_wf_mdecl simp add: wf_mdecl_def) with len have "length vs = length pns" "length Ts = length pns" by simp_all ultimately have red': "P,t \0 \e/es, h\ -\\ \?e'/e#es, h\" by(rule red0Call) moreover from \call e = \(a, M, vs)\\ ha sees' have "\Move0 P h (e, es)" by(fastforce simp add: synthesized_call_def dest: \move0_callD[where P=P and h=h] sees_method_fun) ultimately have "\Red0t P t h (e, es) (?e', e#es)" by auto moreover from bsl bisim fv D length \call e = \(a, M, vs)\\ \call1 E = \(a, M, vs)\\ have "bisim_list (e # es) ((E, xs) # exs2)" by(rule bisim_list.intros) hence "bisim_list1 (?e', e # es) (ex2', (E, xs) # exs2)" unfolding ex2' proof(rule bisim_list1I) from wf_prog_wwf_prog[OF wf] sees' have "wf_mdecl wwf_J_mdecl P D (M,Ts,T,\(pns,Jbody)\)" by(rule sees_wf_mdecl) hence "fv Jbody \ set pns \ {this}" by(auto simp add: wf_mdecl_def) moreover from sees' have "\ contains_insync Jbody" by(auto dest!: sees_wf_mdecl[OF wf] WT_expr_locks simp add: wf_mdecl_def contains_insync_conv) ultimately have "bisim ([] @ this # pns) Jbody (compE1 ([] @ this # pns) Jbody) ?xs2'" by -(rule compE1_bisim, auto) with \length vs = length Ts\ \length Ts = length pns\ body have "bisim [] ?e' (blocks1 (length ([] :: vname list)) (Class D # Ts) body) ?xs2'" by -(rule blocks_bisim, auto simp add: nth_append nth_Cons') thus "bisim [] ?e' ?e2' ?xs2'" by simp from \length vs = length Ts\ \length Ts = length pns\ \fv Jbody \ set pns \ {this}\ show "fv ?e' = {}" by auto from wf sees' have "wf_mdecl wf_J_mdecl P D (M,Ts,T,\(pns,Jbody)\)" by(rule sees_wf_mdecl) hence "\ Jbody \set pns \ {this}\" by(auto simp add: wf_mdecl_def) with \length vs = length Ts\ \length Ts = length pns\ have "\ (blocks pns Ts vs Jbody) \dom [this \ Addr a]\" by(auto) thus "\ ?e' \{}\" by simp from len show "max_vars ?e2' \ length ?xs2'" by(simp add: blocks1_max_vars) qed moreover have "\Move1 P h (ex2, exs2)" using ha \call1 E = \(a, M, vs)\\ sees' by(auto simp add: synthesized_call_def \external'_def dest!: \move1_not_call1[where P=P and h=h] dest: sees_method_fun) ultimately show ?thesis using exs' by(simp del: \Move1.simps \Move0.simps)(rule exI conjI rtranclp.rtrancl_refl|assumption)+ next case (red1Return E' x' E x) note [simp] = \h' = h\ \ta = \\ and ex2 = \ex2 = (E', x')\ and exs2 = \exs2 = (E, x) # exs2'\ and ex2' = \ex2' = (inline_call E' E, x)\ and fin = \final E'\ from bisiml ex2 exs2 obtain e' es' aMvs where es: "es = e' # es'" and bsl: "bisim_list es' exs2'" and bisim: "bisim [] e E' x'" and bisim': "bisim [] e' E x" and fv: "fv e' = {}" and length: "max_vars E \ length x" and D: "\ e' \{}\" and call: "call e' = \aMvs\" and call1: "call1 E = \aMvs\" by(fastforce elim!: bisim_list1_elim) from \final E'\ bisim have fin': "final e" by(auto) hence "P,t \0 \e/e' # es',h\ -\\ \inline_call e e'/es',h\" by(rule red0Return) moreover from bisim fin' bisim' call have "bisim [] (inline_call e e') (inline_call E' E) x" by(rule bisim_inline_call)(simp add: fv) with bsl have "bisim_list1 (inline_call e e', es') (ex2', exs2')" unfolding ex2' proof(rule bisim_list1I) from fin' fv_inline_call[of e e'] fv show "fv (inline_call e e') = {}" by auto from fin' show "\ (inline_call e e') \{}\" using call D by(rule defass_inline_call) from call1_imp_call[OF call1] have "max_vars (inline_call E' E) \ max_vars E + max_vars E'" by(rule inline_call_max_vars) with fin length show "max_vars (inline_call E' E) \ length x" by(auto elim!: final.cases) qed moreover have "\Move1 P h (ex2, exs2) = \Move0 P h (e, es)" using ex2 call1 call fin fin' by(auto) ultimately show ?thesis using es by(simp del: \Move1.simps \Move0.simps) blast qed end sublocale J0_J1_heap_base < red0_Red1': FWdelay_bisimulation_base final_expr0 "mred0 P" final_expr1 "mred1' (compP1 P)" convert_RA "\t. bisim_red0_Red1" bisim_wait01 "\MOVE0 P" "\MOVE1 (compP1 P)" for P . context J0_J1_heap_base begin lemma delay_bisimulation_red0_Red1: assumes wf: "wf_J_prog P" shows "delay_bisimulation_measure (mred0 P t) (mred1' (compP1 P) t) bisim_red0_Red1 (ta_bisim (\t. bisim_red0_Red1)) (\MOVE0 P) (\MOVE1 (compP1 P)) (\es es'. False) (\(((e', xs'), exs'), h') (((e, xs), exs), h). countInitBlock e'< countInitBlock e)" (is "delay_bisimulation_measure _ _ _ _ _ _ ?\1 ?\2") proof(unfold_locales) fix s1 s2 s1' assume "bisim_red0_Red1 s1 s2" "red0_mthr.silent_move P t s1 s1'" moreover obtain ex1 exs1 h1 where s1: "s1 = ((ex1, exs1), h1)" by(cases s1, auto) moreover obtain ex1' exs1' h1' where s1': "s1' = ((ex1', exs1'), h1')" by(cases s1', auto) moreover obtain ex2 exs2 h2 where s2: "s2 = ((ex2, exs2), h2)" by(cases s2, auto) ultimately have bisim: "bisim_list1 (ex1, exs1) (ex2, exs2)" and heap: "h1 = h2" and red: "P,t \0 \ex1/exs1,h1\ -\\ \ex1'/exs1',h1'\" and \: "\Move0 P h1 (ex1, exs1)" by(auto simp add: bisim_red0_Red1_def red0_mthr.silent_move_iff) from Red1_simulates_red0[OF wf red bisim] \ obtain ex2'' exs2'' where bisim': "bisim_list1 (ex1', exs1') (ex2'', exs2'')" and red': "\Red1't (compP1 P) t h1 (ex2, exs2) (ex2'', exs2'')" and [simp]: "h1' = h1" by auto from \Red1't_into_Red1'_\mthr_silent_movet[OF red'] bisim' heap have "\s2'. Red1_mthr.silent_movet False (compP1 P) t s2 s2' \ bisim_red0_Red1 s1' s2'" unfolding s2 s1' by(auto simp add: bisim_red0_Red1_def) thus "bisim_red0_Red1 s1' s2 \ ?\1^++ s1' s1 \ (\s2'. Red1_mthr.silent_movet False (compP1 P) t s2 s2' \ bisim_red0_Red1 s1' s2')" .. next fix s1 s2 s2' assume "bisim_red0_Red1 s1 s2" "Red1_mthr.silent_move False (compP1 P) t s2 s2'" moreover obtain ex1 exs1 h1 where s1: "s1 = ((ex1, exs1), h1)" by(cases s1, auto) moreover obtain ex2 exs2 h2 where s2: "s2 = ((ex2, exs2), h2)" by(cases s2, auto) moreover obtain ex2' exs2' h2' where s2': "s2' = ((ex2', exs2'), h2')" by(cases s2', auto) ultimately have bisim: "bisim_list1 (ex1, exs1) (ex2, exs2)" and heap: "h1 = h2" and red: "False,compP1 P,t \1 \ex2/exs2,h2\ -\\ \ex2'/exs2',h2'\" and \: "\Move1 P h2 (ex2, exs2)" by(fastforce simp add: bisim_red0_Red1_def Red1_mthr.silent_move_iff)+ from red0_simulates_Red1[OF wf red bisim] \ obtain e' es' where bisim': "bisim_list1 (e', es') (ex2', exs2')" and red': "\Red0t P t h2 (ex1, exs1) (e', es') \ countInitBlock (fst ex2') < countInitBlock (fst ex2) \ exs2' = exs2 \ e' = ex1 \ es' = exs1" and [simp]: "h2' = h2" by auto from red' show "bisim_red0_Red1 s1 s2' \ ?\2^++ s2' s2 \ (\s1'. red0_mthr.silent_movet P t s1 s1' \ bisim_red0_Red1 s1' s2')" (is "?refl \ ?step") proof assume "\Red0t P t h2 (ex1, exs1) (e', es')" from \Red0t_into_red0_\mthr_silent_movet[OF this] bisim' heap have ?step unfolding s1 s2' by(auto simp add: bisim_red0_Red1_def) thus ?thesis .. next assume "countInitBlock (fst ex2') < countInitBlock (fst ex2) \ exs2' = exs2 \ e' = ex1 \ es' = exs1" hence ?refl using bisim' heap unfolding s1 s2' s2 by (auto simp add: bisim_red0_Red1_def split_beta) thus ?thesis .. qed next fix s1 s2 ta1 s1' assume "bisim_red0_Red1 s1 s2" and "mred0 P t s1 ta1 s1'" and \: "\ \MOVE0 P s1 ta1 s1'" moreover obtain ex1 exs1 h1 where s1: "s1 = ((ex1, exs1), h1)" by(cases s1, auto) moreover obtain ex1' exs1' h1' where s1': "s1' = ((ex1', exs1'), h1')" by(cases s1', auto) moreover obtain ex2 exs2 h2 where s2: "s2 = ((ex2, exs2), h2)" by(cases s2, auto) ultimately have heap: "h2 = h1" and bisim: "bisim_list1 (ex1, exs1) (ex2, exs2)" and red: "P,t \0 \ex1/exs1,h1\ -ta1\ \ex1'/exs1',h1'\" by(auto simp add: bisim_red0_Red1_def) from \ have "\ \Move0 P h1 (ex1, exs1)" unfolding s1 using red by(auto elim!: red0.cases dest: red_\_taD[where extTA="extTA2J0 P", OF extTA2J0_\]) with Red1_simulates_red0[OF wf red bisim] obtain ex2'' exs2'' ex2' exs2' ta' where bisim': "bisim_list1 (ex1', exs1') (ex2'', exs2'')" and red': "\Red1'r (compP1 P) t h1 (ex2, exs2) (ex2', exs2')" and red'': "False,compP1 P,t \1 \ex2'/exs2',h1\ -ta'\ \ex2''/exs2'',h1'\" and \': "\ \Move1 P h1 (ex2', exs2')" and ta: "ta_bisim01 ta1 ta'" by fastforce from red'' have "mred1' (compP1 P) t ((ex2', exs2'), h1) ta' ((ex2'', exs2''), h1')" by auto moreover from \' have "\ \MOVE1 (compP1 P) ((ex2', exs2'), h1) ta' ((ex2'', exs2''), h1')" by simp moreover from red' have "Red1_mthr.silent_moves False (compP1 P) t s2 ((ex2', exs2'), h1)" unfolding s2 heap by(rule \Red1'r_into_Red1'_\mthr_silent_moves) moreover from bisim' have "bisim_red0_Red1 ((ex1', exs1'), h1') ((ex2'', exs2''), h1')" by(auto simp add: bisim_red0_Red1_def) ultimately show "\s2' s2'' ta2. Red1_mthr.silent_moves False (compP1 P) t s2 s2' \ mred1' (compP1 P) t s2' ta2 s2'' \ \ \MOVE1 (compP1 P) s2' ta2 s2'' \ bisim_red0_Red1 s1' s2'' \ ta_bisim01 ta1 ta2" using ta unfolding s1' by blast next fix s1 s2 ta2 s2' assume "bisim_red0_Red1 s1 s2" and "mred1' (compP1 P) t s2 ta2 s2'" and \: "\ \MOVE1 (compP1 P) s2 ta2 s2'" moreover obtain ex1 exs1 h1 where s1: "s1 = ((ex1, exs1), h1)" by(cases s1, auto) moreover obtain ex2 exs2 h2 where s2: "s2 = ((ex2, exs2), h2)" by(cases s2, auto) moreover obtain ex2' exs2' h2' where s2': "s2' = ((ex2', exs2'), h2')" by(cases s2', auto) ultimately have heap: "h2 = h1" and bisim: "bisim_list1 (ex1, exs1) (ex2, exs2)" and red: "False,compP1 P,t \1 \ex2/exs2,h1\ -ta2\ \ex2'/exs2',h2'\" by(auto simp add: bisim_red0_Red1_def) from \ heap have "\ \Move1 P h2 (ex2, exs2)" unfolding s2 using red by(fastforce elim!: Red1.cases dest: red1_\_taD) with red0_simulates_Red1[OF wf red bisim] obtain e' es' e'' es'' ta' where bisim': "bisim_list1 (e', es') (ex2', exs2')" and red': "\Red0r P t h1 (ex1, exs1) (e'', es'')" and red'': "P,t \0 \e''/es'',h1\ -ta'\ \e'/es',h2'\" and \': "\ \Move0 P h1 (e'', es'')" and ta: "ta_bisim01 ta' ta2" using heap by fastforce from red'' have "mred0 P t ((e'', es''), h1) ta' ((e', es'), h2')" by auto moreover from red' have "red0_mthr.silent_moves P t ((ex1, exs1), h1) ((e'', es''), h1)" by(rule \Red0r_into_red0_\mthr_silent_moves) moreover from \' have "\ \MOVE0 P ((e'', es''), h1) ta' ((e', es'), h2')" by simp moreover from bisim' have "bisim_red0_Red1 ((e', es'), h2') ((ex2', exs2'), h2')" by(auto simp add: bisim_red0_Red1_def) ultimately show "\s1' s1'' ta1. red0_mthr.silent_moves P t s1 s1' \ mred0 P t s1' ta1 s1'' \ \ \MOVE0 P s1' ta1 s1'' \ bisim_red0_Red1 s1'' s2' \ ta_bisim01 ta1 ta2" using ta unfolding s1 s2' by blast next show "wfP ?\1" by auto next have "wf (measure countInitBlock)" .. hence "wf (inv_image (measure countInitBlock) (\(((e', xs'), exs'), h'). e'))" .. also have "inv_image (measure countInitBlock) (\(((e', xs'), exs'), h'). e') = {(s2', s2). ?\2 s2' s2}" by(simp add: inv_image_def split_beta) finally show "wfP ?\2" by(simp add: wfP_def) qed lemma delay_bisimulation_diverge_red0_Red1: assumes "wf_J_prog P" shows "delay_bisimulation_diverge (mred0 P t) (mred1' (compP1 P) t) bisim_red0_Red1 (ta_bisim (\t. bisim_red0_Red1)) (\MOVE0 P) (\MOVE1 (compP1 P))" proof - interpret delay_bisimulation_measure "mred0 P t" "mred1' (compP1 P) t" "bisim_red0_Red1" "ta_bisim (\t. bisim_red0_Red1)" "\MOVE0 P" "\MOVE1 (compP1 P)" "\es es'. False" "\(((e', xs'), exs'), h') (((e, xs), exs), h). countInitBlock e'< countInitBlock e" using assms by(rule delay_bisimulation_red0_Red1) show ?thesis by unfold_locales qed lemma red0_Red1'_FWweak_bisim: assumes wf: "wf_J_prog P" shows "FWdelay_bisimulation_diverge final_expr0 (mred0 P) final_expr1 (mred1' (compP1 P)) (\t. bisim_red0_Red1) bisim_wait01 (\MOVE0 P) (\MOVE1 (compP1 P))" proof - interpret delay_bisimulation_diverge "mred0 P t" "mred1' (compP1 P) t" bisim_red0_Red1 "ta_bisim (\t. bisim_red0_Red1)" "\MOVE0 P" "\MOVE1 (compP1 P)" for t using wf by(rule delay_bisimulation_diverge_red0_Red1) show ?thesis proof fix t and s1 and s2 :: "(('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list) \ 'heap" assume "bisim_red0_Red1 s1 s2" "(\(x1, m). final_expr0 x1) s1" moreover hence "(\(x2, m). final_expr1 x2) s2" by(cases s1)(cases s2,auto simp add: bisim_red0_Red1_def final_iff elim!: bisim_list1_elim elim: bisim_list.cases) ultimately show "\s2'. Red1_mthr.silent_moves False (compP1 P) t s2 s2' \ bisim_red0_Red1 s1 s2' \ (\(x2, m). final_expr1 x2) s2'" by blast next fix t s1 and s2 :: "(('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list) \ 'heap" assume "bisim_red0_Red1 s1 s2" "(\(x2, m). final_expr1 x2) s2" moreover hence "(\(x1, m). final_expr0 x1) s1" by(cases s1)(cases s2,auto simp add: bisim_red0_Red1_def final_iff elim!: bisim_list1_elim elim: bisim_list.cases) ultimately show "\s1'. red0_mthr.silent_moves P t s1 s1' \ bisim_red0_Red1 s1' s2 \ (\(x1, m). final_expr0 x1) s1'" by blast next fix t' x m1 x' m2 t x1 x2 x1' ta1 x1'' m1' x2' ta2 x2'' m2' assume b: "bisim_red0_Red1 (x, m1) (x', m2)" and bo: "bisim_red0_Red1 (x1, m1) (x2, m2)" and \red1: "red0_mthr.silent_moves P t (x1, m1) (x1', m1)" and red1: "mred0 P t (x1', m1) ta1 (x1'', m1')" and \1: "\ \MOVE0 P (x1', m1) ta1 (x1'', m1')" and \red2: "Red1_mthr.silent_moves False (compP1 P) t (x2, m2) (x2', m2)" and red2: "mred1' (compP1 P) t (x2', m2) ta2 (x2'', m2')" and bo': "bisim_red0_Red1 (x1'', m1') (x2'', m2')" and tb: "ta_bisim (\t. bisim_red0_Red1) ta1 ta2" from b have "m1 = m2" by(auto simp add: bisim_red0_Red1_def split_beta) moreover from bo' have "m1' = m2'" by(auto simp add: bisim_red0_Red1_def split_beta) ultimately show "bisim_red0_Red1 (x, m1') (x', m2')" using b by(auto simp add: bisim_red0_Red1_def split_beta) next fix t x1 m1 x2 m2 x1' ta1 x1'' m1' x2' ta2 x2'' m2' w assume "bisim_red0_Red1 (x1, m1) (x2, m2)" and "red0_mthr.silent_moves P t (x1, m1) (x1', m1)" and red0: "mred0 P t (x1', m1) ta1 (x1'', m1')" and "\ \MOVE0 P (x1', m1) ta1 (x1'', m1')" and "Red1_mthr.silent_moves False (compP1 P) t (x2, m2) (x2', m2)" and red1: "mred1' (compP1 P) t (x2', m2) ta2 (x2'', m2')" and "\ \MOVE1 (compP1 P) (x2', m2) ta2 (x2'', m2')" and "bisim_red0_Red1 (x1'', m1') (x2'', m2')" and "ta_bisim01 ta1 ta2" and Suspend: "Suspend w \ set \ta1\\<^bsub>w\<^esub>" "Suspend w \ set \ta2\\<^bsub>w\<^esub>" from red0 red1 Suspend show "bisim_wait01 x1'' x2''" by(cases x2')(cases x2'', auto dest: Red_Suspend_is_call Red1_Suspend_is_call simp add: split_beta bisim_wait01_def is_call_def) next fix t x1 m1 x2 m2 ta1 x1' m1' assume "bisim_red0_Red1 (x1, m1) (x2, m2)" and "bisim_wait01 x1 x2" and "mred0 P t (x1, m1) ta1 (x1', m1')" and wakeup: "Notified \ set \ta1\\<^bsub>w\<^esub> \ WokenUp \ set \ta1\\<^bsub>w\<^esub>" moreover obtain e0 es0 where [simp]: "x1 = (e0, es0)" by(cases x1) moreover obtain e0' es0' where [simp]: "x1' = (e0', es0')" by(cases x1') moreover obtain e1 xs1 exs1 where [simp]: "x2 = ((e1, xs1), exs1)" by(cases x2) auto ultimately have bisim: "bisim_list1 (e0, es0) ((e1, xs1), exs1)" and m1: "m2 = m1" and call: "call e0 \ None" "call1 e1 \ None" and red: "P,t \0 \e0/es0, m1\ -ta1\ \e0'/es0', m1'\" by(auto simp add: bisim_wait01_def bisim_red0_Red1_def) from red wakeup have "\ \Move0 P m1 (e0, es0)" by(auto elim!: red0.cases dest: red_\_taD[where extTA="extTA2J0 P", simplified]) with Red1_simulates_red0[OF wf red bisim] call m1 show "\ta2 x2' m2'. mred1' (compP1 P) t (x2, m2) ta2 (x2', m2') \ bisim_red0_Red1 (x1', m1') (x2', m2') \ ta_bisim01 ta1 ta2" by(auto simp add: bisim_red0_Red1_def) next fix t x1 m1 x2 m2 ta2 x2' m2' assume "bisim_red0_Red1 (x1, m1) (x2, m2)" and "bisim_wait01 x1 x2" and "mred1' (compP1 P) t (x2, m2) ta2 (x2', m2')" and wakeup: "Notified \ set \ta2\\<^bsub>w\<^esub> \ WokenUp \ set \ta2\\<^bsub>w\<^esub>" moreover obtain e0 es0 where [simp]: "x1 = (e0, es0)" by(cases x1) moreover obtain e1 xs1 exs1 where [simp]: "x2 = ((e1, xs1), exs1)" by(cases x2) auto moreover obtain e1' xs1' exs1' where [simp]: "x2' = ((e1', xs1'), exs1')" by(cases x2') auto ultimately have bisim: "bisim_list1 (e0, es0) ((e1, xs1), exs1)" and m1: "m2 = m1" and call: "call e0 \ None" "call1 e1 \ None" and red: "False,compP1 P,t \1 \(e1, xs1)/exs1, m1\ -ta2\ \(e1', xs1')/exs1', m2'\" by(auto simp add: bisim_wait01_def bisim_red0_Red1_def) from red wakeup have "\ \Move1 P m1 ((e1, xs1), exs1)" by(auto elim!: Red1.cases dest: red1_\_taD) with red0_simulates_Red1[OF wf red bisim] m1 call show "\ta1 x1' m1'. mred0 P t (x1, m1) ta1 (x1', m1') \ bisim_red0_Red1 (x1', m1') (x2', m2') \ ta_bisim01 ta1 ta2" by(auto simp add: bisim_red0_Red1_def) next show "(\x. final_expr0 x) \ (\x. final_expr1 x)" by(auto simp add: split_paired_Ex final_iff) qed qed lemma bisim_J0_J1_start: assumes wf: "wf_J_prog P" and start: "wf_start_state P C M vs" shows "red0_Red1'.mbisim (J0_start_state P C M vs) (J1_start_state (compP1 P) C M vs)" proof - from start obtain Ts T pns body D where sees: "P \ C sees M:Ts\T=\(pns,body)\ in D" and conf: "P,start_heap \ vs [:\] Ts" by cases auto from conf have vs: "length vs = length Ts" by(rule list_all2_lengthD) from sees_wf_mdecl[OF wf_prog_wwf_prog[OF wf] sees] have fvbody: "fv body \ {this} \ set pns" and len: "length pns = length Ts" by(auto simp add: wf_mdecl_def) with vs have fv: "fv (blocks pns Ts vs body) \ {this}" by auto have wfCM: "wf_J_mdecl P D (M, Ts, T, pns, body)" using sees_wf_mdecl[OF wf sees] by(auto simp add: wf_mdecl_def) then obtain T' where wtbody: "P,[this # pns [\] Class D # Ts] \ body :: T'" by auto hence elbody: "expr_locks body = (\l. 0)" by(rule WT_expr_locks) hence cisbody: "\ contains_insync body" by(auto simp add: contains_insync_conv) from wfCM len vs have dabody: "\ (blocks pns Ts vs body) \{this}\" by auto from sees have sees1: "compP1 P \ C sees M:Ts\T = \compE1 (this # pns) body\ in D" by(auto dest: sees_method_compP[where f="(\C M Ts T (pns, body). compE1 (this # pns) body)"]) let ?e = "blocks1 0 (Class C#Ts) (compE1 (this # pns) body)" let ?xs = "Null # vs @ replicate (max_vars body) undefined_value" from fvbody cisbody len vs have "bisim [] (blocks (this # pns) (Class D # Ts) (Null # vs) body) (blocks1 (length ([] :: vname list)) (Class D # Ts) (compE1 (this # pns) body)) ?xs" by-(rule blocks_bisim,(fastforce simp add: nth_Cons' nth_append)+) with fv dabody len vs elbody sees sees1 show ?thesis unfolding start_state_def by(auto intro!: red0_Red1'.mbisimI split: if_split_asm)(auto simp add: bisim_red0_Red1_def blocks1_max_vars intro!: bisim_list.intros bisim_list1I wset_thread_okI) qed end end diff --git a/thys/JinjaThreads/Compiler/J1.thy b/thys/JinjaThreads/Compiler/J1.thy --- a/thys/JinjaThreads/Compiler/J1.thy +++ b/thys/JinjaThreads/Compiler/J1.thy @@ -1,1268 +1,1268 @@ (* Title: JinjaThreads/Compiler/J1.thy Author: Andreas Lochbihler *) section \Semantics of the intermediate language\ theory J1 imports J1State J1Heap "../Framework/FWBisimulation" begin abbreviation final_expr1 :: "('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list \ bool" where "final_expr1 \ \(ex, exs). final (fst ex) \ exs = []" definition extNTA2J1 :: "'addr J1_prog \ (cname \ mname \ 'addr) \ (('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list)" where "extNTA2J1 P = (\(C, M, a). let (D, _, _, meth) = method P C M; body = the meth in (({0:Class D=None; body}, Addr a # replicate (max_vars body) undefined_value), []))" lemma extNTA2J1_iff [simp]: "extNTA2J1 P (C, M, a) = (({0:Class (fst (method P C M))=None; the (snd (snd (snd (method P C M))))}, Addr a # replicate (max_vars (the (snd (snd (snd (method P C M)))))) undefined_value), [])" by(simp add: extNTA2J1_def split_beta) abbreviation extTA2J1 :: "'addr J1_prog \ ('addr, 'thread_id, 'heap) external_thread_action \ ('addr, 'thread_id, 'heap) J1_thread_action" where "extTA2J1 P \ convert_extTA (extNTA2J1 P)" abbreviation (input) extRet2J1 :: "'addr expr1 \ 'addr extCallRet \ 'addr expr1" where "extRet2J1 \ extRet2J" lemma max_vars_extRet2J1 [simp]: "max_vars e = 0 \ max_vars (extRet2J1 e va) = 0" by(cases va) simp_all context J1_heap_base begin abbreviation J1_start_state :: "'addr J1_prog \ cname \ mname \ 'addr val list \ ('addr, 'thread_id, 'heap) J1_state" where "J1_start_state \ start_state (\C M Ts T body vs. ((blocks1 0 (Class C # Ts) body, Null # vs @ replicate (max_vars body) undefined_value), []))" inductive red1 :: "bool \ 'addr J1_prog \ 'thread_id \ 'addr expr1 \ 'heap \ 'addr locals1 \ ('addr, 'thread_id, 'heap) external_thread_action \ 'addr expr1 \ 'heap \ 'addr locals1 \ bool" ("_,_,_ \1 ((1\_,/_\) -_\/ (1\_,/_\))" [51,51,0,0,0,0,0,0] 81) and reds1 :: "bool \ 'addr J1_prog \ 'thread_id \ 'addr expr1 list \ 'heap \ 'addr locals1 \ ('addr, 'thread_id, 'heap) external_thread_action \ 'addr expr1 list \ 'heap \ 'addr locals1 \ bool" ("_,_,_ \1 ((1\_,/_\) [-_\]/ (1\_,/_\))" [51,51,0,0,0,0,0,0] 81) for uf :: bool and P :: "'addr J1_prog" and t :: 'thread_id where Red1New: "(h', a) \ allocate h (Class_type C) \ uf,P,t \1 \new C, (h, l)\ -\NewHeapElem a (Class_type C)\\ \addr a, (h', l)\" | Red1NewFail: "allocate h (Class_type C) = {} \ uf,P,t \1 \new C, (h, l)\ -\\ \THROW OutOfMemory, (h, l)\" | New1ArrayRed: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \newA T\e\, s\ -ta\ \newA T\e'\, s'\" | Red1NewArray: "\ 0 <=s i; (h', a) \ allocate h (Array_type T (nat (sint i))) \ \ uf,P,t \1 \newA T\Val (Intg i)\, (h, l)\ -\NewHeapElem a (Array_type T (nat (sint i)))\\ \addr a, (h', l)\" | Red1NewArrayNegative: "i uf,P,t \1 \newA T\Val (Intg i)\, s\ -\\ \THROW NegativeArraySize, s\" | Red1NewArrayFail: "\ 0 <=s i; allocate h (Array_type T (nat (sint i))) = {} \ \ uf,P,t \1 \newA T\Val (Intg i)\, (h, l)\ -\\ \THROW OutOfMemory, (h, l)\" | Cast1Red: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \Cast C e, s\ -ta\ \Cast C e', s'\" | Red1Cast: "\ typeof\<^bsub>hp s\<^esub> v = \U\; P \ U \ T \ \ uf,P,t \1 \Cast T (Val v), s\ -\\ \Val v, s\" | Red1CastFail: "\ typeof\<^bsub>hp s\<^esub> v = \U\; \ P \ U \ T \ \ uf,P,t \1 \Cast T (Val v), s\ -\\ \THROW ClassCast, s\" | InstanceOf1Red: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \e instanceof T, s\ -ta\ \e' instanceof T, s'\" | Red1InstanceOf: "\ typeof\<^bsub>hp s\<^esub> v = \U\; b \ v \ Null \ P \ U \ T \ \ uf,P,t \1 \(Val v) instanceof T, s\ -\\ \Val (Bool b), s\" | Bin1OpRed1: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \e \bop\ e2, s\ -ta\ \e' \bop\ e2, s'\" | Bin1OpRed2: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \(Val v) \bop\ e, s\ -ta\ \(Val v) \bop\ e', s'\" | Red1BinOp: "binop bop v1 v2 = Some (Inl v) \ uf,P,t \1 \(Val v1) \bop\ (Val v2), s\ -\\ \Val v, s\" | Red1BinOpFail: "binop bop v1 v2 = Some (Inr a) \ uf,P,t \1 \(Val v1) \bop\ (Val v2), s\ -\\ \Throw a, s\" | Red1Var: "\ (lcl s)!V = v; V < size (lcl s) \ \ uf,P,t \1 \Var V, s\ -\\ \Val v, s\" | LAss1Red: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \V:=e, s\ -ta\ \V:=e', s'\" | Red1LAss: "V < size l \ uf,P,t \1 \V:=(Val v), (h, l)\ -\\ \unit, (h, l[V := v])\" | AAcc1Red1: "uf,P,t \1 \a, s\ -ta\ \a', s'\ \ uf,P,t \1 \a\i\, s\ -ta\ \a'\i\, s'\" | AAcc1Red2: "uf,P,t \1 \i, s\ -ta\ \i', s'\ \ uf,P,t \1 \(Val a)\i\, s\ -ta\ \(Val a)\i'\, s'\" | Red1AAccNull: "uf,P,t \1 \null\Val i\, s\ -\\ \THROW NullPointer, s\" | Red1AAccBounds: "\ typeof_addr (hp s) a = \Array_type T n\; i sint i \ int n \ \ uf,P,t \1 \(addr a)\Val (Intg i)\, s\ -\\ \THROW ArrayIndexOutOfBounds, s\" | Red1AAcc: "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; heap_read h a (ACell (nat (sint i))) v \ \ uf,P,t \1 \(addr a)\Val (Intg i)\, (h, xs)\ -\ReadMem a (ACell (nat (sint i))) v\\ \Val v, (h, xs)\" | AAss1Red1: "uf,P,t \1 \a, s\ -ta\ \a', s'\ \ uf,P,t \1 \a\i\ := e, s\ -ta\ \a'\i\ := e, s'\" | AAss1Red2: "uf,P,t \1 \i, s\ -ta\ \i', s'\ \ uf,P,t \1 \(Val a)\i\ := e, s\ -ta\ \(Val a)\i'\ := e, s'\" | AAss1Red3: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \AAss (Val a) (Val i) e, s\ -ta\ \(Val a)\Val i\ := e', s'\" | Red1AAssNull: "uf,P,t \1 \AAss null (Val i) (Val e), s\ -\\ \THROW NullPointer, s\" | Red1AAssBounds: "\ typeof_addr (hp s) a = \Array_type T n\; i sint i \ int n \ \ uf,P,t \1 \AAss (addr a) (Val (Intg i)) (Val e), s\ -\\ \THROW ArrayIndexOutOfBounds, s\" | Red1AAssStore: "\ typeof_addr (hp s) a = \Array_type T n\; 0 <=s i; sint i < int n; typeof\<^bsub>hp s\<^esub> w = \U\; \ (P \ U \ T) \ \ uf,P,t \1 \AAss (addr a) (Val (Intg i)) (Val w), s\ -\\ \THROW ArrayStore, s\" | Red1AAss: "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; typeof\<^bsub>h\<^esub> w = Some U; P \ U \ T; heap_write h a (ACell (nat (sint i))) w h' \ \ uf,P,t \1 \AAss (addr a) (Val (Intg i)) (Val w), (h, l)\ -\WriteMem a (ACell (nat (sint i))) w\\ \unit, (h', l)\" | ALength1Red: "uf,P,t \1 \a, s\ -ta\ \a', s'\ \ uf,P,t \1 \a\length, s\ -ta\ \a'\length, s'\" | Red1ALength: "typeof_addr h a = \Array_type T n\ - \ uf,P,t \1 \addr a\length, (h, xs)\ -\\ \Val (Intg (word_of_int (int n))), (h, xs)\" + \ uf,P,t \1 \addr a\length, (h, xs)\ -\\ \Val (Intg (word_of_nat n)), (h, xs)\" | Red1ALengthNull: "uf,P,t \1 \null\length, s\ -\\ \THROW NullPointer, s\" | FAcc1Red: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \e\F{D}, s\ -ta\ \e'\F{D}, s'\" | Red1FAcc: "heap_read h a (CField D F) v \ uf,P,t \1 \(addr a)\F{D}, (h, xs)\ -\ReadMem a (CField D F) v\\ \Val v, (h, xs)\" | Red1FAccNull: "uf,P,t \1 \null\F{D}, s\ -\\ \THROW NullPointer, s\" | FAss1Red1: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \e\F{D}:=e2, s\ -ta\ \e'\F{D}:=e2, s'\" | FAss1Red2: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \FAss (Val v) F D e, s\ -ta\ \Val v\F{D}:=e', s'\" | Red1FAss: "heap_write h a (CField D F) v h' \ uf,P,t \1 \FAss (addr a) F D (Val v), (h, l)\ -\WriteMem a (CField D F) v\\ \unit, (h', l)\" | Red1FAssNull: "uf,P,t \1 \FAss null F D (Val v), s\ -\\ \THROW NullPointer, s\" | CAS1Red1: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \e\compareAndSwap(D\F, e2, e3), s\ -ta\ \e'\compareAndSwap(D\F, e2, e3), s'\" | CAS1Red2: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \Val v\compareAndSwap(D\F, e, e3), s\ -ta\ \Val v\compareAndSwap(D\F, e', e3), s'\" | CAS1Red3: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \Val v\compareAndSwap(D\F, Val v', e), s\ -ta\ \Val v\compareAndSwap(D\F, Val v', e'), s'\" | CAS1Null: "uf,P,t \1 \null\compareAndSwap(D\F, Val v, Val v'), s\ -\\ \THROW NullPointer, s\" | Red1CASSucceed: "\ heap_read h a (CField D F) v; heap_write h a (CField D F) v' h' \ \ uf,P,t \1 \addr a\compareAndSwap(D\F, Val v, Val v'), (h, l)\ -\ReadMem a (CField D F) v, WriteMem a (CField D F) v'\\ \true, (h', l)\" | Red1CASFail: "\ heap_read h a (CField D F) v''; v \ v'' \ \ uf,P,t \1 \addr a\compareAndSwap(D\F, Val v, Val v'), (h, l)\ -\ReadMem a (CField D F) v''\\ \false, (h, l)\" | Call1Obj: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \e\M(es), s\ -ta\ \e'\M(es), s'\" | Call1Params: "uf,P,t \1 \es, s\ [-ta\] \es',s'\ \ uf,P,t \1 \(Val v)\M(es),s\ -ta\ \(Val v)\M(es'),s'\" | Red1CallExternal: "\ typeof_addr (hp s) a = \T\; P \ class_type_of T sees M:Ts\Tr = Native in D; P,t \ \a\M(vs), hp s\ -ta\ext \va, h'\; e' = extRet2J1 ((addr a)\M(map Val vs)) va; s' = (h', lcl s) \ \ uf,P,t \1 \(addr a)\M(map Val vs), s\ -ta\ \e', s'\" | Red1CallNull: "uf,P,t \1 \null\M(map Val vs), s\ -\\ \THROW NullPointer, s\" | Block1Some: "V < length x \ uf,P,t \1 \{V:T=\v\; e}, (h, x)\ -\\ \{V:T=None; e}, (h, x[V := v])\" | Block1Red: "uf,P,t \1 \e, (h, x)\ -ta\ \e', (h', x')\ \ uf,P,t \1 \{V:T=None; e}, (h, x)\ -ta\ \{V:T=None; e'}, (h', x')\" | Red1Block: "uf,P,t \1 \{V:T=None; Val u}, s\ -\\ \Val u, s\" | Synchronized1Red1: "uf,P,t \1 \o', s\ -ta\ \o'', s'\ \ uf,P,t \1 \sync\<^bsub>V\<^esub> (o') e, s\ -ta\ \sync\<^bsub>V\<^esub> (o'') e, s'\" | Synchronized1Null: "V < length xs \ uf,P,t \1 \sync\<^bsub>V\<^esub> (null) e, (h, xs)\ -\\ \THROW NullPointer, (h, xs[V := Null])\" | Lock1Synchronized: "V < length xs \ uf,P,t \1 \sync\<^bsub>V\<^esub> (addr a) e, (h, xs)\ -\Lock\a, SyncLock a\\ \insync\<^bsub>V\<^esub> (a) e, (h, xs[V := Addr a])\" | Synchronized1Red2: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \insync\<^bsub>V\<^esub> (a) e, s\ -ta\ \insync\<^bsub>V\<^esub> (a) e', s'\" | Unlock1Synchronized: "\ xs ! V = Addr a'; V < length xs \ \ uf,P,t \1 \insync\<^bsub>V\<^esub> (a) (Val v), (h, xs)\ -\Unlock\a', SyncUnlock a'\\ \Val v, (h, xs)\" | Unlock1SynchronizedNull: "\ xs ! V = Null; V < length xs \ \ uf,P,t \1 \insync\<^bsub>V\<^esub> (a) (Val v), (h, xs)\ -\\ \THROW NullPointer, (h, xs)\" | Unlock1SynchronizedFail: "\ uf; xs ! V = Addr a'; V < length xs \ \ uf,P,t \1 \insync\<^bsub>V\<^esub> (a) (Val v), (h, xs)\ -\UnlockFail\a'\\ \THROW IllegalMonitorState, (h, xs)\" | Seq1Red: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \e;;e2, s\ -ta\ \e';;e2, s'\" | Red1Seq: "uf,P,t \1 \Seq (Val v) e, s\ -\\ \e, s\" | Cond1Red: "uf,P,t \1 \b, s\ -ta\ \b', s'\ \ uf,P,t \1 \if (b) e1 else e2, s\ -ta\ \if (b') e1 else e2, s'\" | Red1CondT: "uf,P,t \1 \if (true) e1 else e2, s\ -\\ \e1, s\" | Red1CondF: "uf,P,t \1 \if (false) e1 else e2, s\ -\\ \e2, s\" | Red1While: "uf,P,t \1 \while(b) c, s\ -\\ \if (b) (c;;while(b) c) else unit, s\" | Throw1Red: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \throw e, s\ -ta\ \throw e', s'\" | Red1ThrowNull: "uf,P,t \1 \throw null, s\ -\\ \THROW NullPointer, s\" | Try1Red: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ uf,P,t \1 \try e catch(C V) e2, s\ -ta\ \try e' catch(C V) e2, s'\" | Red1Try: "uf,P,t \1 \try (Val v) catch(C V) e2, s\ -\\ \Val v, s\" | Red1TryCatch: "\ typeof_addr h a = \Class_type D\; P \ D \\<^sup>* C; V < length x \ \ uf,P,t \1 \try (Throw a) catch(C V) e2, (h, x)\ -\\ \{V:Class C=None; e2}, (h, x[V := Addr a])\" | Red1TryFail: "\ typeof_addr (hp s) a = \Class_type D\; \ P \ D \\<^sup>* C \ \ uf,P,t \1 \try (Throw a) catch(C V) e2, s\ -\\ \Throw a, s\" | List1Red1: "uf,P,t \1 \e,s\ -ta\ \e',s'\ \ uf,P,t \1 \e#es,s\ [-ta\] \e'#es,s'\" | List1Red2: "uf,P,t \1 \es,s\ [-ta\] \es',s'\ \ uf,P,t \1 \Val v # es,s\ [-ta\] \Val v # es',s'\" | New1ArrayThrow: "uf,P,t \1 \newA T\Throw a\, s\ -\\ \Throw a, s\" | Cast1Throw: "uf,P,t \1 \Cast C (Throw a), s\ -\\ \Throw a, s\" | InstanceOf1Throw: "uf,P,t \1 \(Throw a) instanceof T, s\ -\\ \Throw a, s\" | Bin1OpThrow1: "uf,P,t \1 \(Throw a) \bop\ e\<^sub>2, s\ -\\ \Throw a, s\" | Bin1OpThrow2: "uf,P,t \1 \(Val v\<^sub>1) \bop\ (Throw a), s\ -\\ \Throw a, s\" | LAss1Throw: "uf,P,t \1 \V:=(Throw a), s\ -\\ \Throw a, s\" | AAcc1Throw1: "uf,P,t \1 \(Throw a)\i\, s\ -\\ \Throw a, s\" | AAcc1Throw2: "uf,P,t \1 \(Val v)\Throw a\, s\ -\\ \Throw a, s\" | AAss1Throw1: "uf,P,t \1 \(Throw a)\i\ := e, s\ -\\ \Throw a, s\" | AAss1Throw2: "uf,P,t \1 \(Val v)\Throw a\ := e, s\ -\\ \Throw a, s\" | AAss1Throw3: "uf,P,t \1 \AAss (Val v) (Val i) (Throw a), s\ -\\ \Throw a, s\" | ALength1Throw: "uf,P,t \1 \(Throw a)\length, s\ -\\ \Throw a, s\" | FAcc1Throw: "uf,P,t \1 \(Throw a)\F{D}, s\ -\\ \Throw a, s\" | FAss1Throw1: "uf,P,t \1 \(Throw a)\F{D}:=e\<^sub>2, s\ -\\ \Throw a, s\" | FAss1Throw2: "uf,P,t \1 \FAss (Val v) F D (Throw a), s\ -\\ \Throw a, s\" | CAS1Throw: "uf,P,t \1 \Throw a\compareAndSwap(D\F, e2, e3), s\ -\\ \Throw a, s\" | CAS1Throw2: "uf,P,t \1 \Val v\compareAndSwap(D\F, Throw a, e3), s\ -\\ \Throw a, s\" | CAS1Throw3: "uf,P,t \1 \Val v\compareAndSwap(D\F, Val v', Throw a), s\ -\\ \Throw a, s\" | Call1ThrowObj: "uf,P,t \1 \(Throw a)\M(es), s\ -\\ \Throw a, s\" | Call1ThrowParams: "\ es = map Val vs @ Throw a # es' \ \ uf,P,t \1 \(Val v)\M(es), s\ -\\ \Throw a, s\" | Block1Throw: "uf,P,t \1 \{V:T=None; Throw a}, s\ -\\ \Throw a, s\" | Synchronized1Throw1: "uf,P,t \1 \sync\<^bsub>V\<^esub> (Throw a) e, s\ -\\ \Throw a, s\" | Synchronized1Throw2: "\ xs ! V = Addr a'; V < length xs \ \ uf,P,t \1 \insync\<^bsub>V\<^esub> (a) Throw ad, (h, xs)\ -\Unlock\a', SyncUnlock a'\\ \Throw ad, (h, xs)\" | Synchronized1Throw2Fail: "\ uf; xs ! V = Addr a'; V < length xs \ \ uf,P,t \1 \insync\<^bsub>V\<^esub> (a) Throw ad, (h, xs)\ -\UnlockFail\a'\\ \THROW IllegalMonitorState, (h, xs)\" | Synchronized1Throw2Null: "\ xs ! V = Null; V < length xs \ \ uf,P,t \1 \insync\<^bsub>V\<^esub> (a) Throw ad, (h, xs)\ -\\ \THROW NullPointer, (h, xs)\" | Seq1Throw: "uf,P,t \1 \(Throw a);;e\<^sub>2, s\ -\\ \Throw a, s\" | Cond1Throw: "uf,P,t \1 \if (Throw a) e\<^sub>1 else e\<^sub>2, s\ -\\ \Throw a, s\" | Throw1Throw: "uf,P,t \1 \throw(Throw a), s\ -\\ \Throw a, s\" inductive_cases red1_cases: "uf,P,t \1 \new C, s\ -ta\ \e', s'\" "uf,P,t \1 \new T\e\, s\ -ta\ \e', s'\" "uf,P,t \1 \e \bop\ e', s\ -ta\ \e'', s'\" "uf,P,t \1 \Var V, s\ -ta\ \e', s'\" "uf,P,t \1 \V:=e, s\ -ta\ \e', s'\" "uf,P,t \1 \a\i\, s\ -ta\ \e', s'\" "uf,P,t \1 \a\i\ := e, s\ -ta\ \e', s'\" "uf,P,t \1 \a\length, s\ -ta\ \e', s'\" "uf,P,t \1 \e\F{D}, s\ -ta\ \e', s'\" "uf,P,t \1 \e\F{D} := e2, s\ -ta\ \e', s'\" "uf,P,t \1 \e\compareAndSwap(D\F, e', e''), s\ -ta\ \e''', s'\" "uf,P,t \1 \e\M(es), s\ -ta\ \e', s'\" "uf,P,t \1 \{V:T=vo; e}, s\ -ta\ \e', s'\" "uf,P,t \1 \sync\<^bsub>V\<^esub> (o') e, s\ -ta\ \e', s'\" "uf,P,t \1 \insync\<^bsub>V\<^esub> (a) e, s\ -ta\ \e', s'\" "uf,P,t \1 \e;;e', s\ -ta\ \e'', s'\" "uf,P,t \1 \throw e, s \ -ta\ \e', s'\" "uf,P,t \1 \try e catch(C V) e'', s \ -ta\ \e', s'\" inductive Red1 :: "bool \ 'addr J1_prog \ 'thread_id \ ('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list \ 'heap \ ('addr, 'thread_id, 'heap) J1_thread_action \ ('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list \ 'heap \ bool" ("_,_,_ \1 ((1\_'/_,/_\) -_\/ (1\_'/_,/_\))" [51,51,0,0,0,0,0,0,0,0] 81) for uf :: bool and P :: "'addr J1_prog" and t :: 'thread_id where red1Red: "uf,P,t \1 \e, (h, x)\ -ta\ \e', (h', x')\ \ uf,P,t \1 \(e, x)/exs, h\ -extTA2J1 P ta\ \(e', x')/exs, h'\" | red1Call: "\ call1 e = \(a, M, vs)\; typeof_addr h a = \U\; P \ class_type_of U sees M:Ts\T = \body\ in D; size vs = size Ts \ \ uf,P,t \1 \(e, x)/exs, h\ -\\ \(blocks1 0 (Class D#Ts) body, Addr a # vs @ replicate (max_vars body) undefined_value)/(e, x)#exs, h\" | red1Return: "final e' \ uf,P,t \1 \(e', x')/(e, x)#exs, h\ -\\ \(inline_call e' e, x)/exs, h\" abbreviation mred1g :: "bool \ 'addr J1_prog \ ('addr,'thread_id,('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list,'heap,'addr,('addr, 'thread_id) obs_event) semantics" where "mred1g uf P \ \t ((ex, exs), h) ta ((ex', exs'), h'). uf,P,t \1 \ex/exs, h\ -ta\ \ex'/exs', h'\" abbreviation mred1' :: "'addr J1_prog \ ('addr,'thread_id,('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list,'heap,'addr,('addr, 'thread_id) obs_event) semantics" where "mred1' \ mred1g False" abbreviation mred1 :: "'addr J1_prog \ ('addr,'thread_id,('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list,'heap,'addr,('addr, 'thread_id) obs_event) semantics" where "mred1 \ mred1g True" lemma red1_preserves_len: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ length (lcl s') = length (lcl s)" and reds1_preserves_len: "uf,P,t \1 \es, s\ [-ta\] \es', s'\ \ length (lcl s') = length (lcl s)" by(induct rule: red1_reds1.inducts)(auto) lemma reds1_preserves_elen: "uf,P,t \1 \es, s\ [-ta\] \es', s'\ \ length es' = length es" by(induct es arbitrary: es')(auto elim: reds1.cases) lemma red1_Val_iff [iff]: "\ uf,P,t \1 \Val v, s\ -ta\ \e', s'\" by(auto elim: red1.cases) lemma red1_Throw_iff [iff]: "\ uf,P,t \1 \Throw a, xs\ -ta\ \e', s'\" by(auto elim: red1.cases) lemma reds1_Nil_iff [iff]: "\ uf,P,t \1 \[], s\ [-ta\] \es', s'\" by(auto elim: reds1.cases) lemma reds1_Val_iff [iff]: "\ uf,P,t \1 \map Val vs, s\ [-ta\] \es', s'\" by(induct vs arbitrary: es')(auto elim: reds1.cases) lemma reds1_map_Val_Throw_iff [iff]: "\ uf,P,t \1 \map Val vs @ Throw a # es, s\ [-ta\] \es', s'\" by(induct vs arbitrary: es')(auto elim: reds1.cases elim!: red1_cases) lemma red1_max_vars_decr: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ max_vars e' \ max_vars e" and reds1_max_varss_decr: "uf,P,t \1 \es, s\ [-ta\] \es', s'\ \ max_varss es' \ max_varss es" by(induct rule: red1_reds1.inducts)(auto) lemma red1_new_thread_heap: "\uf,P,t \1 \e, s\ -ta\ \e', s'\; NewThread t' ex h \ set \ta\\<^bsub>t\<^esub> \ \ h = hp s'" and reds1_new_thread_heap: "\uf,P,t \1 \es, s\ [-ta\] \es', s'\; NewThread t' ex h \ set \ta\\<^bsub>t\<^esub> \ \ h = hp s'" apply(induct rule: red1_reds1.inducts) apply(fastforce dest: red_ext_new_thread_heap simp add: ta_upd_simps)+ done lemma red1_new_threadD: "\ uf,P,t \1 \e, s\ -ta\ \e', s'\; NewThread t' x H \ set \ta\\<^bsub>t\<^esub> \ \ \a M vs va T Ts Tr D. P,t \ \a\M(vs), hp s\ -ta\ext \va, hp s'\ \ typeof_addr (hp s) a = \T\ \ P \ class_type_of T sees M:Ts\Tr = Native in D" and reds1_new_threadD: "\ uf,P,t \1 \es, s\ [-ta\] \es', s'\; NewThread t' x H \ set \ta\\<^bsub>t\<^esub> \ \ \a M vs va T Ts Tr D. P,t \ \a\M(vs), hp s\ -ta\ext \va, hp s'\ \ typeof_addr (hp s) a = \T\ \ P \ class_type_of T sees M:Ts\Tr = Native in D" by(induct rule: red1_reds1.inducts)(fastforce simp add: ta_upd_simps)+ lemma red1_call_synthesized: "\ uf,P,t \1 \e, s\ -ta\ \e', s'\; call1 e = \aMvs\ \ \ synthesized_call P (hp s) aMvs" and reds1_calls_synthesized: "\ uf,P,t \1 \es, s\ [-ta\] \es', s'\; calls1 es = \aMvs\ \ \ synthesized_call P (hp s) aMvs" apply(induct rule: red1_reds1.inducts) apply(auto split: if_split_asm simp add: is_vals_conv append_eq_map_conv synthesized_call_conv) apply blast done lemma red1_preserves_B: "\ uf,P,t \1 \e, s\ -ta\ \e', s'\; \ e n\ \ \ e' n" and reds1_preserves_Bs: "\ uf,P,t \1 \es, s\ [-ta\] \es', s'\; \s es n\ \ \s es' n" by(induct arbitrary: n and n rule: red1_reds1.inducts)(auto) end context J1_heap begin lemma red1_hext_incr: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ hext (hp s) (hp s')" and reds1_hext_incr: "uf,P,t \1 \es, s\ [-ta\] \es', s'\ \ hext (hp s) (hp s')" by(induct rule: red1_reds1.inducts)(auto intro: hext_heap_ops red_external_hext) lemma Red1_hext_incr: "uf,P,t \1 \ex/exs,h\ -ta\ \ex'/exs',h'\ \ h \ h'" by(auto elim!: Red1.cases dest: red1_hext_incr) end subsection \Silent moves\ context J1_heap_base begin primrec \move1 :: "'m prog \ 'heap \ ('a, 'b, 'addr) exp \ bool" and \moves1 :: "'m prog \ 'heap \ ('a, 'b, 'addr) exp list \ bool" where "\move1 P h (new C) \ False" | "\move1 P h (newA T\e\) \ \move1 P h e \ (\a. e = Throw a)" | "\move1 P h (Cast U e) \ \move1 P h e \ final e" | "\move1 P h (e instanceof T) \ \move1 P h e \ final e" | "\move1 P h (e \bop\ e') \ \move1 P h e \ (\a. e = Throw a) \ (\v. e = Val v \ (\move1 P h e' \ final e'))" | "\move1 P h (Val v) \ False" | "\move1 P h (Var V) \ True" | "\move1 P h (V := e) \ \move1 P h e \ final e" | "\move1 P h (a\i\) \ \move1 P h a \ (\ad. a = Throw ad) \ (\v. a = Val v \ (\move1 P h i \ (\a. i = Throw a)))" | "\move1 P h (AAss a i e) \ \move1 P h a \ (\ad. a = Throw ad) \ (\v. a = Val v \ (\move1 P h i \ (\a. i = Throw a) \ (\v. i = Val v \ (\move1 P h e \ (\a. e = Throw a)))))" | "\move1 P h (a\length) \ \move1 P h a \ (\ad. a = Throw ad)" | "\move1 P h (e\F{D}) \ \move1 P h e \ (\a. e = Throw a)" | "\move1 P h (FAss e F D e') \ \move1 P h e \ (\a. e = Throw a) \ (\v. e = Val v \ (\move1 P h e' \ (\a. e' = Throw a)))" | "\move1 P h (e\compareAndSwap(D\F, e', e'')) \ \move1 P h e \ (\a. e = Throw a) \ (\v. e = Val v \ (\move1 P h e' \ (\a. e' = Throw a) \ (\v. e' = Val v \ (\move1 P h e'' \ (\a. e'' = Throw a)))))" | "\move1 P h (e\M(es)) \ \move1 P h e \ (\a. e = Throw a) \ (\v. e = Val v \ (\moves1 P h es \ (\vs a es'. es = map Val vs @ Throw a # es') \ (\vs. es = map Val vs \ (v = Null \ (\T C Ts Tr D. typeof\<^bsub>h\<^esub> v = \T\ \ class_type_of' T = \C\ \ P \ C sees M:Ts\Tr = Native in D \ \external_defs D M)))))" | "\move1 P h ({V:T=vo; e}) \ vo \ None \ \move1 P h e \ final e" | "\move1 P h (sync\<^bsub>V'\<^esub>(e) e') \ \move1 P h e \ (\a. e = Throw a)" | "\move1 P h (insync\<^bsub>V'\<^esub>(ad) e) \ \move1 P h e" | "\move1 P h (e;;e') \ \move1 P h e \ final e" | "\move1 P h (if (e) e' else e'') \ \move1 P h e \ final e" | "\move1 P h (while (e) e') = True" | "\move1 P h (throw e) \ \move1 P h e \ (\a. e = Throw a) \ e = null" | "\move1 P h (try e catch(C V) e') \ \move1 P h e \ final e" | "\moves1 P h [] \ False" | "\moves1 P h (e # es) \ \move1 P h e \ (\v. e = Val v \ \moves1 P h es)" fun \Move1 :: "'m prog \ 'heap \ (('a, 'b, 'addr) exp \ 'c) \ (('a, 'b, 'addr) exp \ 'd) list \ bool" where "\Move1 P h ((e, x), exs) = (\move1 P h e \ final e)" definition \red1g :: "bool \ 'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) \ bool" where "\red1g uf P t h exs e'xs' = (uf,P,t \1 \fst exs, (h, snd exs)\ -\\ \fst e'xs', (h, snd e'xs')\ \ \move1 P h (fst exs))" definition \reds1g :: "bool \ 'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 list \ 'addr locals1) \ ('addr expr1 list \ 'addr locals1) \ bool" where "\reds1g uf P t h esxs es'xs' = (uf,P,t \1 \fst esxs, (h, snd esxs)\ [-\\] \fst es'xs', (h, snd es'xs')\ \ \moves1 P h (fst esxs))" abbreviation \red1gt :: "bool \ 'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) \ bool" where "\red1gt uf P t h \ (\red1g uf P t h)^++" abbreviation \reds1gt :: "bool \ 'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 list \ 'addr locals1) \ ('addr expr1 list \ 'addr locals1) \ bool" where "\reds1gt uf P t h \ (\reds1g uf P t h)^++" abbreviation \red1gr :: "bool \ 'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) \ bool" where "\red1gr uf P t h \ (\red1g uf P t h)^**" abbreviation \reds1gr :: "bool \ 'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 list \ 'addr locals1) \ ('addr expr1 list \ 'addr locals1) \ bool" where "\reds1gr uf P t h \ (\reds1g uf P t h)^**" definition \Red1g :: "bool \ 'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ bool" where "\Red1g uf P t h exexs ex'exs' = (uf,P,t \1 \fst exexs/snd exexs, h\ -\\ \fst ex'exs'/snd ex'exs', h\ \ \Move1 P h exexs)" abbreviation \Red1gt :: "bool \ 'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ bool" where "\Red1gt uf P t h \ (\Red1g uf P t h)^++" abbreviation \Red1gr :: "bool \ 'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ bool" where "\Red1gr uf P t h \ (\Red1g uf P t h)^**" abbreviation \red1 :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) \ bool" where "\red1 \ \red1g True" abbreviation \reds1 :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 list \ 'addr locals1) \ ('addr expr1 list \ 'addr locals1) \ bool" where "\reds1 \ \reds1g True" abbreviation \red1t :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) \ bool" where "\red1t \ \red1gt True" abbreviation \reds1t :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 list \ 'addr locals1) \ ('addr expr1 list \ 'addr locals1) \ bool" where "\reds1t \ \reds1gt True" abbreviation \red1r :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) \ bool" where "\red1r \ \red1gr True" abbreviation \reds1r :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 list \ 'addr locals1) \ ('addr expr1 list \ 'addr locals1) \ bool" where "\reds1r \ \reds1gr True" abbreviation \Red1 :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ bool" where "\Red1 \ \Red1g True" abbreviation \Red1t :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ bool" where "\Red1t \ \Red1gt True" abbreviation \Red1r :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ bool" where "\Red1r \ \Red1gr True" abbreviation \red1' :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) \ bool" where "\red1' \ \red1g False" abbreviation \reds1' :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 list \ 'addr locals1) \ ('addr expr1 list \ 'addr locals1) \ bool" where "\reds1' \ \reds1g False" abbreviation \red1't :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) \ bool" where "\red1't \ \red1gt False" abbreviation \reds1't :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 list \ 'addr locals1) \ ('addr expr1 list \ 'addr locals1) \ bool" where "\reds1't \ \reds1gt False" abbreviation \red1'r :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) \ bool" where "\red1'r \ \red1gr False" abbreviation \reds1'r :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 list \ 'addr locals1) \ ('addr expr1 list \ 'addr locals1) \ bool" where "\reds1'r \ \reds1gr False" abbreviation \Red1' :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ bool" where "\Red1' \ \Red1g False" abbreviation \Red1't :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ bool" where "\Red1't \ \Red1gt False" abbreviation \Red1'r :: "'addr J1_prog \ 'thread_id \ 'heap \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ ('addr expr1 \ 'addr locals1) \ (('addr expr1 \ 'addr locals1) list) \ bool" where "\Red1'r \ \Red1gr False" abbreviation \MOVE1 :: "'m prog \ ((('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list) \ 'heap, ('addr, 'thread_id, 'heap) J1_thread_action) trsys" where "\MOVE1 P \ \(exexs, h) ta s. \Move1 P h exexs \ ta = \" lemma \move1_\moves1_intros: fixes e :: "('a, 'b, 'addr) exp" and es :: "('a, 'b, 'addr) exp list" shows \move1NewArray: "\move1 P h e \ \move1 P h (newA T\e\)" and \move1Cast: "\move1 P h e \ \move1 P h (Cast U e)" and \move1CastRed: "\move1 P h (Cast U (Val v))" and \move1InstanceOf: "\move1 P h e \ \move1 P h (e instanceof U)" and \move1InstanceOfRed: "\move1 P h ((Val v) instanceof U)" and \move1BinOp1: "\move1 P h e \ \move1 P h (e\bop\e')" and \move1BinOp2: "\move1 P h e \ \move1 P h (Val v\bop\e)" and \move1BinOp: "\move1 P h (Val v\bop\Val v')" and \move1Var: "\move1 P h (Var V)" and \move1LAss: "\move1 P h e \ \move1 P h (V := e)" and \move1LAssRed: "\move1 P h (V := Val v)" and \move1AAcc1: "\move1 P h e \ \move1 P h (e\e'\)" and \move1AAcc2: "\move1 P h e \ \move1 P h (Val v\e\)" and \move1AAss1: "\move1 P h e \ \move1 P h (AAss e e' e'')" and \move1AAss2: "\move1 P h e \ \move1 P h (AAss (Val v) e e')" and \move1AAss3: "\move1 P h e \ \move1 P h (AAss (Val v) (Val v') e)" and \move1ALength: "\move1 P h e \ \move1 P h (e\length)" and \move1FAcc: "\move1 P h e \ \move1 P h (e\F{D})" and \move1FAss1: "\move1 P h e \ \move1 P h (FAss e F D e')" and \move1FAss2: "\move1 P h e \ \move1 P h (FAss (Val v) F D e)" and \move1CAS1: "\move1 P h e \ \move1 P h (e\compareAndSwap(D\F, e', e''))" and \move1CAS2: "\move1 P h e \ \move1 P h (Val v\compareAndSwap(D\F, e, e''))" and \move1CAS3: "\move1 P h e \ \move1 P h (Val v\compareAndSwap(D\F, Val v', e))" and \move1CallObj: "\move1 P h obj \ \move1 P h (obj\M(ps))" and \move1CallParams: "\moves1 P h ps \ \move1 P h (Val v\M(ps))" and \move1Call: "(\T C Ts Tr D. \ typeof\<^bsub>h\<^esub> v = \T\; class_type_of' T = \C\; P \ C sees M:Ts\Tr = Native in D \ \ \external_defs D M) \ \move1 P h (Val v\M(map Val vs))" and \move1BlockSome: "\move1 P h {V:T=\v\; e}" and \move1Block: "\move1 P h e \ \move1 P h {V:T=None; e}" and \move1BlockRed: "\move1 P h {V:T=None; Val v}" and \move1Sync: "\move1 P h e \ \move1 P h (sync\<^bsub>V'\<^esub> (e) e')" and \move1InSync: "\move1 P h e \ \move1 P h (insync\<^bsub>V'\<^esub> (a) e)" and \move1Seq: "\move1 P h e \ \move1 P h (e;;e')" and \move1SeqRed: "\move1 P h (Val v;; e)" and \move1Cond: "\move1 P h e \ \move1 P h (if (e) e1 else e2)" and \move1CondRed: "\move1 P h (if (Val v) e1 else e2)" and \move1WhileRed: "\move1 P h (while (c) e)" and \move1Throw: "\move1 P h e \ \move1 P h (throw e)" and \move1ThrowNull: "\move1 P h (throw null)" and \move1Try: "\move1 P h e \ \move1 P h (try e catch(C V) e'')" and \move1TryRed: "\move1 P h (try Val v catch(C V) e)" and \move1TryThrow: "\move1 P h (try Throw a catch(C V) e)" and \move1NewArrayThrow: "\move1 P h (newA T\Throw a\)" and \move1CastThrow: "\move1 P h (Cast T (Throw a))" and \move1InstanceOfThrow: "\move1 P h ((Throw a) instanceof T)" and \move1BinOpThrow1: "\move1 P h (Throw a \bop\ e2)" and \move1BinOpThrow2: "\move1 P h (Val v \bop\ Throw a)" and \move1LAssThrow: "\move1 P h (V:=(Throw a))" and \move1AAccThrow1: "\move1 P h (Throw a\e\)" and \move1AAccThrow2: "\move1 P h (Val v\Throw a\)" and \move1AAssThrow1: "\move1 P h (AAss (Throw a) e e')" and \move1AAssThrow2: "\move1 P h (AAss (Val v) (Throw a) e')" and \move1AAssThrow3: "\move1 P h (AAss (Val v) (Val v') (Throw a))" and \move1ALengthThrow: "\move1 P h (Throw a\length)" and \move1FAccThrow: "\move1 P h (Throw a\F{D})" and \move1FAssThrow1: "\move1 P h (Throw a\F{D} := e)" and \move1FAssThrow2: "\move1 P h (FAss (Val v) F D (Throw a))" and \move1CASThrow1: "\move1 P h (CompareAndSwap (Throw a) D F e e')" and \move1CASThrow2: "\move1 P h (CompareAndSwap (Val v) D F (Throw a) e')" and \move1CASThrow3: "\move1 P h (CompareAndSwap (Val v) D F (Val v') (Throw a))" and \move1CallThrowObj: "\move1 P h (Throw a\M(es))" and \move1CallThrowParams: "\move1 P h (Val v\M(map Val vs @ Throw a # es))" and \move1BlockThrow: "\move1 P h {V:T=None; Throw a}" and \move1SyncThrow: "\move1 P h (sync\<^bsub>V'\<^esub> (Throw a) e)" and \move1SeqThrow: "\move1 P h (Throw a;;e)" and \move1CondThrow: "\move1 P h (if (Throw a) e1 else e2)" and \move1ThrowThrow: "\move1 P h (throw (Throw a))" and \moves1Hd: "\move1 P h e \ \moves1 P h (e # es)" and \moves1Tl: "\moves1 P h es \ \moves1 P h (Val v # es)" by fastforce+ lemma \moves1_map_Val [dest!]: "\moves1 P h (map Val es) \ False" by(induct es)(auto) lemma \moves1_map_Val_ThrowD [simp]: "\moves1 P h (map Val vs @ Throw a # es) = False" by(induct vs)(fastforce)+ lemma fixes e :: "('a, 'b, 'addr) exp" and es :: "('a, 'b, 'addr) exp list" shows \move1_not_call1: "call1 e = \(a, M, vs)\ \ \move1 P h e \ (synthesized_call P h (a, M, vs) \ \external' P h a M)" and \moves1_not_calls1: "calls1 es = \(a, M, vs)\ \ \moves1 P h es \ (synthesized_call P h (a, M, vs) \ \external' P h a M)" apply(induct e and es rule: call1.induct calls1.induct) apply(auto split: if_split_asm simp add: is_vals_conv) apply(fastforce simp add: synthesized_call_def map_eq_append_conv \external'_def \external_def dest: sees_method_fun)+ done lemma red1_\_taD: "\ uf,P,t \1 \e, s\ -ta\ \e', s'\; \move1 P (hp s) e \ \ ta = \" and reds1_\_taD: "\ uf,P,t \1 \es, s\ [-ta\] \es', s'\; \moves1 P (hp s) es \ \ ta = \" apply(induct rule: red1_reds1.inducts) apply(fastforce simp add: map_eq_append_conv \external'_def \external_def dest: \external'_red_external_TA_empty)+ done lemma \move1_heap_unchanged: "\ uf,P,t \1 \e, s\ -ta\ \e', s'\; \move1 P (hp s) e \ \ hp s' = hp s" and \moves1_heap_unchanged: "\ uf,P,t \1 \es, s\ [-ta\] \es', s'\; \moves1 P (hp s) es \ \ hp s' = hp s" apply(induct rule: red1_reds1.inducts) apply(auto) apply(fastforce simp add: map_eq_append_conv \external'_def \external_def dest: \external'_red_external_heap_unchanged)+ done lemma \Move1_iff: "\Move1 P h exexs \ (let ((e, _), _) = exexs in \move1 P h e \ final e)" by(cases exexs)(auto) lemma \red1_iff [iff]: "\red1g uf P t h (e, xs) (e', xs') = (uf,P,t \1 \e, (h, xs)\ -\\ \e', (h, xs')\ \ \move1 P h e)" by(simp add: \red1g_def) lemma \reds1_iff [iff]: "\reds1g uf P t h (es, xs) (es', xs') = (uf,P,t \1 \es, (h, xs)\ [-\\] \es', (h, xs')\ \ \moves1 P h es)" by(simp add: \reds1g_def) lemma \red1t_1step: "\ uf,P,t \1 \e, (h, xs)\ -\\ \e', (h, xs')\; \move1 P h e \ \ \red1gt uf P t h (e, xs) (e', xs')" by(blast intro: tranclp.r_into_trancl) lemma \red1t_2step: "\ uf,P,t \1 \e, (h, xs)\ -\\ \e', (h, xs')\; \move1 P h e; uf,P,t \1 \e', (h, xs')\ -\\ \e'', (h, xs'')\; \move1 P h e' \ \ \red1gt uf P t h (e, xs) (e'', xs'')" by(blast intro: tranclp.trancl_into_trancl[OF \red1t_1step]) lemma \red1t_3step: "\ uf,P,t \1 \e, (h, xs)\ -\\ \e', (h, xs')\; \move1 P h e; uf,P,t \1 \e', (h, xs')\ -\\ \e'', (h, xs'')\; \move1 P h e'; uf,P,t \1 \e'', (h, xs'')\ -\\ \e''', (h, xs''')\; \move1 P h e'' \ \ \red1gt uf P t h (e, xs) (e''', xs''')" by(blast intro: tranclp.trancl_into_trancl[OF \red1t_2step]) lemma \reds1t_1step: "\ uf,P,t \1 \es, (h, xs)\ [-\\] \es', (h, xs')\; \moves1 P h es \ \ \reds1gt uf P t h (es, xs) (es', xs')" by(blast intro: tranclp.r_into_trancl) lemma \reds1t_2step: "\ uf,P,t \1 \es, (h, xs)\ [-\\] \es', (h, xs')\; \moves1 P h es; uf,P,t \1 \es', (h, xs')\ [-\\] \es'', (h, xs'')\; \moves1 P h es' \ \ \reds1gt uf P t h (es, xs) (es'', xs'')" by(blast intro: tranclp.trancl_into_trancl[OF \reds1t_1step]) lemma \reds1t_3step: "\ uf,P,t \1 \es, (h, xs)\ [-\\] \es', (h, xs')\; \moves1 P h es; uf,P,t \1 \es', (h, xs')\ [-\\] \es'', (h, xs'')\; \moves1 P h es'; uf,P,t \1 \es'', (h, xs'')\ [-\\] \es''', (h, xs''')\; \moves1 P h es'' \ \ \reds1gt uf P t h (es, xs) (es''', xs''')" by(blast intro: tranclp.trancl_into_trancl[OF \reds1t_2step]) lemma \red1r_1step: "\ uf,P,t \1 \e, (h, xs)\ -\\ \e', (h, xs')\; \move1 P h e \ \ \red1gr uf P t h (e, xs) (e', xs')" by(blast intro: r_into_rtranclp) lemma \red1r_2step: "\ uf,P,t \1 \e, (h, xs)\ -\\ \e', (h, xs')\; \move1 P h e; uf,P,t \1 \e', (h, xs')\ -\\ \e'', (h, xs'')\; \move1 P h e' \ \ \red1gr uf P t h (e, xs) (e'', xs'')" by(blast intro: rtranclp.rtrancl_into_rtrancl[OF \red1r_1step]) lemma \red1r_3step: "\ uf,P,t \1 \e, (h, xs)\ -\\ \e', (h, xs')\; \move1 P h e; uf,P,t \1 \e', (h, xs')\ -\\ \e'', (h, xs'')\; \move1 P h e'; uf,P,t \1 \e'', (h, xs'')\ -\\ \e''', (h, xs''')\; \move1 P h e'' \ \ \red1gr uf P t h (e, xs) (e''', xs''')" by(blast intro: rtranclp.rtrancl_into_rtrancl[OF \red1r_2step]) lemma \reds1r_1step: "\ uf,P,t \1 \es, (h, xs)\ [-\\] \es', (h, xs')\; \moves1 P h es \ \ \reds1gr uf P t h (es, xs) (es', xs')" by(blast intro: r_into_rtranclp) lemma \reds1r_2step: "\ uf,P,t \1 \es, (h, xs)\ [-\\] \es', (h, xs')\; \moves1 P h es; uf,P,t \1 \es', (h, xs')\ [-\\] \es'', (h, xs'')\; \moves1 P h es' \ \ \reds1gr uf P t h (es, xs) (es'', xs'')" by(blast intro: rtranclp.rtrancl_into_rtrancl[OF \reds1r_1step]) lemma \reds1r_3step: "\ uf,P,t \1 \es, (h, xs)\ [-\\] \es', (h, xs')\; \moves1 P h es; uf,P,t \1 \es', (h, xs')\ [-\\] \es'', (h, xs'')\; \moves1 P h es'; uf,P,t \1 \es'', (h, xs'')\ [-\\] \es''', (h, xs''')\; \moves1 P h es'' \ \ \reds1gr uf P t h (es, xs) (es''', xs''')" by(blast intro: rtranclp.rtrancl_into_rtrancl[OF \reds1r_2step]) lemma \red1t_preserves_len: "\red1gt uf P t h (e, xs) (e', xs') \ length xs' = length xs" by(induct rule: tranclp_induct2)(auto dest: red1_preserves_len) lemma \red1r_preserves_len: "\red1gr uf P t h (e, xs) (e', xs') \ length xs' = length xs" by(induct rule: rtranclp_induct2)(auto dest: red1_preserves_len) lemma \red1t_inj_\reds1t: "\red1gt uf P t h (e, xs) (e', xs') \ \reds1gt uf P t h (e # es, xs) (e' # es, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl List1Red1 \moves1Hd) lemma \reds1t_cons_\reds1t: "\reds1gt uf P t h (es, xs) (es', xs') \ \reds1gt uf P t h (Val v # es, xs) (Val v # es', xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl List1Red2 \moves1Tl) lemma \red1r_inj_\reds1r: "\red1gr uf P t h (e, xs) (e', xs') \ \reds1gr uf P t h (e # es, xs) (e' # es, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl List1Red1 \moves1Hd) lemma \reds1r_cons_\reds1r: "\reds1gr uf P t h (es, xs) (es', xs') \ \reds1gr uf P t h (Val v # es, xs) (Val v # es', xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl List1Red2 \moves1Tl) lemma NewArray_\red1t_xt: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (newA T\e\, xs) (newA T\e'\, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl New1ArrayRed \move1NewArray) lemma Cast_\red1t_xt: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (Cast T e, xs) (Cast T e', xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl Cast1Red \move1Cast) lemma InstanceOf_\red1t_xt: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (e instanceof T, xs) (e' instanceof T, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl InstanceOf1Red \move1InstanceOf) lemma BinOp_\red1t_xt1: "\red1gt uf P t h (e1, xs) (e1', xs') \ \red1gt uf P t h (e1 \bop\ e2, xs) (e1' \bop\ e2, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl Bin1OpRed1 \move1BinOp1) lemma BinOp_\red1t_xt2: "\red1gt uf P t h (e2, xs) (e2', xs') \ \red1gt uf P t h (Val v \bop\ e2, xs) (Val v \bop\ e2', xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl Bin1OpRed2 \move1BinOp2) lemma LAss_\red1t: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (V := e, xs) (V := e', xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl LAss1Red \move1LAss) lemma AAcc_\red1t_xt1: "\red1gt uf P t h (a, xs) (a', xs') \ \red1gt uf P t h (a\i\, xs) (a'\i\, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl AAcc1Red1 \move1AAcc1) lemma AAcc_\red1t_xt2: "\red1gt uf P t h (i, xs) (i', xs') \ \red1gt uf P t h (Val a\i\, xs) (Val a\i'\, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl AAcc1Red2 \move1AAcc2) lemma AAss_\red1t_xt1: "\red1gt uf P t h (a, xs) (a', xs') \ \red1gt uf P t h (a\i\ := e, xs) (a'\i\ := e, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl AAss1Red1 \move1AAss1) lemma AAss_\red1t_xt2: "\red1gt uf P t h (i, xs) (i', xs') \ \red1gt uf P t h (Val a\i\ := e, xs) (Val a\i'\ := e, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl AAss1Red2 \move1AAss2) lemma AAss_\red1t_xt3: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (Val a\Val i\ := e, xs) (Val a\Val i\ := e', xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl AAss1Red3 \move1AAss3) lemma ALength_\red1t_xt: "\red1gt uf P t h (a, xs) (a', xs') \ \red1gt uf P t h (a\length, xs) (a'\length, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl ALength1Red \move1ALength) lemma FAcc_\red1t_xt: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (e\F{D}, xs) (e'\F{D}, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl FAcc1Red \move1FAcc) lemma FAss_\red1t_xt1: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (e\F{D} := e2, xs) (e'\F{D} := e2, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl FAss1Red1 \move1FAss1) lemma FAss_\red1t_xt2: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (Val v\F{D} := e, xs) (Val v\F{D} := e', xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl FAss1Red2 \move1FAss2) lemma CAS_\red1t_xt1: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (e\compareAndSwap(D\F, e2, e3), xs) (e'\compareAndSwap(D\F, e2, e3), xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl CAS1Red1) lemma CAS_\red1t_xt2: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (Val v\compareAndSwap(D\F, e, e3), xs) (Val v\compareAndSwap(D\F, e', e3), xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl CAS1Red2) lemma CAS_\red1t_xt3: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (Val v\compareAndSwap(D\F, Val v', e), xs) (Val v\compareAndSwap(D\F, Val v', e'), xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl CAS1Red3) lemma Call_\red1t_obj: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (e\M(ps), xs) (e'\M(ps), xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl Call1Obj \move1CallObj) lemma Call_\red1t_param: "\reds1gt uf P t h (es, xs) (es', xs') \ \red1gt uf P t h (Val v\M(es), xs) (Val v\M(es'), xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl Call1Params \move1CallParams) lemma Block_None_\red1t_xt: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h ({V:T=None; e}, xs) ({V:T=None; e'}, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl \move1Block elim!: Block1Red) lemma Block_\red1t_Some: "\ \red1gt uf P t h (e, xs[V := v]) (e', xs'); V < length xs \ \ \red1gt uf P t h ({V:Ty=\v\; e}, xs) ({V:Ty=None; e'}, xs')" by(blast intro: tranclp_into_tranclp2 Block1Some \move1BlockSome Block_None_\red1t_xt) lemma Sync_\red1t_xt: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (sync\<^bsub>V\<^esub> (e) e2, xs) (sync\<^bsub>V\<^esub> (e') e2, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl Synchronized1Red1 \move1Sync) lemma InSync_\red1t_xt: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (insync\<^bsub>V\<^esub> (a) e, xs) (insync\<^bsub>V\<^esub> (a) e', xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl Synchronized1Red2 \move1InSync) lemma Seq_\red1t_xt: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (e;;e2, xs) (e';;e2, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl Seq1Red \move1Seq) lemma Cond_\red1t_xt: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (if (e) e1 else e2, xs) (if (e') e1 else e2, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl Cond1Red \move1Cond) lemma Throw_\red1t_xt: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (throw e, xs) (throw e', xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl Throw1Red \move1Throw) lemma Try_\red1t_xt: "\red1gt uf P t h (e, xs) (e', xs') \ \red1gt uf P t h (try e catch(C V) e2, xs) (try e' catch(C V) e2, xs')" by(induct rule: tranclp_induct2)(auto intro: tranclp.trancl_into_trancl Try1Red \move1Try) lemma NewArray_\red1r_xt: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (newA T\e\, xs) (newA T\e'\, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl New1ArrayRed \move1NewArray) lemma Cast_\red1r_xt: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (Cast T e, xs) (Cast T e', xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl Cast1Red \move1Cast) lemma InstanceOf_\red1r_xt: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (e instanceof T, xs) (e' instanceof T, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl InstanceOf1Red \move1InstanceOf) lemma BinOp_\red1r_xt1: "\red1gr uf P t h (e1, xs) (e1', xs') \ \red1gr uf P t h (e1 \bop\ e2, xs) (e1' \bop\ e2, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl Bin1OpRed1 \move1BinOp1) lemma BinOp_\red1r_xt2: "\red1gr uf P t h (e2, xs) (e2', xs') \ \red1gr uf P t h (Val v \bop\ e2, xs) (Val v \bop\ e2', xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl Bin1OpRed2 \move1BinOp2) lemma LAss_\red1r: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (V := e, xs) (V := e', xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl LAss1Red \move1LAss) lemma AAcc_\red1r_xt1: "\red1gr uf P t h (a, xs) (a', xs') \ \red1gr uf P t h (a\i\, xs) (a'\i\, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl AAcc1Red1 \move1AAcc1) lemma AAcc_\red1r_xt2: "\red1gr uf P t h (i, xs) (i', xs') \ \red1gr uf P t h (Val a\i\, xs) (Val a\i'\, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl AAcc1Red2 \move1AAcc2) lemma AAss_\red1r_xt1: "\red1gr uf P t h (a, xs) (a', xs') \ \red1gr uf P t h (a\i\ := e, xs) (a'\i\ := e, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl AAss1Red1 \move1AAss1) lemma AAss_\red1r_xt2: "\red1gr uf P t h (i, xs) (i', xs') \ \red1gr uf P t h (Val a\i\ := e, xs) (Val a\i'\ := e, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl AAss1Red2 \move1AAss2) lemma AAss_\red1r_xt3: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (Val a\Val i\ := e, xs) (Val a\Val i\ := e', xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl AAss1Red3 \move1AAss3) lemma ALength_\red1r_xt: "\red1gr uf P t h (a, xs) (a', xs') \ \red1gr uf P t h (a\length, xs) (a'\length, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl ALength1Red \move1ALength) lemma FAcc_\red1r_xt: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (e\F{D}, xs) (e'\F{D}, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl FAcc1Red \move1FAcc) lemma FAss_\red1r_xt1: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (e\F{D} := e2, xs) (e'\F{D} := e2, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl FAss1Red1 \move1FAss1) lemma FAss_\red1r_xt2: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (Val v\F{D} := e, xs) (Val v\F{D} := e', xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl FAss1Red2 \move1FAss2) lemma CAS_\red1r_xt1: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (e\compareAndSwap(D\F, e2, e3), xs) (e'\compareAndSwap(D\F, e2, e3), xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl CAS1Red1) lemma CAS_\red1r_xt2: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (Val v\compareAndSwap(D\F, e, e3), xs) (Val v\compareAndSwap(D\F, e', e3), xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl CAS1Red2) lemma CAS_\red1r_xt3: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (Val v\compareAndSwap(D\F, Val v', e), xs) (Val v\compareAndSwap(D\F, Val v', e'), xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl CAS1Red3) lemma Call_\red1r_obj: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (e\M(ps), xs) (e'\M(ps), xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl Call1Obj \move1CallObj) lemma Call_\red1r_param: "\reds1gr uf P t h (es, xs) (es', xs') \ \red1gr uf P t h (Val v\M(es), xs) (Val v\M(es'), xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl Call1Params \move1CallParams) lemma Block_None_\red1r_xt: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h ({V:T=None; e}, xs) ({V:T=None; e'}, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl \move1Block elim!: Block1Red) lemma Block_\red1r_Some: "\ \red1gr uf P t h (e, xs[V := v]) (e', xs'); V < length xs \ \ \red1gr uf P t h ({V:Ty=\v\; e}, xs) ({V:Ty=None; e'}, xs')" by(blast intro: converse_rtranclp_into_rtranclp Block1Some \move1BlockSome Block_None_\red1r_xt) lemma Sync_\red1r_xt: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (sync\<^bsub>V\<^esub> (e) e2, xs) (sync\<^bsub>V\<^esub> (e') e2, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl Synchronized1Red1 \move1Sync) lemma InSync_\red1r_xt: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (insync\<^bsub>V\<^esub> (a) e, xs) (insync\<^bsub>V\<^esub> (a) e', xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl Synchronized1Red2 \move1InSync) lemma Seq_\red1r_xt: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (e;;e2, xs) (e';;e2, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl Seq1Red \move1Seq) lemma Cond_\red1r_xt: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (if (e) e1 else e2, xs) (if (e') e1 else e2, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl Cond1Red \move1Cond) lemma Throw_\red1r_xt: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (throw e, xs) (throw e', xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl Throw1Red \move1Throw) lemma Try_\red1r_xt: "\red1gr uf P t h (e, xs) (e', xs') \ \red1gr uf P t h (try e catch(C V) e2, xs) (try e' catch(C V) e2, xs')" by(induct rule: rtranclp_induct2)(auto intro: rtranclp.rtrancl_into_rtrancl Try1Red \move1Try) lemma \red1t_ThrowD [dest]: "\red1gt uf P t h (Throw a, xs) (e'', xs'') \ e'' = Throw a \ xs'' = xs" by(induct rule: tranclp_induct2)(auto) lemma \red1r_ThrowD [dest]: "\red1gr uf P t h (Throw a, xs) (e'', xs'') \ e'' = Throw a \ xs'' = xs" by(induct rule: rtranclp_induct2)(auto) lemma \Red1_conv [iff]: "\Red1g uf P t h (ex, exs) (ex', exs') = (uf,P,t \1 \ex/exs, h\ -\\ \ex'/exs', h\ \ \Move1 P h (ex, exs))" by(simp add: \Red1g_def) lemma \red1t_into_\Red1t: "\red1gt uf P t h (e, xs) (e'', xs'') \ \Red1gt uf P t h ((e, xs), exs) ((e'', xs''), exs)" by(induct rule: tranclp_induct2)(fastforce dest: red1Red intro: \move1Block tranclp.intros)+ lemma \red1r_into_\Red1r: "\red1gr uf P t h (e, xs) (e'', xs'') \ \Red1gr uf P t h ((e, xs), exs) ((e'', xs''), exs)" by(induct rule: rtranclp_induct2)(fastforce dest: red1Red intro: \move1Block rtranclp.intros)+ lemma red1_max_vars: "uf,P,t \1 \e, s\ -ta\ \e', s'\ \ max_vars e' \ max_vars e" and reds1_max_varss: "uf,P,t \1 \es, s\ [-ta\] \es', s'\ \ max_varss es' \ max_varss es" by(induct rule: red1_reds1.inducts) auto lemma \red1t_max_vars: "\red1gt uf P t h (e, xs) (e', xs') \ max_vars e' \ max_vars e" by(induct rule: tranclp_induct2)(auto dest: red1_max_vars) lemma \red1r_max_vars: "\red1gr uf P t h (e, xs) (e', xs') \ max_vars e' \ max_vars e" by(induct rule: rtranclp_induct2)(auto dest: red1_max_vars) lemma \red1r_Val: "\red1gr uf P t h (Val v, xs) s' \ s' = (Val v, xs)" proof assume "\red1gr uf P t h (Val v, xs) s'" thus "s' = (Val v, xs)" by induct(auto) qed auto lemma \red1t_Val: "\red1gt uf P t h (Val v, xs) s' \ False" proof assume "\red1gt uf P t h (Val v, xs) s'" thus False by induct auto qed auto lemma \reds1r_map_Val: "\reds1gr uf P t h (map Val vs, xs) s' \ s' = (map Val vs, xs)" proof assume "\reds1gr uf P t h (map Val vs, xs) s'" thus "s' = (map Val vs, xs)" by induct auto qed auto lemma \reds1t_map_Val: "\reds1gt uf P t h (map Val vs, xs) s' \ False" proof assume "\reds1gt uf P t h (map Val vs, xs) s'" thus "False" by induct auto qed auto lemma \reds1r_map_Val_Throw: "\reds1gr uf P t h (map Val vs @ Throw a # es, xs) s' \ s' = (map Val vs @ Throw a # es, xs)" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by induct auto qed auto lemma \reds1t_map_Val_Throw: "\reds1gt uf P t h (map Val vs @ Throw a # es, xs) s' \ False" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by induct auto qed auto lemma \red1r_Throw: "\red1gr uf P t h (Throw a, xs) s' \ s' = (Throw a, xs)" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by induct auto qed simp lemma \red1t_Throw: "\red1gt uf P t h (Throw a, xs) s' \ False" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by induct auto qed simp lemma red1_False_into_red1_True: "False,P,t \1 \e, s\ -ta\ \e', s'\ \ True,P,t \1 \e, s\ -ta\ \e', s'\" and reds1_False_into_reds1_True: "False,P,t \1 \es, s\ [-ta\] \es', s'\ \ True,P,t \1 \es, s\ [-ta\] \es', s'\" -by(induct rule: red1_reds1.inducts)(auto intro: red1_reds1.intros) + by (induct rule: red1_reds1.inducts) (auto intro: red1_reds1.intros) lemma Red1_False_into_Red1_True: assumes "False,P,t \1 \ex/exs,shr s\ -ta\ \ex'/exs',m'\" shows "True,P,t \1 \ex/exs,shr s\ -ta\ \ex'/exs',m'\" using assms by(cases)(auto dest: Red1.intros red1_False_into_red1_True) lemma red1_Suspend_is_call: "\ uf,P,t \1 \e, s\ -ta\ \e', s'\; Suspend w \ set \ta\\<^bsub>w\<^esub> \ \ call1 e' \ None" and reds_Suspend_is_calls: "\ uf,P,t \1 \es, s\ [-ta\] \es', s'\; Suspend w \ set \ta\\<^bsub>w\<^esub> \ \ calls1 es' \ None" by(induct rule: red1_reds1.inducts)(auto dest: red_external_Suspend_StaySame) lemma Red1_Suspend_is_call: "\ uf,P,t \1 \(e, xs)/exs, h\ -ta\ \(e', xs')/exs', h'\; Suspend w \ set \ta\\<^bsub>w\<^esub> \ \ call1 e' \ None" by(auto elim!: Red1.cases dest: red1_Suspend_is_call) lemma Red1_mthr: "multithreaded final_expr1 (mred1g uf P)" by(unfold_locales)(fastforce elim!: Red1.cases dest: red1_new_thread_heap)+ lemma red1_\move1_heap_unchanged: "\ uf,P,t \1 \e, s\ -ta\ \e', s'\; \move1 P (hp s) e \ \ hp s' = hp s" and red1_\moves1_heap_unchanged: "\ uf,P,t \1 \es, s\ [-ta\] \es', s'\; \moves1 P (hp s) es \ \ hp s' = hp s" apply(induct rule: red1_reds1.inducts) apply(fastforce simp add: map_eq_append_conv \external'_def \external_def dest: \external'_red_external_heap_unchanged)+ done lemma Red1_\mthr_wf: "\multithreaded_wf final_expr1 (mred1g uf P) (\MOVE1 P)" proof - interpret multithreaded final_expr1 "mred1g uf P" convert_RA by(rule Red1_mthr) show ?thesis proof fix x1 m1 t ta1 x1' m1' assume "mred1g uf P t (x1, m1) ta1 (x1', m1')" "\MOVE1 P (x1, m1) ta1 (x1', m1')" thus "m1 = m1'" by(cases x1)(fastforce elim!: Red1.cases dest: red1_\move1_heap_unchanged) next fix s ta s' assume "\MOVE1 P s ta s'" thus "ta = \" by(simp add: split_beta) qed qed end sublocale J1_heap_base < Red1_mthr: \multithreaded_wf final_expr1 "mred1g uf P" convert_RA "\MOVE1 P" for uf P by(rule Red1_\mthr_wf) context J1_heap_base begin lemma \Red1't_into_Red1'_\mthr_silent_movet: "\Red1gt uf P t h (ex2, exs2) (ex2'', exs2'') \ Red1_mthr.silent_movet uf P t ((ex2, exs2), h) ((ex2'', exs2''), h)" apply(induct rule: tranclp_induct2) apply clarsimp apply(rule tranclp.r_into_trancl) apply(simp add: Red1_mthr.silent_move_iff) apply(erule tranclp.trancl_into_trancl) apply(simp add: Red1_mthr.silent_move_iff) done lemma \Red1t_into_Red1'_\mthr_silent_moves: "\Red1gt uf P t h (ex2, exs2) (ex2'', exs2'') \ Red1_mthr.silent_moves uf P t ((ex2, exs2), h) ((ex2'', exs2''), h)" by(rule tranclp_into_rtranclp)(rule \Red1't_into_Red1'_\mthr_silent_movet) lemma \Red1'r_into_Red1'_\mthr_silent_moves: "\Red1gr uf P t h (ex, exs) (ex', exs') \ Red1_mthr.silent_moves uf P t ((ex, exs), h) ((ex', exs'), h)" apply(induct rule: rtranclp_induct2) apply blast apply(erule rtranclp.rtrancl_into_rtrancl) apply(simp add: Red1_mthr.silent_move_iff) done lemma \Red1r_rtranclpD: "\Red1gr uf P t h s s' \ \trsys.silent_moves (mred1g uf P t) (\MOVE1 P) (s, h) (s', h)" apply(induct rule: rtranclp_induct) apply(auto elim!: rtranclp.rtrancl_into_rtrancl intro: \trsys.silent_move.intros) done lemma \Red1t_tranclpD: "\Red1gt uf P t h s s' \ \trsys.silent_movet (mred1g uf P t) (\MOVE1 P) (s, h) (s', h)" apply(induct rule: tranclp_induct) apply(rule tranclp.r_into_trancl) apply(auto elim!: tranclp.trancl_into_trancl intro!: \trsys.silent_move.intros simp: \Red1g_def split_def) done lemma \mreds1_Val_Nil: "\trsys.silent_moves (mred1g uf P t) (\MOVE1 P) (((Val v, xs), []), h) s \ s = (((Val v, xs), []), h)" proof assume "\trsys.silent_moves (mred1g uf P t) (\MOVE1 P) (((Val v, xs), []), h) s" thus "s = (((Val v, xs), []), h)" by induct(auto elim!: Red1_mthr.silent_move.cases Red1.cases) qed auto lemma \mreds1_Throw_Nil: "\trsys.silent_moves (mred1g uf P t) (\MOVE1 P) (((Throw a, xs), []), h) s \ s = (((Throw a, xs), []), h)" proof assume "\trsys.silent_moves (mred1g uf P t) (\MOVE1 P) (((Throw a, xs), []), h) s" thus "s = (((Throw a, xs), []), h)" by induct(auto elim!: Red1_mthr.silent_move.cases Red1.cases) qed auto end end diff --git a/thys/JinjaThreads/J/SmallStep.thy b/thys/JinjaThreads/J/SmallStep.thy --- a/thys/JinjaThreads/J/SmallStep.thy +++ b/thys/JinjaThreads/J/SmallStep.thy @@ -1,740 +1,740 @@ (* Title: JinjaThreads/J/SmallStep.thy Author: Tobias Nipkow, Andreas Lochbihler *) section \Small Step Semantics\ theory SmallStep imports Expr State JHeap begin type_synonym ('addr, 'thread_id, 'heap) J_thread_action = "('addr, 'thread_id, 'addr expr \ 'addr locals,'heap) Jinja_thread_action" type_synonym ('addr, 'thread_id, 'heap) J_state = "('addr,'thread_id,'addr expr \ 'addr locals,'heap,'addr) state" (* pretty printing for J_thread_action type *) print_translation \ let fun tr' [a1, t , Const (@{type_syntax "prod"}, _) $ (Const (@{type_syntax "exp"}, _) $ Const (@{type_syntax "String.literal"}, _) $ Const (@{type_syntax "unit"}, _) $ a2) $ (Const (@{type_syntax "fun"}, _) $ Const (@{type_syntax "String.literal"}, _) $ (Const (@{type_syntax "option"}, _) $ (Const (@{type_syntax "val"}, _) $ a3))) , h] = if a1 = a2 andalso a2 = a3 then Syntax.const @{type_syntax "J_thread_action"} $ a1 $ t $ h else raise Match; in [(@{type_syntax "Jinja_thread_action"}, K tr')] end \ typ "('addr,'thread_id,'heap) J_thread_action" (* pretty printing for J_state type *) print_translation \ let fun tr' [a1, t , Const (@{type_syntax "prod"}, _) $ (Const (@{type_syntax "exp"}, _) $ Const (@{type_syntax "String.literal"}, _) $ Const (@{type_syntax "unit"}, _) $ a2) $ (Const (@{type_syntax "fun"}, _) $ Const (@{type_syntax "String.literal"}, _) $ (Const (@{type_syntax "option"}, _) $ (Const (@{type_syntax "val"}, _) $ a3))) , h, a4] = if a1 = a2 andalso a2 = a3 andalso a3 = a4 then Syntax.const @{type_syntax "J_state"} $ a1 $ t $ h else raise Match; in [(@{type_syntax "state"}, K tr')] end \ typ "('addr, 'thread_id, 'heap) J_state" definition extNTA2J :: "'addr J_prog \ (cname \ mname \ 'addr) \ 'addr expr \ 'addr locals" where "extNTA2J P = (\(C, M, a). let (D,Ts,T,meth) = method P C M; (pns,body) = the meth in ({this:Class D=\Addr a\; body}, Map.empty))" abbreviation J_local_start :: "cname \ mname \ ty list \ ty \ 'addr J_mb \ 'addr val list \ 'addr expr \ 'addr locals" where "J_local_start \ \C M Ts T (pns, body) vs. (blocks (this # pns) (Class C # Ts) (Null # vs) body, Map.empty)" abbreviation (in J_heap_base) J_start_state :: "'addr J_prog \ cname \ mname \ 'addr val list \ ('addr, 'thread_id, 'heap) J_state" where "J_start_state \ start_state J_local_start" lemma extNTA2J_iff [simp]: "extNTA2J P (C, M, a) = ({this:Class (fst (method P C M))=\Addr a\; snd (the (snd (snd (snd (method P C M)))))}, Map.empty)" by(simp add: extNTA2J_def split_beta) abbreviation extTA2J :: "'addr J_prog \ ('addr, 'thread_id, 'heap) external_thread_action \ ('addr, 'thread_id, 'heap) J_thread_action" where "extTA2J P \ convert_extTA (extNTA2J P)" lemma extTA2J_\: "extTA2J P \ = \" by(simp) text\Locking mechanism: The expression on which the thread is synchronized is evaluated first to a value. If this expression evaluates to null, a null pointer expression is thrown. If this expression evaluates to an address, a lock must be obtained on this address, the sync expression is rewritten to insync. For insync expressions, the body expression may be evaluated. If the body expression is only a value or a thrown exception, the lock is released and the synchronized expression reduces to the body's expression. This is the normal Java semantics, not the one as presented in LNCS 1523, Cenciarelli/Knapp/Reus/Wirsing. There the expression on which the thread synchronized is evaluated except for the last step. If the thread can obtain the lock on the object immediately after the last evaluation step, the evaluation is done and the lock acquired. If the lock cannot be obtained, the evaluation step is discarded. If another thread changes the evaluation result of this last step, the thread then will try to synchronize on the new object.\ context J_heap_base begin inductive red :: "(('addr, 'thread_id, 'heap) external_thread_action \ ('addr, 'thread_id, 'x,'heap) Jinja_thread_action) \ 'addr J_prog \ 'thread_id \ 'addr expr \ ('addr, 'heap) Jstate \ ('addr, 'thread_id, 'x,'heap) Jinja_thread_action \ 'addr expr \ ('addr, 'heap) Jstate \ bool" ("_,_,_ \ ((1\_,/_\) -_\/ (1\_,/_\))" [51,51,0,0,0,0,0,0] 81) and reds :: "(('addr, 'thread_id, 'heap) external_thread_action \ ('addr, 'thread_id, 'x,'heap) Jinja_thread_action) \ 'addr J_prog \ 'thread_id \ 'addr expr list \ ('addr, 'heap) Jstate \ ('addr, 'thread_id, 'x,'heap) Jinja_thread_action \ 'addr expr list \ ('addr, 'heap) Jstate \ bool" ("_,_,_ \ ((1\_,/_\) [-_\]/ (1\_,/_\))" [51,51,0,0,0,0,0,0] 81) for extTA :: "('addr, 'thread_id, 'heap) external_thread_action \ ('addr, 'thread_id, 'x, 'heap) Jinja_thread_action" and P :: "'addr J_prog" and t :: 'thread_id where RedNew: "(h', a) \ allocate h (Class_type C) \ extTA,P,t \ \new C, (h, l)\ -\NewHeapElem a (Class_type C)\\ \addr a, (h', l)\" | RedNewFail: "allocate h (Class_type C) = {} \ extTA,P,t \ \new C, (h, l)\ -\\ \THROW OutOfMemory, (h, l)\" | NewArrayRed: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \newA T\e\, s\ -ta\ \newA T\e'\, s'\" | RedNewArray: "\ 0 <=s i; (h', a) \ allocate h (Array_type T (nat (sint i))) \ \ extTA,P,t \ \newA T\Val (Intg i)\, (h, l)\ -\NewHeapElem a (Array_type T (nat (sint i)))\\ \addr a, (h', l)\" | RedNewArrayNegative: "i extTA,P,t \ \newA T\Val (Intg i)\, s\ -\\ \THROW NegativeArraySize, s\" | RedNewArrayFail: "\ 0 <=s i; allocate h (Array_type T (nat (sint i))) = {} \ \ extTA,P,t \ \newA T\Val (Intg i)\, (h, l)\ -\\ \THROW OutOfMemory, (h, l)\" | CastRed: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \Cast C e, s\ -ta\ \Cast C e', s'\" | RedCast: "\ typeof\<^bsub>hp s\<^esub> v = \U\; P \ U \ T \ \ extTA,P,t \ \Cast T (Val v), s\ -\\ \Val v, s\" | RedCastFail: "\ typeof\<^bsub>hp s\<^esub> v = \U\; \ P \ U \ T \ \ extTA,P,t \ \Cast T (Val v), s\ -\\ \THROW ClassCast, s\" | InstanceOfRed: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \e instanceof T, s\ -ta\ \e' instanceof T, s'\" | RedInstanceOf: "\ typeof\<^bsub>hp s\<^esub> v = \U\; b \ v \ Null \ P \ U \ T \ \ extTA,P,t \ \(Val v) instanceof T, s\ -\\ \Val (Bool b), s\" | BinOpRed1: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \e \bop\ e2, s\ -ta\ \e' \bop\ e2, s'\" | BinOpRed2: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \(Val v) \bop\ e, s\ -ta\ \(Val v) \bop\ e', s'\" | RedBinOp: "binop bop v1 v2 = Some (Inl v) \ extTA,P,t \ \(Val v1) \bop\ (Val v2), s\ -\\ \Val v, s\" | RedBinOpFail: "binop bop v1 v2 = Some (Inr a) \ extTA,P,t \ \(Val v1) \bop\ (Val v2), s\ -\\ \Throw a, s\" | RedVar: "lcl s V = Some v \ extTA,P,t \ \Var V, s\ -\\ \Val v, s\" | LAssRed: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \V:=e, s\ -ta\ \V:=e', s'\" | RedLAss: "extTA,P,t \ \V:=(Val v), (h, l)\ -\\ \unit, (h, l(V \ v))\" | AAccRed1: "extTA,P,t \ \a, s\ -ta\ \a', s'\ \ extTA,P,t \ \a\i\, s\ -ta\ \a'\i\, s'\" | AAccRed2: "extTA,P,t \ \i, s\ -ta\ \i', s'\ \ extTA,P,t \ \(Val a)\i\, s\ -ta\ \(Val a)\i'\, s'\" | RedAAccNull: "extTA,P,t \ \null\Val i\, s\ -\\ \THROW NullPointer, s\" | RedAAccBounds: "\ typeof_addr (hp s) a = \Array_type T n\; i sint i \ int n \ \ extTA,P,t \ \(addr a)\Val (Intg i)\, s\ -\\ \THROW ArrayIndexOutOfBounds, s\" | RedAAcc: "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; heap_read h a (ACell (nat (sint i))) v \ \ extTA,P,t \ \(addr a)\Val (Intg i)\, (h, l)\ -\ReadMem a (ACell (nat (sint i))) v\\ \Val v, (h, l)\" | AAssRed1: "extTA,P,t \ \a, s\ -ta\ \a', s'\ \ extTA,P,t \ \a\i\ := e, s\ -ta\ \a'\i\ := e, s'\" | AAssRed2: "extTA,P,t \ \i, s\ -ta\ \i', s'\ \ extTA,P,t \ \(Val a)\i\ := e, s\ -ta\ \(Val a)\i'\ := e, s'\" | AAssRed3: "extTA,P,t \ \(e::'addr expr), s\ -ta\ \e', s'\ \ extTA,P,t \ \(Val a)\Val i\ := e, s\ -ta\ \(Val a)\Val i\ := e', s'\" | RedAAssNull: "extTA,P,t \ \null\Val i\ := (Val e::'addr expr), s\ -\\ \THROW NullPointer, s\" | RedAAssBounds: "\ typeof_addr (hp s) a = \Array_type T n\; i sint i \ int n \ \ extTA,P,t \ \(addr a)\Val (Intg i)\ := (Val e::'addr expr), s\ -\\ \THROW ArrayIndexOutOfBounds, s\" | RedAAssStore: "\ typeof_addr (hp s) a = \Array_type T n\; 0 <=s i; sint i < int n; typeof\<^bsub>hp s\<^esub> w = \U\; \ (P \ U \ T) \ \ extTA,P,t \ \(addr a)\Val (Intg i)\ := (Val w::'addr expr), s\ -\\ \THROW ArrayStore, s\" | RedAAss: "\ typeof_addr h a = \Array_type T n\; 0 <=s i; sint i < int n; typeof\<^bsub>h\<^esub> w = Some U; P \ U \ T; heap_write h a (ACell (nat (sint i))) w h' \ \ extTA,P,t \ \(addr a)\Val (Intg i)\ := Val w::'addr expr, (h, l)\ -\WriteMem a (ACell (nat (sint i))) w\\ \unit, (h', l)\" | ALengthRed: "extTA,P,t \ \a, s\ -ta\ \a', s'\ \ extTA,P,t \ \a\length, s\ -ta\ \a'\length, s'\" | RedALength: "typeof_addr h a = \Array_type T n\ - \ extTA,P,t \ \addr a\length, (h, l)\ -\\ \Val (Intg (word_of_int (int n))), (h, l)\" + \ extTA,P,t \ \addr a\length, (h, l)\ -\\ \Val (Intg (word_of_nat n)), (h, l)\" | RedALengthNull: "extTA,P,t \ \null\length, s\ -\\ \THROW NullPointer, s\" | FAccRed: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \e\F{D}, s\ -ta\ \e'\F{D}, s'\" | RedFAcc: "heap_read h a (CField D F) v \ extTA,P,t \ \(addr a)\F{D}, (h, l)\ -\ReadMem a (CField D F) v\\ \Val v, (h, l)\" | RedFAccNull: "extTA,P,t \ \null\F{D}, s\ -\\ \THROW NullPointer, s\" | FAssRed1: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \e\F{D}:=e2, s\ -ta\ \e'\F{D}:=e2, s'\" | FAssRed2: "extTA,P,t \ \(e::'addr expr), s\ -ta\ \e', s'\ \ extTA,P,t \ \Val v\F{D}:=e, s\ -ta\ \Val v\F{D}:=e', s'\" | RedFAss: "heap_write h a (CField D F) v h' \ extTA,P,t \ \(addr a)\F{D}:= Val v, (h, l)\ -\WriteMem a (CField D F) v\\ \unit, (h', l)\" | RedFAssNull: "extTA,P,t \ \null\F{D}:=Val v::'addr expr, s\ -\\ \THROW NullPointer, s\" | CASRed1: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \e\compareAndSwap(D\F, e2, e3), s\ -ta\ \e'\compareAndSwap(D\F, e2, e3), s'\" | CASRed2: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \Val v\compareAndSwap(D\F, e, e3), s\ -ta\ \Val v\compareAndSwap(D\F, e', e3), s'\" | CASRed3: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \Val v\compareAndSwap(D\F, Val v', e), s\ -ta\ \Val v\compareAndSwap(D\F, Val v', e'), s'\" | CASNull: "extTA,P,t \ \null\compareAndSwap(D\F, Val v, Val v'), s\ -\\ \THROW NullPointer, s\" | RedCASSucceed: "\ heap_read h a (CField D F) v; heap_write h a (CField D F) v' h' \ \ extTA,P,t \ \addr a\compareAndSwap(D\F, Val v, Val v'), (h, l)\ -\ReadMem a (CField D F) v, WriteMem a (CField D F) v'\\ \true, (h', l)\" | RedCASFail: "\ heap_read h a (CField D F) v''; v \ v'' \ \ extTA,P,t \ \addr a\compareAndSwap(D\F, Val v, Val v'), (h, l)\ -\ReadMem a (CField D F) v''\\ \false, (h, l)\" | CallObj: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \e\M(es), s\ -ta\ \e'\M(es), s'\" | CallParams: "extTA,P,t \ \es, s\ [-ta\] \es',s'\ \ extTA,P,t \ \(Val v)\M(es),s\ -ta\ \(Val v)\M(es'),s'\" | RedCall: "\ typeof_addr (hp s) a = \hU\; P \ class_type_of hU sees M:Ts\T = \(pns,body)\ in D; size vs = size pns; size Ts = size pns \ \ extTA,P,t \ \(addr a)\M(map Val vs), s\ -\\ \blocks (this # pns) (Class D # Ts) (Addr a # vs) body, s\" | RedCallExternal: "\ typeof_addr (hp s) a = \hU\; P \ class_type_of hU sees M:Ts\T = Native in D; P,t \ \a\M(vs), hp s\ -ta\ext \va, h'\; ta' = extTA ta; e' = extRet2J ((addr a)\M(map Val vs)) va; s' = (h', lcl s) \ \ extTA,P,t \ \(addr a)\M(map Val vs), s\ -ta'\ \e', s'\" | RedCallNull: "extTA,P,t \ \null\M(map Val vs), s\ -\\ \THROW NullPointer, s\" | BlockRed: "extTA,P,t \ \e, (h, l(V:=vo))\ -ta\ \e', (h', l')\ \ extTA,P,t \ \{V:T=vo; e}, (h, l)\ -ta\ \{V:T=l' V; e'}, (h', l'(V := l V))\" | RedBlock: "extTA,P,t \ \{V:T=vo; Val u}, s\ -\\ \Val u, s\" | SynchronizedRed1: "extTA,P,t \ \o', s\ -ta\ \o'', s'\ \ extTA,P,t \ \sync(o') e, s\ -ta\ \sync(o'') e, s'\" | SynchronizedNull: "extTA,P,t \ \sync(null) e, s\ -\\ \THROW NullPointer, s\" | LockSynchronized: "extTA,P,t \ \sync(addr a) e, s\ -\Lock\a, SyncLock a\\ \insync(a) e, s\" | SynchronizedRed2: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \insync(a) e, s\ -ta\ \insync(a) e', s'\" | UnlockSynchronized: "extTA,P,t \ \insync(a) (Val v), s\ -\Unlock\a, SyncUnlock a\\ \Val v, s\" | SeqRed: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \e;;e2, s\ -ta\ \e';;e2, s'\" | RedSeq: "extTA,P,t \ \(Val v);;e, s\ -\\ \e, s\" | CondRed: "extTA,P,t \ \b, s\ -ta\ \b', s'\ \ extTA,P,t \ \if (b) e1 else e2, s\ -ta\ \if (b') e1 else e2, s'\" | RedCondT: "extTA,P,t \ \if (true) e1 else e2, s\ -\\ \e1, s\" | RedCondF: "extTA,P,t \ \if (false) e1 else e2, s\ -\\ \e2, s\" | RedWhile: "extTA,P,t \ \while(b) c, s\ -\\ \if (b) (c;;while(b) c) else unit, s\" | ThrowRed: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \throw e, s\ -ta\ \throw e', s'\" | RedThrowNull: "extTA,P,t \ \throw null, s\ -\\ \THROW NullPointer, s\" | TryRed: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \try e catch(C V) e2, s\ -ta\ \try e' catch(C V) e2, s'\" | RedTry: "extTA,P,t \ \try (Val v) catch(C V) e2, s\ -\\ \Val v, s\" | RedTryCatch: "\ typeof_addr (hp s) a = \Class_type D\; P \ D \\<^sup>* C \ \ extTA,P,t \ \try (Throw a) catch(C V) e2, s\ -\\ \{V:Class C=\Addr a\; e2}, s\" | RedTryFail: "\ typeof_addr (hp s) a = \Class_type D\; \ P \ D \\<^sup>* C \ \ extTA,P,t \ \try (Throw a) catch(C V) e2, s\ -\\ \Throw a, s\" | ListRed1: "extTA,P,t \ \e,s\ -ta\ \e',s'\ \ extTA,P,t \ \e#es,s\ [-ta\] \e'#es,s'\" | ListRed2: "extTA,P,t \ \es,s\ [-ta\] \es',s'\ \ extTA,P,t \ \Val v # es,s\ [-ta\] \Val v # es',s'\" \ \Exception propagation\ | NewArrayThrow: "extTA,P,t \ \newA T\Throw a\, s\ -\\ \Throw a, s\" | CastThrow: "extTA,P,t \ \Cast C (Throw a), s\ -\\ \Throw a, s\" | InstanceOfThrow: "extTA,P,t \ \(Throw a) instanceof T, s\ -\\ \Throw a, s\" | BinOpThrow1: "extTA,P,t \ \(Throw a) \bop\ e\<^sub>2, s\ -\\ \Throw a, s\" | BinOpThrow2: "extTA,P,t \ \(Val v\<^sub>1) \bop\ (Throw a), s\ -\\ \Throw a, s\" | LAssThrow: "extTA,P,t \ \V:=(Throw a), s\ -\\ \Throw a, s\" | AAccThrow1: "extTA,P,t \ \(Throw a)\i\, s\ -\\ \Throw a, s\" | AAccThrow2: "extTA,P,t \ \(Val v)\Throw a\, s\ -\\ \Throw a, s\" | AAssThrow1: "extTA,P,t \ \(Throw a)\i\ := e, s\ -\\ \Throw a, s\" | AAssThrow2: "extTA,P,t \ \(Val v)\Throw a\ := e, s\ -\\ \Throw a, s\" | AAssThrow3: "extTA,P,t \ \(Val v)\Val i\ := Throw a :: 'addr expr, s\ -\\ \Throw a, s\" | ALengthThrow: "extTA,P,t \ \(Throw a)\length, s\ -\\ \Throw a, s\" | FAccThrow: "extTA,P,t \ \(Throw a)\F{D}, s\ -\\ \Throw a, s\" | FAssThrow1: "extTA,P,t \ \(Throw a)\F{D}:=e\<^sub>2, s\ -\\ \Throw a, s\" | FAssThrow2: "extTA,P,t \ \Val v\F{D}:=(Throw a::'addr expr), s\ -\\ \Throw a, s\" | CASThrow: "extTA,P,t \ \Throw a\compareAndSwap(D\F, e2, e3), s\ -\\ \Throw a, s\" | CASThrow2: "extTA,P,t \ \Val v\compareAndSwap(D\F, Throw a, e3), s\ -\\ \Throw a, s\" | CASThrow3: "extTA,P,t \ \Val v\compareAndSwap(D\F, Val v', Throw a), s\ -\\ \Throw a, s\" | CallThrowObj: "extTA,P,t \ \(Throw a)\M(es), s\ -\\ \Throw a, s\" | CallThrowParams: "\ es = map Val vs @ Throw a # es' \ \ extTA,P,t \ \(Val v)\M(es), s\ -\\ \Throw a, s\" | BlockThrow: "extTA,P,t \ \{V:T=vo; Throw a}, s\ -\\ \Throw a, s\" | SynchronizedThrow1: "extTA,P,t \ \sync(Throw a) e, s\ -\\ \Throw a, s\" | SynchronizedThrow2: "extTA,P,t \ \insync(a) Throw ad, s\ -\Unlock\a, SyncUnlock a\\ \Throw ad, s\" | SeqThrow: "extTA,P,t \ \(Throw a);;e\<^sub>2, s\ -\\ \Throw a, s\" | CondThrow: "extTA,P,t \ \if (Throw a) e\<^sub>1 else e\<^sub>2, s\ -\\ \Throw a, s\" | ThrowThrow: "extTA,P,t \ \throw(Throw a), s\ -\\ \Throw a, s\" inductive_cases red_cases: "extTA,P,t \ \new C, s\ -ta\ \e', s'\" "extTA,P,t \ \newA T\e\, s\ -ta\ \e', s'\" "extTA,P,t \ \Cast T e, s\ -ta\ \e', s'\" "extTA,P,t \ \e instanceof T, s\ -ta\ \e', s'\" "extTA,P,t \ \e \bop\ e', s\ -ta\ \e'', s'\" "extTA,P,t \ \Var V, s\ -ta\ \e', s'\" "extTA,P,t \ \V:=e, s\ -ta\ \e', s'\" "extTA,P,t \ \a\i\, s\ -ta\ \e', s'\" "extTA,P,t \ \a\i\ := e, s\ -ta\ \e', s'\" "extTA,P,t \ \a\length, s\ -ta\ \e', s'\" "extTA,P,t \ \e\F{D}, s\ -ta\ \e', s'\" "extTA,P,t \ \e\F{D} := e', s\ -ta\ \e'', s'\" "extTA,P,t \ \e\compareAndSwap(D\F, e', e''), s\ -ta\ \e''', s'\" "extTA,P,t \ \e\M(es), s\ -ta\ \e', s'\" "extTA,P,t \ \{V:T=vo; e}, s\ -ta\ \e', s'\" "extTA,P,t \ \sync(o') e, s\ -ta\ \e', s'\" "extTA,P,t \ \insync(a) e, s\ -ta\ \e', s'\" "extTA,P,t \ \e;;e', s\ -ta\ \e'', s'\" "extTA,P,t \ \if (b) e1 else e2, s \ -ta\ \e', s'\" "extTA,P,t \ \while (b) e, s \ -ta\ \e', s'\" "extTA,P,t \ \throw e, s \ -ta\ \e', s'\" "extTA,P,t \ \try e catch(C V) e', s\ -ta\ \e'', s'\" inductive_cases reds_cases: "extTA,P,t \ \e # es, s\ [-ta\] \es', s'\" abbreviation red' :: "'addr J_prog \ 'thread_id \ 'addr expr \ ('heap \ 'addr locals) \ ('addr, 'thread_id, 'heap) J_thread_action \ 'addr expr \ ('heap \ 'addr locals) \ bool" ("_,_ \ ((1\_,/_\) -_\/ (1\_,/_\))" [51,0,0,0,0,0,0] 81) where "red' P \ red (extTA2J P) P" abbreviation reds' :: "'addr J_prog \ 'thread_id \ 'addr expr list \ ('heap \ 'addr locals) \ ('addr, 'thread_id, 'heap) J_thread_action \ 'addr expr list \ ('heap \ 'addr locals) \ bool" ("_,_ \ ((1\_,/_\) [-_\]/ (1\_,/_\))" [51,0,0,0,0,0,0] 81) where "reds' P \ reds (extTA2J P) P" subsection\Some easy lemmas\ lemma [iff]: "\ extTA,P,t \ \Val v, s\ -ta\ \e', s'\" by(fastforce elim:red.cases) lemma red_no_val [dest]: "\ extTA,P,t \ \e, s\ -tas\ \e', s'\; is_val e \ \ False" by(auto) lemma [iff]: "\ extTA,P,t \ \Throw a, s\ -ta\ \e', s'\" by(fastforce elim: red_cases) lemma reds_map_Val_Throw: "extTA,P,t \ \map Val vs @ Throw a # es, s\ [-ta\] \es', s'\ \ False" by(induct vs arbitrary: es')(auto elim!: reds_cases) lemma reds_preserves_len: "extTA,P,t \ \es, s\ [-ta\] \es', s'\ \ length es' = length es" by(induct es arbitrary: es')(auto elim: reds.cases) lemma red_lcl_incr: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ dom (lcl s) \ dom (lcl s')" and reds_lcl_incr: "extTA,P,t \ \es, s\ [-ta\] \es', s'\ \ dom (lcl s) \ dom (lcl s')" apply(induct rule:red_reds.inducts) apply(auto simp del: fun_upd_apply split: if_split_asm) done lemma red_lcl_add_aux: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ extTA,P,t \ \e, (hp s, l0 ++ lcl s)\ -ta\ \e', (hp s', l0 ++ lcl s')\" and reds_lcl_add_aux: "extTA,P,t \ \es, s\ [-ta\] \es', s'\ \ extTA,P,t \ \es, (hp s, l0 ++ lcl s)\ [-ta\] \es', (hp s', l0 ++ lcl s')\" proof (induct arbitrary: l0 and l0 rule:red_reds.inducts) case (BlockRed e h x V vo ta e' h' x' T) note IH = \\l0. extTA,P,t \ \e,(hp (h, x(V := vo)), l0 ++ lcl (h, x(V := vo)))\ -ta\ \e',(hp (h', x'), l0 ++ lcl (h', x'))\\[simplified] have lrew: "\x x'. x(V := vo) ++ x'(V := vo) = (x ++ x')(V := vo)" by(simp add:fun_eq_iff map_add_def) have lrew1: "\X X' X'' vo. (X(V := vo) ++ X')(V := (X ++ X'') V) = X ++ X'(V := X'' V)" by(simp add: fun_eq_iff map_add_def) have lrew2: "\X X'. (X(V := None) ++ X') V = X' V" by(simp add: map_add_def) show ?case proof(cases vo) case None from IH[of "l0(V := vo)"] show ?thesis apply(simp del: fun_upd_apply add: lrew) apply(drule red_reds.BlockRed) by(simp only: lrew1 None lrew2) next case (Some v) with \extTA,P,t \ \e,(h, x(V := vo))\ -ta\ \e',(h', x')\\ have "x' V \ None" by -(drule red_lcl_incr, auto split: if_split_asm) with IH[of "l0(V := vo)"] show ?thesis apply(clarsimp simp del: fun_upd_apply simp add: lrew) apply(drule red_reds.BlockRed) by(simp add: lrew1 Some del: fun_upd_apply) qed next case RedTryFail thus ?case by(auto intro: red_reds.RedTryFail) qed(fastforce intro:red_reds.intros simp del: fun_upd_apply)+ lemma red_lcl_add: "extTA,P,t \ \e, (h, l)\ -ta\ \e', (h', l')\ \ extTA,P,t \ \e, (h, l0 ++ l)\ -ta\ \e', (h', l0 ++ l')\" and reds_lcl_add: "extTA,P,t \ \es, (h, l)\ [-ta\] \es', (h', l')\ \ extTA,P,t \ \es, (h, l0 ++ l)\ [-ta\] \es', (h', l0 ++ l')\" by(auto dest:red_lcl_add_aux reds_lcl_add_aux) lemma reds_no_val [dest]: "\ extTA,P,t \ \es, s\ [-ta\] \es', s'\; is_vals es \ \ False" apply(induct es arbitrary: s ta es' s') apply(blast elim: reds.cases) apply(erule reds.cases) apply(auto, blast) done lemma red_no_Throw [dest!]: "extTA,P,t \ \Throw a, s\ -ta\ \e', s'\ \ False" by(auto elim!: red_cases) lemma red_lcl_sub: "\ extTA,P,t \ \e, s\ -ta\ \e', s'\; fv e \ W \ \ extTA,P,t \ \e, (hp s, (lcl s)|`W)\ -ta\ \e', (hp s', (lcl s')|`W)\" and reds_lcl_sub: "\ extTA,P,t \ \es, s\ [-ta\] \es', s'\; fvs es \ W \ \ extTA,P,t \ \es, (hp s, (lcl s)|`W)\ [-ta\] \es', (hp s', (lcl s')|`W)\" proof(induct arbitrary: W and W rule: red_reds.inducts) case (RedLAss V v h l W) have "extTA,P,t \ \V:=Val v,(h, l |` W)\ -\\ \unit,(h, (l |`W)(V \ v))\" by(rule red_reds.RedLAss) with RedLAss show ?case by(simp del: fun_upd_apply) next case (BlockRed e h x V vo ta e' h' x' T) have IH: "\W. fv e \ W \ extTA,P,t \ \e,(hp (h, x(V := vo)), lcl (h, x(V := vo)) |` W)\ -ta\ \e',(hp (h', x'), lcl (h', x') |` W)\" by fact from \fv {V:T=vo; e} \ W\ have fve: "fv e \ insert V W" by auto show ?case proof(cases "V \ W") case True with fve have "fv e \ W" by auto from True IH[OF this] have "extTA,P,t \ \e,(h, (x |` W )(V := vo))\ -ta\ \e',(h', x' |` W)\" by(simp) with True have "extTA,P,t \ \{V:T=vo; e},(h, x |` W)\ -ta\ \{V:T=x' V; e'},(h', (x' |` W)(V := x V))\" by -(drule red_reds.BlockRed[where T=T], simp) with True show ?thesis by(simp del: fun_upd_apply) next case False with IH[OF fve] have "extTA,P,t \ \e,(h, (x |` W)(V := vo))\ -ta\ \e',(h', x' |` insert V W)\" by(simp) with False have "extTA,P,t \ \{V:T=vo; e},(h, x |` W)\ -ta\ \{V:T=x' V; e'},(h', (x' |` W))\" by -(drule red_reds.BlockRed[where T=T],simp) with False show ?thesis by(simp del: fun_upd_apply) qed next case RedTryFail thus ?case by(auto intro: red_reds.RedTryFail) qed(fastforce intro: red_reds.intros)+ lemma red_notfree_unchanged: "\ extTA,P,t \ \e, s\ -ta\ \e', s'\; V \ fv e \ \ lcl s' V = lcl s V" and reds_notfree_unchanged: "\ extTA,P,t \ \es, s\ [-ta\] \es', s'\; V \ fvs es \ \ lcl s' V = lcl s V" apply(induct rule: red_reds.inducts) apply(fastforce)+ done lemma red_dom_lcl: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ dom (lcl s') \ dom (lcl s) \ fv e" and reds_dom_lcl: "extTA,P,t \ \es, s\ [-ta\] \es', s'\ \ dom (lcl s') \ dom (lcl s) \ fvs es" proof (induct rule:red_reds.inducts) case (BlockRed e h x V vo ta e' h' x' T) thus ?case by(clarsimp)(fastforce split:if_split_asm) qed auto lemma red_Suspend_is_call: "\ convert_extTA extNTA,P,t \ \e, s\ -ta\ \e', s'\; Suspend w \ set \ta\\<^bsub>w\<^esub> \ \ \a vs hT Ts Tr D. call e' = \(a, wait, vs)\ \ typeof_addr (hp s) a = \hT\ \ P \ class_type_of hT sees wait:Ts\Tr = Native in D" and reds_Suspend_is_calls: "\ convert_extTA extNTA,P,t \ \es, s\ [-ta\] \es', s'\; Suspend w \ set \ta\\<^bsub>w\<^esub> \ \ \a vs hT Ts Tr D. calls es' = \(a, wait, vs)\ \ typeof_addr (hp s) a = \hT\ \ P \ class_type_of hT sees wait:Ts\Tr = Native in D" proof(induct rule: red_reds.inducts) case RedCallExternal thus ?case apply clarsimp apply(frule red_external_Suspend_StaySame, simp) apply(drule red_external_Suspend_waitD, fastforce+) done qed auto end context J_heap begin lemma red_hext_incr: "extTA,P,t \ \e, s\ -ta\ \e', s'\ \ hp s \ hp s'" and reds_hext_incr: "extTA,P,t \ \es, s\ [-ta\] \es', s'\ \ hp s \ hp s'" by(induct rule:red_reds.inducts)(auto intro: hext_heap_ops red_external_hext) lemma red_preserves_tconf: "\ extTA,P,t \ \e, s\ -ta\ \e', s'\; P,hp s \ t \t \ \ P,hp s' \ t \t" by(drule red_hext_incr)(rule tconf_hext_mono) lemma reds_preserves_tconf: "\ extTA,P,t \ \es, s\ [-ta\] \es', s'\; P,hp s \ t \t \ \ P,hp s' \ t \t" by(drule reds_hext_incr)(rule tconf_hext_mono) end subsection \Code generation\ context J_heap_base begin lemma RedCall_code: "\ is_vals es; typeof_addr (hp s) a = \hU\; P \ class_type_of hU sees M:Ts\T = \(pns,body)\ in D; size es = size pns; size Ts = size pns \ \ extTA,P,t \ \(addr a)\M(es), s\ -\\ \blocks (this # pns) (Class D # Ts) (Addr a # map the_Val es) body, s\" and RedCallExternal_code: "\ is_vals es; typeof_addr (hp s) a = \hU\; P \ class_type_of hU sees M:Ts\T = Native in D; P,t \ \a\M(map the_Val es), hp s\ -ta\ext \va, h'\ \ \ extTA,P,t \ \(addr a)\M(es), s\ -extTA ta\ \extRet2J ((addr a)\M(es)) va, (h', lcl s)\" and RedCallNull_code: "is_vals es \ extTA,P,t \ \null\M(es), s\ -\\ \THROW NullPointer, s\" and CallThrowParams_code: "is_Throws es \ extTA,P,t \ \(Val v)\M(es), s\ -\\ \hd (dropWhile is_val es), s\" apply(auto simp add: is_vals_conv is_Throws_conv o_def intro: RedCall RedCallExternal RedCallNull simp del: blocks.simps) apply(subst dropWhile_append2) apply(auto intro: CallThrowParams) done end lemmas [code_pred_intro] = J_heap_base.RedNew[folded Predicate_Compile.contains_def] J_heap_base.RedNewFail J_heap_base.NewArrayRed J_heap_base.RedNewArray[folded Predicate_Compile.contains_def] J_heap_base.RedNewArrayNegative J_heap_base.RedNewArrayFail J_heap_base.CastRed J_heap_base.RedCast J_heap_base.RedCastFail J_heap_base.InstanceOfRed J_heap_base.RedInstanceOf J_heap_base.BinOpRed1 J_heap_base.BinOpRed2 J_heap_base.RedBinOp J_heap_base.RedBinOpFail J_heap_base.RedVar J_heap_base.LAssRed J_heap_base.RedLAss J_heap_base.AAccRed1 J_heap_base.AAccRed2 J_heap_base.RedAAccNull J_heap_base.RedAAccBounds J_heap_base.RedAAcc J_heap_base.AAssRed1 J_heap_base.AAssRed2 J_heap_base.AAssRed3 J_heap_base.RedAAssNull J_heap_base.RedAAssBounds J_heap_base.RedAAssStore J_heap_base.RedAAss J_heap_base.ALengthRed J_heap_base.RedALength J_heap_base.RedALengthNull J_heap_base.FAccRed J_heap_base.RedFAcc J_heap_base.RedFAccNull J_heap_base.FAssRed1 J_heap_base.FAssRed2 J_heap_base.RedFAss J_heap_base.RedFAssNull J_heap_base.CASRed1 J_heap_base.CASRed2 J_heap_base.CASRed3 J_heap_base.CASNull J_heap_base.RedCASSucceed J_heap_base.RedCASFail J_heap_base.CallObj J_heap_base.CallParams declare J_heap_base.RedCall_code[code_pred_intro RedCall_code] J_heap_base.RedCallExternal_code[code_pred_intro RedCallExternal_code] J_heap_base.RedCallNull_code[code_pred_intro RedCallNull_code] lemmas [code_pred_intro] = J_heap_base.BlockRed J_heap_base.RedBlock J_heap_base.SynchronizedRed1 J_heap_base.SynchronizedNull J_heap_base.LockSynchronized J_heap_base.SynchronizedRed2 J_heap_base.UnlockSynchronized J_heap_base.SeqRed J_heap_base.RedSeq J_heap_base.CondRed J_heap_base.RedCondT J_heap_base.RedCondF J_heap_base.RedWhile J_heap_base.ThrowRed declare J_heap_base.RedThrowNull[code_pred_intro RedThrowNull'] lemmas [code_pred_intro] = J_heap_base.TryRed J_heap_base.RedTry J_heap_base.RedTryCatch J_heap_base.RedTryFail J_heap_base.ListRed1 J_heap_base.ListRed2 J_heap_base.NewArrayThrow J_heap_base.CastThrow J_heap_base.InstanceOfThrow J_heap_base.BinOpThrow1 J_heap_base.BinOpThrow2 J_heap_base.LAssThrow J_heap_base.AAccThrow1 J_heap_base.AAccThrow2 J_heap_base.AAssThrow1 J_heap_base.AAssThrow2 J_heap_base.AAssThrow3 J_heap_base.ALengthThrow J_heap_base.FAccThrow J_heap_base.FAssThrow1 J_heap_base.FAssThrow2 J_heap_base.CASThrow J_heap_base.CASThrow2 J_heap_base.CASThrow3 J_heap_base.CallThrowObj declare J_heap_base.CallThrowParams_code[code_pred_intro CallThrowParams_code] lemmas [code_pred_intro] = J_heap_base.BlockThrow J_heap_base.SynchronizedThrow1 J_heap_base.SynchronizedThrow2 J_heap_base.SeqThrow J_heap_base.CondThrow declare J_heap_base.ThrowThrow[code_pred_intro ThrowThrow'] code_pred (modes: J_heap_base.red: i \ i \ i \ i \ i \ i \ (i \ i \ i \ o \ bool) \ (i \ i \ i \ i \ o \ bool) \ i \ i \ i \ i \ i \ o \ o \ o \ bool and J_heap_base.reds: i \ i \ i \ i \ i \ i \ (i \ i \ i \ o \ bool) \ (i \ i \ i \ i \ o \ bool) \ i \ i \ i \ i \ i \ o \ o \ o \ bool) [detect_switches, skip_proof] \ \proofs are possible, but take veeerry long\ J_heap_base.red proof - case red from red.prems show thesis proof(cases rule: J_heap_base.red.cases[consumes 1, case_names RedNew RedNewFail NewArrayRed RedNewArray RedNewArrayNegative RedNewArrayFail CastRed RedCast RedCastFail InstanceOfRed RedInstanceOf BinOpRed1 BinOpRed2 RedBinOp RedBinOpFail RedVar LAssRed RedLAss AAccRed1 AAccRed2 RedAAccNull RedAAccBounds RedAAcc AAssRed1 AAssRed2 AAssRed3 RedAAssNull RedAAssBounds RedAAssStore RedAAss ALengthRed RedALength RedALengthNull FAccRed RedFAcc RedFAccNull FAssRed1 FAssRed2 RedFAss RedFAssNull CASRed1 CASRed2 CASRed3 RedCASNull RedCASSucceed RedCASFail CallObj CallParams RedCall RedCallExternal RedCallNull BlockRed RedBlock SynchronizedRed1 SynchronizedNull LockSynchronized SynchronizedRed2 UnlockSynchronized SeqRed RedSeq CondRed RedCondT RedCondF RedWhile ThrowRed RedThrowNull TryRed RedTry RedTryCatch RedTryFail NewArrayThrow CastThrow InstanceOfThrow BinOpThrow1 BinOpThrow2 LAssThrow AAccThrow1 AAccThrow2 AAssThrow1 AAssThrow2 AAssThrow3 ALengthThrow FAccThrow FAssThrow1 FAssThrow2 CASThrow CASThrow2 CASThrow3 CallThrowObj CallThrowParams BlockThrow SynchronizedThrow1 SynchronizedThrow2 SeqThrow CondThrow ThrowThrow]) case (RedCall s a U M Ts T pns body D vs) with red.RedCall_code[OF refl refl refl refl refl refl refl refl refl refl refl, of a M "map Val vs" s pns D Ts body U T] show ?thesis by(simp add: o_def) next case (RedCallExternal s a U M Ts T D vs ta va h' ta' e' s') with red.RedCallExternal_code[OF refl refl refl refl refl refl refl refl refl refl refl, of a M "map Val vs" s ta va h' U Ts T D] show ?thesis by(simp add: o_def) next case (RedCallNull M vs s) with red.RedCallNull_code[OF refl refl refl refl refl refl refl refl refl refl refl, of M "map Val vs" s] show ?thesis by(simp add: o_def) next case (CallThrowParams es vs a es' v M s) with red.CallThrowParams_code[OF refl refl refl refl refl refl refl refl refl refl refl, of v M "map Val vs @ Throw a # es'" s] show ?thesis apply(auto simp add: is_Throws_conv) apply(erule meta_impE) apply(subst dropWhile_append2) apply auto done next case RedThrowNull thus ?thesis by-(erule (4) red.RedThrowNull'[OF refl refl refl refl refl refl refl refl refl refl refl]) next case ThrowThrow thus ?thesis by-(erule (4) red.ThrowThrow'[OF refl refl refl refl refl refl refl refl refl refl refl]) qed(assumption|erule (4) red.that[unfolded Predicate_Compile.contains_def, OF refl refl refl refl refl refl refl refl refl refl refl])+ next case reds from reds.prems show thesis by(rule J_heap_base.reds.cases)(assumption|erule (4) reds.that[OF refl refl refl refl refl refl refl refl refl refl refl])+ qed end diff --git a/thys/Native_Word/Code_Target_Word_Base.thy b/thys/Native_Word/Code_Target_Word_Base.thy --- a/thys/Native_Word/Code_Target_Word_Base.thy +++ b/thys/Native_Word/Code_Target_Word_Base.thy @@ -1,445 +1,449 @@ (* 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-Word.Word" Bits_Integer begin text \More lemmas\ lemma nat_div_eq_Suc_0_iff: "n div m = Suc 0 \ m \ n \ n < 2 * m" apply auto using div_greater_zero_iff apply fastforce apply (metis One_nat_def div_greater_zero_iff dividend_less_div_times mult.right_neutral mult_Suc mult_numeral_1 numeral_2_eq_2 zero_less_numeral) apply (simp add: div_nat_eqI) done 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) 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 unat_p2: "n < LENGTH('a :: len) \ unat (2 ^ n :: 'a word) = 2 ^ n" proof(induct n) case 0 thus ?case by simp next case (Suc n) then obtain n' where "LENGTH('a) = Suc n'" by(cases "LENGTH('a)") simp_all with Suc show ?case by (simp add: unat_word_ariths bintrunc_mod2p) qed lemma word_div_lt_eq_0: "x < y \ x div y = 0" for x :: "'a :: len word" by (simp add: word_eq_iff word_less_def word_test_bit_def uint_div) 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 nat_div_eq_Suc_0_iff[symmetric]) apply(rule word_unat.Abs_inject) apply(simp only: unat_div[symmetric] word_unat.Rep) apply(simp add: unats_def Suc_0_lt_2p_len_of) done lemma div_half_word: fixes x y :: "'a :: len word" assumes "y \ 0" shows "(x div y, x mod y) = (let q = (x >> 1) div y << 1; 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 = "(x >> 1) div y << 1" 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 - apply (auto simp add: shiftr_word_eq drop_bit_eq_div shiftl_t2n word_arith_nat_div uno_simps) - apply (metis \unat (of_nat n) = n\ uno_simps(2)) - done + by (auto simp add: shiftr_word_eq drop_bit_eq_div shiftl_t2n 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 (metis div_mult2_eq dtle mult.assoc mult.left_commute) 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 - by (simp_all add: word_sub_wi word_mult_def uint_nat unat_of_nat of_nat_mult [symmetric] word_of_nat[symmetric] of_nat_diff word_le_nat_alt del: of_nat_mult) - (metis diff_diff_left less_imp_diff_less of_nat_diff of_nat_inverse word_of_nat) - thus ?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 uint_nat unat_of_nat zmod_int[symmetric] zdiv_int[symmetric] word_of_nat[symmetric])(simp add: Let_def split del: if_split split: if_split_asm) + apply (simp_all add: of_nat_diff) + apply (subst of_nat_diff) + apply (simp_all add: word_le_nat_alt take_bit_nat_eq_self unat_sub_if' unat_word_ariths) + apply (cases \2 \ LENGTH('a)\) + apply (simp_all add: unat_word_ariths take_bit_nat_eq_self) + 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: "(BITS n. f n :: 'a :: len word) !! n \ n < LENGTH('a) \ f n" by (simp add: test_bit_eq_bit bit_set_bits_word_iff) lemma word_of_int_conv_set_bits: "word_of_int i = (BITS n. i !! n)" by (rule word_eqI) (auto simp add: word_test_bit_set_bits test_bit.eq_norm) lemma word_and_mask_or_conv_and_mask: "n !! index \ (n AND mask index) OR (1 << index) = n AND mask (index + 1)" for n :: \'a::len word\ by(rule word_eqI)(auto simp add: word_ao_nth word_size nth_shiftl simp del: shiftl_1) lemma uint_and_mask_or_full: fixes n :: "'a :: len word" assumes "n !! (LENGTH('a) - 1)" and "mask1 = mask (LENGTH('a) - 1)" and "mask2 = 1 << LENGTH('a) - 1" shows "uint (n AND mask1) OR mask2 = uint n" proof - have "mask2 = uint (1 << LENGTH('a) - 1 :: 'a word)" using assms by (simp add: uint_shiftl word_size bintrunc_shiftl del: shiftl_1) hence "uint (n AND mask1) OR mask2 = uint (n AND mask1 OR (1 << LENGTH('a) - 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 text \Division on @{typ "'a word"} is unsigned, but Scala and OCaml only have signed division and modulus.\ 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))" lemma sdiv_smod_id: "(a sdiv b) * b + (a smod b) = a" proof - note [simp] = word_sdiv_def word_smod_def have F5: "\u::'a word. - (- u) = u" by (metis minus_minus) have F7: "\v u::'a word. u + v = v + u" by(metis add.left_commute add_0_right) 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 (metis word_sint.Rep_inverse wi_hom_syms(1) wi_hom_syms(3)) have "\u. u = - sint b \ word_of_int (sint a mod u + - (- u * (sint a div u))) = a" using F5 by (metis minus_minus word_sint.Rep_inverse' mult_minus_left add.commute mult_div_mod_eq [symmetric]) 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(metis mult.commute mult_minus_left) hence eq: "word_of_int (- (sint a div - sint b)) * b + word_of_int (sint a mod - sint b) = a" using F7 by metis 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 (metis wi_hom_add wi_hom_mult add.commute mult.commute word_sint.Rep_inverse add.commute mult_div_mod_eq [symmetric]) qed qed 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 1 << (LENGTH('a) - 1) \ y then if x < y then (0, x) else (1, x - y) else let q = ((x >> 1) sdiv y) << 1; r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" proof(cases "1 << (LENGTH('a) - 1) \ y") case True note y = this show ?thesis proof(cases "x < y") case True then have "x mod y = x" by (cases x, cases y) (simp add: word_less_def word_mod_def) 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(simp add: word_le_nat_alt unat_of_nat unat_p2) + 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 not_less word_le_nat_alt unat_of_nat) + 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)") (simp_all, simp only: of_nat_numeral [where ?'a = int, symmetric] zdiv_int [symmetric] of_nat_power [symmetric]) with y n have "sint (x >> 1) = uint (x >> 1)" - by (simp add: sint_uint sbintrunc_mod2p shiftr_div_2n) (simp add: uint_nat unat_of_nat) + by (simp add: sint_uint sbintrunc_mod2p shiftr_div_2n take_bit_nat_eq_self) moreover have "uint y + 2 ^ (LENGTH('a) - Suc 0) < 2 ^ LENGTH('a)" using y by (cases "LENGTH('a)") (simp_all add: not_le word_2p_lem word_size) then have "sint y = uint y" by (simp add: sint_uint sbintrunc_mod2p) ultimately show ?thesis using y by(subst div_half_word[OF assms])(simp add: word_sdiv_def uint_div[symmetric]) qed text \More implementations tailored towards target-language implementations\ context includes integer.lifting begin lift_definition word_of_integer :: "integer \ 'a :: len word" is word_of_int . lemma word_of_integer_code [code]: "word_of_integer n = word_of_int (int_of_integer n)" by(simp add: word_of_integer.rep_eq) end -lemma word_of_int_code [code abstract]: +lemma word_of_int_code: "uint (word_of_int x :: 'a word) = x AND mask (LENGTH('a :: len))" -by(simp add: uint_word_of_int and_bin_mask_conv_mod) + by (simp add: take_bit_eq_mask) context fixes f :: "nat \ bool" begin definition set_bits_aux :: \'a word \ nat \ 'a :: len word\ where \set_bits_aux w n = push_bit n w OR take_bit n (set_bits f)\ lemma set_bits_aux_conv: \set_bits_aux w n = (w << n) OR (set_bits f AND mask n)\ for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: set_bits_aux_def shiftl_word_eq bit_and_iff bit_or_iff bit_push_bit_iff bit_take_bit_iff bit_mask_iff bit_set_bits_word_iff exp_eq_zero_iff) corollary set_bits_conv_set_bits_aux: \set_bits f = (set_bits_aux 0 (LENGTH('a)) :: 'a :: len word)\ by (simp add: set_bits_aux_conv) lemma set_bits_aux_0 [simp]: \set_bits_aux w 0 = w\ by (simp add: set_bits_aux_conv) lemma set_bits_aux_Suc [simp]: \set_bits_aux w (Suc n) = set_bits_aux ((w << 1) OR (if f n then 1 else 0)) n\ by (simp add: set_bits_aux_def shiftl_word_eq bit_eq_iff bit_or_iff bit_push_bit_iff bit_take_bit_iff bit_set_bits_word_iff) (auto simp add: bit_exp_iff not_less bit_1_iff less_Suc_eq_le exp_eq_zero_iff) lemma set_bits_aux_simps [code]: \set_bits_aux w 0 = w\ \set_bits_aux w (Suc n) = set_bits_aux ((w << 1) OR (if f n then 1 else 0)) n\ by simp_all end lemma word_of_int_via_signed: fixes mask assumes mask_def: "mask = Bit_Operations.mask (LENGTH('a))" and shift_def: "shift = 1 << LENGTH('a)" and index_def: "index = LENGTH('a) - 1" and overflow_def:"overflow = 1 << (LENGTH('a) - 1)" and least_def: "least = - overflow" shows "(word_of_int i :: 'a :: len word) = (let i' = i AND mask in if 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: bin_mask_p1_conv_shift) hence "i' < shift" by(simp add: mask_def i'_def int_and_le) show ?thesis proof(cases "i' !! index") case True hence unf: "i' = overflow OR i'" unfolding assms i'_def by(auto intro!: bin_eqI simp add: bin_nth_ops) have "overflow \ i'" by(subst unf)(rule le_int_or, simp add: bin_sign_and assms i'_def) hence "i' - shift < least \ False" unfolding assms by(cases "LENGTH('a)")(simp_all add: not_less) moreover have "overflow \ i' - shift \ False" using \i' < shift\ unfolding assms by(cases "LENGTH('a)")(auto simp add: not_le elim: less_le_trans) moreover have "word_of_int (i' - shift) = (word_of_int i :: 'a word)" using \i' < shift\ - by(auto intro!: word_eqI simp add: i'_def shift_def mask_def bin_nth_ops bin_nth_minus_p2 bin_sign_and) + by (simp add: i'_def shift_def mask_def shiftl_eq_push_bit push_bit_of_1 flip: take_bit_eq_mask) ultimately show ?thesis using True by(simp add: Let_def i'_def) next case False hence "i' = i AND Bit_Operations.mask (LENGTH('a) - 1)" unfolding assms i'_def by(clarsimp simp add: i'_def bin_nth_ops intro!: bin_eqI)(cases "LENGTH('a)", auto simp add: less_Suc_eq) also have "\ \ Bit_Operations.mask (LENGTH('a) - 1)" by(rule int_and_le) simp also have "\ < overflow" unfolding overflow_def by(simp add: bin_mask_p1_conv_shift[symmetric]) also have "least \ 0" unfolding least_def overflow_def by simp - have "0 \ i'" by(simp add: i'_def mask_def bin_mask_ge0) + 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(rule word_eqI)(auto simp add: i'_def bin_nth_ops mask_def) ultimately show ?thesis using False by(simp add: Let_def i'_def) qed qed text \Quickcheck conversion functions\ notation scomp (infixl "\\" 60) 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)))" no_notation scomp (infixl "\\" 60) 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 export_code signed_take_bit \mask :: nat \ int\ in SML module_name Code 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,618 +1,618 @@ (* Title: Uint16.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 16 bits\ theory Uint16 imports Code_Target_Word_Base 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. \ declare prod.Quotient[transfer_rule] section \Type definition and primitive operations\ typedef uint16 = "UNIV :: 16 word set" .. setup_lifting type_definition_uint16 text \Use an abstract type for code generation to disable pattern matching on @{term Abs_uint16}.\ declare Rep_uint16_inverse[code abstype] declare Quotient_uint16[transfer_rule] instantiation uint16 :: comm_ring_1 begin lift_definition zero_uint16 :: uint16 is "0 :: 16 word" . lift_definition one_uint16 :: uint16 is "1" . lift_definition plus_uint16 :: "uint16 \ uint16 \ uint16" is "(+) :: 16 word \ _" . lift_definition minus_uint16 :: "uint16 \ uint16 \ uint16" is "(-)" . lift_definition uminus_uint16 :: "uint16 \ uint16" is uminus . lift_definition times_uint16 :: "uint16 \ uint16 \ uint16" is "(*)" . instance by (standard; transfer) (simp_all add: algebra_simps) end instantiation uint16 :: semiring_modulo begin lift_definition divide_uint16 :: "uint16 \ uint16 \ uint16" is "(div)" . lift_definition modulo_uint16 :: "uint16 \ uint16 \ uint16" is "(mod)" . instance by (standard; transfer) (fact word_mod_div_equality) end instantiation uint16 :: linorder begin lift_definition less_uint16 :: "uint16 \ uint16 \ bool" is "(<)" . lift_definition less_eq_uint16 :: "uint16 \ uint16 \ bool" is "(\)" . instance by (standard; transfer) (simp_all add: less_le_not_le linear) end lemmas [code] = less_uint16.rep_eq less_eq_uint16.rep_eq context includes lifting_syntax notes transfer_rule_of_bool [transfer_rule] transfer_rule_numeral [transfer_rule] begin lemma [transfer_rule]: "((=) ===> cr_uint16) of_bool of_bool" by transfer_prover lemma transfer_rule_numeral_uint [transfer_rule]: "((=) ===> cr_uint16) numeral numeral" by transfer_prover lemma [transfer_rule]: \(cr_uint16 ===> (\)) even ((dvd) 2 :: uint16 \ bool)\ by (unfold dvd_def) transfer_prover end instantiation uint16 :: semiring_bits begin lift_definition bit_uint16 :: \uint16 \ nat \ bool\ is bit . instance by (standard; transfer) (fact bit_iff_odd even_iff_mod_2_eq_zero odd_iff_mod_2_eq_one odd_one bits_induct bits_div_0 bits_div_by_1 bits_mod_div_trivial even_succ_div_2 even_mask_div_iff exp_div_exp_eq div_exp_eq mod_exp_eq mult_exp_mod_exp_eq div_exp_mod_exp_eq even_mult_exp_div_exp_iff)+ end instantiation uint16 :: semiring_bit_shifts begin lift_definition push_bit_uint16 :: \nat \ uint16 \ uint16\ is push_bit . lift_definition drop_bit_uint16 :: \nat \ uint16 \ uint16\ is drop_bit . lift_definition take_bit_uint16 :: \nat \ uint16 \ uint16\ is take_bit . instance by (standard; transfer) (fact push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod)+ end instantiation uint16 :: ring_bit_operations begin 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 mask_uint16 :: \nat \ uint16\ is mask . instance by (standard; transfer) (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff minus_eq_not_minus_1 mask_eq_decr_exp) end lemma [code]: \take_bit n a = a AND mask n\ for a :: uint16 by (fact take_bit_eq_mask) lemma [code]: \mask (Suc n) = push_bit n (1 :: uint16) OR mask n\ \mask 0 = (0 :: uint16)\ by (simp_all add: mask_Suc_exp push_bit_of_1) instantiation uint16:: semiring_bit_syntax begin lift_definition test_bit_uint16 :: \uint16 \ nat \ bool\ is test_bit . lift_definition shiftl_uint16 :: \uint16 \ nat \ uint16\ is shiftl . lift_definition shiftr_uint16 :: \uint16 \ nat \ uint16\ is shiftr . instance by (standard; transfer) (fact test_bit_eq_bit shiftl_word_eq shiftr_word_eq)+ end instantiation uint16 :: lsb begin lift_definition lsb_uint16 :: \uint16 \ bool\ is lsb . instance by (standard; transfer) (fact lsb_odd) end instantiation uint16 :: msb begin lift_definition msb_uint16 :: \uint16 \ bool\ is msb . instance .. end instantiation uint16 :: set_bit begin lift_definition set_bit_uint16 :: \uint16 \ nat \ bool \ uint16\ is set_bit . instance apply standard apply (unfold Bit_Operations.set_bit_def unset_bit_def) apply transfer apply (simp add: set_bit_eq Bit_Operations.set_bit_def unset_bit_def) done end instantiation uint16 :: bit_comprehension begin lift_definition set_bits_uint16 :: "(nat \ bool) \ uint16" is "set_bits" . instance .. end lemmas [code] = test_bit_uint16.rep_eq lsb_uint16.rep_eq msb_uint16.rep_eq instantiation uint16 :: equal begin lift_definition equal_uint16 :: "uint16 \ uint16 \ bool" is "equal_class.equal" . instance by standard (transfer, simp add: equal_eq) end lemmas [code] = equal_uint16.rep_eq instantiation uint16 :: size begin lift_definition size_uint16 :: "uint16 \ nat" is "size" . instance .. end lemmas [code] = size_uint16.rep_eq lift_definition sshiftr_uint16 :: "uint16 \ nat \ uint16" (infixl ">>>" 55) is sshiftr . lift_definition uint16_of_int :: "int \ uint16" is "word_of_int" . definition uint16_of_nat :: "nat \ uint16" where "uint16_of_nat = uint16_of_int \ int" lift_definition int_of_uint16 :: "uint16 \ int" is "uint" . lift_definition nat_of_uint16 :: "uint16 \ nat" is "unat" . definition integer_of_uint16 :: "uint16 \ integer" where "integer_of_uint16 = integer_of_int o int_of_uint16" text \Use pretty numerals from integer for pretty printing\ context includes integer.lifting begin lift_definition Uint16 :: "integer \ uint16" is "word_of_int" . lemma Rep_uint16_numeral [simp]: "Rep_uint16 (numeral n) = numeral n" by(induction n)(simp_all add: one_uint16_def Abs_uint16_inverse numeral.simps plus_uint16_def) lemma Rep_uint16_neg_numeral [simp]: "Rep_uint16 (- numeral n) = - numeral n" by(simp only: uminus_uint16_def)(simp add: Abs_uint16_inverse) lemma numeral_uint16_transfer [transfer_rule]: "(rel_fun (=) cr_uint16) numeral numeral" by(auto simp add: cr_uint16_def) lemma numeral_uint16 [code_unfold]: "numeral n = Uint16 (numeral n)" by transfer simp lemma neg_numeral_uint16 [code_unfold]: "- numeral n = Uint16 (- numeral n)" by transfer(simp add: cr_uint16_def) end lemma Abs_uint16_numeral [code_post]: "Abs_uint16 (numeral n) = numeral n" by(induction n)(simp_all add: one_uint16_def numeral.simps plus_uint16_def Abs_uint16_inverse) lemma Abs_uint16_0 [code_post]: "Abs_uint16 0 = 0" by(simp add: zero_uint16_def) lemma Abs_uint16_1 [code_post]: "Abs_uint16 1 = 1" by(simp add: one_uint16_def) 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. x !! n)" by transfer simp 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 Uin16_code [code abstract]: "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 \ _" \ (SML_word) "Word16.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (Scala) "(_ & _).toChar" | constant "(OR) :: uint16 \ _" \ (SML_word) "Word16.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (Scala) "(_ | _).toChar" | constant "(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 abstract]: "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 abstract]: "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 (test_bit :: uint16 \ _) x n else x !! (nat_of_integer n))" lemma test_bit_eq_bit_uint16 [code]: \test_bit = (bit :: uint16 \ _)\ by (rule ext)+ (transfer, transfer, simp) 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 (test_bit :: uint16 \ _) w n else Rep_uint16 w !! nat_of_integer n)" unfolding uint16_test_bit_def by(simp add: test_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 abstract]: "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" lift_definition uint16_set_bits :: "(nat \ bool) \ uint16 \ nat \ uint16" is set_bits_aux . lemma uint16_set_bits_code [code]: "uint16_set_bits f w n = (if n = 0 then w else let n' = n - 1 in uint16_set_bits f ((w << 1) OR (if f n' then 1 else 0)) n')" by(transfer fixing: n)(cases n, simp_all) lemma set_bits_uint16 [code]: "(BITS n. f n) = uint16_set_bits f 0 16" by transfer(simp add: set_bits_conv_set_bits_aux) lemma lsb_code [code]: fixes x :: uint16 shows "lsb x = x !! 0" by transfer(simp add: word_lsb_def word_test_bit_def) definition uint16_shiftl :: "uint16 \ integer \ uint16" where [code del]: "uint16_shiftl x n = (if n < 0 \ 16 \ n then undefined (shiftl :: uint16 \ _) x n else x << (nat_of_integer n))" lemma shiftl_uint16_code [code]: "x << n = (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 add: not_less shiftl_zero_size word_size) lemma uint16_shiftl_code [code abstract]: "Rep_uint16 (uint16_shiftl w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined (shiftl :: uint16 \ _) w n) else Rep_uint16 w << nat_of_integer n)" 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 (shiftr :: uint16 \ _) x n else x >> (nat_of_integer n))" lemma shiftr_uint16_code [code]: "x >> n = (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 add: not_less shiftr_zero_size word_size) lemma uint16_shiftr_code [code abstract]: "Rep_uint16 (uint16_shiftr w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined (shiftr :: uint16 \ _) w n) else Rep_uint16 w >> nat_of_integer n)" 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 sshiftr_uint16 x n else sshiftr_uint16 x (nat_of_integer n))" lemma sshiftr_beyond: fixes x :: "'a :: len word" shows "size x \ n \ x >>> n = (if x !! (size x - 1) then -1 else 0)" by(rule word_eqI)(simp add: nth_sshiftr word_size) lemma sshiftr_uint16_code [code]: "x >>> n = (if n < 16 then uint16_sshiftr x (integer_of_nat n) else if x !! 15 then -1 else 0)" including undefined_transfer integer.lifting unfolding uint16_sshiftr_def by transfer (simp add: not_less sshiftr_beyond word_size) lemma uint16_sshiftr_code [code abstract]: "Rep_uint16 (uint16_sshiftr w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined sshiftr_uint16 w n) else Rep_uint16 w >>> nat_of_integer n)" 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 \ (x :: uint16) !! 15" by transfer(simp add: msb_nth) 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: 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 add: unat_def) +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 no_notation sshiftr_uint16 (infixl ">>>" 55) 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,748 +1,748 @@ (* Title: Uint32.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 32 bits\ theory Uint32 imports Code_Target_Word_Base begin declare prod.Quotient[transfer_rule] section \Type definition and primitive operations\ typedef uint32 = "UNIV :: 32 word set" .. setup_lifting type_definition_uint32 text \Use an abstract type for code generation to disable pattern matching on @{term Abs_uint32}.\ declare Rep_uint32_inverse[code abstype] declare Quotient_uint32[transfer_rule] instantiation uint32 :: comm_ring_1 begin lift_definition zero_uint32 :: uint32 is "0 :: 32 word" . lift_definition one_uint32 :: uint32 is "1" . lift_definition plus_uint32 :: "uint32 \ uint32 \ uint32" is "(+) :: 32 word \ _" . lift_definition minus_uint32 :: "uint32 \ uint32 \ uint32" is "(-)" . lift_definition uminus_uint32 :: "uint32 \ uint32" is uminus . lift_definition times_uint32 :: "uint32 \ uint32 \ uint32" is "(*)" . instance by (standard; transfer) (simp_all add: algebra_simps) end instantiation uint32 :: semiring_modulo begin lift_definition divide_uint32 :: "uint32 \ uint32 \ uint32" is "(div)" . lift_definition modulo_uint32 :: "uint32 \ uint32 \ uint32" is "(mod)" . instance by (standard; transfer) (fact word_mod_div_equality) end instantiation uint32 :: linorder begin lift_definition less_uint32 :: "uint32 \ uint32 \ bool" is "(<)" . lift_definition less_eq_uint32 :: "uint32 \ uint32 \ bool" is "(\)" . instance by (standard; transfer) (simp_all add: less_le_not_le linear) end lemmas [code] = less_uint32.rep_eq less_eq_uint32.rep_eq context includes lifting_syntax notes transfer_rule_of_bool [transfer_rule] transfer_rule_numeral [transfer_rule] begin lemma [transfer_rule]: "((=) ===> cr_uint32) of_bool of_bool" by transfer_prover lemma transfer_rule_numeral_uint [transfer_rule]: "((=) ===> cr_uint32) numeral numeral" by transfer_prover lemma [transfer_rule]: \(cr_uint32 ===> (\)) even ((dvd) 2 :: uint32 \ bool)\ by (unfold dvd_def) transfer_prover end instantiation uint32:: semiring_bits begin lift_definition bit_uint32 :: \uint32 \ nat \ bool\ is bit . instance by (standard; transfer) (fact bit_iff_odd even_iff_mod_2_eq_zero odd_iff_mod_2_eq_one odd_one bits_induct bits_div_0 bits_div_by_1 bits_mod_div_trivial even_succ_div_2 even_mask_div_iff exp_div_exp_eq div_exp_eq mod_exp_eq mult_exp_mod_exp_eq div_exp_mod_exp_eq even_mult_exp_div_exp_iff)+ end instantiation uint32 :: semiring_bit_shifts begin lift_definition push_bit_uint32 :: \nat \ uint32 \ uint32\ is push_bit . lift_definition drop_bit_uint32 :: \nat \ uint32 \ uint32\ is drop_bit . lift_definition take_bit_uint32 :: \nat \ uint32 \ uint32\ is take_bit . instance by (standard; transfer) (fact push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod)+ end instantiation uint32 :: ring_bit_operations begin 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 mask_uint32 :: \nat \ uint32\ is mask . instance by (standard; transfer) (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff minus_eq_not_minus_1 mask_eq_decr_exp) end lemma [code]: \take_bit n a = a AND mask n\ for a :: uint32 by (fact take_bit_eq_mask) lemma [code]: \mask (Suc n) = push_bit n (1 :: uint32) OR mask n\ \mask 0 = (0 :: uint32)\ by (simp_all add: mask_Suc_exp push_bit_of_1) instantiation uint32:: semiring_bit_syntax begin lift_definition test_bit_uint32 :: \uint32 \ nat \ bool\ is test_bit . lift_definition shiftl_uint32 :: \uint32 \ nat \ uint32\ is shiftl . lift_definition shiftr_uint32 :: \uint32 \ nat \ uint32\ is shiftr . instance by (standard; transfer) (fact test_bit_eq_bit shiftl_word_eq shiftr_word_eq)+ end instantiation uint32 :: lsb begin lift_definition lsb_uint32 :: \uint32 \ bool\ is lsb . instance by (standard; transfer) (fact lsb_odd) end instantiation uint32 :: msb begin lift_definition msb_uint32 :: \uint32 \ bool\ is msb . instance .. end instantiation uint32 :: set_bit begin lift_definition set_bit_uint32 :: \uint32 \ nat \ bool \ uint32\ is set_bit . instance apply standard apply (unfold Bit_Operations.set_bit_def unset_bit_def) apply transfer apply (simp add: set_bit_eq Bit_Operations.set_bit_def unset_bit_def) done end instantiation uint32 :: bit_comprehension begin lift_definition set_bits_uint32 :: "(nat \ bool) \ uint32" is "set_bits" . instance .. end lemmas [code] = test_bit_uint32.rep_eq lsb_uint32.rep_eq msb_uint32.rep_eq instantiation uint32 :: equal begin lift_definition equal_uint32 :: "uint32 \ uint32 \ bool" is "equal_class.equal" . instance by standard (transfer, simp add: equal_eq) end lemmas [code] = equal_uint32.rep_eq instantiation uint32 :: size begin lift_definition size_uint32 :: "uint32 \ nat" is "size" . instance .. end lemmas [code] = size_uint32.rep_eq lift_definition sshiftr_uint32 :: "uint32 \ nat \ uint32" (infixl ">>>" 55) is sshiftr . lift_definition uint32_of_int :: "int \ uint32" is "word_of_int" . definition uint32_of_nat :: "nat \ uint32" where "uint32_of_nat = uint32_of_int \ int" lift_definition int_of_uint32 :: "uint32 \ int" is "uint" . lift_definition nat_of_uint32 :: "uint32 \ nat" is "unat" . definition integer_of_uint32 :: "uint32 \ integer" where "integer_of_uint32 = integer_of_int o int_of_uint32" text \Use pretty numerals from integer for pretty printing\ context includes integer.lifting begin lift_definition Uint32 :: "integer \ uint32" is "word_of_int" . lemma Rep_uint32_numeral [simp]: "Rep_uint32 (numeral n) = numeral n" by(induction n)(simp_all add: one_uint32_def Abs_uint32_inverse numeral.simps plus_uint32_def) lemma numeral_uint32_transfer [transfer_rule]: "(rel_fun (=) cr_uint32) numeral numeral" by(auto simp add: cr_uint32_def) lemma numeral_uint32 [code_unfold]: "numeral n = Uint32 (numeral n)" by transfer simp lemma Rep_uint32_neg_numeral [simp]: "Rep_uint32 (- numeral n) = - numeral n" by(simp only: uminus_uint32_def)(simp add: Abs_uint32_inverse) lemma neg_numeral_uint32 [code_unfold]: "- numeral n = Uint32 (- numeral n)" by transfer(simp add: cr_uint32_def) end lemma Abs_uint32_numeral [code_post]: "Abs_uint32 (numeral n) = numeral n" by(induction n)(simp_all add: one_uint32_def numeral.simps plus_uint32_def Abs_uint32_inverse) lemma Abs_uint32_0 [code_post]: "Abs_uint32 0 = 0" by(simp add: zero_uint32_def) lemma Abs_uint32_1 [code_post]: "Abs_uint32 1 = 1" by(simp add: one_uint32_def) 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. \ 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 i' !! 31 then Uint32_signed (i' - 0x100000000) else Uint32_signed i')" including undefined_transfer integer.lifting unfolding Uint32_signed_def by transfer(rule word_of_int_via_signed, simp_all add: bin_mask_numeral) lemma Uint32_signed_code [code abstract]: "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) 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. x !! n)" by transfer simp 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 \ _" \ (SML) "Word32.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (OCaml) "Int32.logand" and (Scala) infixl 3 "&" | constant "(OR) :: uint32 \ _" \ (SML) "Word32.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (OCaml) "Int32.logor" and (Scala) infixl 1 "|" | constant "(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 = (uint32_sdiv (x >> 1) y) << 1; 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 by transfer(simp add: divmod_via_sdivmod) lemma uint32_sdiv_code [code abstract]: "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 (test_bit :: uint32 \ _) x n else x !! (nat_of_integer n))" lemma test_bit_eq_bit_uint32 [code]: \test_bit = (bit :: uint32 \ _)\ by (rule ext)+ (transfer, transfer, simp) 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 (test_bit :: uint32 \ _) w n else Rep_uint32 w !! nat_of_integer n)" unfolding uint32_test_bit_def by(simp add: test_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 abstract]: "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)" lift_definition uint32_set_bits :: "(nat \ bool) \ uint32 \ nat \ uint32" is set_bits_aux . lemma uint32_set_bits_code [code]: "uint32_set_bits f w n = (if n = 0 then w else let n' = n - 1 in uint32_set_bits f ((w << 1) OR (if f n' then 1 else 0)) n')" by(transfer fixing: n)(cases n, simp_all) lemma set_bits_uint32 [code]: "(BITS n. f n) = uint32_set_bits f 0 32" by transfer(simp add: set_bits_conv_set_bits_aux) lemma lsb_code [code]: fixes x :: uint32 shows "lsb x = x !! 0" by transfer(simp add: word_lsb_def word_test_bit_def) definition uint32_shiftl :: "uint32 \ integer \ uint32" where [code del]: "uint32_shiftl x n = (if n < 0 \ 32 \ n then undefined (shiftl :: uint32 \ _) x n else x << (nat_of_integer n))" lemma shiftl_uint32_code [code]: "x << n = (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 add: not_less shiftl_zero_size word_size) lemma uint32_shiftl_code [code abstract]: "Rep_uint32 (uint32_shiftl w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined (shiftl :: uint32 \ _) w n) else Rep_uint32 w << (nat_of_integer n))" 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 (shiftr :: uint32 \ _) x n else x >> (nat_of_integer n))" lemma shiftr_uint32_code [code]: "x >> n = (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 add: not_less shiftr_zero_size word_size) lemma uint32_shiftr_code [code abstract]: "Rep_uint32 (uint32_shiftr w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined (shiftr :: uint32 \ _) w n) else Rep_uint32 w >> nat_of_integer n)" 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 sshiftr_uint32 x n else sshiftr_uint32 x (nat_of_integer n))" lemma sshiftr_beyond: fixes x :: "'a :: len word" shows "size x \ n \ x >>> n = (if x !! (size x - 1) then -1 else 0)" by(rule word_eqI)(simp add: nth_sshiftr word_size) lemma sshiftr_uint32_code [code]: "x >>> n = (if n < 32 then uint32_sshiftr x (integer_of_nat n) else if x !! 31 then -1 else 0)" including undefined_transfer integer.lifting unfolding uint32_sshiftr_def by transfer(simp add: not_less sshiftr_beyond word_size) lemma uint32_sshiftr_code [code abstract]: "Rep_uint32 (uint32_sshiftr w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined sshiftr_uint32 w n) else Rep_uint32 w >>> (nat_of_integer n))" 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)" lemma uint32_msb_test_bit: "msb x \ (x :: uint32) !! 31" by transfer(simp add: msb_nth) 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)" by(simp add: integer_of_uint32_def) 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 add: unat_def) +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 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 n !! 31 then undefined integer_of_uint32 n else integer_of_int (uint (Rep_uint32' n)))" unfolding integer_of_uint32_signed_def integer_of_uint32_def including undefined_transfer by transfer simp lemma integer_of_uint32_code [code]: "integer_of_uint32 n = (if n !! 31 then integer_of_uint32_signed (n AND 0x7FFFFFFF) OR 0x80000000 else integer_of_uint32_signed n)" unfolding integer_of_uint32_def integer_of_uint32_signed_def o_def including undefined_transfer integer.lifting by transfer(auto simp add: word_ao_nth uint_and_mask_or_full mask_numeral mask_Suc_0 intro!: uint_and_mask_or_full[symmetric]) 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 no_notation sshiftr_uint32 (infixl ">>>" 55) 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,949 +1,949 @@ (* Title: Uint64.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 64 bits\ theory Uint64 imports Code_Target_Word_Base 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. \ declare prod.Quotient[transfer_rule] section \Type definition and primitive operations\ typedef uint64 = "UNIV :: 64 word set" .. setup_lifting type_definition_uint64 text \Use an abstract type for code generation to disable pattern matching on @{term Abs_uint64}.\ declare Rep_uint64_inverse[code abstype] declare Quotient_uint64[transfer_rule] instantiation uint64 :: comm_ring_1 begin lift_definition zero_uint64 :: uint64 is "0 :: 64 word" . lift_definition one_uint64 :: uint64 is "1" . lift_definition plus_uint64 :: "uint64 \ uint64 \ uint64" is "(+) :: 64 word \ _" . lift_definition minus_uint64 :: "uint64 \ uint64 \ uint64" is "(-)" . lift_definition uminus_uint64 :: "uint64 \ uint64" is uminus . lift_definition times_uint64 :: "uint64 \ uint64 \ uint64" is "(*)" . instance by (standard; transfer) (simp_all add: algebra_simps) end instantiation uint64 :: semiring_modulo begin lift_definition divide_uint64 :: "uint64 \ uint64 \ uint64" is "(div)" . lift_definition modulo_uint64 :: "uint64 \ uint64 \ uint64" is "(mod)" . instance by (standard; transfer) (fact word_mod_div_equality) end instantiation uint64 :: linorder begin lift_definition less_uint64 :: "uint64 \ uint64 \ bool" is "(<)" . lift_definition less_eq_uint64 :: "uint64 \ uint64 \ bool" is "(\)" . instance by (standard; transfer) (simp_all add: less_le_not_le linear) end lemmas [code] = less_uint64.rep_eq less_eq_uint64.rep_eq context includes lifting_syntax notes transfer_rule_of_bool [transfer_rule] transfer_rule_numeral [transfer_rule] begin lemma [transfer_rule]: "((=) ===> cr_uint64) of_bool of_bool" by transfer_prover lemma transfer_rule_numeral_uint [transfer_rule]: "((=) ===> cr_uint64) numeral numeral" by transfer_prover lemma [transfer_rule]: \(cr_uint64 ===> (\)) even ((dvd) 2 :: uint64 \ bool)\ by (unfold dvd_def) transfer_prover end instantiation uint64 :: semiring_bits begin lift_definition bit_uint64 :: \uint64 \ nat \ bool\ is bit . instance by (standard; transfer) (fact bit_iff_odd even_iff_mod_2_eq_zero odd_iff_mod_2_eq_one odd_one bits_induct bits_div_0 bits_div_by_1 bits_mod_div_trivial even_succ_div_2 even_mask_div_iff exp_div_exp_eq div_exp_eq mod_exp_eq mult_exp_mod_exp_eq div_exp_mod_exp_eq even_mult_exp_div_exp_iff)+ end instantiation uint64 :: semiring_bit_shifts begin lift_definition push_bit_uint64 :: \nat \ uint64 \ uint64\ is push_bit . lift_definition drop_bit_uint64 :: \nat \ uint64 \ uint64\ is drop_bit . lift_definition take_bit_uint64 :: \nat \ uint64 \ uint64\ is take_bit . instance by (standard; transfer) (fact push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod)+ end instantiation uint64 :: ring_bit_operations begin 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 mask_uint64 :: \nat \ uint64\ is mask . instance by (standard; transfer) (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff minus_eq_not_minus_1 mask_eq_decr_exp) end lemma [code]: \take_bit n a = a AND mask n\ for a :: uint64 by (fact take_bit_eq_mask) lemma [code]: \mask (Suc n) = push_bit n (1 :: uint64) OR mask n\ \mask 0 = (0 :: uint64)\ by (simp_all add: mask_Suc_exp push_bit_of_1) instantiation uint64:: semiring_bit_syntax begin lift_definition test_bit_uint64 :: \uint64 \ nat \ bool\ is test_bit . lift_definition shiftl_uint64 :: \uint64 \ nat \ uint64\ is shiftl . lift_definition shiftr_uint64 :: \uint64 \ nat \ uint64\ is shiftr . instance by (standard; transfer) (fact test_bit_eq_bit shiftl_word_eq shiftr_word_eq)+ end instantiation uint64 :: lsb begin lift_definition lsb_uint64 :: \uint64 \ bool\ is lsb . instance by (standard; transfer) (fact lsb_odd) end instantiation uint64 :: msb begin lift_definition msb_uint64 :: \uint64 \ bool\ is msb . instance .. end instantiation uint64 :: set_bit begin lift_definition set_bit_uint64 :: \uint64 \ nat \ bool \ uint64\ is set_bit . instance apply standard apply (unfold Bit_Operations.set_bit_def unset_bit_def) apply transfer apply (simp add: set_bit_eq Bit_Operations.set_bit_def unset_bit_def) done end instantiation uint64 :: bit_comprehension begin lift_definition set_bits_uint64 :: "(nat \ bool) \ uint64" is "set_bits" . instance .. end lemmas [code] = test_bit_uint64.rep_eq lsb_uint64.rep_eq msb_uint64.rep_eq instantiation uint64 :: equal begin lift_definition equal_uint64 :: "uint64 \ uint64 \ bool" is "equal_class.equal" . instance by standard (transfer, simp add: equal_eq) end lemmas [code] = equal_uint64.rep_eq instantiation uint64 :: size begin lift_definition size_uint64 :: "uint64 \ nat" is "size" . instance .. end lemmas [code] = size_uint64.rep_eq lift_definition sshiftr_uint64 :: "uint64 \ nat \ uint64" (infixl ">>>" 55) is sshiftr . lift_definition uint64_of_int :: "int \ uint64" is "word_of_int" . definition uint64_of_nat :: "nat \ uint64" where "uint64_of_nat = uint64_of_int \ int" lift_definition int_of_uint64 :: "uint64 \ int" is "uint" . lift_definition nat_of_uint64 :: "uint64 \ nat" is "unat" . definition integer_of_uint64 :: "uint64 \ integer" where "integer_of_uint64 = integer_of_int o int_of_uint64" text \Use pretty numerals from integer for pretty printing\ context includes integer.lifting begin lift_definition Uint64 :: "integer \ uint64" is "word_of_int" . lemma Rep_uint64_numeral [simp]: "Rep_uint64 (numeral n) = numeral n" by(induction n)(simp_all add: one_uint64_def Abs_uint64_inverse numeral.simps plus_uint64_def) lemma numeral_uint64_transfer [transfer_rule]: "(rel_fun (=) cr_uint64) numeral numeral" by(auto simp add: cr_uint64_def) lemma numeral_uint64 [code_unfold]: "numeral n = Uint64 (numeral n)" by transfer simp lemma Rep_uint64_neg_numeral [simp]: "Rep_uint64 (- numeral n) = - numeral n" by(simp only: uminus_uint64_def)(simp add: Abs_uint64_inverse) lemma neg_numeral_uint64 [code_unfold]: "- numeral n = Uint64 (- numeral n)" by transfer(simp add: cr_uint64_def) end lemma Abs_uint64_numeral [code_post]: "Abs_uint64 (numeral n) = numeral n" by(induction n)(simp_all add: one_uint64_def numeral.simps plus_uint64_def Abs_uint64_inverse) lemma Abs_uint64_0 [code_post]: "Abs_uint64 0 = 0" by(simp add: zero_uint64_def) lemma Abs_uint64_1 [code_post]: "Abs_uint64 1 = 1" by(simp add: one_uint64_def) 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. \ 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 i' !! 63 then Uint64_signed (i' - 0x10000000000000000) else Uint64_signed i')" including undefined_transfer integer.lifting unfolding Uint64_signed_def by transfer(rule word_of_int_via_signed, simp_all add: bin_mask_numeral) lemma Uint64_signed_code [code abstract]: "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) 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. x !! n)" by transfer simp 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 \ _" \ (SML) "Uint64.andb" and (Haskell) infixl 7 "Data_Bits..&." and (OCaml) "Int64.logand" and (Scala) infixl 3 "&" | constant "(OR) :: uint64 \ _" \ (SML) "Uint64.orb" and (Haskell) infixl 5 "Data_Bits..|." and (OCaml) "Int64.logor" and (Scala) infixl 1 "|" | constant "(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 = (uint64_sdiv (x >> 1) y) << 1; 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 by transfer(simp add: divmod_via_sdivmod) lemma uint64_sdiv_code [code abstract]: "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 (test_bit :: uint64 \ _) x n else x !! (nat_of_integer n))" lemma test_bit_eq_bit_uint64 [code]: \test_bit = (bit :: uint64 \ _)\ by (rule ext)+ (transfer, transfer, simp) 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, simp, transfer, simp) lemma uint64_test_bit_code [code]: "uint64_test_bit w n = (if n < 0 \ 63 < n then undefined (test_bit :: uint64 \ _) w n else Rep_uint64 w !! nat_of_integer n)" unfolding uint64_test_bit_def by(simp add: test_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 abstract]: "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)" lift_definition uint64_set_bits :: "(nat \ bool) \ uint64 \ nat \ uint64" is set_bits_aux . lemma uint64_set_bits_code [code]: "uint64_set_bits f w n = (if n = 0 then w else let n' = n - 1 in uint64_set_bits f ((w << 1) OR (if f n' then 1 else 0)) n')" by(transfer fixing: n)(cases n, simp_all) lemma set_bits_uint64 [code]: "(BITS n. f n) = uint64_set_bits f 0 64" by transfer(simp add: set_bits_conv_set_bits_aux) lemma lsb_code [code]: fixes x :: uint64 shows "lsb x = x !! 0" by transfer(simp add: word_lsb_def word_test_bit_def) definition uint64_shiftl :: "uint64 \ integer \ uint64" where [code del]: "uint64_shiftl x n = (if n < 0 \ 64 \ n then undefined (shiftl :: uint64 \ _) x n else x << (nat_of_integer n))" lemma shiftl_uint64_code [code]: "x << n = (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 add: not_less shiftl_zero_size word_size) lemma uint64_shiftl_code [code abstract]: "Rep_uint64 (uint64_shiftl w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined (shiftl :: uint64 \ _) w n) else Rep_uint64 w << (nat_of_integer n))" 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 (shiftr :: uint64 \ _) x n else x >> (nat_of_integer n))" lemma shiftr_uint64_code [code]: "x >> n = (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 add: not_less shiftr_zero_size word_size) lemma uint64_shiftr_code [code abstract]: "Rep_uint64 (uint64_shiftr w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined (shiftr :: uint64 \ _) w n) else Rep_uint64 w >> nat_of_integer n)" 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 sshiftr_uint64 x n else sshiftr_uint64 x (nat_of_integer n))" lemma sshiftr_beyond: fixes x :: "'a :: len word" shows "size x \ n \ x >>> n = (if x !! (size x - 1) then -1 else 0)" by(rule word_eqI)(simp add: nth_sshiftr word_size) lemma sshiftr_uint64_code [code]: "x >>> n = (if n < 64 then uint64_sshiftr x (integer_of_nat n) else if x !! 63 then -1 else 0)" including undefined_transfer integer.lifting unfolding uint64_sshiftr_def by transfer(simp add: not_less sshiftr_beyond word_size) lemma uint64_sshiftr_code [code abstract]: "Rep_uint64 (uint64_sshiftr w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined sshiftr_uint64 w n) else Rep_uint64 w >>> (nat_of_integer n))" 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)" lemma uint64_msb_test_bit: "msb x \ (x :: uint64) !! 63" by transfer(simp add: msb_nth) 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)" by(simp add: integer_of_uint64_def) 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 add: unat_def) +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 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 n !! 63 then undefined integer_of_uint64 n else integer_of_int (uint (Rep_uint64' n)))" unfolding integer_of_uint64_signed_def integer_of_uint64_def including undefined_transfer by transfer simp lemma integer_of_uint64_code [code]: "integer_of_uint64 n = (if n !! 63 then integer_of_uint64_signed (n AND 0x7FFFFFFFFFFFFFFF) OR 0x8000000000000000 else integer_of_uint64_signed n)" unfolding integer_of_uint64_def integer_of_uint64_signed_def o_def including undefined_transfer integer.lifting by transfer(auto simp add: word_ao_nth uint_and_mask_or_full mask_numeral mask_Suc_0 intro!: uint_and_mask_or_full[symmetric]) 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 no_notation sshiftr_uint64 (infixl ">>>" 55) 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,672 +1,672 @@ (* Title: Uint8.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 8 bits\ theory Uint8 imports Code_Target_Word_Base 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"}. \ declare prod.Quotient[transfer_rule] section \Type definition and primitive operations\ typedef uint8 = "UNIV :: 8 word set" .. setup_lifting type_definition_uint8 text \Use an abstract type for code generation to disable pattern matching on @{term Abs_uint8}.\ declare Rep_uint8_inverse[code abstype] declare Quotient_uint8[transfer_rule] instantiation uint8 :: comm_ring_1 begin lift_definition zero_uint8 :: uint8 is "0 :: 8 word" . lift_definition one_uint8 :: uint8 is "1" . lift_definition plus_uint8 :: "uint8 \ uint8 \ uint8" is "(+) :: 8 word \ _" . lift_definition minus_uint8 :: "uint8 \ uint8 \ uint8" is "(-)" . lift_definition uminus_uint8 :: "uint8 \ uint8" is uminus . lift_definition times_uint8 :: "uint8 \ uint8 \ uint8" is "(*)" . instance by (standard; transfer) (simp_all add: algebra_simps) end instantiation uint8 :: semiring_modulo begin lift_definition divide_uint8 :: "uint8 \ uint8 \ uint8" is "(div)" . lift_definition modulo_uint8 :: "uint8 \ uint8 \ uint8" is "(mod)" . instance by (standard; transfer) (fact word_mod_div_equality) end instantiation uint8 :: linorder begin lift_definition less_uint8 :: "uint8 \ uint8 \ bool" is "(<)" . lift_definition less_eq_uint8 :: "uint8 \ uint8 \ bool" is "(\)" . instance by (standard; transfer) (simp_all add: less_le_not_le linear) end lemmas [code] = less_uint8.rep_eq less_eq_uint8.rep_eq context includes lifting_syntax notes transfer_rule_of_bool [transfer_rule] transfer_rule_numeral [transfer_rule] begin lemma [transfer_rule]: "((=) ===> cr_uint8) of_bool of_bool" by transfer_prover lemma transfer_rule_numeral_uint [transfer_rule]: "((=) ===> cr_uint8) numeral numeral" by transfer_prover lemma [transfer_rule]: \(cr_uint8 ===> (\)) even ((dvd) 2 :: uint8 \ bool)\ by (unfold dvd_def) transfer_prover end instantiation uint8 :: semiring_bits begin lift_definition bit_uint8 :: \uint8 \ nat \ bool\ is bit . instance by (standard; transfer) (fact bit_iff_odd even_iff_mod_2_eq_zero odd_iff_mod_2_eq_one odd_one bits_induct bits_div_0 bits_div_by_1 bits_mod_div_trivial even_succ_div_2 even_mask_div_iff exp_div_exp_eq div_exp_eq mod_exp_eq mult_exp_mod_exp_eq div_exp_mod_exp_eq even_mult_exp_div_exp_iff)+ end instantiation uint8 :: semiring_bit_shifts begin lift_definition push_bit_uint8 :: \nat \ uint8 \ uint8\ is push_bit . lift_definition drop_bit_uint8 :: \nat \ uint8 \ uint8\ is drop_bit . lift_definition take_bit_uint8 :: \nat \ uint8 \ uint8\ is take_bit . instance by (standard; transfer) (fact push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod)+ end instantiation uint8 :: ring_bit_operations begin 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 mask_uint8 :: \nat \ uint8\ is mask . instance by (standard; transfer) (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff minus_eq_not_minus_1 mask_eq_decr_exp) end lemma [code]: \take_bit n a = a AND mask n\ for a :: uint8 by (fact take_bit_eq_mask) lemma [code]: \mask (Suc n) = push_bit n (1 :: uint8) OR mask n\ \mask 0 = (0 :: uint8)\ by (simp_all add: mask_Suc_exp push_bit_of_1) instantiation uint8:: semiring_bit_syntax begin lift_definition test_bit_uint8 :: \uint8 \ nat \ bool\ is test_bit . lift_definition shiftl_uint8 :: \uint8 \ nat \ uint8\ is shiftl . lift_definition shiftr_uint8 :: \uint8 \ nat \ uint8\ is shiftr . instance by (standard; transfer) (fact test_bit_eq_bit shiftl_word_eq shiftr_word_eq)+ end instantiation uint8 :: lsb begin lift_definition lsb_uint8 :: \uint8 \ bool\ is lsb . instance by (standard; transfer) (fact lsb_odd) end instantiation uint8 :: msb begin lift_definition msb_uint8 :: \uint8 \ bool\ is msb . instance .. end instantiation uint8 :: set_bit begin lift_definition set_bit_uint8 :: \uint8 \ nat \ bool \ uint8\ is set_bit . instance apply standard apply (unfold Bit_Operations.set_bit_def unset_bit_def) apply transfer apply (simp add: set_bit_eq Bit_Operations.set_bit_def unset_bit_def) done end instantiation uint8 :: bit_comprehension begin lift_definition set_bits_uint8 :: "(nat \ bool) \ uint8" is "set_bits" . instance .. end lemmas [code] = test_bit_uint8.rep_eq lsb_uint8.rep_eq msb_uint8.rep_eq instantiation uint8 :: equal begin lift_definition equal_uint8 :: "uint8 \ uint8 \ bool" is "equal_class.equal" . instance by standard (transfer, simp add: equal_eq) end lemmas [code] = equal_uint8.rep_eq instantiation uint8 :: size begin lift_definition size_uint8 :: "uint8 \ nat" is "size" . instance .. end lemmas [code] = size_uint8.rep_eq lift_definition sshiftr_uint8 :: "uint8 \ nat \ uint8" (infixl ">>>" 55) is sshiftr . lift_definition uint8_of_int :: "int \ uint8" is "word_of_int" . definition uint8_of_nat :: "nat \ uint8" where "uint8_of_nat = uint8_of_int \ int" lift_definition int_of_uint8 :: "uint8 \ int" is "uint" . lift_definition nat_of_uint8 :: "uint8 \ nat" is "unat" . definition integer_of_uint8 :: "uint8 \ integer" where "integer_of_uint8 = integer_of_int o int_of_uint8" text \Use pretty numerals from integer for pretty printing\ context includes integer.lifting begin lift_definition Uint8 :: "integer \ uint8" is "word_of_int" . lemma Rep_uint8_numeral [simp]: "Rep_uint8 (numeral n) = numeral n" by(induction n)(simp_all add: one_uint8_def Abs_uint8_inverse numeral.simps plus_uint8_def) lemma numeral_uint8_transfer [transfer_rule]: "(rel_fun (=) cr_uint8) numeral numeral" by(auto simp add: cr_uint8_def) lemma numeral_uint8 [code_unfold]: "numeral n = Uint8 (numeral n)" by transfer simp lemma Rep_uint8_neg_numeral [simp]: "Rep_uint8 (- numeral n) = - numeral n" by(simp only: uminus_uint8_def)(simp add: Abs_uint8_inverse) lemma neg_numeral_uint8 [code_unfold]: "- numeral n = Uint8 (- numeral n)" by transfer(simp add: cr_uint8_def) end lemma Abs_uint8_numeral [code_post]: "Abs_uint8 (numeral n) = numeral n" by(induction n)(simp_all add: one_uint8_def numeral.simps plus_uint8_def Abs_uint8_inverse) lemma Abs_uint8_0 [code_post]: "Abs_uint8 0 = 0" by(simp add: zero_uint8_def) lemma Abs_uint8_1 [code_post]: "Abs_uint8 1 = 1" by(simp add: one_uint8_def) 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. x !! n)" by transfer simp 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 abstract]: "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 \ _" \ (SML) "Word8.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (Scala) "(_ & _).toByte" | constant "(OR) :: uint8 \ _" \ (SML) "Word8.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (Scala) "(_ | _).toByte" | constant "(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 = (uint8_sdiv (x >> 1) y) << 1; 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 by transfer(simp add: divmod_via_sdivmod) lemma uint8_sdiv_code [code abstract]: "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 (test_bit :: uint8 \ _) x n else x !! (nat_of_integer n))" lemma test_bit_eq_bit_uint8 [code]: \test_bit = (bit :: uint8 \ _)\ by (rule ext)+ (transfer, transfer, simp) lemma test_bit_uint8_code [code]: "test_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 (test_bit :: uint8 \ _) w n else Rep_uint8 w !! nat_of_integer n)" unfolding uint8_test_bit_def by(simp add: test_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 abstract]: "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)" lift_definition uint8_set_bits :: "(nat \ bool) \ uint8 \ nat \ uint8" is set_bits_aux . lemma uint8_set_bits_code [code]: "uint8_set_bits f w n = (if n = 0 then w else let n' = n - 1 in uint8_set_bits f ((w << 1) OR (if f n' then 1 else 0)) n')" by(transfer fixing: n)(cases n, simp_all) lemma set_bits_uint8 [code]: "(BITS n. f n) = uint8_set_bits f 0 8" by transfer(simp add: set_bits_conv_set_bits_aux) lemma lsb_code [code]: fixes x :: uint8 shows "lsb x = x !! 0" by transfer(simp add: word_lsb_def word_test_bit_def) definition uint8_shiftl :: "uint8 \ integer \ uint8" where [code del]: "uint8_shiftl x n = (if n < 0 \ 8 \ n then undefined (shiftl :: uint8 \ _) x n else x << (nat_of_integer n))" lemma shiftl_uint8_code [code]: "x << n = (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 add: not_less shiftl_zero_size word_size) lemma uint8_shiftl_code [code abstract]: "Rep_uint8 (uint8_shiftl w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined (shiftl :: uint8 \ _) w n) else Rep_uint8 w << nat_of_integer n)" 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 (shiftr :: uint8 \ _) x n else x >> (nat_of_integer n))" lemma shiftr_uint8_code [code]: "x >> n = (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 add: not_less shiftr_zero_size word_size) lemma uint8_shiftr_code [code abstract]: "Rep_uint8 (uint8_shiftr w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined (shiftr :: uint8 \ _) w n) else Rep_uint8 w >> nat_of_integer n)" 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 sshiftr_uint8 x n else sshiftr_uint8 x (nat_of_integer n))" lemma sshiftr_beyond: fixes x :: "'a :: len word" shows "size x \ n \ x >>> n = (if x !! (size x - 1) then -1 else 0)" by(rule word_eqI)(simp add: nth_sshiftr word_size) lemma sshiftr_uint8_code [code]: "x >>> n = (if n < 8 then uint8_sshiftr x (integer_of_nat n) else if x !! 7 then -1 else 0)" including undefined_transfer integer.lifting unfolding uint8_sshiftr_def by transfer (simp add: not_less sshiftr_beyond word_size) lemma uint8_sshiftr_code [code abstract]: "Rep_uint8 (uint8_sshiftr w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined sshiftr_uint8 w n) else Rep_uint8 w >>> nat_of_integer n)" 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)" lemma uint8_msb_test_bit: "msb x \ (x :: uint8) !! 7" by transfer(simp add: msb_nth) 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: 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 add: unat_def) +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 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 n !! 7 then undefined integer_of_uint8 n else integer_of_int (uint (Rep_uint8' n)))" unfolding integer_of_uint8_signed_def integer_of_uint8_def including undefined_transfer by transfer simp lemma integer_of_uint8_code [code]: "integer_of_uint8 n = (if n !! 7 then integer_of_uint8_signed (n AND 0x7F) OR 0x80 else integer_of_uint8_signed n)" unfolding integer_of_uint8_def integer_of_uint8_signed_def o_def including undefined_transfer integer.lifting by transfer(auto simp add: word_ao_nth uint_and_mask_or_full mask_numeral mask_Suc_0 intro!: uint_and_mask_or_full[symmetric]) 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 no_notation sshiftr_uint8 (infixl ">>>" 55) 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,8578 +1,8583 @@ (* * 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) 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 add: ucast_id) 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" apply (simp add: get_S_def) by (metis (mono_tags) ucast_id zero_neq_one) 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: -assumes a1: "low_equal s1 s2" -shows "get_WIM_bit (unat (((word_of_int ((uint (fst (fst (get_curr_win () s1))) - 1) mod NWINDOWS)))::word5)) - (cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = -get_WIM_bit (unat (((word_of_int ((uint (fst (fst (get_curr_win () s2))) -1) mod NWINDOWS)))::word5)) - (cpu_reg_val WIM (snd (fst (get_curr_win () s2))))" + \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 a1 have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" + 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 a1 have "(fst (fst (get_curr_win () s1))) = (fst (fst (get_curr_win () s2)))" + 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 get_WIM_bit_low_equal2: -assumes a1: "low_equal s1 s2" -shows "get_WIM_bit (unat (((word_of_int ((uint (fst (fst (get_curr_win () s1))) + 1) mod NWINDOWS)))::word5)) - (cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = -get_WIM_bit (unat (((word_of_int ((uint (fst (fst (get_curr_win () s2))) + 1) mod NWINDOWS)))::word5)) - (cpu_reg_val WIM (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 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 a1 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 - by metis + 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 ucast_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 ucast_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 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,910 +1,895 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Word Alignment" theory Aligned imports Word_Lib HOL_Lemmas More_Divides begin lift_definition is_aligned :: \'a::len word \ nat \ bool\ is \\k n. take_bit (min LENGTH('a) n) k = 0\ by (simp only: ac_simps flip: take_bit_take_bit) lemma is_aligned_iff_dvd_int: \is_aligned ptr n \ 2 ^ n dvd uint ptr\ by transfer (simp add: take_bit_eq_mod dvd_eq_mod_eq_0 mod_exp_eq) lemma is_aligned_iff_dvd_nat: \is_aligned ptr n \ 2 ^ n dvd unat ptr\ (is \?P \ ?Q\) 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_mask: \is_aligned w n \ w && mask n = 0\ by transfer (simp flip: take_bit_eq_mask) lemma is_aligned_to_bl: "is_aligned (w :: 'a :: len word) n = (True \ set (drop (size w - n) (to_bl w)))" apply (simp add: is_aligned_mask eq_zero_set_bl) apply (clarsimp simp: in_set_conv_nth word_size) apply (simp add: to_bl_nth word_size cong: conj_cong) apply (simp add: diff_diff_less) apply safe apply (case_tac "n \ LENGTH('a)") prefer 2 apply (rule_tac x=i in exI) apply clarsimp apply (subgoal_tac "\j < LENGTH('a). j < n \ LENGTH('a) - n + j = i") apply (erule exE) apply (rule_tac x=j in exI) apply clarsimp apply (thin_tac "w !! t" for t) apply (rule_tac x="i + n - LENGTH('a)" in exI) apply clarsimp apply arith apply (rule_tac x="LENGTH('a) - n + i" in exI) apply clarsimp apply arith done lemma unat_power_lower [simp]: assumes nv: "n < LENGTH('a::len)" shows "unat ((2::'a::len word) ^ n) = 2 ^ n" using assms by transfer simp lemma power_overflow: "n \ LENGTH('a) \ 2 ^ n = (0 :: 'a::len word)" by simp lemma is_alignedI [intro?]: fixes x::"'a::len word" assumes xv: "x = 2 ^ n * k" shows "is_aligned x n" proof cases assume nv: "n < LENGTH('a)" show ?thesis unfolding is_aligned_iff_dvd_nat proof (rule dvdI [where k = "unat k mod 2 ^ (LENGTH('a) - n)"]) from xv have "unat x = (unat (2::word32) ^ n * unat k) mod 2 ^ LENGTH('a)" using nv by (subst (asm) word_unat.Rep_inject [symmetric], simp, subst unat_word_ariths, simp) also have "\ = 2 ^ n * (unat k mod 2 ^ (LENGTH('a) - n))" using nv by (simp add: mult_mod_right power_add [symmetric] add_diff_inverse) finally show "unat x = 2 ^ n * (unat k mod 2 ^ (LENGTH('a) - n))" . qed next assume "\ n < LENGTH('a)" with xv show ?thesis by (simp add: not_less power_overflow is_aligned_iff_dvd_nat) 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 nat_power_less_diff: assumes lt: "(2::nat) ^ n * q < 2 ^ m" shows "q < 2 ^ (m - n)" using lt proof (induct n arbitrary: m) case 0 then show ?case by simp next case (Suc n) have ih: "\m. 2 ^ n * q < 2 ^ m \ q < 2 ^ (m - n)" and prem: "2 ^ Suc n * q < 2 ^ m" by fact+ show ?case proof (cases m) case 0 then show ?thesis using Suc by simp next case (Suc m') then show ?thesis using prem by (simp add: ac_simps ih) qed qed 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)" proof - from aligned obtain q where wv: "unat w = 2 ^ n * q" unfolding is_aligned_iff_dvd_nat .. show ?thesis proof (rule exI, intro conjI) show "q < 2 ^ (LENGTH('a) - n)" proof (rule nat_power_less_diff) have "unat w < 2 ^ size w" unfolding word_size .. then have "unat w < 2 ^ LENGTH('a)" by simp with wv show "2 ^ n * q < 2 ^ LENGTH('a)" by simp qed have r: "of_nat (2 ^ n) = (2::word32) ^ n" by (induct n) simp+ from wv have "of_nat (unat w) = of_nat (2 ^ n * q)" by simp then have "w = of_nat (2 ^ n * q)" by (subst word_unat.Rep_inverse [symmetric]) then show "w = 2 ^ n * (of_nat q)" by (simp add: r) qed qed lemma is_alignedE: "\is_aligned (w::'a::len word) n; \q. \w = 2 ^ n * (of_nat q); q < 2 ^ (LENGTH('a) - n)\ \ R\ \ R" by (auto dest: is_alignedE_pre) 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" proof - from nv have rl: "\q. q < 2 ^ (LENGTH('a) - n) \ to_bl (2 ^ n * (of_nat q :: 'a word)) = drop n (to_bl (of_nat q :: 'a word)) @ replicate n False" by (metis bl_shiftl le_antisym min_def shiftl_t2n wsst_TYs(3)) show ?thesis using aligned by (auto simp: rl elim: is_alignedE) qed 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 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" by (metis (no_types) \x = of_nat (2 ^ (m + k) * q1)\ l1 nat_mod_lem word_unat.inverse_norm zero_less_numeral zero_less_power) have "unat y = 2 ^ m * q2" by (metis (no_types) \y = of_nat (2 ^ m * q2)\ l2 nat_mod_lem word_unat.inverse_norm zero_less_numeral zero_less_power) 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 (k << m) 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 (k << m)" 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 (k << m) = unat (2 ^ m * k)" by (simp add: shiftl_t2n) also have "\ = (2 ^ m * unat k) mod (2 ^ LENGTH('a))" using mv by (subst unat_word_ariths(2))+ simp also have "\ = 2 ^ m * (unat k mod 2 ^ q)" by (subst mq [symmetric], subst power_add, subst mod_mult2_eq) simp finally show "unat (k << m) = 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 shiftl_zero_size 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 nat_add_offset_less: 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" and "0 < qy" by (auto dest: less_imp_add_positive) have "x * 2 ^ n + y < x * 2 ^ n + 2 ^ n" by simp fact+ 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: 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 clarsimp + then show ?thesis using off ptrq qv + by (simp add: take_bit_nat_less_exp) 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 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 (subst shiftl_1 [symmetric]) apply (subst bl_shiftl) apply (simp add: to_bl_1 min_def word_size) done 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 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)" proof - have x: "to_bl x = take (LENGTH('a)-Suc n) (to_bl x) @ drop (LENGTH('a)-Suc n) (to_bl x)" by simp show ?thesis apply simp apply (rule conjI) apply (clarsimp simp: word_size) apply (simp add: bl_word_xor to_bl_2p) apply (subst x) apply (subst zip_append) apply simp apply (simp add: map_zip_replicate_False_xor drop_minus) apply (auto simp add: word_size nth_w2p intro!: word_eqI) done qed lemma aligned_add_xor: assumes al: "is_aligned (x::'a::len word) n'" and le: "n < n'" shows "(x + 2^n) xor 2^n = x" proof cases assume "n' < LENGTH('a)" with assms show ?thesis apply - apply (rule word_bl.Rep_eqD) apply (subst xor_2p_to_bl) apply simp apply (subst is_aligned_add_conv, simp, simp add: word_less_nat_alt)+ apply (simp add: to_bl_2p nth_append) apply (cases "n' = Suc n") apply simp apply (subst is_aligned_replicate [where n="Suc n", simplified, symmetric]; simp) apply (subgoal_tac "\ LENGTH('a) - Suc n \ LENGTH('a) - n'") prefer 2 apply arith apply (subst replicate_Suc [symmetric]) apply (subst replicate_add [symmetric]) apply (simp add: is_aligned_replicate [simplified, symmetric]) done next assume "\ n' < LENGTH('a)" show ?thesis using al apply (rule is_alignedE) using \\ n' < LENGTH('a)\ by auto qed lemma is_aligned_0 [simp]: "is_aligned p 0" by (simp add: is_aligned_iff_dvd_nat) 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 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 unat_of_nat_len: "x < 2 ^ LENGTH('a) \ unat (of_nat x :: 'a::len word) = x" - by (simp add: word_size unat_of_nat) + by (simp add: take_bit_nat_eq_self_iff) 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 (subst unat_of_nat) - apply (rule mod_less_eq_dividend) + 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" apply (erule is_aligned_get_word_bits) apply (frule is_aligned_add_conv [rotated, where w=0]) apply (simp add: is_aligned_iff_dvd_nat) apply simp apply (drule is_aligned_replicateD) apply simp apply (clarsimp simp: word_size) apply (subst (asm) replicate_add [symmetric]) apply (drule arg_cong[where f="of_bl :: bool list \ 'a::len word"]) apply simp apply (simp only: replicate.simps[symmetric, where x=False] drop_replicate) done 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 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 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_nth: "is_aligned p m = (\n < m. \p !! n)" apply (clarsimp simp: is_aligned_mask bang_eq word_size) apply (rule iffI) apply clarsimp apply (case_tac "n < size p") apply (simp add: word_size) apply (drule test_bit_size) apply simp apply clarsimp 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" -proof cases - assume "n < LENGTH('a)" - with al - show ?thesis - apply (simp add: is_aligned_iff_dvd_nat dvd_eq_mod_eq_0 word_arith_nat_mod) - apply (erule of_nat_neq_0) - apply (rule order_less_trans) - apply (rule mod_less_divisor) - apply simp - apply simp - done -next - assume "\ n < LENGTH('a)" - with al - show ?thesis - by transfer simp -qed + 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 - apply (clarsimp) - apply (erule le_SucE) - apply (simp add: unat_of_nat) - apply (simp add: less_eq_Suc_le [symmetric] unat_of_nat) - done + 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 && ~~ (mask n)) m" by (metis and_not_mask is_aligned_shift is_aligned_weaken) 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 \ is_aligned (- p) n" apply (clarsimp simp: is_aligned_iff_dvd_nat unat_minus word_size word_neq_0_conv) apply (rule dvd_diff_nat, simp_all) apply (rule le_imp_power_dvd) apply (fold is_aligned_iff_dvd_nat) apply (erule_tac Q="0is_aligned (x :: 'a :: len word) n; \n' \ n. n' < LENGTH('a) \ \ p !! n'\ \ x + p && ~~ (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 apply (rule word_eqI) apply (clarsimp simp: word_size is_aligned_nth word_ops_nth_size le_def) apply blast done lemma is_aligned_andI1: "is_aligned x n \ is_aligned (x && y) n" by (simp add: is_aligned_nth) lemma is_aligned_andI2: "is_aligned y n \ is_aligned (x && y) n" by (simp add: is_aligned_nth) lemma is_aligned_shiftl: "is_aligned w (n - m) \ is_aligned (w << m) n" by (simp add: is_aligned_nth nth_shiftl) lemma is_aligned_shiftr: "is_aligned w (n + m) \ is_aligned (w >> m) n" by (simp add: is_aligned_nth nth_shiftr) lemma is_aligned_shiftl_self: "is_aligned (p << n) n" by (rule is_aligned_shift) lemma is_aligned_neg_mask_eq: "is_aligned p n \ p && ~~ (mask n) = p" by (metis add.left_neutral is_aligned_mask word_plus_and_or_coroll2) lemma is_aligned_shiftr_shiftl: "is_aligned w n \ w >> n << n = w" by (metis and_not_mask is_aligned_neg_mask_eq) lemma aligned_shiftr_mask_shiftl: "is_aligned x n \ ((x >> n) && mask v) << n = x && mask (v + n)" apply (rule word_eqI) apply (simp add: word_size nth_shiftl nth_shiftr) apply (subgoal_tac "\m. 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) done lemma mask_zero: "is_aligned x a \ x && mask a = 0" by (metis is_aligned_mask) lemma is_aligned_neg_mask_eq_concrete: "\ is_aligned p n; msk && ~~(mask n) = ~~(mask n) \ \ p && 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 && 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 && ~~(mask m) = p" using is_aligned_neg_mask_eq is_aligned_weaken by blast lemma is_aligned_neg_mask2[simp]: "is_aligned (a && ~~(mask n)) n" by (simp add: and_not_mask is_aligned_shift) 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,6174 +1,6157 @@ (* * 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 "HOL-Word.Misc_lsb" Word_EqI Word_Enum "HOL-Library.Sublist" begin lemmas is_aligned_def = is_aligned_iff_dvd_nat lemma word_plus_mono_left: fixes x :: "'a :: len word" shows "\y \ z; x \ x + z\ \ y + x \ z + x" by unat_arith 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 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 (subst (asm) bang_eq) (fastforce simp: nth_ucast word_size intro: word_eqI) 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" by (fastforce intro: word_eqI simp: bang_eq nth_ucast word_size) lemma ucast_0_I: "x = 0 \ ucast x = 0" by simp 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 same_length_is_parallel: assumes len: "\y \ set as. length y = x" shows "\x \ set as. \y \ set as - {x}. x \ y" proof (rule, rule) fix x y assume xi: "x \ set as" and yi: "y \ set as - {x}" from len obtain q where len': "\y \ set as. length y = q" .. show "x \ y" proof (rule not_equal_is_parallel) from xi yi show "x \ y" by auto from xi yi len' show "length x = length y" by (auto dest: bspec) qed qed text \Lemmas about words\ lemmas and_bang = word_and_nth lemma of_drop_to_bl: "of_bl (drop n (to_bl x)) = (x && mask (size x - n))" by (simp add: of_bl_drop word_size_bl) 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 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_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_unat.Rep_eqD) 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 set_enum_word8_def: "(set enum::word8 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 eval lemma set_strip_insert: "\ x \ insert a S; x \ a \ \ x \ S" by simp lemma word8_exhaust: fixes x :: word8 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 lemma upto_enum_red': assumes lt: "1 \ X" shows "[(0::'a :: len word) .e. X - 1] = map of_nat [0 ..< unat X]" proof - have lt': "unat X < 2 ^ LENGTH('a)" by (rule unat_lt2p) show ?thesis apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_unat_diff_1 [OF lt]) apply (rule map_cong [OF refl]) apply (rule toEnum_of_nat) apply simp apply (erule order_less_trans [OF _ lt']) done qed lemma upto_enum_red2: assumes szv: "sz < LENGTH('a :: len)" shows "[(0:: 'a :: len word) .e. 2 ^ sz - 1] = map of_nat [0 ..< 2 ^ sz]" using szv apply (subst unat_power_lower[OF szv, symmetric]) apply (rule upto_enum_red') apply (subst word_le_nat_alt, simp) done lemma upto_enum_step_red: assumes szv: "sz < LENGTH('a)" and usszv: "us \ sz" shows "[0 :: 'a :: len word , 2 ^ us .e. 2 ^ sz - 1] = map (\x. of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" using szv unfolding upto_enum_step_def apply (subst if_not_P) apply (rule leD) apply (subst word_le_nat_alt) apply (subst unat_minus_one) apply simp apply simp apply simp apply (subst upto_enum_red) apply (simp del: upt.simps) apply (subst Suc_div_unat_helper [where 'a = 'a, OF szv usszv, symmetric]) apply clarsimp apply (subst toEnum_of_nat) apply (erule order_less_trans) using szv apply simp apply simp done lemma upto_enum_word: "[x .e. y] = map of_nat [unat x ..< Suc (unat y)]" apply (subst upto_enum_red) apply clarsimp apply (subst toEnum_of_nat) prefer 2 apply (rule refl) apply (erule disjE, simp) apply clarsimp apply (erule order_less_trans) apply simp done lemma word_upto_Cons_eq: "x < y \ [x::'a::len word .e. y] = x # [x + 1 .e. y]" apply (subst upto_enum_red) apply (subst upt_conv_Cons, unat_arith) apply (simp only: list.map list.inject upto_enum_red to_from_enum simp_thms) apply (rule map_cong[OF _ refl]) apply (rule arg_cong2[where f = "\x y. [x ..< y]"], unat_arith) apply (rule refl) done lemma distinct_enum_upto: "distinct [(0 :: 'a::len word) .e. b]" proof - have "\(b::'a word). [0 .e. b] = nths enum {..< Suc (fromEnum b)}" apply (subst upto_enum_red) apply (subst nths_upt_eq_take) apply (subst enum_word_def) apply (subst take_map) apply (subst take_upt) apply (simp only: add_0 fromEnum_unat) apply (rule order_trans [OF _ order_eq_refl]) apply (rule Suc_leI [OF unat_lt2p]) apply simp apply clarsimp apply (rule toEnum_of_nat) apply (erule order_less_trans [OF _ unat_lt2p]) done then show ?thesis by (rule ssubst) (rule distinct_nthsI, simp) qed lemma upto_enum_set_conv [simp]: fixes a :: "'a :: len word" shows "set [a .e. b] = {x. a \ x \ x \ b}" apply (subst upto_enum_red) apply (subst set_map) apply safe apply simp apply clarsimp apply (erule disjE) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply clarsimp apply (erule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply simp apply (erule iffD2 [OF word_le_nat_alt]) apply simp apply clarsimp apply (erule disjE) apply simp apply clarsimp apply (rule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]]) apply assumption apply simp apply (erule order_less_imp_le [OF iffD2 [OF word_less_nat_alt]]) apply clarsimp apply (rule_tac x="fromEnum x" in image_eqI) apply clarsimp apply clarsimp apply (rule conjI) apply (subst word_le_nat_alt [symmetric]) apply simp apply safe apply (simp add: word_le_nat_alt [symmetric]) apply (simp add: word_less_nat_alt [symmetric]) done lemma upto_enum_less: assumes xin: "x \ set [(a::'a::len word).e.2 ^ n - 1]" and nv: "n < LENGTH('a::len)" shows "x < 2 ^ n" proof (cases n) case 0 then show ?thesis using xin by simp next case (Suc m) show ?thesis using xin nv by simp qed lemma upto_enum_len_less: "\ n \ length [a, b .e. c]; n \ 0 \ \ a \ c" unfolding upto_enum_step_def by (simp split: if_split_asm) lemma length_upto_enum_step: fixes x :: "'a :: len word" shows "x \ z \ length [x , y .e. z] = (unat ((z - x) div (y - x))) + 1" unfolding upto_enum_step_def by (simp add: upto_enum_red) lemma map_length_unfold_one: fixes x :: "'a::len word" assumes xv: "Suc (unat x) < 2 ^ LENGTH('a)" and ax: "a < x" shows "map f [a .e. x] = f a # map f [a + 1 .e. x]" by (subst word_upto_Cons_eq, auto, fact+) lemma upto_enum_set_conv2: fixes a :: "'a::len word" shows "set [a .e. b] = {a .. b}" by auto 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 mask_shift: "(x && ~~ (mask y)) >> y = x >> y" by word_eqI 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 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_unat.Rep_eqD) 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 (case_tac "size w \ n", clarsimp simp: shiftl_zero_size) (clarsimp simp: not_le ucast_bl bl_shiftl bang_eq test_bit_of_bl rev_nth nth_append) 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 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 length_upto_enum [simp]: fixes a :: "'a :: len word" shows "length [a .e. b] = Suc (unat b) - unat a" apply (simp add: word_le_nat_alt upto_enum_red) apply (clarsimp simp: Suc_diff_le) done lemma length_upto_enum_cases: fixes a :: "'a::len word" shows "length [a .e. b] = (if a \ b then Suc (unat b) - unat a else 0)" apply (case_tac "a \ b") apply (clarsimp) apply (clarsimp simp: upto_enum_def) apply unat_arith done lemma length_upto_enum_less_one: "\a \ b; b \ 0\ \ length [a .e. b - 1] = unat (b - a)" apply clarsimp apply (subst unat_sub[symmetric], assumption) apply clarsimp done lemma drop_upto_enum: "drop (unat n) [0 .e. m] = [n .e. m]" apply (clarsimp simp: upto_enum_def) apply (induct m, simp) by (metis drop_map drop_upt plus_nat.add_0) lemma distinct_enum_upto' [simp]: "distinct [a::'a::len word .e. b]" apply (subst drop_upto_enum [symmetric]) apply (rule distinct_drop) apply (rule distinct_enum_upto) done lemma length_interval: "\set xs = {x. (a::'a::len word) \ x \ x \ b}; distinct xs\ \ length xs = Suc (unat b) - unat a" apply (frule distinct_card) apply (subgoal_tac "set xs = set [a .e. b]") apply (cut_tac distinct_card [where xs="[a .e. b]"]) apply (subst (asm) length_upto_enum) apply clarsimp apply (rule distinct_enum_upto') apply simp done 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 declare of_nat_power [simp del] (* TODO: move to word *) 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 word_unat_power 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 is_aligned_0'[simp]: "is_aligned 0 n" by (simp add: is_aligned_def) lemma p_assoc_help: fixes p :: "'a::{ring,power,numeral,one}" shows "p + 2^sz - 1 = p + (2^sz - 1)" by simp 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_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_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 add: word_unat.Rep_inject [symmetric]) apply simp done 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 int_not_emptyD: "A \ B \ {} \ \x. x \ A \ x \ B" by (erule contrapos_np, clarsimp simp: disjoint_iff_not_equal) 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 - assume szv: "sz < LENGTH('a::len)" - show ?thesis - proof (cases "sz = 0") - case True - then show ?thesis using kv szv - by (simp add: unat_of_nat) - next - case False - then have sne: "0 < sz" .. - - have uk: "unat (of_nat k :: 'a word) = k" - apply (subst unat_of_nat) - apply (simp add: nat_mod_eq less_trans[OF kv] sne) - done - - show ?thesis using szv - apply (subst iffD1 [OF unat_mult_lem]) - apply (simp add: uk nat_less_power_trans[OF kv order_less_imp_le [OF szv]])+ - done - qed +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 - assume "\ sz < LENGTH('a)" - with kv show ?thesis by (simp add: not_less power_overflow) + case False + with assms show ?thesis + by simp qed 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 - apply (erule is_alignedE) apply (subst word_unat.Rep_inject [symmetric]) apply (subst unat_mod) apply (subst iffD1 [OF unat_add_lem], rule ux) apply simp apply (subst unat_mult_power_lem, assumption+) apply (simp) apply (rule mod_less[OF less_le_trans[OF unat_mono], OF kv]) apply (erule eq_imp_le[OF unat_power_lower]) 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 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 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 (simp add: word_size_bl) apply (erule iffD1 [OF word_less_sub_le[OF 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 less_two_pow_divD: "\ (x :: nat) < 2 ^ n div 2 ^ m \ \ n \ m \ (x < 2 ^ (n - m))" apply (rule context_conjI) apply (rule ccontr) apply (simp add: power_strict_increasing) apply (simp add: power_sub) done lemma less_two_pow_divI: "\ (x :: nat) < 2 ^ (n - m); m \ n \ \ x < 2 ^ n div 2 ^ m" by (simp add: power_sub) 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 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) (* 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 (simp add: ucast_eq test_bit_bin word_ubin.eq_norm nth_bintr word_size) - (fast elim!: bin_nth_uint_imp) + by transfer (simp add: bit_take_bit_iff ac_simps) lemma ucast_less: "LENGTH('b) < LENGTH('a) \ (ucast (x :: 'b :: len word) :: ('a :: len word)) < 2 ^ LENGTH('b)" by (meson Word.nth_ucast test_bit_conj_lt le_def upper_bits_unset_is_l2p) 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) by word_eqI_solve 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 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 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)" +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: word_unat_power) 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)" - apply (simp add: word_less_nat_alt) - apply (subst unat_of_nat) - apply (subst mod_less) - apply (erule order_less_le_trans) - apply simp+ + 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) - apply (subst (asm) unat_of_nat) + apply (simp add: word_less_nat_alt take_bit_eq_mod) apply (subst (asm) mod_less) - apply (rule order_le_less_trans) - apply (rule diff_le_self) - apply (erule order_less_le_trans) + apply auto + apply (rule order_le_less_trans) + apply (rule diff_le_self) + apply (erule order_less_le_trans) apply simp - apply assumption done - ultimately show ?thesis by auto + ultimately show ?thesis + by auto qed lemma word_less_sub_1: "x < (y :: 'a :: len word) \ x \ y - 1" - apply (erule udvd_minus_le') - apply (simp add: udvd_def)+ - done + 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 sublist_equal_part: "prefix xs ys \ take (length xs) ys = xs" by (clarsimp simp: prefix_def) 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 prefix_length_less: "strict_prefix xs ys \ length xs < length ys" apply (clarsimp simp: strict_prefix_def) apply (frule prefix_length_le) apply (rule ccontr, simp) apply (clarsimp simp: prefix_def) done lemmas take_less = take_strict_prefix lemma not_prefix_longer: "\ length xs > length ys \ \ \ prefix xs ys" by (clarsimp dest!: prefix_length_le) 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 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 (subst unat_of_nat_eq[where x=n, symmetric], simp+) + by transfer + (auto simp add: take_bit_of_nat nat_eq_iff take_bit_nat_eq_self_iff intro: sym) 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: unat_of_nat) + 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 (drule unat_of_nat_eq, simp) + by transfer (simp add: take_bit_eq_mod) 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 of_nat_inj: "\x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\ \ (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)" by (simp add: word_unat.norm_eq_iff [symmetric]) lemma map_prefixI: "prefix xs ys \ prefix (map f xs) (map f ys)" by (clarsimp simp: prefix_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 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 list_all2_induct_suffixeq [consumes 1, case_names Nil Cons]: assumes lall: "list_all2 Q as bs" and nilr: "P [] []" and consr: "\x xs y ys. \list_all2 Q xs ys; Q x y; P xs ys; suffix (x # xs) as; suffix (y # ys) bs\ \ P (x # xs) (y # ys)" shows "P as bs" proof - define as' where "as' == as" define bs' where "bs' == bs" have "suffix as as' \ suffix bs bs'" unfolding as'_def bs'_def by simp then show ?thesis using lall proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]]) case 1 show ?case by 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" using "2.hyps" "2.prems" by (auto dest: suffix_ConsD) from "2.prems" show "suffix (x # xs) as" and "suffix (y # ys) bs" by (auto simp: as'_def bs'_def) qed qed qed 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 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 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 shiftr_less_t2n': "\ x && 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]) apply word_eqI apply (erule_tac x="na + n" in allE) apply fastforce 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 shiftr_not_mask_0: "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) && ~~ (mask m) = 0" apply (simp add: and_not_mask shiftr_less_t2n shiftr_shiftr) apply (subgoal_tac "w >> n + m = 0", simp) apply (simp add: le_mask_iff[symmetric] mask_eq_decr_exp le_def) apply (subst (asm) p2_gt_0[symmetric]) apply (simp add: power_add not_less) 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]) apply word_eqI apply (erule_tac x="na - n" in allE) apply auto 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 ucast_ucast_mask: "(ucast :: 'a :: len word \ 'b :: len word) (ucast x) = x && mask (len_of TYPE ('a))" by word_eqI 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('b \ 'a) x) = 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 sints_subset: "m \ n \ sints m \ sints n" apply (simp add: sints_num) apply clarsimp apply (rule conjI) apply (erule order_trans[rotated]) apply simp apply (erule order_less_le_trans) apply simp done lemma up_scast_inj: "\ scast x = (scast y :: 'b :: len word); size x \ LENGTH('b) \ \ x = y" - apply (simp add: scast_eq) + apply (unfold scast_eq) apply (subst(asm) word_sint.Abs_inject) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply (erule subsetD [OF sints_subset]) apply (simp add: word_size) apply simp 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 nth_bounded: "\(x :: 'a :: len word) !! n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" apply (frule test_bit_size) apply (clarsimp simp: test_bit_bl word_size) apply (simp add: rev_nth) apply (subst(asm) is_aligned_add_conv[OF is_aligned_0', simplified add_0_left, rotated]) apply assumption+ apply (simp only: to_bl_0) apply (simp add: nth_append split: if_split_asm) done lemma is_aligned_add_or: "\is_aligned p n; d < 2 ^ n\ \ p + d = p || d" by (rule word_plus_and_or_coroll, word_eqI) blast lemma two_power_increasing: "\ n \ m; m < LENGTH('a) \ \ (2 :: 'a :: len word) ^ n \ 2 ^ m" by (simp add: word_le_nat_alt) 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 lemmas mask_inner_mask = mask_eqs(1) lemma mask_add_aligned: "is_aligned p n \ (p + q) && mask n = q && mask n" apply (simp add: is_aligned_mask) apply (subst mask_inner_mask [symmetric]) apply simp done lemma take_prefix: "(take (length xs) ys = xs) = prefix xs ys" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case Cons then show ?case by (cases ys) auto qed lemma cart_singleton_empty: "(S \ {e} = {}) = (S = {})" by blast lemma word_div_1: "(n :: 'a :: len word) div 1 = n" by (simp add: word_div_def) lemma word_minus_one_le: "-1 \ (x :: 'a :: len word) = (x = -1)" apply (insert word_n1_ge[where y=x]) apply safe apply (erule(1) order_antisym) done lemma mask_out_sub_mask: "(x && ~~ (mask n)) = x - (x && (mask n))" by (simp add: field_simps word_plus_and_or_coroll2) 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 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 shiftl_mask_is_0[simp]: "(x << n) && mask n = 0" apply (rule iffD1 [OF is_aligned_mask]) apply (rule is_aligned_shiftl_self) done 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_and_helper: "(If x v v') && v'' = If x (v && v'') (v' && v'')" by (rule if_distrib) 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 lemmas word_unat_Rep_inject1 = word_unat.Rep_inject[where y=1] lemmas unat_eq_1 = unat_eq_0 word_unat_Rep_inject1[simplified] lemma rshift_sub_mask_eq: "(a >> (size a - b)) && mask b = a >> (size a - b)" 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) && mask (size a - c)" apply (cases "b = c") apply (simp add: shiftl_shiftr1) apply (simp add: shiftl_shiftr2) done lemma and_mask_shiftr_comm: "m\size w \ (w && mask m) >> n = (w >> n) && mask (m-n)" by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) lemma and_mask_shiftl_comm: "m+n \ size w \ (w && mask m) << n = (w << n) && mask (m+n)" 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 and_not_mask_twice: "(w && ~~ (mask n)) && ~~ (mask m) = w && ~~ (mask (max m n))" apply (simp add: and_not_mask) apply (case_tac "n 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 eq_eqI: "a = b \ (a = x) = (b = x)" by simp lemma mask_and_mask: "mask a && mask b = mask (min a b)" by word_eqI lemma mask_eq_0_eq_x: "(x && w = 0) = (x && ~~ w = x)" using word_plus_and_or_coroll2[where x=x and w=w] by auto lemma mask_eq_x_eq_0: "(x && w = x) = (x && ~~ w = 0)" using word_plus_and_or_coroll2[where x=x and w=w] by auto definition "limited_and (x :: 'a :: len word) y = (x && y = x)" lemma limited_and_eq_0: "\ limited_and x z; y && ~~ z = y \ \ x && y = 0" unfolding limited_and_def apply (subst arg_cong2[where f="(&&)"]) 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 && z = z \ \ x && 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 (simp add: shiftl_over_and_dist[symmetric]) lemma rshift_limited_and: "limited_and x z \ limited_and (x >> n) (z >> n)" unfolding limited_and_def by (simp add: shiftr_over_and_dist[symmetric]) 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] lemma compl_of_1: "~~ 1 = (-2 :: 'a :: len word)" by simp 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] compl_of_1 shiftl_shiftr1[unfolded word_size mask_eq_decr_exp] shiftl_shiftr2[unfolded word_size mask_eq_decr_exp] lemma split_word_eq_on_mask: "(x = y) = (x && m = y && m \ x && ~~ m = y && ~~ m)" by safe word_eqI_solve 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 word_and_1_shiftl: "x && (1 << n) = (if x !! n then (1 << n) else 0)" for x :: "'a :: len word" by word_eqI_solve 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 && (mask n << m) = ((x >> m) && mask n) << m" by word_eqI_solve 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 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 (rule sym, subst word_unat.inverse_norm) - apply (simp add: ucast_eq word_of_int[symmetric] - of_nat_nat[symmetric]) - apply (simp add: unat_of_nat) + apply (simp add: ucast_eq of_nat_nat[symmetric] take_bit_eq_mod) 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 range_subset_eq2: "{a :: 'a :: len word .. b} \ {} \ ({a .. b} \ {c .. d}) = (c \ a \ b \ d)" by simp 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 mod_mod_power: fixes k :: nat shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" proof (cases "m \ n") case True then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ m" apply - apply (subst mod_less [where n = "2 ^ n"]) apply (rule order_less_le_trans [OF mod_less_divisor]) apply simp+ done also have "\ = k mod 2 ^ (min m n)" using True by simp finally show ?thesis . next case False then have "n < m" by simp then obtain d where md: "m = n + d" by (auto dest: less_imp_add_positive) then have "k mod 2 ^ m = 2 ^ n * (k div 2 ^ n mod 2 ^ d) + k mod 2 ^ n" by (simp add: mod_mult2_eq power_add) then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ n" by (simp add: mod_add_left_eq) then show ?thesis using False by simp qed 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 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 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 power_mod_div: fixes x :: "nat" shows "x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" (is "?LHS = ?RHS") proof (cases "n \ m") case True then have "?LHS = 0" apply - apply (rule div_less) apply (rule order_less_le_trans [OF mod_less_divisor]; simp) done also have "\ = ?RHS" using True by simp finally show ?thesis . next case False then have lt: "m < n" by simp then obtain q where nv: "n = m + q" and "0 < q" by (auto dest: less_imp_Suc_add) then have "x mod 2 ^ n = 2 ^ m * (x div 2 ^ m mod 2 ^ q) + x mod 2 ^ m" by (simp add: power_add mod_mult2_eq) then have "?LHS = x div 2 ^ m mod 2 ^ q" by (simp add: div_add1_eq) also have "\ = ?RHS" using nv by simp finally show ?thesis . qed 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 lemmas if_fun_split = if_apply_def2 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 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 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 auto lemma emptyE2: "\ S = {}; x \ S \ \ P" by simp lemma mod_div_equality_div_eq: "a div b * b = (a - (a mod b) :: int)" by (simp add: field_simps) lemma zmod_helper: "n mod m = k \ ((n :: int) + a) mod m = (k + a) mod m" by (metis add.commute mod_add_right_eq) lemma int_div_sub_1: "\ m \ 1 \ \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)" apply (subgoal_tac "m = 0 \ (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)") apply fastforce apply (subst mult_cancel_right[symmetric]) apply (simp only: left_diff_distrib split: if_split) apply (simp only: mod_div_equality_div_eq) apply (clarsimp simp: field_simps) apply (clarsimp simp: dvd_eq_mod_eq_0) apply (cases "m = 1") apply simp apply (subst mod_diff_eq[symmetric], simp add: zmod_minus1) apply clarsimp apply (subst diff_add_cancel[where b=1, symmetric]) apply (subst mod_add_eq[symmetric]) apply (simp add: field_simps) apply (rule mod_pos_pos_trivial) apply (subst add_0_right[where a=0, symmetric]) apply (rule add_mono) apply simp apply simp apply (cases "(n - 1) mod m = m - 1") apply (drule zmod_helper[where a=1]) apply simp apply (subgoal_tac "1 + (n - 1) mod m \ m") apply simp apply (subst field_simps, rule zless_imp_add1_zle) apply simp done 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" apply (simp add: image_def) apply (erule rev_bexI) apply (rule arg_cong[where f="ptr_add ptr"]) apply simp 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 word_plus_strict_mono_right: fixes x :: "'a :: len word" shows "\y < z; x \ x + z\ \ x + y < x + z" by unat_arith lemma replicate_minus: "k < n \ replicate n False = replicate (n - k) False @ replicate k False" by (subst replicate_add [symmetric]) simp 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 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 add: word_less_nat_alt) - apply (subst unat_of_nat_eq) - apply (erule order_less_trans) - apply simp+ - done + 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 enum_word_div: fixes v :: "'a :: len word" shows "\xs ys. enum = xs @ [v] @ ys \ (\x \ set xs. x < v) \ (\y \ set ys. v < y)" apply (simp only: enum_word_def) apply (subst upt_add_eq_append'[where j="unat v"]) apply simp apply (rule order_less_imp_le, simp) apply (simp add: upt_conv_Cons) apply (intro exI conjI) apply fastforce apply clarsimp apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp apply (clarsimp simp: Suc_le_eq) apply (drule of_nat_mono_maybe[rotated, where 'a='a]) apply simp apply simp done lemma of_bool_nth: "of_bool (x !! v) = (x >> v) && 1" by (simp add: test_bit_word_eq shiftr_word_eq bit_eq_iff) (auto simp add: bit_1_iff bit_and_iff bit_drop_bit_eq intro: ccontr) 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) 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 && 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 shiftr_mask_eq: "(x >> n) && mask (size x - n) = x >> n" for x :: "'a :: len word" by word_eqI_solve lemma shiftr_mask_eq': "m = (size x - n) \ (x >> n) && mask m = x >> n" for x :: "'a :: len word" by (simp add: shiftr_mask_eq) 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) 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 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 (simp add: word_unat.norm_eq_iff [symmetric]) done lemma mask_AND_NOT_mask: "(w && ~~ (mask n)) && mask n = 0" by word_eqI lemma AND_NOT_mask_plus_AND_mask_eq: "(w && ~~ (mask n)) + (w && mask n) = w" by (subst word_plus_and_or_coroll; word_eqI_solve) lemma mask_eqI: fixes x :: "'a :: len word" assumes m1: "x && mask n = y && mask n" and m2: "x && ~~ (mask n) = y && ~~ (mask n)" shows "x = y" proof (subst bang_eq, rule allI) fix m show "x !! m = y !! m" proof (cases "m < n") case True then have "x !! m = ((x && mask n) !! m)" by (simp add: word_size test_bit_conj_lt) also have "\ = ((y && mask n) !! m)" using m1 by simp also have "\ = y !! m" using True by (simp add: word_size test_bit_conj_lt) finally show ?thesis . next case False then have "x !! m = ((x && ~~ (mask n)) !! m)" by (simp add: neg_mask_test_bit test_bit_conj_lt) also have "\ = ((y && ~~ (mask n)) !! m)" using m2 by simp also have "\ = y !! m" using False by (simp add: neg_mask_test_bit test_bit_conj_lt) finally show ?thesis . qed qed 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 neq_0_no_wrap: fixes x :: "'a :: len word" shows "\ x \ x + y; x \ 0 \ \ x + y \ 0" by clarsimp lemma plus_minus_one_rewrite: "v + (- 1 :: ('a :: {ring, one, uminus})) \ v - 1" by (simp add: field_simps) lemma power_minus_is_div: "b \ a \ (2 :: nat) ^ (a - b) = 2 ^ a div 2 ^ b" apply (induct a arbitrary: b) apply simp apply (erule le_SucE) apply (clarsimp simp:Suc_diff_le le_iff_add power_add) apply simp done lemma two_pow_div_gt_le: "v < 2 ^ n div (2 ^ m :: nat) \ m \ n" by (clarsimp dest!: less_two_pow_divD) 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_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: unat_of_nat) + apply (simp add: take_bit_eq_mod) done 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 + 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 and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x && y) = 0) = (\ (x !! n))" apply safe apply (drule_tac u="(x && (1 << n))" and x=n in word_eqD) apply (simp add: nth_w2p) apply (simp add: test_bit_bin) apply word_eqI done lemmas arg_cong_Not = arg_cong [where f=Not] lemmas and_neq_0_is_nth = arg_cong_Not [OF and_eq_0_is_nth, simplified] lemma nth_is_and_neq_0: "(x::'a::len word) !! n = (x && 2 ^ n \ 0)" by (subst and_neq_0_is_nth; rule refl) lemma mask_Suc_0 : "mask (Suc 0) = (1 :: 'a::len word)" by (simp add: mask_eq_decr_exp) 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 (rule word_unat.Rep_eqD) apply (simp add: unat_ucast unat_word_ariths mod_mod_power min.absorb2 unat_of_nat) apply (subst mod_add_left_eq[symmetric]) apply (simp add: mod_mod_power min.absorb2) apply (subst mod_add_right_eq) apply simp done 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 bool_mask': fixes x :: "'a :: len word" shows "2 < LENGTH('a) \ (0 < x && 1) = (x && 1 = 1)" by (simp add: and_one_eq mod_2_eq_odd) lemma sint_eq_uint: "\ msb x \ sint x = uint x" apply (rule word_uint.Abs_eqD, subst word_sint.Rep_inverse) apply simp_all apply (cut_tac x=x in word_sint.Rep) apply (clarsimp simp add: uints_num sints_num) apply (rule conjI) apply (rule ccontr) apply (simp add: linorder_not_le word_msb_sint[symmetric]) apply (erule order_less_le_trans) apply simp done lemma scast_eq_ucast: "\ msb x \ scast x = ucast x" - by (simp add: scast_eq ucast_eq sint_eq_uint) + 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 test_bit_word_eq) + 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 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 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 lemma word_0_sle_from_less[unfolded word_size]: "\ x < 2 ^ (size x - 1) \ \ 0 <=s x" apply (clarsimp simp: word_sle_msb_le) apply (simp add: word_msb_nth) apply (subst (asm) word_test_bit_def [symmetric]) apply (drule less_mask_eq) apply (drule_tac x="size x - 1" in word_eqD) apply (simp add: word_size) done 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 done lemma distinct_lemma: "f x \ f y \ x \ y" by auto 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" - by unat_arith + 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) (* * Basic signed arithemetic properties. *) lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)" by (metis sint_n1 word_sint.Rep_inverse') lemma sint_0 [simp]: "(sint x = 0) = (x = 0)" by (metis sint_0 word_sint.Rep_inverse') (* 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: - "\ \ 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 \ - \ P" - apply atomize_elim - apply (case_tac "len_of TYPE ('a) = 1") - apply clarsimp - apply (subgoal_tac "(UNIV :: 'a word set) = {0, 1}") - apply (metis UNIV_I insert_iff singletonE) - apply (subst word_unat.univ) - apply (clarsimp simp: unats_def image_def) - apply (rule set_eqI, rule iffI) - apply clarsimp - apply (metis One_nat_def less_2_cases of_nat_1 semiring_1_class.of_nat_0) - apply clarsimp - apply (metis Abs_fnat_hom_0 Suc_1 lessI of_nat_1 zero_less_Suc) - apply clarsimp - apply (metis One_nat_def arith_is_1 le_def len_gt_0) - done + 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 (subst word_sint.Abs_inverse' [where r="- (2 ^ (LENGTH('a) - Suc 0))"]) apply (clarsimp simp: sints_num) apply (clarsimp simp: wi_hom_syms word_of_int_2p) apply clarsimp 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 sbintrunc_eq_in_range: "(sbintrunc n x = x) = (x \ range (sbintrunc n))" "(x = sbintrunc n x) = (x \ range (sbintrunc n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) \ sbintrunc 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 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) (* Basic proofs that signed word div/mod operations are * truncations of their integer counterparts. *) lemma signed_div_arith: "sint ((a::('a::len) word) sdiv b) = sbintrunc (LENGTH('a) - 1) (sint a sdiv sint b)" apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: sdiv_int_def sdiv_word_def) - apply (metis word_ubin.eq_norm) + apply transfer + apply simp done lemma signed_mod_arith: "sint ((a::('a::len) word) smod b) = sbintrunc (LENGTH('a) - 1) (sint a smod sint b)" + apply (subst word_sbin.norm_Rep [symmetric]) apply (subst bin_sbin_eq_iff' [symmetric]) apply simp apply (subst uint_sint [symmetric]) apply (clarsimp simp: smod_int_def smod_word_def) - apply (metis word_ubin.eq_norm) - done + using word_ubin.inverse_norm by force (* Signed word arithmetic overflow constraints. *) 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 sbintrunc_eq_in_range range_sbintrunc) 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)+ 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: "sbintrunc (size a - 1) (sint a * sint b) \ range (sbintrunc (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) - show ?thesis + 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 (simp add: sint_word_ariths scast_eq) - apply (simp add: wi_hom_mult) - apply (subst word_sint.Abs_inject, simp_all) - apply (simp add: sints_def range_sbintrunc abs_less_iff) apply clarsimp - apply (simp add: sints_def range_sbintrunc word_size) - apply (auto elim: order_less_le_trans order_trans[rotated]) + apply (transfer fixing: r s) + apply (auto simp add: signed_take_bit_int_eq_self simp flip: signed_take_bit_eq_iff_take_bit_eq) done qed (* Properties about signed division. *) lemma int_sdiv_simps [simp]: "(a :: int) sdiv 1 = a" "(a :: int) sdiv 0 = 0" "(a :: int) sdiv -1 = -a" apply (auto simp: sdiv_int_def sgn_if) done lemma sgn_div_eq_sgn_mult: "a div b \ 0 \ sgn ((a :: int) div b) = sgn (a * b)" apply (clarsimp simp: sgn_if zero_le_mult_iff neg_imp_zdiv_nonneg_iff not_less) apply (metis less_le mult_le_0_iff neg_imp_zdiv_neg_iff not_less pos_imp_zdiv_neg_iff zdiv_eq_0_iff) done lemma sgn_sdiv_eq_sgn_mult: "a sdiv b \ 0 \ sgn ((a :: int) sdiv b) = sgn (a * b)" by (auto simp: sdiv_int_def sgn_div_eq_sgn_mult sgn_mult) lemma int_sdiv_same_is_1 [simp]: "a \ 0 \ ((a :: int) sdiv b = a) = (b = 1)" apply (rule iffI) apply (clarsimp simp: sdiv_int_def) apply (subgoal_tac "b > 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if) apply (clarsimp simp: algebra_split_simps not_less) apply (metis int_div_same_is_1 le_neq_trans minus_minus neg_0_le_iff_le neg_equal_0_iff_equal) apply (case_tac "a > 0") apply (case_tac "b = 0") apply clarsimp apply (rule classical) apply (clarsimp simp: sgn_mult not_less) apply (metis le_less neg_0_less_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (rule classical) apply (clarsimp simp: algebra_split_simps sgn_mult not_less sgn_if split: if_splits) apply (metis antisym less_le neg_imp_zdiv_nonneg_iff) apply (clarsimp simp: sdiv_int_def sgn_if) done lemma int_sdiv_negated_is_minus1 [simp]: "a \ 0 \ ((a :: int) sdiv b = - a) = (b = -1)" apply (clarsimp simp: sdiv_int_def) apply (rule iffI) apply (subgoal_tac "b < 0") apply (case_tac "a > 0") apply (clarsimp simp: sgn_if algebra_split_simps not_less) apply (case_tac "sgn (a * b) = -1") apply (clarsimp simp: not_less algebra_split_simps) apply (clarsimp simp: algebra_split_simps not_less) apply (rule classical) apply (case_tac "b = 0") apply (clarsimp simp: not_less sgn_mult) apply (case_tac "a > 0") apply (clarsimp simp: not_less sgn_mult) apply (metis less_le neg_less_0_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff) apply (clarsimp simp: not_less sgn_mult) apply (metis antisym_conv div_minus_right neg_imp_zdiv_nonneg_iff neg_le_0_iff_le not_less) apply (clarsimp simp: sgn_if) done lemma sdiv_int_range: "(a :: int) sdiv b \ { - (abs a) .. (abs a) }" apply (unfold sdiv_int_def) apply (subgoal_tac "(abs a) div (abs b) \ (abs a)") apply (clarsimp simp: sgn_if) apply (meson abs_ge_zero neg_le_0_iff_le nonneg_mod_div order_trans) apply (metis abs_eq_0 abs_ge_zero div_by_0 zdiv_le_dividend zero_less_abs_iff) done lemma word_sdiv_div1 [simp]: "(a :: ('a::len) word) sdiv 1 = a" apply (rule sint_1_cases [where a=a]) apply (clarsimp simp: sdiv_word_def sdiv_int_def) apply (clarsimp simp: sdiv_word_def sdiv_int_def simp del: sint_minus1) apply (clarsimp simp: sdiv_word_def) done lemma sdiv_int_div_0 [simp]: "(x :: int) sdiv 0 = 0" by (clarsimp simp: sdiv_int_def) lemma sdiv_int_0_div [simp]: "0 sdiv (x :: int) = 0" by (clarsimp simp: sdiv_int_def) lemma word_sdiv_div0 [simp]: "(a :: ('a::len) word) sdiv 0 = 0" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) done lemma word_sdiv_div_minus1 [simp]: "(a :: ('a::len) word) sdiv -1 = -a" apply (auto simp: sdiv_word_def sdiv_int_def sgn_if) - apply (metis wi_hom_neg word_sint.Rep_inverse') done lemmas word_sdiv_0 = word_sdiv_div0 lemma sdiv_word_min: "- (2 ^ (size a - 1)) \ sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word)" apply (clarsimp simp: word_size) apply (cut_tac sint_range' [where x=a]) apply (cut_tac sint_range' [where x=b]) apply clarsimp apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: max_def abs_if split: if_split_asm) done 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) have result_range: "sint a sdiv sint b \ (sints (size a)) \ {2 ^ (size a - 1)}" apply (cut_tac sdiv_int_range [where a="sint a" and b="sint b"]) apply (erule rev_subsetD) using sint_range' [where x=a] sint_range' [where x=b] apply (auto simp: max_def abs_if word_size sints_num) done 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: sdiv_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 (insert word_sint.Rep [where x="a"])[1] apply (clarsimp simp: minus_le_iff word_size abs_if sints_num split: if_split_asm) apply (metis minus_minus sint_int_min word_sint.Rep_inject) done have result_range_simple: "(sint a sdiv sint b \ (sints (size a))) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: word_size sints_num 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] lemmas word_sdiv_numerals_lhs = sdiv_word_def[where a="numeral x" for x] sdiv_word_def[where a=0] sdiv_word_def[where a=1] lemmas word_sdiv_numerals = word_sdiv_numerals_lhs[where b="numeral y" for y] word_sdiv_numerals_lhs[where b=0] word_sdiv_numerals_lhs[where b=1] (* * Signed modulo properties. *) lemma smod_int_alt_def: "(a::int) smod b = sgn (a) * (abs a mod abs b)" apply (clarsimp simp: smod_int_def sdiv_int_def) apply (clarsimp simp: minus_div_mult_eq_mod [symmetric] abs_sgn sgn_mult sgn_if algebra_split_simps) done lemma smod_int_range: "b \ 0 \ (a::int) smod b \ { - abs b + 1 .. abs b - 1 }" apply (case_tac "b > 0") apply (insert pos_mod_conj [where a=a and b=b])[1] apply (insert pos_mod_conj [where a="-a" and b=b])[1] apply (auto simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute])[1] apply (metis add_nonneg_nonneg int_one_le_iff_zero_less le_less less_add_same_cancel2 not_le pos_mod_conj) apply (metis (full_types) add.inverse_inverse eucl_rel_int eucl_rel_int_iff le_less_trans neg_0_le_iff_le) apply (insert neg_mod_conj [where a=a and b="b"])[1] apply (insert neg_mod_conj [where a="-a" and b="b"])[1] apply (clarsimp simp: smod_int_alt_def algebra_simps sgn_if abs_if not_less add1_zle_eq [simplified add.commute]) apply (metis neg_0_less_iff_less neg_mod_conj not_le not_less_iff_gr_or_eq order_trans pos_mod_conj) done lemma smod_int_compares: "\ 0 \ a; 0 < b \ \ (a :: int) smod b < b" "\ 0 \ a; 0 < b \ \ 0 \ (a :: int) smod b" "\ a \ 0; 0 < b \ \ -b < (a :: int) smod b" "\ a \ 0; 0 < b \ \ (a :: int) smod b \ 0" "\ 0 \ a; b < 0 \ \ (a :: int) smod b < - b" "\ 0 \ a; b < 0 \ \ 0 \ (a :: int) smod b" "\ a \ 0; b < 0 \ \ (a :: int) smod b \ 0" "\ a \ 0; b < 0 \ \ b \ (a :: int) smod b" apply (insert smod_int_range [where a=a and b=b]) apply (auto simp: add1_zle_eq smod_int_alt_def sgn_if) done lemma smod_int_mod_0 [simp]: "x smod (0 :: int) = x" by (clarsimp simp: smod_int_def) lemma smod_int_0_mod [simp]: "0 smod (x :: int) = 0" by (clarsimp simp: smod_int_alt_def) lemma smod_word_mod_0 [simp]: "x smod (0 :: ('a::len) word) = x" by (clarsimp simp: smod_word_def) lemma smod_word_0_mod [simp]: "0 smod (x :: ('a::len) word) = 0" by (clarsimp simp: smod_word_def) lemma smod_word_max: "sint (a::'a word) smod sint (b::'a word) < 2 ^ (LENGTH('a::len) - Suc 0)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply (clarsimp) apply (insert word_sint.Rep [where x="b", simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if split: if_split_asm) done lemma smod_word_min: "- (2 ^ (LENGTH('a::len) - Suc 0)) \ sint (a::'a word) smod sint (b::'a word)" apply (case_tac "b = 0") apply (insert word_sint.Rep [where x=a, simplified sints_num])[1] apply clarsimp apply (insert word_sint.Rep [where x=b, simplified sints_num])[1] apply (insert smod_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: abs_if add1_zle_eq split: if_split_asm) done lemma smod_word_alt_def: "(a :: ('a::len) word) smod b = a - (a sdiv b) * b" apply (cases \a \ - (2 ^ (LENGTH('a) - 1)) \ b \ - 1\) apply (clarsimp simp: smod_word_def sdiv_word_def smod_int_def simp flip: wi_hom_sub wi_hom_mult) apply (clarsimp simp: smod_word_def smod_int_def) done lemmas word_smod_numerals_lhs = smod_word_def[where a="numeral x" for x] smod_word_def[where a=0] smod_word_def[where a=1] lemmas word_smod_numerals = word_smod_numerals_lhs[where b="numeral y" for y] word_smod_numerals_lhs[where b=0] word_smod_numerals_lhs[where b=1] lemma sint_of_int_eq: "\ - (2 ^ (LENGTH('a) - 1)) \ x; x < 2 ^ (LENGTH('a) - 1) \ \ sint (of_int x :: ('a::len) word) = x" - apply (clarsimp simp: word_of_int int_word_sint) - apply (subst int_mod_eq') - apply simp - apply (subst (2) power_minus_simp) - apply clarsimp - apply clarsimp - apply clarsimp - done - -lemma of_int_sint [simp]: - "of_int (sint a) = a" - apply (insert word_sint.Rep [where x=a]) - apply (clarsimp simp: word_of_int) - done + by (simp add: signed_take_bit_int_eq_self) + +lemma of_int_sint: + "of_int (sint a) = a" + by simp lemma nth_w2p_scast [simp]: "((scast ((2::'a::len signed word) ^ n) :: 'a word) !! m) \ ((((2::'a::len word) ^ n) :: 'a word) !! m)" apply (subst nth_w2p) apply (case_tac "n \ LENGTH('a)") apply (subst power_overflow, simp) apply clarsimp apply (metis nth_w2p scast_eq test_bit_conj_lt len_signed nth_word_of_int word_sint.Rep_inverse) done lemma scast_2_power [simp]: "scast ((2 :: 'a::len signed word) ^ x) = ((2 :: 'a word) ^ x)" by (clarsimp simp: word_eq_iff) lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (clarsimp simp: word_eq_iff) lemma ucast_nat_def': "of_nat (unat x) = (ucast :: 'a :: len word \ ('b :: len) signed word) x" by (fact ucast_nat_def) lemma mod_mod_power_int: fixes k :: int shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)" by (metis bintrunc_bintrunc_min bintrunc_mod2p min.commute) (* Normalise combinations of scast and ucast. *) 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 (clarsimp simp: word_of_int ucast_eq) + apply (simp only: ucast_eq) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') apply (subst (1 2) int_word_uint) - apply (subst word_of_int) apply (subst word_ubin.norm_eq_iff [symmetric]) apply (subst (1 2) bintrunc_mod2p) apply (insert is_down) apply (unfold is_down_def) apply (clarsimp simp: target_size source_size) apply (clarsimp simp: mod_mod_power_int min_def) apply (rule distrib [symmetric]) 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 (clarsimp simp: scast_eq) + 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_bl 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 smod_mod_positive: "\ 0 \ (a :: int); 0 \ b \ \ a smod b = a mod b" by (clarsimp simp: smod_int_alt_def zsgn_def) lemma nat_mult_power_less_eq: "b > 0 \ (a * b ^ n < (b :: nat) ^ m) = (a < b ^ (m - n))" using mult_less_cancel2[where m = a and k = "b ^ n" and n="b ^ (m - n)"] mult_less_cancel2[where m="a * b ^ (n - m)" and k="b ^ m" and n=1] apply (simp only: power_add[symmetric] nat_minus_add_max) apply (simp only: power_add[symmetric] nat_minus_add_max ac_simps) apply (simp add: max_def split: if_split_asm) done lemma signed_shift_guard_to_word: "\ n < len_of TYPE ('a); n > 0 \ \ (unat (x :: 'a :: len word) * 2 ^ y < 2 ^ n) = (x = 0 \ x < (1 << n >> y))" apply (simp only: nat_mult_power_less_eq) apply (cases "y \ n") apply (simp only: shiftl_shiftr1) apply (subst less_mask_eq) apply (simp add: word_less_nat_alt word_size) apply (rule order_less_le_trans[rotated], rule power_increasing[where n=1]) apply simp apply simp apply simp apply (simp add: nat_mult_power_less_eq word_less_nat_alt word_size) apply auto[1] apply (simp only: shiftl_shiftr2, simp add: unat_eq_0) done 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 (subst sint_eq_uint) apply (clarsimp simp: msb_nth nth_ucast is_down) apply (metis Suc_leI Suc_pred len_gt_0) apply (clarsimp simp: uint_up_ucast is_up is_down) 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 (subst int_word_uint) apply (subst mod_pos_pos_trivial) apply simp apply (rule lt2p_lem) apply (clarsimp simp: is_up) apply simp done lemma ucast_mono: "\ (x :: 'b :: len word) < y; y < 2 ^ LENGTH('a) \ \ ucast x < ((ucast y) :: 'a :: len word)" - apply (simp add: ucast_nat_def [symmetric]) + apply (simp only: flip: ucast_nat_def) apply (rule of_nat_mono_maybe) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) 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 add: ucast_nat_def [symmetric]) + apply (simp only: flip: ucast_nat_def) apply (subst of_nat_mono_maybe_le[symmetric]) apply (rule unat_less_helper) apply (simp add: Power.of_nat_power) apply (rule unat_less_helper) apply (erule le_less_trans) apply (simp add: Power.of_nat_power) apply (simp add: word_le_nat_alt) done lemma ucast_mono_le': "\ unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x \ y \ \ UCAST('a \ 'b) x \ UCAST('a \ 'b) y" by (auto simp: word_less_nat_alt intro: ucast_mono_le) lemma zero_sle_ucast_up: "\ is_down (ucast :: 'a word \ 'b signed word) \ (0 <=s ((ucast (b::('a::len) word)) :: ('b::len) signed word))" apply (subgoal_tac "\ msb (ucast b :: 'b signed word)") apply (clarsimp simp: word_sle_msb_le) apply (clarsimp simp: is_down not_le msb_nth nth_ucast) apply (subst (asm) test_bit_conj_lt [symmetric]) apply clarsimp apply arith done lemma word_le_ucast_sless: "\ x \ y; y \ -1; LENGTH('a) < LENGTH('b) \ \ UCAST (('a :: len) \ ('b :: len) signed) x msb (ucast x :: ('a::len) word) = msb (x :: ('b::len) word)" apply (clarsimp simp: word_msb_alt) apply (subst ucast_down_drop [where n=0]) apply (clarsimp simp: source_size_def target_size_def word_size) apply clarsimp done lemma msb_big: "msb (a :: ('a::len) word) = (a \ 2 ^ (LENGTH('a) - Suc 0))" apply (rule iffI) apply (clarsimp simp: msb_nth) apply (drule bang_is_le) apply simp apply (rule ccontr) apply (subgoal_tac "a = a && mask (LENGTH('a) - Suc 0)") apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"]) apply (clarsimp simp: word_not_le [symmetric]) apply clarsimp apply (rule sym, subst and_mask_eq_iff_shiftr_0) apply (clarsimp simp: msb_shift) done lemma zero_sle_ucast: "(0 <=s ((ucast (b::('a::len) word)) :: ('a::len) signed word)) = (uint b < 2 ^ (LENGTH('a) - 1))" apply (case_tac "msb b") apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI) apply (clarsimp simp: msb_big word_le_def uint_2p_alt) done (* to_bool / from_bool. *) 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) definition to_bool :: "'a::len word \ bool" where "to_bool \ (\) 0" lemma to_bool_and_1: "to_bool (x && 1) = (x !! 0)" by (simp add: test_bit_word_eq 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) && 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 (simp add: case_bool_If from_bool_def split: if_split) 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 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 word_rsplit_upt: "\ size x = LENGTH('a :: len) * n; n \ 0 \ \ word_rsplit x = map (\i. ucast (x >> i * len_of TYPE ('a)) :: 'a word) (rev [0 ..< n])" apply (subgoal_tac "length (word_rsplit x :: 'a word list) = n") apply (rule nth_equalityI, simp) apply (intro allI word_eqI impI) apply (simp add: test_bit_rsplit_alt word_size) apply (simp add: nth_ucast nth_shiftr rev_nth field_simps) apply (simp add: length_word_rsplit_exp_size) apply (metis mult.commute given_quot_alt word_size word_size_gt_0) done lemma aligned_shift: "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ x + y >> n = y >> n" by (subst word_plus_and_or_coroll; word_eqI, blast) lemma aligned_shift': "\x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \ LENGTH('a)\ \ y + x >> n = y >> n" by (subst word_plus_and_or_coroll; word_eqI, blast) lemma neg_mask_add_mask: "((x:: 'a :: len word) && ~~ (mask n)) + (2 ^ n - 1) = x || mask n" unfolding mask_2pm1[symmetric] by (subst word_plus_and_or_coroll; word_eqI_solve) lemma subtract_mask: "p - (p && mask n) = (p && ~~ (mask n))" "p - (p && ~~ (mask n)) = (p && mask n)" by (simp add: field_simps word_plus_and_or_coroll2)+ lemma and_neg_mask_plus_mask_mono: "(p && ~~ (mask n)) + mask n \ p" apply (rule word_le_minus_cancel[where x = "p && ~~ (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 && ~~ (mask n)) + (2 ^ n - 1)" by (simp add: and_neg_mask_plus_mask_mono mask_2pm1[symmetric]) 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 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 is_aligned_add_helper: "\ is_aligned p n; d < 2 ^ n \ \ (p + d && mask n = d) \ (p + d && (~~ (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; word_eqI; blast) using word_plus_and_or_coroll2[where x="p + d" and w="mask n"] by simp lemma is_aligned_sub_helper: "\ is_aligned (p - d) n; d < 2 ^ n \ \ (p && mask n = d) \ (p && (~~ (mask n)) = p - d)" by (drule(1) is_aligned_add_helper, simp) lemma mask_twice: "(x && mask n) && mask m = x && mask (min m n)" by word_eqI_solve lemma is_aligned_after_mask: "\is_aligned k m;m\ n\ \ is_aligned (k && mask n) m" by (rule is_aligned_andI1) lemma and_mask_plus: "\is_aligned ptr m; m \ n; a < 2 ^ m\ \ ptr + a && mask n = (ptr && 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 && mask n) && ~~ (mask m) = (ptr + a && ~~ (mask m) ) && 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 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 && mask n) && ~~ (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 && b) && ~~ b = 0" apply (subst word_bw_assocs(1)) apply clarsimp done lemma mask_shift_and_negate[simp]:"(w && mask n << m) && ~~ (mask n << m) = 0" by (clarsimp simp add: mask_eq_decr_exp Parity.bit_eq_iff bit_and_iff bit_not_iff shiftl_word_eq bit_push_bit_iff) 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 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:"~~ a = x \ a = ~~ x" by auto (* 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 && ~~ (mask n << m) || ((y && mask n) << m)) && ~~ (mask n << m) = x && ~~ (mask n << m)" by (induct n arbitrary: m) (auto simp: word_ao_dist) lemma bitfield_op_twice'': "\~~ a = b << c; \x. b = mask x\ \ (x && a || (y && b << c)) && a = x && a" 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 bit_twiddle_min: "(y::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = min x y" by (auto simp add: Parity.bit_eq_iff bit_xor_iff min_def) lemma bit_twiddle_max: "(x::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = max x y" by (auto simp add: Parity.bit_eq_iff bit_xor_iff 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 scast_nop1: +lemma scast_nop1 [simp]: "((scast ((of_int x)::('a::len) word))::'a sword) = of_int x" - apply (clarsimp simp:scast_eq word_of_int) + apply (simp only: scast_eq) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) -lemma scast_nop2: +lemma scast_nop2 [simp]: "((scast ((of_int x)::('a::len) sword))::'a word) = of_int x" - apply (clarsimp simp:scast_eq word_of_int) + apply (simp only: scast_eq) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) -lemmas scast_nop[simp] = scast_nop1 scast_nop2 scast_id +lemmas scast_nop = scast_nop1 scast_nop2 scast_id lemma le_mask_imp_and_mask: "(x::'a::len word) \ mask n \ x && mask n = x" by (metis and_mask_eq_iff_le_mask) lemma or_not_mask_nop: "((x::'a::len word) || ~~ (mask n)) && mask n = x && 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) || y && mask n) && ~~ (mask m) = x && ~~ (mask m)" by (auto simp add: Parity.bit_eq_iff bit_not_iff bit_or_iff bit_and_iff bit_mask_iff) lemma and_mask_0_iff_le_mask: fixes w :: "'a::len word" shows "(w && ~~(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) && mask m) && mask n = x && mask n" by (metis mask_twice min_def) lemma uint_2_id: "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" by simp lemma bintrunc_id: "\m \ of_nat n; 0 < m\ \ bintrunc n m = m" by (simp add: bintrunc_mod2p le_less_trans) lemma shiftr1_unfold: "shiftr1 x = x >> 1" by (metis One_nat_def comp_apply funpow.simps(1) funpow.simps(2) id_apply shiftr_def) lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" by transfer (simp add: drop_bit_Suc) 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 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 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 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 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 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 \ x !! 0" using even_plus_one_iff [of x] by (simp add: test_bit_word_eq) lemma word_lsb_nat:"lsb w = (unat w mod 2 = 1)" apply (simp add: word_lsb_def Groebner_Basis.algebra(31)) apply transfer apply (simp add: even_nat_iff) done lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = x !! 0" apply (simp add:even_iff_mod_2_eq_zero) apply (subst word_lsb_nat[unfolded One_nat_def, symmetric]) apply (rule word_lsb_alt) done 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 simp+ apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") apply (subst (asm) word_unat.norm_eq_iff[symmetric]) apply simp+ done lemma shiftr1_irrelevant_lsb:"(x::('a::len) word) !! 0 \ x >> 1 = (x + 1) >> 1" using word_overflow_unat [of x] apply (simp only: shiftr1_is_div_2 flip: odd_iff_lsb) apply (cases \2 \ LENGTH('a)\) - apply (auto simp add: test_bit_def' uint_nat word_arith_nat_div dest: overflow_imp_lsb) - done + apply (auto simp add: test_bit_def' word_arith_nat_div dest: overflow_imp_lsb) + using odd_iff_lsb overflow_imp_lsb by blast 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':"\((x::('a::len) word) !! 0) \ x >> 1 = (x + 1) >> 1" by (metis shiftr1_irrelevant_lsb) lemma lsb_this_or_next:"\(((x::('a::len) word) + 1) !! 0) \ x !! 0" by (metis (poly_guards_query) even_word_imp_odd_next odd_iff_lsb overflow_imp_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) || (((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) || (((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 mask_or_not_mask: "x && mask n || x && ~~ (mask n) = x" apply (subst word_oa_dist, simp) apply (subst word_oa_dist2, simp) 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 word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" by (metis add.commute add_minus_cancel) lemma neg_mask_add_aligned: "\ is_aligned p n; q < 2 ^ n \ \ (p + q) && ~~ (mask n) = p && ~~ (mask n)" by (metis is_aligned_add_helper is_aligned_neg_mask_eq) 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" from word_unat.Abs_inj_on have "inj_on ?of_nat {i. i < CARD('a word)}" by (simp add: unats_def card_word) 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..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 add: ucast_eq word_of_int_minus) + 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 - by (simp add: ucast_eq word_ubin.eq_norm bintrunc_mod2p) + apply (simp add: word_ubin.eq_norm bintrunc_mod2p unsigned_ucast_eq) + by (metis \x \ 2 ^ LENGTH('b) - 1\ le_def take_bit_word_eq_self ucast_ucast_len unsigned_take_bit_eq word_less_sub_le word_ubin.norm_Rep word_uint_eqI) qed lemma remdups_enum_upto: fixes s::"'a::len word" shows "remdups [s .e. e] = [s .e. e]" by simp lemma card_enum_upto: fixes s::"'a::len word" shows "card (set [s .e. e]) = Suc (unat e) - unat s" by (subst List.card_set) (simp add: remdups_enum_upto) 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 word_shiftr_lt: fixes w :: "'a::len word" shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))" apply (subst shiftr_div_2n') by (metis nat_mod_lem nat_zero_less_power_iff power_mod_div word_unat.Rep_inverse word_unat.eq_norm zero_less_numeral) lemma complement_nth_w2p: shows "n' < LENGTH('a) \ (~~ (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 && 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) && mask m) < 2 ^ m" by (rule word_unat_and_lt) (simp add: unat_mask word_size) 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 td_gal_lt) lemma le_or_mask: "w \ w' \ w || mask x \ w' || mask x" by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) lemma le_shiftr1': "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" apply transfer apply simp done lemma le_shiftr': "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" apply (induct n; simp add: shiftr_def) apply (case_tac "(shiftr1 ^^ n) u = (shiftr1 ^^ n) v", simp) apply (fastforce dest: le_shiftr1') done lemma word_log2_nth_same: "w \ 0 \ w !! word_log2 w" unfolding word_log2_def using nth_length_takeWhile[where P=Not and xs="to_bl w"] apply (simp add: word_clz_def word_size to_bl_nth) apply (fastforce simp: linorder_not_less eq_zero_set_bl dest: takeWhile_take_has_property) done lemma word_log2_nth_not_set: "\ word_log2 w < i ; i < size w \ \ \ w !! i" unfolding word_log2_def word_clz_def using takeWhile_take_has_property_nth[where P=Not and xs="to_bl w" and n="size w - Suc i"] by (fastforce simp add: to_bl_nth word_size) lemma word_log2_highest: assumes a: "w !! i" shows "i \ word_log2 w" proof - from a have "i < size w" by - (rule test_bit_size) with a show ?thesis by - (rule ccontr, simp add: word_log2_nth_not_set) qed lemma word_log2_max: "word_log2 w < size w" unfolding word_log2_def word_clz_def by simp lemma word_clz_0[simp]: "word_clz (0::'a::len word) = LENGTH('a)" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma word_clz_minus_one[simp]: "word_clz (-1::'a::len word) = 0" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma word_add_no_overflow:"(x::'a::len word) < max_word \ 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" unfolding ucast_eq unat_eq_nat_uint apply (subst int_word_uint) apply (subst mod_pos_pos_trivial; simp?) apply (rule lt2p_lem) apply (simp add: assms) 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 add: ucast_nat_def[symmetric]) + 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] (* 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 minus_one_word: "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by simp lemma mask_exceed: "n \ LENGTH('a) \ (x::'a::len word) && ~~ (mask n) = 0" by (simp add: and_not_mask shiftr_eq_0) lemma two_power_strict_part_mono: "strict_part_mono {..LENGTH('a) - 1} (\x. (2 :: 'a :: len word) ^ x)" proof - { fix n have "n < LENGTH('a) \ strict_part_mono {..n} (\x. (2 :: 'a :: len word) ^ x)" proof (induct n) case 0 then show ?case by simp next case (Suc n) from Suc.prems have "2 ^ n < (2 :: 'a :: len word) ^ Suc n" using power_strict_increasing unat_power_lower word_less_nat_alt by fastforce with Suc show ?case by (subst strict_part_mono_by_steps) simp qed } then show ?thesis by simp qed lemma word_shift_by_2: "x * 4 = (x::'a::len word) << 2" by (simp add: shiftl_t2n) lemma le_2p_upper_bits: "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ \n'\n. n' < LENGTH('a) \ \ p !! n'" by (subst upper_bits_unset_is_l2p; simp) lemma le2p_bits_unset: "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ (p::'a::len word) !! n'" using upper_bits_unset_is_l2p [where p=p] by (cases "n < LENGTH('a)") auto 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 lemma word_power_nonzero: "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ \ x * 2 ^ n \ 0" by (metis and_mask_eq_iff_shiftr_0 less_mask_eq p2_gt_0 semiring_normalization_rules(7) shiftl_shiftr_id shiftl_t2n) 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 (rule word_uint.Rep_eqD) apply (simp only: uint_word_ariths uint_div uint_power_lower) apply (subst mod_pos_pos_trivial, fastforce, fastforce)+ apply (subst mod_pos_pos_trivial) apply (simp add: le_diff_eq uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst mod_pos_pos_trivial) apply (simp add: uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst int_div_sub_1; simp add: uint_2p_alt) apply (subst power_0[symmetric]) apply (simp add: uint_2p_alt le_imp_power_dvd power_sub_int) done 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 (rule nat_less_power_trans2[simplified]) apply (simp add: word_less_nat_alt) apply simp 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) unfolding is_aligned_def by (metis le_unat_uoi nat_dvd_not_less order_less_imp_le unat_power_lower) lemma length_upto_enum_one: fixes x :: "'a :: len word" assumes lt1: "x < y" and lt2: "z < y" and lt3: "x \ z" shows "[x , y .e. z] = [x]" unfolding upto_enum_step_def proof (subst upto_enum_red, subst if_not_P [OF leD [OF lt3]], clarsimp, rule conjI) show "unat ((z - x) div (y - x)) = 0" proof (subst unat_div, rule div_less) have syx: "unat (y - x) = unat y - unat x" by (rule unat_sub [OF order_less_imp_le]) fact moreover have "unat (z - x) = unat z - unat x" by (rule unat_sub) fact ultimately show "unat (z - x) < unat (y - x)" using lt2 lt3 unat_mono word_less_minus_mono_left by blast qed then show "(z - x) div (y - x) * (y - x) = 0" by (metis mult_zero_left unat_0 word_unat.Rep_eqD) qed lemma max_word_mask: "(max_word :: 'a::len word) = mask LENGTH('a)" unfolding mask_eq_decr_exp by simp lemmas mask_len_max = max_word_mask[symmetric] lemma is_aligned_alignUp[simp]: "is_aligned (alignUp p n) n" by (simp add: alignUp_def complement_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 complement_mask: "complement (2 ^ n - 1) = ~~ (mask n)" unfolding complement_def mask_eq_decr_exp by simp 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 complement_mask is_aligned_add_helper p_assoc_help power_2_ge_iff) 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 complement_mask) 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 shiftr_div_2n' word_shiftr_lt) 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 min.absorb2) 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 shiftr_div_2n' word_shiftr_lt) 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 (metis unat_div unat_less_helper unat_power_lower) 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, hide_lams) 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 unat_eq_nat_uint zero_neq_numeral) + 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 (clarsimp elim!: range_subset_lower [where x = x]) 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 (clarsimp elim!: range_subset_upper [where x = x]) 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 && ~~ (mask sz)" unfolding alignUp_def[unfolded complement_def] by (simp add:mask_eq_decr_exp[symmetric,unfolded shiftl_t2n,simplified]) lemma mask_out_first_mask_some: "\ x && ~~ (mask n) = y; n \ m \ \ x && ~~ (mask m) = y && ~~ (mask m)" by word_eqI_solve 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 mask_out_add_aligned: assumes al: "is_aligned p n" shows "p + (q && ~~ (mask n)) = (p + q) && ~~ (mask n)" using mask_add_aligned [OF al] by (simp add: mask_out_sub_mask) lemma alignUp_def3: "alignUp a sz = 2^ sz + (a - 1 && ~~ (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 mask_lower_twice: "n \ m \ (x && ~~ (mask n)) && ~~ (mask m) = x && ~~ (mask m)" by word_eqI_solve lemma mask_lower_twice2: "(a && ~~ (mask n)) && ~~ (mask m) = a && ~~ (mask (max n m))" by word_eqI_solve lemma ucast_and_neg_mask: "ucast (x && ~~ (mask n)) = ucast x && ~~ (mask n)" by word_eqI_solve lemma ucast_and_mask: "ucast (x && mask n) = ucast x && mask n" by word_eqI_solve lemma ucast_mask_drop: "LENGTH('a :: len) \ n \ (ucast (x && mask n) :: 'a word) = ucast x" by word_eqI 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 && ~~ (mask sz)) = (p - ((alignUp q sz) && ~~ (mask sz)))" apply (clarsimp simp only:word_and_le2 diff_conv_add_uminus) apply (subst mask_out_add_aligned[symmetric]; simp) apply (rule sum_to_zero) apply (subst add.commute) by (simp add: alignUp_distance and_mask_0_iff_le_mask is_aligned_neg_mask_eq mask_out_add_aligned) lemma map_bits_rev_to_bl: "map ((!!) x) [0.. LENGTH('a) \ x = ucast y \ ucast x = y" for x :: "'a::len word" and y :: "'b::len word" by (simp add: is_down ucast_ucast_a) 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_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 word_le_not_less: "((b::'a::len word) \ a) = (\(a < b))" by fastforce lemma ucast_or_distrib: fixes x :: "'a::len word" fixes y :: "'a::len word" shows "(ucast (x || y) :: ('b::len) word) = ucast x || ucast y" by transfer simp 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 && w' \ 0 \ w \ 0 \ w' \ 0" by auto 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" by (fastforce intro!: list_of_false) 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 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 from_bool_eq_if': "((if P then 1 else 0) = from_bool Q) = (P = Q)" by (simp add: case_bool_If from_bool_def split: if_split) lemma word_exists_nth: "(w::'a::len word) \ 0 \ \i. w !! i" using word_log2_nth_same by blast lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" by (rule word_unat.Rep_eqD) (simp add: shiftr_div_2n') 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) lemma max_word_not_0 [simp]: "- 1 \ (0 :: 'a::len word)" by simp lemma ucast_zero_is_aligned: "UCAST('a::len \ 'b::len) w = 0 \ n \ LENGTH('b) \ is_aligned w n" by (clarsimp simp: is_aligned_mask word_eq_iff word_size nth_ucast) lemma unat_ucast_eq_unat_and_mask: "unat (UCAST('b::len \ 'a::len) w) = unat (w && mask LENGTH('a))" proof - have "unat (UCAST('b \ 'a) w) = unat (UCAST('a \ 'b) (UCAST('b \ 'a) w))" by (cases "LENGTH('a) < LENGTH('b)"; simp add: is_down ucast_ucast_a unat_ucast_up_simp) thus ?thesis using ucast_ucast_mask by simp qed 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 (cases n) case 0 thus ?thesis using eq by simp next case (Suc n') have m_lt: "m < LENGTH('a)" using Suc ws by simp have xylt: "x < 2^m" "y < 2^m" using le m_lt unfolding mask_2pm1 by auto have lenm: "n \ LENGTH('a) - m" using ws by simp show ?thesis using eq xylt apply (fold shiftl_t2n[where n=n, simplified mult.commute]) apply (simp only: word_bl.Rep_inject[symmetric] bl_shiftl) apply (erule ssubst[OF less_is_drop_replicate])+ apply (clarsimp elim!: drop_eq_mono[OF lenm]) 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]) (* Sign extension from bit n. *) lemma sign_extend_bitwise_if: "i < size w \ sign_extend e w !! i \ (if i < e then w !! i else w !! e)" by (simp add: sign_extend_def neg_mask_test_bit word_size) lemma sign_extend_bitwise_disj: "i < size w \ sign_extend e w !! i \ i \ e \ w !! i \ e \ i \ w !! e" by (auto simp: sign_extend_bitwise_if) lemma sign_extend_bitwise_cases: "i < size w \ sign_extend e w !! i \ (i \ e \ w !! i) \ (e \ i \ w !! e)" by (auto simp: sign_extend_bitwise_if) lemmas sign_extend_bitwise_if'[word_eqI_simps] = sign_extend_bitwise_if[simplified word_size] 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 w !! n then w || ~~ (mask (Suc n)) else w && mask (Suc n))" by word_eqI (auto dest: less_antisym) 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 (rule iffI) apply (word_eqI, rename_tac i) apply (case_tac "n < i"; simp add: sign_extended_def word_size) apply (erule subst, rule sign_extended_sign_extend) 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 word_eqI lemma sign_extended_high_bits: "\ sign_extended e p; j < size p; e \ i; i < j \ \ p !! i = p !! j" by (drule (1) sign_extended_weaken; simp add: sign_extended_def) lemma sign_extend_eq: "w && mask (Suc n) = v && mask (Suc n) \ sign_extend n w = sign_extend n v" by word_eqI_solve 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 "\ f !! e" using True e less_2p_is_upper_bits_unset[THEN iffD1, OF f] by (fastforce simp: word_size) hence i: "(p + f) !! e = p !! e" by (simp add: and_or) have fm: "f && mask e = f" by (fastforce intro: subst[where P="\f. f && 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 && ~~ (mask m))" by (fastforce simp: sign_extended_def word_size neg_mask_test_bit) (* Uints *) lemma uints_mono_iff: "uints l \ uints m \ l \ m" using power_increasing_iff[of "2::int" l m] apply (auto simp: uints_num subset_iff simp del: power_increasing_iff) by (meson less_irrefl not_less zle2p) lemmas uints_monoI = uints_mono_iff[THEN iffD2] lemma Bit_in_uints_Suc: "of_bool c + 2 * w \ uints (Suc m)" if "w \ uints m" using that by (auto simp: uints_num) lemma Bit_in_uintsI: "of_bool c + 2 * w \ uints m" if "w \ uints (m - 1)" "m > 0" using Bit_in_uints_Suc[OF that(1)] that(2) by auto lemma bin_cat_in_uintsI: \bin_cat a n b \ uints m\ if \a \ uints l\ \m \ l + n\ proof - from \m \ l + n\ obtain q where \m = l + n + q\ using le_Suc_ex by blast then have \(2::int) ^ m = 2 ^ n * 2 ^ (l + q)\ by (simp add: ac_simps power_add) moreover have \a mod 2 ^ (l + q) = a\ using \a \ uints l\ by (auto simp add: uints_def take_bit_eq_mod power_add Divides.mod_mult2_eq) ultimately have \concat_bit n b a = take_bit m (concat_bit n b a)\ by (simp add: concat_bit_eq take_bit_eq_mod push_bit_eq_mult Divides.mod_mult2_eq) then show ?thesis by (simp add: uints_def) qed lemma bin_cat_cong: "bin_cat a n b = bin_cat c m d" if "n = m" "a = c" "bintrunc m b = bintrunc m d" using that(3) unfolding that(1,2) by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bin_cat_eqD1: "bin_cat a n b = bin_cat c n d \ a = c" by (metis drop_bit_bin_cat_eq) lemma bin_cat_eqD2: "bin_cat a n b = bin_cat c n d \ bintrunc n b = bintrunc n d" by (metis take_bit_bin_cat_eq) lemma bin_cat_inj: "(bin_cat a n b) = bin_cat c n d \ a = c \ bintrunc n b = bintrunc n d" by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) lemma word_of_int_bin_cat_eq_iff: "(word_of_int (bin_cat (uint a) LENGTH('b) (uint b))::'c::len word) = word_of_int (bin_cat (uint c) LENGTH('b) (uint d)) \ b = d \ a = c" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len word" and b::"'b::len word" by (subst word_uint.Abs_inject) (auto simp: bin_cat_inj intro!: that bin_cat_in_uintsI) 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 transfer auto 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 (* 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 unat_minus_one_word: "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by (subst minus_one_word) (subst unat_sub_if', clarsimp) 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_ctz_le: "word_ctz (w :: ('a::len word)) \ LENGTH('a)" apply (clarsimp simp: word_ctz_def) apply (rule nat_le_Suc_less_imp[where y="LENGTH('a) + 1" , simplified]) apply (rule order_le_less_trans[OF List.length_takeWhile_le]) 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) apply (rule order_less_le_trans[OF length_takeWhile_less]) apply fastforce+ done lemma word_ctz_not_minus_1: "1 < LENGTH('a) \ of_nat (word_ctz (w :: 'a :: len word)) \ (- 1 :: 'a::len word)" by (metis (mono_tags) One_nat_def add.right_neutral add_Suc_right le_diff_conv le_less_trans n_less_equal_power_2 not_le suc_le_pow_2 unat_minus_one_word unat_of_nat_len word_ctz_le) 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 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 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 - word_eqI thus ?thesis by simp qed lemma of_bl_max: "(of_bl xs :: 'a::len word) \ mask (length xs)" apply (induct xs) apply simp apply (simp add: of_bl_Cons mask_Suc) apply (rule conjI; clarsimp) apply (erule word_plus_mono_right) apply (rule is_aligned_no_overflow_mask) apply (rule is_aligned_triv) apply (simp add: word_le_nat_alt) apply (subst unat_add_lem') apply (rule is_aligned_mask_offset_unat) apply (rule is_aligned_triv) apply (simp add: mask_eq_decr_exp) apply simp done lemma mask_over_length: "LENGTH('a) \ n \ mask n = (-1::'a::len word)" by (simp add: mask_eq_decr_exp) 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 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 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) apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask) apply (clarsimp simp: is_aligned_def dvd_def word_le_nat_alt) apply (drule le_imp_less_Suc) apply (simp add: Suc_2p_unat_mask) by (metis Groups.mult_ac(2) Suc_leI linorder_not_less mult_le_mono order_refl times_nat.simps(2)) lemma power_2_mult_step_le: "\n' \ n; 2 ^ n' * k' < 2 ^ n * k\ \ 2 ^ n' * (k' + 1) \ 2 ^ n * (k::nat)" apply (cases "n'=n", simp) apply (metis Suc_leI le_refl mult_Suc_right mult_le_mono semiring_normalization_rules(7)) apply (drule (1) le_neq_trans) apply clarsimp apply (subgoal_tac "\m. n = n' + m") prefer 2 apply (simp add: le_Suc_ex) apply (clarsimp simp: power_add) by (metis Suc_leI mult.assoc mult_Suc_right nat_mult_le_cancel_disj) 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: is_aligned_def dvd_def) 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 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 && b = 0" by word_eqI_solve lemma word_and_or_mask_aligned: "\ is_aligned a n; b \ mask n \ \ a + b = a || 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 || 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 (clarsimp simp: word_eqI_simps) 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" by (clarsimp simp: le_mask_high_bits word_size nth_ucast) 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 + (UCAST('a::len \ 'b::len) off << n) = p' + (ucast off' << n); is_aligned p n'; is_aligned p' n'; n' = n + LENGTH('a); n' < LENGTH('b) \ \ p' = p \ off' = off" apply (simp add: word_and_or_mask_aligned le_mask_shiftl_le_mask[where n="LENGTH('a)"] ucast_leq_mask) apply (simp add: is_aligned_nth) apply (rule conjI; word_eqI) apply (metis add.commute test_bit_conj_lt diff_add_inverse le_diff_conv nat_less_le) apply (rename_tac i) apply (erule_tac x="i+n" in allE) apply simp 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 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 max_word_not_less[simp]: "\ max_word < x" by (simp add: not_less) 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 fold_eq_0_to_bool: "(v = 0) = (\ to_bool v)" by (simp add: to_bool_def) 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) && ~~(mask n) = (x && ~~(mask n)) + y * m\ if \m && (2 ^ n - 1) = 0\ by (metis (no_types, hide_lams) 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: 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 && mask n = 0 \ x + y && ~~(mask n) = (x && ~~(mask n)) + y" 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 && mask n = 0; y && 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_def) done lemma sub_right_shift: "\ x && mask n = 0; y && 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 && mask n = mask n \ (x && y) && mask n = x && mask n" by (simp add: ac_simps) lemma and_and_mask_simple_not: "y && mask n = 0 \ (x && y) && mask n = 0" by (simp add: ac_simps) lemma word_and_le': "b \ c \ (a :: 'a :: len word) && b \ c" by (metis word_and_le1 order_trans) lemma word_and_less': "b < c \ (a :: 'a :: len word) && b < c" by (metis word_and_le1 xtr7) 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 ucast_ucast_mask_eq: "\ UCAST('a::len \ 'b::len) x = y; x && mask LENGTH('b) = x \ \ x = ucast y" by word_eqI_solve 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 && mask n = 0; m \ n \ \ x && mask m = 0" by (metis mask_twice2 word_and_notzeroD) lemma mask_len_id [simp]: "(x :: 'a :: len word) && 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) && 1 = x \ \ x = 0 \ x = 1" by (metis word_and_1) lemma pow_mono_leq_imp_lt: "x \ y \ x < 2 ^ y" by (simp add: le_less_trans) lemma unat_of_nat_ctz_mw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len word) = word_ctz w" using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"] pow_mono_leq_imp_lt by simp lemma unat_of_nat_ctz_smw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len sword) = word_ctz w" using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"] pow_mono_leq_imp_lt by (metis le_unat_uoi le_unat_uoi linorder_neqE_nat nat_less_le scast_of_nat word_unat.Rep_inverse) lemma shiftr_and_eq_shiftl: "(w >> n) && x = y \ w && (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 neg_mask_combine: "~~(mask a) && ~~(mask b) = ~~(mask (max a b))" by (auto simp: word_ops_nth_size word_size intro!: word_eqI) lemma neg_mask_twice: "x && ~~(mask n) && ~~(mask m) = x && ~~(mask (max n m))" by (metis neg_mask_combine) lemma multiple_mask_trivia: "n \ m \ (x && ~~(mask n)) + (x && mask n && ~~(mask m)) = x && ~~(mask m)" 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 add_mask_lower_bits': "\ len = LENGTH('a); is_aligned (x :: 'a :: len word) n; \n' \ n. n' < len \ \ p !! n' \ \ x + p && ~~(mask n) = x" using add_mask_lower_bits by auto lemma neg_mask_in_mask_range: "is_aligned ptr bits \ (ptr' && ~~(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 le_word_or2 neg_mask_add_mask and.right_idem) 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 (simp only: is_aligned_add_or flip: neg_mask_in_mask_range) by (metis less_mask_eq mask_subsume) 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 (erule nonemptyE) apply simp 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 (intro range_subsetI) apply (rule is_aligned_no_wrap' [OF al xsz]) 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_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) by (simp add: aligned_add_mask_less_eq is_aligned_weaken algebra_split_simps) lemma aligned_mask_ranges_disjoint: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n'; p && ~~(mask n') \ p'; p' && ~~(mask n) \ p \ \ mask_range p n \ mask_range p' n' = {}" using aligned_mask_range_cases by (auto simp: neg_mask_in_mask_range) lemma aligned_mask_ranges_disjoint2: "\ is_aligned p n; is_aligned ptr bits; n \ m; n < size p; m \ bits; (\y < 2 ^ (n - m). p + (y << m) \ mask_range ptr bits) \ \ mask_range p n \ mask_range ptr bits = {}" apply safe apply (simp only: flip: neg_mask_in_mask_range) apply (drule_tac x="x && mask n >> m" in spec) apply (clarsimp simp: shiftr_less_t2n and_mask_less_size wsst_TYs multiple_mask_trivia word_bw_assocs neg_mask_twice max_absorb2 shiftr_shiftl1) done 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) 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)" - by (clarsimp simp: word_le_def uint_nat of_nat_inverse) + 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 && ~~(mask sz)) + 2 ^ sz - 1" 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, hide_lams) 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) && (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) && mask m) = unat (x && mask m)" by (metis ucast_and_mask unat_ucast_up_simp) lemma small_powers_of_2: "x \ 3 \ x < 2 ^ (x - 1)" by (induct x; simp add: suc_le_pow_2) lemma word_clz_sint_upper[simp]: "LENGTH('a) \ 3 \ sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a sword) \ int (LENGTH('a))" using small_powers_of_2 by (smt One_nat_def diff_less le_less_trans len_gt_0 len_signed lessI n_less_equal_power_2 not_msb_from_less of_nat_mono sint_eq_uint uint_nat unat_of_nat_eq unat_power_lower word_clz_max word_of_nat_less wsst_TYs(3)) 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 small_powers_of_2 uint_nat apply (simp add: order_le_less_trans[OF word_clz_max] not_msb_from_less word_of_nat_less word_size) by (simp add: uint_nat) 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" using less_not_refl3 le_step_down_nat le_trans less_or_eq_imp_le word_shiftr_lt by (metis (no_types, lifting)) 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" by (simp add: mask_shift multi_shift_simps(5) shiftr_shiftr) lemma of_int_uint_ucast: "of_int (uint (x :: 'a::len word)) = (ucast x :: 'b::len word)" - by (simp add: ucast_eq word_of_int) + by (fact Word.of_int_uint) lemma mod_mask_drop: "\ m = 2 ^ n; 0 < m; mask n && msk = mask n \ \ (x mod m) && msk = x mod m" by (simp add: word_mod_2p_is_mask word_bw_assocs) lemma mask_eq_ucast_eq: "\ x && 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 && ~~(mask n)" 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 && ~~(mask sz')) - (ptr && ~~(mask sz)) \ 2 ^ sz - 2 ^ sz'" (is "_ \ ?lhs \ ?rhs") proof - assume lt: "sz' \ sz" hence "?lhs = ptr && (mask sz && ~~(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_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 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 of_bl_length2: "length xs + c < LENGTH('a) \ of_bl xs * 2^c < (2::'a::len word) ^ (length xs + c)" by (simp add: of_bl_length word_less_power_trans2) lemma mask_out_eq_0: "\ idx < 2 ^ sz; sz < LENGTH('a) \ \ (of_nat idx :: 'a :: len word) && ~~(mask sz) = 0" by (simp add: Word_Lemmas.of_nat_power less_mask_eq mask_eq_0_eq_x) lemma is_aligned_neg_mask_eq': "is_aligned ptr sz = (ptr && ~~(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) && ~~(mask sz)) + unat (ptr && 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 (no_types) Word_Lemmas.of_nat_power diff_le_self le_less_trans shiftl_less_t2n unat_less_power word_unat.Rep_inverse) lemma unat_is_aligned_add: "\ is_aligned p n; unat d < 2 ^ n \ \ unat (p + d && mask n) = unat d \ unat (p + d && ~~(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) && ~~(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) && mask high_bits = x >> low_bits" 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 && mask low_bits = 0) = (x = 0)" 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) && mask n = q && 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 from_to_bool_last_bit: "from_bool (to_bool (x && 1)) = x && 1" by (metis from_bool_to_bool_iff word_and_1) 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) && mask sz) + unat z < 2 ^ sz; sz < LENGTH('a) \ \ x \ x + z" apply (subgoal_tac "(x && ~~(mask sz)) + (x && mask sz) \ (x && ~~(mask sz)) + ((x && 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 Word_Lemmas.of_nat_power is_aligned_no_wrap' by force lemma mul_not_mask_eq_neg_shiftl: "~~(mask n) = -1 << n" by (simp add: NOT_mask shiftl_t2n) lemma shiftr_mul_not_mask_eq_and_not_mask: "(x >> n) * ~~(mask n) = - (x && ~~(mask n))" 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) && ~~(mask n) = p + (x && ~~(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) && ~~(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) && mask (m::nat) = mask (min m n)" 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 && mask m = (if n < m then 2 ^ n else 0)" by (rule word_eqI, auto simp add: word_size nth_w2p split: if_split) 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 (max_word :: '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 && mask a >> b) :: 'a :: len word) = p && 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 && mask a >> b) :: 'a :: len word) = unat (p && mask a >> b)" by (metis linear ucast_ucast_mask_shift unat_ucast_up_simp) lemma mask_overlap_zero: "a \ b \ (p && mask a) && ~~(mask b) = 0" by (metis NOT_mask_AND_mask mask_lower_twice2 max_def) lemma mask_shifl_overlap_zero: "a + c \ b \ (p && mask a << c) && ~~(mask b) = 0" 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 && ~~(mask a)) && mask b = 0" 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 && ~~(mask a)) + ((p && mask a) && ~~(mask b)) + (p && mask b) = p" 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 && mask a >> b << b) = (p && mask a) && ~~(mask b)" by (simp add: and_not_mask) lemma mask_shift_sum: "\ a \ b; unat n = unat (p && mask b) \ \ (p && ~~(mask a)) + (p && mask a >> b) * (1 << b) + n = (p :: 'a :: len word)" by (metis and_not_mask mask_rshift_mult_eq_rshift_lshift mask_split_sum_twice word_unat.Rep_eqD) 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 (metis scast_eq word_of_int) + 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 (metis cast_simps(23) scast_scast_id(2)) 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 && mask n \ of_nat ` set [0 ..< 2 ^ n]" -proof - - have "x && mask n \ {0 .. 2 ^ n - 1}" - by (simp add: mask_eq_decr_exp word_and_le1) - also - have "... = of_nat ` {0 .. 2 ^ n - 1}" - apply (rule set_eqI, rule iffI) - apply (clarsimp simp: image_iff) - apply (rule_tac x="unat x" in bexI; simp) - using len - apply (simp add: word_le_nat_alt unat_2tp_if unat_minus_one) - using len - apply (clarsimp simp: word_le_nat_alt unat_2tp_if unat_minus_one) - apply (subst unat_of_nat_eq; simp add: nat_le_Suc_less) - apply (erule less_le_trans) - apply simp - done - also have "{0::nat .. 2^n - 1} = set [0 ..< 2^n]" by (auto simp: nat_le_Suc_less) - finally show ?thesis . -qed + 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_of_nat_ge_zero: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) \ 0" - by (simp add: Word_Lemmas.of_nat_power not_msb_from_less sint_eq_uint) + by (simp add: bit_iff_odd) 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 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)" by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2) nat_power_minus_less of_nat_le_iff sint_eq_uint_2pl uint_nat unat_of_nat_len) lemma int_eq_sint: "x < 2 ^ (LENGTH('a) - 1) \ sint (of_nat x :: 'a :: len word) = int x" by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2) nat_less_le sint_eq_uint_2pl uint_nat unat_lt2p unat_of_nat_len unat_power_lower) 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) by (rule small_powers_of_2, simp) lemma pow_sub_less: "\ a + b \ LENGTH('a); unat (x :: 'a :: len word) = 2 ^ a \ \ unat (x * 2 ^ b - 1) < 2 ^ (a + b)" by (metis (mono_tags) eq_or_less_helperD not_less of_nat_numeral power_add semiring_1_class.of_nat_power unat_pow_le_intro word_unat.Rep_inverse) 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 && mask (size w - n)" by (cases "n \ size w"; clarsimp simp: word_and_le2 and_mask shiftl_zero_size) 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_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 add: mask_eq_decr_exp NOT_eq flip: mul_not_mask_eq_neg_shiftl) lemma ucast_eq_mask: "(UCAST('a::len \ 'b::len) x = UCAST('a \ 'b) y) = (x && mask LENGTH('b) = y && mask LENGTH('b))" by (rule iffI; word_eqI_solve) context fixes w :: "'a::len word" begin private lemma sbintrunc_uint_ucast: assumes "Suc n = LENGTH('b::len)" shows "sbintrunc n (uint (ucast w :: 'b word)) = sbintrunc n (uint w)" by (metis assms sbintrunc_bintrunc ucast_eq word_ubin.eq_norm) private lemma test_bit_sbintrunc: assumes "i < LENGTH('a)" shows "(word_of_int (sbintrunc n (uint w)) :: 'a word) !! i = (if n < i then w !! n else w !! i)" using assms by (simp add: nth_sbintr) (simp add: test_bit_bin) private lemma test_bit_sbintrunc_ucast: assumes len_a: "i < LENGTH('a)" shows "(word_of_int (sbintrunc (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) !! i = (if LENGTH('b::len) \ i then w !! (LENGTH('b) - 1) else w !! i)" apply (subst sbintrunc_uint_ucast) apply simp apply (subst test_bit_sbintrunc) apply (rule len_a) apply (rule if_cong[OF _ refl refl]) using leD less_linear by fastforce lemma scast_ucast_high_bits: - "scast (ucast w :: 'b::len word) = w - \ (\ i \ {LENGTH('b) ..< size w}. w !! i = w !! (LENGTH('b) - 1))" - unfolding scast_eq sint_uint word_size - apply (subst word_eq_iff) - apply (rule iffI) - apply (rule ballI) - apply (drule_tac x=i in spec) - apply (subst (asm) test_bit_sbintrunc_ucast; simp) - apply (rule allI) - apply (case_tac "n < LENGTH('a)") - apply (subst test_bit_sbintrunc_ucast) - apply simp - apply (case_tac "n \ LENGTH('b)") - apply (drule_tac x=n in bspec) - by auto - + \scast (ucast w :: 'b::len word) = w + \ (\ i \ {LENGTH('b) ..< size w}. w !! i = 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) \ ~~(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 && mask LENGTH('b))" by word_eqI_solve lemma ucast_NOT: "ucast (~~x) = ~~(ucast x) && 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) (~~x) = ~~(UCAST('a \ 'b) x)" by word_eqI 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) && ~~(mask n) = a" by (smt add.left_neutral add_diff_cancel_right' add_mask_lower_bits and_mask_plus is_aligned_mask is_aligned_weaken le_less_trans of_bl_length2 subtract_mask(1)) 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 ((a + of_bl b * 2^c) >> c)) = b" apply (subst word_plus_and_or_coroll) apply (erule is_aligned_get_word_bits) apply (rule is_aligned_AND_less_0) apply (simp add: is_aligned_mask) apply (rule order_less_le_trans) apply (rule of_bl_length2) apply simp apply (simp add: two_power_increasing) apply simp apply (rule nth_equalityI) apply (simp only: len_bin_to_bl) apply (clarsimp simp only: len_bin_to_bl nth_bin_to_bl word_test_bit_def[symmetric]) apply (simp add: nth_shiftr nth_shiftl shiftl_t2n[where n=c, simplified mult.commute, simplified, symmetric]) apply (simp add: is_aligned_nth[THEN iffD1, rule_format] test_bit_of_bl rev_nth) apply arith done end diff --git a/thys/Word_Lib/Word_Lemmas_32.thy b/thys/Word_Lib/Word_Lemmas_32.thy --- a/thys/Word_Lib/Word_Lemmas_32.thy +++ b/thys/Word_Lib/Word_Lemmas_32.thy @@ -1,305 +1,312 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Lemmas for Word Length 32" theory Word_Lemmas_32 imports Word_Lemmas Word_Setup_32 begin 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" apply (rule unat_less_helper) apply (rule order_le_less_trans, rule word_and_le1) apply (simp add: mask_eq) done 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)" apply (case_tac "sz < word_bits") apply (case_tac "2\sz") apply (clarsimp simp: word_size_def word_bits_def min_def mask_eq) apply (drule (2) Suc_div_unat_helper [where 'a=32 and sz=sz and us=2, simplified, symmetric]) apply (simp add: not_le word_size_def word_bits_def) apply (case_tac sz, simp add: unat_word_ariths) apply (case_tac nat, simp add: unat_word_ariths unat_mask_word32 min_def word_bits_def) apply simp apply (simp add: unat_word_ariths unat_mask_word32 min_def word_bits_def word_size_def) done 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::word8 assumes a: "a \ 0xFF" shows "ucast a \ (0xFF::word32)" proof assume "ucast a = (0xFF::word32)" also have "(0xFF::word32) = ucast (0xFF::word8)" 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 :: "word8" 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 :: word8))" 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 :: word8]" 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) :: sword16) = (of_nat x)" "(ucast (of_nat x :: word32) :: sword8) = (of_nat x)" "(ucast (of_nat x :: word16) :: sword16) = (of_nat x)" "(ucast (of_nat x :: word16) :: sword8) = (of_nat x)" "(ucast (of_nat x :: word8) :: sword8) = (of_nat x)" - by (auto simp: ucast_of_nat is_down) + apply (simp_all add: of_nat_take_bit) + apply (transfer, simp) + apply (transfer, simp) + apply (transfer, simp) + apply (transfer, simp) + apply (transfer, simp) + apply (transfer, simp) + done 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 word_bitwise clarsimp lemma has_zero_byte: "~~ (((((v::word32) && 0x7f7f7f7f) + 0x7f7f7f7f) || v) || 0x7f7f7f7f) \ 0 \ v && 0xff000000 = 0 \ v && 0xff0000 = 0 \ v && 0xff00 = 0 \ v && 0xff = 0" apply clarsimp apply word_bitwise by metis 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, clarsimp+) - apply (simp add: word_of_int uint_word_of_int) + 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)" apply (subst down_cast_same[symmetric]) apply (simp add:is_down)+ done lemma cast_down_s64: "(scast::64 sword \ 32 word) = (ucast::64 sword \ 32 word)" apply (subst down_cast_same[symmetric]) apply (simp add:is_down)+ done end diff --git a/thys/Word_Lib/Word_Lemmas_64.thy b/thys/Word_Lib/Word_Lemmas_64.thy --- a/thys/Word_Lib/Word_Lemmas_64.thy +++ b/thys/Word_Lib/Word_Lemmas_64.thy @@ -1,294 +1,301 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Lemmas for Word Length 64" theory Word_Lemmas_64 imports Word_Lemmas Word_Setup_64 begin 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" apply (rule unat_less_helper) apply (rule order_le_less_trans, rule word_and_le1) apply (simp add: mask_eq) done 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)" apply (case_tac "sz < word_bits") apply (case_tac "3\sz") apply (clarsimp simp: word_size_def word_bits_def min_def mask_eq) apply (drule (2) Suc_div_unat_helper [where 'a=64 and sz=sz and us=3, simplified, symmetric]) apply (simp add: not_le word_size_def word_bits_def) apply (case_tac sz, simp add: unat_word_ariths) apply (case_tac nat, simp add: unat_word_ariths unat_mask_word64 min_def word_bits_def) apply (case_tac nata, simp add: unat_word_ariths unat_mask_word64 word_bits_def) apply simp apply (simp add: unat_word_ariths unat_mask_word64 min_def word_bits_def word_size_def) done 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::word8 assumes a: "a \ 0xFF" shows "ucast a \ (0xFF::word64)" proof assume "ucast a = (0xFF::word64)" also have "(0xFF::word64) = ucast (0xFF::word8)" 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 :: word8))" 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 :: word8]" 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) :: sword16) = (of_nat x)" "(ucast (of_nat x :: word64) :: sword8) = (of_nat x)" "(ucast (of_nat x :: word16) :: sword16) = (of_nat x)" "(ucast (of_nat x :: word16) :: sword8) = (of_nat x)" "(ucast (of_nat x :: word8) :: sword8) = (of_nat x)" - by (auto simp: ucast_of_nat is_down) + apply (simp_all add: of_nat_take_bit) + apply (transfer, simp) + apply (transfer, simp) + apply (transfer, simp) + apply (transfer, simp) + apply (transfer, simp) + apply (transfer, simp) + done 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 word_bitwise clarsimp 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, clarsimp+) - apply (simp add: word_of_int uint_word_of_int) + 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] (* 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)" apply (subst down_cast_same[symmetric]) apply (simp add:is_down)+ done lemma cast_down_s64: "(scast::64 sword \ 32 word) = (ucast::64 sword \ 32 word)" apply (subst down_cast_same[symmetric]) apply (simp add:is_down)+ done end diff --git a/thys/Word_Lib/Word_Lib.thy b/thys/Word_Lib/Word_Lib.thy --- a/thys/Word_Lib/Word_Lib.thy +++ b/thys/Word_Lib/Word_Lib.thy @@ -1,717 +1,716 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Additional Word Operations" theory Word_Lib imports "HOL-Word.Misc_set_bit" Word_Syntax begin definition ptr_add :: "'a :: len word \ nat \ 'a word" where "ptr_add ptr n \ ptr + of_nat n" definition complement :: "'a :: len word \ 'a word" where "complement x \ ~~ x" definition alignUp :: "'a::len word \ nat \ 'a word" where "alignUp x n \ x + 2 ^ n - 1 && complement (2 ^ n - 1)" (* 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}" (* Haskellish names/syntax *) notation (input) test_bit ("testBit") definition w2byte :: "'a :: len word \ 8 word" where "w2byte \ ucast" (* * Signed division: when the result of a division is negative, * we will round towards zero instead of towards minus infinity. *) class signed_div = fixes sdiv :: "'a \ 'a \ 'a" (infixl "sdiv" 70) fixes smod :: "'a \ 'a \ 'a" (infixl "smod" 70) instantiation int :: signed_div begin definition "(a :: int) sdiv b \ sgn (a * b) * (abs a div abs b)" definition "(a :: int) smod b \ a - (a sdiv b) * b" instance .. end instantiation word :: (len) signed_div begin definition "(a :: ('a::len) word) sdiv b = word_of_int (sint a sdiv sint b)" definition "(a :: ('a::len) word) smod b = word_of_int (sint a smod sint b)" instance .. end (* Tests *) lemma "( 4 :: word32) sdiv 4 = 1" "(-4 :: word32) sdiv 4 = -1" "(-3 :: word32) sdiv 4 = 0" "( 3 :: word32) sdiv -4 = 0" "(-3 :: word32) sdiv -4 = 0" "(-5 :: word32) sdiv -4 = 1" "( 5 :: word32) sdiv -4 = -1" by (simp_all add: sdiv_word_def sdiv_int_def) lemma "( 4 :: word32) smod 4 = 0" "( 3 :: word32) smod 4 = 3" "(-3 :: word32) smod 4 = -3" "( 3 :: word32) smod -4 = 3" "(-3 :: word32) smod -4 = -3" "(-5 :: word32) smod -4 = -1" "( 5 :: word32) smod -4 = 1" by (simp_all add: smod_word_def smod_int_def sdiv_int_def) (* 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)))" 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 w !! n then w || ~~ (mask n) else w && mask n" definition sign_extended :: "nat \ 'a::len word \ bool" where "sign_extended n w \ \i. n < i \ i < size w \ w !! i = w !! n" lemma ptr_add_0 [simp]: "ptr_add ref 0 = ref " unfolding ptr_add_def by simp lemma shiftl_power: "(shiftl1 ^^ x) (y::'a::len word) = 2 ^ x * y" apply (induct x) apply simp apply (simp add: shiftl1_2t) done 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_ops_nth [simp]: shows word_or_nth: "(x || y) !! n = (x !! n \ y !! n)" and word_and_nth: "(x && y) !! n = (x !! n \ y !! n)" and word_xor_nth: "(x xor y) !! n = (x !! n \ y !! n)" by ((cases "n < size x", auto dest: test_bit_size simp: word_ops_nth_size word_size)[1])+ (* simp del to avoid warning on the simp add in iff *) declare test_bit_1 [simp del, iff] (* test: *) lemma "1 < (1024::32 word) \ 1 \ (1024::32 word)" by simp lemma and_not_mask: "w AND NOT (mask n) = (w >> n) << n" for w :: \'a::len word\ apply (rule word_eqI) apply (simp add : word_ops_nth_size word_size) apply (simp add : nth_shiftr nth_shiftl) by auto lemma and_mask: "w AND mask n = (w << (size w - n)) >> (size w - n)" for w :: \'a::len word\ apply (rule word_eqI) apply (simp add : word_ops_nth_size word_size) apply (simp add : nth_shiftr nth_shiftl) by auto lemma AND_twice [simp]: "(w && m) && m = w && m" by (simp add: word_eqI) lemma word_combine_masks: "w && m = z \ w && m' = z' \ w && (m || m') = (z || z')" by (auto simp: word_eq_iff) lemma nth_w2p_same: "(2^n :: 'a :: len word) !! n = (n < LENGTH('a))" by (simp add : nth_w2p) lemma p2_gt_0: "(0 < (2 ^ n :: 'a :: len word)) = (n < LENGTH('a))" apply (simp add : word_gt_0) apply safe apply (erule swap) apply (rule word_eqI) apply (simp add : nth_w2p) apply (drule word_eqD) apply simp apply (erule notE) apply (erule nth_w2p_same [THEN iffD2]) done lemmas uint_2p_alt = uint_2p [unfolded p2_gt_0] 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 (rule word_uint.Rep_inverse' [THEN sym]) apply (rule shiftr_div_2n) done lemmas less_def = less_eq [symmetric] lemmas le_def = not_less [symmetric, where ?'a = nat] lemmas p2_eq_0 [simp] = trans [OF eq_commute iffD2 [OF Not_eq_iff p2_gt_0, folded le_def, unfolded word_gt_0 not_not]] 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 (simp add : and_not_mask shiftr_div_2n_w shiftl_t2n word_size) lemma neg_mask_is_div: "w AND NOT (mask n) = (w div 2^n) * 2^n" for w :: \'a::len word\ apply (cases "n < size w") apply (erule neg_mask_is_div') apply (simp add: word_size) apply (frule p2_gt_0 [THEN Not_eq_iff [THEN iffD2], THEN iffD2]) apply (simp add: word_gt_0 del: p2_eq_0) apply (rule word_eqI) apply (simp add: word_ops_nth_size word_size) done 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 (simp add: and_mask shiftr_div_2n_w shiftl_t2n word_size mult.commute) lemmas p2len = iffD2 [OF p2_eq_0 order_refl] lemma and_mask_arith: "w AND mask n = (w * 2^(size w - n)) div 2^(size w - n)" for w :: \'a::len word\ apply (cases "0 < n") apply (auto elim!: and_mask_arith') apply (simp add: word_size) done 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 && mask n \ 2 ^ n - 1" by (simp add: mask_2pm1[symmetric] word_and_le1) lemma is_aligned_AND_less_0: "u && mask n = 0 \ v < 2^n \ u && v = 0" apply (drule less_mask_eq) apply (simp add: mask_2pm1) apply (rule word_eqI) apply (clarsimp simp add: word_size) apply (drule_tac x=na in word_eqD) apply (drule_tac x=na in word_eqD) apply simp done lemma le_shiftr1: "u <= v \ shiftr1 u <= shiftr1 v" apply (unfold word_le_def shiftr1_eq word_ubin.eq_norm) apply (unfold bin_rest_trunc_i trans [OF bintrunc_bintrunc_l word_ubin.norm_Rep, unfolded word_ubin.norm_Rep, OF order_refl [THEN le_SucI]]) apply (case_tac "uint u" rule: bin_exhaust) apply (rename_tac bs bit) apply (case_tac "uint v" rule: bin_exhaust) apply (rename_tac bs' bit') apply (case_tac "bit") apply (case_tac "bit'", auto simp: less_eq_int_code)[1] apply (case_tac bit') apply (simp add: less_eq_int_code) apply (simp add: less_eq_int_code) done lemma le_shiftr: "u \ v \ u >> (n :: nat) \ (v :: 'a :: len word) >> n" apply (unfold shiftr_def) apply (induct_tac "n") apply auto apply (erule le_shiftr1) done lemma shiftr_mask_le: "n <= m \ mask n >> m = (0 :: 'a::len word)" apply (rule word_eqI) apply (simp add: word_size nth_shiftr) done lemmas shiftr_mask = order_refl [THEN shiftr_mask_le, simp] lemma word_leI: "(\n. \n < size (u::'a::len word); u !! n \ \ (v::'a::len word) !! n) \ u <= v" apply (rule xtr4) apply (rule word_and_le2) apply (rule word_eqI) apply (simp add: word_ao_nth) apply safe apply assumption apply (erule_tac [2] asm_rl) apply (unfold word_size) by auto 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 xtr3) 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) 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 (drule_tac f = "%u. u !! (x - n)" in arg_cong) apply (simp add : nth_shiftr) apply (case_tac "n <= x") apply auto done lemmas and_mask_eq_iff_le_mask = trans [OF and_mask_eq_iff_shiftr_0 le_mask_iff [THEN sym]] lemma mask_shiftl_decompose: "mask m << n = mask (m + n) && ~~ (mask n)" by (auto intro!: word_eqI simp: and_not_mask nth_shiftl nth_shiftr word_size) lemma one_bit_shiftl: "set_bit 0 n True = (1 :: 'a :: len word) << n" apply (rule word_eqI) apply (auto simp add: test_bit_set_gen nth_shiftl word_size simp del: word_set_bit_0 shiftl_1) done lemmas one_bit_pow = trans [OF one_bit_shiftl shiftl_1] lemmas bin_sc_minus_simps = bin_sc_simps (2,3,4) [THEN [2] trans, OF bin_sc_minus [THEN sym]] 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 lemmas gt0_iff_gem1 = iffD1 [OF iffD1 [OF iff_left_commute le_m1_iff_lt] order_refl] lemmas power_2_ge_iff = trans [OF gt0_iff_gem1 [THEN sym] p2_gt_0] 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]) lemmas mask_lt_2pn = le_mask_iff_lt_2n [THEN iffD1, THEN iffD1, OF _ order_refl] lemma bang_eq: fixes x :: "'a::len word" shows "(x = y) = (\n. x !! n = y !! n)" by (subst test_bit_eq_iff[symmetric]) fastforce 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 shiftl_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: word_ao_nth nth_shiftl, safe) done lemma shiftr_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_shiftr word_ao_nth) done 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)" apply(rule word_eqI) apply(simp add:nth_shiftl word_ao_nth, safe) done lemma shiftr_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_shiftr word_ao_nth) done 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:nth_shiftr nth_shiftl word_size word_ao_nth) 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) done lemma shiftr_shiftl1: fixes a::"'a::len word" shows "c \ b \ a >> b << c = (a >> (b - c)) AND (NOT (mask c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ops_nth_size) done 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:nth_shiftr nth_shiftl word_size word_ops_nth_size) done lemmas multi_shift_simps = shiftl_shiftl shiftr_shiftr shiftl_shiftr1 shiftl_shiftr2 shiftr_shiftl1 shiftr_shiftl2 lemma word_and_max_word: fixes a::"'a::len word" shows "x = max_word \ a AND x = a" by simp lemma word_and_full_mask_simp: \x && Bit_Operations.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 && Bit_Operations.mask LENGTH('a)) n \ bit x n\ by (simp add: bit_and_iff bit_mask_iff) qed lemma word8_and_max_simp: \x && 0xFF = x\ for x :: \8 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) lemma word16_and_max_simp: \x && 0xFFFF = x\ for x :: \16 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) lemma word32_and_max_simp: \x && 0xFFFFFFFF = x\ for x :: \32 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) lemma word64_and_max_simp: \x && 0xFFFFFFFFFFFFFFFF = x\ for x :: \64 word\ using word_and_full_mask_simp [of x] by (simp add: numeral_eq_Suc mask_Suc_exp) lemmas word_and_max_simps = word8_and_max_simp word16_and_max_simp word32_and_max_simp word64_and_max_simp lemma word_and_1_bl: fixes x::"'a::len word" shows "(x AND 1) = of_bl [x !! 0]" by (simp add: mod_2_eq_odd test_bit_word_eq and_one_eq) lemma word_1_and_bl: fixes x::"'a::len word" shows "(1 AND x) = of_bl [x !! 0]" by (simp add: mod_2_eq_odd test_bit_word_eq one_and_eq) lemma scast_scast_id [simp]: "scast (scast x :: ('a::len) signed word) = (x :: 'a word)" "scast (scast y :: ('a::len) word) = (y :: 'a signed word)" by (auto simp: is_up scast_up_scast_id) lemma scast_ucast_id [simp]: "scast (ucast (x :: 'a::len word) :: 'a signed word) = x" by (metis down_cast_same is_down len_signed order_refl scast_scast_id(1)) lemma ucast_scast_id [simp]: "ucast (scast (x :: 'a::len signed word) :: 'a word) = x" by (metis scast_scast_id(2) scast_ucast_id) lemma scast_of_nat [simp]: "scast (of_nat x :: 'a::len signed word) = (of_nat x :: 'a word)" by transfer simp 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 (* 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) lemma word_sint_1 [simp]: "sint (1::'a::len word) = (if LENGTH('a) = 1 then -1 else 1)" by (cases \LENGTH('a)\) (simp_all add: not_le sint_uint le_Suc_eq sbintrunc_minus_simps) lemma scast_1': "(scast (1::'a::len word) :: 'b::len word) = (word_of_int (sbintrunc (LENGTH('a::len) - Suc 0) (1::int)))" by transfer simp -lemma scast_1 [simp]: +lemma scast_1: "(scast (1::'a::len word) :: 'b::len word) = (if LENGTH('a) = 1 then -1 else 1)" - by (clarsimp simp: scast_1') - (metis Suc_pred len_gt_0 nat.exhaust sbintrunc_Suc_numeral(1) uint_1 word_uint.Rep_inverse') + by (fact signed_1) lemma scast_eq_scast_id [simp]: "((scast (a :: 'a::len signed word) :: 'a word) = scast b) = (a = b)" by (metis ucast_scast_id) lemma ucast_eq_ucast_id [simp]: "((ucast (a :: 'a::len word) :: 'a signed word) = ucast b) = (a = b)" by (metis scast_ucast_id) lemma scast_ucast_norm [simp]: "(ucast (a :: 'a::len word) = (b :: 'a signed word)) = (a = scast b)" "((b :: 'a signed word) = ucast (a :: 'a::len word)) = (a = scast b)" by (metis scast_ucast_id ucast_scast_id)+ lemma of_bl_drop: "of_bl (drop n xs) = (of_bl xs && mask (length xs - n))" apply (clarsimp simp: bang_eq test_bit_of_bl rev_nth cong: rev_conj_cong) apply (safe; simp add: word_size to_bl_nth) done -lemma of_int_uint [simp]: +lemma of_int_uint: "of_int (uint x) = x" - by (metis word_of_int word_uint.Rep_inverse') + by (fact word_of_int_uint) lemma shiftr_mask2: "n \ LENGTH('a) \ (mask n >> m :: ('a :: len) word) = mask (n - m)" apply (rule word_eqI) apply (simp add: nth_shiftr word_size) apply arith done corollary word_plus_and_or_coroll: "x && y = 0 \ x + y = x || y" using word_plus_and_or[where x=x and y=y] by simp corollary word_plus_and_or_coroll2: "(x && w) + (x && ~~ w) = x" apply (subst word_plus_and_or_coroll) apply (rule word_eqI, simp add: word_size word_ops_nth_size) apply (rule word_eqI, simp add: word_size word_ops_nth_size) apply blast done 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 lemmas less_le_mult_nat = less_le_mult_nat'[simplified distrib_right, simplified] (* FIXME: these should eventually be moved to HOL/Word. *) 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 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 list_of_false: "True \ set xs \ xs = replicate (length xs) False" by (induct xs, simp_all) lemma eq_zero_set_bl: "(w = 0) = (True \ set (to_bl w))" using list_of_false word_bl.Rep_inject by fastforce lemma diff_diff_less: "(i < m - (m - (n :: nat))) = (i < m \ i < n)" by auto 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 end diff --git a/thys/Word_Lib/Word_Next.thy b/thys/Word_Lib/Word_Next.thy --- a/thys/Word_Lib/Word_Next.thy +++ b/thys/Word_Lib/Word_Next.thy @@ -1,94 +1,118 @@ (* SPDX-License-Identifier: BSD-3-Clause *) section\Increment and Decrement Machine Words Without Wrap-Around\ theory Word_Next imports Aligned begin text\Previous and next words addresses, without wrap around.\ definition word_next :: "'a::len word \ 'a::len word" where "word_next a \ if a = max_word then max_word else a + 1" definition word_prev :: "'a::len word \ 'a::len word" where "word_prev a \ if a = 0 then 0 else a - 1" text\Examples:\ +lemma [code]: + \Word.the_int (word_next w :: 'a::len word) = + (if w = - 1 then Word.the_int w else Word.the_int w + 1)\ + apply (simp add: word_next_def) apply transfer + apply (simp add: take_bit_minus_one_eq_mask mask_eq_exp_minus_1) + subgoal for k + using take_bit_incr_eq [of \LENGTH('a)\ k] + apply (simp add: ac_simps) + done + done + +lemma [code]: + \Word.the_int (word_prev w :: 'a::len word) = + (if w = 0 then Word.the_int w else Word.the_int w - 1)\ + apply (simp add: word_prev_def) + apply transfer + subgoal for k + apply rule + apply (subst take_bit_diff [symmetric]) + using take_bit_int_less_exp [of \LENGTH('a)\ k] + apply (simp add: take_bit_int_eq_self_iff less_le) + done + done + 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 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 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 \ max_word \ (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 \ max_word \ 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 \ max_word \ 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 \ max_word" 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 (auto simp add: le_less) lemma word_adjacent_union: "word_next e = s' \ s \ e \ s' \ e' \ {s..e} \ {s'..e'} = {s .. e'}" apply (simp add: word_next_def ivl_disj_un_two_touch split: if_splits) apply (drule sym) apply simp apply (subst word_atLeastLessThan_Suc_atLeastAtMost_union) apply (simp_all add: word_Suc_le) done end