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,1651 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) subsection \Finite Fields\ text \We provide four implementations for $GF(p)$ -- the field with $p$ elements for some prime $p$ -- one by int, one by integers, one by 32-bit numbers and one 64-bit implementation. Correctness of the implementations is proven by transfer rules to the type-based version of $GF(p)$.\ theory Finite_Field_Record_Based imports Finite_Field Arithmetic_Record_Based Native_Word.Uint32 Native_Word.Uint64 Native_Word.Code_Target_Bits_Int "HOL-Library.Code_Target_Numeral" begin (* mod on standard case which can immediately be mapped to target languages without considering special cases *) definition mod_nonneg_pos :: "integer \ integer \ integer" where "x \ 0 \ y > 0 \ mod_nonneg_pos x y = (x mod y)" code_printing \ \FIXME illusion of partiality\ constant mod_nonneg_pos \ (SML) "IntInf.mod/ ( _,/ _ )" and (Eval) "IntInf.mod/ ( _,/ _ )" and (OCaml) "Z.rem" and (Haskell) "Prelude.mod/ ( _ )/ ( _ )" and (Scala) "!((k: BigInt) => (l: BigInt) =>/ (k '% l))" definition mod_nonneg_pos_int :: "int \ int \ int" where "mod_nonneg_pos_int x y = int_of_integer (mod_nonneg_pos (integer_of_int x) (integer_of_int y))" lemma mod_nonneg_pos_int[simp]: "x \ 0 \ y > 0 \ mod_nonneg_pos_int x y = (x mod y)" unfolding mod_nonneg_pos_int_def using mod_nonneg_pos_def by simp context fixes p :: int begin definition plus_p :: "int \ int \ int" where "plus_p x y \ let z = x + y in if z \ p then z - p else z" definition minus_p :: "int \ int \ int" where "minus_p x y \ if y \ x then x - y else x + p - y" definition uminus_p :: "int \ int" where "uminus_p x = (if x = 0 then 0 else p - x)" definition mult_p :: "int \ int \ int" where "mult_p x y = (mod_nonneg_pos_int (x * y) p)" fun power_p :: "int \ nat \ int" where "power_p x n = (if n = 0 then 1 else let (d,r) = Divides.divmod_nat n 2; rec = power_p (mult_p x x) d in if r = 0 then rec else mult_p rec x)" text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p :: "int \ int" where "inverse_p x = (if x = 0 then 0 else power_p x (nat (p - 2)))" definition divide_p :: "int \ int \ int" where "divide_p x y = mult_p x (inverse_p y)" definition finite_field_ops_int :: "int arith_ops_record" where "finite_field_ops_int \ Arith_Ops_Record 0 1 plus_p mult_p minus_p uminus_p divide_p inverse_p (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) (\ x . x) (\ x . x) (\ x. 0 \ x \ x < p)" end context fixes p :: uint32 begin definition plus_p32 :: "uint32 \ uint32 \ uint32" where "plus_p32 x y \ let z = x + y in if z \ p then z - p else z" definition minus_p32 :: "uint32 \ uint32 \ uint32" where "minus_p32 x y \ if y \ x then x - y else (x + p) - y" definition uminus_p32 :: "uint32 \ uint32" where "uminus_p32 x = (if x = 0 then 0 else p - x)" definition mult_p32 :: "uint32 \ uint32 \ uint32" where "mult_p32 x y = (x * y mod p)" lemma int_of_uint32_shift: "int_of_uint32 (shiftr n k) = (int_of_uint32 n) div (2 ^ k)" 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) (shiftr n 1) in if n AND 1 = 0 then rec else mult_p32 rec x)" by pat_completeness auto termination proof - { fix n :: uint32 assume "n \ 0" with int_of_uint32_ge_0[of n] int_of_uint32_0_iff[of n] have "int_of_uint32 n > 0" by auto hence "0 < int_of_uint32 n" "int_of_uint32 n div 2 < int_of_uint32 n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_uint32 n))", auto simp: int_of_uint32_shift *) qed text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p32 :: "uint32 \ uint32" where "inverse_p32 x = (if x = 0 then 0 else power_p32 x (p - 2))" definition divide_p32 :: "uint32 \ uint32 \ uint32" where "divide_p32 x y = mult_p32 x (inverse_p32 y)" definition finite_field_ops32 :: "uint32 arith_ops_record" where "finite_field_ops32 \ Arith_Ops_Record 0 1 plus_p32 mult_p32 minus_p32 uminus_p32 divide_p32 inverse_p32 (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) uint32_of_int int_of_uint32 (\ x. 0 \ x \ x < p)" end lemma shiftr_uint32_code [code_unfold]: "drop_bit 1 x = (uint32_shiftr x 1)" by (simp add: uint32_shiftr_def shiftr_eq_drop_bit) (* ******************************************************************************** *) subsubsection \Transfer Relation\ locale mod_ring_locale = fixes p :: int and ty :: "'a :: nontriv itself" assumes p: "p = int CARD('a)" begin lemma nat_p: "nat p = CARD('a)" unfolding p by simp lemma p2: "p \ 2" unfolding p using nontriv[where 'a = 'a] by auto lemma p2_ident: "int (CARD('a) - 2) = p - 2" using p2 unfolding p by simp definition mod_ring_rel :: "int \ 'a mod_ring \ bool" where "mod_ring_rel x x' = (x = to_int_mod_ring x')" (* domain transfer rules *) lemma Domainp_mod_ring_rel [transfer_domain_rule]: "Domainp (mod_ring_rel) = (\ v. v \ {0 ..< p})" proof - { fix v :: int assume *: "0 \ v" "v < p" have "Domainp mod_ring_rel v" proof show "mod_ring_rel v (of_int_mod_ring v)" unfolding mod_ring_rel_def using * p by auto qed } note * = this show ?thesis by (intro ext iffI, insert range_to_int_mod_ring[where 'a = 'a] *, auto simp: mod_ring_rel_def p) qed (* left/right/bi-unique *) lemma bi_unique_mod_ring_rel [transfer_rule]: "bi_unique mod_ring_rel" "left_unique mod_ring_rel" "right_unique mod_ring_rel" unfolding mod_ring_rel_def bi_unique_def left_unique_def right_unique_def by auto (* left/right-total *) lemma right_total_mod_ring_rel [transfer_rule]: "right_total mod_ring_rel" unfolding mod_ring_rel_def right_total_def by simp (* ************************************************************************************ *) subsubsection \Transfer Rules\ (* 0 / 1 *) lemma mod_ring_0[transfer_rule]: "mod_ring_rel 0 0" unfolding mod_ring_rel_def by simp lemma mod_ring_1[transfer_rule]: "mod_ring_rel 1 1" unfolding mod_ring_rel_def by simp (* addition *) lemma plus_p_mod_def: assumes x: "x \ {0 ..< p}" and y: "y \ {0 ..< p}" shows "plus_p p x y = ((x + y) mod p)" proof (cases "p \ x + y") case False thus ?thesis using x y unfolding plus_p_def Let_def by auto next case True from True x y have *: "p > 0" "0 \ x + y - p" "x + y - p < p" by auto from True have id: "plus_p p x y = x + y - p" unfolding plus_p_def by auto show ?thesis unfolding id using * using mod_pos_pos_trivial by fastforce qed lemma mod_ring_plus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (plus_p p) (+)" proof - { fix x y :: "'a mod_ring" have "plus_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x + y)" by (transfer, subst plus_p_mod_def, auto, auto simp: p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* subtraction *) lemma minus_p_mod_def: assumes x: "x \ {0 ..< p}" and y: "y \ {0 ..< p}" shows "minus_p p x y = ((x - y) mod p)" proof (cases "x - y < 0") case False thus ?thesis using x y unfolding minus_p_def Let_def by auto next case True from True x y have *: "p > 0" "0 \ x - y + p" "x - y + p < p" by auto from True have id: "minus_p p x y = x - y + p" unfolding minus_p_def by auto show ?thesis unfolding id using * using mod_pos_pos_trivial by fastforce qed lemma mod_ring_minus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (minus_p p) (-)" proof - { fix x y :: "'a mod_ring" have "minus_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x - y)" by (transfer, subst minus_p_mod_def, auto simp: p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* unary minus *) lemma mod_ring_uminus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (uminus_p p) uminus" proof - { fix x :: "'a mod_ring" have "uminus_p p (to_int_mod_ring x) = to_int_mod_ring (uminus x)" by (transfer, auto simp: uminus_p_def p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* multiplication *) lemma mod_ring_mult[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (mult_p p) ((*))" proof - { fix x y :: "'a mod_ring" have "mult_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x * y)" by (transfer, auto simp: mult_p_def p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* equality *) lemma mod_ring_eq[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> (=)) (=) (=)" by (intro rel_funI, auto simp: mod_ring_rel_def) (* power *) lemma mod_ring_power[transfer_rule]: "(mod_ring_rel ===> (=) ===> mod_ring_rel) (power_p p) (^)" proof (intro rel_funI, clarify, unfold binary_power[symmetric], goal_cases) fix x y n assume xy: "mod_ring_rel x y" from xy show "mod_ring_rel (power_p p x n) (binary_power y n)" proof (induct y n arbitrary: x rule: binary_power.induct) case (1 x n y) note 1(2)[transfer_rule] show ?case proof (cases "n = 0") case True thus ?thesis by (simp add: mod_ring_1) next case False obtain d r where id: "Divides.divmod_nat n 2 = (d,r)" by force let ?int = "power_p p (mult_p p y y) d" let ?gfp = "binary_power (x * x) d" from False have id': "?thesis = (mod_ring_rel (if r = 0 then ?int else mult_p p ?int y) (if r = 0 then ?gfp else ?gfp * x))" unfolding power_p.simps[of _ _ n] binary_power.simps[of _ n] Let_def id split by simp have [transfer_rule]: "mod_ring_rel ?int ?gfp" by (rule 1(1)[OF False refl id[symmetric]], transfer_prover) show ?thesis unfolding id' by transfer_prover qed qed qed declare power_p.simps[simp del] lemma ring_finite_field_ops_int: "ring_ops (finite_field_ops_int p) mod_ring_rel" by (unfold_locales, auto simp: finite_field_ops_int_def bi_unique_mod_ring_rel right_total_mod_ring_rel mod_ring_plus mod_ring_minus mod_ring_uminus mod_ring_mult mod_ring_eq mod_ring_0 mod_ring_1 Domainp_mod_ring_rel) end locale prime_field = mod_ring_locale p ty for p and ty :: "'a :: prime_card itself" begin lemma prime: "prime p" unfolding p using prime_card[where 'a = 'a] by simp (* mod *) lemma mod_ring_mod[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) ((\ x y. if y = 0 then x else 0)) (mod)" proof - { fix x y :: "'a mod_ring" have "(if to_int_mod_ring y = 0 then to_int_mod_ring x else 0) = to_int_mod_ring (x mod y)" unfolding modulo_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* normalize *) lemma mod_ring_normalize[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) ((\ x. if x = 0 then 0 else 1)) normalize" proof - { fix x :: "'a mod_ring" have "(if to_int_mod_ring x = 0 then 0 else 1) = to_int_mod_ring (normalize x)" unfolding normalize_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* unit_factor *) lemma mod_ring_unit_factor[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (\ x. x) unit_factor" proof - { fix x :: "'a mod_ring" have "to_int_mod_ring x = to_int_mod_ring (unit_factor x)" unfolding unit_factor_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* inverse *) lemma mod_ring_inverse[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (inverse_p p) inverse" proof (intro rel_funI) fix x y assume [transfer_rule]: "mod_ring_rel x y" show "mod_ring_rel (inverse_p p x) (inverse y)" unfolding inverse_p_def inverse_mod_ring_def apply (transfer_prover_start) apply (transfer_step)+ apply (unfold p2_ident) apply (rule refl) done qed (* division *) lemma mod_ring_divide[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (divide_p p) (/)" unfolding divide_p_def[abs_def] divide_mod_ring_def[abs_def] inverse_mod_ring_def[symmetric] by transfer_prover lemma mod_ring_rel_unsafe: assumes "x < CARD('a)" shows "mod_ring_rel (int x) (of_nat x)" "0 < x \ of_nat x \ (0 :: 'a mod_ring)" proof - have id: "of_nat x = (of_int (int x) :: 'a mod_ring)" by simp show "mod_ring_rel (int x) (of_nat x)" "0 < x \ of_nat x \ (0 :: 'a mod_ring)" unfolding id unfolding mod_ring_rel_def proof (auto simp add: assms of_int_of_int_mod_ring) assume "0 < x" with assms have "of_int_mod_ring (int x) \ (0 :: 'a mod_ring)" by (metis (no_types) less_imp_of_nat_less less_irrefl of_nat_0_le_iff of_nat_0_less_iff to_int_mod_ring_hom.hom_zero to_int_mod_ring_of_int_mod_ring) thus "of_int_mod_ring (int x) = (0 :: 'a mod_ring) \ False" by blast qed qed lemma finite_field_ops_int: "field_ops (finite_field_ops_int p) mod_ring_rel" by (unfold_locales, auto simp: finite_field_ops_int_def bi_unique_mod_ring_rel right_total_mod_ring_rel mod_ring_divide mod_ring_plus mod_ring_minus mod_ring_uminus mod_ring_inverse mod_ring_mod mod_ring_unit_factor mod_ring_normalize mod_ring_mult mod_ring_eq mod_ring_0 mod_ring_1 Domainp_mod_ring_rel) end text \Once we have proven the soundness of the implementation, we do not care any longer that @{typ "'a mod_ring"} has been defined internally via lifting. Disabling the transfer-rules will hide the internal definition in further applications of transfer.\ lifting_forget mod_ring.lifting text \For soundness of the 32-bit implementation, we mainly prove that this implementation implements the int-based implementation of the mod-ring.\ context mod_ring_locale begin context fixes pp :: "uint32" assumes ppp: "p = int_of_uint32 pp" and small: "p \ 65535" begin lemmas uint32_simps = int_of_uint32_0 int_of_uint32_plus int_of_uint32_minus int_of_uint32_mult definition urel32 :: "uint32 \ int \ bool" where "urel32 x y = (y = int_of_uint32 x \ y < p)" definition mod_ring_rel32 :: "uint32 \ 'a mod_ring \ bool" where "mod_ring_rel32 x y = (\ z. urel32 x z \ mod_ring_rel z y)" lemma urel32_0: "urel32 0 0" unfolding urel32_def using p2 by (simp, transfer, simp) lemma urel32_1: "urel32 1 1" unfolding urel32_def using p2 by (simp, transfer, simp) lemma le_int_of_uint32: "(x \ y) = (int_of_uint32 x \ int_of_uint32 y)" by (transfer, simp add: word_le_def) lemma urel32_plus: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (plus_p32 pp x x') (plus_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" let ?p = "int_of_uint32 pp" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel32_def by auto have le: "(pp \ x + x') = (?p \ ?x + ?x')" unfolding le_int_of_uint32 using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?p \ ?x + ?x'") case True hence True: "(?p \ ?x + ?x') = True" by simp show ?thesis unfolding id using small rel unfolding plus_p32_def plus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?p \ ?x + ?x') = False" by simp show ?thesis unfolding id using small rel unfolding plus_p32_def plus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_minus: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (minus_p32 pp x x') (minus_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel32_def by auto have le: "(x' \ x) = (?x' \ ?x)" unfolding le_int_of_uint32 using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?x' \ ?x") case True hence True: "(?x' \ ?x) = True" by simp show ?thesis unfolding id using small rel unfolding minus_p32_def minus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?x' \ ?x) = False" by simp show ?thesis unfolding id using small rel unfolding minus_p32_def minus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_uminus: assumes "urel32 x y" shows "urel32 (uminus_p32 pp x) (uminus_p p y)" proof - let ?x = "int_of_uint32 x" from assms int_of_uint32_ge_0 have id: "y = ?x" and rel: "0 \ ?x" "?x < p" unfolding urel32_def by auto have le: "(x = 0) = (?x = 0)" unfolding int_of_uint32_0_iff using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?x = 0") case True hence True: "(?x = 0) = True" by simp show ?thesis unfolding id using small rel unfolding uminus_p32_def uminus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?x = 0) = False" by simp show ?thesis unfolding id using small rel unfolding uminus_p32_def uminus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_mult: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (mult_p32 pp x x') (mult_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' < p" unfolding urel32_def by auto from rel have "?x * ?x' < p * p" by (metis mult_strict_mono') also have "\ \ 65536 * 65536" by (rule mult_mono, insert p2 small, auto) finally have le: "?x * ?x' < 4294967296" by simp show ?thesis unfolding id using small rel unfolding mult_p32_def mult_p_def Let_def urel32_def unfolding ppp by (auto simp: uint32_simps, unfold int_of_uint32_mod int_of_uint32_mult, subst mod_pos_pos_trivial[of _ 4294967296], insert le, auto) qed lemma urel32_eq: assumes "urel32 x y" "urel32 x' y'" shows "(x = x') = (y = y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" unfolding urel32_def by auto - show ?thesis unfolding id by (transfer, auto) + 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 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 done from urel32_eq[OF this urel32_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel32 (shiftr y 1) (int d')" unfolding d' using y unfolding urel32_def using small unfolding ppp apply transfer apply transfer apply (auto simp add: drop_bit_Suc) 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) show ?thesis unfolding inverse_p32_def inverse_p_def urel32_eq[OF x urel32_0] using urel32_0 urel32_power[OF x p] by auto qed lemma mod_ring_0_32: "mod_ring_rel32 0 0" using urel32_0 mod_ring_0 unfolding mod_ring_rel32_def by blast lemma mod_ring_1_32: "mod_ring_rel32 1 1" using urel32_1 mod_ring_1 unfolding mod_ring_rel32_def by blast lemma mod_ring_uminus32: "(mod_ring_rel32 ===> mod_ring_rel32) (uminus_p32 pp) uminus" using urel32_uminus mod_ring_uminus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_plus32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (plus_p32 pp) (+)" using urel32_plus mod_ring_plus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_minus32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (minus_p32 pp) (-)" using urel32_minus mod_ring_minus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_mult32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (mult_p32 pp) ((*))" using urel32_mult mod_ring_mult unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_eq32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> (=)) (=) (=)" using urel32_eq mod_ring_eq unfolding mod_ring_rel32_def rel_fun_def by blast lemma urel32_inj: "urel32 x y \ urel32 x z \ y = z" using urel32_eq[of x y x z] by auto lemma urel32_inj': "urel32 x z \ urel32 y z \ x = y" using urel32_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel32: "bi_unique mod_ring_rel32" "left_unique mod_ring_rel32" "right_unique mod_ring_rel32" using bi_unique_mod_ring_rel urel32_inj' unfolding mod_ring_rel32_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel32_def) lemma right_total_mod_ring_rel32: "right_total mod_ring_rel32" unfolding mod_ring_rel32_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel32 (uint32_of_int z) z" unfolding urel32_def using small unfolding ppp by (auto simp: int_of_uint32_inv) with zy show "\ x z. urel32 x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel32: "Domainp mod_ring_rel32 = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel32 x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel32_def proof let ?i = "int_of_uint32" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" using small unfolding ppp by (transfer, auto simp: word_less_def) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel32 a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel32 x (?i x)" unfolding urel32_def using small * unfolding ppp by (transfer, auto simp: word_less_def) qed next assume "\a b. x = a \ (\z. urel32 a z \ mod_ring_rel z b)" then obtain b z where xz: "urel32 x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel32_def using small unfolding ppp by (transfer, auto simp: word_less_def) qed qed lemma ring_finite_field_ops32: "ring_ops (finite_field_ops32 pp) mod_ring_rel32" by (unfold_locales, auto simp: finite_field_ops32_def bi_unique_mod_ring_rel32 right_total_mod_ring_rel32 mod_ring_plus32 mod_ring_minus32 mod_ring_uminus32 mod_ring_mult32 mod_ring_eq32 mod_ring_0_32 mod_ring_1_32 Domainp_mod_ring_rel32) end end context prime_field begin context fixes pp :: "uint32" assumes *: "p = int_of_uint32 pp" "p \ 65535" begin lemma mod_ring_normalize32: "(mod_ring_rel32 ===> mod_ring_rel32) (\x. if x = 0 then 0 else 1) normalize" using urel32_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_mod32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (\x y. if y = 0 then x else 0) (mod)" using urel32_mod[OF *] mod_ring_mod unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor32: "(mod_ring_rel32 ===> mod_ring_rel32) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_inverse32: "(mod_ring_rel32 ===> mod_ring_rel32) (inverse_p32 pp) inverse" using urel32_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_divide32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (divide_p32 pp) (/)" using mod_ring_inverse32 mod_ring_mult32[OF *] unfolding divide_p32_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops32: "field_ops (finite_field_ops32 pp) mod_ring_rel32" by (unfold_locales, insert ring_finite_field_ops32[OF *], auto simp: ring_ops_def finite_field_ops32_def mod_ring_divide32 mod_ring_inverse32 mod_ring_mod32 mod_ring_normalize32) end end (* now there is 64-bit time *) context fixes p :: uint64 begin definition plus_p64 :: "uint64 \ uint64 \ uint64" where "plus_p64 x y \ let z = x + y in if z \ p then z - p else z" definition minus_p64 :: "uint64 \ uint64 \ uint64" where "minus_p64 x y \ if y \ x then x - y else (x + p) - y" definition uminus_p64 :: "uint64 \ uint64" where "uminus_p64 x = (if x = 0 then 0 else p - x)" definition mult_p64 :: "uint64 \ uint64 \ uint64" where "mult_p64 x y = (x * y mod p)" lemma int_of_uint64_shift: "int_of_uint64 (shiftr n k) = (int_of_uint64 n) div (2 ^ k)" 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) (shiftr n 1) in if n AND 1 = 0 then rec else mult_p64 rec x)" by pat_completeness auto termination proof - { fix n :: uint64 assume "n \ 0" with int_of_uint64_ge_0[of n] int_of_uint64_0_iff[of n] have "int_of_uint64 n > 0" by auto hence "0 < int_of_uint64 n" "int_of_uint64 n div 2 < int_of_uint64 n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_uint64 n))", auto simp: int_of_uint64_shift *) qed text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p64 :: "uint64 \ uint64" where "inverse_p64 x = (if x = 0 then 0 else power_p64 x (p - 2))" definition divide_p64 :: "uint64 \ uint64 \ uint64" where "divide_p64 x y = mult_p64 x (inverse_p64 y)" definition finite_field_ops64 :: "uint64 arith_ops_record" where "finite_field_ops64 \ Arith_Ops_Record 0 1 plus_p64 mult_p64 minus_p64 uminus_p64 divide_p64 inverse_p64 (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) uint64_of_int int_of_uint64 (\ x. 0 \ x \ x < p)" end lemma shiftr_uint64_code [code_unfold]: "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, 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 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 (metis (full_types) even_take_bit_eq le_less_trans odd_iff_mod_2_eq_one take_bit_nonnegative zero_neq_numeral zmod_le_nonneg_dividend) apply (auto simp add: less_le) apply (simp add: le_less) done from urel64_eq[OF this urel64_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel64 (shiftr y 1) (int d')" unfolding d' using y unfolding urel64_def using small unfolding ppp apply transfer apply transfer apply (auto simp add: drop_bit_Suc) 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) show ?thesis unfolding inverse_p64_def inverse_p_def urel64_eq[OF x urel64_0] using urel64_0 urel64_power[OF x p] by auto qed lemma mod_ring_0_64: "mod_ring_rel64 0 0" using urel64_0 mod_ring_0 unfolding mod_ring_rel64_def by blast lemma mod_ring_1_64: "mod_ring_rel64 1 1" using urel64_1 mod_ring_1 unfolding mod_ring_rel64_def by blast lemma mod_ring_uminus64: "(mod_ring_rel64 ===> mod_ring_rel64) (uminus_p64 pp) uminus" using urel64_uminus mod_ring_uminus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_plus64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (plus_p64 pp) (+)" using urel64_plus mod_ring_plus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_minus64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (minus_p64 pp) (-)" using urel64_minus mod_ring_minus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_mult64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (mult_p64 pp) ((*))" using urel64_mult mod_ring_mult unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_eq64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> (=)) (=) (=)" using urel64_eq mod_ring_eq unfolding mod_ring_rel64_def rel_fun_def by blast lemma urel64_inj: "urel64 x y \ urel64 x z \ y = z" using urel64_eq[of x y x z] by auto lemma urel64_inj': "urel64 x z \ urel64 y z \ x = y" using urel64_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel64: "bi_unique mod_ring_rel64" "left_unique mod_ring_rel64" "right_unique mod_ring_rel64" using bi_unique_mod_ring_rel urel64_inj' unfolding mod_ring_rel64_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel64_def) lemma right_total_mod_ring_rel64: "right_total mod_ring_rel64" unfolding mod_ring_rel64_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel64 (uint64_of_int z) z" unfolding urel64_def using small unfolding ppp by (auto simp: int_of_uint64_inv) with zy show "\ x z. urel64 x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel64: "Domainp mod_ring_rel64 = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel64 x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel64_def proof let ?i = "int_of_uint64" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" using small unfolding ppp by (transfer, auto simp: word_less_def) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel64 a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel64 x (?i x)" unfolding urel64_def using small * unfolding ppp by (transfer, auto simp: word_less_def) qed next assume "\a b. x = a \ (\z. urel64 a z \ mod_ring_rel z b)" then obtain b z where xz: "urel64 x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel64_def using small unfolding ppp by (transfer, auto simp: word_less_def) qed qed lemma ring_finite_field_ops64: "ring_ops (finite_field_ops64 pp) mod_ring_rel64" by (unfold_locales, auto simp: finite_field_ops64_def bi_unique_mod_ring_rel64 right_total_mod_ring_rel64 mod_ring_plus64 mod_ring_minus64 mod_ring_uminus64 mod_ring_mult64 mod_ring_eq64 mod_ring_0_64 mod_ring_1_64 Domainp_mod_ring_rel64) end end context prime_field begin context fixes pp :: "uint64" assumes *: "p = int_of_uint64 pp" "p \ 4294967295" begin lemma mod_ring_normalize64: "(mod_ring_rel64 ===> mod_ring_rel64) (\x. if x = 0 then 0 else 1) normalize" using urel64_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_mod64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (\x y. if y = 0 then x else 0) (mod)" using urel64_mod[OF *] mod_ring_mod unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor64: "(mod_ring_rel64 ===> mod_ring_rel64) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_inverse64: "(mod_ring_rel64 ===> mod_ring_rel64) (inverse_p64 pp) inverse" using urel64_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_divide64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (divide_p64 pp) (/)" using mod_ring_inverse64 mod_ring_mult64[OF *] unfolding divide_p64_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops64: "field_ops (finite_field_ops64 pp) mod_ring_rel64" by (unfold_locales, insert ring_finite_field_ops64[OF *], auto simp: ring_ops_def finite_field_ops64_def mod_ring_divide64 mod_ring_inverse64 mod_ring_mod64 mod_ring_normalize64) end end (* and a final implementation via integer *) context fixes p :: integer begin definition plus_p_integer :: "integer \ integer \ integer" where "plus_p_integer x y \ let z = x + y in if z \ p then z - p else z" definition minus_p_integer :: "integer \ integer \ integer" where "minus_p_integer x y \ if y \ x then x - y else (x + p) - y" definition uminus_p_integer :: "integer \ integer" where "uminus_p_integer x = (if x = 0 then 0 else p - x)" definition mult_p_integer :: "integer \ integer \ integer" where "mult_p_integer x y = (x * y mod p)" lemma int_of_integer_0_iff: "int_of_integer n = 0 \ n = 0" using integer_eqI by auto lemma int_of_integer_0: "int_of_integer 0 = 0" unfolding int_of_integer_0_iff by simp lemma int_of_integer_plus: "int_of_integer (x + y) = (int_of_integer x + int_of_integer y)" by simp lemma int_of_integer_minus: "int_of_integer (x - y) = (int_of_integer x - int_of_integer y)" by simp lemma int_of_integer_mult: "int_of_integer (x * y) = (int_of_integer x * int_of_integer y)" by simp lemma int_of_integer_mod: "int_of_integer (x mod y) = (int_of_integer x mod int_of_integer y)" by simp lemma int_of_integer_inv: "int_of_integer (integer_of_int x) = x" by simp lemma int_of_integer_shift: "int_of_integer (shiftr n k) = (int_of_integer n) div (2 ^ k)" by 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) (shiftr n 1) in if n AND 1 = 0 then rec else mult_p_integer rec x)" by pat_completeness auto termination proof - { fix n :: integer assume "\ (n \ 0)" hence "n > 0" by auto hence "int_of_integer n > 0" by (simp add: less_integer.rep_eq) hence "0 < int_of_integer n" "int_of_integer n div 2 < int_of_integer n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_integer n))", auto simp: * int_of_integer_shift) qed text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p_integer :: "integer \ integer" where "inverse_p_integer x = (if x = 0 then 0 else power_p_integer x (p - 2))" definition divide_p_integer :: "integer \ integer \ integer" where "divide_p_integer x y = mult_p_integer x (inverse_p_integer y)" definition finite_field_ops_integer :: "integer arith_ops_record" where "finite_field_ops_integer \ Arith_Ops_Record 0 1 plus_p_integer mult_p_integer minus_p_integer uminus_p_integer divide_p_integer inverse_p_integer (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) integer_of_int int_of_integer (\ x. 0 \ x \ x < p)" end lemma shiftr_integer_code [code_unfold]: "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 (shiftr y 1) (int d')" unfolding d' using y unfolding urel_integer_def unfolding ppp shiftr_integer_conv_div_pow2 by auto from id have "y' \ 0" by auto note IH = 1(1)[OF this refl dr'[symmetric] urel_integer_mult[OF x x] div] show ?thesis unfolding power_p.simps[of _ _ "y'"] power_p_integer.simps[of _ _ y] dr' id if_False rem using IH urel_integer_mult[OF IH x] by (auto simp: Let_def) qed qed lemma urel_integer_inverse: assumes x: "urel_integer x x'" shows "urel_integer (inverse_p_integer pp x) (inverse_p p x')" proof - have p: "urel_integer (pp - 2) (int (nat (p - 2)))" using p2 unfolding urel_integer_def unfolding ppp by auto show ?thesis unfolding inverse_p_integer_def inverse_p_def urel_integer_eq[OF x urel_integer_0] using urel_integer_0 urel_integer_power[OF x p] by auto qed lemma mod_ring_0__integer: "mod_ring_rel_integer 0 0" using urel_integer_0 mod_ring_0 unfolding mod_ring_rel_integer_def by blast lemma mod_ring_1__integer: "mod_ring_rel_integer 1 1" using urel_integer_1 mod_ring_1 unfolding mod_ring_rel_integer_def by blast lemma mod_ring_uminus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (uminus_p_integer pp) uminus" using urel_integer_uminus mod_ring_uminus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_plus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (plus_p_integer pp) (+)" using urel_integer_plus mod_ring_plus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_minus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (minus_p_integer pp) (-)" using urel_integer_minus mod_ring_minus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_mult_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (mult_p_integer pp) ((*))" using urel_integer_mult mod_ring_mult unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_eq_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> (=)) (=) (=)" using urel_integer_eq mod_ring_eq unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma urel_integer_inj: "urel_integer x y \ urel_integer x z \ y = z" using urel_integer_eq[of x y x z] by auto lemma urel_integer_inj': "urel_integer x z \ urel_integer y z \ x = y" using urel_integer_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel_integer: "bi_unique mod_ring_rel_integer" "left_unique mod_ring_rel_integer" "right_unique mod_ring_rel_integer" using bi_unique_mod_ring_rel urel_integer_inj' unfolding mod_ring_rel_integer_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel_integer_def) lemma right_total_mod_ring_rel_integer: "right_total mod_ring_rel_integer" unfolding mod_ring_rel_integer_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel_integer (integer_of_int z) z" unfolding urel_integer_def unfolding ppp by auto with zy show "\ x z. urel_integer x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel_integer: "Domainp mod_ring_rel_integer = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel_integer x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel_integer_def proof let ?i = "int_of_integer" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel_integer a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel_integer x (?i x)" unfolding urel_integer_def using * unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) qed next assume "\a b. x = a \ (\z. urel_integer a z \ mod_ring_rel z b)" then obtain b z where xz: "urel_integer x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel_integer_def unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) qed qed lemma ring_finite_field_ops_integer: "ring_ops (finite_field_ops_integer pp) mod_ring_rel_integer" by (unfold_locales, auto simp: finite_field_ops_integer_def bi_unique_mod_ring_rel_integer right_total_mod_ring_rel_integer mod_ring_plus_integer mod_ring_minus_integer mod_ring_uminus_integer mod_ring_mult_integer mod_ring_eq_integer mod_ring_0__integer mod_ring_1__integer Domainp_mod_ring_rel_integer) end end context prime_field begin context fixes pp :: "integer" assumes *: "p = int_of_integer pp" begin lemma mod_ring_normalize_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (\x. if x = 0 then 0 else 1) normalize" using urel_integer_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_mod_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (\x y. if y = 0 then x else 0) (mod)" using urel_integer_mod[OF *] mod_ring_mod unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_inverse_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (inverse_p_integer pp) inverse" using urel_integer_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_divide_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (divide_p_integer pp) (/)" using mod_ring_inverse_integer mod_ring_mult_integer[OF *] unfolding divide_p_integer_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops_integer: "field_ops (finite_field_ops_integer pp) mod_ring_rel_integer" by (unfold_locales, insert ring_finite_field_ops_integer[OF *], auto simp: ring_ops_def finite_field_ops_integer_def mod_ring_divide_integer mod_ring_inverse_integer mod_ring_mod_integer mod_ring_normalize_integer) end end context prime_field begin (* four implementations of modular integer arithmetic for finite fields *) thm finite_field_ops64 finite_field_ops32 finite_field_ops_integer finite_field_ops_int end context mod_ring_locale begin (* four implementations of modular integer arithmetic for finite rings *) thm ring_finite_field_ops64 ring_finite_field_ops32 ring_finite_field_ops_integer ring_finite_field_ops_int end no_notation shiftr (infixl ">>" 55) (* to avoid conflict with bind *) end diff --git a/thys/Ordinary_Differential_Equations/Numerics/Refine_Rigorous_Numerics_Aform.thy b/thys/Ordinary_Differential_Equations/Numerics/Refine_Rigorous_Numerics_Aform.thy --- a/thys/Ordinary_Differential_Equations/Numerics/Refine_Rigorous_Numerics_Aform.thy +++ b/thys/Ordinary_Differential_Equations/Numerics/Refine_Rigorous_Numerics_Aform.thy @@ -1,1743 +1,1740 @@ theory Refine_Rigorous_Numerics_Aform imports Refine_Rigorous_Numerics "HOL-Types_To_Sets.Types_To_Sets" begin lemma Joints_ne_empty[simp]: "Joints xs \ {}" "{} \ Joints xs" by (auto simp: Joints_def valuate_def) lemma Inf_aform_le_Affine: "x \ Affine X \ Inf_aform X \ x" by (auto simp: Affine_def valuate_def intro!: Inf_aform) lemma Sup_aform_ge_Affine: "x \ Affine X \ x \ Sup_aform X" by (auto simp: Affine_def valuate_def intro!: Sup_aform) lemmas Inf_aform'_Affine_le = order_trans[OF Inf_aform' Inf_aform_le_Affine] lemmas Sup_aform'_Affine_ge = order_trans[OF Sup_aform_ge_Affine Sup_aform'] fun approx_form_aform :: "nat \ form \ real aform list \ bool" where "approx_form_aform prec (Less a b) bs = (case (approx_floatariths prec [a - b] bs) of Some [r] \ Sup_aform' prec r < 0 | _ \ False)" | "approx_form_aform prec (LessEqual a b) bs = (case (approx_floatariths prec [a - b] bs) of Some [r] \ Sup_aform' prec r \ 0 | _ \ False)" | "approx_form_aform prec (AtLeastAtMost a b c) bs = (case (approx_floatariths prec [a - b, a - c] bs) of Some [r, s] \ 0 \ Inf_aform' prec r \ Sup_aform' prec s \ 0 | _ \ False)" | "approx_form_aform prec (Bound a b c d) bs = False" | "approx_form_aform prec (Assign a b c) bs = False" | "approx_form_aform prec (Conj a b) bs \ approx_form_aform prec a bs \ approx_form_aform prec b bs" | "approx_form_aform prec (Disj a b) bs \ approx_form_aform prec a bs \ approx_form_aform prec b bs" lemma in_Joints_intervalD: "x \ {Inf_aform' p X .. Sup_aform' p X} \ xs \ Joints XS" if "x#xs \ Joints (X#XS)" using that by (auto simp: Joints_def valuate_def Affine_def intro!: Inf_aform'_Affine_le Sup_aform'_Affine_ge) lemma approx_form_aform: "interpret_form f vs" if "approx_form_aform p f VS" "vs \ Joints VS" using that by (induction f) (auto split: option.splits list.splits simp: algebra_simps dest!: approx_floatariths_outer dest!: in_Joints_intervalD[where p=p]) fun msum_aforms::"nat \ real aform list \ real aform list \ real aform list" where "msum_aforms d (x#xs) (y#ys) = msum_aform d x y#msum_aforms d xs ys" | "msum_aforms d _ _ = []" definition [simp]: "degree_aforms_real = (degree_aforms::real aform list \ nat)" abbreviation "msum_aforms' \ \X Y. msum_aforms (degree_aforms_real (X @ Y)) X Y" lemma aform_val_msum_aforms: assumes "degree_aforms xs \ d" shows "aform_vals e (msum_aforms d xs ys) = List.map2 (+) (aform_vals e xs) (aform_vals (\i. e (i + d)) ys)" using assms proof (induction xs ys rule: msum_aforms.induct) case (1 d x xs y ys) from 1 have "degree_aforms xs \ d" by (auto simp: degrees_def) from 1(1)[OF this] 1 have "aform_vals e (msum_aforms d xs ys) = List.map2 (+) (aform_vals e xs) (aform_vals (\i. e (i + d)) ys)" by simp then show ?case using 1 by (simp add: aform_vals_def aform_val_msum_aform degrees_def) qed (auto simp: aform_vals_def) lemma Joints_msum_aforms: assumes "degree_aforms xs \ d" assumes "degree_aforms ys \ d" shows "Joints (msum_aforms d xs ys) = {List.map2 (+) a b |a b. a \ Joints xs \ b \ Joints ys}" apply (auto simp: Joints_def valuate_def aform_vals_def[symmetric] aform_val_msum_aforms assms) apply force subgoal for e e' apply (rule image_eqI[where x="\i. if i < d then e i else e' (i - d)"]) apply (auto simp: Pi_iff) proof goal_cases case 1 have "(aform_vals e xs) = aform_vals (\i. if i < d then e i else e' (i - d)) xs" using assms by (auto intro!: simp: aform_vals_def aform_val_def degrees_def intro!: pdevs_val_degree_cong) then show ?case by simp qed done definition "split_aforms_largest_uncond_take n X = (let (i, x) = max_pdev (abssum_of_pdevs_list (map snd (take n X))) in split_aforms X i)" text \intersection with plane\ definition "project_coord x b n = (\i\Basis_list. (if i = b then n else if i = -b then -n else x \ i) *\<^sub>R i)" lemma inner_eq_abs_times_sgn: "u \ b = abs u \ b * sgn (u \ b)" if "b \ Basis" for b::"'a::ordered_euclidean_space" by (subst sgn_mult_abs[symmetric]) (auto simp: that abs_inner ) lemma inner_Basis_eq_zero_absI: "x \ Basis \ abs u \ Basis \ x \ \u\ \ x \ 0 \ u \ x = 0" for x::"'a::ordered_euclidean_space" apply (subst euclidean_inner) apply (auto simp: inner_Basis if_distribR if_distrib cong: if_cong) apply (subst inner_eq_abs_times_sgn) by (auto simp: inner_Basis) lemma abs_in_BasisE: fixes u::"'a::ordered_euclidean_space" assumes "abs u \ Basis" obtains "abs u = u" | "abs u = - u" proof (cases "u \ abs u = 1") case True have "abs u = (\i\Basis. (if i = abs u then (abs u \ i) *\<^sub>R i else 0))" using assms by (auto simp: ) also have "\ = (\i\Basis. (if i = abs u then (u \ i) *\<^sub>R i else 0))" by (rule sum.cong) (auto simp: True) also have "\ = (\i\Basis. (u \ i) *\<^sub>R i)" by (rule sum.cong) (auto simp: inner_Basis_eq_zero_absI assms) also have "\ = u" by (simp add: euclidean_representation) finally show ?thesis .. next case False then have F: "u \ \u\ = -1" apply (subst inner_eq_abs_times_sgn[OF assms]) apply (subst (asm) inner_eq_abs_times_sgn[OF assms]) using assms apply (auto simp: inner_Basis sgn_if) by (metis (full_types) abs_0_eq euclidean_all_zero_iff inner_Basis_eq_zero_absI nonzero_Basis) have "abs u = (\i\Basis. (if i = abs u then (abs u \ i) *\<^sub>R i else 0))" using assms by (auto simp: ) also have "\ = (\i\Basis. (if i = abs u then (- u \ i) *\<^sub>R i else 0))" by (rule sum.cong) (auto simp: F) also have "\ = (\i\Basis. (- u \ i) *\<^sub>R i)" by (rule sum.cong) (auto simp: inner_Basis_eq_zero_absI assms) also have "\ = - u" by (subst euclidean_representation) simp finally show ?thesis .. qed lemma abs_in_Basis_iff: fixes u::"'a::ordered_euclidean_space" shows "abs u \ Basis \ u \ Basis \ - u \ Basis" proof - have u: "u = (\i\Basis. (u \ i) *\<^sub>R i)" by (simp add: euclidean_representation) have u': "- u = (\i\Basis. (- (u \ i)) *\<^sub>R i)" by (subst u) (simp add: sum_negf) have au: "abs u = (\i\Basis. \u \ i\ *\<^sub>R i)" by (simp add: eucl_abs[where 'a='a]) have "(u \ Basis \ - u \ Basis)" if "(\u\ \ Basis)" apply (rule abs_in_BasisE[OF that]) using that by auto show ?thesis apply safe subgoal premises prems using prems(1) apply (rule abs_in_BasisE) using prems by simp_all subgoal apply (subst au) apply (subst (asm) u) apply (subst sum.cong[OF refl]) apply (subst abs_inner[symmetric]) apply auto using u apply auto[1] done subgoal apply (subst au) apply (subst (asm) u') apply (subst sum.cong[OF refl]) apply (subst abs_inner[symmetric]) apply auto using u' apply auto by (metis Basis_nonneg abs_of_nonpos inner_minus_left neg_0_le_iff_le scaleR_left.minus) done qed lemma abs_inner_Basis: fixes u v::"'a::ordered_euclidean_space" assumes "abs u \ Basis" "abs v \ Basis" shows "inner u v = (if u = v then 1 else if u = -v then -1 else 0)" proof - define i where "i \ if u \ Basis then u else -u" define j where "j \ if v \ Basis then v else -v" have u: "u = (if u \ Basis then i else - i)" and v: "v = (if v \ Basis then j else - j)" by (auto simp: i_def j_def) have "i \ Basis" "j \ Basis" using assms by (auto simp: i_def j_def abs_in_Basis_iff) then show ?thesis apply (subst u) apply (subst v) apply (auto simp: inner_Basis) apply (auto simp: j_def i_def split: if_splits) using Basis_nonneg abs_of_nonpos by fastforce qed lemma project_coord_inner_Basis: assumes "i \ Basis" shows "(project_coord x b n) \ i = (if i = b then n else if i = -b then -n else x \ i)" proof - have "project_coord x b n \ i = (\j\Basis. (if j = b then n else if j = -b then -n else x \ j) * (if j = i then 1 else 0))" using assms by (auto simp: project_coord_def inner_sum_left inner_Basis) also have "\ = (\j\Basis. (if j = i then (if j = b then n else if j = -b then -n else x \ j) else 0))" by (rule sum.cong) auto also have "\ = (if i = b then n else if i = -b then -n else x \ i)" using assms by (subst sum.delta) auto finally show ?thesis by simp qed lemma project_coord_inner: assumes "abs i \ Basis" shows "(project_coord x b n) \ i = (if i = b then n else if i = -b then -n else x \ i)" proof - consider "i \ Basis" | "- i \ Basis" using assms by (auto simp: abs_in_Basis_iff) then show ?thesis proof cases case 1 then show ?thesis by (simp add: project_coord_inner_Basis) next case 2 have "project_coord x b n \ i = - (project_coord x b n \ - i)" by simp also have "\ = - (if - i = b then n else if - i = -b then -n else x \ - i)" using 2 by (subst project_coord_inner_Basis) simp_all also have "\ = (if i = b then n else if i = -b then -n else x \ i)" using 2 by auto (metis Basis_nonneg abs_le_zero_iff abs_of_nonneg neg_le_0_iff_le nonzero_Basis) finally show ?thesis . qed qed lift_definition project_coord_pdevs::"'a::executable_euclidean_space sctn \ 'a pdevs \ 'a pdevs" is "\sctn xs i. project_coord (xs i) (normal sctn) 0" by (rule finite_subset) (auto simp: project_coord_def cong: if_cong) definition "project_coord_aform sctn cxs = (project_coord (fst cxs) (normal sctn) (pstn sctn), project_coord_pdevs sctn (snd cxs))" definition project_coord_aform_lv :: "real aform list \ nat \ real \ real aform list" where "project_coord_aform_lv xs b n = xs[b:=(n, zero_pdevs)]" definition "project_ncoord_aform x b n = project_coord_aform (Sctn (Basis_list ! b) n) x" lemma finite_sum_subset_aux_lemma: assumes "finite s" shows " {i. (\x\s. f x i) \ 0} \ {i. \x\s. f x i \ 0}" proof - have "{i. (\x\s. f x i) \ 0} = - {i. (\x\s. f x i) = 0}" by auto also have "\ \ - {i. \x \ s. f x i = 0}" by auto also have "\ = {i. \x \ s. f x i \ 0}" by auto finally show ?thesis by simp qed lift_definition sum_pdevs::"('i \ 'a::comm_monoid_add pdevs) \ 'i set \ 'a pdevs" is "\f X. if finite X then (\i. \x\X. f x i) else (\_. 0)" apply auto apply (rule finite_subset) apply (rule finite_sum_subset_aux_lemma) apply auto done lemma pdevs_apply_sum_pdevs[simp]: "pdevs_apply (sum_pdevs f X) i = (\x\X. pdevs_apply (f x) i)" by transfer auto lemma sum_pdevs_empty[simp]: "sum_pdevs f {} = zero_pdevs" by transfer auto lemma sum_pdevs_insert[simp]: "finite xs \ sum_pdevs f (insert a xs) = (if a \ xs then sum_pdevs f xs else add_pdevs (f a) (sum_pdevs f xs))" by (auto simp: insert_absorb intro!: pdevs_eqI) lemma sum_pdevs_infinite[simp]: "infinite X \ sum_pdevs f X = zero_pdevs" by transfer auto lemma compute_sum_pdevs[code]: "sum_pdevs f (set XS) = foldr (\a b. add_pdevs (f a) b) (remdups XS) zero_pdevs" by (induction XS) (auto simp: ) lemma degree_sum_pdevs_le: "degree (sum_pdevs f X) \ Max (degree ` f ` X)" apply (rule degree_le) apply auto apply (cases "X = {}") subgoal by (simp add: ) subgoal by (cases "finite X") simp_all done lemma pdevs_val_sum_pdevs: "pdevs_val e (sum_pdevs f X) = (\x\X. pdevs_val e (f x))" apply (cases "finite X") subgoal apply (subst pdevs_val_sum_le) apply (rule degree_sum_pdevs_le) apply (auto simp: scaleR_sum_right) apply (subst sum.swap) apply (rule sum.cong[OF refl]) apply (subst pdevs_val_sum_le[where d="Max (degree ` f ` X)"]) apply (auto simp: Max_ge_iff) done subgoal by simp done definition eucl_of_list_aform :: "(real \ real pdevs) list \ 'a::executable_euclidean_space aform" where "eucl_of_list_aform xs = (eucl_of_list (map fst xs), sum_pdevs (\i. pdevs_scaleR (snd (xs ! index Basis_list i)) i) Basis)" definition lv_aforms_rel::"(((real \ real pdevs) list) \ ('a::executable_euclidean_space aform)) set" where "lv_aforms_rel = br eucl_of_list_aform (\xs. length xs = DIM('a))" definition "inner_aforms' p X Y = (let fas = [inner_floatariths (map floatarith.Var [0..d x y. Some (F d x y)) f" assumes "[x, y] \ Joints [X, Y]" assumes "d \ degree_aform X" assumes "d \ degree_aform Y" shows "f x y \ Affine (F d X Y)" proof - from assms(2) obtain e where e: "e \ funcset UNIV {-1 .. 1}" "x = aform_val e X" "y = aform_val e Y" by (auto simp: Joints_def valuate_def) from affine_extension2E[OF assms(1) refl e(1) assms(3) assms(4)] obtain e' where "e' \ funcset UNIV {- 1..1}" "f (aform_val e X) (aform_val e Y) = aform_val e' (F d X Y)" "\n. n < d \ e' n = e n" "aform_val e X = aform_val e' X" "aform_val e Y = aform_val e' Y" by auto then show ?thesis by (auto simp: Affine_def valuate_def e) qed lemma aform_val_zero_pdevs[simp]: "aform_val e (x, zero_pdevs) = x" by (auto simp: aform_val_def) lemma neg_equal_zero_eucl[simp]: "-a = a \ a = 0" for a::"'a::euclidean_space" by (auto simp: algebra_simps euclidean_eq_iff[where 'a='a]) context includes autoref_syntax begin lemma Option_bind_param[param, autoref_rules]: "((\), (\)) \ \S\option_rel \ (S \ \R\option_rel) \ \R\option_rel" unfolding Option.bind_def by parametricity lemma zip_Basis_list_pat[autoref_op_pat_def]: "\(b, m)\zip Basis_list ms. m *\<^sub>R b \ OP eucl_of_list $ ms" proof (rule eq_reflection) have z: "zip ms (Basis_list::'a list) = map (\(x, y). (y, x)) (zip Basis_list ms)" by (subst zip_commute) simp show "(\(b, m)\zip (Basis_list::'a list) ms. m *\<^sub>R b) = OP eucl_of_list $ ms" unfolding eucl_of_list_def autoref_tag_defs apply (subst z) apply (subst map_map) apply (auto cong: map_cong simp: split_beta') done qed lemma length_aforms_of_ivls: "length (aforms_of_ivls a b) = min (length a) (length b)" by (auto simp: aforms_of_ivls_def) lemma length_lv_rel: "(ys, y::'a) \ lv_rel \ length ys = DIM('a::executable_euclidean_space)" by (auto simp: lv_rel_def br_def) lemma lv_rel_nth[simp]: "(xs, x::'a::executable_euclidean_space) \ lv_rel \ b \ Basis \ xs ! index (Basis_list) b = x \ b" unfolding lv_rel_def by (auto simp: br_def eucl_of_list_inner) lemma aform_of_ivl_autoref[autoref_rules]: "(aforms_of_ivls, aform_of_ivl) \ lv_rel \ lv_rel \ lv_aforms_rel" apply (auto simp: lv_aforms_rel_def br_def eucl_of_list_aform_def length_aforms_of_ivls length_lv_rel) subgoal for xs x ys y apply (auto simp: aform_of_ivl_def aforms_of_ivls_def o_def eucl_of_list_inner inner_simps pdevs_apply_pdevs_of_ivl length_lv_rel intro!: euclidean_eqI[where 'a='a] pdevs_eqI) by (metis index_Basis_list_nth inner_not_same_Basis length_Basis_list nth_Basis_list_in_Basis) done lemma bound_intersect_2d_ud[autoref_rules]: shows "(bound_intersect_2d_ud, bound_intersect_2d_ud) \ nat_rel \ ((rnv_rel \\<^sub>r rnv_rel) \\<^sub>r Id) \ rnv_rel \ \rnv_rel \\<^sub>r rnv_rel\option_rel" by auto lemma eucl_of_list_autoref[autoref_rules]: includes autoref_syntax assumes "SIDE_PRECOND (length xs = DIM('a::executable_euclidean_space))" assumes "(xsi, xs) \ \rnv_rel\list_rel" shows "(xsi, eucl_of_list $ xs::'a) \ lv_rel" using assms unfolding lv_rel_def by (auto simp: br_def) definition "inner2s x b c = (inner_lv_rel x b, inner_lv_rel x c)" lemma inner_lv_rel_autoref[autoref_rules]: "(inner_lv_rel, (\)) \ lv_rel \ lv_rel \ rnv_rel" using lv_rel_inner[unfolded inner_lv_rel_def[symmetric]] by auto lemma inner_lv_rel_eq: "\length xs = DIM('a::executable_euclidean_space); (xa, x'a) \ lv_rel\ \ inner_lv_rel xs xa = eucl_of_list xs \ (x'a::'a)" using inner_lv_rel_autoref[param_fo, of "xs" "eucl_of_list xs" xa x'a] unfolding lv_rel_def by (auto simp: br_def) definition "inner_pdevs xs ys = sum_pdevs (\i. scaleR_pdevs (ys ! i) (xs ! i)) {..xs a b. (inner2s (map fst xs) a b, pdevs_inner2s (map snd xs) a b)), inner2_aform) \ lv_aforms_rel \ lv_rel \ lv_rel \ ((rnv_rel \\<^sub>r rnv_rel)\\<^sub>rId)" unfolding inner2_aform_def apply (auto simp: lv_aforms_rel_def br_def eucl_of_list_aform_def inner2_aform_def ) apply (auto simp: inner2s_def inner2_def inner_lv_rel_eq pdevs_inner2s_def inner_pdevs_def sum_Basis_sum_nth_Basis_list inner_sum_left inner_Basis mult.commute intro!: pdevs_eqI) subgoal for a b c d e f apply (rule sum.cong) apply force subgoal for n by (auto dest!: lv_rel_nth[where b="Basis_list ! n"] simp: inner_commute) done subgoal for a b c d e f apply (rule sum.cong) apply force subgoal for n by (auto dest!: lv_rel_nth[where b="Basis_list ! n"] simp: inner_commute) done done lemma RETURN_inter_aform_plane_ortho_annot: "RETURN (inter_aform_plane_ortho p (Z::'a aform) n g) = (case those (map (\b. bound_intersect_2d_ud p (inner2_aform Z n b) g) Basis_list) of Some mMs \ do { ASSERT (length mMs = DIM('a::executable_euclidean_space)); let l = (\(b, m)\zip Basis_list (map fst mMs). m *\<^sub>R b); let u = (\(b, M)\zip Basis_list (map snd mMs). M *\<^sub>R b); RETURN (Some (aform_of_ivl l u)) } | None \ RETURN None)" apply (auto simp: inter_aform_plane_ortho_def split: option.splits) subgoal premises prems for x2 proof - have "length x2 = DIM('a)" using prems using map_eq_imp_length_eq by (force simp: those_eq_Some_map_Some_iff) then show ?thesis using prems by auto qed done definition "inter_aform_plane_ortho_nres p Z n g = RETURN (inter_aform_plane_ortho p Z n g)" schematic_goal inter_aform_plane_ortho_lv: fixes Z::"'a::executable_euclidean_space aform" assumes [autoref_rules_raw]: "DIM_precond TYPE('a) D" assumes [autoref_rules]: "(pp, p) \ nat_rel" "(Zi, Z) \ lv_aforms_rel" "(ni, n) \ lv_rel" "(gi, g) \ rnv_rel" notes [autoref_rules] = those_param param_map shows "(nres_of (?f::?'a dres), inter_aform_plane_ortho_nres $ p $ Z $ n $ g) \ ?R" unfolding autoref_tag_defs unfolding RETURN_inter_aform_plane_ortho_annot inter_aform_plane_ortho_nres_def including art by autoref_monadic concrete_definition inter_aform_plane_ortho_lv for pp Zi ni gi uses inter_aform_plane_ortho_lv lemmas [autoref_rules] = inter_aform_plane_ortho_lv.refine lemma project_coord_lv[autoref_rules]: assumes "(xs, x::'a::executable_euclidean_space aform) \ lv_aforms_rel" "(ni, n) \ nat_rel" assumes "SIDE_PRECOND (n < DIM('a))" shows "(project_coord_aform_lv xs ni, project_ncoord_aform $ x $ n) \ rnv_rel \ lv_aforms_rel" using assms apply (auto simp: project_coord_aform_lv_def project_ncoord_aform_def lv_rel_def project_coord_aform_def eucl_of_list_aform_def lv_aforms_rel_def br_def) subgoal apply (auto intro!: euclidean_eqI[where 'a='a]) apply (subst project_coord_inner_Basis) apply (auto simp: eucl_of_list_inner ) apply (subst nth_list_update) apply (auto simp: ) using in_Basis_index_Basis_list apply force using inner_not_same_Basis nth_Basis_list_in_Basis apply fastforce apply (auto simp: nth_list_update) done subgoal for i apply (auto intro!: pdevs_eqI simp: project_coord_pdevs.rep_eq) apply (auto intro!: euclidean_eqI[where 'a='a]) apply (subst project_coord_inner_Basis) apply (auto simp: eucl_of_list_inner ) apply (subst nth_list_update) apply (auto simp: ) using inner_not_same_Basis nth_Basis_list_in_Basis apply fastforce apply (auto simp: nth_list_update) done done definition inter_aform_plane where "inter_aform_plane prec Xs (sctn::'a sctn) = do { cxs \ inter_aform_plane_ortho_nres (prec) (Xs) (normal sctn) (pstn sctn); case cxs of Some cxs \ (if normal sctn \ set Basis_list then do { ASSERT (index Basis_list (normal sctn) < DIM('a::executable_euclidean_space)); RETURN ((project_ncoord_aform cxs (index Basis_list (normal sctn)) (pstn sctn))) } else if - normal sctn \ set Basis_list then do { ASSERT (index Basis_list (-normal sctn) < DIM('a)); RETURN ((project_ncoord_aform cxs (index Basis_list (-normal sctn)) (- pstn sctn))) } else SUCCEED) | None \ SUCCEED }" lemma [autoref_rules]: assumes [THEN GEN_OP_D, param]: "GEN_OP (=) (=) (A \ A \ bool_rel)" shows "(index, index) \ \A\list_rel \ A \ nat_rel" unfolding index_def find_index_def by parametricity schematic_goal inter_aform_plane_lv: fixes Z::"'a::executable_euclidean_space aform" assumes [autoref_rules_raw]: "DIM_precond TYPE('a) D" assumes [autoref_rules]: "(preci, prec) \ nat_rel" "(Zi, Z) \ lv_aforms_rel" "(sctni, sctn) \ \lv_rel\sctn_rel" notes [autoref_rules] = those_param param_map shows "(nres_of (?f::?'a dres), inter_aform_plane prec Z sctn) \ ?R" unfolding autoref_tag_defs unfolding inter_aform_plane_def including art by (autoref_monadic ) concrete_definition inter_aform_plane_lv for preci Zi sctni uses inter_aform_plane_lv lemmas [autoref_rules] = inter_aform_plane_lv.refine[autoref_higher_order_rule(1)] end lemma Basis_not_uminus: fixes i::"'a::euclidean_space" shows "i \ Basis \ - i \ Basis" by (metis inner_Basis inner_minus_left one_neq_neg_one zero_neq_neg_one) lemma pdevs_apply_project_coord_pdevs: assumes "normal sctn \ Basis \ - normal sctn \ Basis" assumes "i \ Basis" shows "pdevs_apply (project_coord_pdevs sctn cxs) x \ i = (if i = abs (normal sctn) then 0 else pdevs_apply cxs x \ i)" using assms[unfolded abs_in_Basis_iff[symmetric]] apply (transfer fixing: sctn) apply (auto simp: project_coord_inner Basis_not_uminus) using Basis_nonneg apply force using Basis_nonneg assms(1) by force lemma pdevs_domain_project_coord_pdevs_subset: "pdevs_domain (project_coord_pdevs sctn X) \ pdevs_domain X" apply (auto simp: ) apply (auto simp: euclidean_eq_iff[where 'a='a] ) by (metis add.inverse_neutral project_coord_inner_Basis project_coord_pdevs.rep_eq) lemma pdevs_val_project_coord_pdevs_inner_Basis: assumes "b \ Basis" shows "pdevs_val e (project_coord_pdevs sctn X) \ b = (if b = abs (normal sctn) then 0 else pdevs_val e X \ b)" using assms apply (auto simp: ) apply (simp add: pdevs_val_pdevs_domain inner_sum_left ) apply (subst sum.cong[OF refl]) apply (subst pdevs_apply_project_coord_pdevs) apply (simp add: abs_in_Basis_iff) apply simp apply (rule refl) apply auto unfolding pdevs_val_pdevs_domain inner_sum_left apply (rule sum.mono_neutral_cong_left) apply simp apply (rule pdevs_domain_project_coord_pdevs_subset) apply auto apply (metis Basis_nonneg abs_minus_cancel abs_of_nonneg euclidean_all_zero_iff project_coord_inner_Basis project_coord_pdevs.rep_eq) apply (metis Basis_nonneg abs_minus_cancel abs_of_nonneg project_coord_inner_Basis project_coord_pdevs.rep_eq) done lemma inner_in_Sum: "b \ Basis \ x \ b = (\i\Basis. x \ i * (i \ b))" apply (subst euclidean_representation[of x, symmetric]) unfolding inner_sum_left by (auto simp: intro!: ) lemma aform_val_project_coord_aform: "aform_val e (project_coord_aform sctn X) = project_coord (aform_val e X) (normal sctn) (pstn sctn)" apply (auto simp: aform_val_def project_coord_def project_coord_aform_def) apply (rule euclidean_eqI) unfolding inner_add_left inner_sum_left apply (subst pdevs_val_project_coord_pdevs_inner_Basis) apply (auto simp: ) apply (rule sum.cong) apply auto apply (metis abs_in_Basis_iff abs_inner abs_inner_Basis abs_zero inner_commute) apply (subst inner_in_Sum[where x="pdevs_val e (snd X)"], force) unfolding sum.distrib[symmetric] apply (rule sum.cong) apply auto apply (metis Basis_nonneg abs_inner_Basis abs_of_nonneg abs_of_nonpos inner_commute neg_0_le_iff_le) apply (metis Basis_nonneg abs_inner_Basis abs_of_nonneg abs_of_nonpos neg_0_le_iff_le) apply (auto simp: inner_Basis) done lemma project_coord_on_plane: assumes "x \ plane_of (Sctn n c)" shows "project_coord x n c = x" using assms by (auto simp: project_coord_def plane_of_def intro!: euclidean_eqI[where 'a='a]) lemma mem_Affine_project_coord_aformI: assumes "x \ Affine X" assumes "x \ plane_of sctn" shows "x \ Affine (project_coord_aform sctn X)" proof - have "x = project_coord x (normal sctn) (pstn sctn)" using assms by (intro project_coord_on_plane[symmetric]) auto also from assms obtain e where "e \ funcset UNIV {-1 .. 1}" "x = aform_val e X" by (auto simp: Affine_def valuate_def) note this(2) also have "project_coord (aform_val e X) (normal sctn) (pstn sctn) = aform_val e (project_coord_aform sctn X)" by (simp add: aform_val_project_coord_aform) finally have "x = aform_val e (project_coord_aform sctn X)" unfolding \x = aform_val e X\ . then show ?thesis using \e \ _\ by (auto simp: Affine_def valuate_def) qed lemma mem_Affine_project_coord_aformD: assumes "x \ Affine (project_coord_aform sctn X)" assumes "abs (normal sctn) \ Basis" shows "x \ plane_of sctn" using assms by (auto simp: Affine_def valuate_def aform_val_project_coord_aform plane_of_def project_coord_inner) definition "reduce_aform prec t X = summarize_aforms (prec) (collect_threshold (prec) 0 t) (degree_aforms X) X" definition "correct_girard p m N (X::real aform list) = (let Xs = map snd X; M = pdevs_mapping Xs; D = domain_pdevs Xs; diff = m - card D in if 0 < diff then (\_ _. True) else let Ds = sorted_list_of_set D; ortho_indices = map fst (take (diff + N) (sort_key (\(i, r). r) (map (\i. let xs = M i in (i, sum_list' p (map abs xs) - fold max (map abs xs) 0)) Ds))); _ = () in (\i (xs::real list). i \ set ortho_indices))" definition "reduce_aforms prec C X = summarize_aforms (prec) C (degree_aforms X) X" definition "pdevs_of_real x = (x, zero_pdevs)" definition aform_inf_inner where "aform_inf_inner prec X n = (case inner_aforms' (prec) X (map pdevs_of_real n) of Some Xn \ Inf_aform' (prec) (hd Xn))" definition aform_sup_inner where "aform_sup_inner prec X n = (case inner_aforms' (prec) X (map pdevs_of_real n) of Some Xn \ Sup_aform' (prec) (hd Xn))" text \cannot fail\ lemma approx_un_ne_None: "approx_un p (\ivl. Some (f ivl)) (Some r) \ None" by (auto simp: approx_un_def split_beta') lemma approx_un_eq_Some: "approx_un p (\ivl. Some (f ivl)) (Some r) = Some s \ s = ivl_err (real_interval (f (ivl_of_aform_err p r)))" using approx_un_ne_None by (auto simp: approx_un_def split_beta') lemma is_float_uminus: "is_float aa \ is_float (- aa)" by (auto simp: is_float_def) lemma is_float_uminus_iff[simp]: "is_float (- aa) = is_float aa" using is_float_uminus[of aa] is_float_uminus[of "-aa"] by (auto simp: is_float_def) lemma is_float_simps[simp]: "is_float 0" "is_float 1" "is_float (real_of_float x)" "is_float (truncate_down p X)" "is_float (truncate_up p X)" "is_float (eucl_truncate_down p X)" "is_float (eucl_truncate_up p X)" by (auto simp: is_float_def eucl_truncate_down_def eucl_truncate_up_def truncate_down_def truncate_up_def) lemma is_float_sum_list'[simp]: "is_float (sum_list' p xs)" by (induction xs) (auto simp: is_float_def) lemma is_float_inverse_aform': "is_float (fst (fst (inverse_aform' p X)))" unfolding inverse_aform'_def by (simp add: Let_def trunc_bound_eucl_def) lemma is_float_min_range: "min_range_antimono p F D E X Y = Some ((a, b), c) \ is_float a" "min_range_antimono p F D E X Y = Some ((a, b), c) \ is_float c" "min_range_mono p F D E X Y = Some ((a, b), c) \ is_float a" "min_range_mono p F D E X Y = Some ((a, b), c) \ is_float c" by (auto simp: min_range_antimono_def min_range_mono_def Let_def bind_eq_Some_conv prod_eq_iff trunc_bound_eucl_def mid_err_def affine_unop_def split: prod.splits) lemma is_float_ivl_err: assumes "ivl_err x = ((a, b), c)" "is_float (lower x)" "is_float (upper x)" shows "is_float a" "is_float c" proof - define x1 where "x1 = lower x" define x2 where "x2 = upper x" have "a = (x1 + x2) / 2" "c = (x2 - x1) / 2" using assms by (auto simp: ivl_err_def x1_def x2_def) moreover have "(x1 + x2) / 2 \ float" "(x2 - x1) / 2 \ float" using assms by (auto simp: is_float_def x1_def x2_def) ultimately show "is_float a" "is_float c" unfolding is_float_def by blast+ qed lemma is_float_trunc_bound_eucl[simp]: "is_float (fst (trunc_bound_eucl p X))" by (auto simp: trunc_bound_eucl_def Let_def) lemma add_eq_times_2_in_float: "a + b = c * 2 \ is_float a \ is_float b \ is_float c" proof goal_cases case 1 then have "c = (a + b) / 2" by simp also have "\ \ float" using 1(3,2) by (simp add: is_float_def) finally show ?case by (simp add: is_float_def) qed lemma approx_floatarith_Some_is_float: "approx_floatarith p fa XS = Some ((a, b), ba) \ list_all (\((a, b), c). is_float a \ is_float c) XS \ is_float a \ is_float ba" apply (induction fa arbitrary: a b ba) subgoal by (auto simp: bind_eq_Some_conv add_aform'_def Let_def ) subgoal by (auto simp: bind_eq_Some_conv uminus_aform_def Let_def) subgoal by (auto simp: bind_eq_Some_conv mult_aform'_def Let_def ) subgoal by (auto simp: bind_eq_Some_conv inverse_aform_err_def map_aform_err_def prod_eq_iff is_float_inverse_aform' acc_err_def aform_to_aform_err_def aform_err_to_aform_def inverse_aform_def Let_def split: if_splits) subgoal by (auto simp: bind_eq_Some_conv cos_aform_err_def Let_def is_float_min_range is_float_ivl_err split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv arctan_aform_err_def Let_def is_float_min_range is_float_ivl_err split: if_splits prod.splits) subgoal apply (auto simp: bind_eq_Some_conv uminus_aform_def Let_def is_float_min_range is_float_ivl_err set_of_eq real_interval_abs split: if_splits prod.splits) apply (metis is_float_ivl_err(1) is_float_simps(3) lower_real_interval real_interval_abs upper_real_interval) by (metis is_float_ivl_err(2) is_float_simps(3) lower_real_interval real_interval_abs upper_real_interval) subgoal by (auto simp: bind_eq_Some_conv max_aform_err_def Let_def is_float_min_range is_float_ivl_err max_def split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv min_aform_err_def Let_def is_float_min_range is_float_ivl_err min_def split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv arctan_aform_err_def Let_def is_float_min_range is_float_ivl_err split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv sqrt_aform_err_def Let_def is_float_min_range is_float_ivl_err split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv exp_aform_err_def Let_def is_float_min_range is_float_ivl_err split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv powr_aform_err_def approx_bin_def exp_aform_err_def Let_def is_float_min_range is_float_ivl_err mult_aform'_def split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv ln_aform_err_def Let_def is_float_min_range is_float_ivl_err split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv power_aform_err_def Let_def is_float_min_range is_float_ivl_err mid_err_def dest!: add_eq_times_2_in_float split: if_splits prod.splits) subgoal by (auto simp: bind_eq_Some_conv cos_aform_err_def Let_def is_float_min_range approx_un_def is_float_ivl_err split: if_splits) subgoal by (auto simp: bind_eq_Some_conv cos_aform_err_def Let_def is_float_min_range list_all_length split: if_splits) subgoal by (auto simp: bind_eq_Some_conv num_aform_def Let_def) done lemma plain_floatarith_not_None: assumes "plain_floatarith N fa" "N \ length XS" "list_all (\((a, b), c). is_float a \ is_float c) XS" shows "approx_floatarith p fa XS \ None" using assms by (induction fa) (auto simp: Let_def split_beta' approx_un_eq_Some prod_eq_iff approx_floatarith_Some_is_float) lemma plain_floatarith_slp_not_None: assumes "\fa. fa \ set fas \ plain_floatarith N fa" "N \ length XS" "list_all (\((a, b), c). is_float a \ is_float c) XS" shows "approx_slp p fas XS \ None" using assms apply (induction fas arbitrary: XS) subgoal by simp subgoal for fa fas XS using plain_floatarith_not_None[of N fa XS p] by (auto simp: Let_def split_beta' approx_un_eq_Some prod_eq_iff bind_eq_Some_conv approx_floatarith_Some_is_float) done lemma plain_floatarithE: assumes "plain_floatarith N fa" "N \ length XS" "list_all (\((a, b), c). is_float a \ is_float c) XS" obtains R where "approx_floatarith p fa XS = Some R" using plain_floatarith_not_None[OF assms] by force lemma approx_slp_outer_eq_None_iff: "approx_slp_outer p a b XS = None \ approx_slp p b ((map (\x. (x, 0)) XS)) = None" by (auto simp: approx_slp_outer_def Let_def bind_eq_None_conv) lemma approx_slp_sing_eq_None_iff: "approx_slp p [b] xs = None \ approx_floatarith p b xs = None" by (auto simp: approx_slp_outer_def Let_def bind_eq_None_conv) lemma plain_inner_floatariths: "plain_floatarith N (inner_floatariths xs ys)" if "list_all (plain_floatarith N) xs" "list_all (plain_floatarith N) ys" using that by (induction xs ys rule: inner_floatariths.induct) auto lemma aiN: "approx_floatarith p (inner_floatariths xs ys) zs \ None" if "\x. x \ set xs \ approx_floatarith p x zs \ None" and "\x. x \ set ys \ approx_floatarith p x zs \ None" using that apply (induction xs ys rule: inner_floatariths.induct) apply (auto simp: Let_def bind_eq_Some_conv) by (metis old.prod.exhaust) lemma aiVN: "approx_floatarith p (inner_floatariths (map floatarith.Var [0..x. (x, 0)) a @ map (\x. (x, 0)) b) \ None" by (rule aiN) (auto simp: nth_append) lemma iaN: "inner_aforms' p a b \ None" unfolding inner_aforms'_def Let_def approx_slp_outer_def using aiVN[of p a b] by (auto simp: Let_def bind_eq_Some_conv) definition "support_aform prec b X = (let ia = inner_aform X b in fst X \ b + tdev' (prec) (snd ia))" definition "width_aforms prec X = (let t = tdev' (prec) ((abssum_of_pdevs_list (map snd X))) in t)" definition "inf_aforms prec xs = map (Inf_aform' (prec)) xs" definition "sup_aforms prec xs = map (Sup_aform' (prec)) xs" definition "fresh_index_aforms xs = Max (insert 0 (degree_aform ` set xs))" definition "independent_aforms x env = (msum_aform (fresh_index_aforms env) (0, zero_pdevs) x#env)" definition "aform_form_ivl prec form xs = dRETURN (approx_form prec form (ivls_of_aforms prec xs) [])" definition "aform_form prec form xs = dRETURN (approx_form_aform prec form xs)" definition "aform_slp prec n slp xs = dRETURN (approx_slp_outer prec n slp xs)" definition "aform_isFDERIV prec n ns fas xs = dRETURN (isFDERIV_approx prec n ns fas (ivls_of_aforms prec xs))" definition aform_rel_internal: "aform_rel R = br Affine top O \R\set_rel" lemma aform_rel_def: "\rnv_rel\aform_rel = br Affine top" unfolding relAPP_def by (auto simp: aform_rel_internal) definition "aforms_rel = br Joints top" definition aform_rell :: "((real \ real pdevs) list \ real list set) set" where "aform_rell = aforms_rel" definition aforms_relp_internal: "aforms_relp R = aforms_rel O \R\set_rel" lemma aforms_relp_def: "\R\aforms_relp = aforms_rel O \R\set_rel" by (auto simp: aforms_relp_internal relAPP_def) definition "product_aforms x y = x @ msum_aforms (degree_aforms (x)) (replicate (length y) (0, zero_pdevs)) y" lemma eucl_of_list_mem_eucl_of_list_aform: assumes "x \ Joints a" assumes "length a = DIM('a::executable_euclidean_space)" shows "eucl_of_list x \ Affine (eucl_of_list_aform a::'a aform)" using assms by (auto simp: Joints_def Affine_def valuate_def eucl_of_list_aform_def aform_val_def pdevs_val_sum_pdevs inner_simps eucl_of_list_inner intro!: euclidean_eqI[where 'a='a]) lemma in_image_eucl_of_list_eucl_of_list_aform: assumes "length x = DIM('a::executable_euclidean_space)" "xa \ Affine (eucl_of_list_aform x::'a aform)" shows "xa \ eucl_of_list ` Joints x" using assms by (auto intro!: nth_equalityI image_eqI[where x="list_of_eucl xa"] simp: aform_val_def eucl_of_list_aform_def Affine_def valuate_def Joints_def inner_simps pdevs_val_sum_pdevs eucl_of_list_inner) lemma eucl_of_list_image_Joints: assumes "length x = DIM('a::executable_euclidean_space)" shows "eucl_of_list ` Joints x = Affine (eucl_of_list_aform x::'a aform)" using assms by (auto intro!: eucl_of_list_mem_eucl_of_list_aform in_image_eucl_of_list_eucl_of_list_aform) definition "aform_ops = \ appr_of_ivl = aforms_of_ivls, product_appr = product_aforms, msum_appr = msum_aforms', inf_of_appr = \optns. inf_aforms (precision optns), sup_of_appr = \optns. sup_aforms (precision optns), split_appr = split_aforms_largest_uncond_take, appr_inf_inner = \optns. aform_inf_inner (precision optns), appr_sup_inner = \optns. aform_sup_inner (precision optns), inter_appr_plane = \optns xs sctn. inter_aform_plane_lv (length xs) (precision optns) xs sctn, reduce_appr = \optns. reduce_aforms (precision optns), width_appr = \optns. width_aforms (precision optns), approx_slp_dres = \optns. aform_slp (precision optns), approx_euclarithform = \optns. aform_form (precision optns), approx_isFDERIV = \optns. aform_isFDERIV (precision optns) \" lemma Affine_eq_permI: assumes "fst X = fst Y" assumes "map snd (list_of_pdevs (snd X)) <~~> map snd (list_of_pdevs (snd Y))" (is "?perm X Y") shows "Affine X = Affine Y" proof - { fix X Y and e::"nat \ real" assume perm: "?perm X Y" and e: "e \ funcset UNIV {- 1..1}" from pdevs_val_of_list_of_pdevs2[OF e, of "snd X"] obtain e' where e': "pdevs_val e (snd X) = pdevs_val e' (pdevs_of_list (map snd (list_of_pdevs (snd X))))" "e' \ funcset UNIV {- 1..1}" by auto note e'(1) also from pdevs_val_perm[OF perm e'(2)] obtain e'' where e'': "e'' \ funcset UNIV {- 1..1}" "pdevs_val e' (pdevs_of_list (map snd (list_of_pdevs (snd X)))) = pdevs_val e'' (pdevs_of_list (map snd (list_of_pdevs (snd Y))))" by auto note e''(2) also from pdevs_val_of_list_of_pdevs[OF e''(1), of "snd Y", simplified] obtain e''' where e''': "pdevs_val e'' (pdevs_of_list (map snd (list_of_pdevs (snd Y)))) = pdevs_val e''' (snd Y)" "e''' \ funcset UNIV {- 1..1}" by auto note e'''(1) finally have "\e' \ funcset UNIV {-1 .. 1}. pdevs_val e (snd X) = pdevs_val e' (snd Y)" using e'''(2) by auto } note E = this note e1 = E[OF assms(2)] and e2 = E[OF perm_sym[OF assms(2)]] show ?thesis by (auto simp: Affine_def valuate_def aform_val_def assms dest: e1 e2) qed context includes autoref_syntax begin lemma aform_of_ivl_refine: "x \ y \ (aform_of_ivl x y, atLeastAtMost x y) \ \rnv_rel\aform_rel" by (auto simp: aform_rel_def br_def Affine_aform_of_ivl) lemma aforms_of_ivl_leI1: fixes en::real assumes "-1 \ en" "xsn \ ysn" shows "xsn \ (xsn + ysn) / 2 + (ysn - xsn) * en / 2" proof - have "xsn * (1 + en) \ ysn * (1 + en)" using assms mult_right_mono by force then show ?thesis by (auto simp: algebra_simps divide_simps) qed lemma aforms_of_ivl_leI2: fixes en::real assumes "1 \ en" "xsn \ ysn" shows "(xsn + ysn) / 2 + (ysn - xsn) * en / 2 \ ysn" proof - have "xsn * (1 - en) \ ysn * (1 - en)" using assms mult_right_mono by force then show ?thesis by (auto simp: algebra_simps divide_simps) qed lemma Joints_aforms_of_ivlsD1: "zs \ Joints (aforms_of_ivls xs ys) \ list_all2 (\) xs ys \ list_all2 (\) xs zs" by (auto simp: Joints_def valuate_def aforms_of_ivls_def aform_val_def Pi_iff list_all2_conv_all_nth intro!: list_all2_all_nthI aforms_of_ivl_leI1) lemma Joints_aforms_of_ivlsD2: "zs \ Joints (aforms_of_ivls xs ys) \ list_all2 (\) xs ys \ list_all2 (\) zs ys" by (auto simp: Joints_def valuate_def aforms_of_ivls_def aform_val_def Pi_iff list_all2_conv_all_nth intro!: list_all2_all_nthI aforms_of_ivl_leI2) lemma aforms_of_ivls_refine: "list_all2 (\) xrs yrs \ (xri, xrs) \ \rnv_rel\list_rel \ (yri, yrs) \ \rnv_rel\list_rel \ (aforms_of_ivls xri yri, lv_ivl xrs yrs) \ aforms_rel" apply (auto simp: aforms_rel_def br_def list_all2_lengthD lv_ivl_def Joints_aforms_of_ivlsD1 Joints_aforms_of_ivlsD2 intro!: aforms_of_ivls) subgoal by (simp add: list_all2_conv_all_nth) subgoal by (simp add: list_all2_conv_all_nth) done lemma degrees_zero_pdevs[simp]: "degrees (replicate n zero_pdevs) = 0" by (auto simp: degrees_def intro!: Max_eqI) lemma Joints_product_aforms: "Joints (product_aforms a b) = product_listset (Joints a) (Joints b)" apply (auto simp: Joints_def valuate_def product_listset_def product_aforms_def aform_vals_def[symmetric] aform_val_msum_aforms) subgoal for e apply (rule image_eqI[where x="(aform_vals e a, List.map2 (+) (aform_vals e (replicate (length b) (0, zero_pdevs))) (aform_vals (\i. e (i + degree_aforms a)) b))"]) apply (auto simp: split_beta') apply (auto simp: aform_vals_def intro!: nth_equalityI image_eqI[where x="\i. e (i + degree_aforms a)"]) done subgoal for e1 e2 apply (rule image_eqI[where x="\i. if i < degree_aforms a then e1 i else e2 (i - degree_aforms a)"]) apply (auto simp: aform_vals_def aform_val_def Pi_iff intro!: nth_equalityI pdevs_val_degree_cong) subgoal premises prems for x xs k proof - from prems have "degree xs \ degree_aforms a" by (auto simp: degrees_def Max_gr_iff) then show ?thesis using prems by auto qed done done lemma product_aforms_refine: "(product_aforms, product_listset) \ aforms_rel \ aforms_rel \ aforms_rel" by (auto simp: aforms_rel_def br_def Joints_product_aforms) lemma mem_lv_rel_set_rel_iff: fixes z::"'a::executable_euclidean_space set" shows "(y, z) \ \lv_rel\set_rel \ (z = eucl_of_list ` y \ (\x \ y. length x = DIM('a)))" unfolding lv_rel_def by (auto simp: set_rel_def br_def) lemma eucl_of_list_mem_lv_rel: "length x = DIM('a::executable_euclidean_space) \ (x, eucl_of_list x::'a) \ lv_rel" unfolding lv_rel_def by (auto simp: br_def) lemma mem_Joints_msum_aforms'I: "a \ Joints x \ b \ Joints y \ List.map2 (+) a b \ Joints (msum_aforms' x y)" by (auto simp: Joints_msum_aforms degrees_def) lemma mem_Joints_msum_aforms'E: assumes "xa \ Joints (msum_aforms' x y)" obtains a b where "xa = List.map2 (+) a b" "a \ Joints x" "b \ Joints y" using assms by (auto simp: Joints_msum_aforms degrees_def) lemma msum_aforms'_refine_raw: shows "(msum_aforms' x y, {List.map2 (+) a b|a b. a \ Joints x \ b \ Joints y}) \ aforms_rel" unfolding aforms_rel_def br_def by (safe elim!: mem_Joints_msum_aforms'E intro!: mem_Joints_msum_aforms'I) (auto simp: Joints_imp_length_eq) lemma aforms_relD: "(a, b) \ aforms_rel \ b = Joints a" by (auto simp: aforms_rel_def br_def) lemma msum_aforms'_refine: "(msum_aforms', \xs ys. {List.map2 (+) x y |x y. x \ xs \ y \ ys}) \ aforms_rel \ aforms_rel \ aforms_rel" by (safe dest!: aforms_relD intro!: msum_aforms'_refine_raw) lemma length_inf_aforms[simp]: "length (inf_aforms optns x) = length x" and length_sup_aforms[simp]: "length (sup_aforms optns x) = length x" by (auto simp: inf_aforms_def sup_aforms_def) lemma inf_aforms_refine: "(xi, x) \ aforms_rel \ length xi = d \ (RETURN (inf_aforms optns xi), Inf_specs d x) \ \rl_rel\nres_rel" unfolding o_def apply (auto simp: br_def aforms_rel_def mem_lv_rel_set_rel_iff nres_rel_def Inf_specs_def intro!: RETURN_SPEC_refine) unfolding lv_rel_def by (auto simp: aforms_rel_def br_def Joints_imp_length_eq eucl_of_list_inner inf_aforms_def nth_in_AffineI intro!: Inf_aform'_Affine_le list_all2_all_nthI) lemma sup_aforms_refine: "(xi, x) \ aforms_rel \ length xi = d \ (RETURN (sup_aforms optns xi), Sup_specs d x) \ \rl_rel\nres_rel" unfolding o_def apply (auto simp: aforms_relp_def mem_lv_rel_set_rel_iff nres_rel_def Sup_specs_def intro!: RETURN_SPEC_refine) unfolding lv_rel_def by (auto simp: aforms_rel_def br_def Joints_imp_length_eq eucl_of_list_inner sup_aforms_def nth_in_AffineI intro!: Sup_aform'_Affine_ge list_all2_all_nthI) lemma nres_of_THE_DRES_le: assumes "\v. x = Some v \ RETURN v \ y" shows "nres_of (THE_DRES x) \ y" using assms by (auto simp: THE_DRES_def split: option.split) lemma degree_le_fresh_index: "a \ set A \ degree_aform a \ fresh_index_aforms A" by (auto simp: fresh_index_aforms_def intro!: Max_ge) lemma zero_in_JointsI: "xs \ Joints XS \ z = (0, zero_pdevs) \ 0 # xs \ Joints (z # XS)" by (auto simp: Joints_def valuate_def) lemma cancel_nonneg_pos_add_multI: "0 \ c + c * x" if "c \ 0" "1 + x \ 0" for c x::real proof - have "0 \ c + c * x \ 0 \ c * (1 + x)" by (auto simp: algebra_simps) also have "\ \ 0 \ 1 + x" using that by (auto simp: zero_le_mult_iff) finally show ?thesis using that by simp qed lemma Joints_map_split_aform: fixes x::"real aform list" shows "Joints x \ Joints (map (\a. fst (split_aform a i)) x) \ Joints (map (\b. snd (split_aform b i)) x)" unfolding subset_iff apply (intro allI impI) proof goal_cases case (1 xs) then obtain e where e: "e \ funcset UNIV {-1 .. 1}" "xs = map (aform_val e) x" by (auto simp: Joints_def valuate_def) consider "e i \ 0" | "e i \ 0" by arith then show ?case proof cases case 1 let ?e = "e(i:= 2 * e i - 1)" note e(2) also have "map (aform_val e) x = map (aform_val ?e) (map (\b. snd (split_aform b i)) x)" by (auto simp: aform_val_def split_aform_def Let_def divide_simps algebra_simps) also have "\ \ Joints (map (\b. snd (split_aform b i)) x)" using e \0 \ e i\ by (auto simp: Joints_def valuate_def Pi_iff intro!: image_eqI[where x = ?e]) finally show ?thesis by simp next case 2 let ?e = "e(i:= 2 * e i + 1)" note e(2) also have "map (aform_val e) x = map (aform_val ?e) (map (\b. fst (split_aform b i)) x)" by (auto simp: aform_val_def split_aform_def Let_def divide_simps algebra_simps) also have "\ \ Joints (map (\b. fst (split_aform b i)) x)" using e \0 \ e i\ by (auto simp: Joints_def valuate_def Pi_iff real_0_le_add_iff intro!: image_eqI[where x = ?e] cancel_nonneg_pos_add_multI) finally show ?thesis by simp qed qed lemma split_aforms_lemma: fixes x::"real aform list" assumes "(x, y) \ aforms_rel" assumes "z \ y" assumes "l = (length x)" shows "\a b. (split_aforms x i, a, b) \ aforms_rel \\<^sub>r aforms_rel \ env_len a l \ env_len b l \ z \ a \ b" using assms apply (auto simp: split_aforms_def o_def) apply (rule exI[where x="Joints (map (\x. fst (split_aform x i)) x)"]) apply auto subgoal by (auto intro!: brI simp: aforms_rel_def) apply (rule exI[where x="Joints (map (\x. snd (split_aform x i)) x)"]) apply (rule conjI) subgoal by (auto intro!: brI simp: aforms_rel_def) subgoal using Joints_map_split_aform[of x i] by (auto simp: br_def aforms_rel_def env_len_def Joints_imp_length_eq) done lemma length_summarize_pdevs_list[simp]: "length (summarize_pdevs_list a b c d) = length d" by (auto simp: summarize_pdevs_list_def) lemma length_summarize_aforms[simp]: "length (summarize_aforms a b e d) = length d" by (auto simp: summarize_aforms_def Let_def) lemma split_aform_largest_take_refine: "(ni, n) \ nat_rel \ (xi::real aform list, x) \ aforms_rel \ length xi = d \ (RETURN (split_aforms_largest_uncond_take ni xi), split_spec_params d n x) \ \aforms_rel \\<^sub>r aforms_rel\nres_rel" apply (auto simp: split_spec_params_def nres_rel_def aforms_relp_def mem_lv_rel_set_rel_iff split_aforms_largest_uncond_take_def Let_def comps split: prod.splits intro!: RETURN_SPEC_refine) apply (rule split_aforms_lemma) by (auto simp add: aforms_rel_def) lemma aform_val_pdevs_of_real[simp]: "aform_val e (pdevs_of_real x) = x" by (auto simp: pdevs_of_real_def) lemma degree_aform_pdevs_of_real[simp]: "degree_aform (pdevs_of_real x) = 0" by (auto simp: pdevs_of_real_def) lemma interpret_floatarith_inner_eq: shows "interpret_floatarith (inner_floatariths xs ys) vs = (\(x, y) \ (zip xs ys). (interpret_floatarith x vs) * (interpret_floatarith y vs))" by (induction xs ys rule: inner_floatariths.induct) auto lemma approx_slp_outer_sing: "approx_slp_outer p (Suc 0) [fa] XS = Some R \ (\Y. approx_floatarith p fa (map (\x. (x, 0)) XS) = Some Y \ [aform_err_to_aform Y (max (degree_aforms XS) (degree_aform_err Y))] = R)" by (auto simp: approx_slp_outer_def bind_eq_Some_conv degrees_def) lemma aforms_err_aform_valsI: assumes "vs = aform_vals e XS" shows "vs \ aforms_err e (map (\x. (x, 0)) (XS))" by (auto simp: assms aforms_err_def o_def aform_err_def aform_vals_def) lemma aform_val_degree_cong: "b = d \ (\i. i < degree_aform d \ a i = c i) \ aform_val a b = aform_val c d" by (auto simp: aform_val_def intro!: pdevs_val_degree_cong) lemma mem_degree_aformD: "x \ set XS \ degree_aform x \ degree_aforms XS" by (auto simp: degrees_def) lemma degrees_append_leD1: "(degrees xs) \ degrees (xs @ ys)" unfolding degrees_def by (rule Max_mono) (auto simp: degrees_def min_def max_def Max_ge_iff image_Un Max_gr_iff) lemma inner_aforms': assumes "xs \ Joints XS" assumes "inner_aforms' p XS (map pdevs_of_real rs) = Some R" shows "(\(x, y) \ (zip xs rs). x * y) \ Affine (hd R)" (is ?th1) "length R = 1" (is ?th2) proof - from assms obtain e where "e \ funcset UNIV {-1 .. 1}" "xs = aform_vals e XS" by (auto simp: Joints_def valuate_def aform_vals_def) then have e: "xs @ rs = aform_vals e (XS @ map pdevs_of_real rs)" "xs @ rs \ Joints (XS @ map pdevs_of_real rs)" "length xs = length XS" "e \ funcset UNIV {- 1..1}" by (auto simp: aform_vals_def o_def degrees_def Joints_def valuate_def) have "approx_slp_outer p (Suc 0) ([inner_floatariths (map floatarith.Var [0..x. (x, 0)) (XS @ map pdevs_of_real rs)) = Some Y" "R = [aform_err_to_aform Y (max (degree_aforms (XS @ map pdevs_of_real rs)) (degree_aform_err Y))]" unfolding approx_slp_outer_sing by auto let ?m = "(max (degree_aforms (XS @ map pdevs_of_real rs)) (degree_aform_err Y))" from approx_floatarith_Elem[OF Y(1) e(4) aforms_err_aform_valsI[OF e(1)]] have "interpret_floatarith (inner_floatariths (map floatarith.Var [0.. aform_err e Y" "degree_aform_err Y \ max (degree_aforms (XS @ map pdevs_of_real rs)) (degree_aform_err Y)" by auto from aform_err_to_aformE[OF this] obtain err where err: "interpret_floatarith (inner_floatariths (map floatarith.Var [0.. err" "err \ 1" by auto let ?e' = "(e(max (degrees (map snd XS @ map (snd \ pdevs_of_real) rs)) (degree_aform_err Y) := err))" from e(1) have e': "xs @ rs = aform_vals ?e' (XS @ map pdevs_of_real rs)" apply (auto simp: aform_vals_def intro!: aform_val_degree_cong) apply (frule mem_degree_aformD) apply (frule le_less_trans[OF degrees_append_leD1]) apply (auto simp: o_def) done from err have "interpret_floatarith (inner_floatariths (map floatarith.Var [0.. Joints (R @ XS @ map pdevs_of_real rs)" using e(1,3,4) e' apply (auto simp: valuate_def Joints_def intro!: nth_equalityI image_eqI[where x="?e'"]) apply (simp add: Y aform_vals_def; fail) apply (simp add: Y o_def) apply (simp add: nth_append nth_Cons) apply (auto split: nat.splits simp: nth_append nth_Cons aform_vals_def) done then have "(\(x, y)\zip (map floatarith.Var [0.. Joints (R @ XS @ map pdevs_of_real rs)" apply (subst (asm)interpret_floatarith_inner_eq ) apply (auto simp: ) done also have "(\(x, y)\zip (map floatarith.Var [0..(x, y)\zip xs rs. x * y)" by (auto simp: sum_list_sum_nth assms e(3) nth_append intro!: sum.cong) finally show ?th1 ?th2 by (auto simp: Affine_def valuate_def Joints_def Y) qed lemma inner_aforms'_inner_lv_rel: "(a, a') \ aforms_rel \ inner_aforms' prec a (map pdevs_of_real a'a) = Some R \ x \ a' \ inner_lv_rel x a'a \ Affine (hd R)" unfolding mem_lv_rel_set_rel_iff unfolding lv_rel_def aforms_rel_def apply (auto simp: br_def) apply (subst arg_cong2[where f="(\)", OF _ refl]) defer apply (rule inner_aforms') apply (auto simp: br_def Joints_imp_length_eq inner_lv_rel_def) done lemma aform_inf_inner_refine: "(RETURN o2 aform_inf_inner optns, Inf_inners) \ aforms_rel \ rl_rel \ \rnv_rel\nres_rel" by (auto simp: aforms_relp_def nres_rel_def Inf_inners_def aform_inf_inner_def[abs_def] comp2_def iaN intro!: Inf_aform'_Affine_le inner_aforms'_inner_lv_rel split: option.splits list.splits) lemma aform_sup_inner_refine: "(RETURN o2 aform_sup_inner optns, Sup_inners) \ aforms_rel \ rl_rel \ \rnv_rel\nres_rel" by (auto simp: aforms_relp_def nres_rel_def Sup_inners_def aform_sup_inner_def[abs_def] comp2_def iaN intro!: Sup_aform'_Affine_ge inner_aforms'_inner_lv_rel split: option.splits list.splits) lemma lv_aforms_rel_comp_br_Affine_le: "lv_aforms_rel O br Affine top \ \lv_rel\aforms_relp" apply (auto simp: lv_aforms_rel_def aforms_relp_def br_def) apply (rule relcompI) apply (auto simp: aforms_rel_def intro!: brI ) by (auto simp: mem_lv_rel_set_rel_iff Joints_imp_length_eq intro!: eucl_of_list_mem_eucl_of_list_aform intro!: in_image_eucl_of_list_eucl_of_list_aform) lemma bijective_lv_rel[relator_props]: "bijective lv_rel" unfolding lv_rel_def bijective_def apply (auto simp: br_def) by (metis eucl_of_list_inj) lemma sv_lv_rel_inverse[relator_props]: "single_valued (lv_rel\)" using bijective_lv_rel by (rule bijective_imp_sv) lemma list_of_eucl_image_lv_rel_inverse: "(x, list_of_eucl ` x) \ \lv_rel\\set_rel" unfolding set_rel_sv[OF sv_lv_rel_inverse] apply (auto simp: ) apply (rule ImageI) apply (rule converseI) apply (rule lv_relI) apply auto apply (rule image_eqI) prefer 2 apply assumption unfolding lv_rel_def apply (auto simp: br_def) subgoal for y apply (rule exI[where x="list_of_eucl y"]) apply auto done done lemma lv_rel_comp_lv_rel_inverse: "((lv_rel::(_\'a::executable_euclidean_space) set) O lv_rel\) = {(x, y). x = y \ length x = DIM('a)}" apply (auto simp: intro!: lv_relI) unfolding lv_rel_def by (auto simp: br_def intro!: eucl_of_list_inj) lemma inter_aform_plane_refine_aux: "d = CARD('n::enum) \ (xi, x) \ aforms_rel \ (si, s) \ \rl_rel\sctn_rel \ length xi = d \ length (normal si) = d \ (nres_of (inter_aform_plane_lv (length xi) optns xi si), inter_sctn_specs d x s) \ \aforms_rel\nres_rel" proof (goal_cases) case 1 from 1 have sis: "si = s" by (cases si; cases s) (auto simp: sctn_rel_def) have Dp: "DIM_precond TYPE('n rvec) CARD('n)" by auto have a: "(xi, eucl_of_list_aform xi::'n rvec aform) \ lv_aforms_rel" by (auto simp: lv_aforms_rel_def br_def aforms_relp_def aforms_rel_def mem_lv_rel_set_rel_iff 1 Joints_imp_length_eq) have b: "(si, (Sctn (eucl_of_list (normal si)) (pstn si))::'n rvec sctn) \ \lv_rel\sctn_rel" using 1 by (cases si) (auto simp: lv_aforms_rel_def br_def aforms_relp_def aforms_rel_def mem_lv_rel_set_rel_iff Joints_imp_length_eq sctn_rel_def intro!: lv_relI) have a: "(nres_of (inter_aform_plane_lv CARD('n) optns xi si), inter_aform_plane optns (eucl_of_list_aform xi::'n rvec aform) (Sctn (eucl_of_list (normal si)) (pstn si))) \ \lv_aforms_rel\nres_rel" (is "(_, inter_aform_plane _ ?ea ?se) \ _") using inter_aform_plane_lv.refine[OF Dp IdI a b, of optns] by simp have b: "(inter_aform_plane optns ?ea ?se, inter_sctn_spec (Affine ?ea) ?se) \ \br Affine top\nres_rel" using 1 apply (auto simp: inter_sctn_spec_def nres_rel_def inter_aform_plane_def project_ncoord_aform_def inter_aform_plane_ortho_nres_def split: option.splits intro!: RETURN_SPEC_refine dest!: inter_inter_aform_plane_ortho) apply (auto simp: aform_rel_def br_def nres_rel_def comps bind_eq_Some_conv inter_sctn_spec_def inter_aform_plane_def plane_of_def aforms_relp_def aforms_rel_def mem_lv_rel_set_rel_iff intro!: RETURN_SPEC_refine nres_of_THE_DRES_le mem_Affine_project_coord_aformI dest!: inter_inter_aform_plane_ortho split: if_splits) apply (auto dest!: mem_Affine_project_coord_aformD simp: abs_in_Basis_iff plane_of_def) done from relcompI[OF a b] have "(nres_of (inter_aform_plane_lv CARD('n) optns xi si), inter_sctn_spec (Affine ?ea) ?se) \ \lv_aforms_rel\nres_rel O \br Affine top\nres_rel" by auto also have "\lv_aforms_rel\nres_rel O \br Affine top\nres_rel \ \\lv_rel\aforms_relp\nres_rel" unfolding nres_rel_comp apply (rule nres_rel_mono) apply (rule lv_aforms_rel_comp_br_Affine_le) done finally have step1: "(nres_of (inter_aform_plane_lv CARD('n) optns xi si), inter_sctn_spec (Affine (eucl_of_list_aform xi)::'n rvec set) (Sctn (eucl_of_list (normal si)) (pstn si))) \ \\lv_rel\aforms_relp\nres_rel" by simp have step2: "(inter_sctn_spec (Affine (eucl_of_list_aform xi)::'n rvec set) (Sctn (eucl_of_list (normal si)) (pstn si)), inter_sctn_specs CARD('n) (Joints xi) si) \ \\lv_rel\\set_rel\nres_rel" apply (auto simp: inter_sctn_specs_def inter_sctn_spec_def simp: nres_rel_def intro!: SPEC_refine) subgoal for x apply (rule exI[where x="list_of_eucl ` x"]) apply (auto simp: env_len_def plane_ofs_def intro: list_of_eucl_image_lv_rel_inverse) subgoal for y apply (rule image_eqI[where x="eucl_of_list y"]) apply (subst list_of_eucl_eucl_of_list) apply (auto simp: 1 Joints_imp_length_eq) apply (rule subsetD, assumption) apply (auto simp: intro!: eucl_of_list_mem_eucl_of_list_aform 1) using inner_lv_rel_autoref[where 'a="'n rvec", param_fo, OF lv_relI lv_relI, of y "normal si"] by (auto simp: plane_of_def 1 Joints_imp_length_eq) subgoal for y apply (drule subsetD, assumption) using inner_lv_rel_autoref[where 'a="'n rvec", param_fo, OF lv_relI lv_relI, of "list_of_eucl y" "normal si"] by (auto simp: plane_of_def 1 Joints_imp_length_eq) done done from relcompI[OF step1 step2] have "(nres_of (inter_aform_plane_lv CARD('n) optns xi si), inter_sctn_specs CARD('n) (Joints xi) si) \ \\lv_rel::(real list \ 'n rvec)set\aforms_relp\nres_rel O \\lv_rel\\set_rel\nres_rel" by simp also have "\\lv_rel::(real list \ 'n rvec)set\aforms_relp\nres_rel O \\lv_rel\\set_rel\nres_rel \ \aforms_rel O Id\nres_rel" unfolding nres_rel_comp O_assoc aforms_relp_def apply (rule nres_rel_mono) apply (rule relcomp_mono[OF order_refl]) unfolding set_rel_sv[OF sv_lv_rel_inverse] set_rel_sv[OF lv_rel_sv] apply (auto simp: length_lv_rel) unfolding relcomp_Image[symmetric] lv_rel_comp_lv_rel_inverse apply (auto simp: Basis_not_uminus length_lv_rel) unfolding lv_rel_def subgoal for a b c d apply (auto dest!: brD simp: length_lv_rel) using eucl_of_list_inj[where 'a="'n rvec", of d b] by auto done finally show ?case using 1 by (simp add: aforms_rel_def br_def sis) qed end setup \Sign.add_const_constraint (@{const_name "enum_class.enum"}, SOME @{typ "'a list"})\ setup \Sign.add_const_constraint (@{const_name "enum_class.enum_all"}, SOME @{typ "('a \ bool) \ bool"})\ setup \Sign.add_const_constraint (@{const_name "enum_class.enum_ex"}, SOME @{typ "('a \ bool) \ bool"})\ lemmas inter_aform_plane_refine_unoverloaded0 = inter_aform_plane_refine_aux[internalize_sort "'n::enum", unoverload enum_class.enum, unoverload enum_class.enum_all, unoverload enum_class.enum_ex] theorem inter_aform_plane_refine_unoverloaded: "class.enum (enum::'a list) enum_all enum_ex \ d = CARD('a) \ (xi, x) \ aforms_rel \ (si, s) \ \rl_rel\sctn_rel \ length xi = d \ length (normal si) = d \ (nres_of (inter_aform_plane_lv (length xi) optns xi si), inter_sctn_specs d x s) \ \aforms_rel\nres_rel" by (rule inter_aform_plane_refine_unoverloaded0) auto setup \Sign.add_const_constraint (@{const_name "enum_class.enum"}, SOME @{typ "'a::enum list"})\ setup \Sign.add_const_constraint (@{const_name "enum_class.enum_all"}, SOME @{typ "('a::enum \ bool) \ bool"})\ setup \Sign.add_const_constraint (@{const_name "enum_class.enum_ex"}, SOME @{typ "('a::enum \ bool) \ bool"})\ context includes autoref_syntax begin text \TODO: this is a special case of \Cancel_Card_Constraint\ from \AFP/Perron_Frobenius\!\ lemma type_impl_card_n_enum: assumes "\(Rep :: 'a \ nat) Abs. type_definition Rep Abs {0 ..< n :: nat}" obtains enum enum_all enum_ex where "class.enum (enum::'a list) enum_all enum_ex \ n = CARD('a)" proof - - from assms obtain rep :: "'a \ nat" and abs :: "nat \ 'a" where t: "type_definition rep abs {0 ..< n}" by auto - have "card (UNIV :: 'a set) = card {0 ..< n}" using t by (rule type_definition.card) + from assms obtain rep :: "'a \ nat" and abs :: "nat \ 'a" + where t: "type_definition rep abs {0 ..{0 .. . + have "card (UNIV :: 'a set) = card {0 ..< n}" + by (rule card) also have "\ = n" by auto finally have bn: "CARD ('a) = n" . let ?enum = "(map abs [0.. n = CARD('a)" by simp then show ?thesis .. qed lemma inter_aform_plane_refine_ex_typedef: "(xi, x) \ aforms_rel \ (si, s) \ \rl_rel\sctn_rel \ length xi = d \ length (normal si) = d \ (nres_of (inter_aform_plane_lv (length xi) optns xi si), inter_sctn_specs d x s) \ \aforms_rel\nres_rel" if "\(Rep :: 'a \ nat) Abs. type_definition Rep Abs {0 ..< d :: nat}" by (rule type_impl_card_n_enum[OF that]) (rule inter_aform_plane_refine_unoverloaded; assumption) lemma inter_aform_plane_refine: "0 < d \ (xi, x) \ aforms_rel \ (si, s) \ \Id\sctn_rel \ length xi = d \ length (normal si) = d \ (nres_of (inter_aform_plane_lv (length xi) optns xi si), inter_sctn_specs d x s) \ \aforms_rel\nres_rel" by (rule inter_aform_plane_refine_ex_typedef[cancel_type_definition, simplified]) lemma Joints_reduce_aforms: "x \ Joints X \ x \ Joints (reduce_aforms prec t X)" proof (auto simp: reduce_aforms_def summarize_threshold_def[abs_def] Joints_def valuate_def aform_val_def, goal_cases) case (1 e) from summarize_aformsE[OF \e \ _\ order_refl, of "X" "prec" t] guess e' . thus ?case by (auto intro!: image_eqI[where x=e'] simp: aform_vals_def) qed lemma length_reduce_aform[simp]: "length (reduce_aforms optns a x) = length x" by (auto simp: reduce_aforms_def) lemma reduce_aform_refine: "(xi, x) \ aforms_rel \ length xi = d \ (RETURN (reduce_aforms prec C xi), reduce_specs d r x) \ \aforms_rel\nres_rel" apply (auto simp: reduce_specs_def nres_rel_def comps aforms_relp_def mem_lv_rel_set_rel_iff aforms_rel_def env_len_def intro!: RETURN_SPEC_refine) apply (auto simp: br_def env_len_def) by (auto simp: mem_lv_rel_set_rel_iff Joints_imp_length_eq intro!: in_image_eucl_of_list_eucl_of_list_aform Joints_reduce_aforms eucl_of_list_mem_eucl_of_list_aform) lemma aform_euclarithform_refine: "(nres_of o2 aform_form optns, approx_form_spec) \ Id \ aforms_rel \ \bool_rel\nres_rel" by (auto simp: approx_form_spec_def nres_rel_def comps aform_form_def aforms_rel_def br_def approx_form_aform dest!: approx_form_aux intro!: ivls_of_aforms) lemma aform_isFDERIV: "(\N xs fas vs. nres_of (aform_isFDERIV optns N xs fas vs), isFDERIV_spec) \ nat_rel \ \nat_rel\list_rel \ \Id\list_rel \ aforms_rel \ \bool_rel\nres_rel" by (auto simp: isFDERIV_spec_def nres_rel_def comps aform_isFDERIV_def aforms_rel_def br_def dest!: approx_form_aux intro!: ivls_of_aforms isFDERIV_approx) lemma approx_slp_lengthD: "approx_slp p slp a = Some xs \ length xs = length slp + length a" by (induction slp arbitrary: xs a) (auto simp: bind_eq_Some_conv) lemma approx_slp_outer_lengthD: "approx_slp_outer p d slp a = Some xs \ length xs = min d (length slp + length a)" by (auto simp: approx_slp_outer_def Let_def bind_eq_Some_conv o_def aforms_err_to_aforms_def aform_err_to_aform_def dest!: approx_slp_lengthD) lemma approx_slp_refine: "(nres_of o3 aform_slp prec, approx_slp_spec fas) \ nat_rel \ fas_rel \ aforms_rel \ \\aforms_rel\option_rel\nres_rel" apply (auto simp: approx_slp_spec_def comps aform_slp_def nres_rel_def intro!: RETURN_SPEC_refine ASSERT_refine_right) subgoal for a b apply (rule exI[where x = "map_option Joints (approx_slp_outer prec (length fas) (slp_of_fas fas) a)"]) apply (auto simp: option.splits aforms_rel_def br_def env_len_def Joints_imp_length_eq) apply (auto dest!: approx_slp_outer_lengthD)[] using length_slp_of_fas_le trans_le_add1 approx_slp_outer_lengthD apply blast using approx_slp_outer_plain by blast done lemma fresh_index_aforms_Nil[simp]: "fresh_index_aforms [] = 0" by (auto simp: fresh_index_aforms_def) lemma independent_aforms_Nil[simp]: "independent_aforms x [] = [x]" by (auto simp: independent_aforms_def) lemma mem_Joints_zero_iff[simp]: "x # xs \ Joints ((0, zero_pdevs) # XS) \ (x = 0 \ xs \ Joints XS)" by (auto simp: Joints_def valuate_def) lemma Joints_independent_aforms_eq: "Joints (independent_aforms x xs) = set_Cons (Affine x) (Joints xs)" by (simp add: independent_aforms_def Joints_msum_aform degree_le_fresh_index set_Cons_def) lemma independent_aforms_refine: "(independent_aforms, set_Cons) \ \rnv_rel\aform_rel \ aforms_rel \ aforms_rel" by (auto simp: aforms_rel_def br_def aform_rel_def Joints_independent_aforms_eq) end locale aform_approximate_sets = approximate_sets aform_ops Joints aforms_rel begin lemma Joints_in_lv_rel_set_relD: "(Joints xs, X) \ \lv_rel\set_rel \ X = Affine (eucl_of_list_aform xs)" unfolding lv_rel_def set_rel_br by (auto simp: br_def Joints_imp_length_eq eucl_of_list_image_Joints[symmetric]) lemma ncc_precond: "ncc_precond TYPE('a::executable_euclidean_space)" unfolding ncc_precond_def ncc_def appr_rel_def by (auto simp: aforms_rel_def compact_Affine convex_Affine dest!: Joints_in_lv_rel_set_relD brD) lemma fst_eucl_of_list_aform_map: "fst (eucl_of_list_aform (map (\x. (fst x, asdf x)) x)) = fst (eucl_of_list_aform x)" by (auto simp: eucl_of_list_aform_def o_def) lemma Affine_pdevs_of_list:\ \TODO: move!\ "Affine (fst x, pdevs_of_list (map snd (list_of_pdevs (snd x)))) = Affine x" by (auto simp: Affine_def valuate_def aform_val_def elim: pdevs_val_of_list_of_pdevs2[where X = "snd x"] pdevs_val_of_list_of_pdevs[where X = "snd x"]) end lemma aform_approximate_sets: "aform_approximate_sets prec" apply (unfold_locales) unfolding aform_ops_def approximate_set_ops.simps subgoal unfolding relAPP_def aforms_rel_def . subgoal by (force simp: aforms_of_ivls_refine) subgoal by (rule product_aforms_refine) subgoal by (rule msum_aforms'_refine) subgoal by (rule inf_aforms_refine) subgoal by (rule sup_aforms_refine) subgoal by (rule split_aform_largest_take_refine) subgoal by (rule aform_inf_inner_refine) subgoal by (rule aform_sup_inner_refine) subgoal by (rule inter_aform_plane_refine) simp_all subgoal by (auto split: option.splits intro!: reduce_aform_refine) subgoal by (force simp: width_spec_def nres_rel_def) subgoal by (rule approx_slp_refine) subgoal by (rule aform_euclarithform_refine) subgoal by (rule aform_isFDERIV) subgoal by simp subgoal by (auto simp: Joints_imp_length_eq) subgoal by (force simp: Affine_def Joints_def valuate_def intro!:) subgoal by (force simp: Affine_def Joints_def valuate_def intro!:) subgoal by (auto simp: Joints_imp_length_eq) done 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,387 @@ (* 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 + 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)+ definition uint64_of_nat_conv where [simp]: \uint64_of_nat_conv (x :: nat) = x\ lemma less_upper_bintrunc_id: \n < 2 ^b \ n \ 0 \ bintrunc b n = n\ unfolding uint32_of_nat_def by (simp add: no_bintr_alt1) 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) lemma [sepref_fr_rules]: \(return o uint64_of_nat, RETURN o uint64_of_nat_conv) \ [\a. a < 2 ^64]\<^sub>a nat_assn\<^sup>k \ uint64_nat_assn\ by sepref_to_hoare (sep_auto simp: uint64_nat_rel_def br_def nat_of_uint64_uint64_of_nat_id) definition string_rel :: \(String.literal \ string) set\ where \string_rel = {(x, y). y = String.explode x}\ abbreviation string_assn :: \string \ String.literal \ assn\ where \string_assn \ pure string_rel\ lemma eq_string_eq: \((=), (=)) \ string_rel \ string_rel \ bool_rel\ by (auto intro!: frefI simp: string_rel_def String.less_literal_def less_than_char_def rel2p_def literal.explode_inject) lemmas eq_string_eq_hnr = eq_string_eq[sepref_import_param] definition string2_rel :: \(string \ string) set\ where \string2_rel \ \Id\list_rel\ abbreviation string2_assn :: \string \ string \ assn\ where \string2_assn \ pure string2_rel\ abbreviation monom_rel where \monom_rel \ \string_rel\list_rel\ abbreviation monom_assn where \monom_assn \ list_assn string_assn\ abbreviation monomial_rel where \monomial_rel \ monom_rel \\<^sub>r int_rel\ abbreviation monomial_assn where \monomial_assn \ monom_assn \\<^sub>a int_assn\ abbreviation poly_rel where \poly_rel \ \monomial_rel\list_rel\ abbreviation poly_assn where \poly_assn \ list_assn monomial_assn\ lemma poly_assn_alt_def: \poly_assn = pure poly_rel\ by (simp add: list_assn_pure_conv) abbreviation polys_assn where \polys_assn \ hm_fmap_assn uint64_nat_assn poly_assn\ lemma string_rel_string_assn: \(\ ((c, a) \ string_rel)) = string_assn a c\ by (auto simp: pure_app_eq) lemma single_valued_string_rel: \single_valued string_rel\ by (auto simp: single_valued_def string_rel_def) lemma IS_LEFT_UNIQUE_string_rel: \IS_LEFT_UNIQUE string_rel\ by (auto simp: IS_LEFT_UNIQUE_def single_valued_def string_rel_def literal.explode_inject) lemma IS_RIGHT_UNIQUE_string_rel: \IS_RIGHT_UNIQUE string_rel\ by (auto simp: single_valued_def string_rel_def literal.explode_inject) lemma single_valued_monom_rel: \single_valued monom_rel\ by (rule list_rel_sv) (auto intro!: frefI simp: string_rel_def rel2p_def single_valued_def p2rel_def) lemma single_valued_monomial_rel: \single_valued monomial_rel\ using single_valued_monom_rel by (auto intro!: frefI simp: rel2p_def single_valued_def p2rel_def) lemma single_valued_monom_rel': \IS_LEFT_UNIQUE monom_rel\ unfolding IS_LEFT_UNIQUE_def inv_list_rel_eq string2_rel_def by (rule list_rel_sv)+ (auto intro!: frefI simp: string_rel_def rel2p_def single_valued_def p2rel_def literal.explode_inject) lemma single_valued_monomial_rel': \IS_LEFT_UNIQUE monomial_rel\ using single_valued_monom_rel' unfolding IS_LEFT_UNIQUE_def inv_list_rel_eq by (auto intro!: frefI simp: rel2p_def single_valued_def p2rel_def) lemma [safe_constraint_rules]: \Sepref_Constraints.CONSTRAINT single_valued string_rel\ \Sepref_Constraints.CONSTRAINT IS_LEFT_UNIQUE string_rel\ by (auto simp: CONSTRAINT_def single_valued_def string_rel_def IS_LEFT_UNIQUE_def literal.explode_inject) lemma eq_string_monom_hnr[sepref_fr_rules]: \(uncurry (return oo (=)), uncurry (RETURN oo (=))) \ monom_assn\<^sup>k *\<^sub>a monom_assn\<^sup>k \\<^sub>a bool_assn\ using single_valued_monom_rel' single_valued_monom_rel unfolding list_assn_pure_conv by sepref_to_hoare (sep_auto simp: list_assn_pure_conv string_rel_string_assn single_valued_def IS_LEFT_UNIQUE_def dest!: mod_starD simp flip: inv_list_rel_eq) definition term_order_rel' where [simp]: \term_order_rel' x y = ((x, y) \ term_order_rel)\ lemma term_order_rel[def_pat_rules]: \(\)$(x,y)$term_order_rel \ term_order_rel'$x$y\ by auto lemma term_order_rel_alt_def: \term_order_rel = lexord (p2rel char.lexordp)\ by (auto simp: p2rel_def char.lexordp_conv_lexord var_order_rel_def intro!: arg_cong[of _ _ lexord]) instantiation char :: linorder begin definition less_char where [symmetric, simp]: "less_char = PAC_Polynomials_Term.less_char" definition less_eq_char where [symmetric, simp]: "less_eq_char = PAC_Polynomials_Term.less_eq_char" instance apply standard using char.linorder_axioms by (auto simp: class.linorder_def class.order_def class.preorder_def less_eq_char_def less_than_char_def class.order_axioms_def class.linorder_axioms_def p2rel_def less_char_def) end instantiation list :: (linorder) linorder begin definition less_list where "less_list = lexordp (<)" definition less_eq_list where "less_eq_list = lexordp_eq" instance proof standard have [dest]: \\x y :: 'a :: linorder list. (x, y) \ lexord {(x, y). x < y} \ lexordp_eq y x \ False\ by (metis lexordp_antisym lexordp_conv_lexord lexordp_eq_conv_lexord) have [simp]: \\x y :: 'a :: linorder list. lexordp_eq x y \ \ lexordp_eq y x \ (x, y) \ lexord {(x, y). x < y}\ using lexordp_conv_lexord lexordp_conv_lexordp_eq by blast show \(x < y) = Restricted_Predicates.strict (\) x y\ \x \ x\ \x \ y \ y \ z \ x \ z\ \x \ y \ y \ x \ x = y\ \x \ y \ y \ x\ for x y z :: \'a :: linorder list\ by (auto simp: less_list_def less_eq_list_def List.lexordp_def lexordp_conv_lexord lexordp_into_lexordp_eq lexordp_antisym antisym_def lexordp_eq_refl lexordp_eq_linear intro: lexordp_eq_trans dest: lexordp_eq_antisym) qed end lemma term_order_rel'_alt_def_lexord: \term_order_rel' x y = ord_class.lexordp x y\ and term_order_rel'_alt_def: \term_order_rel' x y \ x < y\ proof - show \term_order_rel' x y = ord_class.lexordp x y\ \term_order_rel' x y \ x < y\ unfolding less_than_char_of_char[symmetric, abs_def] by (auto simp: lexordp_conv_lexord less_eq_list_def less_list_def lexordp_def var_order_rel_def rel2p_def term_order_rel_alt_def p2rel_def) qed lemma list_rel_list_rel_order_iff: assumes \(a, b) \ \string_rel\list_rel\ \(a', b') \ \string_rel\list_rel\ shows \a < a' \ b < b'\ proof have H: \(a, b) \ \string_rel\list_rel \ (a, cs) \ \string_rel\list_rel \ b = cs\ for cs using single_valued_monom_rel' IS_RIGHT_UNIQUE_string_rel unfolding string2_rel_def by (subst (asm)list_rel_sv_iff[symmetric]) (auto simp: single_valued_def) assume \a < a'\ then consider u u' where \a' = a @ u # u'\ | u aa v w aaa where \a = u @ aa # v\ \a' = u @ aaa # w\ \aa < aaa\ by (subst (asm) less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) then show \b < b'\ proof cases case 1 then show \b < b'\ using assms by (subst less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff dest: H) next case 2 then obtain u' aa' v' w' aaa' where \b = u' @ aa' # v'\ \b' = u' @ aaa' # w'\ \(aa, aa') \ string_rel\ \(aaa, aaa') \ string_rel\ using assms by (smt list_rel_append1 list_rel_split_right_iff single_valued_def single_valued_monom_rel) with \aa < aaa\ have \aa' < aaa'\ by (auto simp: string_rel_def less_literal.rep_eq less_list_def lexordp_conv_lexord lexordp_def char.lexordp_conv_lexord simp flip: less_char_def PAC_Polynomials_Term.less_char_def) then show \b < b'\ using \b = u' @ aa' # v'\ \b' = u' @ aaa' # w'\ by (subst less_list_def) (fastforce simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) qed next have H: \(a, b) \ \string_rel\list_rel \ (a', b) \ \string_rel\list_rel \ a = a'\ for a a' b using single_valued_monom_rel' by (auto simp: single_valued_def IS_LEFT_UNIQUE_def simp flip: inv_list_rel_eq) assume \b < b'\ then consider u u' where \b' = b @ u # u'\ | u aa v w aaa where \b = u @ aa # v\ \b' = u @ aaa # w\ \aa < aaa\ by (subst (asm) less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) then show \a < a'\ proof cases case 1 then show \a < a'\ using assms by (subst less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append2 list_rel_split_left_iff dest: H) next case 2 then obtain u' aa' v' w' aaa' where \a = u' @ aa' # v'\ \a' = u' @ aaa' # w'\ \(aa', aa) \ string_rel\ \(aaa', aaa) \ string_rel\ using assms by (auto simp: lexord_def List.lexordp_def list_rel_append2 list_rel_split_left_iff dest: H) with \aa < aaa\ have \aa' < aaa'\ by (auto simp: string_rel_def less_literal.rep_eq less_list_def lexordp_conv_lexord lexordp_def char.lexordp_conv_lexord simp flip: less_char_def PAC_Polynomials_Term.less_char_def) then show \a < a'\ using \a = u' @ aa' # v'\ \a' = u' @ aaa' # w'\ by (subst less_list_def) (fastforce simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) qed qed lemma string_rel_le[sepref_import_param]: shows \((<), (<)) \ \string_rel\list_rel \ \string_rel\list_rel \ bool_rel\ by (auto intro!: fun_relI simp: list_rel_list_rel_order_iff) (* TODO Move *) lemma [sepref_import_param]: assumes \CONSTRAINT IS_LEFT_UNIQUE R\ \CONSTRAINT IS_RIGHT_UNIQUE R\ shows \(remove1, remove1) \ R \ \R\list_rel \ \R\list_rel\ apply (intro fun_relI) subgoal premises p for x y xs ys using p(2) p(1) assms by (induction xs ys rule: list_rel_induct) (auto simp: IS_LEFT_UNIQUE_def single_valued_def) done instantiation pac_step :: (heap, heap, heap) heap begin instance proof standard obtain f :: \'a \ nat\ where f: \inj f\ by blast obtain g :: \nat \ nat \ nat \ nat \ nat \ nat\ where g: \inj g\ by blast obtain h :: \'b \ nat\ where h: \inj h\ by blast obtain i :: \'c \ nat\ where i: \inj i\ by blast have [iff]: \g a = g b \ a = b\\h a'' = h b'' \ a'' = b''\ \f a' = f b' \ a' = b'\ \i a''' = i b''' \ a''' = b'''\ for a b a' b' a'' b'' a''' b''' using f g h i unfolding inj_def by blast+ let ?f = \\x :: ('a, 'b, 'c) pac_step. g (case x of Add a b c d \ (0, i a, i b, i c, f d) | Del a \ (1, i a, 0, 0, 0) | Mult a b c d \ (2, i a, f b, i c, f d) | Extension a b c \ (3, i a, f c, 0, h b))\ have \inj ?f\ apply (auto simp: inj_def) apply (case_tac x; case_tac y) apply auto done then show \\f :: ('a, 'b, 'c) pac_step \ nat. inj f\ by blast qed end end \ No newline at end of file diff --git a/thys/VerifyThis2018/Challenge1_short.thy b/thys/VerifyThis2018/Challenge1_short.thy --- a/thys/VerifyThis2018/Challenge1_short.thy +++ b/thys/VerifyThis2018/Challenge1_short.thy @@ -1,198 +1,197 @@ section \Shorter Solution\ theory Challenge1_short imports "lib/VTcomp" begin text \Small specification of textbuffer ADT, and its implementation by a gap buffer. Annotated and elaborated version of just the challenge requirements. \ subsection \Abstract Specification\ datatype 'a textbuffer = BUF ("pos": "nat") ("text": "'a list") \ \Note that we do not model the abstract invariant --- pos in range --- here, as it is not strictly required for the challenge spec.\ text \These are the operations that were specified in the challenge. Note: Isabelle has type inference, so we do not need to specify types. Note: We exploit that, in Isabelle, we have @{lemma "(0::nat)-1=0" by simp}. \ primrec move_left where "move_left (BUF p t) = BUF (p-1) t" primrec move_right where "move_right (BUF p t) = BUF (min (length t) (p+1)) t" primrec insert where "insert x (BUF p t) = BUF (p+1) (take p t@x#drop p t)" primrec delete where "delete (BUF p t) = BUF (p-1) (take (p-1) t@drop p t)" subsection \Refinement 1: List with Gap\ subsection \Implementation on List-Level\ type_synonym 'a gap_buffer = "nat \ nat \ 'a list" subsubsection \Abstraction Relation\ text \We define an invariant on the concrete gap-buffer, and its mapping to the abstract model. From these two, we define a relation \gap_rel\ between concrete and abstract buffers. \ definition "gap_\ \ \(l,r,buf). BUF l (take l buf @ drop r buf)" definition "gap_invar \ \(l,r,buf). l\r \ r\length buf" abbreviation "gap_rel \ br gap_\ gap_invar" subsubsection \Left\ text \For the operations, we insert assertions. These are not required to prove the list-level specification correct (during the proof, they are inferred easily). However, they are required in the subsequent automatic refinement step to arrays, to give our tool the information that all indexes are, indeed, in bounds.\ definition "move_left1 \ \(l,r,buf). doN { if l\0 then doN { ASSERT(r-1 l-1 gap_rel \ \gap_rel\nres_rel" apply clarsimp unfolding move_left1_def apply refine_vcg apply (auto simp: in_br_conv gap_\_def gap_invar_def move_left1_def split: prod.splits) (* sledgehammer! *) by (smt Cons_nth_drop_Suc Suc_pred append.assoc append_Cons append_Nil diff_Suc_less drop_update_cancel hd_drop_conv_nth length_list_update less_le_trans nth_list_update_eq take_hd_drop) subsubsection \Right\ definition "move_right1 \ \(l,r,buf). doN { if r gap_rel \ \gap_rel\nres_rel" apply clarsimp unfolding move_right1_def apply refine_vcg unfolding gap_\_def gap_invar_def apply (auto simp: in_br_conv split: prod.split) - (* sledgehammer *) - by (metis Cons_eq_appendI Cons_nth_drop_Suc append_eq_append_conv2 - atd_lem drop_0 dual_order.strict_trans2 take_eq_Nil take_update_last) - (* Manual: by (simp add: hd_drop_conv_nth take_update_last Cons_nth_drop_Suc) *) - + apply (rule nth_equalityI) + apply (simp_all add: Cons_nth_drop_Suc take_update_last) + done + subsubsection \Insert and Grow\ definition "can_insert \ \(l,r,buf). l \(l,r,buf). doN { let b = op_array_replicate (length buf + K) default; b \ mop_list_blit buf 0 b 0 l; b \ mop_list_blit buf r b (r+K) (length buf - r); RETURN (l,r+K,b) }" \ \Note: Most operations have also a variant prefixed with \mop\. These are defined in the refinement monad and already contain the assertion of their precondition. The backside is that they cannot be easily used in as part of expressions, e.g., in @{term "buf[l:=buf!r]"}, we would have to explicitly bind each intermediate value: @{term "doN { v \ mop_list_get buf r; mop_list_set buf l v }"}. \ lemma grow1_correct[THEN SPEC_trans, refine_vcg]: \ \Declares this as a rule to be used by the VCG\ assumes "gap_invar gb" shows "grow1 K gb \ (SPEC (\gb'. gap_invar gb' \ gap_\ gb' = gap_\ gb \ (K>0 \ can_insert gb')))" unfolding grow1_def apply refine_vcg using assms unfolding gap_\_def gap_invar_def can_insert_def apply (auto simp: op_list_blit_def) done definition "insert1 x \ \(l,r,buf). doN { (l,r,buf) \ if (l=r) then grow1 (length buf+1) (l,r,buf) else RETURN (l,r,buf); ASSERT (l Id \ gap_rel \ \gap_rel\nres_rel" apply clarsimp unfolding insert1_def apply refine_vcg \ \VCG knows the rule for grow1 already\ unfolding insert_def gap_\_def gap_invar_def can_insert_def apply (auto simp: in_br_conv take_update_last split: prod.split) done subsubsection \Delete\ definition "delete1 \ \(l,r,buf). if l>0 then RETURN (l-1,r,buf) else RETURN (l,r,buf)" lemma delete1_correct: "(delete1,RETURN o delete) \ gap_rel \ \gap_rel\nres_rel" apply clarsimp unfolding delete1_def apply refine_vcg unfolding gap_\_def gap_invar_def by (auto simp: in_br_conv butlast_take split: prod.split) subsection \Imperative Arrays\ text \The following indicates how we will further refine the gap-buffer: The list will become an array, the indices and the content will not be refined (expressed by @{const nat_assn} and @{const id_assn}). \ abbreviation "gap_impl_assn \ nat_assn \\<^sub>a nat_assn \\<^sub>a array_assn id_assn" sepref_definition move_left_impl is move_left1 :: "gap_impl_assn\<^sup>d\\<^sub>agap_impl_assn" unfolding move_left1_def by sepref sepref_definition move_right_impl is move_right1 :: "gap_impl_assn\<^sup>d\\<^sub>agap_impl_assn" unfolding move_right1_def by sepref sepref_definition insert_impl is "uncurry insert1" :: "id_assn\<^sup>k*\<^sub>agap_impl_assn\<^sup>d\\<^sub>agap_impl_assn" unfolding insert1_def grow1_def by sepref \ \We inline @{const grow1} here\ sepref_definition delete_impl is delete1 :: "gap_impl_assn\<^sup>d\\<^sub>agap_impl_assn" unfolding delete1_def by sepref text \Finally, we combine the two refinement steps, to get overall correctness theorems\ definition "gap_assn \ hr_comp gap_impl_assn gap_rel" \ \@{const hr_comp} is composition of refinement relations\ context notes gap_assn_def[symmetric,fcomp_norm_unfold] begin lemmas move_left_impl_correct = move_left_impl.refine[FCOMP move_left1_correct] and move_right_impl_correct = move_right_impl.refine[FCOMP move_right1_correct] and insert_impl_correct = insert_impl.refine[FCOMP insert1_correct] and delete_impl_correct = delete_impl.refine[FCOMP delete1_correct] text \Proves: @{thm [display] move_left_impl_correct} @{thm [display] move_right_impl_correct} @{thm [display] insert_impl_correct} @{thm [display] delete_impl_correct} \ end subsection \Executable Code\ text \Isabelle/HOL can generate code in various target languages.\ export_code move_left_impl move_right_impl insert_impl delete_impl in SML_imp module_name Gap_Buffer in OCaml_imp module_name Gap_Buffer in Haskell module_name Gap_Buffer in Scala module_name Gap_Buffer end diff --git a/thys/WebAssembly/Wasm_Ast.thy b/thys/WebAssembly/Wasm_Ast.thy --- a/thys/WebAssembly/Wasm_Ast.thy +++ b/thys/WebAssembly/Wasm_Ast.thy @@ -1,180 +1,185 @@ section \WebAssembly Core AST\ -theory Wasm_Ast imports Main "Native_Word.Uint8" begin +theory Wasm_Ast + imports + Main + "Native_Word.Uint8" + "Word_Lib.Reversed_Bit_Lists" +begin type_synonym \ \immediate\ i = nat type_synonym \ \static offset\ off = nat type_synonym \ \alignment exponent\ a = nat \ \primitive types\ typedecl i32 typedecl i64 typedecl f32 typedecl f64 \ \memory\ type_synonym byte = uint8 typedef bytes = "UNIV :: (byte list) set" .. setup_lifting type_definition_bytes declare Quotient_bytes[transfer_rule] lift_definition bytes_takefill :: "byte \ nat \ bytes \ bytes" is "(\a n as. takefill (Abs_uint8 a) n as)" . lift_definition bytes_replicate :: "nat \ byte \ bytes" is "(\n b. replicate n (Abs_uint8 b))" . definition msbyte :: "bytes \ byte" where "msbyte bs = last (Rep_bytes bs)" typedef mem = "UNIV :: (byte list) set" .. setup_lifting type_definition_mem declare Quotient_mem[transfer_rule] lift_definition read_bytes :: "mem \ nat \ nat \ bytes" is "(\m n l. take l (drop n m))" . lift_definition write_bytes :: "mem \ nat \ bytes \ mem" is "(\m n bs. (take n m) @ bs @ (drop (n + length bs) m))" . lift_definition mem_append :: "mem \ bytes \ mem" is append . \ \host\ typedecl host typedecl host_state datatype \ \value types\ t = T_i32 | T_i64 | T_f32 | T_f64 datatype \ \packed types\ tp = Tp_i8 | Tp_i16 | Tp_i32 datatype \ \mutability\ mut = T_immut | T_mut record tg = \ \global types\ tg_mut :: mut tg_t :: t datatype \ \function types\ tf = Tf "t list" "t list" ("_ '_> _" 60) (* TYPING *) record t_context = types_t :: "tf list" func_t :: "tf list" global :: "tg list" table :: "nat option" memory :: "nat option" local :: "t list" label :: "(t list) list" return :: "(t list) option" record s_context = s_inst :: "t_context list" s_funcs :: "tf list" s_tab :: "nat list" s_mem :: "nat list" s_globs :: "tg list" datatype sx = S | U datatype unop_i = Clz | Ctz | Popcnt datatype unop_f = Neg | Abs | Ceil | Floor | Trunc | Nearest | Sqrt datatype binop_i = Add | Sub | Mul | Div sx | Rem sx | And | Or | Xor | Shl | Shr sx | Rotl | Rotr datatype binop_f = Addf | Subf | Mulf | Divf | Min | Max | Copysign datatype testop = Eqz datatype relop_i = Eq | Ne | Lt sx | Gt sx | Le sx | Ge sx datatype relop_f = Eqf | Nef | Ltf | Gtf | Lef | Gef datatype cvtop = Convert | Reinterpret datatype \ \values\ v = ConstInt32 i32 | ConstInt64 i64 | ConstFloat32 f32 | ConstFloat64 f64 datatype \ \basic instructions\ b_e = Unreachable | Nop | Drop | Select | Block tf "b_e list" | Loop tf "b_e list" | If tf "b_e list" "b_e list" | Br i | Br_if i | Br_table "i list" i | Return | Call i | Call_indirect i | Get_local i | Set_local i | Tee_local i | Get_global i | Set_global i | Load t "(tp \ sx) option" a off | Store t "tp option" a off | Current_memory | Grow_memory | EConst v ("C _" 60) | Unop_i t unop_i | Unop_f t unop_f | Binop_i t binop_i | Binop_f t binop_f | Testop t testop | Relop_i t relop_i | Relop_f t relop_f | Cvtop t cvtop t "sx option" datatype cl = \ \function closures\ Func_native i tf "t list" "b_e list" | Func_host tf host record inst = \ \instances\ types :: "tf list" funcs :: "i list" tab :: "i option" mem :: "i option" globs :: "i list" type_synonym tabinst = "(cl option) list" record global = g_mut :: mut g_val :: v record s = \ \store\ inst :: "inst list" funcs :: "cl list" tab :: "tabinst list" mem :: "mem list" globs :: "global list" datatype e = \ \administrative instruction\ Basic b_e ("$_" 60) | Trap | Callcl cl | Label nat "e list" "e list" | Local nat i "v list" "e list" datatype Lholed = \ \L0 = v* [] e*\ LBase "e list" "e list" \ \L(i+1) = v* (label n {e* } Li) e*\ | LRec "e list" nat "e list" Lholed "e list" end diff --git a/thys/Word_Lib/Aligned.thy b/thys/Word_Lib/Aligned.thy --- a/thys/Word_Lib/Aligned.thy +++ b/thys/Word_Lib/Aligned.thy @@ -1,1081 +1,929 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Word Alignment" theory Aligned imports - "HOL-Library.Word" - Reversed_Bit_Lists - More_Word - Word_EqI + "HOL-Library.Word" + More_Word + Word_EqI + Typedef_Morphisms begin lift_definition is_aligned :: \'a::len word \ nat \ bool\ is \\k n. 2 ^ n dvd take_bit LENGTH('a) k\ by simp lemma is_aligned_iff_udvd: \is_aligned w n \ 2 ^ n udvd w\ by transfer (simp flip: take_bit_eq_0_iff add: min_def) lemma is_aligned_iff_take_bit_eq_0: \is_aligned w n \ take_bit n w = 0\ by (simp add: is_aligned_iff_udvd take_bit_eq_0_iff exp_dvd_iff_exp_udvd) lemma is_aligned_iff_dvd_int: \is_aligned ptr n \ 2 ^ n dvd uint ptr\ by transfer simp lemma is_aligned_iff_dvd_nat: \is_aligned ptr n \ 2 ^ n dvd unat ptr\ proof - have \unat ptr = nat \uint ptr\\ by transfer simp then have \2 ^ n dvd unat ptr \ 2 ^ n dvd uint ptr\ by (simp only: dvd_nat_abs_iff) simp then show ?thesis by (simp add: is_aligned_iff_dvd_int) qed lemma is_aligned_0 [simp]: \is_aligned 0 n\ by transfer simp lemma is_aligned_at_0 [simp]: \is_aligned w 0\ by transfer simp lemma is_aligned_beyond_length: \is_aligned w n \ w = 0\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that apply (simp add: is_aligned_iff_udvd) apply transfer apply auto done lemma is_alignedI [intro?]: \is_aligned x n\ if \x = 2 ^ n * k\ for x :: \'a::len word\ proof (unfold is_aligned_iff_udvd) from that show \2 ^ n udvd x\ using dvd_triv_left exp_dvd_iff_exp_udvd by blast qed -lemma is_alignedE [elim?]: +lemma is_alignedE: fixes w :: \'a::len word\ assumes \is_aligned w n\ obtains q where \w = 2 ^ n * word_of_nat q\ \q < 2 ^ (LENGTH('a) - n)\ proof (cases \n < LENGTH('a)\) case False with assms have \w = 0\ by (simp add: is_aligned_beyond_length) with that [of 0] show thesis by simp next case True moreover define m where \m = LENGTH('a) - n\ ultimately have l: \LENGTH('a) = n + m\ and \m \ 0\ by simp_all from \n < LENGTH('a)\ have *: \unat (2 ^ n :: 'a word) = 2 ^ n\ by transfer simp from assms have \2 ^ n udvd w\ by (simp add: is_aligned_iff_udvd) then obtain v :: \'a word\ where \unat w = unat (2 ^ n :: 'a word) * unat v\ .. moreover define q where \q = unat v\ ultimately have unat_w: \unat w = 2 ^ n * q\ by (simp add: *) then have \word_of_nat (unat w) = (word_of_nat (2 ^ n * q) :: 'a word)\ by simp then have w: \w = 2 ^ n * word_of_nat q\ by simp moreover have \q < 2 ^ (LENGTH('a) - n)\ proof (rule ccontr) assume \\ q < 2 ^ (LENGTH('a) - n)\ then have \2 ^ (LENGTH('a) - n) \ q\ by simp then have \2 ^ LENGTH('a) \ 2 ^ n * q\ by (simp add: l power_add) with unat_w [symmetric] show False by (metis le_antisym nat_less_le unsigned_less) qed ultimately show thesis using that by blast qed +lemma is_alignedE' [elim?]: + fixes w :: \'a::len word\ + assumes \is_aligned w n\ + obtains q where \w = push_bit n (word_of_nat q)\ \q < 2 ^ (LENGTH('a) - n)\ +proof - + from assms + obtain q where \w = 2 ^ n * word_of_nat q\ \q < 2 ^ (LENGTH('a) - n)\ + by (rule is_alignedE) + then have \w = push_bit n (word_of_nat q)\ + by (simp add: push_bit_eq_mult) + with that show thesis + using \q < 2 ^ (LENGTH('a) - n)\ . +qed + lemma is_aligned_mask: \is_aligned w n \ w AND mask n = 0\ by (simp add: is_aligned_iff_take_bit_eq_0 take_bit_eq_mask) +lemma is_aligned_imp_not_bit: + \\ bit w m\ if \is_aligned w n\ and \m < n\ + for w :: \'a::len word\ +proof - + from \is_aligned w n\ + obtain q where \w = push_bit n (word_of_nat q)\ \q < 2 ^ (LENGTH('a) - n)\ .. + moreover have \\ bit (push_bit n (word_of_nat q :: 'a word)) m\ + using \m < n\ by (simp add: bit_simps) + ultimately show ?thesis + by simp +qed + lemma is_aligned_weaken: "\ is_aligned w x; x \ y \ \ is_aligned w y" unfolding is_aligned_iff_dvd_nat by (erule dvd_trans [rotated]) (simp add: le_imp_power_dvd) -lemma nat_power_less_diff: - assumes lt: "(2::nat) ^ n * q < 2 ^ m" - shows "q < 2 ^ (m - n)" - using lt -proof (induct n arbitrary: m) - case 0 - then show ?case by simp -next - case (Suc n) - - have ih: "\m. 2 ^ n * q < 2 ^ m \ q < 2 ^ (m - n)" - and prem: "2 ^ Suc n * q < 2 ^ m" by fact+ - - show ?case - proof (cases m) - case 0 - then show ?thesis using Suc by simp - next - case (Suc m') - then show ?thesis using prem - by (simp add: ac_simps ih) - qed -qed - lemma is_alignedE_pre: fixes w::"'a::len word" assumes aligned: "is_aligned w n" shows rl: "\q. w = 2 ^ n * (of_nat q) \ q < 2 ^ (LENGTH('a) - n)" using aligned is_alignedE by blast -lemma is_aligned_to_bl: - "is_aligned (w :: 'a :: len word) n = (True \ set (drop (size w - n) (to_bl w)))" - apply (simp add: is_aligned_mask eq_zero_set_bl) - apply (clarsimp simp: in_set_conv_nth word_size) - apply (simp add: to_bl_nth word_size cong: conj_cong) - apply (simp add: diff_diff_less) - apply safe - apply (case_tac "n \ LENGTH('a)") - prefer 2 - apply (rule_tac x=i in exI) - apply clarsimp - apply (subgoal_tac "\j < LENGTH('a). j < n \ LENGTH('a) - n + j = i") - apply (erule exE) - apply (rule_tac x=j in exI) - apply clarsimp - apply (thin_tac "w !! t" for t) - apply (rule_tac x="i + n - LENGTH('a)" in exI) - apply clarsimp - apply arith - apply (rule_tac x="LENGTH('a) - n + i" in exI) - apply clarsimp - apply arith - done - -lemma is_aligned_replicate: - fixes w::"'a::len word" - assumes aligned: "is_aligned w n" - and nv: "n \ LENGTH('a)" - shows "to_bl w = (take (LENGTH('a) - n) (to_bl w)) @ replicate n False" -proof - - from nv have rl: "\q. q < 2 ^ (LENGTH('a) - n) \ - to_bl (2 ^ n * (of_nat q :: 'a word)) = - drop n (to_bl (of_nat q :: 'a word)) @ replicate n False" - by (metis bl_shiftl le_antisym min_def shiftl_t2n wsst_TYs(3)) - show ?thesis using aligned - by (auto simp: rl elim: is_alignedE) -qed - -lemma is_aligned_drop: - fixes w::"'a::len word" - assumes "is_aligned w n" "n \ LENGTH('a)" - shows "drop (LENGTH('a) - n) (to_bl w) = replicate n False" -proof - - have "to_bl w = take (LENGTH('a) - n) (to_bl w) @ replicate n False" - by (rule is_aligned_replicate) fact+ - then have "drop (LENGTH('a) - n) (to_bl w) = drop (LENGTH('a) - n) \" by simp - also have "\ = replicate n False" by simp - finally show ?thesis . -qed - -lemma less_is_drop_replicate: - fixes x::"'a::len word" - assumes lt: "x < 2 ^ n" - shows "to_bl x = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl x)" - by (metis assms bl_and_mask' less_mask_eq) - -lemma is_aligned_add_conv: - fixes off::"'a::len word" - assumes aligned: "is_aligned w n" - and offv: "off < 2 ^ n" - shows "to_bl (w + off) = - (take (LENGTH('a) - n) (to_bl w)) @ (drop (LENGTH('a) - n) (to_bl off))" -proof cases - assume nv: "n \ LENGTH('a)" - show ?thesis - proof (subst aligned_bl_add_size, simp_all only: word_size) - show "drop (LENGTH('a) - n) (to_bl w) = replicate n False" - by (subst is_aligned_replicate [OF aligned nv]) (simp add: word_size) - - from offv show "take (LENGTH('a) - n) (to_bl off) = - replicate (LENGTH('a) - n) False" - by (subst less_is_drop_replicate, assumption) simp - qed fact -next - assume "\ n \ LENGTH('a)" - with offv show ?thesis by (simp add: power_overflow) -qed - lemma aligned_add_aligned: fixes x::"'a::len word" assumes aligned1: "is_aligned x n" and aligned2: "is_aligned y m" and lt: "m \ n" shows "is_aligned (x + y) m" proof cases assume nlt: "n < LENGTH('a)" show ?thesis unfolding is_aligned_iff_dvd_nat dvd_def proof - from aligned2 obtain q2 where yv: "y = 2 ^ m * of_nat q2" and q2v: "q2 < 2 ^ (LENGTH('a) - m)" by (auto elim: is_alignedE) from lt obtain k where kv: "m + k = n" by (auto simp: le_iff_add) with aligned1 obtain q1 where xv: "x = 2 ^ (m + k) * of_nat q1" and q1v: "q1 < 2 ^ (LENGTH('a) - (m + k))" by (auto elim: is_alignedE) have l1: "2 ^ (m + k) * q1 < 2 ^ LENGTH('a)" by (rule nat_less_power_trans [OF q1v]) (subst kv, rule order_less_imp_le [OF nlt]) have l2: "2 ^ m * q2 < 2 ^ LENGTH('a)" by (rule nat_less_power_trans [OF q2v], rule order_less_imp_le [OF order_le_less_trans]) fact+ have "x = of_nat (2 ^ (m + k) * q1)" using xv by simp moreover have "y = of_nat (2 ^ m * q2)" using yv by simp ultimately have upls: "unat x + unat y = 2 ^ m * (2 ^ k * q1 + q2)" proof - have f1: "unat x = 2 ^ (m + k) * q1" by (metis (no_types) \x = of_nat (2 ^ (m + k) * q1)\ l1 nat_mod_lem word_unat.inverse_norm zero_less_numeral zero_less_power) have "unat y = 2 ^ m * q2" by (metis (no_types) \y = of_nat (2 ^ m * q2)\ l2 nat_mod_lem word_unat.inverse_norm zero_less_numeral zero_less_power) then show ?thesis using f1 by (simp add: power_add semiring_normalization_rules(34)) qed (* (2 ^ k * q1 + q2) *) show "\d. unat (x + y) = 2 ^ m * d" proof (cases "unat x + unat y < 2 ^ LENGTH('a)") case True have "unat (x + y) = unat x + unat y" by (subst unat_plus_if', rule if_P) fact also have "\ = 2 ^ m * (2 ^ k * q1 + q2)" by (rule upls) finally show ?thesis .. next case False then have "unat (x + y) = (unat x + unat y) mod 2 ^ LENGTH('a)" by (subst unat_word_ariths(1)) simp also have "\ = (2 ^ m * (2 ^ k * q1 + q2)) mod 2 ^ LENGTH('a)" by (subst upls, rule refl) also have "\ = 2 ^ m * ((2 ^ k * q1 + q2) mod 2 ^ (LENGTH('a) - m))" proof - have "m \ len_of (TYPE('a))" by (meson le_trans less_imp_le_nat lt nlt) then show ?thesis by (metis mult_mod_right ordered_cancel_comm_monoid_diff_class.add_diff_inverse power_add) qed finally show ?thesis .. qed qed next assume "\ n < LENGTH('a)" with assms show ?thesis by (simp add: is_aligned_mask not_less take_bit_eq_mod power_overflow word_arith_nat_defs(7) flip: take_bit_eq_mask) qed corollary aligned_sub_aligned: "\is_aligned (x::'a::len word) n; is_aligned y m; m \ n\ \ is_aligned (x - y) m" apply (simp del: add_uminus_conv_diff add:diff_conv_add_uminus) apply (erule aligned_add_aligned, simp_all) apply (erule is_alignedE) apply (rule_tac k="- of_nat q" in is_alignedI) apply simp done lemma is_aligned_shift: fixes k::"'a::len word" shows "is_aligned (k << m) m" proof cases assume mv: "m < LENGTH('a)" from mv obtain q where mq: "m + q = LENGTH('a)" and "0 < q" by (auto dest: less_imp_add_positive) have "(2::nat) ^ m dvd unat (k << m)" proof have kv: "(unat k div 2 ^ q) * 2 ^ q + unat k mod 2 ^ q = unat k" by (rule div_mult_mod_eq) have "unat (k << m) = unat (2 ^ m * k)" by (simp add: shiftl_t2n) also have "\ = (2 ^ m * unat k) mod (2 ^ LENGTH('a))" using mv by (simp add: unat_word_ariths(2)) also have "\ = 2 ^ m * (unat k mod 2 ^ q)" by (subst mq [symmetric], subst power_add, subst mod_mult2_eq) simp finally show "unat (k << m) = 2 ^ m * (unat k mod 2 ^ q)" . qed then show ?thesis by (unfold is_aligned_iff_dvd_nat) next assume "\ m < LENGTH('a)" then show ?thesis by (simp add: not_less power_overflow is_aligned_mask shiftl_zero_size word_size) qed lemma word_mod_by_0: "k mod (0::'a::len word) = k" by (simp add: word_arith_nat_mod) lemma aligned_mod_eq_0: fixes p::"'a::len word" assumes al: "is_aligned p sz" shows "p mod 2 ^ sz = 0" proof cases assume szv: "sz < LENGTH('a)" with al show ?thesis unfolding is_aligned_iff_dvd_nat by (simp add: and_mask_dvd_nat p2_gt_0 word_mod_2p_is_mask) next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: is_aligned_mask flip: take_bit_eq_mask take_bit_eq_mod) qed lemma is_aligned_triv: "is_aligned (2 ^ n ::'a::len word) n" by (rule is_alignedI [where k = 1], simp) lemma is_aligned_mult_triv1: "is_aligned (2 ^ n * x ::'a::len word) n" by (rule is_alignedI [OF refl]) lemma is_aligned_mult_triv2: "is_aligned (x * 2 ^ n ::'a::len word) n" by (subst mult.commute, simp add: is_aligned_mult_triv1) lemma word_power_less_0_is_0: fixes x :: "'a::len word" shows "x < a ^ 0 \ x = 0" by simp lemma is_aligned_no_wrap: fixes off :: "'a::len word" fixes ptr :: "'a::len word" assumes al: "is_aligned ptr sz" and off: "off < 2 ^ sz" shows "unat ptr + unat off < 2 ^ LENGTH('a)" proof - have szv: "sz < LENGTH('a)" using off p2_gt_0 word_neq_0_conv by fastforce from al obtain q where ptrq: "ptr = 2 ^ sz * of_nat q" and qv: "q < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis proof (cases "sz = 0") case True then show ?thesis using off ptrq qv by simp next case False then have sne: "0 < sz" .. show ?thesis proof - have uq: "unat (of_nat q ::'a::len word) = q" apply (subst unat_of_nat) apply (rule mod_less) apply (rule order_less_trans [OF qv]) apply (rule power_strict_increasing [OF diff_less [OF sne]]) apply (simp_all) done have uptr: "unat ptr = 2 ^ sz * q" apply (subst ptrq) apply (subst iffD1 [OF unat_mult_lem]) apply (subst unat_power_lower [OF szv]) apply (subst uq) apply (rule nat_less_power_trans [OF qv order_less_imp_le [OF szv]]) apply (subst uq) apply (subst unat_power_lower [OF szv]) apply simp done show "unat ptr + unat off < 2 ^ LENGTH('a)" using szv apply (subst uptr) apply (subst mult.commute, rule nat_add_offset_less [OF _ qv]) apply (rule order_less_le_trans [OF unat_mono [OF off] order_eq_refl]) apply simp_all done qed qed qed lemma is_aligned_no_wrap': fixes ptr :: "'a::len word" assumes al: "is_aligned ptr sz" and off: "off < 2 ^ sz" shows "ptr \ ptr + off" by (subst no_plus_overflow_unat_size, subst word_size, rule is_aligned_no_wrap) fact+ lemma is_aligned_no_overflow': fixes p :: "'a::len word" assumes al: "is_aligned p n" shows "p \ p + (2 ^ n - 1)" proof cases assume "n n ptr \ ptr + 2^sz - 1" by (drule is_aligned_no_overflow') (simp add: field_simps) lemma replicate_not_True: "\n. xs = replicate n False \ True \ set xs" by (induct xs) auto -lemma is_aligned_replicateI: - "to_bl p = addr @ replicate n False \ is_aligned (p::'a::len word) n" - apply (simp add: is_aligned_to_bl word_size) - apply (subgoal_tac "length addr = LENGTH('a) - n") - apply (simp add: replicate_not_True) - apply (drule arg_cong [where f=length]) - apply simp - done - -lemma to_bl_2p: - "n < LENGTH('a) \ - to_bl ((2::'a::len word) ^ n) = - replicate (LENGTH('a) - Suc n) False @ True # replicate n False" - apply (subst shiftl_1 [symmetric]) - apply (subst bl_shiftl) - apply (simp add: to_bl_1 min_def word_size) - done - lemma map_zip_replicate_False_xor: "n = length xs \ map (\(x, y). x = (\ y)) (zip xs (replicate n False)) = xs" by (induct xs arbitrary: n, auto) lemma drop_minus_lem: "\ n \ length xs; 0 < n; n' = length xs \ \ drop (n' - n) xs = rev xs ! (n - 1) # drop (Suc (n' - n)) xs" proof (induct xs arbitrary: n n') case Nil then show ?case by simp next case (Cons y ys) from Cons.prems show ?case apply simp apply (cases "n = Suc (length ys)") apply (simp add: nth_append) apply (simp add: Suc_diff_le Cons.hyps nth_append) apply clarsimp apply arith done qed lemma drop_minus: "\ n < length xs; n' = length xs \ \ drop (n' - Suc n) xs = rev xs ! n # drop (n' - n) xs" apply (subst drop_minus_lem) apply simp apply simp apply simp apply simp apply (cases "length xs", simp) apply (simp add: Suc_diff_le) done -lemma xor_2p_to_bl: - fixes x::"'a::len word" - shows "to_bl (x XOR 2^n) = - (if n < LENGTH('a) - then take (LENGTH('a)-Suc n) (to_bl x) @ (\rev (to_bl x)!n) # drop (LENGTH('a)-n) (to_bl x) - else to_bl x)" -proof - - have x: "to_bl x = take (LENGTH('a)-Suc n) (to_bl x) @ drop (LENGTH('a)-Suc n) (to_bl x)" - by simp - - show ?thesis - apply simp - apply (rule conjI) - apply (clarsimp simp: word_size) - apply (simp add: bl_word_xor to_bl_2p) - apply (subst x) - apply (subst zip_append) - apply simp - apply (simp add: map_zip_replicate_False_xor drop_minus) - apply (auto simp add: word_size nth_w2p intro!: word_eqI) - done -qed - lemma aligned_add_xor: - assumes al: "is_aligned (x::'a::len word) n'" and le: "n < n'" - shows "(x + 2^n) XOR 2^n = x" -proof cases - assume "n' < LENGTH('a)" - with assms show ?thesis - apply - - apply (rule word_bl.Rep_eqD) - apply (subst xor_2p_to_bl) - apply simp - apply (subst is_aligned_add_conv, simp, simp add: word_less_nat_alt)+ - apply (simp add: to_bl_2p nth_append) - apply (cases "n' = Suc n") - apply simp - apply (subst is_aligned_replicate [where n="Suc n", simplified, symmetric]; simp) - apply (subgoal_tac "\ LENGTH('a) - Suc n \ LENGTH('a) - n'") - prefer 2 - apply arith - apply (subst replicate_Suc [symmetric]) - apply (subst replicate_add [symmetric]) - apply (simp add: is_aligned_replicate [simplified, symmetric]) - done -next - assume "\ n' < LENGTH('a)" - show ?thesis - using al apply (rule is_alignedE) - using \\ n' < LENGTH('a)\ by auto + \(x + 2 ^ n) XOR 2 ^ n = x\ + if al: \is_aligned (x::'a::len word) n'\ and le: \n < n'\ +proof - + have \\ bit x n\ + using that by (rule is_aligned_imp_not_bit) + then have \x + 2 ^ n = x OR 2 ^ n\ + by (subst disjunctive_add) (auto simp add: bit_simps disjunctive_add) + moreover have \(x OR 2 ^ n) XOR 2 ^ n = x\ + by (rule bit_word_eqI) (auto simp add: bit_simps \\ bit x n\) + ultimately show ?thesis + by simp qed -lemma is_aligned_replicateD: - "\ is_aligned (w::'a::len word) n; n \ LENGTH('a) \ - \ \xs. to_bl w = xs @ replicate n False - \ length xs = size w - n" - apply (subst is_aligned_replicate, assumption+) - apply (rule exI, rule conjI, rule refl) - apply (simp add: word_size) - done - lemma is_aligned_add_mult_multI: fixes p :: "'a::len word" shows "\is_aligned p m; n \ m; n' = n\ \ is_aligned (p + x * 2 ^ n * z) n'" apply (erule aligned_add_aligned) apply (auto intro: is_alignedI [where k="x*z"]) done lemma is_aligned_add_multI: fixes p :: "'a::len word" shows "\is_aligned p m; n \ m; n' = n\ \ is_aligned (p + x * 2 ^ n) n'" apply (erule aligned_add_aligned) apply (auto intro: is_alignedI [where k="x"]) done lemma is_aligned_no_wrap''': fixes ptr :: "'a::len word" shows"\ is_aligned ptr sz; sz < LENGTH('a); off < 2 ^ sz \ \ unat ptr + off < 2 ^ LENGTH('a)" apply (drule is_aligned_no_wrap[where off="of_nat off"]) apply (simp add: word_less_nat_alt) apply (erule order_le_less_trans[rotated]) apply (simp add: take_bit_eq_mod) apply (subst(asm) unat_of_nat_len) apply (erule order_less_trans) apply (erule power_strict_increasing) apply simp apply assumption done lemma is_aligned_get_word_bits: fixes p :: "'a::len word" shows "\ is_aligned p n; \ is_aligned p n; n < LENGTH('a) \ \ P; \ p = 0; n \ LENGTH('a) \ \ P \ \ P" apply (cases "n < LENGTH('a)") apply simp apply simp apply (erule meta_mp) apply (simp add: is_aligned_mask power_add power_overflow not_less flip: take_bit_eq_mask) apply (metis take_bit_length_eq take_bit_of_0 take_bit_tightened) done lemma aligned_small_is_0: "\ is_aligned x n; x < 2 ^ n \ \ x = 0" - apply (erule is_aligned_get_word_bits) - apply (frule is_aligned_add_conv [rotated, where w=0]) - apply (simp add: is_aligned_iff_dvd_nat) - apply simp - apply (drule is_aligned_replicateD) - apply simp - apply (clarsimp simp: word_size) - apply (subst (asm) replicate_add [symmetric]) - apply (drule arg_cong[where f="of_bl :: bool list \ 'a::len word"]) - apply simp - apply (simp only: replicate.simps[symmetric, where x=False] - drop_replicate) - done + by (simp add: is_aligned_mask less_mask_eq) corollary is_aligned_less_sz: "\is_aligned a sz; a \ 0\ \ \ a < 2 ^ sz" by (rule notI, drule(1) aligned_small_is_0, erule(1) notE) lemma aligned_at_least_t2n_diff: "\is_aligned x n; is_aligned y n; x < y\ \ x \ y - 2 ^ n" apply (erule is_aligned_get_word_bits[where p=y]) apply (rule ccontr) apply (clarsimp simp: linorder_not_le) apply (subgoal_tac "y - x = 0") apply clarsimp apply (rule aligned_small_is_0) apply (erule(1) aligned_sub_aligned) apply simp apply unat_arith apply simp done lemma is_aligned_no_overflow'': "\is_aligned x n; x + 2 ^ n \ 0\ \ x \ x + 2 ^ n" apply (frule is_aligned_no_overflow') apply (erule order_trans) apply (simp add: field_simps) apply (erule word_sub_1_le) done lemma is_aligned_nth [word_eqI_simps]: "is_aligned p m = (\n < m. \p !! n)" apply (clarsimp simp: is_aligned_mask bang_eq word_size) apply (rule iffI) apply clarsimp apply (case_tac "n < size p") apply (simp add: word_size) apply (drule test_bit_size) apply simp apply clarsimp done lemma range_inter: "({a..b} \ {c..d} = {}) = (\x. \(a \ x \ x \ b \ c \ x \ x \ d))" by auto lemma aligned_inter_non_empty: "\ {p..p + (2 ^ n - 1)} \ {p..p + 2 ^ m - 1} = {}; is_aligned p n; is_aligned p m\ \ False" apply (clarsimp simp only: range_inter) apply (erule_tac x=p in allE) apply simp apply (erule impE) apply (erule is_aligned_no_overflow') apply (erule notE) apply (erule is_aligned_no_overflow) done lemma not_aligned_mod_nz: assumes al: "\ is_aligned a n" shows "a mod 2 ^ n \ 0" apply (rule ccontr) using al apply (rule notE) apply simp apply (rule is_alignedI [of _ _ \a div 2 ^ n\]) apply (metis add.right_neutral mult.commute word_mod_div_equality) done lemma nat_add_offset_le: fixes x :: nat assumes yv: "y \ 2 ^ n" and xv: "x < 2 ^ m" and mn: "sz = m + n" shows "x * 2 ^ n + y \ 2 ^ sz" proof (subst mn) from yv obtain qy where "y + qy = 2 ^ n" by (auto simp: le_iff_add) have "x * 2 ^ n + y \ x * 2 ^ n + 2 ^ n" using yv xv by simp also have "\ = (x + 1) * 2 ^ n" by simp also have "\ \ 2 ^ (m + n)" using xv by (subst power_add) (rule mult_le_mono1, simp) finally show "x * 2 ^ n + y \ 2 ^ (m + n)" . qed lemma is_aligned_no_wrap_le: fixes ptr::"'a::len word" assumes al: "is_aligned ptr sz" and szv: "sz < LENGTH('a)" and off: "off \ 2 ^ sz" shows "unat ptr + off \ 2 ^ LENGTH('a)" proof - from al obtain q where ptrq: "ptr = 2 ^ sz * of_nat q" and qv: "q < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis proof (cases "sz = 0") case True then show ?thesis using off ptrq qv by (auto simp add: le_Suc_eq Suc_le_eq) (simp add: le_less) next case False then have sne: "0 < sz" .. show ?thesis proof - have uq: "unat (of_nat q :: 'a word) = q" apply (subst unat_of_nat) apply (rule mod_less) apply (rule order_less_trans [OF qv]) apply (rule power_strict_increasing [OF diff_less [OF sne]]) apply simp_all done have uptr: "unat ptr = 2 ^ sz * q" apply (subst ptrq) apply (subst iffD1 [OF unat_mult_lem]) apply (subst unat_power_lower [OF szv]) apply (subst uq) apply (rule nat_less_power_trans [OF qv order_less_imp_le [OF szv]]) apply (subst uq) apply (subst unat_power_lower [OF szv]) apply simp done show "unat ptr + off \ 2 ^ LENGTH('a)" using szv apply (subst uptr) apply (subst mult.commute, rule nat_add_offset_le [OF off qv]) apply simp done qed qed qed lemma is_aligned_neg_mask: "m \ n \ is_aligned (x AND NOT (mask n)) m" by (metis and_not_mask is_aligned_shift is_aligned_weaken) lemma unat_minus: "unat (- (x :: 'a :: len word)) = (if x = 0 then 0 else 2 ^ size x - unat x)" using unat_sub_if_size[where x="2 ^ size x" and y=x] by (simp add: unat_eq_0 word_size) lemma is_aligned_minus: \is_aligned (- p) n\ if \is_aligned p n\ for p :: \'a::len word\ using that apply (cases \n < LENGTH('a)\) apply (simp_all add: not_less is_aligned_beyond_length) apply transfer apply (simp flip: take_bit_eq_0_iff) apply (subst take_bit_minus [symmetric]) apply simp done lemma add_mask_lower_bits: "\is_aligned (x :: 'a :: len word) n; \n' \ n. n' < LENGTH('a) \ \ p !! n'\ \ x + p AND NOT (mask n) = x" apply (subst word_plus_and_or_coroll) apply (rule word_eqI) apply (clarsimp simp: word_size is_aligned_nth) apply (erule_tac x=na in allE)+ apply simp apply (rule bit_word_eqI) apply (auto simp add: bit_simps not_less test_bit_eq_bit) apply (metis is_aligned_nth not_le test_bit_eq_bit) done lemma is_aligned_andI1: "is_aligned x n \ is_aligned (x AND y) n" by (simp add: is_aligned_nth) lemma is_aligned_andI2: "is_aligned y n \ is_aligned (x AND y) n" by (simp add: is_aligned_nth) lemma is_aligned_shiftl: "is_aligned w (n - m) \ is_aligned (w << m) n" by (simp add: is_aligned_nth nth_shiftl) lemma is_aligned_shiftr: "is_aligned w (n + m) \ is_aligned (w >> m) n" by (simp add: is_aligned_nth nth_shiftr) lemma is_aligned_shiftl_self: "is_aligned (p << n) n" by (rule is_aligned_shift) lemma is_aligned_neg_mask_eq: "is_aligned p n \ p AND NOT (mask n) = p" by (metis add.left_neutral is_aligned_mask word_plus_and_or_coroll2) lemma is_aligned_shiftr_shiftl: "is_aligned w n \ w >> n << n = w" by (metis and_not_mask is_aligned_neg_mask_eq) lemma aligned_shiftr_mask_shiftl: "is_aligned x n \ ((x >> n) AND mask v) << n = x AND mask (v + n)" apply (rule word_eqI) apply (simp add: word_size nth_shiftl nth_shiftr) apply (subgoal_tac "\m. x !! m \ m \ n") apply auto[1] apply (clarsimp simp: is_aligned_mask) apply (drule_tac x=m in word_eqD) apply (frule test_bit_size) apply (simp add: word_size) done lemma mask_zero: "is_aligned x a \ x AND mask a = 0" by (metis is_aligned_mask) lemma is_aligned_neg_mask_eq_concrete: "\ is_aligned p n; msk AND NOT (mask n) = NOT (mask n) \ \ p AND msk = p" by (metis word_bw_assocs(1) word_bw_comms(1) is_aligned_neg_mask_eq) lemma is_aligned_and_not_zero: "\ is_aligned n k; n \ 0 \ \ 2 ^ k \ n" using is_aligned_less_sz leI by blast lemma is_aligned_and_2_to_k: "(n AND 2 ^ k - 1) = 0 \ is_aligned (n :: 'a :: len word) k" by (simp add: is_aligned_mask mask_eq_decr_exp) lemma is_aligned_power2: "b \ a \ is_aligned (2 ^ a) b" by (metis is_aligned_triv is_aligned_weaken) lemma aligned_sub_aligned': "\ is_aligned (a :: 'a :: len word) n; is_aligned b n; n < LENGTH('a) \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma is_aligned_neg_mask_weaken: "\ is_aligned p n; m \ n \ \ p AND NOT (mask m) = p" using is_aligned_neg_mask_eq is_aligned_weaken by blast lemma is_aligned_neg_mask2 [simp]: "is_aligned (a AND NOT (mask n)) n" by (simp add: and_not_mask is_aligned_shift) lemma is_aligned_0': "is_aligned 0 n" by (fact is_aligned_0) lemma aligned_add_offset_no_wrap: fixes off :: "('a::len) word" and x :: "'a word" assumes al: "is_aligned x sz" and offv: "off < 2 ^ sz" shows "unat x + unat off < 2 ^ LENGTH('a)" proof cases assume szv: "sz < LENGTH('a)" from al obtain k where xv: "x = 2 ^ sz * (of_nat k)" and kl: "k < 2 ^ (LENGTH('a) - sz)" by (auto elim: is_alignedE) show ?thesis using szv apply (subst xv) apply (subst unat_mult_power_lem[OF kl]) apply (subst mult.commute, rule nat_add_offset_less) apply (rule less_le_trans[OF unat_mono[OF offv, simplified]]) apply (erule eq_imp_le[OF unat_power_lower]) apply (rule kl) apply simp done next assume "\ sz < LENGTH('a)" with offv show ?thesis by (simp add: not_less power_overflow ) qed lemma aligned_add_offset_mod: fixes x :: "('a::len) word" assumes al: "is_aligned x sz" and kv: "k < 2 ^ sz" shows "(x + k) mod 2 ^ sz = k" proof cases assume szv: "sz < LENGTH('a)" have ux: "unat x + unat k < 2 ^ LENGTH('a)" by (rule aligned_add_offset_no_wrap) fact+ show ?thesis using al szv apply - apply (erule is_alignedE) apply (subst word_unat.Rep_inject [symmetric]) apply (subst unat_mod) apply (subst iffD1 [OF unat_add_lem], rule ux) apply simp apply (subst unat_mult_power_lem, assumption+) apply (simp) apply (rule mod_less[OF less_le_trans[OF unat_mono], OF kv]) apply (erule eq_imp_le[OF unat_power_lower]) done next assume "\ sz < LENGTH('a)" with al show ?thesis by (simp add: not_less power_overflow is_aligned_mask mask_eq_decr_exp word_mod_by_0) qed lemma aligned_neq_into_no_overlap: fixes x :: "'a::len word" assumes neq: "x \ y" and alx: "is_aligned x sz" and aly: "is_aligned y sz" shows "{x .. x + (2 ^ sz - 1)} \ {y .. y + (2 ^ sz - 1)} = {}" proof cases assume szv: "sz < LENGTH('a)" show ?thesis proof (rule equals0I, clarsimp) fix z assume xb: "x \ z" and xt: "z \ x + (2 ^ sz - 1)" and yb: "y \ z" and yt: "z \ y + (2 ^ sz - 1)" have rl: "\(p::'a word) k w. \uint p + uint k < 2 ^ LENGTH('a); w = p + k; w \ p + (2 ^ sz - 1) \ \ k < 2 ^ sz" apply - apply simp apply (subst (asm) add.commute, subst (asm) add.commute, drule word_plus_mcs_4) - apply (subst add.commute, subst no_plus_overflow_uint_size) - apply (simp add: word_size_bl) + apply (subst add.commute, subst no_plus_overflow_uint_size) + apply transfer + apply simp apply (auto simp add: le_less power_2_ge_iff szv) apply (metis le_less_trans mask_eq_decr_exp mask_lt_2pn order_less_imp_le szv) done from xb obtain kx where kx: "z = x + kx" and kxl: "uint x + uint kx < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') from yb obtain ky where ky: "z = y + ky" and kyl: "uint y + uint ky < 2 ^ LENGTH('a)" by (clarsimp dest!: word_le_exists') have "x = y" proof - have "kx = z mod 2 ^ sz" proof (subst kx, rule sym, rule aligned_add_offset_mod) show "kx < 2 ^ sz" by (rule rl) fact+ qed fact+ also have "\ = ky" proof (subst ky, rule aligned_add_offset_mod) show "ky < 2 ^ sz" using kyl ky yt by (rule rl) qed fact+ finally have kxky: "kx = ky" . moreover have "x + kx = y + ky" by (simp add: kx [symmetric] ky [symmetric]) ultimately show ?thesis by simp qed then show False using neq by simp qed next assume "\ sz < LENGTH('a)" with neq alx aly have False by (simp add: is_aligned_mask mask_eq_decr_exp power_overflow) then show ?thesis .. qed lemma is_aligned_add_helper: "\ is_aligned p n; d < 2 ^ n \ \ (p + d AND mask n = d) \ (p + d AND (NOT (mask n)) = p)" apply (subst (asm) is_aligned_mask) apply (drule less_mask_eq) apply (rule context_conjI) apply (subst word_plus_and_or_coroll) apply (simp_all flip: take_bit_eq_mask) apply (metis take_bit_eq_mask word_bw_lcs(1) word_log_esimps(1)) apply (metis add.commute add_left_imp_eq take_bit_eq_mask word_plus_and_or_coroll2) done lemmas mask_inner_mask = mask_eqs(1) lemma mask_add_aligned: "is_aligned p n \ (p + q) AND mask n = q AND mask n" apply (simp add: is_aligned_mask) apply (subst mask_inner_mask [symmetric]) apply simp done lemma mask_out_add_aligned: assumes al: "is_aligned p n" shows "p + (q AND NOT (mask n)) = (p + q) AND NOT (mask n)" using mask_add_aligned [OF al] by (simp add: mask_out_sub_mask) lemma is_aligned_add_or: "\is_aligned p n; d < 2 ^ n\ \ p + d = p OR d" apply (subst disjunctive_add) apply (simp_all add: is_aligned_iff_take_bit_eq_0) apply (simp add: bit_eq_iff) apply (auto simp add: bit_simps) subgoal for m apply (cases \m < n\) apply (auto simp add: not_less) apply (metis bit_take_bit_iff less_mask_eq take_bit_eq_mask) done done lemma not_greatest_aligned: "\ x < y; is_aligned x n; is_aligned y n \ \ x + 2 ^ n \ 0" by (metis NOT_mask add_diff_cancel_right' diff_0 is_aligned_neg_mask_eq not_le word_and_le1) lemma neg_mask_mono_le: "x \ y \ x AND NOT(mask n) \ y AND NOT(mask n)" for x :: "'a :: len word" proof (rule ccontr, simp add: linorder_not_le, cases "n < LENGTH('a)") case False then show "y AND NOT(mask n) < x AND NOT(mask n) \ False" by (simp add: mask_eq_decr_exp linorder_not_less power_overflow) next case True assume a: "x \ y" and b: "y AND NOT(mask n) < x AND NOT(mask n)" have word_bits: "n < LENGTH('a)" by fact have "y \ (y AND NOT(mask n)) + (y AND mask n)" by (simp add: word_plus_and_or_coroll2 add.commute) also have "\ \ (y AND NOT(mask n)) + 2 ^ n" apply (rule word_plus_mono_right) apply (rule order_less_imp_le, rule and_mask_less_size) apply (simp add: word_size word_bits) apply (rule is_aligned_no_overflow'', simp add: is_aligned_neg_mask word_bits) apply (rule not_greatest_aligned, rule b; simp add: is_aligned_neg_mask) done also have "\ \ x AND NOT(mask n)" using b apply (subst add.commute) apply (rule le_plus) apply (rule aligned_at_least_t2n_diff; simp add: is_aligned_neg_mask) apply (rule ccontr, simp add: linorder_not_le) apply (drule aligned_small_is_0[rotated]; simp add: is_aligned_neg_mask) done also have "\ \ x" by (rule word_and_le2) also have "x \ y" by fact finally show "False" using b by simp qed lemma and_neg_mask_eq_iff_not_mask_le: "w AND NOT(mask n) = NOT(mask n) \ NOT(mask n) \ w" for w :: \'a::len word\ by (metis eq_iff neg_mask_mono_le word_and_le1 word_and_le2 word_bw_same(1)) lemma neg_mask_le_high_bits [word_eqI_simps]: "NOT(mask n) \ w \ (\i \ {n ..< size w}. w !! i)" for w :: \'a::len word\ by (auto simp: word_size and_neg_mask_eq_iff_not_mask_le[symmetric] word_eq_iff neg_mask_test_bit) end diff --git a/thys/Word_Lib/More_Arithmetic.thy b/thys/Word_Lib/More_Arithmetic.thy --- a/thys/Word_Lib/More_Arithmetic.thy +++ b/thys/Word_Lib/More_Arithmetic.thy @@ -1,290 +1,314 @@ section \Arithmetic lemmas\ theory More_Arithmetic imports Main begin declare iszero_0 [intro] declare min.absorb1 [simp] min.absorb2 [simp] lemma n_less_equal_power_2 [simp]: "n < 2 ^ n" by (induction n) simp_all lemma min_pm [simp]: "min a b + (a - b) = a" for a b :: nat by arith lemma min_pm1 [simp]: "a - b + min a b = a" for a b :: nat by arith lemma rev_min_pm [simp]: "min b a + (a - b) = a" for a b :: nat by arith lemma rev_min_pm1 [simp]: "a - b + min b a = a" for a b :: nat by arith lemma min_minus [simp]: "min m (m - k) = m - k" for m k :: nat by arith lemma min_minus' [simp]: "min (m - k) m = m - k" for m k :: nat by arith lemma nat_min_simps: "(a::nat) \ b \ min b a = a" "a \ b \ min a b = a" by simp_all lemma iszero_minus: \iszero (- z) \ iszero z\ by (simp add: iszero_def) lemma diff_le_eq': "a - b \ c \ a \ b + c" for a b c :: int by arith lemma zless2: "0 < (2 :: int)" by (fact zero_less_numeral) lemma zless2p: "0 < (2 ^ n :: int)" by arith lemma zle2p: "0 \ (2 ^ n :: int)" by arith lemma ex_eq_or: "(\m. n = Suc m \ (m = k \ P m)) \ n = Suc k \ (\m. n = Suc m \ P m)" by auto lemma power_minus_simp: "0 < n \ a ^ n = a * a ^ (n - 1)" by (auto dest: gr0_implies_Suc) lemma n2s_ths: \2 + n = Suc (Suc n)\ \n + 2 = Suc (Suc n)\ by (fact add_2_eq_Suc add_2_eq_Suc')+ lemma s2n_ths: \Suc (Suc n) = 2 + n\ \Suc (Suc n) = n + 2\ by simp_all lemma gt_or_eq_0: "0 < y \ 0 = y" for y :: nat by arith lemmas nat_simps = diff_add_inverse2 diff_add_inverse lemmas nat_iffs = le_add1 le_add2 lemma sum_imp_diff: "j = k + i \ j - i = k" for k :: nat by simp lemma le_diff_eq': "a \ c - b \ b + a \ c" for a b c :: int by arith lemma less_diff_eq': "a < c - b \ b + a < c" for a b c :: int by arith lemma diff_less_eq': "a - b < c \ a < b + c" for a b c :: int by arith lemma axxbyy: "a + m + m = b + n + n \ a = 0 \ a = 1 \ b = 0 \ b = 1 \ a = b \ m = n" for a b m n :: int by arith lemmas zadd_diff_inverse = trans [OF diff_add_cancel [symmetric] add.commute] lemmas add_diff_cancel2 = add.commute [THEN diff_eq_eq [THEN iffD2]] lemmas mcl = mult_cancel_left [THEN iffD1, THEN make_pos_rule] lemma pl_pl_rels: "a + b = c + d \ a \ c \ b \ d \ a \ c \ b \ d" for a b c d :: nat by arith lemmas pl_pl_rels' = add.commute [THEN [2] trans, THEN pl_pl_rels] lemma minus_eq: "m - k = m \ k = 0 \ m = 0" for k m :: nat by arith lemma pl_pl_mm: "a + b = c + d \ a - c = d - b" for a b c d :: nat by arith lemmas pl_pl_mm' = add.commute [THEN [2] trans, THEN pl_pl_mm] lemma less_le_mult': "w * c < b * c \ 0 \ c \ (w + 1) * c \ b * c" for b c w :: int apply (rule mult_right_mono) apply (rule zless_imp_add1_zle) apply (erule (1) mult_right_less_imp_less) apply assumption done lemma less_le_mult: "w * c < b * c \ 0 \ c \ w * c + c \ b * c" for b c w :: int using less_le_mult' [of w c b] by (simp add: algebra_simps) lemmas less_le_mult_minus = iffD2 [OF le_diff_eq less_le_mult, simplified left_diff_distrib] lemma gen_minus: "0 < n \ f n = f (Suc (n - 1))" by auto lemma mpl_lem: "j \ i \ k < j \ i - j + k < i" for i j k :: nat by arith lemmas dme = div_mult_mod_eq lemmas dtle = div_times_less_eq_dividend lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] div_times_less_eq_dividend] lemma td_gal: "0 < c \ a \ b * c \ a div c \ b" for a b c :: nat apply safe apply (erule (1) xtrans(4) [OF div_le_mono div_mult_self_is_m]) apply (erule th2) done lemmas td_gal_lt = td_gal [simplified not_less [symmetric], simplified] lemmas sdl = div_nat_eqI lemma given_quot: "f > 0 \ (f * l + (f - 1)) div f = l" for f l :: nat by (rule div_nat_eqI) (simp_all) lemma given_quot_alt: "f > 0 \ (l * f + f - Suc 0) div f = l" for f l :: nat apply (frule given_quot) apply (rule trans) prefer 2 apply (erule asm_rl) apply (rule_tac f="\n. n div f" in arg_cong) apply (simp add : ac_simps) done lemma nat_less_power_trans: fixes n :: nat assumes nv: "n < 2 ^ (m - k)" and kv: "k \ m" shows "2 ^ k * n < 2 ^ m" proof (rule order_less_le_trans) show "2 ^ k * n < 2 ^ k * 2 ^ (m - k)" by (rule mult_less_mono2 [OF nv zero_less_power]) simp show "(2::nat) ^ k * 2 ^ (m - k) \ 2 ^ m" using nv kv by (subst power_add [symmetric]) simp qed lemma nat_le_power_trans: fixes n :: nat shows "\n \ 2 ^ (m - k); k \ m\ \ 2 ^ k * n \ 2 ^ m" by (metis le_add_diff_inverse mult_le_mono2 semiring_normalization_rules(26)) lemma x_power_minus_1: fixes x :: "'a :: {ab_group_add, power, numeral, one}" shows "x + (2::'a) ^ n - (1::'a) = x + (2 ^ n - 1)" by simp lemma nat_diff_add: fixes i :: nat shows "\ i + j = k \ \ i = k - j" by arith lemma pow_2_gt: "n \ 2 \ (2::int) < 2 ^ n" by (induct n) auto lemma sum_to_zero: "(a :: 'a :: ring) + b = 0 \ a = (- b)" by (drule arg_cong[where f="\ x. x - a"], simp) lemma arith_is_1: "\ x \ Suc 0; x > 0 \ \ x = 1" by arith lemma suc_le_pow_2: "1 < (n::nat) \ Suc n < 2 ^ n" by (induct n; clarsimp) (case_tac "n = 1"; clarsimp) lemma nat_le_Suc_less_imp: "x < y \ x \ y - Suc 0" by arith lemma power_sub_int: "\ m \ n; 0 < b \ \ b ^ n div b ^ m = (b ^ (n - m) :: int)" apply (subgoal_tac "\n'. n = m + n'") apply (clarsimp simp: power_add) apply (rule exI[where x="n - m"]) apply simp done lemma nat_Suc_less_le_imp: "(k::nat) < Suc n \ k \ n" by auto lemma nat_add_less_by_max: "\ (x::nat) \ xmax ; y < k - xmax \ \ x + y < k" by simp lemma nat_le_Suc_less: "0 < y \ (x \ y - Suc 0) = (x < y)" by arith lemma nat_power_minus_less: "a < 2 ^ (x - n) \ (a :: nat) < 2 ^ x" by (erule order_less_le_trans) simp lemma less_le_mult_nat': "w * c < b * c ==> 0 \ c ==> Suc w * c \ b * (c::nat)" apply (rule mult_right_mono) apply (rule Suc_leI) apply (erule (1) mult_right_less_imp_less) apply assumption done lemma less_le_mult_nat: \0 < c \ w < b \ c + w * c \ b * c\ for b c w :: nat using less_le_mult_nat' [of w c b] by simp lemma diff_diff_less: "(i < m - (m - (n :: nat))) = (i < m \ i < n)" by auto lemma nat_add_offset_less: fixes x :: nat assumes yv: "y < 2 ^ n" and xv: "x < 2 ^ m" and mn: "sz = m + n" shows "x * 2 ^ n + y < 2 ^ sz" proof (subst mn) from yv obtain qy where "y + qy = 2 ^ n" and "0 < qy" by (auto dest: less_imp_add_positive) have "x * 2 ^ n + y < x * 2 ^ n + 2 ^ n" by simp fact+ also have "\ = (x + 1) * 2 ^ n" by simp also have "\ \ 2 ^ (m + n)" using xv by (subst power_add) (rule mult_le_mono1, simp) finally show "x * 2 ^ n + y < 2 ^ (m + n)" . qed lemma p_assoc_help: fixes p :: "'a::{ring,power,numeral,one}" shows "p + 2^sz - 1 = p + (2^sz - 1)" by simp +lemma nat_power_less_diff: + assumes lt: "(2::nat) ^ n * q < 2 ^ m" + shows "q < 2 ^ (m - n)" + using lt +proof (induct n arbitrary: m) + case 0 + then show ?case by simp +next + case (Suc n) + + have ih: "\m. 2 ^ n * q < 2 ^ m \ q < 2 ^ (m - n)" + and prem: "2 ^ Suc n * q < 2 ^ m" by fact+ + + show ?case + proof (cases m) + case 0 + then show ?thesis using Suc by simp + next + case (Suc m') + then show ?thesis using prem + by (simp add: ac_simps ih) + qed +qed + end diff --git a/thys/Word_Lib/More_Word_Operations.thy b/thys/Word_Lib/More_Word_Operations.thy --- a/thys/Word_Lib/More_Word_Operations.thy +++ b/thys/Word_Lib/More_Word_Operations.thy @@ -1,663 +1,664 @@ section \Misc word operations\ theory More_Word_Operations imports "HOL-Library.Word" Aligned + Reversed_Bit_Lists More_Misc begin definition ptr_add :: "'a :: len word \ nat \ 'a word" where "ptr_add ptr n \ ptr + of_nat n" definition alignUp :: "'a::len word \ nat \ 'a word" where "alignUp x n \ x + 2 ^ n - 1 AND NOT (2 ^ n - 1)" (* standard notation for blocks of 2^n-1 words, usually aligned; abbreviation so it simplifies directly *) abbreviation mask_range :: "'a::len word \ nat \ 'a word set" where "mask_range p n \ {p .. p + mask n}" definition w2byte :: "'a :: len word \ 8 word" where "w2byte \ ucast" (* Count leading zeros *) definition word_clz :: "'a::len word \ nat" where "word_clz w \ length (takeWhile Not (to_bl w))" (* Count trailing zeros *) definition word_ctz :: "'a::len word \ nat" where "word_ctz w \ length (takeWhile Not (rev (to_bl w)))" lemma word_ctz_le: "word_ctz (w :: ('a::len word)) \ LENGTH('a)" apply (clarsimp simp: word_ctz_def) apply (rule nat_le_Suc_less_imp[where y="LENGTH('a) + 1" , simplified]) apply (rule order_le_less_trans[OF List.length_takeWhile_le]) apply simp done lemma word_ctz_less: "w \ 0 \ word_ctz (w :: ('a::len word)) < LENGTH('a)" apply (clarsimp simp: word_ctz_def eq_zero_set_bl) apply (rule order_less_le_trans[OF length_takeWhile_less]) apply fastforce+ done lemma word_ctz_not_minus_1: "1 < LENGTH('a) \ of_nat (word_ctz (w :: 'a :: len word)) \ (- 1 :: 'a::len word)" by (metis (mono_tags) One_nat_def add.right_neutral add_Suc_right le_diff_conv le_less_trans n_less_equal_power_2 not_le suc_le_pow_2 unat_minus_one_word unat_of_nat_len word_ctz_le) definition word_log2 :: "'a::len word \ nat" where "word_log2 (w::'a::len word) \ size w - 1 - word_clz w" (* Bit population count. Equivalent of __builtin_popcount. *) definition pop_count :: "('a::len) word \ nat" where "pop_count w \ length (filter id (to_bl w))" (* Sign extension from bit n *) definition sign_extend :: "nat \ 'a::len word \ 'a word" where "sign_extend n w \ if w !! n then w OR NOT (mask n) else w AND mask n" lemma sign_extend_eq_signed_take_bit: \sign_extend = signed_take_bit\ proof (rule ext)+ fix n and w :: \'a::len word\ show \sign_extend n w = signed_take_bit n w\ proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ then show \bit (sign_extend n w) q \ bit (signed_take_bit n w) q\ by (auto simp add: test_bit_eq_bit bit_signed_take_bit_iff sign_extend_def bit_and_iff bit_or_iff bit_not_iff bit_mask_iff not_less exp_eq_0_imp_not_bit not_le min_def) qed qed definition sign_extended :: "nat \ 'a::len word \ bool" where "sign_extended n w \ \i. n < i \ i < size w \ w !! i = w !! n" lemma ptr_add_0 [simp]: "ptr_add ref 0 = ref " unfolding ptr_add_def by simp lemma pop_count_0[simp]: "pop_count 0 = 0" by (clarsimp simp:pop_count_def) lemma pop_count_1[simp]: "pop_count 1 = 1" by (clarsimp simp:pop_count_def to_bl_1) lemma pop_count_0_imp_0: "(pop_count w = 0) = (w = 0)" apply (rule iffI) apply (clarsimp simp:pop_count_def) apply (subst (asm) filter_empty_conv) apply (clarsimp simp:eq_zero_set_bl) apply fast apply simp done lemma word_log2_nth_same: "w \ 0 \ w !! word_log2 w" unfolding word_log2_def using nth_length_takeWhile[where P=Not and xs="to_bl w"] apply (simp add: word_clz_def word_size to_bl_nth) apply (fastforce simp: linorder_not_less eq_zero_set_bl dest: takeWhile_take_has_property) done lemma word_log2_nth_not_set: "\ word_log2 w < i ; i < size w \ \ \ w !! i" unfolding word_log2_def word_clz_def using takeWhile_take_has_property_nth[where P=Not and xs="to_bl w" and n="size w - Suc i"] by (fastforce simp add: to_bl_nth word_size) lemma word_log2_highest: assumes a: "w !! i" shows "i \ word_log2 w" proof - from a have "i < size w" by - (rule test_bit_size) with a show ?thesis by - (rule ccontr, simp add: word_log2_nth_not_set) qed lemma word_log2_max: "word_log2 w < size w" unfolding word_log2_def word_clz_def by simp lemma word_clz_0[simp]: "word_clz (0::'a::len word) = LENGTH('a)" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma word_clz_minus_one[simp]: "word_clz (-1::'a::len word) = 0" unfolding word_clz_def by (simp add: takeWhile_replicate) lemma is_aligned_alignUp[simp]: "is_aligned (alignUp p n) n" by (simp add: alignUp_def is_aligned_mask mask_eq_decr_exp word_bw_assocs) lemma alignUp_le[simp]: "alignUp p n \ p + 2 ^ n - 1" unfolding alignUp_def by (rule word_and_le2) lemma alignUp_idem: fixes a :: "'a::len word" assumes "is_aligned a n" "n < LENGTH('a)" shows "alignUp a n = a" using assms unfolding alignUp_def by (metis mask_eq_decr_exp is_aligned_add_helper p_assoc_help power_2_ge_iff) lemma alignUp_not_aligned_eq: fixes a :: "'a :: len word" assumes al: "\ is_aligned a n" and sz: "n < LENGTH('a)" shows "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" proof - have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) fact+ then have um: "unat (a mod 2 ^ n - 1) div 2 ^ n = 0" using sz by (meson Euclidean_Division.div_eq_0_iff le_m1_iff_lt measure_unat order_less_trans unat_less_power word_less_sub_le word_mod_less_divisor) have "a + 2 ^ n - 1 = (a div 2 ^ n) * 2 ^ n + (a mod 2 ^ n) + 2 ^ n - 1" by (simp add: word_mod_div_equality) also have "\ = (a mod 2 ^ n - 1) + (a div 2 ^ n + 1) * 2 ^ n" by (simp add: field_simps) finally show "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" using sz unfolding alignUp_def apply (subst mask_eq_decr_exp [symmetric]) apply (erule ssubst) apply (subst neg_mask_is_div) apply (simp add: word_arith_nat_div) apply (subst unat_word_ariths(1) unat_word_ariths(2))+ apply (subst uno_simps) apply (subst unat_1) apply (subst mod_add_right_eq) apply simp apply (subst power_mod_div) apply (subst div_mult_self1) apply simp apply (subst um) apply simp apply (subst mod_mod_power) apply simp apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst mult_mod_left) apply (subst power_add [symmetric]) apply simp apply (subst Abs_fnat_hom_1) apply (subst Abs_fnat_hom_add) apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst word_unat.Rep_inverse[symmetric], subst Abs_fnat_hom_mult) apply simp done qed lemma alignUp_ge: fixes a :: "'a :: len word" assumes sz: "n < LENGTH('a)" and nowrap: "alignUp a n \ 0" shows "a \ alignUp a n" proof (cases "is_aligned a n") case True then show ?thesis using sz by (subst alignUp_idem, simp_all) next case False have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis shiftr_div_2n' word_shiftr_lt) have"2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans nat_less_le) moreover have "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using nowrap sz apply - apply (erule contrapos_nn) apply (subst alignUp_not_aligned_eq [OF False sz]) apply (subst unat_arith_simps) apply (subst unat_word_ariths) apply (subst unat_word_ariths) apply simp apply (subst mult_mod_left) apply (simp add: unat_div field_simps power_add[symmetric] mod_mod_power min.absorb2) done ultimately have lt: "2 ^ n * (unat a div 2 ^ n + 1) < 2 ^ LENGTH('a)" by simp have "a = a div 2 ^ n * 2 ^ n + a mod 2 ^ n" by (rule word_mod_div_equality [symmetric]) also have "\ < (a div 2 ^ n + 1) * 2 ^ n" using sz lt apply (simp add: field_simps) apply (rule word_add_less_mono1) apply (rule word_mod_less_divisor) apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (simp add: unat_div) done also have "\ = alignUp a n" by (rule alignUp_not_aligned_eq [symmetric]) fact+ finally show ?thesis by (rule order_less_imp_le) qed lemma alignUp_le_greater_al: fixes x :: "'a :: len word" assumes le: "a \ x" and sz: "n < LENGTH('a)" and al: "is_aligned x n" shows "alignUp a n \ x" proof (cases "is_aligned a n") case True then show ?thesis using sz le by (simp add: alignUp_idem) next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) from al obtain k where xk: "x = 2 ^ n * of_nat k" and kv: "k < 2 ^ (LENGTH('a) - n)" by (auto elim!: is_alignedE) then have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst mult.commute) apply simp apply (rule nat_less_power_trans) apply simp apply simp done have au: "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ also have "\ \ of_nat k * 2 ^ n" proof (rule word_mult_le_mono1 [OF inc_le _ kn]) show "a div 2 ^ n < of_nat k" using kv xk le sz anz by (simp add: alignUp_div_helper) show "(0:: 'a word) < 2 ^ n" using sz by (simp add: p2_gt_0 sz) qed finally show ?thesis using xk by (simp add: field_simps) qed lemma alignUp_is_aligned_nz: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and ax: "a \ x" and az: "a \ 0" shows "alignUp (a::'a :: len word) n \ 0" proof (cases "is_aligned a n") case True then have "alignUp a n = a" using sz by (simp add: alignUp_idem) then show ?thesis using az by simp next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) { assume asm: "alignUp a n = 0" have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis shiftr_div_2n' word_shiftr_lt) have leq: "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans order_less_imp_le) from al obtain k where kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" by (auto elim!: is_alignedE) then have "a div 2 ^ n < of_nat k" using ax sz anz by (rule alignUp_div_helper) then have r: "unat a div 2 ^ n < k" using sz by (simp flip: drop_bit_eq_div unat_drop_bit_eq) (metis leI le_unat_uoi unat_mono) have "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ then have "\ = 0" using asm by simp then have "2 ^ LENGTH('a) dvd 2 ^ n * (unat a div 2 ^ n + 1)" using sz by (simp add: unat_arith_simps ac_simps) (simp add: unat_word_ariths mod_simps mod_eq_0_iff_dvd) with leq have "2 ^ n * (unat a div 2 ^ n + 1) = 2 ^ LENGTH('a)" by (force elim!: le_SucE) then have "unat a div 2 ^ n = 2 ^ LENGTH('a) div 2 ^ n - 1" by (metis (no_types, hide_lams) Groups.add_ac(2) add.right_neutral add_diff_cancel_left' div_le_dividend div_mult_self4 gr_implies_not0 le_neq_implies_less power_eq_0_iff zero_neq_numeral) then have "unat a div 2 ^ n = 2 ^ (LENGTH('a) - n) - 1" using sz by (simp add: power_sub) then have "2 ^ (LENGTH('a) - n) - 1 < k" using r by simp then have False using kv by simp } then show ?thesis by clarsimp qed lemma alignUp_ar_helper: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and sub: "{x..x + 2 ^ n - 1} \ {a..b}" and anz: "a \ 0" shows "a \ alignUp a n \ alignUp a n + 2 ^ n - 1 \ b" proof from al have xl: "x \ x + 2 ^ n - 1" by (simp add: is_aligned_no_overflow) from xl sub have ax: "a \ x" by (clarsimp elim!: range_subset_lower [where x = x]) show "a \ alignUp a n" proof (rule alignUp_ge) show "alignUp a n \ 0" using al sz ax anz by (rule alignUp_is_aligned_nz) qed fact+ show "alignUp a n + 2 ^ n - 1 \ b" proof (rule order_trans) from xl show tp: "x + 2 ^ n - 1 \ b" using sub by (clarsimp elim!: range_subset_upper [where x = x]) from ax have "alignUp a n \ x" by (rule alignUp_le_greater_al) fact+ then have "alignUp a n + (2 ^ n - 1) \ x + (2 ^ n - 1)" using xl al is_aligned_no_overflow' olen_add_eqv word_plus_mcs_3 by blast then show "alignUp a n + 2 ^ n - 1 \ x + 2 ^ n - 1" by (simp add: field_simps) qed qed lemma alignUp_def2: "alignUp a sz = a + 2 ^ sz - 1 AND NOT (mask sz)" by (simp add: alignUp_def flip: mask_eq_decr_exp) lemma alignUp_def3: "alignUp a sz = 2^ sz + (a - 1 AND NOT (mask sz))" by (simp add: alignUp_def2 is_aligned_triv field_simps mask_out_add_aligned) lemma alignUp_plus: "is_aligned w us \ alignUp (w + a) us = w + alignUp a us" by (clarsimp simp: alignUp_def2 mask_out_add_aligned field_simps) lemma alignUp_distance: "alignUp (q :: 'a :: len word) sz - q \ mask sz" by (metis (no_types) add.commute add_diff_cancel_left alignUp_def2 diff_add_cancel mask_2pm1 subtract_mask(2) word_and_le1 word_sub_le_iff) lemma is_aligned_diff_neg_mask: "is_aligned p sz \ (p - q AND NOT (mask sz)) = (p - ((alignUp q sz) AND NOT (mask sz)))" apply (clarsimp simp only:word_and_le2 diff_conv_add_uminus) apply (subst mask_out_add_aligned[symmetric]; simp) apply (rule sum_to_zero) apply (subst add.commute) apply (simp add: alignUp_distance is_aligned_neg_mask_eq mask_out_add_aligned) apply (metis add_cancel_right_right alignUp_distance and_mask_eq_iff_le_mask word_plus_and_or_coroll2) done lemma word_clz_max: "word_clz w \ size (w::'a::len word)" unfolding word_clz_def by (metis length_takeWhile_le word_size_bl) lemma word_clz_nonzero_max: fixes w :: "'a::len word" assumes nz: "w \ 0" shows "word_clz w < size (w::'a::len word)" proof - { assume a: "word_clz w = size (w::'a::len word)" hence "length (takeWhile Not (to_bl w)) = length (to_bl w)" by (simp add: word_clz_def word_size) hence allj: "\j\set(to_bl w). \ j" by (metis a length_takeWhile_less less_irrefl_nat word_clz_def) hence "to_bl w = replicate (length (to_bl w)) False" by (fastforce intro!: list_of_false) hence "w = 0" by (metis to_bl_0 word_bl.Rep_eqD word_bl_Rep') with nz have False by simp } thus ?thesis using word_clz_max by (fastforce intro: le_neq_trans) qed (* Sign extension from bit n. *) lemma sign_extend_bitwise_if: "i < size w \ sign_extend e w !! i \ (if i < e then w !! i else w !! e)" by (simp add: sign_extend_def neg_mask_test_bit word_size) lemma sign_extend_bitwise_if' [word_eqI_simps]: \i < LENGTH('a) \ sign_extend e w !! i \ (if i < e then w !! i else w !! e)\ for w :: \'a::len word\ using sign_extend_bitwise_if [of i w e] by (simp add: word_size) lemma sign_extend_bitwise_disj: "i < size w \ sign_extend e w !! i \ i \ e \ w !! i \ e \ i \ w !! e" by (auto simp: sign_extend_bitwise_if) lemma sign_extend_bitwise_cases: "i < size w \ sign_extend e w !! i \ (i \ e \ w !! i) \ (e \ i \ w !! e)" by (auto simp: sign_extend_bitwise_if) lemmas sign_extend_bitwise_disj' = sign_extend_bitwise_disj[simplified word_size] lemmas sign_extend_bitwise_cases' = sign_extend_bitwise_cases[simplified word_size] (* Often, it is easier to reason about an operation which does not overwrite the bit which determines which mask operation to apply. *) lemma sign_extend_def': "sign_extend n w = (if w !! n then w OR NOT (mask (Suc n)) else w AND mask (Suc n))" by (rule bit_word_eqI) (auto simp add: bit_simps sign_extend_eq_signed_take_bit min_def test_bit_eq_bit less_Suc_eq_le) lemma sign_extended_sign_extend: "sign_extended n (sign_extend n w)" by (clarsimp simp: sign_extended_def word_size sign_extend_bitwise_if) lemma sign_extended_iff_sign_extend: "sign_extended n w \ sign_extend n w = w" apply auto apply (auto simp add: bit_eq_iff) apply (simp_all add: bit_simps sign_extend_eq_signed_take_bit not_le min_def sign_extended_def test_bit_eq_bit word_size split: if_splits) using le_imp_less_or_eq apply auto[1] apply (metis bit_imp_le_length nat_less_le) apply (metis Suc_leI Suc_n_not_le_n le_trans nat_less_le) done lemma sign_extended_weaken: "sign_extended n w \ n \ m \ sign_extended m w" unfolding sign_extended_def by (cases "n < m") auto lemma sign_extend_sign_extend_eq: "sign_extend m (sign_extend n w) = sign_extend (min m n) w" by (rule bit_word_eqI) (simp add: sign_extend_eq_signed_take_bit bit_simps) lemma sign_extended_high_bits: "\ sign_extended e p; j < size p; e \ i; i < j \ \ p !! i = p !! j" by (drule (1) sign_extended_weaken; simp add: sign_extended_def) lemma sign_extend_eq: "w AND mask (Suc n) = v AND mask (Suc n) \ sign_extend n w = sign_extend n v" by (simp flip: take_bit_eq_mask add: sign_extend_eq_signed_take_bit signed_take_bit_eq_iff_take_bit_eq) lemma sign_extended_add: assumes p: "is_aligned p n" assumes f: "f < 2 ^ n" assumes e: "n \ e" assumes "sign_extended e p" shows "sign_extended e (p + f)" proof (cases "e < size p") case True note and_or = is_aligned_add_or[OF p f] have "\ f !! e" using True e less_2p_is_upper_bits_unset[THEN iffD1, OF f] by (fastforce simp: word_size) hence i: "(p + f) !! e = p !! e" by (simp add: and_or) have fm: "f AND mask e = f" by (fastforce intro: subst[where P="\f. f AND mask e = f", OF less_mask_eq[OF f]] simp: mask_twice e) show ?thesis using assms apply (simp add: sign_extended_iff_sign_extend sign_extend_def i) apply (simp add: and_or word_bw_comms[of p f]) apply (clarsimp simp: word_ao_dist fm word_bw_assocs split: if_splits) done next case False thus ?thesis by (simp add: sign_extended_def word_size) qed lemma sign_extended_neq_mask: "\sign_extended n ptr; m \ n\ \ sign_extended n (ptr AND NOT (mask m))" by (fastforce simp: sign_extended_def word_size neg_mask_test_bit) definition "limited_and (x :: 'a :: len word) y \ (x AND y = x)" lemma limited_and_eq_0: "\ limited_and x z; y AND NOT z = y \ \ x AND y = 0" unfolding limited_and_def apply (subst arg_cong2[where f="(AND)"]) apply (erule sym)+ apply (simp(no_asm) add: word_bw_assocs word_bw_comms word_bw_lcs) done lemma limited_and_eq_id: "\ limited_and x z; y AND z = z \ \ x AND y = x" unfolding limited_and_def by (erule subst, fastforce simp: word_bw_lcs word_bw_assocs word_bw_comms) lemma lshift_limited_and: "limited_and x z \ limited_and (x << n) (z << n)" unfolding limited_and_def by (simp add: shiftl_over_and_dist[symmetric]) lemma rshift_limited_and: "limited_and x z \ limited_and (x >> n) (z >> n)" unfolding limited_and_def by (simp add: shiftr_over_and_dist[symmetric]) lemmas limited_and_simps1 = limited_and_eq_0 limited_and_eq_id lemmas is_aligned_limited_and = is_aligned_neg_mask_eq[unfolded mask_eq_decr_exp, folded limited_and_def] lemmas limited_and_simps = limited_and_simps1 limited_and_simps1[OF is_aligned_limited_and] limited_and_simps1[OF lshift_limited_and] limited_and_simps1[OF rshift_limited_and] limited_and_simps1[OF rshift_limited_and, OF is_aligned_limited_and] not_one shiftl_shiftr1[unfolded word_size mask_eq_decr_exp] shiftl_shiftr2[unfolded word_size mask_eq_decr_exp] definition from_bool :: "bool \ 'a::len word" where "from_bool b \ case b of True \ of_nat 1 | False \ of_nat 0" lemma from_bool_eq: \from_bool = of_bool\ by (simp add: fun_eq_iff from_bool_def) lemma from_bool_0: "(from_bool x = 0) = (\ x)" by (simp add: from_bool_def split: bool.split) lemma from_bool_eq_if': "((if P then 1 else 0) = from_bool Q) = (P = Q)" by (cases Q) (simp_all add: from_bool_def) definition to_bool :: "'a::len word \ bool" where "to_bool \ (\) 0" lemma to_bool_and_1: "to_bool (x AND 1) = (x !! 0)" by (simp add: test_bit_word_eq to_bool_def and_one_eq mod_2_eq_odd) lemma to_bool_from_bool [simp]: "to_bool (from_bool r) = r" unfolding from_bool_def to_bool_def by (simp split: bool.splits) lemma from_bool_neq_0 [simp]: "(from_bool b \ 0) = b" by (simp add: from_bool_def split: bool.splits) lemma from_bool_mask_simp [simp]: "(from_bool r :: 'a::len word) AND 1 = from_bool r" unfolding from_bool_def by (clarsimp split: bool.splits) lemma from_bool_1 [simp]: "(from_bool P = 1) = P" by (simp add: from_bool_def split: bool.splits) lemma ge_0_from_bool [simp]: "(0 < from_bool P) = P" by (simp add: from_bool_def split: bool.splits) lemma limited_and_from_bool: "limited_and (from_bool b) 1" by (simp add: from_bool_def limited_and_def split: bool.split) lemma to_bool_1 [simp]: "to_bool 1" by (simp add: to_bool_def) lemma to_bool_0 [simp]: "\to_bool 0" by (simp add: to_bool_def) lemma from_bool_eq_if: "(from_bool Q = (if P then 1 else 0)) = (P = Q)" by (cases Q) (simp_all add: from_bool_def) lemma to_bool_eq_0: "(\ to_bool x) = (x = 0)" by (simp add: to_bool_def) lemma to_bool_neq_0: "(to_bool x) = (x \ 0)" by (simp add: to_bool_def) lemma from_bool_all_helper: "(\bool. from_bool bool = val \ P bool) = ((\bool. from_bool bool = val) \ P (val \ 0))" by (auto simp: from_bool_0) lemma fold_eq_0_to_bool: "(v = 0) = (\ to_bool v)" by (simp add: to_bool_def) lemma from_bool_to_bool_iff: "w = from_bool b \ to_bool w = b \ (w = 0 \ w = 1)" by (cases b) (auto simp: from_bool_def to_bool_def) lemma from_bool_eqI: "from_bool x = from_bool y \ x = y" unfolding from_bool_def by (auto split: bool.splits) end \ No newline at end of file diff --git a/thys/Word_Lib/Reversed_Bit_Lists.thy b/thys/Word_Lib/Reversed_Bit_Lists.thy --- a/thys/Word_Lib/Reversed_Bit_Lists.thy +++ b/thys/Word_Lib/Reversed_Bit_Lists.thy @@ -1,1965 +1,2094 @@ (* Author: Jeremy Dawson, NICTA *) section \Bit values as reversed lists of bools\ theory Reversed_Bit_Lists imports "HOL-Library.Word" Typedef_Morphisms Least_significant_bit Most_significant_bit Even_More_List "HOL-Library.Sublist" + Aligned begin lemma horner_sum_of_bool_2_concat: \horner_sum of_bool 2 (concat (map (\x. map (bit x) [0.. for ws :: \'a::len word list\ proof (induction ws) case Nil then show ?case by simp next case (Cons w ws) moreover have \horner_sum of_bool 2 (map (bit w) [0.. proof transfer fix k :: int have \map (\n. n < LENGTH('a) \ bit k n) [0.. by simp then show \horner_sum of_bool 2 (map (\n. n < LENGTH('a) \ bit k n) [0.. by (simp only: horner_sum_bit_eq_take_bit) qed ultimately show ?case by (simp add: horner_sum_append) qed subsection \Implicit augmentation of list prefixes\ primrec takefill :: "'a \ nat \ 'a list \ 'a list" where Z: "takefill fill 0 xs = []" | Suc: "takefill fill (Suc n) xs = (case xs of [] \ fill # takefill fill n xs | y # ys \ y # takefill fill n ys)" lemma nth_takefill: "m < n \ takefill fill n l ! m = (if m < length l then l ! m else fill)" apply (induct n arbitrary: m l) apply clarsimp apply clarsimp apply (case_tac m) apply (simp split: list.split) apply (simp split: list.split) done lemma takefill_alt: "takefill fill n l = take n l @ replicate (n - length l) fill" by (induct n arbitrary: l) (auto split: list.split) lemma takefill_replicate [simp]: "takefill fill n (replicate m fill) = replicate n fill" by (simp add: takefill_alt replicate_add [symmetric]) lemma takefill_le': "n = m + k \ takefill x m (takefill x n l) = takefill x m l" by (induct m arbitrary: l n) (auto split: list.split) lemma length_takefill [simp]: "length (takefill fill n l) = n" by (simp add: takefill_alt) lemma take_takefill': "n = k + m \ take k (takefill fill n w) = takefill fill k w" by (induct k arbitrary: w n) (auto split: list.split) lemma drop_takefill: "drop k (takefill fill (m + k) w) = takefill fill m (drop k w)" by (induct k arbitrary: w) (auto split: list.split) lemma takefill_le [simp]: "m \ n \ takefill x m (takefill x n l) = takefill x m l" by (auto simp: le_iff_add takefill_le') lemma take_takefill [simp]: "m \ n \ take m (takefill fill n w) = takefill fill m w" by (auto simp: le_iff_add take_takefill') lemma takefill_append: "takefill fill (m + length xs) (xs @ w) = xs @ (takefill fill m w)" by (induct xs) auto lemma takefill_same': "l = length xs \ takefill fill l xs = xs" by (induct xs arbitrary: l) auto lemmas takefill_same [simp] = takefill_same' [OF refl] lemma tf_rev: "n + k = m + length bl \ takefill x m (rev (takefill y n bl)) = rev (takefill y m (rev (takefill x k (rev bl))))" apply (rule nth_equalityI) apply (auto simp add: nth_takefill rev_nth) apply (rule_tac f = "\n. bl ! n" in arg_cong) apply arith done lemma takefill_minus: "0 < n \ takefill fill (Suc (n - 1)) w = takefill fill n w" by auto lemmas takefill_Suc_cases = list.cases [THEN takefill.Suc [THEN trans]] lemmas takefill_Suc_Nil = takefill_Suc_cases (1) lemmas takefill_Suc_Cons = takefill_Suc_cases (2) lemmas takefill_minus_simps = takefill_Suc_cases [THEN [2] takefill_minus [symmetric, THEN trans]] lemma takefill_numeral_Nil [simp]: "takefill fill (numeral k) [] = fill # takefill fill (pred_numeral k) []" by (simp add: numeral_eq_Suc) lemma takefill_numeral_Cons [simp]: "takefill fill (numeral k) (x # xs) = x # takefill fill (pred_numeral k) xs" by (simp add: numeral_eq_Suc) subsection \Range projection\ definition bl_of_nth :: "nat \ (nat \ 'a) \ 'a list" where "bl_of_nth n f = map f (rev [0.. rev (bl_of_nth n f) ! m = f m" by (simp add: bl_of_nth_def rev_map) lemma bl_of_nth_inj: "(\k. k < n \ f k = g k) \ bl_of_nth n f = bl_of_nth n g" by (simp add: bl_of_nth_def) lemma bl_of_nth_nth_le: "n \ length xs \ bl_of_nth n (nth (rev xs)) = drop (length xs - n) xs" apply (induct n arbitrary: xs) apply clarsimp apply clarsimp apply (rule trans [OF _ hd_Cons_tl]) apply (frule Suc_le_lessD) apply (simp add: rev_nth trans [OF drop_Suc drop_tl, symmetric]) apply (subst hd_drop_conv_nth) apply force apply simp_all apply (rule_tac f = "\n. drop n xs" in arg_cong) apply simp done lemma bl_of_nth_nth [simp]: "bl_of_nth (length xs) ((!) (rev xs)) = xs" by (simp add: bl_of_nth_nth_le) subsection \More\ definition rotater1 :: "'a list \ 'a list" where "rotater1 ys = (case ys of [] \ [] | x # xs \ last ys # butlast ys)" definition rotater :: "nat \ 'a list \ 'a list" where "rotater n = rotater1 ^^ n" lemmas rotater_0' [simp] = rotater_def [where n = "0", simplified] lemma rotate1_rl': "rotater1 (l @ [a]) = a # l" by (cases l) (auto simp: rotater1_def) lemma rotate1_rl [simp] : "rotater1 (rotate1 l) = l" apply (unfold rotater1_def) apply (cases "l") apply (case_tac [2] "list") apply auto done lemma rotate1_lr [simp] : "rotate1 (rotater1 l) = l" by (cases l) (auto simp: rotater1_def) lemma rotater1_rev': "rotater1 (rev xs) = rev (rotate1 xs)" by (cases "xs") (simp add: rotater1_def, simp add: rotate1_rl') lemma rotater_rev': "rotater n (rev xs) = rev (rotate n xs)" by (induct n) (auto simp: rotater_def intro: rotater1_rev') lemma rotater_rev: "rotater n ys = rev (rotate n (rev ys))" using rotater_rev' [where xs = "rev ys"] by simp lemma rotater_drop_take: "rotater n xs = drop (length xs - n mod length xs) xs @ take (length xs - n mod length xs) xs" by (auto simp: rotater_rev rotate_drop_take rev_take rev_drop) lemma rotater_Suc [simp]: "rotater (Suc n) xs = rotater1 (rotater n xs)" unfolding rotater_def by auto lemma nth_rotater: \rotater m xs ! n = xs ! ((n + (length xs - m mod length xs)) mod length xs)\ if \n < length xs\ using that by (simp add: rotater_drop_take nth_append not_less less_diff_conv ac_simps le_mod_geq) lemma nth_rotater1: \rotater1 xs ! n = xs ! ((n + (length xs - 1)) mod length xs)\ if \n < length xs\ using that nth_rotater [of n xs 1] by simp lemma rotate_inv_plus [rule_format]: "\k. k = m + n \ rotater k (rotate n xs) = rotater m xs \ rotate k (rotater n xs) = rotate m xs \ rotater n (rotate k xs) = rotate m xs \ rotate n (rotater k xs) = rotater m xs" by (induct n) (auto simp: rotater_def rotate_def intro: funpow_swap1 [THEN trans]) lemmas rotate_inv_rel = le_add_diff_inverse2 [symmetric, THEN rotate_inv_plus] lemmas rotate_inv_eq = order_refl [THEN rotate_inv_rel, simplified] lemmas rotate_lr [simp] = rotate_inv_eq [THEN conjunct1] lemmas rotate_rl [simp] = rotate_inv_eq [THEN conjunct2, THEN conjunct1] lemma rotate_gal: "rotater n xs = ys \ rotate n ys = xs" by auto lemma rotate_gal': "ys = rotater n xs \ xs = rotate n ys" by auto lemma length_rotater [simp]: "length (rotater n xs) = length xs" by (simp add : rotater_rev) lemma rotate_eq_mod: "m mod length xs = n mod length xs \ rotate m xs = rotate n xs" apply (rule box_equals) defer apply (rule rotate_conv_mod [symmetric])+ apply simp done lemma restrict_to_left: "x = y \ x = z \ y = z" by simp lemmas rotate_eqs = trans [OF rotate0 [THEN fun_cong] id_apply] rotate_rotate [symmetric] rotate_id rotate_conv_mod rotate_eq_mod lemmas rrs0 = rotate_eqs [THEN restrict_to_left, simplified rotate_gal [symmetric] rotate_gal' [symmetric]] lemmas rrs1 = rrs0 [THEN refl [THEN rev_iffD1]] lemmas rotater_eqs = rrs1 [simplified length_rotater] lemmas rotater_0 = rotater_eqs (1) lemmas rotater_add = rotater_eqs (2) lemma butlast_map: "xs \ [] \ butlast (map f xs) = map f (butlast xs)" by (induct xs) auto lemma rotater1_map: "rotater1 (map f xs) = map f (rotater1 xs)" by (cases xs) (auto simp: rotater1_def last_map butlast_map) lemma rotater_map: "rotater n (map f xs) = map f (rotater n xs)" by (induct n) (auto simp: rotater_def rotater1_map) lemma but_last_zip [rule_format] : "\ys. length xs = length ys \ xs \ [] \ last (zip xs ys) = (last xs, last ys) \ butlast (zip xs ys) = zip (butlast xs) (butlast ys)" apply (induct xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ done lemma but_last_map2 [rule_format] : "\ys. length xs = length ys \ xs \ [] \ last (map2 f xs ys) = f (last xs) (last ys) \ butlast (map2 f xs ys) = map2 f (butlast xs) (butlast ys)" apply (induct xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ done lemma rotater1_zip: "length xs = length ys \ rotater1 (zip xs ys) = zip (rotater1 xs) (rotater1 ys)" apply (unfold rotater1_def) apply (cases xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv but_last_zip)[1])+ done lemma rotater1_map2: "length xs = length ys \ rotater1 (map2 f xs ys) = map2 f (rotater1 xs) (rotater1 ys)" by (simp add: rotater1_map rotater1_zip) lemmas lrth = box_equals [OF asm_rl length_rotater [symmetric] length_rotater [symmetric], THEN rotater1_map2] lemma rotater_map2: "length xs = length ys \ rotater n (map2 f xs ys) = map2 f (rotater n xs) (rotater n ys)" by (induct n) (auto intro!: lrth) lemma rotate1_map2: "length xs = length ys \ rotate1 (map2 f xs ys) = map2 f (rotate1 xs) (rotate1 ys)" by (cases xs; cases ys) auto lemmas lth = box_equals [OF asm_rl length_rotate [symmetric] length_rotate [symmetric], THEN rotate1_map2] lemma rotate_map2: "length xs = length ys \ rotate n (map2 f xs ys) = map2 f (rotate n xs) (rotate n ys)" by (induct n) (auto intro!: lth) subsection \Explicit bit representation of \<^typ>\int\\ primrec bl_to_bin_aux :: "bool list \ int \ int" where Nil: "bl_to_bin_aux [] w = w" | Cons: "bl_to_bin_aux (b # bs) w = bl_to_bin_aux bs (of_bool b + 2 * w)" definition bl_to_bin :: "bool list \ int" where "bl_to_bin bs = bl_to_bin_aux bs 0" primrec bin_to_bl_aux :: "nat \ int \ bool list \ bool list" where Z: "bin_to_bl_aux 0 w bl = bl" | Suc: "bin_to_bl_aux (Suc n) w bl = bin_to_bl_aux n (bin_rest w) ((bin_last w) # bl)" definition bin_to_bl :: "nat \ int \ bool list" where "bin_to_bl n w = bin_to_bl_aux n w []" lemma bin_to_bl_aux_zero_minus_simp [simp]: "0 < n \ bin_to_bl_aux n 0 bl = bin_to_bl_aux (n - 1) 0 (False # bl)" by (cases n) auto lemma bin_to_bl_aux_minus1_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (- 1) bl = bin_to_bl_aux (n - 1) (- 1) (True # bl)" by (cases n) auto lemma bin_to_bl_aux_one_minus_simp [simp]: "0 < n \ bin_to_bl_aux n 1 bl = bin_to_bl_aux (n - 1) 0 (True # bl)" by (cases n) auto lemma bin_to_bl_aux_Bit0_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (numeral (Num.Bit0 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (False # bl)" by (cases n) simp_all lemma bin_to_bl_aux_Bit1_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (numeral (Num.Bit1 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (True # bl)" by (cases n) simp_all lemma bl_to_bin_aux_append: "bl_to_bin_aux (bs @ cs) w = bl_to_bin_aux cs (bl_to_bin_aux bs w)" by (induct bs arbitrary: w) auto lemma bin_to_bl_aux_append: "bin_to_bl_aux n w bs @ cs = bin_to_bl_aux n w (bs @ cs)" by (induct n arbitrary: w bs) auto lemma bl_to_bin_append: "bl_to_bin (bs @ cs) = bl_to_bin_aux cs (bl_to_bin bs)" unfolding bl_to_bin_def by (rule bl_to_bin_aux_append) lemma bin_to_bl_aux_alt: "bin_to_bl_aux n w bs = bin_to_bl n w @ bs" by (simp add: bin_to_bl_def bin_to_bl_aux_append) lemma bin_to_bl_0 [simp]: "bin_to_bl 0 bs = []" by (auto simp: bin_to_bl_def) lemma size_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" by (induct n arbitrary: w bs) auto lemma size_bin_to_bl [simp]: "length (bin_to_bl n w) = n" by (simp add: bin_to_bl_def size_bin_to_bl_aux) lemma bl_bin_bl': "bin_to_bl (n + length bs) (bl_to_bin_aux bs w) = bin_to_bl_aux n w bs" apply (induct bs arbitrary: w n) apply auto apply (simp_all only: add_Suc [symmetric]) apply (auto simp add: bin_to_bl_def) done lemma bl_bin_bl [simp]: "bin_to_bl (length bs) (bl_to_bin bs) = bs" unfolding bl_to_bin_def apply (rule box_equals) apply (rule bl_bin_bl') prefer 2 apply (rule bin_to_bl_aux.Z) apply simp done lemma bl_to_bin_inj: "bl_to_bin bs = bl_to_bin cs \ length bs = length cs \ bs = cs" apply (rule_tac box_equals) defer apply (rule bl_bin_bl) apply (rule bl_bin_bl) apply simp done lemma bl_to_bin_False [simp]: "bl_to_bin (False # bl) = bl_to_bin bl" by (auto simp: bl_to_bin_def) lemma bl_to_bin_Nil [simp]: "bl_to_bin [] = 0" by (auto simp: bl_to_bin_def) lemma bin_to_bl_zero_aux: "bin_to_bl_aux n 0 bl = replicate n False @ bl" by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) lemma bin_to_bl_zero: "bin_to_bl n 0 = replicate n False" by (simp add: bin_to_bl_def bin_to_bl_zero_aux) lemma bin_to_bl_minus1_aux: "bin_to_bl_aux n (- 1) bl = replicate n True @ bl" by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) lemma bin_to_bl_minus1: "bin_to_bl n (- 1) = replicate n True" by (simp add: bin_to_bl_def bin_to_bl_minus1_aux) subsection \Semantic interpretation of \<^typ>\bool list\ as \<^typ>\int\\ lemma bin_bl_bin': "bl_to_bin (bin_to_bl_aux n w bs) = bl_to_bin_aux bs (bintrunc n w)" by (induct n arbitrary: w bs) (auto simp: bl_to_bin_def take_bit_Suc ac_simps mod_2_eq_odd) lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = bintrunc n w" by (auto simp: bin_to_bl_def bin_bl_bin') lemma bl_to_bin_rep_F: "bl_to_bin (replicate n False @ bl) = bl_to_bin bl" by (simp add: bin_to_bl_zero_aux [symmetric] bin_bl_bin') (simp add: bl_to_bin_def) lemma bin_to_bl_trunc [simp]: "n \ m \ bin_to_bl n (bintrunc m w) = bin_to_bl n w" by (auto intro: bl_to_bin_inj) lemma bin_to_bl_aux_bintr: "bin_to_bl_aux n (bintrunc m bin) bl = replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl" apply (induct n arbitrary: m bin bl) apply clarsimp apply clarsimp apply (case_tac "m") apply (clarsimp simp: bin_to_bl_zero_aux) apply (erule thin_rl) apply (induct_tac n) apply (auto simp add: take_bit_Suc) done lemma bin_to_bl_bintr: "bin_to_bl n (bintrunc m bin) = replicate (n - m) False @ bin_to_bl (min n m) bin" unfolding bin_to_bl_def by (rule bin_to_bl_aux_bintr) lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = 0" by (induct n) auto lemma len_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" by (fact size_bin_to_bl_aux) lemma len_bin_to_bl: "length (bin_to_bl n w) = n" by (fact size_bin_to_bl) (* FIXME: duplicate *) lemma sign_bl_bin': "bin_sign (bl_to_bin_aux bs w) = bin_sign w" by (induction bs arbitrary: w) (simp_all add: bin_sign_def) lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = 0" by (simp add: bl_to_bin_def sign_bl_bin') lemma bl_sbin_sign_aux: "hd (bin_to_bl_aux (Suc n) w bs) = (bin_sign (sbintrunc n w) = -1)" by (induction n arbitrary: w bs) (auto simp add: bin_sign_def even_iff_mod_2_eq_zero bit_Suc) lemma bl_sbin_sign: "hd (bin_to_bl (Suc n) w) = (bin_sign (sbintrunc n w) = -1)" unfolding bin_to_bl_def by (rule bl_sbin_sign_aux) lemma bin_nth_of_bl_aux: "bin_nth (bl_to_bin_aux bl w) n = (n < size bl \ rev bl ! n \ n \ length bl \ bin_nth w (n - size bl))" apply (induction bl arbitrary: w) apply simp_all apply safe apply (simp_all add: not_le nth_append bit_double_iff even_bit_succ_iff split: if_splits) done lemma bin_nth_of_bl: "bin_nth (bl_to_bin bl) n = (n < length bl \ rev bl ! n)" by (simp add: bl_to_bin_def bin_nth_of_bl_aux) lemma bin_nth_bl: "n < m \ bin_nth w n = nth (rev (bin_to_bl m w)) n" apply (induct n arbitrary: m w) apply clarsimp apply (case_tac m, clarsimp) apply (clarsimp simp: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt) apply (case_tac m, clarsimp) apply (clarsimp simp: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt bit_Suc) done lemma nth_bin_to_bl_aux: "n < m + length bl \ (bin_to_bl_aux m w bl) ! n = (if n < m then bit w (m - 1 - n) else bl ! (n - m))" apply (induction bl arbitrary: w) apply simp_all apply (simp add: bin_nth_bl [of \m - Suc n\ m] rev_nth flip: bin_to_bl_def) apply (metis One_nat_def Suc_pred add_diff_cancel_left' add_diff_cancel_right' bin_to_bl_aux_alt bin_to_bl_def diff_Suc_Suc diff_is_0_eq diff_zero less_Suc_eq_0_disj less_antisym less_imp_Suc_add list.size(3) nat_less_le nth_append size_bin_to_bl_aux) done lemma nth_bin_to_bl: "n < m \ (bin_to_bl m w) ! n = bin_nth w (m - Suc n)" by (simp add: bin_to_bl_def nth_bin_to_bl_aux) lemma takefill_bintrunc: "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))" apply (rule nth_equalityI) apply simp apply (clarsimp simp: nth_takefill rev_nth nth_bin_to_bl bin_nth_of_bl) done lemma bl_bin_bl_rtf: "bin_to_bl n (bl_to_bin bl) = rev (takefill False n (rev bl))" by (simp add: takefill_bintrunc) lemma bl_to_bin_lt2p_aux: "bl_to_bin_aux bs w < (w + 1) * (2 ^ length bs)" proof (induction bs arbitrary: w) case Nil then show ?case by simp next case (Cons b bs) from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] show ?case apply (auto simp add: algebra_simps) apply (subst mult_2 [of \2 ^ length bs\]) apply (simp only: add.assoc) apply (rule pos_add_strict) apply simp_all done qed lemma bl_to_bin_lt2p_drop: "bl_to_bin bs < 2 ^ length (dropWhile Not bs)" proof (induct bs) case Nil then show ?case by simp next case (Cons b bs) with bl_to_bin_lt2p_aux[where w=1] show ?case by (simp add: bl_to_bin_def) qed lemma bl_to_bin_lt2p: "bl_to_bin bs < 2 ^ length bs" by (metis bin_bl_bin bintr_lt2p bl_bin_bl) lemma bl_to_bin_ge2p_aux: "bl_to_bin_aux bs w \ w * (2 ^ length bs)" proof (induction bs arbitrary: w) case Nil then show ?case by simp next case (Cons b bs) from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] show ?case apply (auto simp add: algebra_simps) apply (rule add_le_imp_le_left [of \2 ^ length bs\]) apply (rule add_increasing) apply simp_all done qed lemma bl_to_bin_ge0: "bl_to_bin bs \ 0" apply (unfold bl_to_bin_def) apply (rule xtrans(4)) apply (rule bl_to_bin_ge2p_aux) apply simp done lemma butlast_rest_bin: "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (bin_rest w)" apply (unfold bin_to_bl_def) apply (cases n, clarsimp) apply clarsimp apply (auto simp add: bin_to_bl_aux_alt) done lemma butlast_bin_rest: "butlast bl = bin_to_bl (length bl - Suc 0) (bin_rest (bl_to_bin bl))" using butlast_rest_bin [where w="bl_to_bin bl" and n="length bl"] by simp lemma butlast_rest_bl2bin_aux: "bl \ [] \ bl_to_bin_aux (butlast bl) w = bin_rest (bl_to_bin_aux bl w)" by (induct bl arbitrary: w) auto lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bin_rest (bl_to_bin bl)" by (cases bl) (auto simp: bl_to_bin_def butlast_rest_bl2bin_aux) lemma trunc_bl2bin_aux: "bintrunc m (bl_to_bin_aux bl w) = bl_to_bin_aux (drop (length bl - m) bl) (bintrunc (m - length bl) w)" proof (induct bl arbitrary: w) case Nil show ?case by simp next case (Cons b bl) show ?case proof (cases "m - length bl") case 0 then have "Suc (length bl) - m = Suc (length bl - m)" by simp with Cons show ?thesis by simp next case (Suc n) then have "m - Suc (length bl) = n" by simp with Cons Suc show ?thesis by (simp add: take_bit_Suc ac_simps) qed qed lemma trunc_bl2bin: "bintrunc m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" by (simp add: bl_to_bin_def trunc_bl2bin_aux) lemma trunc_bl2bin_len [simp]: "bintrunc (length bl) (bl_to_bin bl) = bl_to_bin bl" by (simp add: trunc_bl2bin) lemma bl2bin_drop: "bl_to_bin (drop k bl) = bintrunc (length bl - k) (bl_to_bin bl)" apply (rule trans) prefer 2 apply (rule trunc_bl2bin [symmetric]) apply (cases "k \ length bl") apply auto done lemma take_rest_power_bin: "m \ n \ take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^^ (n - m)) w)" apply (rule nth_equalityI) apply simp apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin) done lemma last_bin_last': "size xs > 0 \ last xs \ bin_last (bl_to_bin_aux xs w)" by (induct xs arbitrary: w) auto lemma last_bin_last: "size xs > 0 \ last xs \ bin_last (bl_to_bin xs)" unfolding bl_to_bin_def by (erule last_bin_last') lemma bin_last_last: "bin_last w \ last (bin_to_bl (Suc n) w)" by (simp add: bin_to_bl_def) (auto simp: bin_to_bl_aux_alt) lemma drop_bin2bl_aux: "drop m (bin_to_bl_aux n bin bs) = bin_to_bl_aux (n - m) bin (drop (m - n) bs)" apply (induction n arbitrary: m bin bs) apply auto apply (case_tac "m \ n") apply (auto simp add: not_le Suc_diff_le) apply (case_tac "m - n") apply auto apply (use Suc_diff_Suc in fastforce) done lemma drop_bin2bl: "drop m (bin_to_bl n bin) = bin_to_bl (n - m) bin" by (simp add: bin_to_bl_def drop_bin2bl_aux) lemma take_bin2bl_lem1: "take m (bin_to_bl_aux m w bs) = bin_to_bl m w" apply (induct m arbitrary: w bs) apply clarsimp apply clarsimp apply (simp add: bin_to_bl_aux_alt) apply (simp add: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt) done lemma take_bin2bl_lem: "take m (bin_to_bl_aux (m + n) w bs) = take m (bin_to_bl (m + n) w)" by (induct n arbitrary: w bs) (simp_all (no_asm) add: bin_to_bl_def take_bin2bl_lem1, simp) lemma bin_split_take: "bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl (m + n) c)" apply (induct n arbitrary: b c) apply clarsimp apply (clarsimp simp: Let_def split: prod.split_asm) apply (simp add: bin_to_bl_def) apply (simp add: take_bin2bl_lem drop_bit_Suc) done lemma bin_to_bl_drop_bit: "k = m + n \ bin_to_bl m (drop_bit n c) = take m (bin_to_bl k c)" using bin_split_take by simp lemma bin_split_take1: "k = m + n \ bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl k c)" using bin_split_take by simp lemma bl_bin_bl_rep_drop: "bin_to_bl n (bl_to_bin bl) = replicate (n - length bl) False @ drop (length bl - n) bl" by (simp add: bl_to_bin_inj bl_to_bin_rep_F trunc_bl2bin) lemma bl_to_bin_aux_cat: "bl_to_bin_aux bs (bin_cat w nv v) = bin_cat w (nv + length bs) (bl_to_bin_aux bs v)" by (rule bit_eqI) (auto simp add: bin_nth_of_bl_aux bin_nth_cat algebra_simps) lemma bin_to_bl_aux_cat: "bin_to_bl_aux (nv + nw) (bin_cat v nw w) bs = bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)" by (induction nw arbitrary: w bs) (simp_all add: concat_bit_Suc) lemma bl_to_bin_aux_alt: "bl_to_bin_aux bs w = bin_cat w (length bs) (bl_to_bin bs)" using bl_to_bin_aux_cat [where nv = "0" and v = "0"] by (simp add: bl_to_bin_def [symmetric]) lemma bin_to_bl_cat: "bin_to_bl (nv + nw) (bin_cat v nw w) = bin_to_bl_aux nv v (bin_to_bl nw w)" by (simp add: bin_to_bl_def bin_to_bl_aux_cat) lemmas bl_to_bin_aux_app_cat = trans [OF bl_to_bin_aux_append bl_to_bin_aux_alt] lemmas bin_to_bl_aux_cat_app = trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt] lemma bl_to_bin_app_cat: "bl_to_bin (bsa @ bs) = bin_cat (bl_to_bin bsa) (length bs) (bl_to_bin bs)" by (simp only: bl_to_bin_aux_app_cat bl_to_bin_def) lemma bin_to_bl_cat_app: "bin_to_bl (n + nw) (bin_cat w nw wa) = bin_to_bl n w @ bin_to_bl nw wa" by (simp only: bin_to_bl_def bin_to_bl_aux_cat_app) text \\bl_to_bin_app_cat_alt\ and \bl_to_bin_app_cat\ are easily interderivable.\ lemma bl_to_bin_app_cat_alt: "bin_cat (bl_to_bin cs) n w = bl_to_bin (cs @ bin_to_bl n w)" by (simp add: bl_to_bin_app_cat) lemma mask_lem: "(bl_to_bin (True # replicate n False)) = bl_to_bin (replicate n True) + 1" apply (unfold bl_to_bin_def) apply (induct n) apply simp apply (simp only: Suc_eq_plus1 replicate_add append_Cons [symmetric] bl_to_bin_aux_append) apply simp done lemma bin_exhaust: "(\x b. bin = of_bool b + 2 * x \ Q) \ Q" for bin :: int apply (cases \even bin\) apply (auto elim!: evenE oddE) apply fastforce apply fastforce done primrec rbl_succ :: "bool list \ bool list" where Nil: "rbl_succ Nil = Nil" | Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)" primrec rbl_pred :: "bool list \ bool list" where Nil: "rbl_pred Nil = Nil" | Cons: "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)" primrec rbl_add :: "bool list \ bool list \ bool list" where \ \result is length of first arg, second arg may be longer\ Nil: "rbl_add Nil x = Nil" | Cons: "rbl_add (y # ys) x = (let ws = rbl_add ys (tl x) in (y \ hd x) # (if hd x \ y then rbl_succ ws else ws))" primrec rbl_mult :: "bool list \ bool list \ bool list" where \ \result is length of first arg, second arg may be longer\ Nil: "rbl_mult Nil x = Nil" | Cons: "rbl_mult (y # ys) x = (let ws = False # rbl_mult ys x in if y then rbl_add ws x else ws)" lemma size_rbl_pred: "length (rbl_pred bl) = length bl" by (induct bl) auto lemma size_rbl_succ: "length (rbl_succ bl) = length bl" by (induct bl) auto lemma size_rbl_add: "length (rbl_add bl cl) = length bl" by (induct bl arbitrary: cl) (auto simp: Let_def size_rbl_succ) lemma size_rbl_mult: "length (rbl_mult bl cl) = length bl" by (induct bl arbitrary: cl) (auto simp add: Let_def size_rbl_add) lemmas rbl_sizes [simp] = size_rbl_pred size_rbl_succ size_rbl_add size_rbl_mult lemmas rbl_Nils = rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil lemma rbl_add_app2: "length blb \ length bla \ rbl_add bla (blb @ blc) = rbl_add bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def) done lemma rbl_add_take2: "length blb \ length bla \ rbl_add bla (take (length bla) blb) = rbl_add bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def) done lemma rbl_mult_app2: "length blb \ length bla \ rbl_mult bla (blb @ blc) = rbl_mult bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def rbl_add_app2) done lemma rbl_mult_take2: "length blb \ length bla \ rbl_mult bla (take (length bla) blb) = rbl_mult bla blb" apply (rule trans) apply (rule rbl_mult_app2 [symmetric]) apply simp apply (rule_tac f = "rbl_mult bla" in arg_cong) apply (rule append_take_drop_id) done lemma rbl_add_split: "P (rbl_add (y # ys) (x # xs)) = (\ws. length ws = length ys \ ws = rbl_add ys xs \ (y \ ((x \ P (False # rbl_succ ws)) \ (\ x \ P (True # ws)))) \ (\ y \ P (x # ws)))" by (cases y) (auto simp: Let_def) lemma rbl_mult_split: "P (rbl_mult (y # ys) xs) = (\ws. length ws = Suc (length ys) \ ws = False # rbl_mult ys xs \ (y \ P (rbl_add ws xs)) \ (\ y \ P ws))" by (auto simp: Let_def) lemma rbl_pred: "rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin - 1))" proof (unfold bin_to_bl_def, induction n arbitrary: bin) case 0 then show ?case by simp next case (Suc n) obtain b k where \bin = of_bool b + 2 * k\ using bin_exhaust by blast moreover have \(2 * k - 1) div 2 = k - 1\ using even_succ_div_2 [of \2 * (k - 1)\] by simp ultimately show ?case using Suc [of \bin div 2\] by simp (simp add: bin_to_bl_aux_alt) qed lemma rbl_succ: "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))" apply (unfold bin_to_bl_def) apply (induction n arbitrary: bin) apply simp_all apply (case_tac bin rule: bin_exhaust) apply simp apply (simp add: bin_to_bl_aux_alt ac_simps) done lemma rbl_add: "\bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = rev (bin_to_bl n (bina + binb))" apply (unfold bin_to_bl_def) apply (induct n) apply simp apply clarsimp apply (case_tac bina rule: bin_exhaust) apply (case_tac binb rule: bin_exhaust) apply (case_tac b) apply (case_tac [!] "ba") apply (auto simp: rbl_succ bin_to_bl_aux_alt Let_def ac_simps) done lemma rbl_add_long: "m \ n \ rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = rev (bin_to_bl n (bina + binb))" apply (rule box_equals [OF _ rbl_add_take2 rbl_add]) apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong) apply (rule rev_swap [THEN iffD1]) apply (simp add: rev_take drop_bin2bl) apply simp done lemma rbl_mult_gt1: "m \ length bl \ rbl_mult bl (rev (bin_to_bl m binb)) = rbl_mult bl (rev (bin_to_bl (length bl) binb))" apply (rule trans) apply (rule rbl_mult_take2 [symmetric]) apply simp_all apply (rule_tac f = "rbl_mult bl" in arg_cong) apply (rule rev_swap [THEN iffD1]) apply (simp add: rev_take drop_bin2bl) done lemma rbl_mult_gt: "m > n \ rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))" by (auto intro: trans [OF rbl_mult_gt1]) lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt] lemma rbbl_Cons: "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (of_bool b + 2 * x))" by (simp add: bin_to_bl_def) (simp add: bin_to_bl_aux_alt) lemma rbl_mult: "rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = rev (bin_to_bl n (bina * binb))" apply (induct n arbitrary: bina binb) apply simp_all apply (unfold bin_to_bl_def) apply clarsimp apply (case_tac bina rule: bin_exhaust) apply (case_tac binb rule: bin_exhaust) apply simp apply (simp add: bin_to_bl_aux_alt) apply (simp add: rbbl_Cons rbl_mult_Suc rbl_add algebra_simps) done lemma sclem: "size (concat (map (bin_to_bl n) xs)) = length xs * n" by (induct xs) auto lemma bin_cat_foldl_lem: "foldl (\u. bin_cat u n) x xs = bin_cat x (size xs * n) (foldl (\u. bin_cat u n) y xs)" apply (induct xs arbitrary: x) apply simp apply (simp (no_asm)) apply (frule asm_rl) apply (drule meta_spec) apply (erule trans) apply (drule_tac x = "bin_cat y n a" in meta_spec) apply (simp add: bin_cat_assoc_sym min.absorb2) done lemma bin_rcat_bl: "bin_rcat n wl = bl_to_bin (concat (map (bin_to_bl n) wl))" apply (unfold bin_rcat_eq_foldl) apply (rule sym) apply (induct wl) apply (auto simp add: bl_to_bin_append) apply (simp add: bl_to_bin_aux_alt sclem) apply (simp add: bin_cat_foldl_lem [symmetric]) done lemma bin_last_bl_to_bin: "bin_last (bl_to_bin bs) \ bs \ [] \ last bs" by(cases "bs = []")(auto simp add: bl_to_bin_def last_bin_last'[where w=0]) lemma bin_rest_bl_to_bin: "bin_rest (bl_to_bin bs) = bl_to_bin (butlast bs)" by(cases "bs = []")(simp_all add: bl_to_bin_def butlast_rest_bl2bin_aux) lemma bl_xor_aux_bin: "map2 (\x y. x \ y) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v XOR w) (map2 (\x y. x \ y) bs cs)" apply (induction n arbitrary: v w bs cs) apply auto apply (case_tac v rule: bin_exhaust) apply (case_tac w rule: bin_exhaust) apply clarsimp done lemma bl_or_aux_bin: "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v OR w) (map2 (\) bs cs)" by (induct n arbitrary: v w bs cs) simp_all lemma bl_and_aux_bin: "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v AND w) (map2 (\) bs cs)" by (induction n arbitrary: v w bs cs) simp_all lemma bl_not_aux_bin: "map Not (bin_to_bl_aux n w cs) = bin_to_bl_aux n (NOT w) (map Not cs)" by (induct n arbitrary: w cs) auto lemma bl_not_bin: "map Not (bin_to_bl n w) = bin_to_bl n (NOT w)" by (simp add: bin_to_bl_def bl_not_aux_bin) lemma bl_and_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v AND w)" by (simp add: bin_to_bl_def bl_and_aux_bin) lemma bl_or_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v OR w)" by (simp add: bin_to_bl_def bl_or_aux_bin) lemma bl_xor_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v XOR w)" using bl_xor_aux_bin by (simp add: bin_to_bl_def) subsection \Type \<^typ>\'a word\\ lift_definition of_bl :: \bool list \ 'a::len word\ is bl_to_bin . lift_definition to_bl :: \'a::len word \ bool list\ is \bin_to_bl LENGTH('a)\ by (simp add: bl_to_bin_inj) lemma to_bl_eq: \to_bl w = bin_to_bl (LENGTH('a)) (uint w)\ for w :: \'a::len word\ by transfer simp lemma bit_of_bl_iff [bit_simps]: \bit (of_bl bs :: 'a word) n \ rev bs ! n \ n < LENGTH('a::len) \ n < length bs\ by transfer (simp add: bin_nth_of_bl ac_simps) lemma rev_to_bl_eq: \rev (to_bl w) = map (bit w) [0.. for w :: \'a::len word\ apply (rule nth_equalityI) apply (simp add: to_bl.rep_eq) apply (simp add: bin_nth_bl bit_word.rep_eq to_bl.rep_eq) done lemma to_bl_eq_rev: \to_bl w = map (bit w) (rev [0.. for w :: \'a::len word\ using rev_to_bl_eq [of w] apply (subst rev_is_rev_conv [symmetric]) apply (simp add: rev_map) done lemma of_bl_rev_eq: \of_bl (rev bs) = horner_sum of_bool 2 bs\ apply (rule bit_word_eqI) apply (simp add: bit_of_bl_iff) apply transfer apply (simp add: bit_horner_sum_bit_iff ac_simps) done lemma of_bl_eq: \of_bl bs = horner_sum of_bool 2 (rev bs)\ using of_bl_rev_eq [of \rev bs\] by simp lemma bshiftr1_eq: \bshiftr1 b w = of_bl (b # butlast (to_bl w))\ apply transfer apply auto apply (subst bl_to_bin_app_cat [of \[True]\, simplified]) apply simp apply (metis One_nat_def add.commute bin_bl_bin bin_last_bl_to_bin bin_rest_bl_to_bin butlast_bin_rest concat_bit_eq last.simps list.distinct(1) list.size(3) list.size(4) odd_iff_mod_2_eq_one plus_1_eq_Suc power_Suc0_right push_bit_of_1 size_bin_to_bl take_bit_eq_mod trunc_bl2bin_len) apply (simp add: butlast_rest_bl2bin) done lemma length_to_bl_eq: \length (to_bl w) = LENGTH('a)\ for w :: \'a::len word\ by transfer simp lemma word_rotr_eq: \word_rotr n w = of_bl (rotater n (to_bl w))\ apply (rule bit_word_eqI) subgoal for n apply (cases \n < LENGTH('a)\) apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff rotater_rev length_to_bl_eq nth_rotate rev_to_bl_eq ac_simps) done done lemma word_rotl_eq: \word_rotl n w = of_bl (rotate n (to_bl w))\ proof - have \rotate n (to_bl w) = rev (rotater n (rev (to_bl w)))\ by (simp add: rotater_rev') then show ?thesis apply (simp add: word_rotl_eq_word_rotr bit_of_bl_iff length_to_bl_eq rev_to_bl_eq) apply (rule bit_word_eqI) subgoal for n apply (cases \n < LENGTH('a)\) apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff nth_rotater) done done qed lemma to_bl_def': "(to_bl :: 'a::len word \ bool list) = bin_to_bl (LENGTH('a)) \ uint" by transfer (simp add: fun_eq_iff) \ \type definitions theorem for in terms of equivalent bool list\ lemma td_bl: "type_definition (to_bl :: 'a::len word \ bool list) of_bl {bl. length bl = LENGTH('a)}" apply (standard; transfer) apply (auto dest: sym) done interpretation word_bl: type_definition "to_bl :: 'a::len word \ bool list" of_bl "{bl. length bl = LENGTH('a::len)}" by (fact td_bl) lemmas word_bl_Rep' = word_bl.Rep [unfolded mem_Collect_eq, iff] lemma word_size_bl: "size w = size (to_bl w)" by (auto simp: word_size) lemma to_bl_use_of_bl: "to_bl w = bl \ w = of_bl bl \ length bl = length (to_bl w)" by (fastforce elim!: word_bl.Abs_inverse [unfolded mem_Collect_eq]) lemma length_bl_gt_0 [iff]: "0 < length (to_bl x)" for x :: "'a::len word" unfolding word_bl_Rep' by (rule len_gt_0) lemma bl_not_Nil [iff]: "to_bl x \ []" for x :: "'a::len word" by (fact length_bl_gt_0 [unfolded length_greater_0_conv]) lemma length_bl_neq_0 [iff]: "length (to_bl x) \ 0" for x :: "'a::len word" by (fact length_bl_gt_0 [THEN gr_implies_not0]) lemma hd_bl_sign_sint: "hd (to_bl w) = (bin_sign (sint w) = -1)" apply transfer apply (auto simp add: bin_sign_def) using bin_sign_lem bl_sbin_sign apply fastforce using bin_sign_lem bl_sbin_sign apply force done lemma of_bl_drop': "lend = length bl - LENGTH('a::len) \ of_bl (drop lend bl) = (of_bl bl :: 'a word)" by transfer (simp flip: trunc_bl2bin) lemma test_bit_of_bl: "(of_bl bl::'a::len word) !! n = (rev bl ! n \ n < LENGTH('a) \ n < length bl)" by transfer (simp add: bin_nth_of_bl ac_simps) lemma no_of_bl: "(numeral bin ::'a::len word) = of_bl (bin_to_bl (LENGTH('a)) (numeral bin))" by transfer simp lemma uint_bl: "to_bl w = bin_to_bl (size w) (uint w)" by transfer simp lemma to_bl_bin: "bl_to_bin (to_bl w) = uint w" by (simp add: uint_bl word_size) lemma to_bl_of_bin: "to_bl (word_of_int bin::'a::len word) = bin_to_bl (LENGTH('a)) bin" by (auto simp: uint_bl word_ubin.eq_norm word_size) lemma to_bl_numeral [simp]: "to_bl (numeral bin::'a::len word) = bin_to_bl (LENGTH('a)) (numeral bin)" unfolding word_numeral_alt by (rule to_bl_of_bin) lemma to_bl_neg_numeral [simp]: "to_bl (- numeral bin::'a::len word) = bin_to_bl (LENGTH('a)) (- numeral bin)" unfolding word_neg_numeral_alt by (rule to_bl_of_bin) lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w" by (simp add: uint_bl word_size) lemma uint_bl_bin: "bl_to_bin (bin_to_bl (LENGTH('a)) (uint x)) = uint x" for x :: "'a::len word" by (rule trans [OF bin_bl_bin word_ubin.norm_Rep]) lemma ucast_bl: "ucast w = of_bl (to_bl w)" by transfer simp lemma ucast_down_bl: \(ucast :: 'a::len word \ 'b::len word) (of_bl bl) = of_bl bl\ if \is_down (ucast :: 'a::len word \ 'b::len word)\ using that by transfer simp lemma of_bl_append_same: "of_bl (X @ to_bl w) = w" by transfer (simp add: bl_to_bin_app_cat) lemma ucast_of_bl_up: \ucast (of_bl bl :: 'a::len word) = of_bl bl\ if \size bl \ size (of_bl bl :: 'a::len word)\ using that apply transfer apply (rule bit_eqI) apply (auto simp add: bit_take_bit_iff) apply (subst (asm) trunc_bl2bin_len [symmetric]) apply (auto simp only: bit_take_bit_iff) done lemma word_rev_tf: "to_bl (of_bl bl::'a::len word) = rev (takefill False (LENGTH('a)) (rev bl))" by transfer (simp add: bl_bin_bl_rtf) lemma word_rep_drop: "to_bl (of_bl bl::'a::len word) = replicate (LENGTH('a) - length bl) False @ drop (length bl - LENGTH('a)) bl" by (simp add: word_rev_tf takefill_alt rev_take) lemma to_bl_ucast: "to_bl (ucast (w::'b::len word) ::'a::len word) = replicate (LENGTH('a) - LENGTH('b)) False @ drop (LENGTH('b) - LENGTH('a)) (to_bl w)" apply (unfold ucast_bl) apply (rule trans) apply (rule word_rep_drop) apply simp done lemma ucast_up_app: \to_bl (ucast w :: 'b::len word) = replicate n False @ (to_bl w)\ if \source_size (ucast :: 'a word \ 'b word) + n = target_size (ucast :: 'a word \ 'b word)\ for w :: \'a::len word\ using that by (auto simp add : source_size target_size to_bl_ucast) lemma ucast_down_drop [OF refl]: "uc = ucast \ source_size uc = target_size uc + n \ to_bl (uc w) = drop n (to_bl w)" by (auto simp add : source_size target_size to_bl_ucast) lemma scast_down_drop [OF refl]: "sc = scast \ source_size sc = target_size sc + n \ to_bl (sc w) = drop n (to_bl w)" apply (subgoal_tac "sc = ucast") apply safe apply simp apply (erule ucast_down_drop) apply (rule down_cast_same [symmetric]) apply (simp add : source_size target_size is_down) done lemma word_0_bl [simp]: "of_bl [] = 0" by transfer simp lemma word_1_bl: "of_bl [True] = 1" by transfer (simp add: bl_to_bin_def) lemma of_bl_0 [simp]: "of_bl (replicate n False) = 0" by transfer (simp add: bl_to_bin_rep_False) lemma to_bl_0 [simp]: "to_bl (0::'a::len word) = replicate (LENGTH('a)) False" by (simp add: uint_bl word_size bin_to_bl_zero) \ \links with \rbl\ operations\ lemma word_succ_rbl: "to_bl w = bl \ to_bl (word_succ w) = rev (rbl_succ (rev bl))" by transfer (simp add: rbl_succ) lemma word_pred_rbl: "to_bl w = bl \ to_bl (word_pred w) = rev (rbl_pred (rev bl))" by transfer (simp add: rbl_pred) lemma word_add_rbl: "to_bl v = vbl \ to_bl w = wbl \ to_bl (v + w) = rev (rbl_add (rev vbl) (rev wbl))" apply transfer apply (drule sym) apply (drule sym) apply (simp add: rbl_add) done lemma word_mult_rbl: "to_bl v = vbl \ to_bl w = wbl \ to_bl (v * w) = rev (rbl_mult (rev vbl) (rev wbl))" apply transfer apply (drule sym) apply (drule sym) apply (simp add: rbl_mult) done lemma rtb_rbl_ariths: "rev (to_bl w) = ys \ rev (to_bl (word_succ w)) = rbl_succ ys" "rev (to_bl w) = ys \ rev (to_bl (word_pred w)) = rbl_pred ys" "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v * w)) = rbl_mult ys xs" "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v + w)) = rbl_add ys xs" by (auto simp: rev_swap [symmetric] word_succ_rbl word_pred_rbl word_mult_rbl word_add_rbl) lemma of_bl_length_less: \(of_bl x :: 'a::len word) < 2 ^ k\ if \length x = k\ \k < LENGTH('a)\ proof - from that have \length x < LENGTH('a)\ by simp then have \(of_bl x :: 'a::len word) < 2 ^ length x\ apply (simp add: of_bl_eq) apply transfer apply (simp add: take_bit_horner_sum_bit_eq) apply (subst length_rev [symmetric]) apply (simp only: horner_sum_of_bool_2_less) done with that show ?thesis by simp qed lemma word_eq_rbl_eq: "x = y \ rev (to_bl x) = rev (to_bl y)" by simp lemma bl_word_not: "to_bl (NOT w) = map Not (to_bl w)" by transfer (simp add: bl_not_bin) lemma bl_word_xor: "to_bl (v XOR w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_xor_bin) lemma bl_word_or: "to_bl (v OR w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_or_bin) lemma bl_word_and: "to_bl (v AND w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_and_bin) lemma bin_nth_uint': "bin_nth (uint w) n \ rev (bin_to_bl (size w) (uint w)) ! n \ n < size w" apply (unfold word_size) apply (safe elim!: bin_nth_uint_imp) apply (frule bin_nth_uint_imp) apply (fast dest!: bin_nth_bl)+ done lemmas bin_nth_uint = bin_nth_uint' [unfolded word_size] lemma test_bit_bl: "w !! n \ rev (to_bl w) ! n \ n < size w" by transfer (auto simp add: bin_nth_bl) lemma to_bl_nth: "n < size w \ to_bl w ! n = w !! (size w - Suc n)" by (simp add: word_size rev_nth test_bit_bl) lemma map_bit_interval_eq: \map (bit w) [0.. for w :: \'a::len word\ proof (rule nth_equalityI) show \length (map (bit w) [0.. by simp fix m assume \m < length (map (bit w) [0.. then have \m < n\ by simp then have \bit w m \ takefill False n (rev (to_bl w)) ! m\ by (auto simp add: nth_takefill not_less rev_nth to_bl_nth word_size test_bit_word_eq dest: bit_imp_le_length) with \m < n \show \map (bit w) [0.. takefill False n (rev (to_bl w)) ! m\ by simp qed lemma to_bl_unfold: \to_bl w = rev (map (bit w) [0.. for w :: \'a::len word\ by (simp add: map_bit_interval_eq takefill_bintrunc to_bl_def flip: bin_to_bl_def) lemma nth_rev_to_bl: \rev (to_bl w) ! n \ bit w n\ if \n < LENGTH('a)\ for w :: \'a::len word\ using that by (simp add: to_bl_unfold) lemma nth_to_bl: \to_bl w ! n \ bit w (LENGTH('a) - Suc n)\ if \n < LENGTH('a)\ for w :: \'a::len word\ using that by (simp add: to_bl_unfold rev_nth) lemma of_bl_rep_False: "of_bl (replicate n False @ bs) = of_bl bs" by (auto simp: of_bl_def bl_to_bin_rep_F) lemma [code abstract]: \uint (of_bl bs :: 'a word) = horner_sum of_bool 2 (take LENGTH('a::len) (rev bs))\ apply (simp add: of_bl_eq flip: take_bit_horner_sum_bit_eq) apply transfer apply simp done lemma [code]: \to_bl w = map (bit w) (rev [0.. for w :: \'a::len word\ by (simp add: to_bl_unfold rev_map) lemma word_reverse_eq_of_bl_rev_to_bl: \word_reverse w = of_bl (rev (to_bl w))\ by (rule bit_word_eqI) (auto simp add: bit_word_reverse_iff bit_of_bl_iff nth_to_bl) lemmas word_reverse_no_def [simp] = word_reverse_eq_of_bl_rev_to_bl [of "numeral w"] for w lemma to_bl_word_rev: "to_bl (word_reverse w) = rev (to_bl w)" by (rule nth_equalityI) (simp_all add: nth_rev_to_bl word_reverse_def word_rep_drop flip: of_bl_eq) lemma to_bl_n1 [simp]: "to_bl (-1::'a::len word) = replicate (LENGTH('a)) True" apply (rule word_bl.Abs_inverse') apply simp apply (rule word_eqI) apply (clarsimp simp add: word_size) apply (auto simp add: word_bl.Abs_inverse test_bit_bl word_size) done lemma rbl_word_or: "rev (to_bl (x OR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_or rev_map) lemma rbl_word_and: "rev (to_bl (x AND y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_and rev_map) lemma rbl_word_xor: "rev (to_bl (x XOR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_xor rev_map) lemma rbl_word_not: "rev (to_bl (NOT x)) = map Not (rev (to_bl x))" by (simp add: bl_word_not rev_map) lemma bshiftr1_numeral [simp]: \bshiftr1 b (numeral w :: 'a word) = of_bl (b # butlast (bin_to_bl LENGTH('a::len) (numeral w)))\ by (simp add: bshiftr1_eq) lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)" unfolding bshiftr1_eq by (rule word_bl.Abs_inverse) simp lemma shiftl1_of_bl: "shiftl1 (of_bl bl) = of_bl (bl @ [False])" by transfer (simp add: bl_to_bin_append) lemma shiftl1_bl: "shiftl1 w = of_bl (to_bl w @ [False])" for w :: "'a::len word" proof - have "shiftl1 w = shiftl1 (of_bl (to_bl w))" by simp also have "\ = of_bl (to_bl w @ [False])" by (rule shiftl1_of_bl) finally show ?thesis . qed lemma bl_shiftl1: "to_bl (shiftl1 w) = tl (to_bl w) @ [False]" for w :: "'a::len word" by (simp add: shiftl1_bl word_rep_drop drop_Suc drop_Cons') (fast intro!: Suc_leI) \ \Generalized version of \bl_shiftl1\. Maybe this one should replace it?\ lemma bl_shiftl1': "to_bl (shiftl1 w) = tl (to_bl w @ [False])" by (simp add: shiftl1_bl word_rep_drop drop_Suc del: drop_append) lemma shiftr1_bl: \shiftr1 w = of_bl (butlast (to_bl w))\ proof (rule bit_word_eqI) fix n assume \n < LENGTH('a)\ show \bit (shiftr1 w) n \ bit (of_bl (butlast (to_bl w)) :: 'a word) n\ proof (cases \n = LENGTH('a) - 1\) case True then show ?thesis by (simp add: bit_shiftr1_iff bit_of_bl_iff) next case False with \n < LENGTH('a)\ have \n < LENGTH('a) - 1\ by simp with \n < LENGTH('a)\ show ?thesis by (simp add: bit_shiftr1_iff bit_of_bl_iff rev_nth nth_butlast word_size test_bit_word_eq to_bl_nth) qed qed lemma bl_shiftr1: "to_bl (shiftr1 w) = False # butlast (to_bl w)" for w :: "'a::len word" by (simp add: shiftr1_bl word_rep_drop len_gt_0 [THEN Suc_leI]) \ \Generalized version of \bl_shiftr1\. Maybe this one should replace it?\ lemma bl_shiftr1': "to_bl (shiftr1 w) = butlast (False # to_bl w)" apply (rule word_bl.Abs_inverse') apply (simp del: butlast.simps) apply (simp add: shiftr1_bl of_bl_def) done lemma bl_sshiftr1: "to_bl (sshiftr1 w) = hd (to_bl w) # butlast (to_bl w)" for w :: "'a::len word" proof (rule nth_equalityI) fix n assume \n < length (to_bl (sshiftr1 w))\ then have \n < LENGTH('a)\ by simp then show \to_bl (sshiftr1 w) ! n \ (hd (to_bl w) # butlast (to_bl w)) ! n\ apply (cases n) apply (simp_all add: to_bl_nth word_size hd_conv_nth test_bit_eq_bit bit_sshiftr1_iff nth_butlast Suc_diff_Suc nth_to_bl) done qed simp lemma drop_shiftr: "drop n (to_bl (w >> n)) = take (size w - n) (to_bl w)" for w :: "'a::len word" apply (unfold shiftr_def) apply (induct n) prefer 2 apply (simp add: drop_Suc bl_shiftr1 butlast_drop [symmetric]) apply (rule butlast_take [THEN trans]) apply (auto simp: word_size) done lemma drop_sshiftr: "drop n (to_bl (w >>> n)) = take (size w - n) (to_bl w)" for w :: "'a::len word" apply (simp_all add: word_size sshiftr_eq) apply (rule nth_equalityI) apply (simp_all add: word_size nth_to_bl bit_signed_drop_bit_iff) done lemma take_shiftr: "n \ size w \ take n (to_bl (w >> n)) = replicate n False" apply (unfold shiftr_def) apply (induct n) prefer 2 apply (simp add: bl_shiftr1' length_0_conv [symmetric] word_size) apply (rule take_butlast [THEN trans]) apply (auto simp: word_size) done lemma take_sshiftr': "n \ size w \ hd (to_bl (w >>> n)) = hd (to_bl w) \ take n (to_bl (w >>> n)) = replicate n (hd (to_bl w))" for w :: "'a::len word" apply (auto simp add: sshiftr_eq hd_bl_sign_sint bin_sign_def not_le word_size sint_signed_drop_bit_eq) apply (rule nth_equalityI) apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff) apply (rule nth_equalityI) apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff) done lemmas hd_sshiftr = take_sshiftr' [THEN conjunct1] lemmas take_sshiftr = take_sshiftr' [THEN conjunct2] lemma atd_lem: "take n xs = t \ drop n xs = d \ xs = t @ d" by (auto intro: append_take_drop_id [symmetric]) lemmas bl_shiftr = atd_lem [OF take_shiftr drop_shiftr] lemmas bl_sshiftr = atd_lem [OF take_sshiftr drop_sshiftr] lemma shiftl_of_bl: "of_bl bl << n = of_bl (bl @ replicate n False)" by (induct n) (auto simp: shiftl_def shiftl1_of_bl replicate_app_Cons_same) lemma shiftl_bl: "w << n = of_bl (to_bl w @ replicate n False)" for w :: "'a::len word" proof - have "w << n = of_bl (to_bl w) << n" by simp also have "\ = of_bl (to_bl w @ replicate n False)" by (rule shiftl_of_bl) finally show ?thesis . qed lemma bl_shiftl: "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False" by (simp add: shiftl_bl word_rep_drop word_size) lemma shiftr1_bl_of: "length bl \ LENGTH('a) \ shiftr1 (of_bl bl::'a::len word) = of_bl (butlast bl)" by transfer (simp add: butlast_rest_bl2bin trunc_bl2bin) lemma shiftr_bl_of: "length bl \ LENGTH('a) \ (of_bl bl::'a::len word) >> n = of_bl (take (length bl - n) bl)" apply (unfold shiftr_def) apply (induct n) apply clarsimp apply clarsimp apply (subst shiftr1_bl_of) apply simp apply (simp add: butlast_take) done lemma shiftr_bl: "x >> n \ of_bl (take (LENGTH('a) - n) (to_bl x))" for x :: "'a::len word" using shiftr_bl_of [where 'a='a, of "to_bl x"] by simp lemma aligned_bl_add_size [OF refl]: "size x - n = m \ n \ size x \ drop m (to_bl x) = replicate n False \ take m (to_bl y) = replicate m False \ to_bl (x + y) = take m (to_bl x) @ drop m (to_bl y)" for x :: \'a::len word\ apply (subgoal_tac "x AND y = 0") prefer 2 apply (rule word_bl.Rep_eqD) apply (simp add: bl_word_and) apply (rule align_lem_and [THEN trans]) apply (simp_all add: word_size)[5] apply simp apply (subst word_plus_and_or [symmetric]) apply (simp add : bl_word_or) apply (rule align_lem_or) apply (simp_all add: word_size) done lemma mask_bl: "mask n = of_bl (replicate n True)" by (auto simp add : test_bit_of_bl word_size intro: word_eqI) lemma bl_and_mask': "to_bl (w AND mask n :: 'a::len word) = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl w)" apply (rule nth_equalityI) apply simp apply (clarsimp simp add: to_bl_nth word_size) apply (auto simp add: word_size test_bit_bl nth_append rev_nth) done lemma slice1_eq_of_bl: \(slice1 n w :: 'b::len word) = of_bl (takefill False n (to_bl w))\ for w :: \'a::len word\ proof (rule bit_word_eqI) fix m assume \m < LENGTH('b)\ show \bit (slice1 n w :: 'b::len word) m \ bit (of_bl (takefill False n (to_bl w)) :: 'b word) m\ by (cases \m \ n\; cases \LENGTH('a) \ n\) (auto simp add: bit_slice1_iff bit_of_bl_iff not_less rev_nth not_le nth_takefill nth_to_bl algebra_simps) qed lemma slice1_no_bin [simp]: "slice1 n (numeral w :: 'b word) = of_bl (takefill False n (bin_to_bl (LENGTH('b::len)) (numeral w)))" by (simp add: slice1_eq_of_bl) (* TODO: neg_numeral *) lemma slice_no_bin [simp]: "slice n (numeral w :: 'b word) = of_bl (takefill False (LENGTH('b::len) - n) (bin_to_bl (LENGTH('b::len)) (numeral w)))" by (simp add: slice_def) (* TODO: neg_numeral *) lemma slice_take': "slice n w = of_bl (take (size w - n) (to_bl w))" by (simp add: slice_def word_size slice1_eq_of_bl takefill_alt) lemmas slice_take = slice_take' [unfolded word_size] \ \shiftr to a word of the same size is just slice, slice is just shiftr then ucast\ lemmas shiftr_slice = trans [OF shiftr_bl [THEN meta_eq_to_obj_eq] slice_take [symmetric]] lemma slice1_down_alt': "sl = slice1 n w \ fs = size sl \ fs + k = n \ to_bl sl = takefill False fs (drop k (to_bl w))" apply (simp add: slice1_eq_of_bl) apply transfer apply (simp add: bl_bin_bl_rep_drop) using drop_takefill apply force done lemma slice1_up_alt': "sl = slice1 n w \ fs = size sl \ fs = n + k \ to_bl sl = takefill False fs (replicate k False @ (to_bl w))" apply (simp add: slice1_eq_of_bl) apply transfer apply (simp add: bl_bin_bl_rep_drop flip: takefill_append) apply (metis diff_add_inverse) done lemmas sd1 = slice1_down_alt' [OF refl refl, unfolded word_size] lemmas su1 = slice1_up_alt' [OF refl refl, unfolded word_size] lemmas slice1_down_alt = le_add_diff_inverse [THEN sd1] lemmas slice1_up_alts = le_add_diff_inverse [symmetric, THEN su1] le_add_diff_inverse2 [symmetric, THEN su1] lemma slice1_tf_tf': "to_bl (slice1 n w :: 'a::len word) = rev (takefill False (LENGTH('a)) (rev (takefill False n (to_bl w))))" unfolding slice1_eq_of_bl by (rule word_rev_tf) lemmas slice1_tf_tf = slice1_tf_tf' [THEN word_bl.Rep_inverse', symmetric] lemma revcast_eq_of_bl: \(revcast w :: 'b::len word) = of_bl (takefill False (LENGTH('b)) (to_bl w))\ for w :: \'a::len word\ by (simp add: revcast_def slice1_eq_of_bl) lemmas revcast_no_def [simp] = revcast_eq_of_bl [where w="numeral w", unfolded word_size] for w lemma to_bl_revcast: "to_bl (revcast w :: 'a::len word) = takefill False (LENGTH('a)) (to_bl w)" apply (rule nth_equalityI) apply simp apply (cases \LENGTH('a) \ LENGTH('b)\) apply (auto simp add: nth_to_bl nth_takefill bit_revcast_iff) done lemma word_cat_bl: "word_cat a b = of_bl (to_bl a @ to_bl b)" apply (rule bit_word_eqI) apply (simp add: bit_word_cat_iff bit_of_bl_iff nth_append not_less nth_rev_to_bl) apply (meson bit_word.rep_eq less_diff_conv2 nth_rev_to_bl) done lemma of_bl_append: "(of_bl (xs @ ys) :: 'a::len word) = of_bl xs * 2^(length ys) + of_bl ys" apply transfer apply (simp add: bl_to_bin_app_cat bin_cat_num) done lemma of_bl_False [simp]: "of_bl (False#xs) = of_bl xs" by (rule word_eqI) (auto simp: test_bit_of_bl nth_append) lemma of_bl_True [simp]: "(of_bl (True # xs) :: 'a::len word) = 2^length xs + of_bl xs" by (subst of_bl_append [where xs="[True]", simplified]) (simp add: word_1_bl) lemma of_bl_Cons: "of_bl (x#xs) = of_bool x * 2^length xs + of_bl xs" by (cases x) simp_all lemma word_split_bl': "std = size c - size b \ (word_split c = (a, b)) \ (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c)))" apply (simp add: word_split_def) apply transfer apply (cases \LENGTH('b) \ LENGTH('a)\) apply (auto simp add: drop_bit_take_bit drop_bin2bl bin_to_bl_drop_bit [symmetric, of \LENGTH('a)\ \LENGTH('a) - LENGTH('b)\ \LENGTH('b)\] min_absorb2) done lemma word_split_bl: "std = size c - size b \ (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c))) \ word_split c = (a, b)" apply (rule iffI) defer apply (erule (1) word_split_bl') apply (case_tac "word_split c") apply (auto simp add: word_size) apply (frule word_split_bl' [rotated]) apply (auto simp add: word_size) done lemma word_split_bl_eq: "(word_split c :: ('c::len word \ 'd::len word)) = (of_bl (take (LENGTH('a::len) - LENGTH('d::len)) (to_bl c)), of_bl (drop (LENGTH('a) - LENGTH('d)) (to_bl c)))" for c :: "'a::len word" apply (rule word_split_bl [THEN iffD1]) apply (unfold word_size) apply (rule refl conjI)+ done lemma word_rcat_bl: \word_rcat wl = of_bl (concat (map to_bl wl))\ proof - define ws where \ws = rev wl\ moreover have \word_rcat (rev ws) = of_bl (concat (map to_bl (rev ws)))\ apply (simp add: word_rcat_def of_bl_eq rev_concat rev_map comp_def rev_to_bl_eq flip: horner_sum_of_bool_2_concat) apply transfer apply simp done ultimately show ?thesis by simp qed lemma size_rcat_lem': "size (concat (map to_bl wl)) = length wl * size (hd wl)" by (induct wl) (auto simp: word_size) lemmas size_rcat_lem = size_rcat_lem' [unfolded word_size] lemma nth_rcat_lem: "n < length (wl::'a word list) * LENGTH('a::len) \ rev (concat (map to_bl wl)) ! n = rev (to_bl (rev wl ! (n div LENGTH('a)))) ! (n mod LENGTH('a))" apply (induct wl) apply clarsimp apply (clarsimp simp add : nth_append size_rcat_lem) apply (simp flip: mult_Suc minus_div_mult_eq_mod add: less_Suc_eq_le not_less) apply (metis (no_types, lifting) diff_is_0_eq div_le_mono len_not_eq_0 less_Suc_eq less_mult_imp_div_less nonzero_mult_div_cancel_right not_le nth_Cons_0) done lemma foldl_eq_foldr: "foldl (+) x xs = foldr (+) (x # xs) 0" for x :: "'a::comm_monoid_add" by (induct xs arbitrary: x) (auto simp: add.assoc) lemmas word_cat_bl_no_bin [simp] = word_cat_bl [where a="numeral a" and b="numeral b", unfolded to_bl_numeral] for a b (* FIXME: negative numerals, 0 and 1 *) lemmas word_split_bl_no_bin [simp] = word_split_bl_eq [where c="numeral c", unfolded to_bl_numeral] for c lemmas word_rot_defs = word_roti_eq_word_rotr_word_rotl word_rotr_eq word_rotl_eq lemma to_bl_rotl: "to_bl (word_rotl n w) = rotate n (to_bl w)" by (simp add: word_rotl_eq to_bl_use_of_bl) lemmas blrs0 = rotate_eqs [THEN to_bl_rotl [THEN trans]] lemmas word_rotl_eqs = blrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotl [symmetric]] lemma to_bl_rotr: "to_bl (word_rotr n w) = rotater n (to_bl w)" by (simp add: word_rotr_eq to_bl_use_of_bl) lemmas brrs0 = rotater_eqs [THEN to_bl_rotr [THEN trans]] lemmas word_rotr_eqs = brrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotr [symmetric]] declare word_rotr_eqs (1) [simp] declare word_rotl_eqs (1) [simp] lemmas abl_cong = arg_cong [where f = "of_bl"] locale word_rotate begin lemmas word_rot_defs' = to_bl_rotl to_bl_rotr lemmas blwl_syms [symmetric] = bl_word_not bl_word_and bl_word_or bl_word_xor lemmas lbl_lbl = trans [OF word_bl_Rep' word_bl_Rep' [symmetric]] lemmas ths_map2 [OF lbl_lbl] = rotate_map2 rotater_map2 lemmas ths_map [where xs = "to_bl v"] = rotate_map rotater_map for v lemmas th1s [simplified word_rot_defs' [symmetric]] = ths_map2 ths_map end lemmas bl_word_rotl_dt = trans [OF to_bl_rotl rotate_drop_take, simplified word_bl_Rep'] lemmas bl_word_rotr_dt = trans [OF to_bl_rotr rotater_drop_take, simplified word_bl_Rep'] lemma bl_word_roti_dt': "n = nat ((- i) mod int (size (w :: 'a::len word))) \ to_bl (word_roti i w) = drop n (to_bl w) @ take n (to_bl w)" apply (unfold word_roti_eq_word_rotr_word_rotl) apply (simp add: bl_word_rotl_dt bl_word_rotr_dt word_size) apply safe apply (simp add: zmod_zminus1_eq_if) apply safe apply (simp add: nat_mult_distrib) apply (simp add: nat_diff_distrib [OF pos_mod_sign pos_mod_conj [THEN conjunct2, THEN order_less_imp_le]] nat_mod_distrib) apply (simp add: nat_mod_distrib) done lemmas bl_word_roti_dt = bl_word_roti_dt' [unfolded word_size] lemmas word_rotl_dt = bl_word_rotl_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_rotr_dt = bl_word_rotr_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_roti_dt = bl_word_roti_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_rotr_dt_no_bin' [simp] = word_rotr_dt [where w="numeral w", unfolded to_bl_numeral] for w (* FIXME: negative numerals, 0 and 1 *) lemmas word_rotl_dt_no_bin' [simp] = word_rotl_dt [where w="numeral w", unfolded to_bl_numeral] for w (* FIXME: negative numerals, 0 and 1 *) lemma max_word_bl: "to_bl (max_word::'a::len word) = replicate LENGTH('a) True" by (fact to_bl_n1) lemma to_bl_mask: "to_bl (mask n :: 'a::len word) = replicate (LENGTH('a) - n) False @ replicate (min (LENGTH('a)) n) True" by (simp add: mask_bl word_rep_drop min_def) lemma map_replicate_True: "n = length xs \ map (\(x,y). x \ y) (zip xs (replicate n True)) = xs" by (induct xs arbitrary: n) auto lemma map_replicate_False: "n = length xs \ map (\(x,y). x \ y) (zip xs (replicate n False)) = replicate n False" by (induct xs arbitrary: n) auto lemma bl_and_mask: fixes w :: "'a::len word" and n :: nat defines "n' \ LENGTH('a) - n" shows "to_bl (w AND mask n) = replicate n' False @ drop n' (to_bl w)" proof - note [simp] = map_replicate_True map_replicate_False have "to_bl (w AND mask n) = map2 (\) (to_bl w) (to_bl (mask n::'a::len word))" by (simp add: bl_word_and) also have "to_bl w = take n' (to_bl w) @ drop n' (to_bl w)" by simp also have "map2 (\) \ (to_bl (mask n::'a::len word)) = replicate n' False @ drop n' (to_bl w)" unfolding to_bl_mask n'_def by (subst zip_append) auto finally show ?thesis . qed lemma drop_rev_takefill: "length xs \ n \ drop (n - length xs) (rev (takefill False n (rev xs))) = xs" by (simp add: takefill_alt rev_take) declare bin_to_bl_def [simp] lemmas of_bl_reasoning = to_bl_use_of_bl of_bl_append lemma uint_of_bl_is_bl_to_bin_drop: "length (dropWhile Not l) \ LENGTH('a) \ uint (of_bl l :: 'a::len word) = bl_to_bin l" apply transfer apply (simp add: take_bit_eq_mod) apply (rule Divides.mod_less) apply (rule bl_to_bin_ge0) using bl_to_bin_lt2p_drop apply (rule order.strict_trans2) apply simp done corollary uint_of_bl_is_bl_to_bin: "length l\LENGTH('a) \ uint ((of_bl::bool list\ ('a :: len) word) l) = bl_to_bin l" apply(rule uint_of_bl_is_bl_to_bin_drop) using le_trans length_dropWhile_le by blast lemma bin_to_bl_or: "bin_to_bl n (a OR b) = map2 (\) (bin_to_bl n a) (bin_to_bl n b)" using bl_or_aux_bin[where n=n and v=a and w=b and bs="[]" and cs="[]"] by simp lemma word_and_1_bl: fixes x::"'a::len word" shows "(x AND 1) = of_bl [x !! 0]" by (simp add: mod_2_eq_odd test_bit_word_eq and_one_eq) lemma word_1_and_bl: fixes x::"'a::len word" shows "(1 AND x) = of_bl [x !! 0]" by (simp add: mod_2_eq_odd test_bit_word_eq one_and_eq) lemma of_bl_drop: "of_bl (drop n xs) = (of_bl xs AND mask (length xs - n))" apply (clarsimp simp: bang_eq test_bit_of_bl rev_nth cong: rev_conj_cong) apply (safe; simp add: word_size to_bl_nth) done lemma to_bl_1: "to_bl (1::'a::len word) = replicate (LENGTH('a) - 1) False @ [True]" by (rule nth_equalityI) (auto simp add: to_bl_unfold nth_append rev_nth bit_1_iff not_less not_le) lemma eq_zero_set_bl: "(w = 0) = (True \ set (to_bl w))" using list_of_false word_bl.Rep_inject by fastforce lemma of_drop_to_bl: "of_bl (drop n (to_bl x)) = (x AND mask (size x - n))" by (simp add: of_bl_drop word_size_bl) lemma unat_of_bl_length: "unat (of_bl xs :: 'a::len word) < 2 ^ (length xs)" proof (cases "length xs < LENGTH('a)") case True then have "(of_bl xs::'a::len word) < 2 ^ length xs" by (simp add: of_bl_length_less) with True show ?thesis by (simp add: word_less_nat_alt word_unat_power unat_of_nat del: of_nat_power) next case False have "unat (of_bl xs::'a::len word) < 2 ^ LENGTH('a)" by (simp split: unat_split) also from False have "LENGTH('a) \ length xs" by simp then have "2 ^ LENGTH('a) \ (2::nat) ^ length xs" by (rule power_increasing) simp finally show ?thesis . qed lemma word_msb_alt: "msb w \ hd (to_bl w)" for w :: "'a::len word" apply (simp add: msb_word_eq) apply (subst hd_conv_nth) apply simp apply (subst nth_to_bl) apply simp apply simp done lemma word_lsb_last: \lsb w \ last (to_bl w)\ for w :: \'a::len word\ using nth_to_bl [of \LENGTH('a) - Suc 0\ w] by (simp add: lsb_odd last_conv_nth) +lemma is_aligned_to_bl: + "is_aligned (w :: 'a :: len word) n = (True \ set (drop (size w - n) (to_bl w)))" + apply (simp add: is_aligned_mask eq_zero_set_bl) + apply (clarsimp simp: in_set_conv_nth word_size) + apply (simp add: to_bl_nth word_size cong: conj_cong) + apply (simp add: diff_diff_less) + apply safe + apply (case_tac "n \ LENGTH('a)") + prefer 2 + apply (rule_tac x=i in exI) + apply clarsimp + apply (subgoal_tac "\j < LENGTH('a). j < n \ LENGTH('a) - n + j = i") + apply (erule exE) + apply (rule_tac x=j in exI) + apply clarsimp + apply (thin_tac "w !! t" for t) + apply (rule_tac x="i + n - LENGTH('a)" in exI) + apply clarsimp + apply arith + apply (rule_tac x="LENGTH('a) - n + i" in exI) + apply clarsimp + apply arith + done + +lemma is_aligned_replicate: + fixes w::"'a::len word" + assumes aligned: "is_aligned w n" + and nv: "n \ LENGTH('a)" + shows "to_bl w = (take (LENGTH('a) - n) (to_bl w)) @ replicate n False" +proof - + from nv have rl: "\q. q < 2 ^ (LENGTH('a) - n) \ + to_bl (2 ^ n * (of_nat q :: 'a word)) = + drop n (to_bl (of_nat q :: 'a word)) @ replicate n False" + by (metis bl_shiftl le_antisym min_def shiftl_t2n wsst_TYs(3)) + show ?thesis using aligned + by (auto simp: rl elim: is_alignedE) +qed + +lemma is_aligned_drop: + fixes w::"'a::len word" + assumes "is_aligned w n" "n \ LENGTH('a)" + shows "drop (LENGTH('a) - n) (to_bl w) = replicate n False" +proof - + have "to_bl w = take (LENGTH('a) - n) (to_bl w) @ replicate n False" + by (rule is_aligned_replicate) fact+ + then have "drop (LENGTH('a) - n) (to_bl w) = drop (LENGTH('a) - n) \" by simp + also have "\ = replicate n False" by simp + finally show ?thesis . +qed + +lemma less_is_drop_replicate: + fixes x::"'a::len word" + assumes lt: "x < 2 ^ n" + shows "to_bl x = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl x)" + by (metis assms bl_and_mask' less_mask_eq) + +lemma is_aligned_add_conv: + fixes off::"'a::len word" + assumes aligned: "is_aligned w n" + and offv: "off < 2 ^ n" + shows "to_bl (w + off) = + (take (LENGTH('a) - n) (to_bl w)) @ (drop (LENGTH('a) - n) (to_bl off))" +proof cases + assume nv: "n \ LENGTH('a)" + show ?thesis + proof (subst aligned_bl_add_size, simp_all only: word_size) + show "drop (LENGTH('a) - n) (to_bl w) = replicate n False" + by (subst is_aligned_replicate [OF aligned nv]) (simp add: word_size) + + from offv show "take (LENGTH('a) - n) (to_bl off) = + replicate (LENGTH('a) - n) False" + by (subst less_is_drop_replicate, assumption) simp + qed fact +next + assume "\ n \ LENGTH('a)" + with offv show ?thesis by (simp add: power_overflow) +qed + +lemma is_aligned_replicateI: + "to_bl p = addr @ replicate n False \ is_aligned (p::'a::len word) n" + apply (simp add: is_aligned_to_bl word_size) + apply (subgoal_tac "length addr = LENGTH('a) - n") + apply (simp add: replicate_not_True) + apply (drule arg_cong [where f=length]) + apply simp + done + +lemma to_bl_2p: + "n < LENGTH('a) \ + to_bl ((2::'a::len word) ^ n) = + replicate (LENGTH('a) - Suc n) False @ True # replicate n False" + apply (subst shiftl_1 [symmetric]) + apply (subst bl_shiftl) + apply (simp add: to_bl_1 min_def word_size) + done + +lemma xor_2p_to_bl: + fixes x::"'a::len word" + shows "to_bl (x XOR 2^n) = + (if n < LENGTH('a) + then take (LENGTH('a)-Suc n) (to_bl x) @ (\rev (to_bl x)!n) # drop (LENGTH('a)-n) (to_bl x) + else to_bl x)" +proof - + have x: "to_bl x = take (LENGTH('a)-Suc n) (to_bl x) @ drop (LENGTH('a)-Suc n) (to_bl x)" + by simp + + show ?thesis + apply simp + apply (rule conjI) + apply (clarsimp simp: word_size) + apply (simp add: bl_word_xor to_bl_2p) + apply (subst x) + apply (subst zip_append) + apply simp + apply (simp add: map_zip_replicate_False_xor drop_minus) + apply (auto simp add: word_size nth_w2p intro!: word_eqI) + done +qed + +lemma is_aligned_replicateD: + "\ is_aligned (w::'a::len word) n; n \ LENGTH('a) \ + \ \xs. to_bl w = xs @ replicate n False + \ length xs = size w - n" + apply (subst is_aligned_replicate, assumption+) + apply (rule exI, rule conjI, rule refl) + apply (simp add: word_size) + done + text \right-padding a word to a certain length\ definition "bl_pad_to bl sz \ bl @ (replicate (sz - length bl) False)" lemma bl_pad_to_length: assumes lbl: "length bl \ sz" shows "length (bl_pad_to bl sz) = sz" using lbl by (simp add: bl_pad_to_def) lemma bl_pad_to_prefix: "prefix bl (bl_pad_to bl sz)" by (simp add: bl_pad_to_def) end