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,1651 +1,1640 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) subsection \Finite Fields\ text \We provide four implementations for $GF(p)$ -- the field with $p$ elements for some prime $p$ -- one by int, one by integers, one by 32-bit numbers and one 64-bit implementation. Correctness of the implementations is proven by transfer rules to the type-based version of $GF(p)$.\ theory Finite_Field_Record_Based imports Finite_Field Arithmetic_Record_Based Native_Word.Uint32 Native_Word.Uint64 Native_Word.Code_Target_Bits_Int "HOL-Library.Code_Target_Numeral" begin (* mod on standard case which can immediately be mapped to target languages without considering special cases *) definition mod_nonneg_pos :: "integer \ integer \ integer" where "x \ 0 \ y > 0 \ mod_nonneg_pos x y = (x mod y)" code_printing \ \FIXME illusion of partiality\ constant mod_nonneg_pos \ (SML) "IntInf.mod/ ( _,/ _ )" and (Eval) "IntInf.mod/ ( _,/ _ )" and (OCaml) "Z.rem" and (Haskell) "Prelude.mod/ ( _ )/ ( _ )" and (Scala) "!((k: BigInt) => (l: BigInt) =>/ (k '% l))" definition mod_nonneg_pos_int :: "int \ int \ int" where "mod_nonneg_pos_int x y = int_of_integer (mod_nonneg_pos (integer_of_int x) (integer_of_int y))" lemma mod_nonneg_pos_int[simp]: "x \ 0 \ y > 0 \ mod_nonneg_pos_int x y = (x mod y)" unfolding mod_nonneg_pos_int_def using mod_nonneg_pos_def by simp context fixes p :: int begin definition plus_p :: "int \ int \ int" where "plus_p x y \ let z = x + y in if z \ p then z - p else z" definition minus_p :: "int \ int \ int" where "minus_p x y \ if y \ x then x - y else x + p - y" definition uminus_p :: "int \ int" where "uminus_p x = (if x = 0 then 0 else p - x)" definition mult_p :: "int \ int \ int" where "mult_p x y = (mod_nonneg_pos_int (x * y) p)" fun power_p :: "int \ nat \ int" where "power_p x n = (if n = 0 then 1 else let (d,r) = Divides.divmod_nat n 2; rec = power_p (mult_p x x) d in if r = 0 then rec else mult_p rec x)" text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p :: "int \ int" where "inverse_p x = (if x = 0 then 0 else power_p x (nat (p - 2)))" definition divide_p :: "int \ int \ int" where "divide_p x y = mult_p x (inverse_p y)" definition finite_field_ops_int :: "int arith_ops_record" where "finite_field_ops_int \ Arith_Ops_Record 0 1 plus_p mult_p minus_p uminus_p divide_p inverse_p (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) (\ x . x) (\ x . x) (\ x. 0 \ x \ x < p)" end context fixes p :: uint32 begin definition plus_p32 :: "uint32 \ uint32 \ uint32" where "plus_p32 x y \ let z = x + y in if z \ p then z - p else z" definition minus_p32 :: "uint32 \ uint32 \ uint32" where "minus_p32 x y \ if y \ x then x - y else (x + p) - y" definition uminus_p32 :: "uint32 \ uint32" where "uminus_p32 x = (if x = 0 then 0 else p - x)" definition mult_p32 :: "uint32 \ uint32 \ uint32" where "mult_p32 x y = (x * y mod p)" lemma int_of_uint32_shift: "int_of_uint32 (drop_bit k n) = (int_of_uint32 n) div (2 ^ k)" apply transfer apply transfer apply (simp add: take_bit_drop_bit min_def) apply (simp add: drop_bit_eq_div) done lemma int_of_uint32_0_iff: "int_of_uint32 n = 0 \ n = 0" by (transfer, rule uint_0_iff) lemma int_of_uint32_0: "int_of_uint32 0 = 0" unfolding int_of_uint32_0_iff by simp lemma int_of_uint32_ge_0: "int_of_uint32 n \ 0" by (transfer, auto) lemma two_32: "2 ^ LENGTH(32) = (4294967296 :: int)" by simp lemma int_of_uint32_plus: "int_of_uint32 (x + y) = (int_of_uint32 x + int_of_uint32 y) mod 4294967296" by (transfer, unfold uint_word_ariths two_32, rule refl) lemma int_of_uint32_minus: "int_of_uint32 (x - y) = (int_of_uint32 x - int_of_uint32 y) mod 4294967296" by (transfer, unfold uint_word_ariths two_32, rule refl) lemma int_of_uint32_mult: "int_of_uint32 (x * y) = (int_of_uint32 x * int_of_uint32 y) mod 4294967296" by (transfer, unfold uint_word_ariths two_32, rule refl) lemma int_of_uint32_mod: "int_of_uint32 (x mod y) = (int_of_uint32 x mod int_of_uint32 y)" by (transfer, unfold uint_mod two_32, rule refl) lemma int_of_uint32_inv: "0 \ x \ x < 4294967296 \ int_of_uint32 (uint32_of_int x) = x" by transfer (simp add: take_bit_int_eq_self) function power_p32 :: "uint32 \ uint32 \ uint32" where "power_p32 x n = (if n = 0 then 1 else let rec = power_p32 (mult_p32 x x) (drop_bit 1 n) in if n AND 1 = 0 then rec else mult_p32 rec x)" by pat_completeness auto termination proof - { fix n :: uint32 assume "n \ 0" with int_of_uint32_ge_0[of n] int_of_uint32_0_iff[of n] have "int_of_uint32 n > 0" by auto hence "0 < int_of_uint32 n" "int_of_uint32 n div 2 < int_of_uint32 n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_uint32 n))", auto simp: int_of_uint32_shift *) qed text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p32 :: "uint32 \ uint32" where "inverse_p32 x = (if x = 0 then 0 else power_p32 x (p - 2))" definition divide_p32 :: "uint32 \ uint32 \ uint32" where "divide_p32 x y = mult_p32 x (inverse_p32 y)" definition finite_field_ops32 :: "uint32 arith_ops_record" where "finite_field_ops32 \ Arith_Ops_Record 0 1 plus_p32 mult_p32 minus_p32 uminus_p32 divide_p32 inverse_p32 (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) uint32_of_int int_of_uint32 (\ x. 0 \ x \ x < p)" end lemma shiftr_uint32_code [code_unfold]: "drop_bit 1 x = (uint32_shiftr x 1)" by (simp add: uint32_shiftr_def) (* ******************************************************************************** *) subsubsection \Transfer Relation\ locale mod_ring_locale = fixes p :: int and ty :: "'a :: nontriv itself" assumes p: "p = int CARD('a)" begin lemma nat_p: "nat p = CARD('a)" unfolding p by simp lemma p2: "p \ 2" unfolding p using nontriv[where 'a = 'a] by auto lemma p2_ident: "int (CARD('a) - 2) = p - 2" using p2 unfolding p by simp definition mod_ring_rel :: "int \ 'a mod_ring \ bool" where "mod_ring_rel x x' = (x = to_int_mod_ring x')" (* domain transfer rules *) lemma Domainp_mod_ring_rel [transfer_domain_rule]: "Domainp (mod_ring_rel) = (\ v. v \ {0 ..< p})" proof - { fix v :: int assume *: "0 \ v" "v < p" have "Domainp mod_ring_rel v" proof show "mod_ring_rel v (of_int_mod_ring v)" unfolding mod_ring_rel_def using * p by auto qed } note * = this show ?thesis by (intro ext iffI, insert range_to_int_mod_ring[where 'a = 'a] *, auto simp: mod_ring_rel_def p) qed (* left/right/bi-unique *) lemma bi_unique_mod_ring_rel [transfer_rule]: "bi_unique mod_ring_rel" "left_unique mod_ring_rel" "right_unique mod_ring_rel" unfolding mod_ring_rel_def bi_unique_def left_unique_def right_unique_def by auto (* left/right-total *) lemma right_total_mod_ring_rel [transfer_rule]: "right_total mod_ring_rel" unfolding mod_ring_rel_def right_total_def by simp (* ************************************************************************************ *) subsubsection \Transfer Rules\ (* 0 / 1 *) lemma mod_ring_0[transfer_rule]: "mod_ring_rel 0 0" unfolding mod_ring_rel_def by simp lemma mod_ring_1[transfer_rule]: "mod_ring_rel 1 1" unfolding mod_ring_rel_def by simp (* addition *) lemma plus_p_mod_def: assumes x: "x \ {0 ..< p}" and y: "y \ {0 ..< p}" shows "plus_p p x y = ((x + y) mod p)" proof (cases "p \ x + y") case False thus ?thesis using x y unfolding plus_p_def Let_def by auto next case True from True x y have *: "p > 0" "0 \ x + y - p" "x + y - p < p" by auto from True have id: "plus_p p x y = x + y - p" unfolding plus_p_def by auto show ?thesis unfolding id using * using mod_pos_pos_trivial by fastforce qed lemma mod_ring_plus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (plus_p p) (+)" proof - { fix x y :: "'a mod_ring" have "plus_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x + y)" by (transfer, subst plus_p_mod_def, auto, auto simp: p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* subtraction *) lemma minus_p_mod_def: assumes x: "x \ {0 ..< p}" and y: "y \ {0 ..< p}" shows "minus_p p x y = ((x - y) mod p)" proof (cases "x - y < 0") case False thus ?thesis using x y unfolding minus_p_def Let_def by auto next case True from True x y have *: "p > 0" "0 \ x - y + p" "x - y + p < p" by auto from True have id: "minus_p p x y = x - y + p" unfolding minus_p_def by auto show ?thesis unfolding id using * using mod_pos_pos_trivial by fastforce qed lemma mod_ring_minus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (minus_p p) (-)" proof - { fix x y :: "'a mod_ring" have "minus_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x - y)" by (transfer, subst minus_p_mod_def, auto simp: p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* unary minus *) lemma mod_ring_uminus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (uminus_p p) uminus" proof - { fix x :: "'a mod_ring" have "uminus_p p (to_int_mod_ring x) = to_int_mod_ring (uminus x)" by (transfer, auto simp: uminus_p_def p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* multiplication *) lemma mod_ring_mult[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (mult_p p) ((*))" proof - { fix x y :: "'a mod_ring" have "mult_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x * y)" by (transfer, auto simp: mult_p_def p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* equality *) lemma mod_ring_eq[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> (=)) (=) (=)" by (intro rel_funI, auto simp: mod_ring_rel_def) (* power *) lemma mod_ring_power[transfer_rule]: "(mod_ring_rel ===> (=) ===> mod_ring_rel) (power_p p) (^)" proof (intro rel_funI, clarify, unfold binary_power[symmetric], goal_cases) fix x y n assume xy: "mod_ring_rel x y" from xy show "mod_ring_rel (power_p p x n) (binary_power y n)" proof (induct y n arbitrary: x rule: binary_power.induct) case (1 x n y) note 1(2)[transfer_rule] show ?case proof (cases "n = 0") case True thus ?thesis by (simp add: mod_ring_1) next case False obtain d r where id: "Divides.divmod_nat n 2 = (d,r)" by force let ?int = "power_p p (mult_p p y y) d" let ?gfp = "binary_power (x * x) d" from False have id': "?thesis = (mod_ring_rel (if r = 0 then ?int else mult_p p ?int y) (if r = 0 then ?gfp else ?gfp * x))" unfolding power_p.simps[of _ _ n] binary_power.simps[of _ n] Let_def id split by simp have [transfer_rule]: "mod_ring_rel ?int ?gfp" by (rule 1(1)[OF False refl id[symmetric]], transfer_prover) show ?thesis unfolding id' by transfer_prover qed qed qed declare power_p.simps[simp del] lemma ring_finite_field_ops_int: "ring_ops (finite_field_ops_int p) mod_ring_rel" by (unfold_locales, auto simp: finite_field_ops_int_def bi_unique_mod_ring_rel right_total_mod_ring_rel mod_ring_plus mod_ring_minus mod_ring_uminus mod_ring_mult mod_ring_eq mod_ring_0 mod_ring_1 Domainp_mod_ring_rel) end locale prime_field = mod_ring_locale p ty for p and ty :: "'a :: prime_card itself" begin lemma prime: "prime p" unfolding p using prime_card[where 'a = 'a] by simp (* mod *) lemma mod_ring_mod[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) ((\ x y. if y = 0 then x else 0)) (mod)" proof - { fix x y :: "'a mod_ring" have "(if to_int_mod_ring y = 0 then to_int_mod_ring x else 0) = to_int_mod_ring (x mod y)" unfolding modulo_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* normalize *) lemma mod_ring_normalize[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) ((\ x. if x = 0 then 0 else 1)) normalize" proof - { fix x :: "'a mod_ring" have "(if to_int_mod_ring x = 0 then 0 else 1) = to_int_mod_ring (normalize x)" unfolding normalize_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* unit_factor *) lemma mod_ring_unit_factor[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (\ x. x) unit_factor" proof - { fix x :: "'a mod_ring" have "to_int_mod_ring x = to_int_mod_ring (unit_factor x)" unfolding unit_factor_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* inverse *) lemma mod_ring_inverse[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (inverse_p p) inverse" proof (intro rel_funI) fix x y assume [transfer_rule]: "mod_ring_rel x y" show "mod_ring_rel (inverse_p p x) (inverse y)" unfolding inverse_p_def inverse_mod_ring_def apply (transfer_prover_start) apply (transfer_step)+ apply (unfold p2_ident) apply (rule refl) done qed (* division *) lemma mod_ring_divide[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (divide_p p) (/)" unfolding divide_p_def[abs_def] divide_mod_ring_def[abs_def] inverse_mod_ring_def[symmetric] by transfer_prover lemma mod_ring_rel_unsafe: assumes "x < CARD('a)" shows "mod_ring_rel (int x) (of_nat x)" "0 < x \ of_nat x \ (0 :: 'a mod_ring)" proof - have id: "of_nat x = (of_int (int x) :: 'a mod_ring)" by simp show "mod_ring_rel (int x) (of_nat x)" "0 < x \ of_nat x \ (0 :: 'a mod_ring)" unfolding id unfolding mod_ring_rel_def proof (auto simp add: assms of_int_of_int_mod_ring) assume "0 < x" with assms have "of_int_mod_ring (int x) \ (0 :: 'a mod_ring)" by (metis (no_types) less_imp_of_nat_less less_irrefl of_nat_0_le_iff of_nat_0_less_iff to_int_mod_ring_hom.hom_zero to_int_mod_ring_of_int_mod_ring) thus "of_int_mod_ring (int x) = (0 :: 'a mod_ring) \ False" by blast qed qed lemma finite_field_ops_int: "field_ops (finite_field_ops_int p) mod_ring_rel" by (unfold_locales, auto simp: finite_field_ops_int_def bi_unique_mod_ring_rel right_total_mod_ring_rel mod_ring_divide mod_ring_plus mod_ring_minus mod_ring_uminus mod_ring_inverse mod_ring_mod mod_ring_unit_factor mod_ring_normalize mod_ring_mult mod_ring_eq mod_ring_0 mod_ring_1 Domainp_mod_ring_rel) end text \Once we have proven the soundness of the implementation, we do not care any longer that @{typ "'a mod_ring"} has been defined internally via lifting. Disabling the transfer-rules will hide the internal definition in further applications of transfer.\ lifting_forget mod_ring.lifting text \For soundness of the 32-bit implementation, we mainly prove that this implementation implements the int-based implementation of the mod-ring.\ context mod_ring_locale begin context fixes pp :: "uint32" assumes ppp: "p = int_of_uint32 pp" and small: "p \ 65535" begin lemmas uint32_simps = int_of_uint32_0 int_of_uint32_plus int_of_uint32_minus int_of_uint32_mult definition urel32 :: "uint32 \ int \ bool" where "urel32 x y = (y = int_of_uint32 x \ y < p)" definition mod_ring_rel32 :: "uint32 \ 'a mod_ring \ bool" where "mod_ring_rel32 x y = (\ z. urel32 x z \ mod_ring_rel z y)" lemma urel32_0: "urel32 0 0" unfolding urel32_def using p2 by (simp, transfer, simp) lemma urel32_1: "urel32 1 1" unfolding urel32_def using p2 by (simp, transfer, simp) lemma le_int_of_uint32: "(x \ y) = (int_of_uint32 x \ int_of_uint32 y)" by (transfer, simp add: word_le_def) lemma urel32_plus: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (plus_p32 pp x x') (plus_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" let ?p = "int_of_uint32 pp" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel32_def by auto have le: "(pp \ x + x') = (?p \ ?x + ?x')" unfolding le_int_of_uint32 using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?p \ ?x + ?x'") case True hence True: "(?p \ ?x + ?x') = True" by simp show ?thesis unfolding id using small rel unfolding plus_p32_def plus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?p \ ?x + ?x') = False" by simp show ?thesis unfolding id using small rel unfolding plus_p32_def plus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_minus: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (minus_p32 pp x x') (minus_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel32_def by auto have le: "(x' \ x) = (?x' \ ?x)" unfolding le_int_of_uint32 using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?x' \ ?x") case True hence True: "(?x' \ ?x) = True" by simp show ?thesis unfolding id using small rel unfolding minus_p32_def minus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?x' \ ?x) = False" by simp show ?thesis unfolding id using small rel unfolding minus_p32_def minus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_uminus: assumes "urel32 x y" shows "urel32 (uminus_p32 pp x) (uminus_p p y)" proof - let ?x = "int_of_uint32 x" from assms int_of_uint32_ge_0 have id: "y = ?x" and rel: "0 \ ?x" "?x < p" unfolding urel32_def by auto have le: "(x = 0) = (?x = 0)" unfolding int_of_uint32_0_iff using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?x = 0") case True hence True: "(?x = 0) = True" by simp show ?thesis unfolding id using small rel unfolding uminus_p32_def uminus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?x = 0) = False" by simp show ?thesis unfolding id using small rel unfolding uminus_p32_def uminus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_mult: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (mult_p32 pp x x') (mult_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' < p" unfolding urel32_def by auto from rel have "?x * ?x' < p * p" by (metis mult_strict_mono') also have "\ \ 65536 * 65536" by (rule mult_mono, insert p2 small, auto) finally have le: "?x * ?x' < 4294967296" by simp show ?thesis unfolding id using small rel unfolding mult_p32_def mult_p_def Let_def urel32_def unfolding ppp by (auto simp: uint32_simps, unfold int_of_uint32_mod int_of_uint32_mult, subst mod_pos_pos_trivial[of _ 4294967296], insert le, auto) qed lemma urel32_eq: assumes "urel32 x y" "urel32 x' y'" shows "(x = x') = (y = y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" unfolding urel32_def by auto show ?thesis unfolding id by (transfer, transfer) rule qed lemma urel32_normalize: assumes x: "urel32 x y" shows "urel32 (if x = 0 then 0 else 1) (if y = 0 then 0 else 1)" unfolding urel32_eq[OF x urel32_0] using urel32_0 urel32_1 by auto lemma urel32_mod: assumes x: "urel32 x x'" and y: "urel32 y y'" shows "urel32 (if y = 0 then x else 0) (if y' = 0 then x' else 0)" unfolding urel32_eq[OF y urel32_0] using urel32_0 x by auto lemma urel32_power: "urel32 x x' \ urel32 y (int y') \ urel32 (power_p32 pp x y) (power_p p x' y')" proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) case (1 x' y' x y) note x = 1(2) note y = 1(3) show ?case proof (cases "y' = 0") case True hence y: "y = 0" using urel32_eq[OF y urel32_0] by auto show ?thesis unfolding y True by (simp add: power_p.simps urel32_1) next case False hence id: "(y = 0) = False" "(y' = 0) = False" using urel32_eq[OF y urel32_0] by auto + from y have \int y' = int_of_uint32 y\ \int y' < p\ + by (simp_all add: urel32_def) obtain d' r' where dr': "Divides.divmod_nat y' 2 = (d',r')" by force from divmod_nat_def[of y' 2, unfolded dr'] have r': "r' = y' mod 2" and d': "d' = y' div 2" by auto have "urel32 (y AND 1) r'" - 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 + using \int y' < p\ small + apply (simp add: urel32_def and_one_eq r') + apply (auto simp add: ppp and_one_eq) + apply (simp add: of_nat_mod int_of_uint32.rep_eq modulo_uint32.rep_eq uint_mod \int y' = int_of_uint32 y\) done from urel32_eq[OF this urel32_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel32 (drop_bit 1 y) (int d')" unfolding d' using y unfolding urel32_def using small unfolding ppp apply transfer apply transfer apply (auto simp add: drop_bit_Suc take_bit_int_eq_self) done note IH = 1(1)[OF False refl dr'[symmetric] urel32_mult[OF x x] div] show ?thesis unfolding power_p.simps[of _ _ "y'"] power_p32.simps[of _ _ y] dr' id if_False rem using IH urel32_mult[OF IH x] by (auto simp: Let_def) qed qed lemma urel32_inverse: assumes x: "urel32 x x'" shows "urel32 (inverse_p32 pp x) (inverse_p p x')" proof - have p: "urel32 (pp - 2) (int (nat (p - 2)))" using p2 small unfolding urel32_def unfolding ppp - by (transfer, auto simp: uint_word_ariths) + by (simp add: int_of_uint32.rep_eq minus_uint32.rep_eq uint_sub_if') show ?thesis unfolding inverse_p32_def inverse_p_def urel32_eq[OF x urel32_0] using urel32_0 urel32_power[OF x p] by auto qed lemma mod_ring_0_32: "mod_ring_rel32 0 0" using urel32_0 mod_ring_0 unfolding mod_ring_rel32_def by blast lemma mod_ring_1_32: "mod_ring_rel32 1 1" using urel32_1 mod_ring_1 unfolding mod_ring_rel32_def by blast lemma mod_ring_uminus32: "(mod_ring_rel32 ===> mod_ring_rel32) (uminus_p32 pp) uminus" using urel32_uminus mod_ring_uminus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_plus32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (plus_p32 pp) (+)" using urel32_plus mod_ring_plus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_minus32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (minus_p32 pp) (-)" using urel32_minus mod_ring_minus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_mult32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (mult_p32 pp) ((*))" using urel32_mult mod_ring_mult unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_eq32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> (=)) (=) (=)" using urel32_eq mod_ring_eq unfolding mod_ring_rel32_def rel_fun_def by blast lemma urel32_inj: "urel32 x y \ urel32 x z \ y = z" using urel32_eq[of x y x z] by auto lemma urel32_inj': "urel32 x z \ urel32 y z \ x = y" using urel32_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel32: "bi_unique mod_ring_rel32" "left_unique mod_ring_rel32" "right_unique mod_ring_rel32" using bi_unique_mod_ring_rel urel32_inj' unfolding mod_ring_rel32_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel32_def) lemma right_total_mod_ring_rel32: "right_total mod_ring_rel32" unfolding mod_ring_rel32_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel32 (uint32_of_int z) z" unfolding urel32_def using small unfolding ppp by (auto simp: int_of_uint32_inv) with zy show "\ x z. urel32 x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel32: "Domainp mod_ring_rel32 = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel32 x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel32_def proof let ?i = "int_of_uint32" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" using small unfolding ppp by (transfer, auto simp: word_less_def) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel32 a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel32 x (?i x)" unfolding urel32_def using small * unfolding ppp by (transfer, auto simp: word_less_def) qed next assume "\a b. x = a \ (\z. urel32 a z \ mod_ring_rel z b)" then obtain b z where xz: "urel32 x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel32_def using small unfolding ppp by (transfer, auto simp: word_less_def) qed qed lemma ring_finite_field_ops32: "ring_ops (finite_field_ops32 pp) mod_ring_rel32" by (unfold_locales, auto simp: finite_field_ops32_def bi_unique_mod_ring_rel32 right_total_mod_ring_rel32 mod_ring_plus32 mod_ring_minus32 mod_ring_uminus32 mod_ring_mult32 mod_ring_eq32 mod_ring_0_32 mod_ring_1_32 Domainp_mod_ring_rel32) end end context prime_field begin context fixes pp :: "uint32" assumes *: "p = int_of_uint32 pp" "p \ 65535" begin lemma mod_ring_normalize32: "(mod_ring_rel32 ===> mod_ring_rel32) (\x. if x = 0 then 0 else 1) normalize" using urel32_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_mod32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (\x y. if y = 0 then x else 0) (mod)" using urel32_mod[OF *] mod_ring_mod unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor32: "(mod_ring_rel32 ===> mod_ring_rel32) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_inverse32: "(mod_ring_rel32 ===> mod_ring_rel32) (inverse_p32 pp) inverse" using urel32_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_divide32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (divide_p32 pp) (/)" using mod_ring_inverse32 mod_ring_mult32[OF *] unfolding divide_p32_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops32: "field_ops (finite_field_ops32 pp) mod_ring_rel32" by (unfold_locales, insert ring_finite_field_ops32[OF *], auto simp: ring_ops_def finite_field_ops32_def mod_ring_divide32 mod_ring_inverse32 mod_ring_mod32 mod_ring_normalize32) end end (* now there is 64-bit time *) context fixes p :: uint64 begin definition plus_p64 :: "uint64 \ uint64 \ uint64" where "plus_p64 x y \ let z = x + y in if z \ p then z - p else z" definition minus_p64 :: "uint64 \ uint64 \ uint64" where "minus_p64 x y \ if y \ x then x - y else (x + p) - y" definition uminus_p64 :: "uint64 \ uint64" where "uminus_p64 x = (if x = 0 then 0 else p - x)" definition mult_p64 :: "uint64 \ uint64 \ uint64" where "mult_p64 x y = (x * y mod p)" lemma int_of_uint64_shift: "int_of_uint64 (drop_bit k n) = (int_of_uint64 n) div (2 ^ k)" apply transfer apply transfer apply (simp add: take_bit_drop_bit min_def) apply (simp add: drop_bit_eq_div) done lemma int_of_uint64_0_iff: "int_of_uint64 n = 0 \ n = 0" by (transfer, rule uint_0_iff) lemma int_of_uint64_0: "int_of_uint64 0 = 0" unfolding int_of_uint64_0_iff by simp lemma int_of_uint64_ge_0: "int_of_uint64 n \ 0" by (transfer, auto) lemma two_64: "2 ^ LENGTH(64) = (18446744073709551616 :: int)" by simp lemma int_of_uint64_plus: "int_of_uint64 (x + y) = (int_of_uint64 x + int_of_uint64 y) mod 18446744073709551616" by (transfer, unfold uint_word_ariths two_64, rule refl) lemma int_of_uint64_minus: "int_of_uint64 (x - y) = (int_of_uint64 x - int_of_uint64 y) mod 18446744073709551616" by (transfer, unfold uint_word_ariths two_64, rule refl) lemma int_of_uint64_mult: "int_of_uint64 (x * y) = (int_of_uint64 x * int_of_uint64 y) mod 18446744073709551616" by (transfer, unfold uint_word_ariths two_64, rule refl) lemma int_of_uint64_mod: "int_of_uint64 (x mod y) = (int_of_uint64 x mod int_of_uint64 y)" by (transfer, unfold uint_mod two_64, rule refl) lemma int_of_uint64_inv: "0 \ x \ x < 18446744073709551616 \ int_of_uint64 (uint64_of_int x) = x" by transfer (simp add: take_bit_int_eq_self) function power_p64 :: "uint64 \ uint64 \ uint64" where "power_p64 x n = (if n = 0 then 1 else let rec = power_p64 (mult_p64 x x) (drop_bit 1 n) in if n AND 1 = 0 then rec else mult_p64 rec x)" by pat_completeness auto termination proof - { fix n :: uint64 assume "n \ 0" with int_of_uint64_ge_0[of n] int_of_uint64_0_iff[of n] have "int_of_uint64 n > 0" by auto hence "0 < int_of_uint64 n" "int_of_uint64 n div 2 < int_of_uint64 n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_uint64 n))", auto simp: int_of_uint64_shift *) qed text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p64 :: "uint64 \ uint64" where "inverse_p64 x = (if x = 0 then 0 else power_p64 x (p - 2))" definition divide_p64 :: "uint64 \ uint64 \ uint64" where "divide_p64 x y = mult_p64 x (inverse_p64 y)" definition finite_field_ops64 :: "uint64 arith_ops_record" where "finite_field_ops64 \ Arith_Ops_Record 0 1 plus_p64 mult_p64 minus_p64 uminus_p64 divide_p64 inverse_p64 (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) uint64_of_int int_of_uint64 (\ x. 0 \ x \ x < p)" end lemma shiftr_uint64_code [code_unfold]: "drop_bit 1 x = (uint64_shiftr x 1)" by (simp add: uint64_shiftr_def) text \For soundness of the 64-bit implementation, we mainly prove that this implementation implements the int-based implementation of GF(p).\ context mod_ring_locale begin context fixes pp :: "uint64" assumes ppp: "p = int_of_uint64 pp" and small: "p \ 4294967295" begin lemmas uint64_simps = int_of_uint64_0 int_of_uint64_plus int_of_uint64_minus int_of_uint64_mult definition urel64 :: "uint64 \ int \ bool" where "urel64 x y = (y = int_of_uint64 x \ y < p)" definition mod_ring_rel64 :: "uint64 \ 'a mod_ring \ bool" where "mod_ring_rel64 x y = (\ z. urel64 x z \ mod_ring_rel z y)" lemma urel64_0: "urel64 0 0" unfolding urel64_def using p2 by (simp, transfer, simp) lemma urel64_1: "urel64 1 1" unfolding urel64_def using p2 by (simp, transfer, simp) lemma le_int_of_uint64: "(x \ y) = (int_of_uint64 x \ int_of_uint64 y)" by (transfer, simp add: word_le_def) lemma urel64_plus: assumes "urel64 x y" "urel64 x' y'" shows "urel64 (plus_p64 pp x x') (plus_p p y y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" let ?p = "int_of_uint64 pp" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel64_def by auto have le: "(pp \ x + x') = (?p \ ?x + ?x')" unfolding le_int_of_uint64 using rel small by (auto simp: uint64_simps) show ?thesis proof (cases "?p \ ?x + ?x'") case True hence True: "(?p \ ?x + ?x') = True" by simp show ?thesis unfolding id using small rel unfolding plus_p64_def plus_p_def Let_def urel64_def unfolding ppp le True if_True using True by (auto simp: uint64_simps) next case False hence False: "(?p \ ?x + ?x') = False" by simp show ?thesis unfolding id using small rel unfolding plus_p64_def plus_p_def Let_def urel64_def unfolding ppp le False if_False using False by (auto simp: uint64_simps) qed qed lemma urel64_minus: assumes "urel64 x y" "urel64 x' y'" shows "urel64 (minus_p64 pp x x') (minus_p p y y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel64_def by auto have le: "(x' \ x) = (?x' \ ?x)" unfolding le_int_of_uint64 using rel small by (auto simp: uint64_simps) show ?thesis proof (cases "?x' \ ?x") case True hence True: "(?x' \ ?x) = True" by simp show ?thesis unfolding id using small rel unfolding minus_p64_def minus_p_def Let_def urel64_def unfolding ppp le True if_True using True by (auto simp: uint64_simps) next case False hence False: "(?x' \ ?x) = False" by simp show ?thesis unfolding id using small rel unfolding minus_p64_def minus_p_def Let_def urel64_def unfolding ppp le False if_False using False by (auto simp: uint64_simps) qed qed lemma urel64_uminus: assumes "urel64 x y" shows "urel64 (uminus_p64 pp x) (uminus_p p y)" proof - let ?x = "int_of_uint64 x" from assms int_of_uint64_ge_0 have id: "y = ?x" and rel: "0 \ ?x" "?x < p" unfolding urel64_def by auto have le: "(x = 0) = (?x = 0)" unfolding int_of_uint64_0_iff using rel small by (auto simp: uint64_simps) show ?thesis proof (cases "?x = 0") case True hence True: "(?x = 0) = True" by simp show ?thesis unfolding id using small rel unfolding uminus_p64_def uminus_p_def Let_def urel64_def unfolding ppp le True if_True using True by (auto simp: uint64_simps) next case False hence False: "(?x = 0) = False" by simp show ?thesis unfolding id using small rel unfolding uminus_p64_def uminus_p_def Let_def urel64_def unfolding ppp le False if_False using False by (auto simp: uint64_simps) qed qed lemma urel64_mult: assumes "urel64 x y" "urel64 x' y'" shows "urel64 (mult_p64 pp x x') (mult_p p y y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' < p" unfolding urel64_def by auto from rel have "?x * ?x' < p * p" by (metis mult_strict_mono') also have "\ \ 4294967296 * 4294967296" by (rule mult_mono, insert p2 small, auto) finally have le: "?x * ?x' < 18446744073709551616" by simp show ?thesis unfolding id using small rel unfolding mult_p64_def mult_p_def Let_def urel64_def unfolding ppp by (auto simp: uint64_simps, unfold int_of_uint64_mod int_of_uint64_mult, subst mod_pos_pos_trivial[of _ 18446744073709551616], insert le, auto) qed lemma urel64_eq: assumes "urel64 x y" "urel64 x' y'" shows "(x = x') = (y = y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" unfolding urel64_def by auto show ?thesis unfolding id by (transfer, transfer) rule qed lemma urel64_normalize: assumes x: "urel64 x y" shows "urel64 (if x = 0 then 0 else 1) (if y = 0 then 0 else 1)" unfolding urel64_eq[OF x urel64_0] using urel64_0 urel64_1 by auto lemma urel64_mod: assumes x: "urel64 x x'" and y: "urel64 y y'" shows "urel64 (if y = 0 then x else 0) (if y' = 0 then x' else 0)" unfolding urel64_eq[OF y urel64_0] using urel64_0 x by auto lemma urel64_power: "urel64 x x' \ urel64 y (int y') \ urel64 (power_p64 pp x y) (power_p p x' y')" proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) case (1 x' y' x y) note x = 1(2) note y = 1(3) show ?case proof (cases "y' = 0") case True hence y: "y = 0" using urel64_eq[OF y urel64_0] by auto show ?thesis unfolding y True by (simp add: power_p.simps urel64_1) next case False hence id: "(y = 0) = False" "(y' = 0) = False" using urel64_eq[OF y urel64_0] by auto + from y have \int y' = int_of_uint64 y\ \int y' < p\ + by (simp_all add: urel64_def) obtain d' r' where dr': "Divides.divmod_nat y' 2 = (d',r')" by force from divmod_nat_def[of y' 2, unfolded dr'] have r': "r' = y' mod 2" and d': "d' = y' div 2" by auto have "urel64 (y AND 1) r'" - 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 (auto simp add: less_le) - apply (auto simp add: le_less) - apply (metis linorder_neqE_linordered_idom mod_pos_pos_trivial not_take_bit_negative power_0 take_bit_0 take_bit_eq_mod take_bit_nonnegative) - apply (metis even_take_bit_eq mod_pos_pos_trivial neq0_conv numeral_eq_Suc power_0 take_bit_eq_mod take_bit_nonnegative zero_less_Suc) + using \int y' < p\ small + apply (simp add: urel64_def and_one_eq r') + apply (auto simp add: ppp and_one_eq) + apply (simp add: of_nat_mod int_of_uint64.rep_eq modulo_uint64.rep_eq uint_mod \int y' = int_of_uint64 y\) done from urel64_eq[OF this urel64_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel64 (drop_bit 1 y) (int d')" unfolding d' using y unfolding urel64_def using small unfolding ppp apply transfer apply transfer apply (auto simp add: drop_bit_Suc take_bit_int_eq_self) done note IH = 1(1)[OF False refl dr'[symmetric] urel64_mult[OF x x] div] show ?thesis unfolding power_p.simps[of _ _ "y'"] power_p64.simps[of _ _ y] dr' id if_False rem using IH urel64_mult[OF IH x] by (auto simp: Let_def) qed qed lemma urel64_inverse: assumes x: "urel64 x x'" shows "urel64 (inverse_p64 pp x) (inverse_p p x')" proof - have p: "urel64 (pp - 2) (int (nat (p - 2)))" using p2 small unfolding urel64_def unfolding ppp - by (transfer, auto simp: uint_word_ariths) + by (simp add: int_of_uint64.rep_eq minus_uint64.rep_eq uint_sub_if') show ?thesis unfolding inverse_p64_def inverse_p_def urel64_eq[OF x urel64_0] using urel64_0 urel64_power[OF x p] by auto qed lemma mod_ring_0_64: "mod_ring_rel64 0 0" using urel64_0 mod_ring_0 unfolding mod_ring_rel64_def by blast lemma mod_ring_1_64: "mod_ring_rel64 1 1" using urel64_1 mod_ring_1 unfolding mod_ring_rel64_def by blast lemma mod_ring_uminus64: "(mod_ring_rel64 ===> mod_ring_rel64) (uminus_p64 pp) uminus" using urel64_uminus mod_ring_uminus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_plus64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (plus_p64 pp) (+)" using urel64_plus mod_ring_plus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_minus64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (minus_p64 pp) (-)" using urel64_minus mod_ring_minus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_mult64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (mult_p64 pp) ((*))" using urel64_mult mod_ring_mult unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_eq64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> (=)) (=) (=)" using urel64_eq mod_ring_eq unfolding mod_ring_rel64_def rel_fun_def by blast lemma urel64_inj: "urel64 x y \ urel64 x z \ y = z" using urel64_eq[of x y x z] by auto lemma urel64_inj': "urel64 x z \ urel64 y z \ x = y" using urel64_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel64: "bi_unique mod_ring_rel64" "left_unique mod_ring_rel64" "right_unique mod_ring_rel64" using bi_unique_mod_ring_rel urel64_inj' unfolding mod_ring_rel64_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel64_def) lemma right_total_mod_ring_rel64: "right_total mod_ring_rel64" unfolding mod_ring_rel64_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel64 (uint64_of_int z) z" unfolding urel64_def using small unfolding ppp by (auto simp: int_of_uint64_inv) with zy show "\ x z. urel64 x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel64: "Domainp mod_ring_rel64 = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel64 x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel64_def proof let ?i = "int_of_uint64" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" using small unfolding ppp by (transfer, auto simp: word_less_def) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel64 a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel64 x (?i x)" unfolding urel64_def using small * unfolding ppp by (transfer, auto simp: word_less_def) qed next assume "\a b. x = a \ (\z. urel64 a z \ mod_ring_rel z b)" then obtain b z where xz: "urel64 x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel64_def using small unfolding ppp by (transfer, auto simp: word_less_def) qed qed lemma ring_finite_field_ops64: "ring_ops (finite_field_ops64 pp) mod_ring_rel64" by (unfold_locales, auto simp: finite_field_ops64_def bi_unique_mod_ring_rel64 right_total_mod_ring_rel64 mod_ring_plus64 mod_ring_minus64 mod_ring_uminus64 mod_ring_mult64 mod_ring_eq64 mod_ring_0_64 mod_ring_1_64 Domainp_mod_ring_rel64) end end context prime_field begin context fixes pp :: "uint64" assumes *: "p = int_of_uint64 pp" "p \ 4294967295" begin lemma mod_ring_normalize64: "(mod_ring_rel64 ===> mod_ring_rel64) (\x. if x = 0 then 0 else 1) normalize" using urel64_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_mod64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (\x y. if y = 0 then x else 0) (mod)" using urel64_mod[OF *] mod_ring_mod unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor64: "(mod_ring_rel64 ===> mod_ring_rel64) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_inverse64: "(mod_ring_rel64 ===> mod_ring_rel64) (inverse_p64 pp) inverse" using urel64_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_divide64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (divide_p64 pp) (/)" using mod_ring_inverse64 mod_ring_mult64[OF *] unfolding divide_p64_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops64: "field_ops (finite_field_ops64 pp) mod_ring_rel64" by (unfold_locales, insert ring_finite_field_ops64[OF *], auto simp: ring_ops_def finite_field_ops64_def mod_ring_divide64 mod_ring_inverse64 mod_ring_mod64 mod_ring_normalize64) end end (* and a final implementation via integer *) context fixes p :: integer begin definition plus_p_integer :: "integer \ integer \ integer" where "plus_p_integer x y \ let z = x + y in if z \ p then z - p else z" definition minus_p_integer :: "integer \ integer \ integer" where "minus_p_integer x y \ if y \ x then x - y else (x + p) - y" definition uminus_p_integer :: "integer \ integer" where "uminus_p_integer x = (if x = 0 then 0 else p - x)" definition mult_p_integer :: "integer \ integer \ integer" where "mult_p_integer x y = (x * y mod p)" lemma int_of_integer_0_iff: "int_of_integer n = 0 \ n = 0" using integer_eqI by auto lemma int_of_integer_0: "int_of_integer 0 = 0" unfolding int_of_integer_0_iff by simp lemma int_of_integer_plus: "int_of_integer (x + y) = (int_of_integer x + int_of_integer y)" by simp lemma int_of_integer_minus: "int_of_integer (x - y) = (int_of_integer x - int_of_integer y)" by simp lemma int_of_integer_mult: "int_of_integer (x * y) = (int_of_integer x * int_of_integer y)" by simp lemma int_of_integer_mod: "int_of_integer (x mod y) = (int_of_integer x mod int_of_integer y)" by simp lemma int_of_integer_inv: "int_of_integer (integer_of_int x) = x" by simp lemma int_of_integer_shift: "int_of_integer (drop_bit k n) = (int_of_integer n) div (2 ^ k)" by transfer (simp add: int_of_integer_pow shiftr_integer_conv_div_pow2) function power_p_integer :: "integer \ integer \ integer" where "power_p_integer x n = (if n \ 0 then 1 else let rec = power_p_integer (mult_p_integer x x) (drop_bit 1 n) in if n AND 1 = 0 then rec else mult_p_integer rec x)" by pat_completeness auto termination proof - { fix n :: integer assume "\ (n \ 0)" hence "n > 0" by auto hence "int_of_integer n > 0" by (simp add: less_integer.rep_eq) hence "0 < int_of_integer n" "int_of_integer n div 2 < int_of_integer n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_integer n))", auto simp: * int_of_integer_shift) qed text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p_integer :: "integer \ integer" where "inverse_p_integer x = (if x = 0 then 0 else power_p_integer x (p - 2))" definition divide_p_integer :: "integer \ integer \ integer" where "divide_p_integer x y = mult_p_integer x (inverse_p_integer y)" definition finite_field_ops_integer :: "integer arith_ops_record" where "finite_field_ops_integer \ Arith_Ops_Record 0 1 plus_p_integer mult_p_integer minus_p_integer uminus_p_integer divide_p_integer inverse_p_integer (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) integer_of_int int_of_integer (\ x. 0 \ x \ x < p)" end lemma shiftr_integer_code [code_unfold]: "drop_bit 1 x = (integer_shiftr x 1)" unfolding shiftr_integer_code using integer_of_nat_1 by auto text \For soundness of the integer implementation, we mainly prove that this implementation implements the int-based implementation of GF(p).\ context mod_ring_locale begin context fixes pp :: "integer" assumes ppp: "p = int_of_integer pp" begin lemmas integer_simps = int_of_integer_0 int_of_integer_plus int_of_integer_minus int_of_integer_mult definition urel_integer :: "integer \ int \ bool" where "urel_integer x y = (y = int_of_integer x \ y \ 0 \ y < p)" definition mod_ring_rel_integer :: "integer \ 'a mod_ring \ bool" where "mod_ring_rel_integer x y = (\ z. urel_integer x z \ mod_ring_rel z y)" lemma urel_integer_0: "urel_integer 0 0" unfolding urel_integer_def using p2 by simp lemma urel_integer_1: "urel_integer 1 1" unfolding urel_integer_def using p2 by simp lemma le_int_of_integer: "(x \ y) = (int_of_integer x \ int_of_integer y)" by (rule less_eq_integer.rep_eq) lemma urel_integer_plus: assumes "urel_integer x y" "urel_integer x' y'" shows "urel_integer (plus_p_integer pp x x') (plus_p p y y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" let ?p = "int_of_integer pp" from assms have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel_integer_def by auto have le: "(pp \ x + x') = (?p \ ?x + ?x')" unfolding le_int_of_integer using rel by auto show ?thesis proof (cases "?p \ ?x + ?x'") case True hence True: "(?p \ ?x + ?x') = True" by simp show ?thesis unfolding id using rel unfolding plus_p_integer_def plus_p_def Let_def urel_integer_def unfolding ppp le True if_True using True by auto next case False hence False: "(?p \ ?x + ?x') = False" by simp show ?thesis unfolding id using rel unfolding plus_p_integer_def plus_p_def Let_def urel_integer_def unfolding ppp le False if_False using False by auto qed qed lemma urel_integer_minus: assumes "urel_integer x y" "urel_integer x' y'" shows "urel_integer (minus_p_integer pp x x') (minus_p p y y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" from assms have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel_integer_def by auto have le: "(x' \ x) = (?x' \ ?x)" unfolding le_int_of_integer using rel by auto show ?thesis proof (cases "?x' \ ?x") case True hence True: "(?x' \ ?x) = True" by simp show ?thesis unfolding id using rel unfolding minus_p_integer_def minus_p_def Let_def urel_integer_def unfolding ppp le True if_True using True by auto next case False hence False: "(?x' \ ?x) = False" by simp show ?thesis unfolding id using rel unfolding minus_p_integer_def minus_p_def Let_def urel_integer_def unfolding ppp le False if_False using False by auto qed qed lemma urel_integer_uminus: assumes "urel_integer x y" shows "urel_integer (uminus_p_integer pp x) (uminus_p p y)" proof - let ?x = "int_of_integer x" from assms have id: "y = ?x" and rel: "0 \ ?x" "?x < p" unfolding urel_integer_def by auto have le: "(x = 0) = (?x = 0)" unfolding int_of_integer_0_iff using rel by auto show ?thesis proof (cases "?x = 0") case True hence True: "(?x = 0) = True" by simp show ?thesis unfolding id using rel unfolding uminus_p_integer_def uminus_p_def Let_def urel_integer_def unfolding ppp le True if_True using True by auto next case False hence False: "(?x = 0) = False" by simp show ?thesis unfolding id using rel unfolding uminus_p_integer_def uminus_p_def Let_def urel_integer_def unfolding ppp le False if_False using False by auto qed qed lemma pp_pos: "int_of_integer pp > 0" using ppp nontriv[where 'a = 'a] unfolding p by (simp add: less_integer.rep_eq) lemma urel_integer_mult: assumes "urel_integer x y" "urel_integer x' y'" shows "urel_integer (mult_p_integer pp x x') (mult_p p y y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" from assms have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' < p" unfolding urel_integer_def by auto from rel(1,3) have xx: "0 \ ?x * ?x'" by simp show ?thesis unfolding id using rel unfolding mult_p_integer_def mult_p_def Let_def urel_integer_def unfolding ppp mod_nonneg_pos_int[OF xx pp_pos] using xx pp_pos by simp qed lemma urel_integer_eq: assumes "urel_integer x y" "urel_integer x' y'" shows "(x = x') = (y = y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" from assms have id: "y = ?x" "y' = ?x'" unfolding urel_integer_def by auto show ?thesis unfolding id integer_eq_iff .. qed lemma urel_integer_normalize: assumes x: "urel_integer x y" shows "urel_integer (if x = 0 then 0 else 1) (if y = 0 then 0 else 1)" unfolding urel_integer_eq[OF x urel_integer_0] using urel_integer_0 urel_integer_1 by auto lemma urel_integer_mod: assumes x: "urel_integer x x'" and y: "urel_integer y y'" shows "urel_integer (if y = 0 then x else 0) (if y' = 0 then x' else 0)" unfolding urel_integer_eq[OF y urel_integer_0] using urel_integer_0 x by auto lemma urel_integer_power: "urel_integer x x' \ urel_integer y (int y') \ urel_integer (power_p_integer pp x y) (power_p p x' y')" proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) case (1 x' y' x y) note x = 1(2) note y = 1(3) show ?case proof (cases "y' \ 0") case True hence y: "y = 0" "y' = 0" using urel_integer_eq[OF y urel_integer_0] by auto show ?thesis unfolding y True by (simp add: power_p.simps urel_integer_1) next case False hence id: "(y \ 0) = False" "(y' = 0) = False" using False y by (auto simp add: urel_integer_def not_le) (metis of_int_integer_of of_int_of_nat_eq of_nat_0_less_iff) obtain d' r' where dr': "Divides.divmod_nat y' 2 = (d',r')" by force from divmod_nat_def[of y' 2, unfolded dr'] have r': "r' = y' mod 2" and d': "d' = y' div 2" by auto have aux: "\ y'. int (y' mod 2) = int y' mod 2" by presburger have "urel_integer (y AND 1) r'" unfolding r' using y unfolding urel_integer_def unfolding ppp apply (auto simp add: and_one_eq) apply (simp add: of_nat_mod) done from urel_integer_eq[OF this urel_integer_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel_integer (drop_bit 1 y) (int d')" unfolding d' using y unfolding urel_integer_def unfolding ppp shiftr_integer_conv_div_pow2 by auto from id have "y' \ 0" by auto note IH = 1(1)[OF this refl dr'[symmetric] urel_integer_mult[OF x x] div] show ?thesis unfolding power_p.simps[of _ _ "y'"] power_p_integer.simps[of _ _ y] dr' id if_False rem using IH urel_integer_mult[OF IH x] by (auto simp: Let_def) qed qed lemma urel_integer_inverse: assumes x: "urel_integer x x'" shows "urel_integer (inverse_p_integer pp x) (inverse_p p x')" proof - have p: "urel_integer (pp - 2) (int (nat (p - 2)))" using p2 unfolding urel_integer_def unfolding ppp by auto show ?thesis unfolding inverse_p_integer_def inverse_p_def urel_integer_eq[OF x urel_integer_0] using urel_integer_0 urel_integer_power[OF x p] by auto qed lemma mod_ring_0__integer: "mod_ring_rel_integer 0 0" using urel_integer_0 mod_ring_0 unfolding mod_ring_rel_integer_def by blast lemma mod_ring_1__integer: "mod_ring_rel_integer 1 1" using urel_integer_1 mod_ring_1 unfolding mod_ring_rel_integer_def by blast lemma mod_ring_uminus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (uminus_p_integer pp) uminus" using urel_integer_uminus mod_ring_uminus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_plus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (plus_p_integer pp) (+)" using urel_integer_plus mod_ring_plus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_minus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (minus_p_integer pp) (-)" using urel_integer_minus mod_ring_minus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_mult_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (mult_p_integer pp) ((*))" using urel_integer_mult mod_ring_mult unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_eq_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> (=)) (=) (=)" using urel_integer_eq mod_ring_eq unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma urel_integer_inj: "urel_integer x y \ urel_integer x z \ y = z" using urel_integer_eq[of x y x z] by auto lemma urel_integer_inj': "urel_integer x z \ urel_integer y z \ x = y" using urel_integer_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel_integer: "bi_unique mod_ring_rel_integer" "left_unique mod_ring_rel_integer" "right_unique mod_ring_rel_integer" using bi_unique_mod_ring_rel urel_integer_inj' unfolding mod_ring_rel_integer_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel_integer_def) lemma right_total_mod_ring_rel_integer: "right_total mod_ring_rel_integer" unfolding mod_ring_rel_integer_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel_integer (integer_of_int z) z" unfolding urel_integer_def unfolding ppp by auto with zy show "\ x z. urel_integer x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel_integer: "Domainp mod_ring_rel_integer = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel_integer x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel_integer_def proof let ?i = "int_of_integer" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel_integer a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel_integer x (?i x)" unfolding urel_integer_def using * unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) qed next assume "\a b. x = a \ (\z. urel_integer a z \ mod_ring_rel z b)" then obtain b z where xz: "urel_integer x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel_integer_def unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) qed qed lemma ring_finite_field_ops_integer: "ring_ops (finite_field_ops_integer pp) mod_ring_rel_integer" by (unfold_locales, auto simp: finite_field_ops_integer_def bi_unique_mod_ring_rel_integer right_total_mod_ring_rel_integer mod_ring_plus_integer mod_ring_minus_integer mod_ring_uminus_integer mod_ring_mult_integer mod_ring_eq_integer mod_ring_0__integer mod_ring_1__integer Domainp_mod_ring_rel_integer) end end context prime_field begin context fixes pp :: "integer" assumes *: "p = int_of_integer pp" begin lemma mod_ring_normalize_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (\x. if x = 0 then 0 else 1) normalize" using urel_integer_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_mod_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (\x y. if y = 0 then x else 0) (mod)" using urel_integer_mod[OF *] mod_ring_mod unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_inverse_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (inverse_p_integer pp) inverse" using urel_integer_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_divide_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (divide_p_integer pp) (/)" using mod_ring_inverse_integer mod_ring_mult_integer[OF *] unfolding divide_p_integer_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops_integer: "field_ops (finite_field_ops_integer pp) mod_ring_rel_integer" by (unfold_locales, insert ring_finite_field_ops_integer[OF *], auto simp: ring_ops_def finite_field_ops_integer_def mod_ring_divide_integer mod_ring_inverse_integer mod_ring_mod_integer mod_ring_normalize_integer) end end context prime_field begin (* four implementations of modular integer arithmetic for finite fields *) thm finite_field_ops64 finite_field_ops32 finite_field_ops_integer finite_field_ops_int end context mod_ring_locale begin (* four implementations of modular integer arithmetic for finite rings *) thm ring_finite_field_ops64 ring_finite_field_ops32 ring_finite_field_ops_integer ring_finite_field_ops_int end end diff --git a/thys/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,408 +1,409 @@ (* Title: Code_Target_Word_Base.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Common base for target language implementations of word types\ theory Code_Target_Word_Base imports "HOL-Library.Word" "Word_Lib.Signed_Division_Word" Bits_Integer begin text \More lemmas\ lemma div_half_nat: fixes x y :: nat assumes "y \ 0" shows "(x div y, x mod y) = (let q = 2 * (x div 2 div y); r = x - q * y in if y \ r then (q + 1, r - y) else (q, r))" proof - let ?q = "2 * (x div 2 div y)" have q: "?q = x div y - x div y mod 2" by(metis div_mult2_eq mult.commute minus_mod_eq_mult_div [symmetric]) let ?r = "x - ?q * y" have r: "?r = x mod y + x div y mod 2 * y" by(simp add: q diff_mult_distrib minus_mod_eq_div_mult [symmetric])(metis diff_diff_cancel mod_less_eq_dividend mod_mult2_eq add.commute mult.commute) show ?thesis proof(cases "y \ x - ?q * y") case True with assms q have "x div y mod 2 \ 0" unfolding r by (metis Nat.add_0_right diff_0_eq_0 diff_Suc_1 le_div_geq mod2_gr_0 mod_div_trivial mult_0 neq0_conv numeral_1_eq_Suc_0 numerals(1)) hence "x div y = ?q + 1" unfolding q by simp moreover hence "x mod y = ?r - y" by simp(metis minus_div_mult_eq_mod [symmetric] diff_commute diff_diff_left mult_Suc) ultimately show ?thesis using True by(simp add: Let_def) next case False hence "x div y mod 2 = 0" unfolding r by(simp add: not_le)(metis Nat.add_0_right assms div_less div_mult_self2 mod_div_trivial mult.commute) hence "x div y = ?q" unfolding q by simp moreover hence "x mod y = ?r" by (metis minus_div_mult_eq_mod [symmetric]) ultimately show ?thesis using False by(simp add: Let_def) qed qed lemma div_half_word: fixes x y :: "'a :: len word" assumes "y \ 0" shows "(x div y, x mod y) = (let q = push_bit 1 (drop_bit 1 x div y); r = x - q * y in if y \ r then (q + 1, r - y) else (q, r))" proof - obtain n where n: "x = of_nat n" "n < 2 ^ LENGTH('a)" by (rule that [of \unat x\]) simp_all moreover obtain m where m: "y = of_nat m" "m < 2 ^ LENGTH('a)" by (rule that [of \unat y\]) simp_all ultimately have [simp]: \unat (of_nat n :: 'a word) = n\ \unat (of_nat m :: 'a word) = m\ by (transfer, simp add: take_bit_of_nat take_bit_nat_eq_self_iff)+ let ?q = "push_bit 1 (drop_bit 1 x div y)" let ?q' = "2 * (n div 2 div m)" have "n div 2 div m < 2 ^ LENGTH('a)" using n by (metis of_nat_inverse unat_lt2p uno_simps(2)) hence q: "?q = of_nat ?q'" using n m by (auto simp add: drop_bit_eq_div word_arith_nat_div uno_simps take_bit_nat_eq_self) from assms have "m \ 0" using m by -(rule notI, simp) from n have "2 * (n div 2 div m) < 2 ^ LENGTH('a)" by(metis mult.commute div_mult2_eq minus_mod_eq_mult_div [symmetric] less_imp_diff_less of_nat_inverse unat_lt2p uno_simps(2)) moreover have "2 * (n div 2 div m) * m < 2 ^ LENGTH('a)" using n unfolding div_mult2_eq[symmetric] by(subst (2) mult.commute)(simp add: minus_mod_eq_div_mult [symmetric] diff_mult_distrib minus_mod_eq_mult_div [symmetric] div_mult2_eq) moreover have "2 * (n div 2 div m) * m \ n" by (simp flip: div_mult2_eq ac_simps) ultimately have r: "x - ?q * y = of_nat (n - ?q' * m)" and "y \ x - ?q * y \ of_nat (n - ?q' * m) - y = of_nat (n - ?q' * m - m)" using n m unfolding q apply (simp_all add: of_nat_diff) apply (subst of_nat_diff) apply (simp_all add: word_le_nat_alt take_bit_nat_eq_self unat_sub_if' unat_word_ariths) done then show ?thesis using n m div_half_nat [OF \m \ 0\, of n] unfolding q by (simp add: word_le_nat_alt word_div_def word_mod_def Let_def take_bit_nat_eq_self flip: zdiv_int zmod_int split del: if_split split: if_split_asm) qed lemma word_test_bit_set_bits: "bit (BITS n. f n :: 'a :: len word) n \ n < LENGTH('a) \ f n" by (fact bit_set_bits_word_iff) lemma word_of_int_conv_set_bits: "word_of_int i = (BITS n. bit i n)" by (rule word_eqI) (auto simp add: word_test_bit_set_bits bit_simps) lemma word_and_mask_or_conv_and_mask: "bit n index \ (n AND mask index) OR (push_bit index 1) = n AND mask (index + 1)" for n :: \'a::len word\ by(rule word_eqI)(auto simp add: bit_simps) lemma uint_and_mask_or_full: fixes n :: "'a :: len word" assumes "bit n (LENGTH('a) - 1)" and "mask1 = mask (LENGTH('a) - 1)" and "mask2 = push_bit (LENGTH('a) - 1) 1" shows "uint (n AND mask1) OR mask2 = uint n" proof - have "mask2 = uint (push_bit (LENGTH('a) - 1) 1 :: 'a word)" using assms by transfer (simp add: take_bit_push_bit) hence "uint (n AND mask1) OR mask2 = uint (n AND mask1 OR (push_bit (LENGTH('a) - 1) 1 :: 'a word))" by(simp add: uint_or) also have "\ = uint (n AND mask (LENGTH('a) - 1 + 1))" using assms by(simp only: word_and_mask_or_conv_and_mask) also have "\ = uint n" by simp finally show ?thesis . qed text \Division on @{typ "'a word"} is unsigned, but Scala and OCaml only have signed division and modulus.\ lemmas word_sdiv_def = sdiv_word_def lemmas word_smod_def = smod_word_def lemma [code]: "x sdiv y = (let x' = sint x; y' = sint y; negative = (x' < 0) \ (y' < 0); result = abs x' div abs y' in word_of_int (if negative then -result else result))" for x y :: \'a::len word\ by (simp add: sdiv_word_def signed_divide_int_def sgn_if Let_def not_less not_le) lemma [code]: "x smod y = (let x' = sint x; y' = sint y; negative = (x' < 0); result = abs x' mod abs y' in word_of_int (if negative then -result else result))" for x y :: \'a::len word\ proof - have *: \k mod l = k - k div l * l\ for k l :: int by (simp add: minus_div_mult_eq_mod) show ?thesis by (simp add: smod_word_def signed_modulo_int_def signed_divide_int_def * sgn_if Let_def) qed text \ This algorithm implements unsigned division in terms of signed division. Taken from Hacker's Delight. \ lemma divmod_via_sdivmod: fixes x y :: "'a :: len word" assumes "y \ 0" shows "(x div y, x mod y) = (if push_bit (LENGTH('a) - 1) 1 \ y then if x < y then (0, x) else (1, x - y) else let q = (push_bit 1 (drop_bit 1 x sdiv y)); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" proof(cases "push_bit (LENGTH('a) - 1) 1 \ y") case True note y = this show ?thesis proof(cases "x < y") case True then have "x mod y = x" by transfer simp thus ?thesis using True y by(simp add: word_div_lt_eq_0) next case False obtain n where n: "y = of_nat n" "n < 2 ^ LENGTH('a)" by (rule that [of \unat y\]) simp_all have "unat x < 2 ^ LENGTH('a)" by(rule unat_lt2p) also have "\ = 2 * 2 ^ (LENGTH('a) - 1)" by(metis Suc_pred len_gt_0 power_Suc One_nat_def) also have "\ \ 2 * n" using y n by transfer (simp add: push_bit_of_1 take_bit_eq_mod) finally have div: "x div of_nat n = 1" using False n by (simp add: word_div_eq_1_iff take_bit_nat_eq_self) moreover have "x mod y = x - x div y * y" by (simp add: minus_div_mult_eq_mod) with div n have "x mod y = x - y" by simp ultimately show ?thesis using False y n by simp qed next case False note y = this obtain n where n: "x = of_nat n" "n < 2 ^ LENGTH('a)" by (rule that [of \unat x\]) simp_all hence "int n div 2 + 2 ^ (LENGTH('a) - Suc 0) < 2 ^ LENGTH('a)" by (cases \LENGTH('a)\) (auto dest: less_imp_of_nat_less [where ?'a = int]) with y n have "sint (drop_bit 1 x) = uint (drop_bit 1 x)" by (cases \LENGTH('a)\) (auto simp add: sint_uint drop_bit_eq_div take_bit_nat_eq_self uint_div_distrib signed_take_bit_int_eq_self_iff) moreover have "uint y + 2 ^ (LENGTH('a) - Suc 0) < 2 ^ LENGTH('a)" using y by (cases \LENGTH('a)\) (simp_all add: not_le push_bit_of_1 word_less_alt uint_power_lower) then have "sint y = uint y" apply (cases \LENGTH('a)\) apply (auto simp add: sint_uint signed_take_bit_int_eq_self_iff) using uint_ge_0 [of y] by linarith ultimately show ?thesis using y apply (subst div_half_word [OF assms]) apply (simp add: sdiv_word_def signed_divide_int_def flip: uint_div) done qed text \More implementations tailored towards target-language implementations\ context includes integer.lifting 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: "uint (word_of_int x :: 'a word) = x AND mask (LENGTH('a :: len))" by (simp add: take_bit_eq_mask) context fixes f :: "nat \ bool" begin -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)\ +definition set_bits_aux :: \nat \ 'a word \ 'a::len word\ + where \set_bits_aux n w = push_bit n w OR take_bit n (set_bits f)\ lemma bit_set_bit_aux [bit_simps]: - \bit (set_bits_aux w n) m \ m < LENGTH('a) \ + \bit (set_bits_aux n w) m \ m < LENGTH('a) \ (if m < n then f m else bit w (m - n))\ for w :: \'a::len word\ by (auto simp add: bit_simps set_bits_aux_def) -lemma set_bits_aux_conv: - \set_bits_aux w n = (push_bit n w) OR (set_bits f AND mask n)\ - for w :: \'a::len word\ +corollary set_bits_conv_set_bits_aux: + \set_bits f = (set_bits_aux LENGTH('a) 0 :: 'a :: len word)\ by (rule bit_word_eqI) (simp add: bit_simps) -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) + \set_bits_aux 0 w = w\ + by (simp add: set_bits_aux_def) lemma set_bits_aux_Suc [simp]: - \set_bits_aux w (Suc n) = set_bits_aux (push_bit 1 w OR (if f n then 1 else 0)) n\ - by (rule bit_word_eqI) (auto simp add: bit_simps not_less le_less_Suc_eq mult.commute [of _ 2]) + \set_bits_aux (Suc n) w = set_bits_aux n (push_bit 1 w OR (if f n then 1 else 0))\ + by (rule bit_word_eqI) (auto simp add: bit_simps le_less_Suc_eq mult.commute [of _ 2]) lemma set_bits_aux_simps [code]: - \set_bits_aux w 0 = w\ - \set_bits_aux w (Suc n) = set_bits_aux (push_bit 1 w OR (if f n then 1 else 0)) n\ + \set_bits_aux 0 w = w\ + \set_bits_aux (Suc n) w = set_bits_aux n (push_bit 1 w OR (if f n then 1 else 0))\ by simp_all +lemma set_bits_aux_rec: + \set_bits_aux n w = + (if n = 0 then w + else let n' = n - 1 in set_bits_aux n' (push_bit 1 w OR (if f n' then 1 else 0)))\ + by (cases n) simp_all + end lemma word_of_int_via_signed: fixes mask assumes mask_def: "mask = Bit_Operations.mask (LENGTH('a))" and shift_def: "shift = push_bit LENGTH('a) 1" and index_def: "index = LENGTH('a) - 1" and overflow_def:"overflow = push_bit (LENGTH('a) - 1) 1" and least_def: "least = - overflow" shows "(word_of_int i :: 'a :: len word) = (let i' = i AND mask in if bit i' index then if i' - shift < least \ overflow \ i' - shift then arbitrary1 i' else word_of_int (i' - shift) else if i' < least \ overflow \ i' then arbitrary2 i' else word_of_int i')" proof - define i' where "i' = i AND mask" have "shift = mask + 1" unfolding assms by (simp add: mask_eq_exp_minus_1 push_bit_of_1) hence "i' < shift" by (simp add: mask_def i'_def) show ?thesis proof(cases "bit i' index") case True then have unf: "i' = overflow OR i'" apply (simp add: assms i'_def push_bit_of_1 flip: take_bit_eq_mask) apply (rule bit_eqI) apply (auto simp add: bit_take_bit_iff bit_or_iff bit_exp_iff) done have \overflow \ overflow OR i'\ by (simp add: i'_def mask_def or_greater_eq) then have "overflow \ i'" by (subst unf) hence "i' - shift < least \ False" unfolding assms by(cases "LENGTH('a)")(simp_all add: not_less push_bit_of_1) moreover have "overflow \ i' - shift \ False" using \i' < shift\ unfolding assms by(cases "LENGTH('a)")(auto simp add: not_le push_bit_of_1 elim: less_le_trans) moreover have "word_of_int (i' - shift) = (word_of_int i :: 'a word)" using \i' < shift\ by (simp add: i'_def shift_def mask_def push_bit_of_1 word_of_int_eq_iff flip: take_bit_eq_mask) ultimately show ?thesis using True by(simp add: Let_def i'_def) next case False have "i' = i AND Bit_Operations.mask (LENGTH('a) - 1)" apply (rule bit_eqI) apply (use False in \auto simp add: bit_simps assms i'_def\) apply (auto simp add: less_le) done also have "\ \ Bit_Operations.mask (LENGTH('a) - 1)" using AND_upper2 mask_nonnegative_int by blast also have "\ < overflow" by (simp add: mask_int_def push_bit_of_1 overflow_def) also have "least \ 0" unfolding least_def overflow_def by simp have "0 \ i'" by (simp add: i'_def mask_def) hence "least \ i'" using \least \ 0\ by simp moreover have "word_of_int i' = (word_of_int i :: 'a word)" by (simp add: i'_def mask_def of_int_and_eq of_int_mask_eq) ultimately show ?thesis using False by(simp add: Let_def i'_def) qed qed text \Quickcheck conversion functions\ context includes state_combinator_syntax begin definition qc_random_cnv :: "(natural \ 'a::term_of) \ natural \ Random.seed \ ('a \ (unit \ Code_Evaluation.term)) \ Random.seed" where "qc_random_cnv a_of_natural i = Random.range (i + 1) \\ (\k. Pair ( let n = a_of_natural k in (n, \_. Code_Evaluation.term_of n)))" end definition qc_exhaustive_cnv :: "(natural \ 'a) \ ('a \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "qc_exhaustive_cnv a_of_natural f d = Quickcheck_Exhaustive.exhaustive (%x. f (a_of_natural x)) d" definition qc_full_exhaustive_cnv :: "(natural \ ('a::term_of)) \ ('a \ (unit \ term) \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "qc_full_exhaustive_cnv a_of_natural f d = Quickcheck_Exhaustive.full_exhaustive (%(x, xt). f (a_of_natural x, %_. Code_Evaluation.term_of (a_of_natural x))) d" declare [[quickcheck_narrowing_ghc_options = "-XTypeSynonymInstances"]] definition qc_narrowing_drawn_from :: "'a list \ integer \ _" where "qc_narrowing_drawn_from xs = foldr Quickcheck_Narrowing.sum (map Quickcheck_Narrowing.cons (butlast xs)) (Quickcheck_Narrowing.cons (last xs))" locale quickcheck_narrowing_samples = fixes a_of_integer :: "integer \ 'a \ 'a :: {partial_term_of, term_of}" and zero :: "'a" and tr :: "typerep" begin function narrowing_samples :: "integer \ 'a list" where "narrowing_samples i = (if i > 0 then let (a, a') = a_of_integer i in narrowing_samples (i - 1) @ [a, a'] else [zero])" by pat_completeness auto termination including integer.lifting proof(relation "measure nat_of_integer") fix i :: integer assume "0 < i" thus "(i - 1, i) \ measure nat_of_integer" by simp(transfer, simp) qed simp definition partial_term_of_sample :: "integer \ 'a" where "partial_term_of_sample i = (if i < 0 then undefined else if i = 0 then zero else if i mod 2 = 0 then snd (a_of_integer (i div 2)) else fst (a_of_integer (i div 2 + 1)))" lemma partial_term_of_code: "partial_term_of (ty :: 'a itself) (Quickcheck_Narrowing.Narrowing_variable p t) \ Code_Evaluation.Free (STR ''_'') tr" "partial_term_of (ty :: 'a itself) (Quickcheck_Narrowing.Narrowing_constructor i []) \ Code_Evaluation.term_of (partial_term_of_sample i)" by (rule partial_term_of_anything)+ end lemmas [code] = quickcheck_narrowing_samples.narrowing_samples.simps quickcheck_narrowing_samples.partial_term_of_sample_def text \ The separate code target \SML_word\ collects setups for the code generator that PolyML does not provide. \ setup \Code_Target.add_derived_target ("SML_word", [(Code_ML.target_SML, I)])\ code_identifier code_module Code_Target_Word_Base \ (SML) Word and (Haskell) Word and (OCaml) Word and (Scala) Word end diff --git a/thys/Native_Word/Native_Word_Test.thy b/thys/Native_Word/Native_Word_Test.thy --- a/thys/Native_Word/Native_Word_Test.thy +++ b/thys/Native_Word/Native_Word_Test.thy @@ -1,485 +1,487 @@ (* Title: Native_Word_Test.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Test cases\ theory Native_Word_Test imports Uint64 Uint32 Uint16 Uint8 Uint Native_Cast_Uint "HOL-Library.Code_Test" "Word_Lib.Bit_Shifts_Infix_Syntax" begin section \Tests for @{typ uint32}\ -notation sshiftr_uint32 (infixl ">>>" 55) +abbreviation (input) sshiftr_uint32 (infixl ">>>" 55) + where \w >>> n \ signed_drop_bit_uint32 n w\ definition test_uint32 where "test_uint32 \ (([ 0x100000001, -1, -4294967291, 0xFFFFFFFF, 0x12345678 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + (- 6), 0xFFFFFFFFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x12345678 * 0x87654321 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 32 True, set_bit 5 32 False, set_bit (- 5) 32 True, set_bit (- 5) 32 False , 1 << 2, -1 << 3, 1 << 32, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 32, -100 >> 32 , 100 >>> 3, -100 >>> 3, 100 >>> 32, -100 >>> 32] :: uint32 list) = [ 1, 4294967295, 5, 4294967295, 305419896 , 18 , 126 , 108 , 4294967205 , 11, 1, 4294967295, 4294967285, 0 , 2, 4294967294 , 15, 4294967281, 20, 1891143032 , 1, 1431655763, 0, 0 , 2, 2, 4294967291, 5 , 21, 4294967295, 4, 4294967289 , 5, 5, 4294967291, 4294967291 , 4, 4294967288, 0, 1 , 12, 536870899, 0, 0 , 12, 4294967283, 0, 4294967295]) \ ([ (0x5 :: uint32) = 0x5, (0x5 :: uint32) = 0x6 , (0x5 :: uint32) < 0x5, (0x5 :: uint32) < 0x6, (-5 :: uint32) < 6, (6 :: uint32) < -5 , (0x5 :: uint32) \ 0x5, (0x5 :: uint32) \ 0x4, (-5 :: uint32) \ 6, (6 :: uint32) \ -5 , (0x7FFFFFFF :: uint32) < 0x80000000, (0xFFFFFFFF :: uint32) < 0, (0x80000000 :: uint32) < 0x7FFFFFFF , bit (0x7FFFFFFF :: uint32) 0, bit (0x7FFFFFFF :: uint32) 31, bit (0x80000000 :: uint32) 31, bit (0x80000000 :: uint32) 32 ] = [ True, False , False, True, False, True , True, False, False, True , True, False, False , True, False, True, False ]) \ ([integer_of_uint32 0, integer_of_uint32 0x7FFFFFFF, integer_of_uint32 0x80000000, integer_of_uint32 0xAAAAAAAA] = [0, 0x7FFFFFFF, 0x80000000, 0xAAAAAAAA])" no_notation sshiftr_uint32 (infixl ">>>" 55) export_code test_uint32 checking SML Haskell? OCaml? Scala notepad begin have test_uint32 by eval have test_uint32 by code_simp have test_uint32 by normalization end definition test_uint32' :: uint32 where "test_uint32' = 0 + 10 - 14 * 3 div 6 mod 3 << 3 >> 2" ML \val 0wx12 = @{code test_uint32'}\ lemma "x AND y = x OR (y :: uint32)" quickcheck[random, expect=counterexample] quickcheck[exhaustive, expect=counterexample] oops lemma "(x :: uint32) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] by transfer simp lemma "(f :: uint32 \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) section \Tests for @{typ uint16}\ -notation sshiftr_uint16 (infixl ">>>" 55) +abbreviation (input) sshiftr_uint16 (infixl ">>>" 55) + where \w >>> n \ signed_drop_bit_uint16 n w\ definition test_uint16 where "test_uint16 \ (([ 0x10001, -1, -65535, 0xFFFF, 0x1234 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + -6, 0xFFFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x1234 * 0x8765 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 32 True, set_bit 5 32 False, set_bit (- 5) 32 True, set_bit (- 5) 32 False , 1 << 2, -1 << 3, 1 << 16, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 16, -100 >> 16 , 100 >>> 3, -100 >>> 3, 100 >>> 16, -100 >>> 16] :: uint16 list) = [ 1, 65535, 1, 65535, 4660 , 18 , 126 , 108 , 65445 , 11, 1, 65535, 65525, 0 , 2, 65534 , 15, 65521, 20, 39556 , 1, 21843, 0, 0 , 2, 2, 65531, 5 , 21, 65535, 4, 65529 , 5, 5, 65531, 65531 , 4, 65528, 0, 1 , 12, 8179, 0, 0 , 12, 65523, 0, 65535]) \ ([ (0x5 :: uint16) = 0x5, (0x5 :: uint16) = 0x6 , (0x5 :: uint16) < 0x5, (0x5 :: uint16) < 0x6, (-5 :: uint16) < 6, (6 :: uint16) < -5 , (0x5 :: uint16) \ 0x5, (0x5 :: uint16) \ 0x4, (-5 :: uint16) \ 6, (6 :: uint16) \ -5 , (0x7FFF :: uint16) < 0x8000, (0xFFFF :: uint16) < 0, (0x8000 :: uint16) < 0x7FFF , bit (0x7FFF :: uint16) 0, bit (0x7FFF :: uint16) 15, bit (0x8000 :: uint16) 15, bit (0x8000 :: uint16) 16 ] = [ True, False , False, True, False, True , True, False, False, True , True, False, False , True, False, True, False ]) \ ([integer_of_uint16 0, integer_of_uint16 0x7FFF, integer_of_uint16 0x8000, integer_of_uint16 0xAAAA] = [0, 0x7FFF, 0x8000, 0xAAAA])" -no_notation sshiftr_uint16 (infixl ">>>" 55) - export_code test_uint16 checking Haskell? Scala -export_code test_uint16 in SML_word +export_code test_uint16 checking SML_word notepad begin have test_uint16 by code_simp have test_uint16 by normalization end lemma "(x :: uint16) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] by transfer simp lemma "(f :: uint16 \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) section \Tests for @{typ uint8}\ -notation sshiftr_uint8 (infixl ">>>" 55) +abbreviation (input) sshiftr_uint8 (infixl ">>>" 55) + where \w >>> n \ signed_drop_bit_uint8 n w\ definition test_uint8 where "test_uint8 \ (([ 0x101, -1, -255, 0xFF, 0x12 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + -6, 0xFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x12 * 0x87 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 32 True, set_bit 5 32 False, set_bit (- 5) 32 True, set_bit (- 5) 32 False , 1 << 2, -1 << 3, 1 << 8, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 8, -100 >> 8 , 100 >>> 3, -100 >>> 3, 100 >>> 8, -100 >>> 8] :: uint8 list) = [ 1, 255, 1, 255, 18 , 18 , 126 , 108 , 165 , 11, 1, 255, 245, 0 , 2, 254 , 15, 241, 20, 126 , 1, 83, 0, 0 , 2, 2, 251, 5 , 21, 255, 4, 249 , 5, 5, 251, 251 , 4, 248, 0, 1 , 12, 19, 0, 0 , 12, 243, 0, 255]) \ ([ (0x5 :: uint8) = 0x5, (0x5 :: uint8) = 0x6 , (0x5 :: uint8) < 0x5, (0x5 :: uint8) < 0x6, (-5 :: uint8) < 6, (6 :: uint8) < -5 , (0x5 :: uint8) \ 0x5, (0x5 :: uint8) \ 0x4, (-5 :: uint8) \ 6, (6 :: uint8) \ -5 , (0x7F :: uint8) < 0x80, (0xFF :: uint8) < 0, (0x80 :: uint8) < 0x7F , bit (0x7F :: uint8) 0, bit (0x7F :: uint8) 7, bit (0x80 :: uint8) 7, bit (0x80 :: uint8) 8 ] = [ True, False , False, True, False, True , True, False, False, True , True, False, False , True, False, True, False ]) \ ([integer_of_uint8 0, integer_of_uint8 0x7F, integer_of_uint8 0x80, integer_of_uint8 0xAA] = [0, 0x7F, 0x80, 0xAA])" no_notation sshiftr_uint8 (infixl ">>>" 55) export_code test_uint8 checking SML Haskell? Scala -export_code test_uint8 in SML +notepad begin -notepad begin have test_uint8 by eval have test_uint8 by code_simp have test_uint8 by normalization end ML_val \val true = @{code test_uint8}\ definition test_uint8' :: uint8 where "test_uint8' = 0 + 10 - 14 * 3 div 6 mod 3 << 3 >> 2" ML \val 0wx12 = @{code test_uint8'}\ lemma "x AND y = x OR (y :: uint8)" quickcheck[random, expect=counterexample] quickcheck[exhaustive, expect=counterexample] oops lemma "(x :: uint8) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] by transfer simp lemma "(f :: uint8 \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) section \Tests for @{typ "uint"}\ -notation sshiftr_uint (infixl ">>>" 55) +abbreviation (input) sshiftr_uint (infixl ">>>" 55) + where \w >>> n \ signed_drop_bit_uint n w\ definition "test_uint \ let test_list1 = (let HS = uint_of_int (2 ^ (dflt_size - 1)) in ([ HS + HS + 1, -1, -HS - HS + 5, HS + (HS - 1), 0x12 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + -6, HS + (HS - 1) + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x12345678 * 0x87654321] @ (if dflt_size > 4 then [ 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 dflt_size True, set_bit 5 dflt_size False, set_bit (- 5) dflt_size True, set_bit (- 5) dflt_size False , 1 << 2, -1 << 3, push_bit dflt_size 1, 1 << 0 , 31 >> 3, -1 >> 3, 31 >> dflt_size, -1 >> dflt_size , 15 >>> 2, -1 >>> 3, 15 >>> dflt_size, -1 >>> dflt_size] else []) :: uint list)); test_list2 = (let S = wivs_shift in ([ 1, -1, -S + 5, S - 1, 0x12 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + -6, 0 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x12345678 * 0x87654321] @ (if dflt_size > 4 then [ 5 div 3, (S - 5) div 3, (S - 5) div (S - 3), 5 div (S - 3) , 5 mod 3, (S - 5) mod 3, (S - 5) mod (S - 3), 5 mod (S - 3) , set_bit 5 4 True, -1, set_bit 5 0 False, -7 , 5, 5, -5, -5 , 4, -8, 0, 1 , 3, (S >> 3) - 1, 0, 0 , 3, (S >> 1) + (S >> 1) - 1, 0, -1] else []) :: int list)); test_list_c1 = (let HS = uint_of_int ((2^(dflt_size - 1))) in [ (0x5 :: uint) = 0x5, (0x5 :: uint) = 0x6 , (0x5 :: uint) < 0x5, (0x5 :: uint) < 0x6, (-5 :: uint) < 6, (6 :: uint) < -5 , (0x5 :: uint) \ 0x5, (0x5 :: uint) \ 0x4, (-5 :: uint) \ 6, (6 :: uint) \ -5 , (HS - 1) < HS, (HS + HS - 1) < 0, HS < HS - 1 , bit (HS - 1) 0, bit (HS - 1 :: uint) (dflt_size - 1), bit (HS :: uint) (dflt_size - 1), bit (HS :: uint) dflt_size ]); test_list_c2 = [ True, False , False, dflt_size\2, dflt_size=3, dflt_size\3 , True, False, dflt_size=3, dflt_size\3 , True, False, False , dflt_size\1, False, True, False ] in test_list1 = map uint_of_int test_list2 \ test_list_c1 = test_list_c2" no_notation sshiftr_uint (infixl ">>>" 55) export_code test_uint checking SML Haskell? OCaml? Scala lemma "test_uint" quickcheck[exhaustive, expect=no_counterexample] oops \ \FIXME: prove correctness of test by reflective means (not yet supported)\ lemma "x AND y = x OR (y :: uint)" quickcheck[random, expect=counterexample] quickcheck[exhaustive, expect=counterexample] oops lemma "(x :: uint) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] by transfer simp lemma "(f :: uint \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) section \ Tests for @{typ uint64} \ -notation sshiftr_uint64 (infixl ">>>" 55) +abbreviation (input) sshiftr_uint64 (infixl ">>>" 55) + where \w >>> n \ signed_drop_bit_uint64 n w\ definition test_uint64 where "test_uint64 \ (([ 0x10000000000000001, -1, -9223372036854775808, 0xFFFFFFFFFFFFFFFF, 0x1234567890ABCDEF , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + (- 6), 0xFFFFFFFFFFFFFFFFFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x1234567890ABCDEF * 0xFEDCBA0987654321 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 64 True, set_bit 5 64 False, set_bit (- 5) 64 True, set_bit (- 5) 64 False , 1 << 2, -1 << 3, 1 << 64, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 64, -100 >> 64 , 100 >>> 3, -100 >>> 3, 100 >>> 64, -100 >>> 64] :: uint64 list) = [ 1, 18446744073709551615, 9223372036854775808, 18446744073709551615, 1311768467294899695 , 18 , 126 , 108 , 18446744073709551525 , 11, 1, 18446744073709551615, 18446744073709551605, 0 , 2, 18446744073709551614 , 15, 18446744073709551601, 20, 14000077364136384719 , 1, 6148914691236517203, 0, 0 , 2, 2, 18446744073709551611, 5 , 21, 18446744073709551615, 4, 18446744073709551609 , 5, 5, 18446744073709551611, 18446744073709551611 , 4, 18446744073709551608, 0, 1 , 12, 2305843009213693939, 0, 0 , 12, 18446744073709551603, 0, 18446744073709551615]) \ ([ (0x5 :: uint64) = 0x5, (0x5 :: uint64) = 0x6 , (0x5 :: uint64) < 0x5, (0x5 :: uint64) < 0x6, (-5 :: uint64) < 6, (6 :: uint64) < -5 , (0x5 :: uint64) \ 0x5, (0x5 :: uint64) \ 0x4, (-5 :: uint64) \ 6, (6 :: uint64) \ -5 , (0x7FFFFFFFFFFFFFFF :: uint64) < 0x8000000000000000, (0xFFFFFFFFFFFFFFFF :: uint64) < 0, (0x8000000000000000 :: uint64) < 0x7FFFFFFFFFFFFFFF , bit (0x7FFFFFFFFFFFFFFF :: uint64) 0, bit (0x7FFFFFFFFFFFFFFF :: uint64) 63, bit (0x8000000000000000 :: uint64) 63, bit (0x8000000000000000 :: uint64) 64 ] = [ True, False , False, True, False, True , True, False, False, True , True, False, False , True, False, True, False ]) \ ([integer_of_uint64 0, integer_of_uint64 0x7FFFFFFFFFFFFFFF, integer_of_uint64 0x8000000000000000, integer_of_uint64 0xAAAAAAAAAAAAAAAA] = [0, 0x7FFFFFFFFFFFFFFF, 0x8000000000000000, 0xAAAAAAAAAAAAAAAA])" value [nbe] "[0x10000000000000001, -1, -9223372036854775808, 0xFFFFFFFFFFFFFFFF, 0x1234567890ABCDEF , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + (- 6), 0xFFFFFFFFFFFFFFFFFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x1234567890ABCDEF * 0xFEDCBA0987654321 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 64 True, set_bit 5 64 False, set_bit (- 5) 64 True, set_bit (- 5) 64 False , 1 << 2, -1 << 3, 1 << 64, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 64, -100 >> 64 , 100 >>> 3, -100 >>> 3, 100 >>> 64, -100 >>> 64] :: uint64 list" no_notation sshiftr_uint64 (infixl ">>>" 55) export_code test_uint64 checking SML Haskell? OCaml? Scala notepad begin have test_uint64 by eval have test_uint64 by code_simp have test_uint64 by normalization end ML_val \val true = @{code test_uint64}\ definition test_uint64' :: uint64 where "test_uint64' = 0 + 10 - 14 * 3 div 6 mod 3 << 3 >> 2" section \Tests for casts\ definition test_casts :: bool where "test_casts \ map uint8_of_uint32 [10, 0, 0xFE, 0xFFFFFFFF] = [10, 0, 0xFE, 0xFF] \ map uint8_of_uint64 [10, 0, 0xFE, 0xFFFFFFFFFFFFFFFF] = [10, 0, 0xFE, 0xFF] \ map uint32_of_uint8 [10, 0, 0xFF] = [10, 0, 0xFF] \ map uint64_of_uint8 [10, 0, 0xFF] = [10, 0, 0xFF]" definition test_casts' :: bool where "test_casts' \ map uint8_of_uint16 [10, 0, 0xFE, 0xFFFF] = [10, 0, 0xFE, 0xFF] \ map uint16_of_uint8 [10, 0, 0xFF] = [10, 0, 0xFF] \ map uint16_of_uint32 [10, 0, 0xFFFE, 0xFFFFFFFF] = [10, 0, 0xFFFE, 0xFFFF] \ map uint16_of_uint64 [10, 0, 0xFFFE, 0xFFFFFFFFFFFFFFFF] = [10, 0, 0xFFFE, 0xFFFF] \ map uint32_of_uint16 [10, 0, 0xFFFF] = [10, 0, 0xFFFF] \ map uint64_of_uint16 [10, 0, 0xFFFF] = [10, 0, 0xFFFF]" definition test_casts'' :: bool where "test_casts'' \ map uint32_of_uint64 [10, 0, 0xFFFFFFFE, 0xFFFFFFFFFFFFFFFF] = [10, 0, 0xFFFFFFFE, 0xFFFFFFFF] \ map uint64_of_uint32 [10, 0, 0xFFFFFFFF] = [10, 0, 0xFFFFFFFF]" export_code test_casts test_casts'' checking SML Haskell? Scala export_code test_casts'' checking OCaml? export_code test_casts' checking Haskell? Scala notepad begin have test_casts by eval have test_casts by normalization have test_casts by code_simp have test_casts' by normalization have test_casts' by code_simp have test_casts'' by eval have test_casts'' by normalization have test_casts'' by code_simp end ML \ val true = @{code test_casts} val true = @{code test_casts''} \ definition test_casts_uint :: bool where "test_casts_uint \ map uint_of_uint32 ([0, 10] @ (if dflt_size < 32 then [push_bit (dflt_size - 1) 1, 0xFFFFFFFF] else [0xFFFFFFFF])) = [0, 10] @ (if dflt_size < 32 then [push_bit (dflt_size - 1) 1, (push_bit dflt_size 1) - 1] else [0xFFFFFFFF]) \ map uint32_of_uint [0, 10, if dflt_size < 32 then push_bit (dflt_size - 1) 1 else 0xFFFFFFFF] = [0, 10, if dflt_size < 32 then push_bit (dflt_size - 1) 1 else 0xFFFFFFFF] \ map uint_of_uint64 [0, 10, push_bit (dflt_size - 1) 1, 0xFFFFFFFFFFFFFFFF] = [0, 10, push_bit (dflt_size - 1) 1, (push_bit dflt_size 1) - 1] \ map uint64_of_uint [0, 10, push_bit (dflt_size - 1) 1] = [0, 10, push_bit (dflt_size - 1) 1]" definition test_casts_uint' :: bool where "test_casts_uint' \ map uint_of_uint16 [0, 10, 0xFFFF] = [0, 10, 0xFFFF] \ map uint16_of_uint [0, 10, 0xFFFF] = [0, 10, 0xFFFF]" definition test_casts_uint'' :: bool where "test_casts_uint'' \ map uint_of_uint8 [0, 10, 0xFF] = [0, 10, 0xFF] \ map uint8_of_uint [0, 10, 0xFF] = [0, 10, 0xFF]" end diff --git a/thys/Native_Word/Uint.thy b/thys/Native_Word/Uint.thy --- a/thys/Native_Word/Uint.thy +++ b/thys/Native_Word/Uint.thy @@ -1,894 +1,821 @@ (* Title: Uint.thy Author: Peter Lammich, TU Munich Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of default size\ theory Uint imports - Code_Target_Word_Base + Code_Target_Word_Base Word_Type_Copies begin text \ This theory provides access to words in the target languages of the code generator whose bit width is the default of the target language. To that end, the type \uint\ models words of width \dflt_size\, but \dflt_size\ is known only to be positive. Usage restrictions: Default-size words (type \uint\) cannot be used for evaluation, because the results depend on the particular choice of word size in the target language and implementation. Symbolic evaluation has not yet been set up for \uint\. \ text \The default size type\ typedecl dflt_size instantiation dflt_size :: typerep begin definition "typerep_class.typerep \ \_ :: dflt_size itself. Typerep.Typerep (STR ''Uint.dflt_size'') []" instance .. end consts dflt_size_aux :: "nat" specification (dflt_size_aux) dflt_size_aux_g0: "dflt_size_aux > 0" by auto hide_fact dflt_size_aux_def instantiation dflt_size :: len begin definition "len_of_dflt_size (_ :: dflt_size itself) \ dflt_size_aux" instance by(intro_classes)(simp add: len_of_dflt_size_def dflt_size_aux_g0) end abbreviation "dflt_size \ len_of (TYPE (dflt_size))" context includes integer.lifting begin lift_definition dflt_size_integer :: integer is "int dflt_size" . declare dflt_size_integer_def[code del] \ \The code generator will substitute a machine-dependent value for this constant\ lemma dflt_size_by_int[code]: "dflt_size = nat_of_integer dflt_size_integer" by transfer simp lemma dflt_size[simp]: "dflt_size > 0" "dflt_size \ Suc 0" "\ dflt_size < Suc 0" using len_gt_0[where 'a=dflt_size] by (simp_all del: len_gt_0) end section \Type definition and primitive operations\ -typedef uint = "UNIV :: dflt_size word set" .. +typedef uint = \UNIV :: dflt_size word set\ .. + +global_interpretation uint: word_type_copy Abs_uint Rep_uint + using type_definition_uint by (rule word_type_copy.intro) setup_lifting type_definition_uint -text \Use an abstract type for code generation to disable pattern matching on @{term Abs_uint}.\ -declare Rep_uint_inverse[code abstype] - -declare Quotient_uint[transfer_rule] - -instantiation uint :: comm_ring_1 -begin -lift_definition zero_uint :: uint is "0 :: dflt_size word" . -lift_definition one_uint :: uint is "1" . -lift_definition plus_uint :: "uint \ uint \ uint" is "(+) :: dflt_size word \ _" . -lift_definition minus_uint :: "uint \ uint \ uint" is "(-)" . -lift_definition uminus_uint :: "uint \ uint" is uminus . -lift_definition times_uint :: "uint \ uint \ uint" is "(*)" . -instance by (standard; transfer) (simp_all add: algebra_simps) -end +declare uint.of_word_of [code abstype] -instantiation uint :: semiring_modulo -begin -lift_definition divide_uint :: "uint \ uint \ uint" is "(div)" . -lift_definition modulo_uint :: "uint \ uint \ uint" is "(mod)" . -instance by (standard; transfer) (fact word_mod_div_equality) -end +declare Quotient_uint [transfer_rule] -instantiation uint :: linorder begin -lift_definition less_uint :: "uint \ uint \ bool" is "(<)" . -lift_definition less_eq_uint :: "uint \ uint \ bool" is "(\)" . -instance by (standard; transfer) (simp_all add: less_le_not_le linear) -end - -lemmas [code] = less_uint.rep_eq less_eq_uint.rep_eq - -context - includes lifting_syntax - notes - transfer_rule_of_bool [transfer_rule] - transfer_rule_numeral [transfer_rule] +instantiation uint :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin -lemma [transfer_rule]: - "((=) ===> cr_uint) of_bool of_bool" - by transfer_prover - -lemma transfer_rule_numeral_uint [transfer_rule]: - "((=) ===> cr_uint) numeral numeral" - by transfer_prover - -lemma [transfer_rule]: - \(cr_uint ===> (\)) even ((dvd) 2 :: uint \ bool)\ - by (unfold dvd_def) transfer_prover - -end - -instantiation uint :: semiring_bits -begin +lift_definition zero_uint :: uint is 0 . +lift_definition one_uint :: uint is 1 . +lift_definition plus_uint :: \uint \ uint \ uint\ is \(+)\ . +lift_definition uminus_uint :: \uint \ uint\ is uminus . +lift_definition minus_uint :: \uint \ uint \ uint\ is \(-)\ . +lift_definition times_uint :: \uint \ uint \ uint\ is \(*)\ . +lift_definition divide_uint :: \uint \ uint \ uint\ is \(div)\ . +lift_definition modulo_uint :: \uint \ uint \ uint\ is \(mod)\ . +lift_definition equal_uint :: \uint \ uint \ bool\ is \HOL.equal\ . +lift_definition less_eq_uint :: \uint \ uint \ bool\ is \(\)\ . +lift_definition less_uint :: \uint \ uint \ bool\ is \(<)\ . -lift_definition bit_uint :: \uint \ 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)+ +global_interpretation uint: word_type_copy_ring Abs_uint Rep_uint + by standard (fact zero_uint.rep_eq one_uint.rep_eq + plus_uint.rep_eq uminus_uint.rep_eq minus_uint.rep_eq + times_uint.rep_eq divide_uint.rep_eq modulo_uint.rep_eq + equal_uint.rep_eq less_eq_uint.rep_eq less_uint.rep_eq)+ -end +instance proof - + show \OFCLASS(uint, comm_ring_1_class)\ + by (rule uint.of_class_comm_ring_1) + show \OFCLASS(uint, semiring_modulo_class)\ + by (fact uint.of_class_semiring_modulo) + show \OFCLASS(uint, equal_class)\ + by (fact uint.of_class_equal) + show \OFCLASS(uint, linorder_class)\ + by (fact uint.of_class_linorder) +qed -instantiation uint :: semiring_bit_shifts -begin -lift_definition push_bit_uint :: \nat \ uint \ uint\ is push_bit . -lift_definition drop_bit_uint :: \nat \ uint \ uint\ is drop_bit . -lift_definition take_bit_uint :: \nat \ uint \ uint\ is take_bit . -instance by (standard; transfer) - (fact push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod)+ end instantiation uint :: ring_bit_operations begin + +lift_definition bit_uint :: \uint \ nat \ bool\ is bit . lift_definition not_uint :: \uint \ uint\ is NOT . lift_definition and_uint :: \uint \ uint \ uint\ is \(AND)\ . lift_definition or_uint :: \uint \ uint \ uint\ is \(OR)\ . lift_definition xor_uint :: \uint \ uint \ uint\ is \(XOR)\ . lift_definition mask_uint :: \nat \ uint\ is mask . -lift_definition set_bit_uint :: \nat \ uint \ uint\ is \Bit_Operations.set_bit\ . -lift_definition unset_bit_uint :: \nat \ uint \ uint\ is \unset_bit\ . -lift_definition flip_bit_uint :: \nat \ uint \ uint\ is \flip_bit\ . -instance by (standard; transfer) - (simp_all add: bit_simps mask_eq_decr_exp minus_eq_not_minus_1 set_bit_def flip_bit_def) -end - -lemma [code]: - \take_bit n a = a AND mask n\ for a :: uint - by (fact take_bit_eq_mask) - -lemma [code]: - \mask (Suc n) = push_bit n (1 :: uint) OR mask n\ - \mask 0 = (0 :: uint)\ - by (simp_all add: mask_Suc_exp push_bit_of_1) - -lemma [code]: - \Bit_Operations.set_bit n w = w OR push_bit n 1\ for w :: uint - by (fact set_bit_eq_or) +lift_definition push_bit_uint :: \nat \ uint \ uint\ is push_bit . +lift_definition drop_bit_uint :: \nat \ uint \ uint\ is drop_bit . +lift_definition signed_drop_bit_uint :: \nat \ uint \ uint\ is signed_drop_bit . +lift_definition take_bit_uint :: \nat \ uint \ uint\ is take_bit . +lift_definition set_bit_uint :: \nat \ uint \ uint\ is Bit_Operations.set_bit . +lift_definition unset_bit_uint :: \nat \ uint \ uint\ is unset_bit . +lift_definition flip_bit_uint :: \nat \ uint \ uint\ is flip_bit . -lemma [code]: - \unset_bit n w = w AND NOT (push_bit n 1)\ for w :: uint - by (fact unset_bit_eq_and_not) - -lemma [code]: - \flip_bit n w = w XOR push_bit n 1\ for w :: uint - by (fact flip_bit_eq_xor) +global_interpretation uint: word_type_copy_bits Abs_uint Rep_uint signed_drop_bit_uint + by standard (fact bit_uint.rep_eq not_uint.rep_eq and_uint.rep_eq or_uint.rep_eq xor_uint.rep_eq + mask_uint.rep_eq push_bit_uint.rep_eq drop_bit_uint.rep_eq signed_drop_bit_uint.rep_eq take_bit_uint.rep_eq + set_bit_uint.rep_eq unset_bit_uint.rep_eq flip_bit_uint.rep_eq)+ -instantiation uint :: lsb -begin -lift_definition lsb_uint :: \uint \ bool\ is lsb . -instance by (standard; transfer) - (fact lsb_odd) -end - -instantiation uint :: msb -begin -lift_definition msb_uint :: \uint \ bool\ is msb . -instance .. -end - -setup \Context.theory_map (Name_Space.map_naming (Name_Space.qualified_path true \<^binding>\Generic\))\ - -instantiation uint :: set_bit -begin -lift_definition set_bit_uint :: \uint \ nat \ bool \ uint\ is set_bit . instance - apply standard - apply transfer - apply (simp add: bit_simps) - done -end - -setup \Context.theory_map (Name_Space.map_naming (Name_Space.parent_path))\ - -instantiation uint :: bit_comprehension begin -lift_definition set_bits_uint :: "(nat \ bool) \ uint" is "set_bits" . -instance by (standard; transfer) (fact set_bits_bit_eq) -end - -lemmas [code] = bit_uint.rep_eq lsb_uint.rep_eq msb_uint.rep_eq - -instantiation uint :: equal begin -lift_definition equal_uint :: "uint \ uint \ bool" is "equal_class.equal" . -instance by standard (transfer, simp add: equal_eq) -end - -lemmas [code] = equal_uint.rep_eq - -instantiation uint :: size begin -lift_definition size_uint :: "uint \ nat" is "size" . -instance .. -end - -lemmas [code] = size_uint.rep_eq - -lift_definition sshiftr_uint :: "uint \ nat \ uint" (infixl ">>>" 55) is \\w n. signed_drop_bit n w\ . - -lift_definition uint_of_int :: "int \ uint" is "word_of_int" . - -text \Use pretty numerals from integer for pretty printing\ - -context includes integer.lifting begin - -lift_definition Uint :: "integer \ uint" is "word_of_int" . - -lemma Rep_uint_numeral [simp]: "Rep_uint (numeral n) = numeral n" -by(induction n)(simp_all add: one_uint_def Abs_uint_inverse numeral.simps plus_uint_def) - -lemma numeral_uint_transfer [transfer_rule]: - "(rel_fun (=) cr_uint) numeral numeral" -by(auto simp add: cr_uint_def) - -lemma numeral_uint [code_unfold]: "numeral n = Uint (numeral n)" -by transfer simp - -lemma Rep_uint_neg_numeral [simp]: "Rep_uint (- numeral n) = - numeral n" -by(simp only: uminus_uint_def)(simp add: Abs_uint_inverse) - -lemma neg_numeral_uint [code_unfold]: "- numeral n = Uint (- numeral n)" -by transfer(simp add: cr_uint_def) + by (fact uint.of_class_ring_bit_operations) end -lemma Abs_uint_numeral [code_post]: "Abs_uint (numeral n) = numeral n" -by(induction n)(simp_all add: one_uint_def numeral.simps plus_uint_def Abs_uint_inverse) +lift_definition uint_of_nat :: \nat \ uint\ + is word_of_nat . -lemma Abs_uint_0 [code_post]: "Abs_uint 0 = 0" -by(simp add: zero_uint_def) +lift_definition nat_of_uint :: \uint \ nat\ + is unat . -lemma Abs_uint_1 [code_post]: "Abs_uint 1 = 1" -by(simp add: one_uint_def) +lift_definition uint_of_int :: \int \ uint\ + is word_of_int . + +lift_definition int_of_uint :: \uint \ int\ + is uint . + +context + includes integer.lifting +begin + +lift_definition Uint :: \integer \ uint\ + is word_of_int . + +lift_definition integer_of_uint :: \uint \ integer\ + is uint . + +end + +global_interpretation uint: word_type_copy_more Abs_uint Rep_uint signed_drop_bit_uint + uint_of_nat nat_of_uint uint_of_int int_of_uint Uint integer_of_uint + apply standard + apply (simp_all add: uint_of_nat.rep_eq nat_of_uint.rep_eq + uint_of_int.rep_eq int_of_uint.rep_eq + Uint.rep_eq integer_of_uint.rep_eq integer_eq_iff) + done + +instantiation uint :: "{size, msb, lsb, set_bit, bit_comprehension}" +begin + +lift_definition size_uint :: \uint \ nat\ is size . + +lift_definition msb_uint :: \uint \ bool\ is msb . +lift_definition lsb_uint :: \uint \ bool\ is lsb . + +text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ + +definition set_bit_uint :: \uint \ nat \ bool \ uint\ + where set_bit_uint_eq: \set_bit_uint a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ + +context + includes lifting_syntax +begin + +lemma set_bit_uint_transfer [transfer_rule]: + \(cr_uint ===> (=) ===> (\) ===> cr_uint) Generic_set_bit.set_bit Generic_set_bit.set_bit\ + by (simp only: set_bit_eq [abs_def] set_bit_uint_eq [abs_def]) transfer_prover + +end + +lift_definition set_bits_uint :: \(nat \ bool) \ uint\ is set_bits . +lift_definition set_bits_aux_uint :: \(nat \ bool) \ nat \ uint \ uint\ is set_bits_aux . + +global_interpretation uint: word_type_copy_misc Abs_uint Rep_uint signed_drop_bit_uint + uint_of_nat nat_of_uint uint_of_int int_of_uint Uint integer_of_uint dflt_size set_bits_aux_uint + by (standard; transfer) simp_all + +instance using uint.of_class_bit_comprehension + uint.of_class_set_bit uint.of_class_lsb + by simp_all standard + +end section \Code setup\ code_printing code_module Uint \ (SML) \ structure Uint : sig val set_bit : Word.word -> IntInf.int -> bool -> Word.word val shiftl : Word.word -> IntInf.int -> Word.word val shiftr : Word.word -> IntInf.int -> Word.word val shiftr_signed : Word.word -> IntInf.int -> Word.word val test_bit : Word.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word.orb (x, mask) else Word.andb (x, Word.notb mask) end fun shiftl x n = Word.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word.andb (x, Word.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word.fromInt 0 end; (* struct Uint *)\ code_reserved SML Uint code_printing code_module Uint \ (Haskell) \module Uint(Int, Word, dflt_size) where import qualified Prelude import Data.Int(Int) import Data.Word(Word) import qualified Data.Bits dflt_size :: Prelude.Integer dflt_size = Prelude.toInteger (bitSize_aux (0::Word)) where bitSize_aux :: (Data.Bits.Bits a, Prelude.Bounded a) => a -> Int bitSize_aux = Data.Bits.bitSize\ and (Haskell_Quickcheck) \module Uint(Int, Word, dflt_size) where import qualified Prelude import Data.Int(Int) import Data.Word(Word) import qualified Data.Bits dflt_size :: Prelude.Int dflt_size = bitSize_aux (0::Word) where bitSize_aux :: (Data.Bits.Bits a, Prelude.Bounded a) => a -> Int bitSize_aux = Data.Bits.bitSize \ code_reserved Haskell Uint dflt_size text \ OCaml and Scala provide only signed bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module "Uint" \ (OCaml) \module Uint : sig type t = int val dflt_size : Z.t val less : t -> t -> bool val less_eq : t -> t -> bool val set_bit : t -> Z.t -> bool -> t val shiftl : t -> Z.t -> t val shiftr : t -> Z.t -> t val shiftr_signed : t -> Z.t -> t val test_bit : t -> Z.t -> bool val int_mask : int val int32_mask : int32 val int64_mask : int64 end = struct type t = int let dflt_size = Z.of_int Sys.int_size;; (* negative numbers have their highest bit set, so they are greater than positive ones *) let less x y = if x<0 then y<0 && x 0;; let int_mask = if Sys.int_size < 32 then lnot 0 else 0xFFFFFFFF;; let int32_mask = if Sys.int_size < 32 then Int32.pred (Int32.shift_left Int32.one Sys.int_size) else Int32.of_string "0xFFFFFFFF";; let int64_mask = if Sys.int_size < 64 then Int64.pred (Int64.shift_left Int64.one Sys.int_size) else Int64.of_string "0xFFFFFFFFFFFFFFFF";; end;; (*struct Uint*)\ code_reserved OCaml Uint code_printing code_module Uint \ (Scala) \object Uint { def dflt_size : BigInt = BigInt(32) def less(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Int, n: BigInt, b: Boolean) : Int = if (b) x | (1 << n.intValue) else x & (1 << n.intValue).unary_~ def shiftl(x: Int, n: BigInt) : Int = x << n.intValue def shiftr(x: Int, n: BigInt) : Int = x >>> n.intValue def shiftr_signed(x: Int, n: BigInt) : Int = x >> n.intValue def test_bit(x: Int, n: BigInt) : Boolean = (x & (1 << n.intValue)) != 0 } /* object Uint */\ code_reserved Scala Uint text \ OCaml's conversion from Big\_int to int demands that the value fits into a signed integer. The following justifies the implementation. \ context includes integer.lifting begin definition wivs_mask :: int where "wivs_mask = 2^ dflt_size - 1" lift_definition wivs_mask_integer :: integer is wivs_mask . lemma [code]: "wivs_mask_integer = 2 ^ dflt_size - 1" by transfer (simp add: wivs_mask_def) definition wivs_shift :: int where "wivs_shift = 2 ^ dflt_size" lift_definition wivs_shift_integer :: integer is wivs_shift . lemma [code]: "wivs_shift_integer = 2 ^ dflt_size" by transfer (simp add: wivs_shift_def) definition wivs_index :: nat where "wivs_index == dflt_size - 1" lift_definition wivs_index_integer :: integer is "int wivs_index". lemma wivs_index_integer_code[code]: "wivs_index_integer = dflt_size_integer - 1" by transfer (simp add: wivs_index_def of_nat_diff) definition wivs_overflow :: int where "wivs_overflow == 2^ (dflt_size - 1)" lift_definition wivs_overflow_integer :: integer is wivs_overflow . lemma [code]: "wivs_overflow_integer = 2 ^ (dflt_size - 1)" by transfer (simp add: wivs_overflow_def) definition wivs_least :: int where "wivs_least == - wivs_overflow" lift_definition wivs_least_integer :: integer is wivs_least . lemma [code]: "wivs_least_integer = - (2 ^ (dflt_size - 1))" by transfer (simp add: wivs_overflow_def wivs_least_def) definition Uint_signed :: "integer \ uint" where "Uint_signed i = (if i < wivs_least_integer \ wivs_overflow_integer \ i then undefined Uint i else Uint i)" lemma Uint_code [code]: "Uint i = (let i' = i AND wivs_mask_integer in if bit i' wivs_index then Uint_signed (i' - wivs_shift_integer) else Uint_signed i')" including undefined_transfer unfolding Uint_signed_def apply transfer apply (subst word_of_int_via_signed) apply (auto simp add: push_bit_of_1 mask_eq_exp_minus_1 word_of_int_via_signed wivs_mask_def wivs_index_def wivs_overflow_def wivs_least_def wivs_shift_def) done -lemma Uint_signed_code [code abstract]: +lemma Uint_signed_code [code]: "Rep_uint (Uint_signed i) = (if i < wivs_least_integer \ i \ wivs_overflow_integer then Rep_uint (undefined Uint i) else word_of_int (int_of_integer_symbolic i))" unfolding Uint_signed_def Uint_def int_of_integer_symbolic_def word_of_integer_def by(simp add: Abs_uint_inverse) end text \ Avoid @{term Abs_uint} in generated code, use @{term Rep_uint'} instead. The symbolic implementations for code\_simp use @{term Rep_uint}. The new destructor @{term Rep_uint'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint} ([code abstract] equations for @{typ uint} may use @{term Rep_uint} because these instances will be folded away.) \ definition Rep_uint' where [simp]: "Rep_uint' = Rep_uint" lemma Rep_uint'_code [code]: "Rep_uint' x = (BITS n. bit x n)" unfolding Rep_uint'_def by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint' :: "dflt_size word \ uint" is "\x :: dflt_size word. x" . lemma Abs_uint'_code [code]: "Abs_uint' x = Uint (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint \ _"]] lemma term_of_uint_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint.uint.Abs_uint'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR (STR ''Uint.dflt_size'') []], TR (STR ''Uint.uint'') []])) (term_of_class.term_of (Rep_uint' x))" by(simp add: term_of_anything) text \Important: We must prevent the reflection oracle (eval-tac) to use our machine-dependent type. \ code_printing type_constructor uint \ (SML) "Word.word" and (Haskell) "Uint.Word" and (OCaml) "Uint.t" and (Scala) "Int" and (Eval) "*** \"Error: Machine dependent type\" ***" and (Quickcheck) "Word.word" | constant dflt_size_integer \ (SML) "(IntInf.fromLarge (Int.toLarge Word.wordSize))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.wordSize" and (Haskell) "Uint.dflt'_size" and (OCaml) "Uint.dflt'_size" and (Scala) "Uint.dflt'_size" | constant Uint \ (SML) "Word.fromLargeInt (IntInf.toLarge _)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.fromInt" and (Haskell) "(Prelude.fromInteger _ :: Uint.Word)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint.Word)" and (Scala) "_.intValue" | constant Uint_signed \ (OCaml) "Z.to'_int" | constant "0 :: uint" \ (SML) "(Word.fromInt 0)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "(Word.fromInt 0)" and (Haskell) "(0 :: Uint.Word)" and (OCaml) "0" and (Scala) "0" | constant "1 :: uint" \ (SML) "(Word.fromInt 1)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "(Word.fromInt 1)" and (Haskell) "(1 :: Uint.Word)" and (OCaml) "1" and (Scala) "1" | constant "plus :: uint \ _ " \ (SML) "Word.+ ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.+ ((_), (_))" and (Haskell) infixl 6 "+" and (OCaml) "Pervasives.(+)" and (Scala) infixl 7 "+" | constant "uminus :: uint \ _" \ (SML) "Word.~" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.~" and (Haskell) "negate" and (OCaml) "Pervasives.(~-)" and (Scala) "!(- _)" | constant "minus :: uint \ _" \ (SML) "Word.- ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.- ((_), (_))" and (Haskell) infixl 6 "-" and (OCaml) "Pervasives.(-)" and (Scala) infixl 7 "-" | constant "times :: uint \ _ \ _" \ (SML) "Word.* ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.* ((_), (_))" and (Haskell) infixl 7 "*" and (OCaml) "Pervasives.( * )" and (Scala) infixl 8 "*" | constant "HOL.equal :: uint \ _ \ bool" \ (SML) "!((_ : Word.word) = _)" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "!((_ : Word.word) = _)" and (Haskell) infix 4 "==" and (OCaml) "(Pervasives.(=):Uint.t -> Uint.t -> bool)" and (Scala) infixl 5 "==" | class_instance uint :: equal \ (Haskell) - | constant "less_eq :: uint \ _ \ bool" \ (SML) "Word.<= ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.<= ((_), (_))" and (Haskell) infix 4 "<=" and (OCaml) "Uint.less'_eq" and (Scala) "Uint.less'_eq" | constant "less :: uint \ _ \ bool" \ (SML) "Word.< ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.< ((_), (_))" and (Haskell) infix 4 "<" and (OCaml) "Uint.less" and (Scala) "Uint.less" | constant "NOT :: uint \ _" \ (SML) "Word.notb" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.notb" and (Haskell) "Data'_Bits.complement" and (OCaml) "Pervasives.lnot" and (Scala) "_.unary'_~" | constant "(AND) :: uint \ _" \ (SML) "Word.andb ((_),/ (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (OCaml) "Pervasives.(land)" and (Scala) infixl 3 "&" | constant "(OR) :: uint \ _" \ (SML) "Word.orb ((_),/ (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (OCaml) "Pervasives.(lor)" and (Scala) infixl 1 "|" | constant "(XOR) :: uint \ _" \ (SML) "Word.xorb ((_),/ (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.xorb ((_),/ (_))" and (Haskell) "Data'_Bits.xor" and (OCaml) "Pervasives.(lxor)" and (Scala) infixl 2 "^" definition uint_divmod :: "uint \ uint \ uint \ uint" where "uint_divmod x y = (if y = 0 then (undefined ((div) :: uint \ _) x (0 :: uint), undefined ((mod) :: uint \ _) x (0 :: uint)) else (x div y, x mod y))" definition uint_div :: "uint \ uint \ uint" where "uint_div x y = fst (uint_divmod x y)" definition uint_mod :: "uint \ uint \ uint" where "uint_mod x y = snd (uint_divmod x y)" lemma div_uint_code [code]: "x div y = (if y = 0 then 0 else uint_div x y)" including undefined_transfer unfolding uint_divmod_def uint_div_def by transfer(simp add: word_div_def) lemma mod_uint_code [code]: "x mod y = (if y = 0 then x else uint_mod x y)" including undefined_transfer unfolding uint_mod_def uint_divmod_def by transfer(simp add: word_mod_def) definition uint_sdiv :: "uint \ uint \ uint" where [code del]: "uint_sdiv x y = (if y = 0 then undefined ((div) :: uint \ _) x (0 :: uint) else Abs_uint (Rep_uint x sdiv Rep_uint y))" definition div0_uint :: "uint \ uint" where [code del]: "div0_uint x = undefined ((div) :: uint \ _) x (0 :: uint)" declare [[code abort: div0_uint]] definition mod0_uint :: "uint \ uint" where [code del]: "mod0_uint x = undefined ((mod) :: uint \ _) x (0 :: uint)" declare [[code abort: mod0_uint]] definition wivs_overflow_uint :: uint where "wivs_overflow_uint \ push_bit (dflt_size - 1) 1" +lemma Rep_uint_wivs_overflow_uint_eq: + \Rep_uint wivs_overflow_uint = 2 ^ (dflt_size - Suc 0)\ + by (simp add: wivs_overflow_uint_def one_uint.rep_eq push_bit_uint.rep_eq uint.word_of_power push_bit_eq_mult) + +lemma wivs_overflow_uint_greater_eq_0: + \wivs_overflow_uint > 0\ + apply (simp add: less_uint.rep_eq zero_uint.rep_eq Rep_uint_wivs_overflow_uint_eq) + apply transfer + apply (simp add: take_bit_push_bit push_bit_eq_mult) + done + lemma uint_divmod_code [code]: "uint_divmod x y = (if wivs_overflow_uint \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint x, mod0_uint x) else let q = push_bit 1 (uint_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" proof (cases \y = 0\) case True moreover have \x \ 0\ by transfer simp - moreover have \wivs_overflow_uint > 0\ - apply (simp add: wivs_overflow_uint_def push_bit_of_1) - apply transfer - apply transfer - apply simp - done + moreover note wivs_overflow_uint_greater_eq_0 ultimately show ?thesis by (auto simp add: uint_divmod_def div0_uint_def mod0_uint_def not_less) next case False then show ?thesis including undefined_transfer unfolding uint_divmod_def uint_sdiv_def div0_uint_def mod0_uint_def wivs_overflow_uint_def apply transfer apply (simp add: divmod_via_sdivmod push_bit_of_1) done qed -lemma uint_sdiv_code [code abstract]: +lemma uint_sdiv_code [code]: "Rep_uint (uint_sdiv x y) = (if y = 0 then Rep_uint (undefined ((div) :: uint \ _) x (0 :: uint)) else Rep_uint x sdiv Rep_uint y)" unfolding uint_sdiv_def by(simp add: Abs_uint_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint_divmod_code} computes both with division only. \ code_printing constant uint_div \ (SML) "Word.div ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.div ((_), (_))" and (Haskell) "Prelude.div" | constant uint_mod \ (SML) "Word.mod ((_), (_))" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Word.mod ((_), (_))" and (Haskell) "Prelude.mod" | constant uint_divmod \ (Haskell) "divmod" | constant uint_sdiv \ (OCaml) "Pervasives.('/)" and (Scala) "_ '/ _" definition uint_test_bit :: "uint \ integer \ bool" where [code del]: "uint_test_bit x n = (if n < 0 \ dflt_size_integer \ n then undefined (bit :: uint \ _) x n else bit x (nat_of_integer n))" lemma test_bit_uint_code [code]: "bit x n \ n < dflt_size \ uint_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint_test_bit_def by (transfer, simp, transfer, simp) lemma uint_test_bit_code [code]: "uint_test_bit w n = (if n < 0 \ dflt_size_integer \ n then undefined (bit :: uint \ _) w n else bit (Rep_uint w) (nat_of_integer n))" unfolding uint_test_bit_def by(simp add: bit_uint.rep_eq) code_printing constant uint_test_bit \ (SML) "Uint.test'_bit" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (OCaml) "Uint.test'_bit" and (Scala) "Uint.test'_bit" definition uint_set_bit :: "uint \ integer \ bool \ uint" where [code del]: "uint_set_bit x n b = (if n < 0 \ dflt_size_integer \ n then undefined (set_bit :: uint \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint_code [code]: "set_bit x n b = (if n < dflt_size then uint_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint_set_bit_def by (transfer) (auto cong: conj_cong simp add: not_less set_bit_beyond word_size) -lemma uint_set_bit_code [code abstract]: +lemma uint_set_bit_code [code]: "Rep_uint (uint_set_bit w n b) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined (set_bit :: uint \ _) w n b) else set_bit (Rep_uint w) (nat_of_integer n) b)" including undefined_transfer integer.lifting unfolding uint_set_bit_def by transfer simp code_printing constant uint_set_bit \ (SML) "Uint.set'_bit" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (OCaml) "Uint.set'_bit" and (Scala) "Uint.set'_bit" -lift_definition uint_set_bits :: "(nat \ bool) \ uint \ nat \ uint" is set_bits_aux . - -lemma uint_set_bits_code [code]: - "uint_set_bits f w n = - (if n = 0 then w - else let n' = n - 1 in uint_set_bits f (push_bit 1 w OR (if f n' then 1 else 0)) n')" - apply (transfer fixing: n) - apply (cases n) - apply simp_all - done - -lemma set_bits_uint [code]: - "(BITS n. f n) = uint_set_bits f 0 dflt_size" - by transfer (simp add: set_bits_conv_set_bits_aux) - -lemma lsb_code [code]: fixes x :: uint shows "lsb x = bit x 0" - by transfer (simp add: lsb_word_eq) - definition uint_shiftl :: "uint \ integer \ uint" where [code del]: "uint_shiftl x n = (if n < 0 \ dflt_size_integer \ n then undefined (push_bit :: nat \ uint \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint_code [code]: "push_bit n x = (if n < dflt_size then uint_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint_shiftl_def by (transfer fixing: n) simp -lemma uint_shiftl_code [code abstract]: +lemma uint_shiftl_code [code]: "Rep_uint (uint_shiftl w n) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined (push_bit :: nat \ uint \ _) w n) else push_bit (nat_of_integer n) (Rep_uint w))" including undefined_transfer integer.lifting unfolding uint_shiftl_def by transfer simp code_printing constant uint_shiftl \ (SML) "Uint.shiftl" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (OCaml) "Uint.shiftl" and (Scala) "Uint.shiftl" definition uint_shiftr :: "uint \ integer \ uint" where [code del]: "uint_shiftr x n = (if n < 0 \ dflt_size_integer \ n then undefined (drop_bit :: nat \ uint \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint_code [code]: "drop_bit n x = (if n < dflt_size then uint_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint_shiftr_def by (transfer fixing: n) simp - -lemma uint_shiftr_code [code abstract]: + +lemma uint_shiftr_code [code]: "Rep_uint (uint_shiftr w n) = (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined (drop_bit :: nat \ uint \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint w))" -including undefined_transfer unfolding uint_shiftr_def by transfer simp + including undefined_transfer integer.lifting unfolding uint_shiftr_def by transfer simp code_printing constant uint_shiftr \ (SML) "Uint.shiftr" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (OCaml) "Uint.shiftr" and (Scala) "Uint.shiftr" definition uint_sshiftr :: "uint \ integer \ uint" where [code del]: "uint_sshiftr x n = - (if n < 0 \ dflt_size_integer \ n then undefined sshiftr_uint x n else sshiftr_uint x (nat_of_integer n))" + (if n < 0 \ dflt_size_integer \ n then undefined signed_drop_bit_uint n x else signed_drop_bit_uint (nat_of_integer n) x)" lemma sshiftr_uint_code [code]: - "x >>> n = + "signed_drop_bit_uint n x = (if n < dflt_size then uint_sshiftr x (integer_of_nat n) else if bit x wivs_index then -1 else 0)" including undefined_transfer integer.lifting unfolding uint_sshiftr_def by transfer(simp add: not_less signed_drop_bit_beyond word_size wivs_index_def) -lemma uint_sshiftr_code [code abstract]: +lemma uint_sshiftr_code [code]: "Rep_uint (uint_sshiftr w n) = - (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined sshiftr_uint w n) else signed_drop_bit (nat_of_integer n) (Rep_uint w))" -including undefined_transfer unfolding uint_sshiftr_def by transfer simp + (if n < 0 \ dflt_size_integer \ n then Rep_uint (undefined signed_drop_bit_uint n w) else signed_drop_bit (nat_of_integer n) (Rep_uint w))" +including undefined_transfer integer.lifting unfolding uint_sshiftr_def by transfer simp code_printing constant uint_sshiftr \ (SML) "Uint.shiftr'_signed" and (Eval) "(raise (Fail \"Machine dependent code\"))" and (Quickcheck) "Uint.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint.Int) _)) :: Uint.Word)" and (OCaml) "Uint.shiftr'_signed" and (Scala) "Uint.shiftr'_signed" lemma uint_msb_test_bit: "msb x \ bit (x :: uint) wivs_index" by transfer (simp add: msb_word_iff_bit wivs_index_def) lemma msb_uint_code [code]: "msb x \ uint_test_bit x wivs_index_integer" apply(simp add: uint_test_bit_def uint_msb_test_bit wivs_index_integer_code dflt_size_integer_def wivs_index_def) by (metis (full_types) One_nat_def dflt_size(2) less_iff_diff_less_0 nat_of_integer_of_nat of_nat_1 of_nat_diff of_nat_less_0_iff wivs_index_def) lemma uint_of_int_code [code]: "uint_of_int i = (BITS n. bit i n)" by transfer (simp add: word_of_int_conv_set_bits) section \Quickcheck setup\ definition uint_of_natural :: "natural \ uint" where "uint_of_natural x \ Uint (integer_of_natural x)" instantiation uint :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint \ qc_random_cnv uint_of_natural" definition "exhaustive_uint \ qc_exhaustive_cnv uint_of_natural" definition "full_exhaustive_uint \ qc_full_exhaustive_cnv uint_of_natural" instance .. end instantiation uint :: narrowing begin interpretation quickcheck_narrowing_samples "\i. (Uint i, Uint (- i))" "0" "Typerep.Typerep (STR ''Uint.uint'') []" . definition "narrowing_uint d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint itself \ _"]] lemmas partial_term_of_uint [code] = partial_term_of_code instance .. end -no_notation sshiftr_uint (infixl ">>>" 55) - 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,620 +1,532 @@ (* Title: Uint16.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 16 bits\ theory Uint16 imports - Code_Target_Word_Base + Code_Target_Word_Base Word_Type_Copies begin text \ Restriction for ML code generation: This theory assumes that the ML system provides a Word16 implementation (mlton does, but PolyML 5.5 does not). Therefore, the code setup lives in the target \SML_word\ rather than \SML\. This ensures that code generation still works as long as \uint16\ is not involved. For the target \SML\ itself, no special code generation - for this type is set up. Nevertheless, it should work by emulation via @{typ "16 word"} + for this type is set up. Nevertheless, it should work by emulation via \<^typ>\16 word\ if the theory \Code_Target_Bits_Int\ is imported. Restriction for OCaml code generation: OCaml does not provide an int16 type, so no special code generation for this type is set up. \ section \Type definition and primitive operations\ -typedef uint16 = "UNIV :: 16 word set" .. +typedef uint16 = \UNIV :: 16 word set\ .. + +global_interpretation uint16: word_type_copy Abs_uint16 Rep_uint16 + using type_definition_uint16 by (rule word_type_copy.intro) setup_lifting type_definition_uint16 -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 +declare uint16.of_word_of [code abstype] -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 +declare Quotient_uint16 [transfer_rule] -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] +instantiation uint16 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ 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 zero_uint16 :: uint16 is 0 . +lift_definition one_uint16 :: uint16 is 1 . +lift_definition plus_uint16 :: \uint16 \ uint16 \ uint16\ is \(+)\ . +lift_definition uminus_uint16 :: \uint16 \ uint16\ is uminus . +lift_definition minus_uint16 :: \uint16 \ uint16 \ uint16\ is \(-)\ . +lift_definition times_uint16 :: \uint16 \ uint16 \ uint16\ is \(*)\ . +lift_definition divide_uint16 :: \uint16 \ uint16 \ uint16\ is \(div)\ . +lift_definition modulo_uint16 :: \uint16 \ uint16 \ uint16\ is \(mod)\ . +lift_definition equal_uint16 :: \uint16 \ uint16 \ bool\ is \HOL.equal\ . +lift_definition less_eq_uint16 :: \uint16 \ uint16 \ bool\ is \(\)\ . +lift_definition less_uint16 :: \uint16 \ uint16 \ bool\ is \(<)\ . -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)+ +global_interpretation uint16: word_type_copy_ring Abs_uint16 Rep_uint16 + by standard (fact zero_uint16.rep_eq one_uint16.rep_eq + plus_uint16.rep_eq uminus_uint16.rep_eq minus_uint16.rep_eq + times_uint16.rep_eq divide_uint16.rep_eq modulo_uint16.rep_eq + equal_uint16.rep_eq less_eq_uint16.rep_eq less_uint16.rep_eq)+ -end +instance proof - + show \OFCLASS(uint16, comm_ring_1_class)\ + by (rule uint16.of_class_comm_ring_1) + show \OFCLASS(uint16, semiring_modulo_class)\ + by (fact uint16.of_class_semiring_modulo) + show \OFCLASS(uint16, equal_class)\ + by (fact uint16.of_class_equal) + show \OFCLASS(uint16, linorder_class)\ + by (fact uint16.of_class_linorder) +qed -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 bit_uint16 :: \uint16 \ nat \ bool\ is bit . lift_definition not_uint16 :: \uint16 \ uint16\ is NOT . lift_definition and_uint16 :: \uint16 \ uint16 \ uint16\ is \(AND)\ . lift_definition or_uint16 :: \uint16 \ uint16 \ uint16\ is \(OR)\ . lift_definition xor_uint16 :: \uint16 \ uint16 \ uint16\ is \(XOR)\ . lift_definition mask_uint16 :: \nat \ uint16\ is mask . -lift_definition set_bit_uint16 :: \nat \ uint16 \ uint16\ is \Bit_Operations.set_bit\ . -lift_definition unset_bit_uint16 :: \nat \ uint16 \ uint16\ is \unset_bit\ . -lift_definition flip_bit_uint16 :: \nat \ uint16 \ uint16\ is \flip_bit\ . -instance by (standard; transfer) - (simp_all add: bit_simps mask_eq_decr_exp minus_eq_not_minus_1 set_bit_def flip_bit_def) -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) - -lemma [code]: - \Bit_Operations.set_bit n w = w OR push_bit n 1\ for w :: uint16 - by (fact set_bit_eq_or) - -lemma [code]: - \unset_bit n w = w AND NOT (push_bit n 1)\ for w :: uint16 - by (fact unset_bit_eq_and_not) - -lemma [code]: - \flip_bit n w = w XOR push_bit n 1\ for w :: uint16 - by (fact flip_bit_eq_xor) - -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 - -setup \Context.theory_map (Name_Space.map_naming (Name_Space.qualified_path true \<^binding>\Generic\))\ - -instantiation uint16 :: set_bit -begin -lift_definition set_bit_uint16 :: \uint16 \ nat \ bool \ uint16\ is set_bit . -instance - apply standard - apply transfer - apply (simp add: bit_simps) - done -end - -setup \Context.theory_map (Name_Space.map_naming (Name_Space.parent_path))\ +lift_definition push_bit_uint16 :: \nat \ uint16 \ uint16\ is push_bit . +lift_definition drop_bit_uint16 :: \nat \ uint16 \ uint16\ is drop_bit . +lift_definition signed_drop_bit_uint16 :: \nat \ uint16 \ uint16\ is signed_drop_bit . +lift_definition take_bit_uint16 :: \nat \ uint16 \ uint16\ is take_bit . +lift_definition set_bit_uint16 :: \nat \ uint16 \ uint16\ is Bit_Operations.set_bit . +lift_definition unset_bit_uint16 :: \nat \ uint16 \ uint16\ is unset_bit . +lift_definition flip_bit_uint16 :: \nat \ uint16 \ uint16\ is flip_bit . -instantiation uint16 :: bit_comprehension begin -lift_definition set_bits_uint16 :: "(nat \ bool) \ uint16" is "set_bits" . -instance by (standard; transfer) (fact set_bits_bit_eq) -end - -lemmas [code] = 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 \\w n. signed_drop_bit n w\ . - -lift_definition uint16_of_int :: "int \ uint16" is "word_of_int" . +global_interpretation uint16: word_type_copy_bits Abs_uint16 Rep_uint16 signed_drop_bit_uint16 + by standard (fact bit_uint16.rep_eq not_uint16.rep_eq and_uint16.rep_eq or_uint16.rep_eq xor_uint16.rep_eq + mask_uint16.rep_eq push_bit_uint16.rep_eq drop_bit_uint16.rep_eq signed_drop_bit_uint16.rep_eq take_bit_uint16.rep_eq + set_bit_uint16.rep_eq unset_bit_uint16.rep_eq flip_bit_uint16.rep_eq)+ -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) +instance + by (fact uint16.of_class_ring_bit_operations) 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) +lift_definition uint16_of_nat :: \nat \ uint16\ + is word_of_nat . -lemma Abs_uint16_0 [code_post]: "Abs_uint16 0 = 0" -by(simp add: zero_uint16_def) +lift_definition nat_of_uint16 :: \uint16 \ nat\ + is unat . -lemma Abs_uint16_1 [code_post]: "Abs_uint16 1 = 1" -by(simp add: one_uint16_def) +lift_definition uint16_of_int :: \int \ uint16\ + is word_of_int . + +lift_definition int_of_uint16 :: \uint16 \ int\ + is uint . + +context + includes integer.lifting +begin + +lift_definition Uint16 :: \integer \ uint16\ + is word_of_int . + +lift_definition integer_of_uint16 :: \uint16 \ integer\ + is uint . + +end + +global_interpretation uint16: word_type_copy_more Abs_uint16 Rep_uint16 signed_drop_bit_uint16 + uint16_of_nat nat_of_uint16 uint16_of_int int_of_uint16 Uint16 integer_of_uint16 + apply standard + apply (simp_all add: uint16_of_nat.rep_eq nat_of_uint16.rep_eq + uint16_of_int.rep_eq int_of_uint16.rep_eq + Uint16.rep_eq integer_of_uint16.rep_eq integer_eq_iff) + done + +instantiation uint16 :: "{size, msb, lsb, set_bit, bit_comprehension}" +begin + +lift_definition size_uint16 :: \uint16 \ nat\ is size . + +lift_definition msb_uint16 :: \uint16 \ bool\ is msb . +lift_definition lsb_uint16 :: \uint16 \ bool\ is lsb . + +text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ + +definition set_bit_uint16 :: \uint16 \ nat \ bool \ uint16\ + where set_bit_uint16_eq: \set_bit_uint16 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ + +context + includes lifting_syntax +begin + +lemma set_bit_uint16_transfer [transfer_rule]: + \(cr_uint16 ===> (=) ===> (\) ===> cr_uint16) Generic_set_bit.set_bit Generic_set_bit.set_bit\ + by (simp only: set_bit_eq [abs_def] set_bit_uint16_eq [abs_def]) transfer_prover + +end + +lift_definition set_bits_uint16 :: \(nat \ bool) \ uint16\ is set_bits . +lift_definition set_bits_aux_uint16 :: \(nat \ bool) \ nat \ uint16 \ uint16\ is set_bits_aux . + +global_interpretation uint16: word_type_copy_misc Abs_uint16 Rep_uint16 signed_drop_bit_uint16 + uint16_of_nat nat_of_uint16 uint16_of_int int_of_uint16 Uint16 integer_of_uint16 16 set_bits_aux_uint16 + by (standard; transfer) simp_all + +instance using uint16.of_class_bit_comprehension + uint16.of_class_set_bit uint16.of_class_lsb + by simp_all standard + +end section \Code setup\ code_printing code_module Uint16 \ (SML_word) \(* Test that words can handle numbers between 0 and 15 *) val _ = if 4 <= Word.wordSize then () else raise (Fail ("wordSize less than 4")); structure Uint16 : sig val set_bit : Word16.word -> IntInf.int -> bool -> Word16.word val shiftl : Word16.word -> IntInf.int -> Word16.word val shiftr : Word16.word -> IntInf.int -> Word16.word val shiftr_signed : Word16.word -> IntInf.int -> Word16.word val test_bit : Word16.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word16.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word16.orb (x, mask) else Word16.andb (x, Word16.notb mask) end fun shiftl x n = Word16.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word16.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word16.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word16.andb (x, Word16.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word16.fromInt 0 end; (* struct Uint16 *)\ code_reserved SML_word Uint16 code_printing code_module Uint16 \ (Haskell) \module Uint16(Int16, Word16) where import Data.Int(Int16) import Data.Word(Word16)\ code_reserved Haskell Uint16 text \Scala provides unsigned 16-bit numbers as Char.\ code_printing code_module Uint16 \ (Scala) \object Uint16 { def set_bit(x: scala.Char, n: BigInt, b: Boolean) : scala.Char = if (b) (x | (1.toChar << n.intValue)).toChar else (x & (1.toChar << n.intValue).unary_~).toChar def shiftl(x: scala.Char, n: BigInt) : scala.Char = (x << n.intValue).toChar def shiftr(x: scala.Char, n: BigInt) : scala.Char = (x >>> n.intValue).toChar def shiftr_signed(x: scala.Char, n: BigInt) : scala.Char = (x.toShort >> n.intValue).toChar def test_bit(x: scala.Char, n: BigInt) : Boolean = (x & (1.toChar << n.intValue)) != 0 } /* object Uint16 */\ code_reserved Scala Uint16 text \ Avoid @{term Abs_uint16} in generated code, use @{term Rep_uint16'} instead. The symbolic implementations for code\_simp use @{term Rep_uint16}. The new destructor @{term Rep_uint16'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint16} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint16} ([code abstract] equations for @{typ uint16} may use @{term Rep_uint16} because these instances will be folded away.) To convert @{typ "16 word"} values into @{typ uint16}, use @{term "Abs_uint16'"}. \ definition Rep_uint16' where [simp]: "Rep_uint16' = Rep_uint16" lemma Rep_uint16'_transfer [transfer_rule]: "rel_fun cr_uint16 (=) (\x. x) Rep_uint16'" unfolding Rep_uint16'_def by(rule uint16.rep_transfer) lemma Rep_uint16'_code [code]: "Rep_uint16' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint16' :: "16 word \ uint16" is "\x :: 16 word. x" . lemma Abs_uint16'_code [code]: "Abs_uint16' x = Uint16 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint16 \ _"]] lemma term_of_uint16_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint16.uint16.Abs_uint16'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]]], TR (STR ''Uint16.uint16'') []])) (term_of_class.term_of (Rep_uint16' x))" by(simp add: term_of_anything) -lemma Uin16_code [code abstract]: "Rep_uint16 (Uint16 i) = word_of_int (int_of_integer_symbolic i)" +lemma Uint16_code [code]: "Rep_uint16 (Uint16 i) = word_of_int (int_of_integer_symbolic i)" unfolding Uint16_def int_of_integer_symbolic_def by(simp add: Abs_uint16_inverse) code_printing type_constructor uint16 \ (SML_word) "Word16.word" and (Haskell) "Uint16.Word16" and (Scala) "scala.Char" | constant Uint16 \ (SML_word) "Word16.fromLargeInt (IntInf.toLarge _)" and (Haskell) "(Prelude.fromInteger _ :: Uint16.Word16)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint16.Word16)" and (Scala) "_.charValue" | constant "0 :: uint16" \ (SML_word) "(Word16.fromInt 0)" and (Haskell) "(0 :: Uint16.Word16)" and (Scala) "0" | constant "1 :: uint16" \ (SML_word) "(Word16.fromInt 1)" and (Haskell) "(1 :: Uint16.Word16)" and (Scala) "1" | constant "plus :: uint16 \ _ \ _" \ (SML_word) "Word16.+ ((_), (_))" and (Haskell) infixl 6 "+" and (Scala) "(_ +/ _).toChar" | constant "uminus :: uint16 \ _" \ (SML_word) "Word16.~" and (Haskell) "negate" and (Scala) "(- _).toChar" | constant "minus :: uint16 \ _" \ (SML_word) "Word16.- ((_), (_))" and (Haskell) infixl 6 "-" and (Scala) "(_ -/ _).toChar" | constant "times :: uint16 \ _ \ _" \ (SML_word) "Word16.* ((_), (_))" and (Haskell) infixl 7 "*" and (Scala) "(_ */ _).toChar" | constant "HOL.equal :: uint16 \ _ \ bool" \ (SML_word) "!((_ : Word16.word) = _)" and (Haskell) infix 4 "==" and (Scala) infixl 5 "==" | class_instance uint16 :: equal \ (Haskell) - | constant "less_eq :: uint16 \ _ \ bool" \ (SML_word) "Word16.<= ((_), (_))" and (Haskell) infix 4 "<=" and (Scala) infixl 4 "<=" | constant "less :: uint16 \ _ \ bool" \ (SML_word) "Word16.< ((_), (_))" and (Haskell) infix 4 "<" and (Scala) infixl 4 "<" | constant "NOT :: uint16 \ _" \ (SML_word) "Word16.notb" and (Haskell) "Data'_Bits.complement" and (Scala) "_.unary'_~.toChar" | constant "(AND) :: uint16 \ _" \ (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]: +lemma uint16_div_code [code]: "Rep_uint16 (uint16_div x y) = (if y = 0 then Rep_uint16 (undefined ((div) :: uint16 \ _) x (0 :: uint16)) else Rep_uint16 x div Rep_uint16 y)" unfolding uint16_div_def by transfer simp -lemma uint16_mod_code [code abstract]: +lemma uint16_mod_code [code]: "Rep_uint16 (uint16_mod x y) = (if y = 0 then Rep_uint16 (undefined ((mod) :: uint16 \ _) x (0 :: uint16)) else Rep_uint16 x mod Rep_uint16 y)" unfolding uint16_mod_def by transfer simp end code_printing constant uint16_div \ (SML_word) "Word16.div ((_), (_))" and (Haskell) "Prelude.div" and (Scala) "(_ '/ _).toChar" | constant uint16_mod \ (SML_word) "Word16.mod ((_), (_))" and (Haskell) "Prelude.mod" and (Scala) "(_ % _).toChar" definition uint16_test_bit :: "uint16 \ integer \ bool" where [code del]: "uint16_test_bit x n = (if n < 0 \ 15 < n then undefined (bit :: uint16 \ _) x n else bit x (nat_of_integer n))" lemma test_bit_uint16_code [code]: "bit x n \ n < 16 \ uint16_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint16_test_bit_def by (transfer, simp, transfer, simp) lemma uint16_test_bit_code [code]: "uint16_test_bit w n = (if n < 0 \ 15 < n then undefined (bit :: uint16 \ _) w n else bit (Rep_uint16 w) (nat_of_integer n))" unfolding uint16_test_bit_def by (simp add: bit_uint16.rep_eq) code_printing constant uint16_test_bit \ (SML_word) "Uint16.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (Scala) "Uint16.test'_bit" definition uint16_set_bit :: "uint16 \ integer \ bool \ uint16" where [code del]: "uint16_set_bit x n b = (if n < 0 \ 15 < n then undefined (set_bit :: uint16 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint16_code [code]: "set_bit x n b = (if n < 16 then uint16_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint16_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) -lemma uint16_set_bit_code [code abstract]: +lemma uint16_set_bit_code [code]: "Rep_uint16 (uint16_set_bit w n b) = (if n < 0 \ 15 < n then Rep_uint16 (undefined (set_bit :: uint16 \ _) w n b) else set_bit (Rep_uint16 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint16_set_bit_def by transfer simp code_printing constant uint16_set_bit \ (SML_word) "Uint16.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (Scala) "Uint16.set'_bit" -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 ((push_bit 1 w) OR (if f n' then 1 else 0)) n')" - apply (transfer fixing: n) - apply (cases n) - apply simp_all - done - -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 \ bit x 0" - by transfer (simp add: lsb_odd) - definition uint16_shiftl :: "uint16 \ integer \ uint16" where [code del]: "uint16_shiftl x n = (if n < 0 \ 16 \ n then undefined (push_bit :: nat \ uint16 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint16_code [code]: "push_bit n x = (if n < 16 then uint16_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint16_shiftl_def by transfer simp -lemma uint16_shiftl_code [code abstract]: +lemma uint16_shiftl_code [code]: "Rep_uint16 (uint16_shiftl w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined (push_bit :: nat \ uint16 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint16 w))" including undefined_transfer unfolding uint16_shiftl_def by transfer simp code_printing constant uint16_shiftl \ (SML_word) "Uint16.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (Scala) "Uint16.shiftl" definition uint16_shiftr :: "uint16 \ integer \ uint16" where [code del]: "uint16_shiftr x n = (if n < 0 \ 16 \ n then undefined (drop_bit :: nat \ uint16 \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint16_code [code]: "drop_bit n x = (if n < 16 then uint16_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint16_shiftr_def by transfer simp -lemma uint16_shiftr_code [code abstract]: +lemma uint16_shiftr_code [code]: "Rep_uint16 (uint16_shiftr w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined (drop_bit :: nat \ uint16 \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint16 w))" including undefined_transfer unfolding uint16_shiftr_def by transfer simp code_printing constant uint16_shiftr \ (SML_word) "Uint16.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (Scala) "Uint16.shiftr" definition uint16_sshiftr :: "uint16 \ integer \ uint16" where [code del]: "uint16_sshiftr x n = - (if n < 0 \ 16 \ n then undefined sshiftr_uint16 x n else sshiftr_uint16 x (nat_of_integer n))" + (if n < 0 \ 16 \ n then undefined signed_drop_bit_uint16 n x else signed_drop_bit_uint16 (nat_of_integer n) x)" lemma sshiftr_uint16_code [code]: - "x >>> n = + "signed_drop_bit_uint16 n x = (if n < 16 then uint16_sshiftr x (integer_of_nat n) else if bit x 15 then -1 else 0)" -including undefined_transfer integer.lifting unfolding uint16_sshiftr_def -by transfer (simp add: not_less signed_drop_bit_beyond word_size) + including undefined_transfer integer.lifting unfolding uint16_sshiftr_def + by transfer (simp add: not_less signed_drop_bit_beyond word_size) -lemma uint16_sshiftr_code [code abstract]: +lemma uint16_sshiftr_code [code]: "Rep_uint16 (uint16_sshiftr w n) = - (if n < 0 \ 16 \ n then Rep_uint16 (undefined sshiftr_uint16 w n) + (if n < 0 \ 16 \ n then Rep_uint16 (undefined signed_drop_bit_uint16 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint16 w))" -including undefined_transfer unfolding uint16_sshiftr_def by transfer simp + including undefined_transfer unfolding uint16_sshiftr_def + by transfer simp code_printing constant uint16_sshiftr \ (SML_word) "Uint16.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint16.Int16) _)) :: Uint16.Word16)" and (Scala) "Uint16.shiftr'_signed" lemma uint16_msb_test_bit: "msb x \ bit (x :: uint16) 15" by transfer (simp add: msb_word_iff_bit) lemma msb_uint16_code [code]: "msb x \ uint16_test_bit x 15" by (simp add: uint16_test_bit_def uint16_msb_test_bit) lemma uint16_of_int_code [code]: "uint16_of_int i = Uint16 (integer_of_int i)" including integer.lifting by transfer simp lemma int_of_uint16_code [code]: "int_of_uint16 x = int_of_integer (integer_of_uint16 x)" -by(simp add: integer_of_uint16_def) + by (simp add: int_of_uint16.rep_eq integer_of_uint16_def) lemma nat_of_uint16_code [code]: "nat_of_uint16 x = nat_of_integer (integer_of_uint16 x)" -unfolding integer_of_uint16_def including integer.lifting by transfer simp + 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,767 +1,674 @@ (* Title: Uint32.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 32 bits\ theory Uint32 imports - Code_Target_Word_Base + Code_Target_Word_Base Word_Type_Copies begin section \Type definition and primitive operations\ -typedef uint32 = "UNIV :: 32 word set" .. +typedef uint32 = \UNIV :: 32 word set\ .. + +global_interpretation uint32: word_type_copy Abs_uint32 Rep_uint32 + using type_definition_uint32 by (rule word_type_copy.intro) setup_lifting type_definition_uint32 -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 +declare uint32.of_word_of [code abstype] -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 +declare Quotient_uint32 [transfer_rule] -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] +instantiation uint32 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ 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 zero_uint32 :: uint32 is 0 . +lift_definition one_uint32 :: uint32 is 1 . +lift_definition plus_uint32 :: \uint32 \ uint32 \ uint32\ is \(+)\ . +lift_definition uminus_uint32 :: \uint32 \ uint32\ is uminus . +lift_definition minus_uint32 :: \uint32 \ uint32 \ uint32\ is \(-)\ . +lift_definition times_uint32 :: \uint32 \ uint32 \ uint32\ is \(*)\ . +lift_definition divide_uint32 :: \uint32 \ uint32 \ uint32\ is \(div)\ . +lift_definition modulo_uint32 :: \uint32 \ uint32 \ uint32\ is \(mod)\ . +lift_definition equal_uint32 :: \uint32 \ uint32 \ bool\ is \HOL.equal\ . +lift_definition less_eq_uint32 :: \uint32 \ uint32 \ bool\ is \(\)\ . +lift_definition less_uint32 :: \uint32 \ uint32 \ bool\ is \(<)\ . -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)+ +global_interpretation uint32: word_type_copy_ring Abs_uint32 Rep_uint32 + by standard (fact zero_uint32.rep_eq one_uint32.rep_eq + plus_uint32.rep_eq uminus_uint32.rep_eq minus_uint32.rep_eq + times_uint32.rep_eq divide_uint32.rep_eq modulo_uint32.rep_eq + equal_uint32.rep_eq less_eq_uint32.rep_eq less_uint32.rep_eq)+ -end +instance proof - + show \OFCLASS(uint32, comm_ring_1_class)\ + by (rule uint32.of_class_comm_ring_1) + show \OFCLASS(uint32, semiring_modulo_class)\ + by (fact uint32.of_class_semiring_modulo) + show \OFCLASS(uint32, equal_class)\ + by (fact uint32.of_class_equal) + show \OFCLASS(uint32, linorder_class)\ + by (fact uint32.of_class_linorder) +qed -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 bit_uint32 :: \uint32 \ nat \ bool\ is bit . lift_definition not_uint32 :: \uint32 \ uint32\ is NOT . lift_definition and_uint32 :: \uint32 \ uint32 \ uint32\ is \(AND)\ . lift_definition or_uint32 :: \uint32 \ uint32 \ uint32\ is \(OR)\ . lift_definition xor_uint32 :: \uint32 \ uint32 \ uint32\ is \(XOR)\ . lift_definition mask_uint32 :: \nat \ uint32\ is mask . -lift_definition set_bit_uint32 :: \nat \ uint32 \ uint32\ is \Bit_Operations.set_bit\ . -lift_definition unset_bit_uint32 :: \nat \ uint32 \ uint32\ is \unset_bit\ . -lift_definition flip_bit_uint32 :: \nat \ uint32 \ uint32\ is \flip_bit\ . -instance by (standard; transfer) - (simp_all add: bit_simps mask_eq_decr_exp minus_eq_not_minus_1 set_bit_def flip_bit_def) -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) - -lemma [code]: - \Bit_Operations.set_bit n w = w OR push_bit n 1\ for w :: uint32 - by (fact set_bit_eq_or) - -lemma [code]: - \unset_bit n w = w AND NOT (push_bit n 1)\ for w :: uint32 - by (fact unset_bit_eq_and_not) - -lemma [code]: - \flip_bit n w = w XOR push_bit n 1\ for w :: uint32 - by (fact flip_bit_eq_xor) - -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 - -setup \Context.theory_map (Name_Space.map_naming (Name_Space.qualified_path true \<^binding>\Generic\))\ - -instantiation uint32 :: set_bit -begin -lift_definition set_bit_uint32 :: \uint32 \ nat \ bool \ uint32\ is set_bit . -instance - apply standard - apply transfer - apply (simp add: bit_simps) - done -end - -setup \Context.theory_map (Name_Space.map_naming (Name_Space.parent_path))\ +lift_definition push_bit_uint32 :: \nat \ uint32 \ uint32\ is push_bit . +lift_definition drop_bit_uint32 :: \nat \ uint32 \ uint32\ is drop_bit . +lift_definition signed_drop_bit_uint32 :: \nat \ uint32 \ uint32\ is signed_drop_bit . +lift_definition take_bit_uint32 :: \nat \ uint32 \ uint32\ is take_bit . +lift_definition set_bit_uint32 :: \nat \ uint32 \ uint32\ is Bit_Operations.set_bit . +lift_definition unset_bit_uint32 :: \nat \ uint32 \ uint32\ is unset_bit . +lift_definition flip_bit_uint32 :: \nat \ uint32 \ uint32\ is flip_bit . -instantiation uint32 :: bit_comprehension begin -lift_definition set_bits_uint32 :: "(nat \ bool) \ uint32" is "set_bits" . -instance by (standard; transfer) (fact set_bits_bit_eq) -end - -lemmas [code] = 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 \\w n. signed_drop_bit n w\ . - -lift_definition uint32_of_int :: "int \ uint32" is "word_of_int" . +global_interpretation uint32: word_type_copy_bits Abs_uint32 Rep_uint32 signed_drop_bit_uint32 + by standard (fact bit_uint32.rep_eq not_uint32.rep_eq and_uint32.rep_eq or_uint32.rep_eq xor_uint32.rep_eq + mask_uint32.rep_eq push_bit_uint32.rep_eq drop_bit_uint32.rep_eq signed_drop_bit_uint32.rep_eq take_bit_uint32.rep_eq + set_bit_uint32.rep_eq unset_bit_uint32.rep_eq flip_bit_uint32.rep_eq)+ -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) +instance + by (fact uint32.of_class_ring_bit_operations) 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) +lift_definition uint32_of_nat :: \nat \ uint32\ + is word_of_nat . -lemma Abs_uint32_0 [code_post]: "Abs_uint32 0 = 0" -by(simp add: zero_uint32_def) +lift_definition nat_of_uint32 :: \uint32 \ nat\ + is unat . -lemma Abs_uint32_1 [code_post]: "Abs_uint32 1 = 1" -by(simp add: one_uint32_def) +lift_definition uint32_of_int :: \int \ uint32\ + is word_of_int . + +lift_definition int_of_uint32 :: \uint32 \ int\ + is uint . + +context + includes integer.lifting +begin + +lift_definition Uint32 :: \integer \ uint32\ + is word_of_int . + +lift_definition integer_of_uint32 :: \uint32 \ integer\ + is uint . + +end + +global_interpretation uint32: word_type_copy_more Abs_uint32 Rep_uint32 signed_drop_bit_uint32 + uint32_of_nat nat_of_uint32 uint32_of_int int_of_uint32 Uint32 integer_of_uint32 + apply standard + apply (simp_all add: uint32_of_nat.rep_eq nat_of_uint32.rep_eq + uint32_of_int.rep_eq int_of_uint32.rep_eq + Uint32.rep_eq integer_of_uint32.rep_eq integer_eq_iff) + done + +instantiation uint32 :: "{size, msb, lsb, set_bit, bit_comprehension}" +begin + +lift_definition size_uint32 :: \uint32 \ nat\ is size . + +lift_definition msb_uint32 :: \uint32 \ bool\ is msb . +lift_definition lsb_uint32 :: \uint32 \ bool\ is lsb . + +text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ + +definition set_bit_uint32 :: \uint32 \ nat \ bool \ uint32\ + where set_bit_uint32_eq: \set_bit_uint32 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ + +context + includes lifting_syntax +begin + +lemma set_bit_uint32_transfer [transfer_rule]: + \(cr_uint32 ===> (=) ===> (\) ===> cr_uint32) Generic_set_bit.set_bit Generic_set_bit.set_bit\ + by (simp only: set_bit_eq [abs_def] set_bit_uint32_eq [abs_def]) transfer_prover + +end + +lift_definition set_bits_uint32 :: \(nat \ bool) \ uint32\ is set_bits . +lift_definition set_bits_aux_uint32 :: \(nat \ bool) \ nat \ uint32 \ uint32\ is set_bits_aux . + +global_interpretation uint32: word_type_copy_misc Abs_uint32 Rep_uint32 signed_drop_bit_uint32 + uint32_of_nat nat_of_uint32 uint32_of_int int_of_uint32 Uint32 integer_of_uint32 32 set_bits_aux_uint32 + by (standard; transfer) simp_all + +instance using uint32.of_class_bit_comprehension + uint32.of_class_set_bit uint32.of_class_lsb + by simp_all standard + +end section \Code setup\ code_printing code_module Uint32 \ (SML) \(* Test that words can handle numbers between 0 and 31 *) val _ = if 5 <= Word.wordSize then () else raise (Fail ("wordSize less than 5")); structure Uint32 : sig val set_bit : Word32.word -> IntInf.int -> bool -> Word32.word val shiftl : Word32.word -> IntInf.int -> Word32.word val shiftr : Word32.word -> IntInf.int -> Word32.word val shiftr_signed : Word32.word -> IntInf.int -> Word32.word val test_bit : Word32.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word32.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word32.orb (x, mask) else Word32.andb (x, Word32.notb mask) end fun shiftl x n = Word32.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word32.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word32.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word32.andb (x, Word32.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word32.fromInt 0 end; (* struct Uint32 *)\ code_reserved SML Uint32 code_printing code_module Uint32 \ (Haskell) \module Uint32(Int32, Word32) where import Data.Int(Int32) import Data.Word(Word32)\ code_reserved Haskell Uint32 text \ OCaml and Scala provide only signed 32bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module "Uint32" \ (OCaml) \module Uint32 : sig val less : int32 -> int32 -> bool val less_eq : int32 -> int32 -> bool val set_bit : int32 -> Z.t -> bool -> int32 val shiftl : int32 -> Z.t -> int32 val shiftr : int32 -> Z.t -> int32 val shiftr_signed : int32 -> Z.t -> int32 val test_bit : int32 -> Z.t -> bool end = struct (* negative numbers have their highest bit set, so they are greater than positive ones *) let less x y = if Int32.compare x Int32.zero < 0 then Int32.compare y Int32.zero < 0 && Int32.compare x y < 0 else Int32.compare y Int32.zero < 0 || Int32.compare x y < 0;; let less_eq x y = if Int32.compare x Int32.zero < 0 then Int32.compare y Int32.zero < 0 && Int32.compare x y <= 0 else Int32.compare y Int32.zero < 0 || Int32.compare x y <= 0;; let set_bit x n b = let mask = Int32.shift_left Int32.one (Z.to_int n) in if b then Int32.logor x mask else Int32.logand x (Int32.lognot mask);; let shiftl x n = Int32.shift_left x (Z.to_int n);; let shiftr x n = Int32.shift_right_logical x (Z.to_int n);; let shiftr_signed x n = Int32.shift_right x (Z.to_int n);; let test_bit x n = Int32.compare (Int32.logand x (Int32.shift_left Int32.one (Z.to_int n))) Int32.zero <> 0;; end;; (*struct Uint32*)\ code_reserved OCaml Uint32 code_printing code_module Uint32 \ (Scala) \object Uint32 { def less(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Int, n: BigInt, b: Boolean) : Int = if (b) x | (1 << n.intValue) else x & (1 << n.intValue).unary_~ def shiftl(x: Int, n: BigInt) : Int = x << n.intValue def shiftr(x: Int, n: BigInt) : Int = x >>> n.intValue def shiftr_signed(x: Int, n: BigInt) : Int = x >> n.intValue def test_bit(x: Int, n: BigInt) : Boolean = (x & (1 << n.intValue)) != 0 } /* object Uint32 */\ code_reserved Scala Uint32 text \ OCaml's conversion from Big\_int to int32 demands that the value fits int a signed 32-bit integer. The following justifies the implementation. \ definition Uint32_signed :: "integer \ uint32" where "Uint32_signed i = (if i < -(0x80000000) \ i \ 0x80000000 then undefined Uint32 i else Uint32 i)" lemma Uint32_code [code]: "Uint32 i = (let i' = i AND 0xFFFFFFFF in if bit i' 31 then Uint32_signed (i' - 0x100000000) else Uint32_signed i')" including undefined_transfer integer.lifting unfolding Uint32_signed_def apply transfer apply (subst word_of_int_via_signed) apply (auto simp add: push_bit_of_1 mask_eq_exp_minus_1 word_of_int_via_signed cong del: if_cong) done -lemma Uint32_signed_code [code abstract]: +lemma Uint32_signed_code [code]: "Rep_uint32 (Uint32_signed i) = (if i < -(0x80000000) \ i \ 0x80000000 then Rep_uint32 (undefined Uint32 i) else word_of_int (int_of_integer_symbolic i))" unfolding Uint32_signed_def Uint32_def int_of_integer_symbolic_def word_of_integer_def by(simp add: Abs_uint32_inverse) text \ Avoid @{term Abs_uint32} in generated code, use @{term Rep_uint32'} instead. The symbolic implementations for code\_simp use @{term Rep_uint32}. The new destructor @{term Rep_uint32'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint32} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint32} ([code abstract] equations for @{typ uint32} may use @{term Rep_uint32} because these instances will be folded away.) To convert @{typ "32 word"} values into @{typ uint32}, use @{term "Abs_uint32'"}. \ definition Rep_uint32' where [simp]: "Rep_uint32' = Rep_uint32" lemma Rep_uint32'_transfer [transfer_rule]: "rel_fun cr_uint32 (=) (\x. x) Rep_uint32'" unfolding Rep_uint32'_def by(rule uint32.rep_transfer) lemma Rep_uint32'_code [code]: "Rep_uint32' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint32' :: "32 word \ uint32" is "\x :: 32 word. x" . lemma Abs_uint32'_code [code]: "Abs_uint32' x = Uint32 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint32 \ _"]] lemma term_of_uint32_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint32.uint32.Abs_uint32'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]]]], TR (STR ''Uint32.uint32'') []])) (term_of_class.term_of (Rep_uint32' x))" by(simp add: term_of_anything) code_printing type_constructor uint32 \ (SML) "Word32.word" and (Haskell) "Uint32.Word32" and (OCaml) "int32" and (Scala) "Int" and (Eval) "Word32.word" | constant Uint32 \ (SML) "Word32.fromLargeInt (IntInf.toLarge _)" and (Haskell) "(Prelude.fromInteger _ :: Uint32.Word32)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint32.Word32)" and (Scala) "_.intValue" | constant Uint32_signed \ (OCaml) "Z.to'_int32" | constant "0 :: uint32" \ (SML) "(Word32.fromInt 0)" and (Haskell) "(0 :: Uint32.Word32)" and (OCaml) "Int32.zero" and (Scala) "0" | constant "1 :: uint32" \ (SML) "(Word32.fromInt 1)" and (Haskell) "(1 :: Uint32.Word32)" and (OCaml) "Int32.one" and (Scala) "1" | constant "plus :: uint32 \ _ " \ (SML) "Word32.+ ((_), (_))" and (Haskell) infixl 6 "+" and (OCaml) "Int32.add" and (Scala) infixl 7 "+" | constant "uminus :: uint32 \ _" \ (SML) "Word32.~" and (Haskell) "negate" and (OCaml) "Int32.neg" and (Scala) "!(- _)" | constant "minus :: uint32 \ _" \ (SML) "Word32.- ((_), (_))" and (Haskell) infixl 6 "-" and (OCaml) "Int32.sub" and (Scala) infixl 7 "-" | constant "times :: uint32 \ _ \ _" \ (SML) "Word32.* ((_), (_))" and (Haskell) infixl 7 "*" and (OCaml) "Int32.mul" and (Scala) infixl 8 "*" | constant "HOL.equal :: uint32 \ _ \ bool" \ (SML) "!((_ : Word32.word) = _)" and (Haskell) infix 4 "==" and (OCaml) "(Int32.compare _ _ = 0)" and (Scala) infixl 5 "==" | class_instance uint32 :: equal \ (Haskell) - | constant "less_eq :: uint32 \ _ \ bool" \ (SML) "Word32.<= ((_), (_))" and (Haskell) infix 4 "<=" and (OCaml) "Uint32.less'_eq" and (Scala) "Uint32.less'_eq" | constant "less :: uint32 \ _ \ bool" \ (SML) "Word32.< ((_), (_))" and (Haskell) infix 4 "<" and (OCaml) "Uint32.less" and (Scala) "Uint32.less" | constant "NOT :: uint32 \ _" \ (SML) "Word32.notb" and (Haskell) "Data'_Bits.complement" and (OCaml) "Int32.lognot" and (Scala) "_.unary'_~" | constant "(AND) :: uint32 \ _" \ (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 = push_bit 1 (uint32_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" including undefined_transfer unfolding uint32_divmod_def uint32_sdiv_def div0_uint32_def mod0_uint32_def - by transfer (simp add: divmod_via_sdivmod ac_simps) + less_eq_uint32.rep_eq + apply transfer + apply (simp add: divmod_via_sdivmod push_bit_eq_mult) + done -lemma uint32_sdiv_code [code abstract]: +lemma uint32_sdiv_code [code]: "Rep_uint32 (uint32_sdiv x y) = (if y = 0 then Rep_uint32 (undefined ((div) :: uint32 \ _) x (0 :: uint32)) else Rep_uint32 x sdiv Rep_uint32 y)" unfolding uint32_sdiv_def by(simp add: Abs_uint32_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint32_divmod_code} computes both with division only. \ code_printing constant uint32_div \ (SML) "Word32.div ((_), (_))" and (Haskell) "Prelude.div" | constant uint32_mod \ (SML) "Word32.mod ((_), (_))" and (Haskell) "Prelude.mod" | constant uint32_divmod \ (Haskell) "divmod" | constant uint32_sdiv \ (OCaml) "Int32.div" and (Scala) "_ '/ _" definition uint32_test_bit :: "uint32 \ integer \ bool" where [code del]: "uint32_test_bit x n = (if n < 0 \ 31 < n then undefined (bit :: uint32 \ _) x n else bit x (nat_of_integer n))" lemma test_bit_uint32_code [code]: "bit x n \ n < 32 \ uint32_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint32_test_bit_def by (transfer, simp, transfer, simp) lemma uint32_test_bit_code [code]: "uint32_test_bit w n = (if n < 0 \ 31 < n then undefined (bit :: uint32 \ _) w n else bit (Rep_uint32 w) (nat_of_integer n))" unfolding uint32_test_bit_def by(simp add: bit_uint32.rep_eq) code_printing constant uint32_test_bit \ (SML) "Uint32.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (OCaml) "Uint32.test'_bit" and (Scala) "Uint32.test'_bit" and (Eval) "(fn w => fn n => if n < 0 orelse 32 <= n then raise (Fail \"argument to uint32'_test'_bit out of bounds\") else Uint32.test'_bit w n)" definition uint32_set_bit :: "uint32 \ integer \ bool \ uint32" where [code del]: "uint32_set_bit x n b = (if n < 0 \ 31 < n then undefined (set_bit :: uint32 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint32_code [code]: "set_bit x n b = (if n < 32 then uint32_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint32_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) -lemma uint32_set_bit_code [code abstract]: +lemma uint32_set_bit_code [code]: "Rep_uint32 (uint32_set_bit w n b) = (if n < 0 \ 31 < n then Rep_uint32 (undefined (set_bit :: uint32 \ _) w n b) else set_bit (Rep_uint32 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint32_set_bit_def by transfer simp code_printing constant uint32_set_bit \ (SML) "Uint32.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (OCaml) "Uint32.set'_bit" and (Scala) "Uint32.set'_bit" and (Eval) "(fn w => fn n => fn b => if n < 0 orelse 32 <= n then raise (Fail \"argument to uint32'_set'_bit out of bounds\") else Uint32.set'_bit x n b)" -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 (push_bit 1 w OR (if f n' then 1 else 0)) n')" - apply (transfer fixing: n) - apply (cases n) - apply simp_all - done - -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 \ bit x 0" - by transfer (simp add: lsb_word_eq) - - definition uint32_shiftl :: "uint32 \ integer \ uint32" where [code del]: "uint32_shiftl x n = (if n < 0 \ 32 \ n then undefined (push_bit :: nat \ uint32 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint32_code [code]: "push_bit n x = (if n < 32 then uint32_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint32_shiftl_def by transfer simp -lemma uint32_shiftl_code [code abstract]: +lemma uint32_shiftl_code [code]: "Rep_uint32 (uint32_shiftl w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined (push_bit :: nat \ uint32 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint32 w))" including undefined_transfer unfolding uint32_shiftl_def by transfer simp code_printing constant uint32_shiftl \ (SML) "Uint32.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (OCaml) "Uint32.shiftl" and (Scala) "Uint32.shiftl" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 32 then raise Fail \"argument to uint32'_shiftl out of bounds\" else Uint32.shiftl x i)" definition uint32_shiftr :: "uint32 \ integer \ uint32" where [code del]: "uint32_shiftr x n = (if n < 0 \ 32 \ n then undefined (drop_bit :: nat \ uint32 \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint32_code [code]: "drop_bit n x = (if n < 32 then uint32_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint32_shiftr_def by transfer simp -lemma uint32_shiftr_code [code abstract]: +lemma uint32_shiftr_code [code]: "Rep_uint32 (uint32_shiftr w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined (drop_bit :: nat \ uint32 \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint32 w))" including undefined_transfer unfolding uint32_shiftr_def by transfer simp code_printing constant uint32_shiftr \ (SML) "Uint32.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (OCaml) "Uint32.shiftr" and (Scala) "Uint32.shiftr" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 32 then raise Fail \"argument to uint32'_shiftr out of bounds\" else Uint32.shiftr x i)" definition uint32_sshiftr :: "uint32 \ integer \ uint32" where [code del]: "uint32_sshiftr x n = - (if n < 0 \ 32 \ n then undefined sshiftr_uint32 x n else sshiftr_uint32 x (nat_of_integer n))" + (if n < 0 \ 32 \ n then undefined signed_drop_bit_uint32 n x else signed_drop_bit_uint32 (nat_of_integer n) x)" lemma sshiftr_uint32_code [code]: - "x >>> n = + "signed_drop_bit_uint32 n x = (if n < 32 then uint32_sshiftr x (integer_of_nat n) else if bit x 31 then - 1 else 0)" including undefined_transfer integer.lifting unfolding uint32_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond) -lemma uint32_sshiftr_code [code abstract]: +lemma uint32_sshiftr_code [code]: "Rep_uint32 (uint32_sshiftr w n) = - (if n < 0 \ 32 \ n then Rep_uint32 (undefined sshiftr_uint32 w n) else signed_drop_bit (nat_of_integer n) (Rep_uint32 w))" + (if n < 0 \ 32 \ n then Rep_uint32 (undefined signed_drop_bit_uint32 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint32 w))" including undefined_transfer unfolding uint32_sshiftr_def by transfer simp code_printing constant uint32_sshiftr \ (SML) "Uint32.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint32.Int32) _)) :: Uint32.Word32)" and (OCaml) "Uint32.shiftr'_signed" and (Scala) "Uint32.shiftr'_signed" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 32 then raise Fail \"argument to uint32'_shiftr'_signed out of bounds\" else Uint32.shiftr'_signed x i)" lemma uint32_msb_test_bit: "msb x \ bit (x :: uint32) 31" by transfer (simp add: msb_word_iff_bit) lemma msb_uint32_code [code]: "msb x \ uint32_test_bit x 31" by (simp add: uint32_test_bit_def uint32_msb_test_bit) lemma uint32_of_int_code [code]: "uint32_of_int i = Uint32 (integer_of_int i)" including integer.lifting by transfer simp -lemma int_of_uint32_code [code]: - "int_of_uint32 x = int_of_integer (integer_of_uint32 x)" -by(simp add: integer_of_uint32_def) +lemma int_of_uint32_code [code]: "int_of_uint32 x = int_of_integer (integer_of_uint32 x)" +including integer.lifting by transfer simp lemma nat_of_uint32_code [code]: "nat_of_uint32 x = nat_of_integer (integer_of_uint32 x)" unfolding integer_of_uint32_def including integer.lifting by transfer simp + definition integer_of_uint32_signed :: "uint32 \ integer" where "integer_of_uint32_signed n = (if bit n 31 then undefined integer_of_uint32 n else integer_of_uint32 n)" lemma integer_of_uint32_signed_code [code]: "integer_of_uint32_signed n = (if bit n 31 then undefined integer_of_uint32 n else integer_of_int (uint (Rep_uint32' n)))" -unfolding integer_of_uint32_signed_def integer_of_uint32_def -including undefined_transfer by transfer simp + by (simp add: integer_of_uint32_signed_def integer_of_uint32_def) lemma integer_of_uint32_code [code]: "integer_of_uint32 n = (if bit n 31 then integer_of_uint32_signed (n AND 0x7FFFFFFF) OR 0x80000000 else integer_of_uint32_signed n)" proof - - have \(0x7FFFFFFF :: uint32) = mask 31\ - by (simp add: mask_eq_exp_minus_1) - then have *: \n AND 0x7FFFFFFF = take_bit 31 n\ - by (simp add: take_bit_eq_mask) - have **: \(0x80000000 :: int) = 2 ^ 31\ - by simp - show ?thesis - unfolding integer_of_uint32_def integer_of_uint32_signed_def o_def * - including undefined_transfer integer.lifting - apply transfer - apply (rule bit_eqI) - apply (simp add: bit_or_iff bit_take_bit_iff bit_uint_iff) - apply (simp only: bit_exp_iff bit_or_iff **) - apply auto - done + have \integer_of_uint32_signed (n AND 0x7FFFFFFF) OR 0x80000000 = Bit_Operations.set_bit 31 (integer_of_uint32_signed (take_bit 31 n))\ + by (simp add: take_bit_eq_mask set_bit_eq_or push_bit_eq_mult mask_eq_exp_minus_1) + moreover have \integer_of_uint32 n = Bit_Operations.set_bit 31 (integer_of_uint32 (take_bit 31 n))\ if \bit n 31\ + proof (rule bit_eqI) + fix m + from that show \bit (integer_of_uint32 n) m = bit (Bit_Operations.set_bit 31 (integer_of_uint32 (take_bit 31 n))) m\ for m + including integer.lifting by transfer (auto simp add: bit_simps dest: bit_imp_le_length) + qed + ultimately show ?thesis + by simp (simp add: integer_of_uint32_signed_def bit_simps) qed 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,966 +1,871 @@ (* Title: Uint64.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 64 bits\ theory Uint64 imports - Code_Target_Word_Base + Code_Target_Word_Base Word_Type_Copies begin text \ PolyML (in version 5.7) provides a Word64 structure only when run in 64-bit mode. Therefore, we by default provide an implementation of 64-bit words using \verb$IntInf.int$ and masking. The code target \texttt{SML\_word} replaces this implementation and maps the operations directly to the \verb$Word64$ structure provided by the Standard ML implementations. The \verb$Eval$ target used by @{command value} and @{method eval} dynamically tests at runtime for the version of PolyML and uses PolyML's Word64 structure if it detects a 64-bit version which does not suffer from a division bug found in PolyML 5.6. \ section \Type definition and primitive operations\ -typedef uint64 = "UNIV :: 64 word set" .. +typedef uint64 = \UNIV :: 64 word set\ .. + +global_interpretation uint64: word_type_copy Abs_uint64 Rep_uint64 + using type_definition_uint64 by (rule word_type_copy.intro) setup_lifting type_definition_uint64 -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 +declare uint64.of_word_of [code abstype] -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 +declare Quotient_uint64 [transfer_rule] -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] +instantiation uint64 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ 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 zero_uint64 :: uint64 is 0 . +lift_definition one_uint64 :: uint64 is 1 . +lift_definition plus_uint64 :: \uint64 \ uint64 \ uint64\ is \(+)\ . +lift_definition uminus_uint64 :: \uint64 \ uint64\ is uminus . +lift_definition minus_uint64 :: \uint64 \ uint64 \ uint64\ is \(-)\ . +lift_definition times_uint64 :: \uint64 \ uint64 \ uint64\ is \(*)\ . +lift_definition divide_uint64 :: \uint64 \ uint64 \ uint64\ is \(div)\ . +lift_definition modulo_uint64 :: \uint64 \ uint64 \ uint64\ is \(mod)\ . +lift_definition equal_uint64 :: \uint64 \ uint64 \ bool\ is \HOL.equal\ . +lift_definition less_eq_uint64 :: \uint64 \ uint64 \ bool\ is \(\)\ . +lift_definition less_uint64 :: \uint64 \ uint64 \ bool\ is \(<)\ . -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)+ +global_interpretation uint64: word_type_copy_ring Abs_uint64 Rep_uint64 + by standard (fact zero_uint64.rep_eq one_uint64.rep_eq + plus_uint64.rep_eq uminus_uint64.rep_eq minus_uint64.rep_eq + times_uint64.rep_eq divide_uint64.rep_eq modulo_uint64.rep_eq + equal_uint64.rep_eq less_eq_uint64.rep_eq less_uint64.rep_eq)+ -end +instance proof - + show \OFCLASS(uint64, comm_ring_1_class)\ + by (rule uint64.of_class_comm_ring_1) + show \OFCLASS(uint64, semiring_modulo_class)\ + by (fact uint64.of_class_semiring_modulo) + show \OFCLASS(uint64, equal_class)\ + by (fact uint64.of_class_equal) + show \OFCLASS(uint64, linorder_class)\ + by (fact uint64.of_class_linorder) +qed -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 bit_uint64 :: \uint64 \ nat \ bool\ is bit . lift_definition not_uint64 :: \uint64 \ uint64\ is NOT . lift_definition and_uint64 :: \uint64 \ uint64 \ uint64\ is \(AND)\ . lift_definition or_uint64 :: \uint64 \ uint64 \ uint64\ is \(OR)\ . lift_definition xor_uint64 :: \uint64 \ uint64 \ uint64\ is \(XOR)\ . lift_definition mask_uint64 :: \nat \ uint64\ is mask . -lift_definition set_bit_uint64 :: \nat \ uint64 \ uint64\ is \Bit_Operations.set_bit\ . -lift_definition unset_bit_uint64 :: \nat \ uint64 \ uint64\ is \unset_bit\ . -lift_definition flip_bit_uint64 :: \nat \ uint64 \ uint64\ is \flip_bit\ . -instance by (standard; transfer) - (simp_all add: bit_simps mask_eq_decr_exp minus_eq_not_minus_1 set_bit_def flip_bit_def) -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) - -lemma [code]: - \Bit_Operations.set_bit n w = w OR push_bit n 1\ for w :: uint64 - by (fact set_bit_eq_or) - -lemma [code]: - \unset_bit n w = w AND NOT (push_bit n 1)\ for w :: uint64 - by (fact unset_bit_eq_and_not) - -lemma [code]: - \flip_bit n w = w XOR push_bit n 1\ for w :: uint64 - by (fact flip_bit_eq_xor) - -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 - -setup \Context.theory_map (Name_Space.map_naming (Name_Space.qualified_path true \<^binding>\Generic\))\ - -instantiation uint64 :: set_bit -begin -lift_definition set_bit_uint64 :: \uint64 \ nat \ bool \ uint64\ is set_bit . -instance - apply standard - apply transfer - apply (simp add: bit_simps) - done -end - -setup \Context.theory_map (Name_Space.map_naming (Name_Space.parent_path))\ +lift_definition push_bit_uint64 :: \nat \ uint64 \ uint64\ is push_bit . +lift_definition drop_bit_uint64 :: \nat \ uint64 \ uint64\ is drop_bit . +lift_definition signed_drop_bit_uint64 :: \nat \ uint64 \ uint64\ is signed_drop_bit . +lift_definition take_bit_uint64 :: \nat \ uint64 \ uint64\ is take_bit . +lift_definition set_bit_uint64 :: \nat \ uint64 \ uint64\ is Bit_Operations.set_bit . +lift_definition unset_bit_uint64 :: \nat \ uint64 \ uint64\ is unset_bit . +lift_definition flip_bit_uint64 :: \nat \ uint64 \ uint64\ is flip_bit . -instantiation uint64 :: bit_comprehension begin -lift_definition set_bits_uint64 :: "(nat \ bool) \ uint64" is "set_bits" . -instance by (standard; transfer) (fact set_bits_bit_eq) -end - -lemmas [code] = 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 \\w n. signed_drop_bit n w\ . - -lift_definition uint64_of_int :: "int \ uint64" is "word_of_int" . +global_interpretation uint64: word_type_copy_bits Abs_uint64 Rep_uint64 signed_drop_bit_uint64 + by standard (fact bit_uint64.rep_eq not_uint64.rep_eq and_uint64.rep_eq or_uint64.rep_eq xor_uint64.rep_eq + mask_uint64.rep_eq push_bit_uint64.rep_eq drop_bit_uint64.rep_eq signed_drop_bit_uint64.rep_eq take_bit_uint64.rep_eq + set_bit_uint64.rep_eq unset_bit_uint64.rep_eq flip_bit_uint64.rep_eq)+ -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) +instance + by (fact uint64.of_class_ring_bit_operations) 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) +lift_definition uint64_of_nat :: \nat \ uint64\ + is word_of_nat . -lemma Abs_uint64_0 [code_post]: "Abs_uint64 0 = 0" -by(simp add: zero_uint64_def) +lift_definition nat_of_uint64 :: \uint64 \ nat\ + is unat . -lemma Abs_uint64_1 [code_post]: "Abs_uint64 1 = 1" -by(simp add: one_uint64_def) +lift_definition uint64_of_int :: \int \ uint64\ + is word_of_int . + +lift_definition int_of_uint64 :: \uint64 \ int\ + is uint . + +context + includes integer.lifting +begin + +lift_definition Uint64 :: \integer \ uint64\ + is word_of_int . + +lift_definition integer_of_uint64 :: \uint64 \ integer\ + is uint . + +end + +global_interpretation uint64: word_type_copy_more Abs_uint64 Rep_uint64 signed_drop_bit_uint64 + uint64_of_nat nat_of_uint64 uint64_of_int int_of_uint64 Uint64 integer_of_uint64 + apply standard + apply (simp_all add: uint64_of_nat.rep_eq nat_of_uint64.rep_eq + uint64_of_int.rep_eq int_of_uint64.rep_eq + Uint64.rep_eq integer_of_uint64.rep_eq integer_eq_iff) + done + +instantiation uint64 :: "{size, msb, lsb, set_bit, bit_comprehension}" +begin + +lift_definition size_uint64 :: \uint64 \ nat\ is size . + +lift_definition msb_uint64 :: \uint64 \ bool\ is msb . +lift_definition lsb_uint64 :: \uint64 \ bool\ is lsb . + +text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ + +definition set_bit_uint64 :: \uint64 \ nat \ bool \ uint64\ + where set_bit_uint64_eq: \set_bit_uint64 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ + +context + includes lifting_syntax +begin + +lemma set_bit_uint64_transfer [transfer_rule]: + \(cr_uint64 ===> (=) ===> (\) ===> cr_uint64) Generic_set_bit.set_bit Generic_set_bit.set_bit\ + by (simp only: set_bit_eq [abs_def] set_bit_uint64_eq [abs_def]) transfer_prover + +end + +lift_definition set_bits_uint64 :: \(nat \ bool) \ uint64\ is set_bits . +lift_definition set_bits_aux_uint64 :: \(nat \ bool) \ nat \ uint64 \ uint64\ is set_bits_aux . + +global_interpretation uint64: word_type_copy_misc Abs_uint64 Rep_uint64 signed_drop_bit_uint64 + uint64_of_nat nat_of_uint64 uint64_of_int int_of_uint64 Uint64 integer_of_uint64 64 set_bits_aux_uint64 + by (standard; transfer) simp_all + +instance using uint64.of_class_bit_comprehension + uint64.of_class_set_bit uint64.of_class_lsb + by simp_all standard + +end section \Code setup\ text \ For SML, we generate an implementation of unsigned 64-bit words using \verb$IntInf.int$. If @{ML "LargeWord.wordSize > 63"} of the Isabelle/ML runtime environment holds, then we assume that there is also a \Word64\ structure available and accordingly replace the implementation for the target \verb$Eval$. \ code_printing code_module "Uint64" \ (SML) \(* Test that words can handle numbers between 0 and 63 *) val _ = if 6 <= Word.wordSize then () else raise (Fail ("wordSize less than 6")); structure Uint64 : sig eqtype uint64; val zero : uint64; val one : uint64; val fromInt : IntInf.int -> uint64; val toInt : uint64 -> IntInf.int; val toLarge : uint64 -> LargeWord.word; val fromLarge : LargeWord.word -> uint64 val plus : uint64 -> uint64 -> uint64; val minus : uint64 -> uint64 -> uint64; val times : uint64 -> uint64 -> uint64; val divide : uint64 -> uint64 -> uint64; val modulus : uint64 -> uint64 -> uint64; val negate : uint64 -> uint64; val less_eq : uint64 -> uint64 -> bool; val less : uint64 -> uint64 -> bool; val notb : uint64 -> uint64; val andb : uint64 -> uint64 -> uint64; val orb : uint64 -> uint64 -> uint64; val xorb : uint64 -> uint64 -> uint64; val shiftl : uint64 -> IntInf.int -> uint64; val shiftr : uint64 -> IntInf.int -> uint64; val shiftr_signed : uint64 -> IntInf.int -> uint64; val set_bit : uint64 -> IntInf.int -> bool -> uint64; val test_bit : uint64 -> IntInf.int -> bool; end = struct type uint64 = IntInf.int; val mask = 0xFFFFFFFFFFFFFFFF : IntInf.int; val zero = 0 : IntInf.int; val one = 1 : IntInf.int; fun fromInt x = IntInf.andb(x, mask); fun toInt x = x fun toLarge x = LargeWord.fromLargeInt (IntInf.toLarge x); fun fromLarge x = IntInf.fromLarge (LargeWord.toLargeInt x); fun plus x y = IntInf.andb(IntInf.+(x, y), mask); fun minus x y = IntInf.andb(IntInf.-(x, y), mask); fun negate x = IntInf.andb(IntInf.~(x), mask); fun times x y = IntInf.andb(IntInf.*(x, y), mask); fun divide x y = IntInf.div(x, y); fun modulus x y = IntInf.mod(x, y); fun less_eq x y = IntInf.<=(x, y); fun less x y = IntInf.<(x, y); fun notb x = IntInf.andb(IntInf.notb(x), mask); fun orb x y = IntInf.orb(x, y); fun andb x y = IntInf.andb(x, y); fun xorb x y = IntInf.xorb(x, y); val maxWord = IntInf.pow (2, Word.wordSize); fun shiftl x n = if n < maxWord then IntInf.andb(IntInf.<< (x, Word.fromLargeInt (IntInf.toLarge n)), mask) else 0; fun shiftr x n = if n < maxWord then IntInf.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) else 0; val msb_mask = 0x8000000000000000 : IntInf.int; fun shiftr_signed x i = if IntInf.andb(x, msb_mask) = 0 then shiftr x i else if i >= 64 then 0xFFFFFFFFFFFFFFFF else let val x' = shiftr x i val m' = IntInf.andb(IntInf.<<(mask, Word.max(0w64 - Word.fromLargeInt (IntInf.toLarge i), 0w0)), mask) in IntInf.orb(x', m') end; fun test_bit x n = if n < maxWord then IntInf.andb (x, IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n))) <> 0 else false; fun set_bit x n b = if n < 64 then if b then IntInf.orb (x, IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n))) else IntInf.andb (x, IntInf.notb (IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n)))) else x; end \ code_reserved SML Uint64 setup \ let val polyml64 = LargeWord.wordSize > 63; (* PolyML 5.6 has bugs in its Word64 implementation. We test for one such bug and refrain from using Word64 in that case. Testing is done with dynamic code evaluation such that the compiler does not choke on the Word64 structure, which need not be present in a 32bit environment. *) val error_msg = "Buggy Word64 structure"; val test_code = "val _ = if Word64.div (0w18446744073709551611 : Word64.word, 0w3) = 0w6148914691236517203 then ()\n" ^ "else raise (Fail \"" ^ error_msg ^ "\");"; val f = Exn.interruptible_capture (fn () => ML_Compiler.eval ML_Compiler.flags Position.none (ML_Lex.tokenize test_code)) val use_Word64 = polyml64 andalso (case f () of Exn.Res _ => true | Exn.Exn (e as ERROR m) => if String.isSuffix error_msg m then false else Exn.reraise e | Exn.Exn e => Exn.reraise e) ; val newline = "\n"; val content = "structure Uint64 : sig" ^ newline ^ " eqtype uint64;" ^ newline ^ " val zero : uint64;" ^ newline ^ " val one : uint64;" ^ newline ^ " val fromInt : IntInf.int -> uint64;" ^ newline ^ " val toInt : uint64 -> IntInf.int;" ^ newline ^ " val toLarge : uint64 -> LargeWord.word;" ^ newline ^ " val fromLarge : LargeWord.word -> uint64" ^ newline ^ " val plus : uint64 -> uint64 -> uint64;" ^ newline ^ " val minus : uint64 -> uint64 -> uint64;" ^ newline ^ " val times : uint64 -> uint64 -> uint64;" ^ newline ^ " val divide : uint64 -> uint64 -> uint64;" ^ newline ^ " val modulus : uint64 -> uint64 -> uint64;" ^ newline ^ " val negate : uint64 -> uint64;" ^ newline ^ " val less_eq : uint64 -> uint64 -> bool;" ^ newline ^ " val less : uint64 -> uint64 -> bool;" ^ newline ^ " val notb : uint64 -> uint64;" ^ newline ^ " val andb : uint64 -> uint64 -> uint64;" ^ newline ^ " val orb : uint64 -> uint64 -> uint64;" ^ newline ^ " val xorb : uint64 -> uint64 -> uint64;" ^ newline ^ " val shiftl : uint64 -> IntInf.int -> uint64;" ^ newline ^ " val shiftr : uint64 -> IntInf.int -> uint64;" ^ newline ^ " val shiftr_signed : uint64 -> IntInf.int -> uint64;" ^ newline ^ " val set_bit : uint64 -> IntInf.int -> bool -> uint64;" ^ newline ^ " val test_bit : uint64 -> IntInf.int -> bool;" ^ newline ^ "end = struct" ^ newline ^ "" ^ newline ^ "type uint64 = Word64.word;" ^ newline ^ "" ^ newline ^ "val zero = (0wx0 : uint64);" ^ newline ^ "" ^ newline ^ "val one = (0wx1 : uint64);" ^ newline ^ "" ^ newline ^ "fun fromInt x = Word64.fromLargeInt (IntInf.toLarge x);" ^ newline ^ "" ^ newline ^ "fun toInt x = IntInf.fromLarge (Word64.toLargeInt x);" ^ newline ^ "" ^ newline ^ "fun fromLarge x = Word64.fromLarge x;" ^ newline ^ "" ^ newline ^ "fun toLarge x = Word64.toLarge x;" ^ newline ^ "" ^ newline ^ "fun plus x y = Word64.+(x, y);" ^ newline ^ "" ^ newline ^ "fun minus x y = Word64.-(x, y);" ^ newline ^ "" ^ newline ^ "fun negate x = Word64.~(x);" ^ newline ^ "" ^ newline ^ "fun times x y = Word64.*(x, y);" ^ newline ^ "" ^ newline ^ "fun divide x y = Word64.div(x, y);" ^ newline ^ "" ^ newline ^ "fun modulus x y = Word64.mod(x, y);" ^ newline ^ "" ^ newline ^ "fun less_eq x y = Word64.<=(x, y);" ^ newline ^ "" ^ newline ^ "fun less x y = Word64.<(x, y);" ^ newline ^ "" ^ newline ^ "fun set_bit x n b =" ^ newline ^ " let val mask = Word64.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ " in if b then Word64.orb (x, mask)" ^ newline ^ " else Word64.andb (x, Word64.notb mask)" ^ newline ^ " end" ^ newline ^ "" ^ newline ^ "fun shiftl x n =" ^ newline ^ " Word64.<< (x, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ "" ^ newline ^ "fun shiftr x n =" ^ newline ^ " Word64.>> (x, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ "" ^ newline ^ "fun shiftr_signed x n =" ^ newline ^ " Word64.~>> (x, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ "" ^ newline ^ "fun test_bit x n =" ^ newline ^ " Word64.andb (x, Word64.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word64.fromInt 0" ^ newline ^ "" ^ newline ^ "val notb = Word64.notb" ^ newline ^ "" ^ newline ^ "fun andb x y = Word64.andb(x, y);" ^ newline ^ "" ^ newline ^ "fun orb x y = Word64.orb(x, y);" ^ newline ^ "" ^ newline ^ "fun xorb x y = Word64.xorb(x, y);" ^ newline ^ "" ^ newline ^ "end (*struct Uint64*)" val target_SML64 = "SML_word"; in (if use_Word64 then Code_Target.set_printings (Code_Symbol.Module ("Uint64", [(Code_Runtime.target, SOME (content, []))])) else I) #> Code_Target.set_printings (Code_Symbol.Module ("Uint64", [(target_SML64, SOME (content, []))])) end \ code_printing code_module Uint64 \ (Haskell) \module Uint64(Int64, Word64) where import Data.Int(Int64) import Data.Word(Word64)\ code_reserved Haskell Uint64 text \ OCaml and Scala provide only signed 64bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module "Uint64" \ (OCaml) \module Uint64 : sig val less : int64 -> int64 -> bool val less_eq : int64 -> int64 -> bool val set_bit : int64 -> Z.t -> bool -> int64 val shiftl : int64 -> Z.t -> int64 val shiftr : int64 -> Z.t -> int64 val shiftr_signed : int64 -> Z.t -> int64 val test_bit : int64 -> Z.t -> bool end = struct (* negative numbers have their highest bit set, so they are greater than positive ones *) let less x y = if Int64.compare x Int64.zero < 0 then Int64.compare y Int64.zero < 0 && Int64.compare x y < 0 else Int64.compare y Int64.zero < 0 || Int64.compare x y < 0;; let less_eq x y = if Int64.compare x Int64.zero < 0 then Int64.compare y Int64.zero < 0 && Int64.compare x y <= 0 else Int64.compare y Int64.zero < 0 || Int64.compare x y <= 0;; let set_bit x n b = let mask = Int64.shift_left Int64.one (Z.to_int n) in if b then Int64.logor x mask else Int64.logand x (Int64.lognot mask);; let shiftl x n = Int64.shift_left x (Z.to_int n);; let shiftr x n = Int64.shift_right_logical x (Z.to_int n);; let shiftr_signed x n = Int64.shift_right x (Z.to_int n);; let test_bit x n = Int64.compare (Int64.logand x (Int64.shift_left Int64.one (Z.to_int n))) Int64.zero <> 0;; end;; (*struct Uint64*)\ code_reserved OCaml Uint64 code_printing code_module Uint64 \ (Scala) \object Uint64 { def less(x: Long, y: Long) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Long, y: Long) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Long, n: BigInt, b: Boolean) : Long = if (b) x | (1L << n.intValue) else x & (1L << n.intValue).unary_~ def shiftl(x: Long, n: BigInt) : Long = x << n.intValue def shiftr(x: Long, n: BigInt) : Long = x >>> n.intValue def shiftr_signed(x: Long, n: BigInt) : Long = x >> n.intValue def test_bit(x: Long, n: BigInt) : Boolean = (x & (1L << n.intValue)) != 0 } /* object Uint64 */\ code_reserved Scala Uint64 text \ OCaml's conversion from Big\_int to int64 demands that the value fits int a signed 64-bit integer. The following justifies the implementation. \ definition Uint64_signed :: "integer \ uint64" where "Uint64_signed i = (if i < -(0x8000000000000000) \ i \ 0x8000000000000000 then undefined Uint64 i else Uint64 i)" lemma Uint64_code [code]: "Uint64 i = (let i' = i AND 0xFFFFFFFFFFFFFFFF in if bit i' 63 then Uint64_signed (i' - 0x10000000000000000) else Uint64_signed i')" including undefined_transfer integer.lifting unfolding Uint64_signed_def apply transfer apply (subst word_of_int_via_signed) apply (auto simp add: push_bit_of_1 mask_eq_exp_minus_1 word_of_int_via_signed cong del: if_cong) done -lemma Uint64_signed_code [code abstract]: +lemma Uint64_signed_code [code]: "Rep_uint64 (Uint64_signed i) = (if i < -(0x8000000000000000) \ i \ 0x8000000000000000 then Rep_uint64 (undefined Uint64 i) else word_of_int (int_of_integer_symbolic i))" unfolding Uint64_signed_def Uint64_def int_of_integer_symbolic_def word_of_integer_def by(simp add: Abs_uint64_inverse) text \ Avoid @{term Abs_uint64} in generated code, use @{term Rep_uint64'} instead. The symbolic implementations for code\_simp use @{term Rep_uint64}. The new destructor @{term Rep_uint64'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint64} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint64} ([code abstract] equations for @{typ uint64} may use @{term Rep_uint64} because these instances will be folded away.) To convert @{typ "64 word"} values into @{typ uint64}, use @{term "Abs_uint64'"}. \ definition Rep_uint64' where [simp]: "Rep_uint64' = Rep_uint64" lemma Rep_uint64'_transfer [transfer_rule]: "rel_fun cr_uint64 (=) (\x. x) Rep_uint64'" unfolding Rep_uint64'_def by(rule uint64.rep_transfer) lemma Rep_uint64'_code [code]: "Rep_uint64' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint64' :: "64 word \ uint64" is "\x :: 64 word. x" . lemma Abs_uint64'_code [code]: "Abs_uint64' x = Uint64 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint64 \ _"]] lemma term_of_uint64_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint64.uint64.Abs_uint64'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]]]]], TR (STR ''Uint64.uint64'') []])) (term_of_class.term_of (Rep_uint64' x))" by(simp add: term_of_anything) code_printing type_constructor uint64 \ (SML) "Uint64.uint64" and (Haskell) "Uint64.Word64" and (OCaml) "int64" and (Scala) "Long" | constant Uint64 \ (SML) "Uint64.fromInt" and (Haskell) "(Prelude.fromInteger _ :: Uint64.Word64)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint64.Word64)" and (Scala) "_.longValue" | constant Uint64_signed \ (OCaml) "Z.to'_int64" | constant "0 :: uint64" \ (SML) "Uint64.zero" and (Haskell) "(0 :: Uint64.Word64)" and (OCaml) "Int64.zero" and (Scala) "0" | constant "1 :: uint64" \ (SML) "Uint64.one" and (Haskell) "(1 :: Uint64.Word64)" and (OCaml) "Int64.one" and (Scala) "1" | constant "plus :: uint64 \ _ " \ (SML) "Uint64.plus" and (Haskell) infixl 6 "+" and (OCaml) "Int64.add" and (Scala) infixl 7 "+" | constant "uminus :: uint64 \ _" \ (SML) "Uint64.negate" and (Haskell) "negate" and (OCaml) "Int64.neg" and (Scala) "!(- _)" | constant "minus :: uint64 \ _" \ (SML) "Uint64.minus" and (Haskell) infixl 6 "-" and (OCaml) "Int64.sub" and (Scala) infixl 7 "-" | constant "times :: uint64 \ _ \ _" \ (SML) "Uint64.times" and (Haskell) infixl 7 "*" and (OCaml) "Int64.mul" and (Scala) infixl 8 "*" | constant "HOL.equal :: uint64 \ _ \ bool" \ (SML) "!((_ : Uint64.uint64) = _)" and (Haskell) infix 4 "==" and (OCaml) "(Int64.compare _ _ = 0)" and (Scala) infixl 5 "==" | class_instance uint64 :: equal \ (Haskell) - | constant "less_eq :: uint64 \ _ \ bool" \ (SML) "Uint64.less'_eq" and (Haskell) infix 4 "<=" and (OCaml) "Uint64.less'_eq" and (Scala) "Uint64.less'_eq" | constant "less :: uint64 \ _ \ bool" \ (SML) "Uint64.less" and (Haskell) infix 4 "<" and (OCaml) "Uint64.less" and (Scala) "Uint64.less" | constant "NOT :: uint64 \ _" \ (SML) "Uint64.notb" and (Haskell) "Data'_Bits.complement" and (OCaml) "Int64.lognot" and (Scala) "_.unary'_~" | constant "(AND) :: uint64 \ _" \ (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 = push_bit 1 (uint64_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" including undefined_transfer unfolding uint64_divmod_def uint64_sdiv_def div0_uint64_def mod0_uint64_def - by transfer (simp add: divmod_via_sdivmod ac_simps) + less_eq_uint64.rep_eq + apply transfer + apply (simp add: divmod_via_sdivmod push_bit_eq_mult) + done -lemma uint64_sdiv_code [code abstract]: +lemma uint64_sdiv_code [code]: "Rep_uint64 (uint64_sdiv x y) = (if y = 0 then Rep_uint64 (undefined ((div) :: uint64 \ _) x (0 :: uint64)) else Rep_uint64 x sdiv Rep_uint64 y)" unfolding uint64_sdiv_def by(simp add: Abs_uint64_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint64_divmod_code} computes both with division only. \ code_printing constant uint64_div \ (SML) "Uint64.divide" and (Haskell) "Prelude.div" | constant uint64_mod \ (SML) "Uint64.modulus" and (Haskell) "Prelude.mod" | constant uint64_divmod \ (Haskell) "divmod" | constant uint64_sdiv \ (OCaml) "Int64.div" and (Scala) "_ '/ _" definition uint64_test_bit :: "uint64 \ integer \ bool" where [code del]: "uint64_test_bit x n = (if n < 0 \ 63 < n then undefined (bit :: uint64 \ _) x n else bit x (nat_of_integer n))" lemma bit_uint64_code [code]: "bit x n \ n < 64 \ uint64_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint64_test_bit_def - by (transfer, simp, transfer, simp) + by transfer (auto dest: bit_imp_le_length) lemma uint64_test_bit_code [code]: "uint64_test_bit w n = (if n < 0 \ 63 < n then undefined (bit :: uint64 \ _) w n else bit (Rep_uint64 w) (nat_of_integer n))" unfolding uint64_test_bit_def by(simp add: bit_uint64.rep_eq) code_printing constant uint64_test_bit \ (SML) "Uint64.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (OCaml) "Uint64.test'_bit" and (Scala) "Uint64.test'_bit" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_test'_bit out of bounds\") else Uint64.test'_bit x i)" definition uint64_set_bit :: "uint64 \ integer \ bool \ uint64" where [code del]: "uint64_set_bit x n b = (if n < 0 \ 63 < n then undefined (set_bit :: uint64 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint64_code [code]: "set_bit x n b = (if n < 64 then uint64_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint64_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) -lemma uint64_set_bit_code [code abstract]: +lemma uint64_set_bit_code [code]: "Rep_uint64 (uint64_set_bit w n b) = (if n < 0 \ 63 < n then Rep_uint64 (undefined (set_bit :: uint64 \ _) w n b) else set_bit (Rep_uint64 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint64_set_bit_def by transfer simp code_printing constant uint64_set_bit \ (SML) "Uint64.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (OCaml) "Uint64.set'_bit" and (Scala) "Uint64.set'_bit" and (Eval) "(fn x => fn i => fn b => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_set'_bit out of bounds\") else Uint64.set'_bit x i b)" -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 (push_bit 1 w OR (if f n' then 1 else 0)) n')" - apply (transfer fixing: n) - apply (cases n) - apply simp_all - done - -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 = bit x 0" - by transfer (simp add: lsb_word_eq) - - definition uint64_shiftl :: "uint64 \ integer \ uint64" where [code del]: "uint64_shiftl x n = (if n < 0 \ 64 \ n then undefined (push_bit :: nat \ uint64 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint64_code [code]: "push_bit n x = (if n < 64 then uint64_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint64_shiftl_def by transfer simp -lemma uint64_shiftl_code [code abstract]: +lemma uint64_shiftl_code [code]: "Rep_uint64 (uint64_shiftl w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined (push_bit :: nat \ uint64 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint64 w))" including undefined_transfer unfolding uint64_shiftl_def by transfer simp code_printing constant uint64_shiftl \ (SML) "Uint64.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (OCaml) "Uint64.shiftl" and (Scala) "Uint64.shiftl" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_shiftl out of bounds\") else Uint64.shiftl x i)" - definition uint64_shiftr :: "uint64 \ integer \ uint64" where [code del]: "uint64_shiftr x n = (if n < 0 \ 64 \ n then undefined (drop_bit :: nat \ uint64 \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint64_code [code]: "drop_bit n x = (if n < 64 then uint64_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint64_shiftr_def by transfer simp -lemma uint64_shiftr_code [code abstract]: +lemma uint64_shiftr_code [code]: "Rep_uint64 (uint64_shiftr w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined (drop_bit :: nat \ uint64 \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint64 w))" including undefined_transfer unfolding uint64_shiftr_def by transfer simp code_printing constant uint64_shiftr \ (SML) "Uint64.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (OCaml) "Uint64.shiftr" and (Scala) "Uint64.shiftr" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_shiftr out of bounds\") else Uint64.shiftr x i)" - definition uint64_sshiftr :: "uint64 \ integer \ uint64" where [code del]: "uint64_sshiftr x n = - (if n < 0 \ 64 \ n then undefined sshiftr_uint64 x n else sshiftr_uint64 x (nat_of_integer n))" + (if n < 0 \ 64 \ n then undefined signed_drop_bit_uint64 n x else signed_drop_bit_uint64 (nat_of_integer n) x)" lemma sshiftr_uint64_code [code]: - "x >>> n = + "signed_drop_bit_uint64 n x = (if n < 64 then uint64_sshiftr x (integer_of_nat n) else if bit x 63 then - 1 else 0)" including undefined_transfer integer.lifting unfolding uint64_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond) -lemma uint64_sshiftr_code [code abstract]: +lemma uint64_sshiftr_code [code]: "Rep_uint64 (uint64_sshiftr w n) = - (if n < 0 \ 64 \ n then Rep_uint64 (undefined sshiftr_uint64 w n) else signed_drop_bit (nat_of_integer n) (Rep_uint64 w))" + (if n < 0 \ 64 \ n then Rep_uint64 (undefined signed_drop_bit_uint64 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint64 w))" including undefined_transfer unfolding uint64_sshiftr_def by transfer simp code_printing constant uint64_sshiftr \ (SML) "Uint64.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint64.Int64) _)) :: Uint64.Word64)" and (OCaml) "Uint64.shiftr'_signed" and (Scala) "Uint64.shiftr'_signed" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_shiftr'_signed out of bounds\") else Uint64.shiftr'_signed x i)" lemma uint64_msb_test_bit: "msb x \ bit (x :: uint64) 63" by transfer (simp add: msb_word_iff_bit) lemma msb_uint64_code [code]: "msb x \ uint64_test_bit x 63" by (simp add: uint64_test_bit_def uint64_msb_test_bit) lemma uint64_of_int_code [code]: "uint64_of_int i = Uint64 (integer_of_int i)" including integer.lifting by transfer simp lemma int_of_uint64_code [code]: "int_of_uint64 x = int_of_integer (integer_of_uint64 x)" -by(simp add: integer_of_uint64_def) +including integer.lifting by transfer simp lemma nat_of_uint64_code [code]: "nat_of_uint64 x = nat_of_integer (integer_of_uint64 x)" unfolding integer_of_uint64_def including integer.lifting by transfer simp definition integer_of_uint64_signed :: "uint64 \ integer" where "integer_of_uint64_signed n = (if bit n 63 then undefined integer_of_uint64 n else integer_of_uint64 n)" lemma integer_of_uint64_signed_code [code]: "integer_of_uint64_signed n = (if bit n 63 then undefined integer_of_uint64 n else integer_of_int (uint (Rep_uint64' n)))" -unfolding integer_of_uint64_signed_def integer_of_uint64_def -including undefined_transfer by transfer simp + by (simp add: integer_of_uint64_signed_def integer_of_uint64_def) lemma integer_of_uint64_code [code]: "integer_of_uint64 n = (if bit n 63 then integer_of_uint64_signed (n AND 0x7FFFFFFFFFFFFFFF) OR 0x8000000000000000 else integer_of_uint64_signed n)" proof - - have \(0x7FFFFFFFFFFFFFFF :: uint64) = mask 63\ - by (simp add: mask_eq_exp_minus_1) - then have *: \n AND 0x7FFFFFFFFFFFFFFF = take_bit 63 n\ - by (simp add: take_bit_eq_mask) - have **: \(0x8000000000000000 :: int) = 2 ^ 63\ - by simp - show ?thesis - unfolding integer_of_uint64_def integer_of_uint64_signed_def o_def * - including undefined_transfer integer.lifting - apply transfer - apply (rule bit_eqI) - apply (simp add: bit_or_iff bit_take_bit_iff bit_uint_iff) - apply (simp only: bit_exp_iff bit_or_iff **) - apply auto - done + have \integer_of_uint64_signed (n AND 0x7FFFFFFFFFFFFFFF) OR 0x8000000000000000 = Bit_Operations.set_bit 63 (integer_of_uint64_signed (take_bit 63 n))\ + by (simp add: take_bit_eq_mask set_bit_eq_or push_bit_eq_mult mask_eq_exp_minus_1) + moreover have \integer_of_uint64 n = Bit_Operations.set_bit 63 (integer_of_uint64 (take_bit 63 n))\ if \bit n 63\ + proof (rule bit_eqI) + fix m + from that show \bit (integer_of_uint64 n) m = bit (Bit_Operations.set_bit 63 (integer_of_uint64 (take_bit 63 n))) m\ for m + including integer.lifting by transfer (auto simp add: bit_simps dest: bit_imp_le_length) + qed + ultimately show ?thesis + by simp (simp add: integer_of_uint64_signed_def bit_simps) qed 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,695 +1,599 @@ (* Title: Uint8.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 8 bits\ theory Uint8 imports - Code_Target_Word_Base + Code_Target_Word_Base Word_Type_Copies begin text \ Restriction for OCaml code generation: OCaml does not provide an int8 type, so no special code generation for this type is set up. If the theory \Code_Target_Bits_Int\ - is imported, the type \uint8\ is emulated via @{typ "8 word"}. + is imported, the type \uint8\ is emulated via \<^typ>\8 word\. \ section \Type definition and primitive operations\ -typedef uint8 = "UNIV :: 8 word set" .. +typedef uint8 = \UNIV :: 8 word set\ .. + +global_interpretation uint8: word_type_copy Abs_uint8 Rep_uint8 + using type_definition_uint8 by (rule word_type_copy.intro) setup_lifting type_definition_uint8 -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 +declare uint8.of_word_of [code abstype] -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 +declare Quotient_uint8 [transfer_rule] -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] +instantiation uint8 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ 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 zero_uint8 :: uint8 is 0 . +lift_definition one_uint8 :: uint8 is 1 . +lift_definition plus_uint8 :: \uint8 \ uint8 \ uint8\ is \(+)\ . +lift_definition uminus_uint8 :: \uint8 \ uint8\ is uminus . +lift_definition minus_uint8 :: \uint8 \ uint8 \ uint8\ is \(-)\ . +lift_definition times_uint8 :: \uint8 \ uint8 \ uint8\ is \(*)\ . +lift_definition divide_uint8 :: \uint8 \ uint8 \ uint8\ is \(div)\ . +lift_definition modulo_uint8 :: \uint8 \ uint8 \ uint8\ is \(mod)\ . +lift_definition equal_uint8 :: \uint8 \ uint8 \ bool\ is \HOL.equal\ . +lift_definition less_eq_uint8 :: \uint8 \ uint8 \ bool\ is \(\)\ . +lift_definition less_uint8 :: \uint8 \ uint8 \ bool\ is \(<)\ . -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)+ +global_interpretation uint8: word_type_copy_ring Abs_uint8 Rep_uint8 + by standard (fact zero_uint8.rep_eq one_uint8.rep_eq + plus_uint8.rep_eq uminus_uint8.rep_eq minus_uint8.rep_eq + times_uint8.rep_eq divide_uint8.rep_eq modulo_uint8.rep_eq + equal_uint8.rep_eq less_eq_uint8.rep_eq less_uint8.rep_eq)+ -end +instance proof - + show \OFCLASS(uint8, comm_ring_1_class)\ + by (rule uint8.of_class_comm_ring_1) + show \OFCLASS(uint8, semiring_modulo_class)\ + by (fact uint8.of_class_semiring_modulo) + show \OFCLASS(uint8, equal_class)\ + by (fact uint8.of_class_equal) + show \OFCLASS(uint8, linorder_class)\ + by (fact uint8.of_class_linorder) +qed -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 bit_uint8 :: \uint8 \ nat \ bool\ is bit . lift_definition not_uint8 :: \uint8 \ uint8\ is NOT . lift_definition and_uint8 :: \uint8 \ uint8 \ uint8\ is \(AND)\ . lift_definition or_uint8 :: \uint8 \ uint8 \ uint8\ is \(OR)\ . lift_definition xor_uint8 :: \uint8 \ uint8 \ uint8\ is \(XOR)\ . lift_definition mask_uint8 :: \nat \ uint8\ is mask . -lift_definition set_bit_uint8 :: \nat \ uint8 \ uint8\ is \Bit_Operations.set_bit\ . -lift_definition unset_bit_uint8 :: \nat \ uint8 \ uint8\ is \unset_bit\ . -lift_definition flip_bit_uint8 :: \nat \ uint8 \ uint8\ is \flip_bit\ . -instance by (standard; transfer) - (simp_all add: bit_simps mask_eq_decr_exp minus_eq_not_minus_1 set_bit_def flip_bit_def) -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) - -lemma [code]: - \Bit_Operations.set_bit n w = w OR push_bit n 1\ for w :: uint8 - by (fact set_bit_eq_or) - -lemma [code]: - \unset_bit n w = w AND NOT (push_bit n 1)\ for w :: uint8 - by (fact unset_bit_eq_and_not) - -lemma [code]: - \flip_bit n w = w XOR push_bit n 1\ for w :: uint8 - by (fact flip_bit_eq_xor) - -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 - -setup \Context.theory_map (Name_Space.map_naming (Name_Space.qualified_path true \<^binding>\Generic\))\ - -instantiation uint8 :: set_bit -begin -lift_definition set_bit_uint8 :: \uint8 \ nat \ bool \ uint8\ is set_bit . -instance - apply standard - apply transfer - apply (simp add: bit_simps) - done -end - -setup \Context.theory_map (Name_Space.map_naming (Name_Space.parent_path))\ +lift_definition push_bit_uint8 :: \nat \ uint8 \ uint8\ is push_bit . +lift_definition drop_bit_uint8 :: \nat \ uint8 \ uint8\ is drop_bit . +lift_definition signed_drop_bit_uint8 :: \nat \ uint8 \ uint8\ is signed_drop_bit . +lift_definition take_bit_uint8 :: \nat \ uint8 \ uint8\ is take_bit . +lift_definition set_bit_uint8 :: \nat \ uint8 \ uint8\ is Bit_Operations.set_bit . +lift_definition unset_bit_uint8 :: \nat \ uint8 \ uint8\ is unset_bit . +lift_definition flip_bit_uint8 :: \nat \ uint8 \ uint8\ is flip_bit . -instantiation uint8 :: bit_comprehension begin -lift_definition set_bits_uint8 :: "(nat \ bool) \ uint8" is "set_bits" . -instance by (standard; transfer) (fact set_bits_bit_eq) -end - -lemmas [code] = 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 \\w n. signed_drop_bit n w\ . - -lift_definition uint8_of_int :: "int \ uint8" is "word_of_int" . +global_interpretation uint8: word_type_copy_bits Abs_uint8 Rep_uint8 signed_drop_bit_uint8 + by standard (fact bit_uint8.rep_eq not_uint8.rep_eq and_uint8.rep_eq or_uint8.rep_eq xor_uint8.rep_eq + mask_uint8.rep_eq push_bit_uint8.rep_eq drop_bit_uint8.rep_eq signed_drop_bit_uint8.rep_eq take_bit_uint8.rep_eq + set_bit_uint8.rep_eq unset_bit_uint8.rep_eq flip_bit_uint8.rep_eq)+ -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) +instance + by (fact uint8.of_class_ring_bit_operations) 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) +lift_definition uint8_of_nat :: \nat \ uint8\ + is word_of_nat . -lemma Abs_uint8_0 [code_post]: "Abs_uint8 0 = 0" -by(simp add: zero_uint8_def) +lift_definition nat_of_uint8 :: \uint8 \ nat\ + is unat . -lemma Abs_uint8_1 [code_post]: "Abs_uint8 1 = 1" -by(simp add: one_uint8_def) +lift_definition uint8_of_int :: \int \ uint8\ + is word_of_int . + +lift_definition int_of_uint8 :: \uint8 \ int\ + is uint . + +context + includes integer.lifting +begin + +lift_definition Uint8 :: \integer \ uint8\ + is word_of_int . + +lift_definition integer_of_uint8 :: \uint8 \ integer\ + is uint . + +end + +global_interpretation uint8: word_type_copy_more Abs_uint8 Rep_uint8 signed_drop_bit_uint8 + uint8_of_nat nat_of_uint8 uint8_of_int int_of_uint8 Uint8 integer_of_uint8 + apply standard + apply (simp_all add: uint8_of_nat.rep_eq nat_of_uint8.rep_eq + uint8_of_int.rep_eq int_of_uint8.rep_eq + Uint8.rep_eq integer_of_uint8.rep_eq integer_eq_iff) + done + +instantiation uint8 :: "{size, msb, lsb, set_bit, bit_comprehension}" +begin + +lift_definition size_uint8 :: \uint8 \ nat\ is size . + +lift_definition msb_uint8 :: \uint8 \ bool\ is msb . +lift_definition lsb_uint8 :: \uint8 \ bool\ is lsb . + +text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ + +definition set_bit_uint8 :: \uint8 \ nat \ bool \ uint8\ + where set_bit_uint8_eq: \set_bit_uint8 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ + +context + includes lifting_syntax +begin + +lemma set_bit_uint8_transfer [transfer_rule]: + \(cr_uint8 ===> (=) ===> (\) ===> cr_uint8) Generic_set_bit.set_bit Generic_set_bit.set_bit\ + by (simp only: set_bit_eq [abs_def] set_bit_uint8_eq [abs_def]) transfer_prover + +end + +lift_definition set_bits_uint8 :: \(nat \ bool) \ uint8\ is set_bits . +lift_definition set_bits_aux_uint8 :: \(nat \ bool) \ nat \ uint8 \ uint8\ is set_bits_aux . + +global_interpretation uint8: word_type_copy_misc Abs_uint8 Rep_uint8 signed_drop_bit_uint8 + uint8_of_nat nat_of_uint8 uint8_of_int int_of_uint8 Uint8 integer_of_uint8 8 set_bits_aux_uint8 + by (standard; transfer) simp_all + +instance using uint8.of_class_bit_comprehension + uint8.of_class_set_bit uint8.of_class_lsb + by simp_all standard + +end section \Code setup\ code_printing code_module Uint8 \ (SML) \(* Test that words can handle numbers between 0 and 3 *) val _ = if 3 <= Word.wordSize then () else raise (Fail ("wordSize less than 3")); structure Uint8 : sig val set_bit : Word8.word -> IntInf.int -> bool -> Word8.word val shiftl : Word8.word -> IntInf.int -> Word8.word val shiftr : Word8.word -> IntInf.int -> Word8.word val shiftr_signed : Word8.word -> IntInf.int -> Word8.word val test_bit : Word8.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word8.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word8.orb (x, mask) else Word8.andb (x, Word8.notb mask) end fun shiftl x n = Word8.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word8.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word8.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word8.andb (x, Word8.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word8.fromInt 0 end; (* struct Uint8 *)\ code_reserved SML Uint8 code_printing code_module Uint8 \ (Haskell) \module Uint8(Int8, Word8) where import Data.Int(Int8) import Data.Word(Word8)\ code_reserved Haskell Uint8 text \ Scala provides only signed 8bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module Uint8 \ (Scala) \object Uint8 { def less(x: Byte, y: Byte) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Byte, y: Byte) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Byte, n: BigInt, b: Boolean) : Byte = if (b) (x | (1 << n.intValue)).toByte else (x & (1 << n.intValue).unary_~).toByte def shiftl(x: Byte, n: BigInt) : Byte = (x << n.intValue).toByte def shiftr(x: Byte, n: BigInt) : Byte = ((x & 255) >>> n.intValue).toByte def shiftr_signed(x: Byte, n: BigInt) : Byte = (x >> n.intValue).toByte def test_bit(x: Byte, n: BigInt) : Boolean = (x & (1 << n.intValue)) != 0 } /* object Uint8 */\ code_reserved Scala Uint8 text \ Avoid @{term Abs_uint8} in generated code, use @{term Rep_uint8'} instead. The symbolic implementations for code\_simp use @{term Rep_uint8}. The new destructor @{term Rep_uint8'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint8} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint8} ([code abstract] equations for @{typ uint8} may use @{term Rep_uint8} because these instances will be folded away.) To convert @{typ "8 word"} values into @{typ uint8}, use @{term "Abs_uint8'"}. \ definition Rep_uint8' where [simp]: "Rep_uint8' = Rep_uint8" lemma Rep_uint8'_transfer [transfer_rule]: "rel_fun cr_uint8 (=) (\x. x) Rep_uint8'" unfolding Rep_uint8'_def by(rule uint8.rep_transfer) lemma Rep_uint8'_code [code]: "Rep_uint8' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint8' :: "8 word \ uint8" is "\x :: 8 word. x" . lemma Abs_uint8'_code [code]: "Abs_uint8' x = Uint8 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint8 \ _"]] lemma term_of_uint8_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint8.uint8.Abs_uint8'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]], TR (STR ''Uint8.uint8'') []])) (term_of_class.term_of (Rep_uint8' x))" by(simp add: term_of_anything) -lemma Uin8_code [code abstract]: "Rep_uint8 (Uint8 i) = word_of_int (int_of_integer_symbolic i)" +lemma Uin8_code [code]: "Rep_uint8 (Uint8 i) = word_of_int (int_of_integer_symbolic i)" unfolding Uint8_def int_of_integer_symbolic_def by(simp add: Abs_uint8_inverse) code_printing type_constructor uint8 \ (SML) "Word8.word" and (Haskell) "Uint8.Word8" and (Scala) "Byte" | constant Uint8 \ (SML) "Word8.fromLargeInt (IntInf.toLarge _)" and (Haskell) "(Prelude.fromInteger _ :: Uint8.Word8)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint8.Word8)" and (Scala) "_.byteValue" | constant "0 :: uint8" \ (SML) "(Word8.fromInt 0)" and (Haskell) "(0 :: Uint8.Word8)" and (Scala) "0.toByte" | constant "1 :: uint8" \ (SML) "(Word8.fromInt 1)" and (Haskell) "(1 :: Uint8.Word8)" and (Scala) "1.toByte" | constant "plus :: uint8 \ _ \ _" \ (SML) "Word8.+ ((_), (_))" and (Haskell) infixl 6 "+" and (Scala) "(_ +/ _).toByte" | constant "uminus :: uint8 \ _" \ (SML) "Word8.~" and (Haskell) "negate" and (Scala) "(- _).toByte" | constant "minus :: uint8 \ _" \ (SML) "Word8.- ((_), (_))" and (Haskell) infixl 6 "-" and (Scala) "(_ -/ _).toByte" | constant "times :: uint8 \ _ \ _" \ (SML) "Word8.* ((_), (_))" and (Haskell) infixl 7 "*" and (Scala) "(_ */ _).toByte" | constant "HOL.equal :: uint8 \ _ \ bool" \ (SML) "!((_ : Word8.word) = _)" and (Haskell) infix 4 "==" and (Scala) infixl 5 "==" | class_instance uint8 :: equal \ (Haskell) - | constant "less_eq :: uint8 \ _ \ bool" \ (SML) "Word8.<= ((_), (_))" and (Haskell) infix 4 "<=" and (Scala) "Uint8.less'_eq" | constant "less :: uint8 \ _ \ bool" \ (SML) "Word8.< ((_), (_))" and (Haskell) infix 4 "<" and (Scala) "Uint8.less" | constant "NOT :: uint8 \ _" \ (SML) "Word8.notb" and (Haskell) "Data'_Bits.complement" and (Scala) "_.unary'_~.toByte" | constant "(AND) :: uint8 \ _" \ (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 + including undefined_transfer unfolding uint8_divmod_def uint8_div_def by transfer (simp add: word_div_def) lemma mod_uint8_code [code]: "x mod y = (if y = 0 then x else uint8_mod x y)" including undefined_transfer unfolding uint8_mod_def uint8_divmod_def by transfer (simp add: word_mod_def) definition uint8_sdiv :: "uint8 \ uint8 \ uint8" where "uint8_sdiv x y = (if y = 0 then undefined ((div) :: uint8 \ _) x (0 :: uint8) else Abs_uint8 (Rep_uint8 x sdiv Rep_uint8 y))" definition div0_uint8 :: "uint8 \ uint8" where [code del]: "div0_uint8 x = undefined ((div) :: uint8 \ _) x (0 :: uint8)" declare [[code abort: div0_uint8]] definition mod0_uint8 :: "uint8 \ uint8" where [code del]: "mod0_uint8 x = undefined ((mod) :: uint8 \ _) x (0 :: uint8)" declare [[code abort: mod0_uint8]] lemma uint8_divmod_code [code]: "uint8_divmod x y = (if 0x80 \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint8 x, mod0_uint8 x) else let q = push_bit 1 (uint8_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" -including undefined_transfer unfolding uint8_divmod_def uint8_sdiv_def div0_uint8_def mod0_uint8_def + including undefined_transfer unfolding uint8_divmod_def uint8_sdiv_def div0_uint8_def mod0_uint8_def + less_eq_uint8.rep_eq apply transfer apply (simp add: divmod_via_sdivmod push_bit_eq_mult) done -lemma uint8_sdiv_code [code abstract]: +lemma uint8_sdiv_code [code]: "Rep_uint8 (uint8_sdiv x y) = (if y = 0 then Rep_uint8 (undefined ((div) :: uint8 \ _) x (0 :: uint8)) else Rep_uint8 x sdiv Rep_uint8 y)" unfolding uint8_sdiv_def by(simp add: Abs_uint8_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint8_divmod_code} computes both with division only. \ code_printing constant uint8_div \ (SML) "Word8.div ((_), (_))" and (Haskell) "Prelude.div" | constant uint8_mod \ (SML) "Word8.mod ((_), (_))" and (Haskell) "Prelude.mod" | constant uint8_divmod \ (Haskell) "divmod" | constant uint8_sdiv \ (Scala) "(_ '/ _).toByte" definition uint8_test_bit :: "uint8 \ integer \ bool" where [code del]: "uint8_test_bit x n = (if n < 0 \ 7 < n then undefined (bit :: uint8 \ _) x n else bit x (nat_of_integer n))" lemma bit_uint8_code [code]: "bit x n \ n < 8 \ uint8_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint8_test_bit_def by (transfer, simp, transfer, simp) lemma uint8_test_bit_code [code]: "uint8_test_bit w n = (if n < 0 \ 7 < n then undefined (bit :: uint8 \ _) w n else bit (Rep_uint8 w) (nat_of_integer n))" unfolding uint8_test_bit_def by (simp add: bit_uint8.rep_eq) code_printing constant uint8_test_bit \ (SML) "Uint8.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (Scala) "Uint8.test'_bit" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_test'_bit out of bounds\") else Uint8.test'_bit x i)" definition uint8_set_bit :: "uint8 \ integer \ bool \ uint8" where [code del]: "uint8_set_bit x n b = (if n < 0 \ 7 < n then undefined (set_bit :: uint8 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint8_code [code]: "set_bit x n b = (if n < 8 then uint8_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint8_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) -lemma uint8_set_bit_code [code abstract]: +lemma uint8_set_bit_code [code]: "Rep_uint8 (uint8_set_bit w n b) = (if n < 0 \ 7 < n then Rep_uint8 (undefined (set_bit :: uint8 \ _) w n b) else set_bit (Rep_uint8 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint8_set_bit_def by transfer simp code_printing constant uint8_set_bit \ (SML) "Uint8.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (Scala) "Uint8.set'_bit" and (Eval) "(fn x => fn i => fn b => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_set'_bit out of bounds\") else Uint8.set'_bit x i b)" -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 (push_bit 1 w OR (if f n' then 1 else 0)) n')" - apply (transfer fixing: n) - apply (cases n) - apply simp_all - done - -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 = bit x 0" - by (simp add: lsb_odd) - - definition uint8_shiftl :: "uint8 \ integer \ uint8" where [code del]: "uint8_shiftl x n = (if n < 0 \ 8 \ n then undefined (push_bit :: nat \ uint8 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint8_code [code]: "push_bit n x = (if n < 8 then uint8_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint8_shiftl_def by transfer simp -lemma uint8_shiftl_code [code abstract]: +lemma uint8_shiftl_code [code]: "Rep_uint8 (uint8_shiftl w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined (push_bit :: nat \ uint8 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint8 w))" including undefined_transfer unfolding uint8_shiftl_def by transfer simp code_printing constant uint8_shiftl \ (SML) "Uint8.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (Scala) "Uint8.shiftl" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_shiftl out of bounds\") else Uint8.shiftl x i)" definition uint8_shiftr :: "uint8 \ integer \ uint8" where [code del]: "uint8_shiftr x n = (if n < 0 \ 8 \ n then undefined (drop_bit :: _ \ _ \ uint8) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint8_code [code]: "drop_bit n x = (if n < 8 then uint8_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint8_shiftr_def by transfer simp -lemma uint8_shiftr_code [code abstract]: +lemma uint8_shiftr_code [code]: "Rep_uint8 (uint8_shiftr w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined (drop_bit :: _ \ _ \ uint8) w n) else drop_bit (nat_of_integer n) (Rep_uint8 w))" including undefined_transfer unfolding uint8_shiftr_def by transfer simp code_printing constant uint8_shiftr \ (SML) "Uint8.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (Scala) "Uint8.shiftr" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_shiftr out of bounds\") else Uint8.shiftr x i)" definition uint8_sshiftr :: "uint8 \ integer \ uint8" where [code del]: "uint8_sshiftr x n = - (if n < 0 \ 8 \ n then undefined sshiftr_uint8 x n else sshiftr_uint8 x (nat_of_integer n))" + (if n < 0 \ 8 \ n then undefined signed_drop_bit_uint8 n x else signed_drop_bit_uint8 (nat_of_integer n) x)" lemma sshiftr_uint8_code [code]: - "x >>> n = + "signed_drop_bit_uint8 n x = (if n < 8 then uint8_sshiftr x (integer_of_nat n) else if bit x 7 then -1 else 0)" including undefined_transfer integer.lifting unfolding uint8_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond word_size) -lemma uint8_sshiftr_code [code abstract]: +lemma uint8_sshiftr_code [code]: "Rep_uint8 (uint8_sshiftr w n) = - (if n < 0 \ 8 \ n then Rep_uint8 (undefined sshiftr_uint8 w n) + (if n < 0 \ 8 \ n then Rep_uint8 (undefined signed_drop_bit_uint8 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint8 w))" including undefined_transfer unfolding uint8_sshiftr_def by transfer simp code_printing constant uint8_sshiftr \ (SML) "Uint8.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint8.Int8) _)) :: Uint8.Word8)" and (Scala) "Uint8.shiftr'_signed" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_sshiftr out of bounds\") else Uint8.shiftr'_signed x i)" lemma uint8_msb_test_bit: "msb x \ bit (x :: uint8) 7" by transfer (simp add: msb_word_iff_bit) lemma msb_uint16_code [code]: "msb x \ uint8_test_bit x 7" by (simp add: uint8_test_bit_def uint8_msb_test_bit) lemma uint8_of_int_code [code]: "uint8_of_int i = Uint8 (integer_of_int i)" including integer.lifting by transfer simp lemma int_of_uint8_code [code]: "int_of_uint8 x = int_of_integer (integer_of_uint8 x)" -by(simp add: integer_of_uint8_def) + by (simp add: int_of_uint8.rep_eq integer_of_uint8_def) lemma nat_of_uint8_code [code]: "nat_of_uint8 x = nat_of_integer (integer_of_uint8 x)" -unfolding integer_of_uint8_def including integer.lifting by transfer simp + unfolding integer_of_uint8_def including integer.lifting by transfer simp definition integer_of_uint8_signed :: "uint8 \ integer" where "integer_of_uint8_signed n = (if bit n 7 then undefined integer_of_uint8 n else integer_of_uint8 n)" lemma integer_of_uint8_signed_code [code]: "integer_of_uint8_signed n = (if bit n 7 then undefined integer_of_uint8 n else integer_of_int (uint (Rep_uint8' n)))" -unfolding integer_of_uint8_signed_def integer_of_uint8_def -including undefined_transfer by transfer simp + by (simp add: integer_of_uint8_signed_def integer_of_uint8_def) lemma integer_of_uint8_code [code]: "integer_of_uint8 n = (if bit n 7 then integer_of_uint8_signed (n AND 0x7F) OR 0x80 else integer_of_uint8_signed n)" proof - - have \(0x7F :: uint8) = mask 7\ - by (simp add: mask_eq_exp_minus_1) - then have *: \n AND 0x7F = take_bit 7 n\ - by (simp only: take_bit_eq_mask) - have **: \(0x80 :: int) = 2 ^ 7\ - by simp - show ?thesis - unfolding integer_of_uint8_def integer_of_uint8_signed_def o_def * - including undefined_transfer integer.lifting - apply transfer - apply (auto simp add: bit_take_bit_iff uint_take_bit_eq) - apply (rule bit_eqI) - apply (simp add: bit_uint_iff bit_or_iff bit_take_bit_iff) - apply (simp only: ** bit_exp_iff) - apply auto - done + have \integer_of_uint8_signed (n AND 0x7F) OR 0x80 = Bit_Operations.set_bit 7 (integer_of_uint8_signed (take_bit 7 n))\ + by (simp add: take_bit_eq_mask set_bit_eq_or push_bit_eq_mult mask_eq_exp_minus_1) + moreover have \integer_of_uint8 n = Bit_Operations.set_bit 7 (integer_of_uint8 (take_bit 7 n))\ if \bit n 7\ + proof (rule bit_eqI) + fix m + from that show \bit (integer_of_uint8 n) m = bit (Bit_Operations.set_bit 7 (integer_of_uint8 (take_bit 7 n))) m\ for m + including integer.lifting by transfer (auto simp add: bit_simps dest: bit_imp_le_length) + qed + ultimately show ?thesis + by simp (simp add: integer_of_uint8_signed_def bit_simps) qed 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/Native_Word/Word_Type_Copies.thy b/thys/Native_Word/Word_Type_Copies.thy new file mode 100644 --- /dev/null +++ b/thys/Native_Word/Word_Type_Copies.thy @@ -0,0 +1,344 @@ +(* Title: Word_Type_Copies.thy + Author: Florian Haftmann, TU Muenchen +*) + +chapter \Systematic approach towards type copies of word type\ + +theory Word_Type_Copies + imports + "HOL-Library.Word" + "Word_Lib.Most_significant_bit" + "Word_Lib.Least_significant_bit" + "Word_Lib.Generic_set_bit" + Code_Target_Word_Base +begin + +lemma int_of_integer_unsigned_eq [simp]: + \int_of_integer (unsigned w) = uint w\ + by transfer simp + +lemma int_of_integer_signed_eq [simp]: + \int_of_integer (signed w) = sint w\ + by transfer simp + +lemma word_of_int_of_integer_eq [simp]: + \word_of_int (int_of_integer k) = word_of_integer k\ + by (simp add: word_of_integer.rep_eq) + + +text \The lifting machinery is not localized, hence the abstract proofs are carried out using morphisms.\ + +locale word_type_copy = + fixes of_word :: \'b::len word \ 'a\ + and word_of :: \'a \ 'b word\ + assumes type_definition: \type_definition word_of of_word UNIV\ +begin + +lemma word_of_word: + \word_of (of_word w) = w\ + using type_definition by (simp add: type_definition_def) + +lemma of_word_of [code abstype]: + \of_word (word_of p) = p\ + \ \Use an abstract type for code generation to disable pattern matching on \<^term>\of_word\.\ + using type_definition by (simp add: type_definition_def) + +lemma word_of_eqI: + \p = q\ if \word_of p = word_of q\ +proof - + from that have \of_word (word_of p) = of_word (word_of q)\ + by simp + then show ?thesis + by (simp add: of_word_of) +qed + +lemma eq_iff_word_of: + \p = q \ word_of p = word_of q\ + by (auto intro: word_of_eqI) + +end + +bundle constraintless +begin + +declaration \ + let + val cs = map (rpair NONE o fst o dest_Const) + [\<^term>\0\, \<^term>\(+)\, \<^term>\uminus\, \<^term>\(-)\, + \<^term>\1\, \<^term>\(*)\, \<^term>\(div)\, \<^term>\(mod)\, + \<^term>\HOL.equal\, \<^term>\(\)\, \<^term>\(<)\, + \<^term>\(dvd)\, \<^term>\of_bool\, \<^term>\numeral\, \<^term>\of_nat\, + \<^term>\bit\, + \<^term>\NOT\, \<^term>\(AND)\, \<^term>\(OR)\, \<^term>\(XOR)\, \<^term>\mask\, + \<^term>\push_bit\, \<^term>\drop_bit\, \<^term>\take_bit\, + \<^term>\Bit_Operations.set_bit\, \<^term>\unset_bit\, \<^term>\flip_bit\, + \<^term>\msb\, \<^term>\lsb\, \<^term>\size\, \<^term>\Generic_set_bit.set_bit\, \<^term>\set_bits\] + in + K (Context.mapping I (fold Proof_Context.add_const_constraint cs)) + end +\ + +end + +locale word_type_copy_ring = word_type_copy + opening constraintless + + constrains word_of :: \'a \ 'b::len word\ + assumes word_of_0 [code]: \word_of 0 = 0\ + and word_of_1 [code]: \word_of 1 = 1\ + and word_of_add [code]: \word_of (p + q) = word_of p + word_of q\ + and word_of_minus [code]: \word_of (- p) = - (word_of p)\ + and word_of_diff [code]: \word_of (p - q) = word_of p - word_of q\ + and word_of_mult [code]: \word_of (p * q) = word_of p * word_of q\ + and word_of_div [code]: \word_of (p div q) = word_of p div word_of q\ + and word_of_mod [code]: \word_of (p mod q) = word_of p mod word_of q\ + and equal_iff_word_of [code]: \HOL.equal p q \ HOL.equal (word_of p) (word_of q)\ + and less_eq_iff_word_of [code]: \p \ q \ word_of p \ word_of q\ + and less_iff_word_of [code]: \p < q \ word_of p < word_of q\ +begin + +lemma of_class_comm_ring_1: + \OFCLASS('a, comm_ring_1_class)\ + by standard (simp_all add: eq_iff_word_of word_of_0 word_of_1 + word_of_add word_of_minus word_of_diff word_of_mult algebra_simps) + +lemma of_class_semiring_modulo: + \OFCLASS('a, semiring_modulo_class)\ + by standard (simp_all add: eq_iff_word_of word_of_0 word_of_1 + word_of_add word_of_minus word_of_diff word_of_mult word_of_mod word_of_div algebra_simps + mod_mult_div_eq) + +lemma of_class_equal: + \OFCLASS('a, equal_class)\ + by standard (simp add: eq_iff_word_of equal_iff_word_of equal) + +lemma of_class_linorder: + \OFCLASS('a, linorder_class)\ + by standard (auto simp add: eq_iff_word_of less_eq_iff_word_of less_iff_word_of) + +end + +locale word_type_copy_bits = word_type_copy_ring + opening constraintless + + constrains word_of :: \'a::{comm_ring_1, semiring_modulo, equal, linorder} \ 'b::len word\ + fixes signed_drop_bit :: \nat \ 'a \ 'a\ + assumes bit_eq_word_of [code]: \bit p = bit (word_of p)\ + and word_of_not [code]: \word_of (NOT p) = NOT (word_of p)\ + and word_of_and [code]: \word_of (p AND q) = word_of p AND word_of q\ + and word_of_or [code]: \word_of (p OR q) = word_of p OR word_of q\ + and word_of_xor [code]: \word_of (p XOR q) = word_of p XOR word_of q\ + and word_of_mask [code]: \word_of (mask n) = mask n\ + and word_of_push_bit [code]: \word_of (push_bit n p) = push_bit n (word_of p)\ + and word_of_drop_bit [code]: \word_of (drop_bit n p) = drop_bit n (word_of p)\ + and word_of_signed_drop_bit [code]: \word_of (signed_drop_bit n p) = Word.signed_drop_bit n (word_of p)\ + and word_of_take_bit [code]: \word_of (take_bit n p) = take_bit n (word_of p)\ + and word_of_set_bit [code]: \word_of (Bit_Operations.set_bit n p) = Bit_Operations.set_bit n (word_of p)\ + and word_of_unset_bit [code]: \word_of (unset_bit n p) = unset_bit n (word_of p)\ + and word_of_flip_bit [code]: \word_of (flip_bit n p) = flip_bit n (word_of p)\ +begin + +lemma word_of_bool: + \word_of (of_bool n) = of_bool n\ + by (simp add: word_of_0 word_of_1) + +lemma word_of_nat: + \word_of (of_nat n) = of_nat n\ + by (induction n) (simp_all add: word_of_0 word_of_1 word_of_add) + +lemma word_of_numeral [simp]: + \word_of (numeral n) = numeral n\ +proof - + have \word_of (of_nat (numeral n)) = of_nat (numeral n)\ + by (simp only: word_of_nat) + then show ?thesis by simp +qed + +lemma word_of_power: + \word_of (p ^ n) = word_of p ^ n\ + by (induction n) (simp_all add: word_of_1 word_of_mult word_of_numeral) + +lemma even_iff_word_of: + \2 dvd p \ 2 dvd word_of p\ (is \?P \ ?Q\) +proof + assume ?P + then obtain q where \p = 2 * q\ .. + then show ?Q by (simp add: word_of_mult word_of_numeral) +next + assume ?Q + then obtain w where \word_of p = 2 * w\ .. + then have \of_word (word_of p) = of_word (2 * w)\ + by simp + then have \p = 2 * of_word w\ + by (simp add: eq_iff_word_of word_of_word word_of_mult word_of_numeral) + then show ?P + by simp +qed + +lemma of_class_ring_bit_operations: + \OFCLASS('a, ring_bit_operations_class)\ +proof - + have induct: \P p\ if stable: \(\p. p div 2 = p \ P p)\ + and rec: \(\p b. P p \ (of_bool b + 2 * p) div 2 = p \ P (of_bool b + 2 * p))\ + for p :: 'a and P + proof - + from stable have stable': \(\p. word_of p div 2 = word_of p \ P p)\ + by (simp add: eq_iff_word_of word_of_div word_of_numeral) + from rec have rec': \(\p b. P p \ (of_bool b + 2 * word_of p) div 2 = word_of p \ P (of_bool b + 2 * p))\ + by (simp add: eq_iff_word_of word_of_add word_of_bool word_of_mult word_of_div word_of_numeral) + define w where \w = word_of p\ + then have \p = of_word w\ + by (simp add: of_word_of) + also have \P (of_word w)\ + proof (induction w rule: bits_induct) + case (stable w) + show ?case + by (rule stable') (simp add: word_of_word stable) + next + case (rec w b) + have \P (of_bool b + 2 * of_word w)\ + by (rule rec') (simp_all add: word_of_word rec) + also have \of_bool b + 2 * of_word w = of_word (of_bool b + 2 * w)\ + by (simp add: eq_iff_word_of word_of_word word_of_add word_of_1 word_of_mult word_of_numeral word_of_0) + finally show ?case . + qed + finally show \P p\ . + qed + have \class.semiring_parity_axioms (+) (0::'a) (*) 1 (mod)\ + by standard (simp_all add: eq_iff_word_of word_of_0 word_of_1 word_of_numeral even_iff_word_of word_of_mod even_iff_mod_2_eq_zero) + with of_class_semiring_modulo have \OFCLASS('a, semiring_parity_class)\ + by (rule semiring_parity_class.intro) + moreover have \class.semiring_bits_axioms (+) (-) (0::'a) (*) 1 (div) (mod) bit\ + apply (standard, fact induct) + apply (simp_all only: eq_iff_word_of word_of_0 word_of_1 word_of_bool word_of_numeral + word_of_add word_of_diff word_of_mult word_of_div word_of_mod word_of_power even_iff_word_of + bit_eq_word_of push_bit_take_bit drop_bit_take_bit + even_drop_bit_iff_not_bit + flip: push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod mask_eq_exp_minus_1) + apply (auto simp add: ac_simps bit_simps drop_bit_exp_eq push_bit_of_1) + done + ultimately have \OFCLASS('a, semiring_bits_class)\ + by (rule semiring_bits_class.intro) + moreover have \class.semiring_bit_shifts_axioms (+) (*) (1::'a) (div) (mod) push_bit drop_bit take_bit\ + by standard (simp_all add: eq_iff_word_of word_of_push_bit word_of_mult word_of_power word_of_numeral + word_of_drop_bit word_of_div word_of_take_bit word_of_mod + flip: push_bit_eq_mult drop_bit_eq_div take_bit_eq_mod) + ultimately have \OFCLASS('a, semiring_bit_shifts_class)\ + by (rule semiring_bit_shifts_class.intro) + moreover have \class.semiring_bit_operations_axioms (+) (-) (*) (1::'a) bit push_bit (AND) (OR) (XOR) mask Bit_Operations.set_bit unset_bit flip_bit\ + by standard (simp_all add: eq_iff_word_of word_of_push_bit word_of_power word_of_numeral + bit_eq_word_of word_of_and word_of_or word_of_xor word_of_mask word_of_diff word_of_1 bit_simps + word_of_set_bit set_bit_eq_or word_of_unset_bit word_of_flip_bit flip_bit_eq_xor + flip: mask_eq_exp_minus_1) + ultimately have \OFCLASS('a, semiring_bit_operations_class)\ + by (rule semiring_bit_operations_class.intro) + moreover have \OFCLASS('a, ring_parity_class)\ + using \OFCLASS('a, semiring_parity_class)\ by (rule ring_parity_class.intro) standard + moreover have \class.ring_bit_operations_axioms (+) (-) (0::'a) (*) 1 bit uminus NOT\ + by standard (simp_all add: eq_iff_word_of word_of_power word_of_numeral + bit_eq_word_of word_of_diff word_of_1 bit_simps + word_of_not word_of_0 + word_of_minus minus_eq_not_minus_1) + ultimately show \OFCLASS('a, ring_bit_operations_class)\ + by (rule ring_bit_operations_class.intro) +qed + +lemma [code]: + \take_bit n a = a AND mask n\ for a :: 'a + by (simp add: eq_iff_word_of word_of_take_bit word_of_and word_of_mask take_bit_eq_mask) + +lemma [code]: + \mask (Suc n) = push_bit n (1 :: 'a) OR mask n\ + \mask 0 = (0 :: 'a)\ + by (simp_all add: eq_iff_word_of word_of_mask word_of_or word_of_push_bit word_of_0 word_of_1 mask_Suc_exp push_bit_of_1) + +lemma [code]: + \Bit_Operations.set_bit n w = w OR push_bit n 1\ for w :: 'a + by (simp add: eq_iff_word_of word_of_set_bit word_of_or word_of_push_bit word_of_1 set_bit_eq_or) + +lemma [code]: + \unset_bit n w = w AND NOT (push_bit n 1)\ for w :: 'a + by (simp add: eq_iff_word_of word_of_unset_bit word_of_and word_of_not word_of_push_bit word_of_1 unset_bit_eq_and_not) + +lemma [code]: + \flip_bit n w = w XOR push_bit n 1\ for w :: 'a + by (simp add: eq_iff_word_of word_of_flip_bit word_of_xor word_of_push_bit word_of_1 flip_bit_eq_xor) + +end + +locale word_type_copy_more = word_type_copy_bits + + constrains word_of :: \'a::{ring_bit_operations, equal, linorder} \ 'b::len word\ + fixes of_nat :: \nat \ 'a\ and nat_of :: \'a \ nat\ + and of_int :: \int \ 'a\ and int_of :: \'a \ int\ + and of_integer :: \integer \ 'a\ and integer_of :: \'a \ integer\ + assumes word_of_nat_eq: \word_of (of_nat n) = word_of_nat n\ + and nat_of_eq_word_of: \nat_of p = unat (word_of p)\ + and word_of_int_eq: \word_of (of_int k) = word_of_int k\ + and int_of_eq_word_of: \int_of p = uint (word_of p)\ + and word_of_integer_eq: \word_of (of_integer l) = word_of_integer l\ + and integer_of_eq_word_of: \integer_of p = unsigned (word_of p)\ +begin + +lemma of_word_numeral [code_post]: + \of_word (numeral n) = numeral n\ + by (simp add: eq_iff_word_of word_of_word word_of_numeral) + +lemma of_word_0 [code_post]: + \of_word 0 = 0\ + by (simp add: eq_iff_word_of word_of_0 word_of_word) + +lemma of_word_1 [code_post]: + \of_word 1 = 1\ + by (simp add: eq_iff_word_of word_of_1 word_of_word) + +text \Use pretty numerals from integer for pretty printing\ + +lemma numeral_eq_integer [code_unfold]: + \numeral n = of_integer (numeral n)\ + by (simp add: eq_iff_word_of word_of_integer_eq word_of_numeral word_of_integer.rep_eq) + +lemma neg_numeral_eq_integer [code_unfold]: + \- numeral n = of_integer (- numeral n)\ + by (simp add: eq_iff_word_of word_of_integer_eq word_of_minus word_of_numeral word_of_integer.rep_eq) + +end + +locale word_type_copy_misc = word_type_copy_more + opening constraintless + + constrains word_of :: \'a::{ring_bit_operations, equal, linorder} \ 'b::len word\ + fixes size :: nat and set_bits_aux :: \(nat \ bool) \ nat \ 'a \ 'a\ + assumes size_eq_length: \size = LENGTH('b::len)\ + and msb_iff_word_of [code]: \msb p \ msb (word_of p)\ + and lsb_iff_word_of [code]: \lsb p \ lsb (word_of p)\ + and size_eq_word_of: \Nat.size (p :: 'a) = Nat.size (word_of p)\ + and word_of_generic_set_bit [code]: \word_of (Generic_set_bit.set_bit p n b) = Generic_set_bit.set_bit (word_of p) n b\ + and word_of_set_bits: \word_of (set_bits P) = set_bits P\ + and word_of_set_bits_aux: \word_of (set_bits_aux P n p) = Code_Target_Word_Base.set_bits_aux P n (word_of p)\ +begin + +lemma size_eq [code]: + \Nat.size p = size\ for p :: 'a + by (simp add: size_eq_length size_eq_word_of word_size) + +lemma set_bits_aux_code [code]: + \set_bits_aux f n w = + (if n = 0 then w + else let n' = n - 1 in set_bits_aux f n' (push_bit 1 w OR (if f n' then 1 else 0)))\ + by (simp add: eq_iff_word_of word_of_set_bits_aux Let_def word_of_mult word_of_or word_of_0 word_of_1 set_bits_aux_rec [of f n]) + +lemma set_bits_code [code]: \set_bits P = set_bits_aux P size 0\ + by (simp add: fun_eq_iff eq_iff_word_of word_of_set_bits word_of_set_bits_aux word_of_0 size_eq_length set_bits_conv_set_bits_aux) + +lemma of_class_lsb: + \OFCLASS('a, lsb_class)\ + by standard (simp add: fun_eq_iff lsb_iff_word_of even_iff_word_of lsb_odd) + +lemma of_class_set_bit: + \OFCLASS('a, set_bit_class)\ + by standard (simp add: eq_iff_word_of word_of_generic_set_bit bit_eq_word_of word_of_power word_of_numeral word_of_0 bit_simps) + +lemma of_class_bit_comprehension: + \OFCLASS('a, bit_comprehension_class)\ + by standard (simp add: eq_iff_word_of word_of_set_bits bit_eq_word_of set_bits_bit_eq) + +end + +end diff --git a/thys/PAC_Checker/PAC_Checker_Relation.thy b/thys/PAC_Checker/PAC_Checker_Relation.thy --- a/thys/PAC_Checker/PAC_Checker_Relation.thy +++ b/thys/PAC_Checker/PAC_Checker_Relation.thy @@ -1,387 +1,382 @@ (* File: PAC_Checker_Relation.thy Author: Mathias Fleury, Daniela Kaufmann, JKU Maintainer: Mathias Fleury, JKU *) theory PAC_Checker_Relation imports PAC_Checker WB_Sort "Native_Word.Uint64" begin section \Various Refinement Relations\ text \When writing this, it was not possible to share the definition with the IsaSAT version.\ definition uint64_nat_rel :: "(uint64 \ nat) set" where \uint64_nat_rel = br nat_of_uint64 (\_. True)\ abbreviation uint64_nat_assn where \uint64_nat_assn \ pure uint64_nat_rel\ instantiation uint32 :: hashable begin definition hashcode_uint32 :: \uint32 \ uint32\ where \hashcode_uint32 n = n\ definition def_hashmap_size_uint32 :: \uint32 itself \ nat\ where \def_hashmap_size_uint32 = (\_. 16)\ \ \same as @{typ nat}\ instance by standard (simp add: def_hashmap_size_uint32_def) end instantiation uint64 :: hashable begin definition hashcode_uint64 :: \uint64 \ uint32\ where \hashcode_uint64 n = (uint32_of_nat (nat_of_uint64 ((n) AND ((2 :: uint64)^32 -1))))\ definition def_hashmap_size_uint64 :: \uint64 itself \ nat\ where \def_hashmap_size_uint64 = (\_. 16)\ \ \same as @{typ nat}\ instance by standard (simp add: def_hashmap_size_uint64_def) end lemma word_nat_of_uint64_Rep_inject[simp]: \nat_of_uint64 ai = nat_of_uint64 bi \ ai = bi\ by transfer (simp add: word_unat_eq_iff) instance uint64 :: heap by standard (auto simp: inj_def exI[of _ nat_of_uint64]) instance uint64 :: semiring_numeral by standard lemma nat_of_uint64_012[simp]: \nat_of_uint64 0 = 0\ \nat_of_uint64 2 = 2\ \nat_of_uint64 1 = 1\ - by (transfer, auto)+ + by (simp_all add: nat_of_uint64.rep_eq zero_uint64.rep_eq one_uint64.rep_eq) definition uint64_of_nat_conv where [simp]: \uint64_of_nat_conv (x :: nat) = x\ lemma less_upper_bintrunc_id: \n < 2 ^b \ n \ 0 \ take_bit b n = n\ for n :: int by (rule take_bit_int_eq_self) lemma nat_of_uint64_uint64_of_nat_id: \n < 2^64 \ nat_of_uint64 (uint64_of_nat n) = n\ - unfolding uint64_of_nat_def - apply simp - apply transfer - apply (subst unat_eq_nat_uint) - apply transfer - by (auto simp: less_upper_bintrunc_id) + by transfer (simp add: take_bit_nat_eq_self) lemma [sepref_fr_rules]: \(return o uint64_of_nat, RETURN o uint64_of_nat_conv) \ [\a. a < 2 ^64]\<^sub>a nat_assn\<^sup>k \ uint64_nat_assn\ by sepref_to_hoare (sep_auto simp: uint64_nat_rel_def br_def nat_of_uint64_uint64_of_nat_id) definition string_rel :: \(String.literal \ string) set\ where \string_rel = {(x, y). y = String.explode x}\ abbreviation string_assn :: \string \ String.literal \ assn\ where \string_assn \ pure string_rel\ lemma eq_string_eq: \((=), (=)) \ string_rel \ string_rel \ bool_rel\ by (auto intro!: frefI simp: string_rel_def String.less_literal_def less_than_char_def rel2p_def literal.explode_inject) lemmas eq_string_eq_hnr = eq_string_eq[sepref_import_param] definition string2_rel :: \(string \ string) set\ where \string2_rel \ \Id\list_rel\ abbreviation string2_assn :: \string \ string \ assn\ where \string2_assn \ pure string2_rel\ abbreviation monom_rel where \monom_rel \ \string_rel\list_rel\ abbreviation monom_assn where \monom_assn \ list_assn string_assn\ abbreviation monomial_rel where \monomial_rel \ monom_rel \\<^sub>r int_rel\ abbreviation monomial_assn where \monomial_assn \ monom_assn \\<^sub>a int_assn\ abbreviation poly_rel where \poly_rel \ \monomial_rel\list_rel\ abbreviation poly_assn where \poly_assn \ list_assn monomial_assn\ lemma poly_assn_alt_def: \poly_assn = pure poly_rel\ by (simp add: list_assn_pure_conv) abbreviation polys_assn where \polys_assn \ hm_fmap_assn uint64_nat_assn poly_assn\ lemma string_rel_string_assn: \(\ ((c, a) \ string_rel)) = string_assn a c\ by (auto simp: pure_app_eq) lemma single_valued_string_rel: \single_valued string_rel\ by (auto simp: single_valued_def string_rel_def) lemma IS_LEFT_UNIQUE_string_rel: \IS_LEFT_UNIQUE string_rel\ by (auto simp: IS_LEFT_UNIQUE_def single_valued_def string_rel_def literal.explode_inject) lemma IS_RIGHT_UNIQUE_string_rel: \IS_RIGHT_UNIQUE string_rel\ by (auto simp: single_valued_def string_rel_def literal.explode_inject) lemma single_valued_monom_rel: \single_valued monom_rel\ by (rule list_rel_sv) (auto intro!: frefI simp: string_rel_def rel2p_def single_valued_def p2rel_def) lemma single_valued_monomial_rel: \single_valued monomial_rel\ using single_valued_monom_rel by (auto intro!: frefI simp: rel2p_def single_valued_def p2rel_def) lemma single_valued_monom_rel': \IS_LEFT_UNIQUE monom_rel\ unfolding IS_LEFT_UNIQUE_def inv_list_rel_eq string2_rel_def by (rule list_rel_sv)+ (auto intro!: frefI simp: string_rel_def rel2p_def single_valued_def p2rel_def literal.explode_inject) lemma single_valued_monomial_rel': \IS_LEFT_UNIQUE monomial_rel\ using single_valued_monom_rel' unfolding IS_LEFT_UNIQUE_def inv_list_rel_eq by (auto intro!: frefI simp: rel2p_def single_valued_def p2rel_def) lemma [safe_constraint_rules]: \Sepref_Constraints.CONSTRAINT single_valued string_rel\ \Sepref_Constraints.CONSTRAINT IS_LEFT_UNIQUE string_rel\ by (auto simp: CONSTRAINT_def single_valued_def string_rel_def IS_LEFT_UNIQUE_def literal.explode_inject) lemma eq_string_monom_hnr[sepref_fr_rules]: \(uncurry (return oo (=)), uncurry (RETURN oo (=))) \ monom_assn\<^sup>k *\<^sub>a monom_assn\<^sup>k \\<^sub>a bool_assn\ using single_valued_monom_rel' single_valued_monom_rel unfolding list_assn_pure_conv by sepref_to_hoare (sep_auto simp: list_assn_pure_conv string_rel_string_assn single_valued_def IS_LEFT_UNIQUE_def dest!: mod_starD simp flip: inv_list_rel_eq) definition term_order_rel' where [simp]: \term_order_rel' x y = ((x, y) \ term_order_rel)\ lemma term_order_rel[def_pat_rules]: \(\)$(x,y)$term_order_rel \ term_order_rel'$x$y\ by auto lemma term_order_rel_alt_def: \term_order_rel = lexord (p2rel char.lexordp)\ by (auto simp: p2rel_def char.lexordp_conv_lexord var_order_rel_def intro!: arg_cong[of _ _ lexord]) instantiation char :: linorder begin definition less_char where [symmetric, simp]: "less_char = PAC_Polynomials_Term.less_char" definition less_eq_char where [symmetric, simp]: "less_eq_char = PAC_Polynomials_Term.less_eq_char" instance apply standard using char.linorder_axioms by (auto simp: class.linorder_def class.order_def class.preorder_def less_eq_char_def less_than_char_def class.order_axioms_def class.linorder_axioms_def p2rel_def less_char_def) end instantiation list :: (linorder) linorder begin definition less_list where "less_list = lexordp (<)" definition less_eq_list where "less_eq_list = lexordp_eq" instance proof standard have [dest]: \\x y :: 'a :: linorder list. (x, y) \ lexord {(x, y). x < y} \ lexordp_eq y x \ False\ by (metis lexordp_antisym lexordp_conv_lexord lexordp_eq_conv_lexord) have [simp]: \\x y :: 'a :: linorder list. lexordp_eq x y \ \ lexordp_eq y x \ (x, y) \ lexord {(x, y). x < y}\ using lexordp_conv_lexord lexordp_conv_lexordp_eq by blast show \(x < y) = Restricted_Predicates.strict (\) x y\ \x \ x\ \x \ y \ y \ z \ x \ z\ \x \ y \ y \ x \ x = y\ \x \ y \ y \ x\ for x y z :: \'a :: linorder list\ by (auto simp: less_list_def less_eq_list_def List.lexordp_def lexordp_conv_lexord lexordp_into_lexordp_eq lexordp_antisym antisym_def lexordp_eq_refl lexordp_eq_linear intro: lexordp_eq_trans dest: lexordp_eq_antisym) qed end lemma term_order_rel'_alt_def_lexord: \term_order_rel' x y = ord_class.lexordp x y\ and term_order_rel'_alt_def: \term_order_rel' x y \ x < y\ proof - show \term_order_rel' x y = ord_class.lexordp x y\ \term_order_rel' x y \ x < y\ unfolding less_than_char_of_char[symmetric, abs_def] by (auto simp: lexordp_conv_lexord less_eq_list_def less_list_def lexordp_def var_order_rel_def rel2p_def term_order_rel_alt_def p2rel_def) qed lemma list_rel_list_rel_order_iff: assumes \(a, b) \ \string_rel\list_rel\ \(a', b') \ \string_rel\list_rel\ shows \a < a' \ b < b'\ proof have H: \(a, b) \ \string_rel\list_rel \ (a, cs) \ \string_rel\list_rel \ b = cs\ for cs using single_valued_monom_rel' IS_RIGHT_UNIQUE_string_rel unfolding string2_rel_def by (subst (asm)list_rel_sv_iff[symmetric]) (auto simp: single_valued_def) assume \a < a'\ then consider u u' where \a' = a @ u # u'\ | u aa v w aaa where \a = u @ aa # v\ \a' = u @ aaa # w\ \aa < aaa\ by (subst (asm) less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) then show \b < b'\ proof cases case 1 then show \b < b'\ using assms by (subst less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff dest: H) next case 2 then obtain u' aa' v' w' aaa' where \b = u' @ aa' # v'\ \b' = u' @ aaa' # w'\ \(aa, aa') \ string_rel\ \(aaa, aaa') \ string_rel\ using assms by (smt list_rel_append1 list_rel_split_right_iff single_valued_def single_valued_monom_rel) with \aa < aaa\ have \aa' < aaa'\ by (auto simp: string_rel_def less_literal.rep_eq less_list_def lexordp_conv_lexord lexordp_def char.lexordp_conv_lexord simp flip: less_char_def PAC_Polynomials_Term.less_char_def) then show \b < b'\ using \b = u' @ aa' # v'\ \b' = u' @ aaa' # w'\ by (subst less_list_def) (fastforce simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) qed next have H: \(a, b) \ \string_rel\list_rel \ (a', b) \ \string_rel\list_rel \ a = a'\ for a a' b using single_valued_monom_rel' by (auto simp: single_valued_def IS_LEFT_UNIQUE_def simp flip: inv_list_rel_eq) assume \b < b'\ then consider u u' where \b' = b @ u # u'\ | u aa v w aaa where \b = u @ aa # v\ \b' = u @ aaa # w\ \aa < aaa\ by (subst (asm) less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) then show \a < a'\ proof cases case 1 then show \a < a'\ using assms by (subst less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append2 list_rel_split_left_iff dest: H) next case 2 then obtain u' aa' v' w' aaa' where \a = u' @ aa' # v'\ \a' = u' @ aaa' # w'\ \(aa', aa) \ string_rel\ \(aaa', aaa) \ string_rel\ using assms by (auto simp: lexord_def List.lexordp_def list_rel_append2 list_rel_split_left_iff dest: H) with \aa < aaa\ have \aa' < aaa'\ by (auto simp: string_rel_def less_literal.rep_eq less_list_def lexordp_conv_lexord lexordp_def char.lexordp_conv_lexord simp flip: less_char_def PAC_Polynomials_Term.less_char_def) then show \a < a'\ using \a = u' @ aa' # v'\ \a' = u' @ aaa' # w'\ by (subst less_list_def) (fastforce simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) qed qed lemma string_rel_le[sepref_import_param]: shows \((<), (<)) \ \string_rel\list_rel \ \string_rel\list_rel \ bool_rel\ by (auto intro!: fun_relI simp: list_rel_list_rel_order_iff) (* TODO Move *) lemma [sepref_import_param]: assumes \CONSTRAINT IS_LEFT_UNIQUE R\ \CONSTRAINT IS_RIGHT_UNIQUE R\ shows \(remove1, remove1) \ R \ \R\list_rel \ \R\list_rel\ apply (intro fun_relI) subgoal premises p for x y xs ys using p(2) p(1) assms by (induction xs ys rule: list_rel_induct) (auto simp: IS_LEFT_UNIQUE_def single_valued_def) done instantiation pac_step :: (heap, heap, heap) heap begin instance proof standard obtain f :: \'a \ nat\ where f: \inj f\ by blast obtain g :: \nat \ nat \ nat \ nat \ nat \ nat\ where g: \inj g\ by blast obtain h :: \'b \ nat\ where h: \inj h\ by blast obtain i :: \'c \ nat\ where i: \inj i\ by blast have [iff]: \g a = g b \ a = b\\h a'' = h b'' \ a'' = b''\ \f a' = f b' \ a' = b'\ \i a''' = i b''' \ a''' = b'''\ for a b a' b' a'' b'' a''' b''' using f g h i unfolding inj_def by blast+ let ?f = \\x :: ('a, 'b, 'c) pac_step. g (case x of Add a b c d \ (0, i a, i b, i c, f d) | Del a \ (1, i a, 0, 0, 0) | Mult a b c d \ (2, i a, f b, i c, f d) | Extension a b c \ (3, i a, f c, 0, h b))\ have \inj ?f\ apply (auto simp: inj_def) apply (case_tac x; case_tac y) apply auto done then show \\f :: ('a, 'b, 'c) pac_step \ nat. inj f\ by blast qed end end \ No newline at end of file