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,1652 @@ (* 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, 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, 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) + apply (auto simp add: le_less) + apply (metis linorder_neqE_linordered_idom mod_pos_pos_trivial not_take_bit_negative power_0 take_bit_0 take_bit_eq_mod take_bit_nonnegative) + apply (metis even_take_bit_eq mod_pos_pos_trivial neq0_conv numeral_eq_Suc power_0 take_bit_eq_mod take_bit_nonnegative zero_less_Suc) 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/Markov_Models/Continuous_Time_Markov_Chain.thy b/thys/Markov_Models/Continuous_Time_Markov_Chain.thy --- a/thys/Markov_Models/Continuous_Time_Markov_Chain.thy +++ b/thys/Markov_Models/Continuous_Time_Markov_Chain.thy @@ -1,1464 +1,1464 @@ (* Author: Johannes Hölzl *) section \Continuous-time Markov chains\ theory Continuous_Time_Markov_Chain imports Discrete_Time_Markov_Process Discrete_Time_Markov_Chain begin subsection \Trace Operations: relate @{typ "('a \ real) stream"} and @{typ "real \ 'a"}\ partial_function (tailrec) trace_at :: "'a \ (real \ 'a) stream \ real \ 'a" where "trace_at s \ j = (case \ of (t', s')##\ \ if t' \ j then trace_at s' \ j else s)" lemma trace_at_simp[simp]: "trace_at s ((t', s')##\) j = (if t' \ j then trace_at s' \ j else s)" by (subst trace_at.simps) simp lemma trace_at_eq: "trace_at s \ j = (case sfirst (\x. j < fst (shd x)) \ of \ \ undefined | enat i \ (s ## smap snd \) !! i)" proof (split enat.split; safe) assume "sfirst (\x. j < fst (shd x)) \ = \" with sfirst_finite[of "\x. j < fst (shd x)" \] have "alw (\x. fst (shd x) \ j) \" by (simp add: not_ev_iff not_less) then show "trace_at s \ j = undefined" by (induction arbitrary: s \ rule: trace_at.fixp_induct) (auto split: stream.split) next show "sfirst (\x. j < fst (shd x)) \ = enat n \ trace_at s \ j = (s ## smap snd \) !! n" for n proof (induction n arbitrary: s \) case 0 then show ?case by (subst trace_at.simps) (auto simp add: enat_0 sfirst_eq_0 split: stream.split) next case (Suc n) show ?case using sfirst.simps[of "\x. j < fst (shd x)" \] Suc.prems Suc.IH[of "stl \" "snd (shd \)"] by (cases \) (auto simp add: eSuc_enat[symmetric] split: stream.split if_split_asm) qed qed lemma trace_at_shift: "trace_at s (smap (\(t, s'). (t + t', s')) \) t = trace_at s \ (t - t')" by (induction arbitrary: s \ rule: trace_at.fixp_induct) (auto split: stream.split) primcorec merge_at :: "(real \ 'a) stream \ real \ (real \ 'a) stream \ (real \ 'a) stream" where "merge_at \ j \' = (case \ of (t, s) ## \ \ if t \ j then (t, s)##merge_at \ j \' else \')" lemma merge_at_simp[simp]: "merge_at (x##\) j \' = (if fst x \ j then x##merge_at \ j \' else \')" by (cases x) (subst merge_at.code; simp) subsection \Exponential Distribution\ definition exponential :: "real \ real measure" where "exponential l = density lborel (exponential_density l)" lemma space_exponential: "space (exponential l) = UNIV" by (simp add: exponential_def) lemma sets_exponential[measurable_cong]: "sets (exponential l) = sets borel" by (simp add: exponential_def) lemma prob_space_exponential: "0 < l \ prob_space (exponential l)" unfolding exponential_def by (intro prob_space_exponential_density) lemma AE_exponential: "0 < l \ AE x in exponential l. 0 < x" unfolding exponential_def using AE_lborel_singleton[of 0] by (auto simp add: AE_density exponential_density_def) lemma emeasure_exponential_Ioi_cutoff: assumes "0 < l" shows "emeasure (exponential l) {x <..} = exp (- (max 0 x) * l)" proof - interpret prob_space "exponential l" unfolding exponential_def using \0 by (rule prob_space_exponential_density) have *: "prob {xa \ space (exponential l). max 0 x < xa} = exp (- max 0 x * l)" apply (rule exponential_distributedD_gt[OF _ _ \0]) apply (auto simp: exponential_def distributed_def) apply (subst (6) distr_id[symmetric]) apply (subst (2) distr_cong) apply simp_all done have "emeasure (exponential l) {x <..} = emeasure (exponential l) {max 0 x <..}" using AE_exponential[OF \0] by (intro emeasure_eq_AE) auto also have "\ = exp (- (max 0 x) * l)" using * unfolding emeasure_eq_measure by (simp add: space_exponential greaterThan_def) finally show ?thesis . qed lemma emeasure_exponential_Ioi: "0 < l \ 0 \ x \ emeasure (exponential l) {x <..} = exp (- x * l)" using emeasure_exponential_Ioi_cutoff[of l x] by simp lemma exponential_eq_stretch: assumes "0 < l" shows "exponential l = distr (exponential 1) borel (\x. (1/l) * x)" proof (intro measure_eqI) fix A assume "A \ sets (exponential l)" then have [measurable]: "A \ sets borel" by (simp add: sets_exponential) then have [measurable]: "(\x. x / l) -` A \ sets borel" by (rule measurable_sets_borel[rotated]) simp have "emeasure (exponential l) A = (\\<^sup>+x. ennreal l * (indicator (((*) (1/l) -` A) \ {0 ..}) (l * x) * ennreal (exp (- (l * x)))) \lborel)" using \0 < l\ by (auto simp: ac_simps emeasure_distr exponential_def emeasure_density exponential_density_def ennreal_mult zero_le_mult_iff intro!: nn_integral_cong split: split_indicator) also have "\ = (\\<^sup>+x. indicator (((*) (1/l) -` A) \ {0 ..}) x * ennreal (exp (- x)) \lborel)" using \0 apply (subst nn_integral_stretch) apply (auto simp: nn_integral_cmult) apply (simp add: ennreal_mult[symmetric] mult.assoc[symmetric]) done also have "\ = emeasure (distr (exponential 1) borel (\x. (1/l) * x)) A" by (auto simp add: emeasure_distr exponential_def emeasure_density exponential_density_def intro!: nn_integral_cong split: split_indicator) finally show "emeasure (exponential l) A = emeasure (distr (exponential 1) borel (\x. (1/l) * x)) A" . qed (simp add: sets_exponential) lemma uniform_measure_exponential: assumes "0 < l" "0 \ t" shows "uniform_measure (exponential l) {t <..} = distr (exponential l) borel ((+) t)" (is "?L = ?R") proof (rule measure_eqI_lessThan) fix x have "0 < emeasure (exponential l) {t<..}" unfolding emeasure_exponential_Ioi[OF assms] by simp with assms show "?L {x<..} < \" by (simp add: ennreal_divide_eq_top_iff less_top[symmetric] lessThan_Int_lessThan emeasure_exponential_Ioi) have *: "((+) t -` {x<..} \ space (exponential l)) = {x - t <..}" by (auto simp: space_exponential) show "?L {x<..} = ?R {x<..}" using assms by (simp add: lessThan_Int_lessThan emeasure_exponential_Ioi divide_ennreal emeasure_distr * emeasure_exponential_Ioi_cutoff exp_diff[symmetric] field_simps split: split_max) qed (auto simp: sets_exponential) lemma emeasure_PiM_exponential_Ioi_finite: assumes "J \ I" "finite J" "\i. i \ I \ 0 < R i" "0 \ x" shows "emeasure (\\<^sub>M i\I. exponential (R i)) (prod_emb I (\i. exponential (R i)) J (\\<^sub>E j\J. {x<..})) = exp (- x * (\i\J. R i))" proof (subst emeasure_PiM_emb) from assms show "(\i\J. emeasure (exponential (R i)) {x<..}) = ennreal (exp (- x * sum R J))" by (subst prod.cong[OF refl emeasure_exponential_Ioi]) (auto simp add: prod_ennreal exp_sum sum_negf[symmetric] sum_distrib_left) qed (insert assms, auto intro!: prob_space_exponential) lemma emeasure_PiM_exponential_Ioi_sequence: assumes R: "summable R" "\i. 0 < R i" "0 \ x" shows "emeasure (\\<^sub>M i\UNIV. exponential (R i)) (\ i\UNIV. {x<..}) = exp (- x * suminf R)" proof - let ?R = "\i. exponential (R i)" let ?P = "\\<^sub>M i\UNIV. ?R i" let ?N = "\n::nat. prod_emb UNIV ?R {..\<^sub>E i\{..\<^sub>M i\UNIV. exponential (R i)) (\n. ?N n) = (INF n. (\\<^sub>M i\UNIV. exponential (R i)) (?N n))" by (intro INF_emeasure_decseq[symmetric] decseq_emb_PiE) (auto simp: incseq_def) also have "\ = (INF n. ennreal (exp (- x * (\i = ennreal (exp (- x * (SUP n. (\i = ennreal (exp (- x * (\i. R i)))" using R by (subst suminf_eq_SUP_real) (auto simp: less_imp_le) also have "(\n. ?N n) = (\ i\UNIV. {x<..})" by (fastforce simp: prod_emb_def Pi_iff PiE_iff space_exponential) finally show ?thesis using R by simp qed lemma emeasure_PiM_exponential_Ioi_countable: assumes R: "J \ I" "countable J" "\i. i \ I \ 0 < R i" "0 \ x" and finite: "integrable (count_space J) R" shows "emeasure (\\<^sub>M i\I. exponential (R i)) (prod_emb I (\i. exponential (R i)) J (\\<^sub>E j\J. {x<..})) = exp (- x * (LINT i|count_space J. R i))" proof cases assume "finite J" with assms show ?thesis by (subst emeasure_PiM_exponential_Ioi_finite) (auto simp: lebesgue_integral_count_space_finite) next assume "infinite J" let ?R = "\i. exponential (R i)" let ?P = "\\<^sub>M i\I. ?R i" define f where "f = from_nat_into J" have J_eq: "J = range f" and f: "inj f" "f \ UNIV \ I" using from_nat_into_inj_infinite[of J] range_from_nat_into[of J] \countable J\ \infinite J\ \J \ I\ by (auto simp: inj_on_def f_def simp del: range_from_nat_into) have Bf: "bij_betw f UNIV J" unfolding J_eq using inj_on_imp_bij_betw[OF f(1)] . have summable_R: "summable (\i. R (f i))" using finite unfolding integrable_bij_count_space[OF Bf, symmetric] integrable_count_space_nat_iff by (rule summable_norm_cancel) have "emeasure (\\<^sub>M i\UNIV. exponential (R (f i))) (\ i\UNIV. {x<..}) = exp (- x * (\i. R (f i)))" using finite assms unfolding J_eq by (intro emeasure_PiM_exponential_Ioi_sequence[OF summable_R]) auto also have "(\\<^sub>M i\UNIV. exponential (R (f i))) = distr ?P (\\<^sub>M i\UNIV. exponential (R (f i))) (\\. \i\UNIV. \ (f i))" using R by (intro distr_PiM_reindex[symmetric, OF _ f] prob_space_exponential) auto also have "\ (\ i\UNIV. {x<..}) = ?P ((\\. \i\UNIV. \ (f i)) -` (\ i\UNIV. {x<..}) \ space ?P)" using f(2) by (intro emeasure_distr infprod_in_sets) (auto simp: Pi_iff) also have "(\\. \i\UNIV. \ (f i)) -` (\ i\UNIV. {x<..}) \ space ?P = prod_emb I ?R J (\\<^sub>E j\J. {x<..})" by (auto simp: prod_emb_def space_PiM space_exponential Pi_iff J_eq) also have "(\i. R (f i)) = (LINT i|count_space J. R i)" using finite by (subst integral_count_space_nat[symmetric]) (auto simp: integrable_bij_count_space[OF Bf] integral_bij_count_space[OF Bf]) finally show ?thesis . qed lemma AE_PiM_exponential_suminf_infty: fixes R :: "nat \ real" assumes R: "\n. 0 < R n" and finite: "(\n. ennreal (1 / R n)) = top" shows "AE \ in \\<^sub>M n\UNIV. exponential (R n). (\n. ereal (\ n)) = \" proof - let ?P = "\\<^sub>M n\UNIV. exponential (R n)" interpret prob_space "exponential (R n)" for n by (intro prob_space_exponential R) interpret product_prob_space "\n. exponential (R n)" UNIV proof qed have AE_pos: "AE \ in ?P. \i. 0 < \ i" unfolding AE_all_countable by (intro AE_PiM_component allI prob_space_exponential R AE_exponential) simp have indep: "indep_vars (\i. borel) (\i x. x i) UNIV" using PiM_component apply (subst P.indep_vars_iff_distr_eq_PiM) apply (auto simp: restrict_UNIV distr_id2) apply (subst distr_id2) apply (intro sets_PiM_cong) apply (auto simp: sets_exponential cong: distr_cong) done have [simp]: "0 \ x + x * R i \ 0 \ x" for x i using zero_le_mult_iff[of x "1 + R i"] R[of i] by (simp add: field_simps) have "(\\<^sup>+\. eexp (\n. - ereal (\ n)) \?P) = (\\<^sup>+\. (INF n. \i i))) \?P)" proof (intro nn_integral_cong_AE, use AE_pos in eventually_elim) fix \ :: "nat \ real" assume \: "\i. 0 < \ i" show "eexp (\n. - ereal (\ n)) = (\n. \i i)))" proof (rule LIMSEQ_unique[OF _ LIMSEQ_INF]) show "(\i. \i i))) \ eexp (\n. - ereal (\ n))" using \ by (intro eexp_suminf summable_minus_ereal summable_ereal_pos) (auto intro: less_imp_le) show "decseq (\n. \i i)))" using \ by (auto simp: decseq_def intro!: prod_mono3 intro: less_imp_le) qed qed also have "\ = (INF n. (\\<^sup>+\. (\i i))) \?P))" proof (intro nn_integral_monotone_convergence_INF_AE') show "AE \ in ?P. (\i i))) \ (\i i)))" for n using AE_pos proof eventually_elim case (elim \) show ?case by (rule prod_mono3) (auto simp: elim le_less) qed qed (auto simp: less_top[symmetric]) also have "\ = (INF n. (\i\<^sup>+\. eexp (- ereal (\ i)) \?P)))" proof (intro INF_cong refl indep_vars_nn_integral) show "indep_vars (\_. borel) (\i \. eexp (- ereal (\ i))) {..i x. eexp(- ereal x)"]) show "indep_vars (\i. borel) (\i x. x i) {.. = (INF n. (\i\<^sup>+x. indicator {0 ..} ((1 + R i) * x) * ennreal (exp (- ((1 + R i) * x))) \lborel)))" by (subst product_nn_integral_component) (auto simp: field_simps exponential_def nn_integral_density ennreal_mult'[symmetric] ennreal_mult''[symmetric] exponential_density_def exp_diff exp_minus nn_integral_cmult[symmetric] intro!: INF_cong prod.cong nn_integral_cong split: split_indicator) also have "\ = (INF n. (\i\<^sup>+ x. indicator {0..} ((1 + R i) * x) * ennreal (exp (- ((1 + R i) * x))) \lborel) = ennreal (R i / (1 + R i))" for i using nn_intergal_power_times_exp_Ici[of 0] \0 < R i\ by (subst nn_integral_stretch[where c="1 + R i"]) (auto simp: mult.assoc[symmetric] ennreal_mult''[symmetric] less_imp_le mult.commute) qed also have "\ = (INF n. ennreal (\i = (INF n. ennreal (inverse (\i = (INF n. inverse (ennreal (\i = inverse (SUP n. ennreal (\iiin. ennreal (\i r" using R unfolding real(2)[symmetric] by (intro LIMSEQ_SUP monoI ennreal_leI prod_mono2) (auto intro!: divide_nonneg_nonneg add_nonneg_nonneg intro: less_imp_le) then have "(\n. (\i r" by (rule tendsto_ennrealD) (use R real in \auto intro!: always_eventually prod_nonneg divide_nonneg_nonneg add_nonneg_nonneg intro: less_imp_le\) moreover have "(1 + R i) / R i = 1 / R i + 1" for i using \0 < R i\ by (auto simp: field_simps) ultimately have "convergent (\n. \ii. 1 / R i)" using R by (subst summable_iff_convergent_prod) (auto intro: less_imp_le) moreover have "0 \ 1 / R i" for i using R by (auto simp: less_imp_le) ultimately show ?thesis using finite ennreal_suminf_neq_top[of "\i. 1 / R i"] by blast qed finally have "(\\<^sup>+\. eexp (\n. - ereal (\ n)) \?P) = 0" by simp then have "AE \ in ?P. eexp (\n. - ereal (\ n)) = 0" by (subst (asm) nn_integral_0_iff_AE) auto then show ?thesis using AE_pos proof eventually_elim show "(\i. 0 < \ i) \ eexp (\n. - ereal (\ n)) = 0 \ (\n. ereal (\ n)) = \" for \ apply (auto simp del: uminus_ereal.simps simp add: uminus_ereal.simps[symmetric] intro!: summable_iff_suminf_neq_top intro: less_imp_le) apply (subst (asm) suminf_minus_ereal) apply (auto intro!: summable_ereal_pos intro: less_imp_le) done qed qed subsection \Transition Rates\ locale transition_rates = fixes R :: "'a \ 'a \ real" assumes R_nonneg[simp]: "\x y. 0 \ R x y" assumes R_diagonal_0[simp]: "\x. R x x = 0" assumes finite_weight: "\x. (\\<^sup>+y. R x y \count_space UNIV) < \" assumes positive_weight: "\x. 0 < (\\<^sup>+y. R x y \count_space UNIV)" begin abbreviation S :: "(real \ 'a) measure" where "S \ (borel \\<^sub>M count_space UNIV)" abbreviation T :: "(real \ 'a) stream measure" where "T \ stream_space S" abbreviation I :: "'a \ 'a set" where "I x \ {y. 0 < R x y}" lemma I_countable: "countable (I x)" proof - let ?P = "point_measure UNIV (R x)" interpret finite_measure ?P proof show "emeasure ?P (space ?P) \ \" using finite_weight by (simp add: emeasure_density point_measure_def less_top) qed from countable_support emeasure_point_measure_finite2[of "{_}" UNIV "R x"] show ?thesis by (simp add: emeasure_eq_measure less_le) qed definition escape_rate :: "'a \ real" where "escape_rate x = \y. R x y \count_space UNIV" lemma ennreal_escape_rate: "ennreal (escape_rate x) = (\\<^sup>+y. R x y \count_space UNIV)" using finite_weight[of x] unfolding escape_rate_def by (intro nn_integral_eq_integral[symmetric]) (auto simp: integrable_iff_bounded) lemma escape_rate_pos: "0 < escape_rate x" using positive_weight unfolding ennreal_escape_rate[symmetric] by simp lemma nonneg_escape_rate[simp]: "0 \ escape_rate x" using escape_rate_pos[THEN less_imp_le] . lemma prob_space_exponential_escape_rate: "prob_space (exponential (escape_rate x))" using escape_rate_pos by (rule prob_space_exponential) lemma measurable_escape_rate[measurable]: "escape_rate \ count_space UNIV \\<^sub>M borel" by auto lemma measurable_exponential_escape_rate[measurable]: "(\x. exponential (escape_rate x)) \ count_space UNIV \\<^sub>M prob_algebra borel" by (auto simp: space_prob_algebra sets_exponential prob_space_exponential_escape_rate) interpretation pmf_as_function . lift_definition J :: "'a \ 'a pmf" is "\x y. R x y / escape_rate x" proof safe show "0 \ R x y / escape_rate x" for x y by (auto intro!: integral_nonneg_AE divide_nonneg_nonneg R_nonneg simp: escape_rate_def) show "(\\<^sup>+y. R x y / escape_rate x \count_space UNIV) = 1" for x using escape_rate_pos[of x] by (auto simp add: divide_ennreal[symmetric] nn_integral_divide ennreal_escape_rate[symmetric] intro!: ennreal_divide_self) qed lemma set_pmf_J: "set_pmf (J x) = I x" using escape_rate_pos[of x] by (auto simp: set_pmf_iff J.rep_eq less_le) interpretation exp_esc: pair_prob_space "distr (exponential (escape_rate x)) borel ((+) t)" "J x" for x proof - interpret prob_space "distr (exponential (escape_rate x)) borel ((+) t)" by (intro prob_space.prob_space_distr prob_space_exponential_escape_rate) simp show "pair_prob_space (distr (exponential (escape_rate x)) borel ((+) t)) (measure_pmf (J x))" by standard qed subsection \Continuous-time Kernel\ definition K :: "(real \ 'a) \ (real \ 'a) measure" where "K = (\(t, x). (distr (exponential (escape_rate x)) borel ((+) t)) \\<^sub>M J x)" interpretation K: discrete_Markov_process "borel \\<^sub>M count_space UNIV" K proof show "K \ borel \\<^sub>M count_space UNIV \\<^sub>M prob_algebra (borel \\<^sub>M count_space UNIV)" unfolding K_def apply measurable apply (rule measurable_snd[THEN measurable_compose]) apply (auto simp: space_prob_algebra prob_space_measure_pmf) done qed interpretation DTMC: MC_syntax J . lemma in_space_S[simp]: "x \ space S" by (simp add: space_pair_measure) lemma in_space_T[simp]: "x \ space T" by (simp add: space_pair_measure space_stream_space) lemma in_space_lim_stream: "\ \ space (K.lim_stream x)" unfolding K.space_lim_stream space_stream_space[symmetric] by simp lemma prob_space_K_lim: "prob_space (K.lim_stream x)" using K.lim_stream[THEN measurable_space] by (simp add: space_prob_algebra) definition select_first :: "'a \ ('a \ real) \ 'a \ bool" where "select_first x p y = (y \ I x \ (\y'\I x - {y}. p y < p y'))" lemma select_firstD1: "select_first x p y \ y \ I x" by (simp add: select_first_def) lemma select_first_unique: assumes y: "select_first x p y1" " select_first x p y2" shows "y1 = y2" proof - have "y1 \ y2 \ p y1 < p y2" "y1 \ y2 \ p y2 < p y1" using y by (auto simp: select_first_def) then show "y1 = y2" by (rule_tac ccontr) auto qed lemma The_select_first[simp]: "select_first x p y \ The (select_first x p) = y" by (intro the_equality select_first_unique) lemma select_first_INF: "select_first x p y \ (INF x\I x. p x) = p y" by (intro antisym cINF_greatest cINF_lower bdd_belowI2[where m="p y"]) (auto simp: select_first_def le_less) lemma measurable_select_first[measurable]: "(\p. select_first x p y) \ (\\<^sub>M y\I x. borel) \\<^sub>M count_space UNIV" using I_countable unfolding select_first_def by (intro measurable_pred_countable pred_intros_conj1') measurable lemma measurable_THE_select_first[measurable]: "(\p. The (select_first x p)) \ (\\<^sub>M y\I x. borel) \\<^sub>M count_space UNIV" by (rule measurable_THE) (auto intro: select_first_unique I_countable dest: select_firstD1) lemma sets_S_eq: "sets S = sigma_sets UNIV { {t ..} \ A | t A. A \ - I x \ (\s\I x. A = {s}) }" proof (subst sets_pair_eq) let ?CI = "\a::real. {a ..}" let ?Ea = "range ?CI" show "?Ea \ Pow (space borel)" "sets borel = sigma_sets (space borel) ?Ea" unfolding borel_Ici by auto show "?CI`Rats \ ?Ea" "(\i\Rats. ?CI i) = space borel" using Rats_dense_in_real[of "x - 1" "x" for x] by (auto intro: less_imp_le) let ?Eb = "Pow (- I x) \ (\s. {s}) ` I x" have "b \ sigma_sets UNIV (Pow (- I x) \ (\s. {s}) ` I x)" for b proof - have "b = (b - I x) \ (\x\b \ I x. {x})" by auto also have "\ \ sigma UNIV (Pow (- I x) \ (\s. {s}) ` I x)" using I_countable by (intro sets.Un sets.countable_UN') auto finally show ?thesis by simp qed then show "sets (count_space UNIV) = sigma_sets (space (count_space UNIV)) ?Eb" by auto show "countable ({- I x} \ (\s\I x. {{s}}))" using I_countable by auto show "sets (sigma (space borel \ space (count_space UNIV)) {a \ b |a b. a \ ?Ea \ b \ ?Eb}) = sigma_sets UNIV {{t ..} \ A |t A. A \ - I x \ (\s\I x. A = {s})}" apply simp apply (intro arg_cong[where f="sigma_sets _"]) apply auto done qed (auto intro: countable_rat) subsection \Kernel equals Parallel Choice\ abbreviation PAR :: "'a \ ('a \ real) measure" where "PAR x \ (\\<^sub>M y\I x. exponential (R x y))" lemma PAR_least: assumes y: "y \ I x" shows "PAR x {p\space (PAR x). t \ p y \ select_first x p y} = emeasure (exponential (escape_rate x)) {t ..} * ennreal (pmf (J x) y)" proof - let ?E = "\y. exponential (R x y)" let ?P' = "\\<^sub>M y\I x - {y}. ?E y" interpret P': prob_space ?P' by (intro prob_space_PiM prob_space_exponential) simp have *: "PAR x = (\\<^sub>M y\insert y (I x - {y}). ?E y)" using y by (intro PiM_cong) auto have "0 < R x y" using y by simp have **: "(\(x, X). X(y := x)) \ exponential (R x y) \\<^sub>M Pi\<^sub>M (I x - {y}) (\i. exponential (R x i)) \\<^sub>M PAR x" using y apply (subst measurable_cong_sets[OF sets_pair_measure_cong[OF sets_exponential sets_PiM_cong[OF refl sets_exponential]] sets_PiM_cong[OF refl sets_exponential]]) apply measurable apply (rule measurable_fun_upd[where J="I x - {y}"]) apply auto done have "PAR x {p\space (PAR x). t \ p y \ (\y'\I x-{y}. p y < p y')} = (\\<^sup>+ty. indicator {t..} ty * ?P' {p\space ?P'. \y'\I x-{y}. ty < p y'} \?E y)" unfolding * using \y \ I x\ apply (subst distr_pair_PiM_eq_PiM[symmetric]) apply (auto intro!: prob_space_exponential simp: emeasure_distr insert_absorb) apply (subst emeasure_distr[OF **]) subgoal using I_countable by (auto simp: pred_def[symmetric]) apply (subst P'.emeasure_pair_measure_alt) subgoal using I_countable[of x] apply (intro measurable_sets[OF **]) apply (auto simp: pred_def[symmetric]) done apply (auto intro!: nn_integral_cong arg_cong2[where f=emeasure] split: split_indicator if_split_asm simp: space_exponential space_PiM space_pair_measure PiE_iff extensional_def) done also have "\ = (\\<^sup>+ty. indicator {t..} ty * ennreal (exp (- ty * (escape_rate x - R x y))) \?E y)" apply (intro nn_integral_cong_AE) using AE_exponential[OF \0 < R x y\] proof eventually_elim fix ty :: real assume "0 < ty" have "escape_rate x = (\\<^sup>+y'. R x y' * indicator {y} y' \count_space UNIV) + (\\<^sup>+y'. R x y' * indicator (I x - {y}) y' \count_space UNIV)" unfolding ennreal_escape_rate by (subst nn_integral_add[symmetric]) (auto simp: less_le split: split_indicator intro!: nn_integral_cong) also have "\ = R x y + (\\<^sup>+y'. R x y' \count_space (I x - {y}))" by (auto simp add: nn_integral_count_space_indicator less_le simp del: nn_integral_indicator_singleton intro!: arg_cong2[where f="(+)"] nn_integral_cong split: split_indicator) finally have "(\\<^sup>+y'. R x y' \count_space (I x - {y})) = escape_rate x - R x y \ R x y \ escape_rate x" using escape_rate_pos[THEN less_imp_le] by (cases "(\\<^sup>+y'. R x y' \count_space (I x - {y}))") (auto simp: add_top ennreal_plus[symmetric] simp del: ennreal_plus) then have "integrable (count_space (I x - {y})) (R x)" "(LINT y'|count_space (I x - {y}). R x y') = escape_rate x - R x y" by (auto simp: nn_integral_eq_integrable) then have "?P' (prod_emb (I x-{y}) ?E (I x-{y}) (\\<^sub>E j\(I x-{y}). {ty<..})) = exp (- ty * (escape_rate x - R x y))" using I_countable \0 < ty\ by (subst emeasure_PiM_exponential_Ioi_countable) auto also have "prod_emb (I x-{y}) ?E (I x-{y}) (\\<^sub>E j\(I x-{y}). {ty<..}) = {p\space ?P'. \y'\I x-{y}. ty < p y'}" by (simp add: set_eq_iff prod_emb_def space_PiM space_exponential ac_simps Pi_iff) finally show "indicator {t..} ty * ?P' {p\space ?P'. \y'\I x-{y}. ty < p y'} = indicator {t..} ty * ennreal (exp (- ty * (escape_rate x - R x y)))" by simp qed also have "\ = (\\<^sup>+ty. ennreal (R x y) * (ennreal (exp (- ty * escape_rate x)) * indicator {max 0 t..} ty) \lborel)" by (auto simp add: exponential_def exponential_density_def nn_integral_density ennreal_mult[symmetric] exp_add[symmetric] field_simps intro!: nn_integral_cong split: split_indicator) also have "\ = (R x y / escape_rate x) * emeasure (exponential (escape_rate x)) {max 0 t..}" using escape_rate_pos[of x] by (auto simp: exponential_def exponential_density_def emeasure_density nn_integral_cmult[symmetric] ennreal_mult[symmetric] split: split_indicator intro!: nn_integral_cong ) also have "\ = pmf (J x) y * emeasure (exponential (escape_rate x)) {t..}" using AE_exponential[OF escape_rate_pos[of x]] by (intro arg_cong2[where f="(*)"] emeasure_eq_AE) (auto simp: J.rep_eq ) finally show ?thesis using assms by (simp add: mult_ac select_first_def) qed lemma AE_PAR_least: "AE p in PAR x. \y\I x. select_first x p y" proof - have D: "disjoint_family_on (\y. {p \ space (PAR x). select_first x p y}) (I x)" by (auto simp: disjoint_family_on_def dest: select_first_unique) have "PAR x {p\space (PAR x). \y\I x. select_first x p y} = PAR x (\y\I x. {p\space (PAR x). select_first x p y})" by (auto intro!: arg_cong2[where f=emeasure]) also have "\ = (\\<^sup>+y. PAR x {p\space (PAR x). select_first x p y} \count_space (I x))" using I_countable by (intro emeasure_UN_countable D) auto also have "\ = (\\<^sup>+y. PAR x {p\space (PAR x). 0 \ p y \ select_first x p y} \count_space (I x))" proof (intro nn_integral_cong emeasure_eq_AE, goal_cases) case (1 y) with AE_PiM_component[of "I x" "\y. exponential (R x y)" y "(<) 0"] AE_exponential[of "R x y"] show ?case by (auto simp: prob_space_exponential) qed (insert I_countable, auto) also have "\ = (\\<^sup>+y. emeasure (exponential (escape_rate x)) {0 ..} * ennreal (pmf (J x) y) \count_space (I x))" by (auto simp add: PAR_least intro!: nn_integral_cong) also have "\ = (\\<^sup>+y. emeasure (exponential (escape_rate x)) {0 ..} \J x)" by (auto simp: nn_integral_measure_pmf nn_integral_count_space_indicator ac_simps pmf_eq_0_set_pmf set_pmf_J simp del: nn_integral_const intro!: nn_integral_cong split: split_indicator) also have "\ = 1" using AE_exponential[of "escape_rate x"] by (auto intro!: prob_space.emeasure_eq_1_AE prob_space_exponential simp: escape_rate_pos less_imp_le) finally show ?thesis using I_countable by (subst prob_space.AE_iff_emeasure_eq_1 prob_space_PiM prob_space_exponential) (auto intro!: prob_space_PiM prob_space_exponential simp del: Set.bex_simps(6)) qed lemma K_alt: "K (t, x) = distr (\\<^sub>M y\I x. exponential (R x y)) S (\p. (t + (INF y\I x. p y), The (select_first x p)))" (is "_ = ?R") proof (rule measure_eqI_generator_eq_countable) let ?E = "{ {t ..} \ A | (t::real) A. A \ - I x \ (\s\I x. A = {s}) }" show "Int_stable ?E" apply (auto simp: Int_stable_def) subgoal for t1 A1 t2 A2 by (intro exI[of _ "max t1 t2"] exI[of _ "A1 \ A2"]) auto subgoal for t1 t2 y1 y2 by (intro exI[of _ "max t1 t2"] exI[of _ "{y1} \ {y2}"]) auto done show "sets (K (t, x)) = sigma_sets UNIV ?E" unfolding K.sets_K[OF in_space_S] by (subst sets_S_eq) rule show "sets ?R = sigma_sets UNIV ?E" using sets_S_eq by simp show "countable ((\(t, A). {t ..} \ A) ` (\ \ ({- I x} \ (\s. {s}) ` I x)))" by (intro countable_image countable_SIGMA countable_rat countable_Un I_countable) auto have *: "(+) t -` {t'..} \ space (exponential (escape_rate x)) = {t' - t..}" for t' by (auto simp: space_exponential) { fix X assume "X \ ?E" then consider t' s where "s \ I x" "X = {t' ..} \ {s}" | t' A where "A \ - I x" "X = {t' ..} \ A" by auto then show "K (t, x) X = ?R X" proof cases case 1 have "AE p in PAR x. (t' - t \ p s \ select_first x p s) = (t' \ t + (\x\I x. p x) \ The (select_first x p) = s)" using AE_PAR_least by eventually_elim (auto dest: select_first_unique simp: select_first_INF) with 1 I_countable show ?thesis by (auto simp add: K_def measure_pmf.emeasure_pair_measure_Times emeasure_distr emeasure_pmf_single * PAR_least[symmetric] intro!: emeasure_eq_AE) next case 2 moreover then have "emeasure (measure_pmf (J x)) A = 0" by (subst AE_iff_measurable[symmetric, where P="\x. x \ A"]) (auto simp: AE_measure_pmf_iff set_pmf_J subset_eq) moreover have "PAR x ((\p. (t + \(p ` (I x)), The (select_first x p))) -` ({t'..} \ A) \ space (PAR x)) = 0" using \A \ - I x\ AE_PAR_least[of x] I_countable by (subst AE_iff_measurable[symmetric, where P="\p. (t + \(p ` (I x)), The (select_first x p)) \ {t'..} \ A"]) (auto simp del: all_simps(5) simp add: imp_ex imp_conjL subset_eq) ultimately show ?thesis using I_countable by (simp add: K_def measure_pmf.emeasure_pair_measure_Times emeasure_distr *) qed } interpret prob_space "K ts" for ts by (rule K.prob_space_K) simp show "emeasure (K (t, x)) a \ \" for a using emeasure_finite by simp qed (insert Rats_dense_in_real[of "x - 1" x for x], auto, blast intro: less_imp_le) lemma AE_K: "AE y in K x. fst x < fst y \ snd y \ J (snd x)" unfolding K_def split_beta apply (subst exp_esc.AE_pair_iff[symmetric]) apply measurable apply (simp_all add: AE_distr_iff AE_measure_pmf_iff exponential_def AE_density exponential_density_def cong del: AE_cong) using AE_lborel_singleton[of 0] apply eventually_elim apply simp done lemma AE_lim_stream: "AE \ in K.lim_stream x. \i. snd ((x ## \) !! i) \ DTMC.acc``{snd x} \ snd (\ !! i) \ J (snd ((x ## \) !! i)) \ fst ((x ## \) !! i) < fst (\ !! i)" (is "AE \ in K.lim_stream x. \i. ?P \ i") unfolding AE_all_countable proof let ?F = "\i x \. fst ((x ## \) !! i)" and ?S = "\i x \. snd ((x ## \) !! i)" fix i show "AE \ in K.lim_stream x. ?P \ i" proof (induction i arbitrary: x) case 0 with AE_K[of x] show ?case by (subst K.AE_lim_stream) (auto simp add: space_pair_measure cong del: AE_cong) next case (Suc i) show ?case proof (subst K.AE_lim_stream, goal_cases) case 2 show ?case using DTMC.countable_reachable by (intro measurable_compose_countable_restrict[where f="?S (Suc i) x"]) (simp_all del: Image_singleton_iff) next case 3 show ?case apply (simp del: AE_conj_iff cong del: AE_cong) using AE_K[of x] apply eventually_elim subgoal premises K_prems for y using Suc by eventually_elim (insert K_prems, auto intro: converse_rtrancl_into_rtrancl) done qed (simp add: space_pair_measure) qed qed lemma measurable_merge_at[measurable]: "(\(\, \'). merge_at \ j \') \ (T \\<^sub>M T) \\<^sub>M T" proof (rule measurable_stream_space2) define F where "F x n = (case x of (\::(real \ 'a) stream, \') \ merge_at \ j \') !! n" for x n fix n have "(\x. F x n) \ stream_space S \\<^sub>M stream_space S \\<^sub>M S" proof (induction n) case 0 then show ?case by (simp add: F_def split_beta' stream.case_eq_if) next case (Suc n) from Suc[measurable] have eq: "F x (Suc n) = (case fst x of (t, s) ## \ \ if t \ j then F (\, snd x) n else snd x !! Suc n)" for x by (auto simp: F_def split: prod.split stream.split) show ?case unfolding eq stream.case_eq_if by measurable qed then show "(\x. (case x of (\, \') \ merge_at \ j \') !! n) \ stream_space S \\<^sub>M stream_space S \\<^sub>M S" unfolding F_def by auto qed lemma measurable_trace_at[measurable]: "(\(s, \). trace_at s \ j) \ (count_space UNIV \\<^sub>M T) \\<^sub>M count_space UNIV" unfolding trace_at_eq by measurable lemma measurable_trace_at': "(\((s, j), \). trace_at s \ j) \ ((count_space UNIV \\<^sub>M borel) \\<^sub>M T) \\<^sub>M count_space UNIV" unfolding trace_at_eq split_beta' by measurable lemma K_time_split: assumes "t \ j" and [measurable]: "f \ S \\<^sub>M borel" shows "(\\<^sup>+x. f x * indicator {j <..} (fst x) \K (t, s)) = (\\<^sup>+x. f x \K (j, s)) * exponential (escape_rate s) {j - t <..}" proof - have "(\\<^sup>+ y. \\<^sup>+ x. f (t + x, y) * indicator {j<..} (t + x) \exponential (escape_rate s) \J s) = (\\<^sup>+ y. \\<^sup>+ x. f (t + x, y) * indicator {j - t<..} x \exponential (escape_rate s) \J s)" by (intro nn_integral_cong) (auto split: split_indicator) also have "\ = (\\<^sup>+ y. \\<^sup>+ x. f (t + x, y) \uniform_measure (exponential (escape_rate s)) {j-t <..} \J s) * emeasure (exponential (escape_rate s)) {j - t <..}" using \t \ j\ escape_rate_pos by (subst nn_integral_uniform_measure) (auto simp: nn_integral_divide ennreal_divide_times emeasure_exponential_Ioi) also have "\ = (\\<^sup>+ y. \\<^sup>+ x. f (j + x, y) \exponential (escape_rate s) \J s) * emeasure (exponential (escape_rate s)) {j - t <..}" using \t \ j\ escape_rate_pos by (simp add: uniform_measure_exponential nn_integral_distr) finally show ?thesis by (simp add: K_def exp_esc.nn_integral_snd[symmetric] nn_integral_distr) qed lemma K_in_space[simp]: "K x \ space (prob_algebra S)" by (rule measurable_space [OF K.K]) simp lemma L_in_space[simp]: "K.lim_stream x \ space (prob_algebra T)" by (rule measurable_space [OF K.lim_stream]) simp subsection \Markov Chain Property\ lemma lim_time_split: "t \ j \ K.lim_stream (t, s) = do { \ \ K.lim_stream (t, s) ; \' \ K.lim_stream (j, trace_at s \ j) ; return T (merge_at \ j \')}" (is "_ \ _ = ?DO t s") proof (coinduction arbitrary: t s rule: K.lim_stream_eq_coinduct) case step let ?L = K.lim_stream note measurable_compose[OF measurable_prob_algebraD measurable_emeasure_subprob_algebra, measurable (raw)] define B' where "B' = (\(t', s). if t' \ j then ?DO t' s else ?L (t', s))" show ?case proof (intro bexI conjI AE_I2) show [measurable]: "B' \ S \\<^sub>M prob_algebra T" unfolding B'_def by measurable show "(\t s. y = (t, s) \ B' y = ?DO t s \ t \ j) \ ?L y = B' y" for y by (cases y; cases "fst y \ j") (auto simp: B'_def) let ?C = "\x. do { \ \ ?L x; \' \ ?L (j, trace_at s (x##\) j); return T (merge_at (x##\) j \') }" have "?DO t s = do { x \ K (t, s); ?C x }" apply (subst K.lim_stream_eq[OF in_space_S]) apply (subst bind_assoc[OF measurable_prob_algebraD measurable_prob_algebraD]) apply (subst measurable_cong_sets[OF K.sets_K[OF in_space_S] refl]) apply measurable apply (subst bind_assoc[OF measurable_prob_algebraD measurable_prob_algebraD]) apply measurable apply (subst bind_cong[OF refl bind_cong[OF refl bind_return[OF measurable_prob_algebraD]]]) apply measurable done also have "\ = K (t, s) \ (\y. B' y \ (\\. return T (y ## \)))" (is "?DO' = ?R") proof (rule measure_eqI) have "sets ?DO' = sets T" by (intro sets_bind'[OF K_in_space]) measurable moreover have "sets ?R = sets T" by (intro sets_bind'[OF K_in_space]) measurable ultimately show "sets ?DO' = sets ?R" by simp fix A assume "A \ sets ?DO'" then have A[measurable]: "A \ T" unfolding \sets ?DO' = sets T\ . have "?DO' A = (\\<^sup>+x. ?C x A \K (t, s))" by (subst emeasure_bind_prob_algebra[OF K_in_space]) measurable also have "\ = (\\<^sup>+x. ?C x A * indicator {.. j} (fst x) \K (t, s)) + (\\<^sup>+x. ?C x A * indicator {j <..} (fst x) \K (t, s))" by (subst nn_integral_add[symmetric]) (auto intro!: nn_integral_cong split: split_indicator) also have "(\\<^sup>+x. ?C x A * indicator {.. j} (fst x) \K (t, s)) = (\\<^sup>+y. emeasure (B' y \ (\\. return T (y ## \))) A * indicator {.. j} (fst y) \K (t, s))" proof (intro nn_integral_cong ennreal_mult_right_cong refl arg_cong2[where f=emeasure]) fix x :: "real \ 'a" assume "indicator {..j} (fst x) \ (0::ennreal)" then have "fst x \ j" by (auto split: split_indicator_asm) then show "?C x = (B' x \ (\\. return T (x ## \)))" apply (cases x) apply (simp add: B'_def) apply (subst bind_assoc[OF measurable_prob_algebraD measurable_prob_algebraD]) apply measurable apply (subst bind_assoc[OF measurable_prob_algebraD measurable_prob_algebraD]) apply measurable apply (subst bind_return) apply measurable done qed also have "(\\<^sup>+x. ?C x A * indicator {j <..} (fst x) \K (t, s)) = (\\<^sup>+y. emeasure (B' y \ (\\. return T (y ## \))) A * indicator {j <..} (fst y) \K (t, s))" proof - have *: "(+) t -` {j<..} = {j - t <..}" by auto have "(\\<^sup>+x. ?C x A * indicator {j <..} (fst x) \K (t, s)) = (\\<^sup>+x. ?L (j, s) A * indicator {j <..} (fst x) \K (t, s))" by (intro nn_integral_cong ennreal_mult_right_cong refl arg_cong2[where f=emeasure]) (auto simp: K.sets_lim_stream bind_return'' bind_const' prob_space_K_lim prob_space_imp_subprob_space split: split_indicator_asm) also have "\ = ?L (j, s) A * exponential (escape_rate s) {j - t <..}" by (subst nn_integral_cmult) (simp_all add: K_def exp_esc.nn_integral_snd[symmetric] emeasure_distr space_exponential *) also have "\ = (\\<^sup>+x. emeasure (?L x \ (\\. return T (x ## \))) A \K (j, s)) * exponential (escape_rate s) {j - t <..}" by (subst K.lim_stream_eq) (auto simp: emeasure_bind_prob_algebra[OF K_in_space _ A]) also have "\ = (\\<^sup>+y. emeasure (?L y \ (\\. return T (y ## \))) A * indicator {j <..} (fst y) \K (t, s))" using \t \ j\ by (rule K_time_split[symmetric]) measurable also have "\ = (\\<^sup>+y. emeasure (B' y \ (\\. return T (y ## \))) A * indicator {j <..} (fst y) \K (t, s))" by (intro nn_integral_cong ennreal_mult_right_cong refl arg_cong2[where f=emeasure]) (auto simp add: B'_def split: split_indicator_asm) finally show ?thesis . qed also have "(\\<^sup>+y. emeasure (B' y \ (\\. return T (y ## \))) A * indicator {.. j} (fst y) \K (t, s)) + (\\<^sup>+y. emeasure (B' y \ (\\. return T (y ## \))) A * indicator {j <..} (fst y) \K (t, s)) = (\\<^sup>+y. emeasure (B' y \ (\\. return T (y ## \))) A \K (t, s))" by (subst nn_integral_add[symmetric]) (auto intro!: nn_integral_cong split: split_indicator) also have "\ = emeasure (K (t, s) \ (\y. B' y \ (\\. return T (y ## \)))) A" by (rule emeasure_bind_prob_algebra[symmetric, OF K_in_space _ A]) auto finally show "?DO' A = emeasure (K (t, s) \ (\y. B' y \ (\\. return T (y ## \)))) A" . qed finally show "?DO t s = K (t, s) \ (\y. B' y \ (\\. return T (y ## \)))" . qed qed (simp add: space_pair_measure) lemma K_eq: "K (t, s) = distr (exponential (escape_rate s) \\<^sub>M J s) S (\(t', s). (t + t', s))" proof - have "distr (exponential (escape_rate s)) borel ((+) t) \\<^sub>M distr (J s) (J s) (\x. x) = distr (exponential (escape_rate s) \\<^sub>M J s) (borel \\<^sub>M J s) (\(x, y). (t + x, y))" proof (intro pair_measure_distr) interpret prob_space "distr (measure_pmf (J s)) (measure_pmf (J s)) (\x. x)" by (intro measure_pmf.prob_space_distr) simp show "sigma_finite_measure (distr (measure_pmf (J s)) (measure_pmf (J s)) (\x. x))" by unfold_locales qed auto also have "\ = distr (exponential (escape_rate s) \\<^sub>M J s) S (\(x, y). (t + x, y))" by (intro distr_cong refl sets_pair_measure_cong) simp finally show ?thesis by (simp add: K_def) qed lemma K_shift: "K (t + t', s) = distr (K (t, s)) S (\(t, s). (t + t', s))" unfolding K_eq by (subst distr_distr) (auto simp: comp_def split_beta' ac_simps) lemma K_not_empty: "space (K x) \ {}" by (simp add: K_def space_pair_measure split: prod.split) lemma lim_stream_not_empty: "space (K.lim_stream x) \ {}" by (simp add: K.space_lim_stream space_pair_measure split: prod.split) lemma lim_shift: \ \Generalize to bijective function on @{const K.lim_stream} invariant on @{const K}\ "K.lim_stream (t + t', s) = distr (K.lim_stream (t, s)) T (smap (\(t, s). (t + t', s)))" (is "_ = ?D t s") proof (coinduction arbitrary: t s rule: K.lim_stream_eq_coinduct) case step then show ?case proof (intro bexI[of _ "\(t, s). ?D (t - t') s"] conjI) show "?D t s = K (t + t', s) \ (\y. (case y of (t, s) \ ?D (t - t') s) \ (\\. return T (y ## \)))" apply (subst K.lim_stream_eq[OF in_space_S]) apply (subst K_shift) apply (subst distr_bind[OF measurable_prob_algebraD K_not_empty]) apply (measurable; fail) apply (measurable; fail) apply (subst bind_distr[OF _ measurable_prob_algebraD K_not_empty]) apply (measurable; fail) apply (measurable; fail) apply (intro bind_cong refl) apply (subst distr_bind[OF measurable_prob_algebraD lim_stream_not_empty]) apply (measurable; fail) apply (measurable; fail) apply (simp add: distr_return split_beta) apply (subst bind_distr[OF _ measurable_prob_algebraD lim_stream_not_empty]) apply (measurable; fail) apply (measurable; fail) apply (simp add: split_beta') done qed (auto cong: conj_cong intro!: exI[of _ "_ - t'"]) qed simp lemma lim_0: "K.lim_stream (t, s) = distr (K.lim_stream (0, s)) T (smap (\(t', s). (t' + t, s)))" using lim_shift[of 0 t s] by simp subsection \Explosion time\ definition explosion :: "(real \ 'a) stream \ ereal" where "explosion \ = (SUP i. ereal (fst (\ !! i)))" lemma ball_less_Suc_eq: "(\i (P 0 \ (\i_. borel)) (\\ i. escape_rate (snd ((ts##\) !! i)) * (fst (\ !! i) - fst ((ts##\) !! i))) = PiM UNIV (\_. exponential 1)" (is "?D = ?P") proof (rule measure_eqI_PiM_sequence) show "sets ?D = sets (PiM UNIV (\_. borel))" "sets ?P = sets (PiM UNIV (\_. borel))" by (auto intro!: sets_PiM_cong simp: sets_exponential) have [measurable]: "ts \ space S" by auto { interpret prob_space ?D by (intro prob_space.prob_space_distr K.prob_space_lim_stream measurable_abs_UNIV) auto show "finite_measure ?D" by unfold_locales } interpret E: prob_space "exponential 1" by (rule prob_space_exponential) simp interpret P: product_prob_space "\_. exponential 1" UNIV by unfold_locales let "distr _ _ (?f ts)" = ?D fix A :: "nat \ real set" and n :: nat assume A[measurable]: "\i. A i \ sets borel" define n' where "n' = Suc n" have "emeasure ?D (prod_emb UNIV (\_. borel) {..n} (Pi\<^sub>E {..n} A)) = emeasure (K.lim_stream ts) {\\space (stream_space S). \i i \ A i}" apply (subst emeasure_distr) apply (auto intro!: measurable_abs_UNIV arg_cong[where f="emeasure _"]) apply (auto simp: prod_emb_def K.space_lim_stream space_pair_measure n'_def) done also have "\ = (\i space S" by auto have "emeasure (K.lim_stream ts) {\ \ space (stream_space S). \i i \ A i} = (\\<^sup>+ts'. indicator (A 0) (escape_rate (snd ts) * (fst ts' - fst ts)) * emeasure (K.lim_stream ts') {\ \ space (stream_space S). \i i \ A (Suc i)} \K ts)" apply (subst K.emeasure_lim_stream) apply simp apply measurable apply (auto intro!: nn_integral_cong arg_cong2[where f=emeasure] split: split_indicator simp: ball_less_Suc_eq) done also have "\ = (\\<^sup>+ts'. indicator (A 0) (escape_rate (snd ts) * (fst ts' - fst ts)) \K ts) * (\i\<^sup>+ts'. indicator (A 0) (escape_rate (snd ts) * (fst ts' - fst ts)) \K ts) = (\\<^sup>+t. indicator (A 0) (escape_rate (snd ts) * t) \exponential (escape_rate (snd ts)))" by (simp add: K_def exp_esc.nn_integral_snd[symmetric] nn_integral_distr split: prod.split) also have "\ = emeasure (exponential 1) (A 0)" using escape_rate_pos[of "snd ts"] by (subst exponential_eq_stretch) (simp_all add: nn_integral_distr) also have "emeasure (exponential 1) (A 0) * (\ii = emeasure ?P (prod_emb UNIV (\_. borel) {..E {.._. borel) {..n} (Pi\<^sub>E {..n} A)) = emeasure ?P (prod_emb UNIV (\_. borel) {..n} (Pi\<^sub>E {..n} A))" by (simp add: n'_def lessThan_Suc_atMost) qed lemma AE_explosion_infty: assumes bdd: "bdd_above (range escape_rate)" shows "AE \ in K.lim_stream x. explosion \ = \" proof - have "escape_rate undefined \ (SUP x. escape_rate x)" using bdd by (intro cSUP_upper) auto then have SUP_escape_pos: "0 < (SUP x. escape_rate x)" using escape_rate_pos[of undefined] by simp then have SUP_escape_nonneg: "0 \ (SUP x. escape_rate x)" by (rule less_imp_le) have [measurable]: "x \ space S" by auto have "(\i. 1::ennreal) = top" by (rule sums_unique[symmetric]) (auto simp: sums_def of_nat_tendsto_top_ennreal) then have "AE \ in (PiM UNIV (\_. exponential 1)). (\i. ereal (\ i)) = \" by (intro AE_PiM_exponential_suminf_infty) auto then have "AE \ in K.lim_stream x. (\i. ereal (escape_rate (snd ((x##\) !! i)) * (fst (\ !! i) - fst ((x##\) !! i)))) = \" apply (subst (asm) lim_stream_timediff_eq_exponential_1[symmetric, of x]) apply (subst (asm) AE_distr_iff) apply (auto intro!: measurable_abs_UNIV) done then show ?thesis using AE_lim_stream proof eventually_elim case (elim \) then have le: "fst ((x##\) !! n) \ fst ((x ## \) !! m)" if "n \ m" for n m by (intro lift_Suc_mono_le[OF _ \n \ m\, of "\i. fst ((x ## \) !! i)"]) (auto intro: less_imp_le) have [simp]: "fst x \ fst ((x##\) !! i)" "fst ((x##\) !! i) \ fst (\ !! i)" for i using le[of "i" "Suc i"] le[of 0 i] by auto have "(\i. ereal (escape_rate (snd ((x ## \) !! i)) * (fst (\ !! i) - fst ((x ## \) !! i)))) = (SUP n. \i) !! i)) * (fst (\ !! i) - fst ((x ## \) !! i))))" by (intro suminf_ereal_eq_SUP) (auto intro!: mult_nonneg_nonneg) also have "\ \ (SUP n. (SUP x. escape_rate x) * (ereal (fst ((x ## \) !! n)) - ereal (fst x)))" proof (intro SUP_least SUP_upper2) fix n have "(\i) !! i)) * (fst (\ !! i) - fst ((x ## \) !! i)))) \ (\i !! i) - fst ((x ## \) !! i))))" using elim bdd by (intro sum_mono) (auto intro!: cSUP_upper) also have "\ = (SUP i. escape_rate i) * (\i) !! Suc i) - fst ((x ## \) !! i))" using elim bdd by (subst sum_ereal) (auto simp: sum_distrib_left) also have "\ = (SUP i. escape_rate i) * (fst ((x ## \) !! n) - fst x)" by (subst sum_lessThan_telescope) simp finally show "(\i) !! i)) * (fst (\ !! i) - fst ((x ## \) !! i)))) \ (SUP x. escape_rate x) * (ereal (fst ((x ## \) !! n)) - ereal (fst x))" by simp qed simp also have "\ = (SUP x. escape_rate x) * ((SUP n. ereal (fst ((x ## \) !! n))) - ereal (fst x))" using elim SUP_escape_nonneg by (subst SUP_ereal_mult_left) (auto simp: SUP_ereal_minus_left[symmetric]) also have "(SUP n. ereal (fst ((x ## \) !! n))) = explosion \" unfolding explosion_def apply (intro SUP_eq) subgoal for i by (intro bexI[of _ i]) auto subgoal for i by (intro bexI[of _ "Suc i"]) auto done finally show "explosion \ = \" using elim SUP_escape_pos by (cases "explosion \") (auto split: if_splits) qed qed subsection \Transition probability $p_t$\ context begin declare [[inductive_internals = true]] inductive trace_in :: "'a set \ real \ 'a \ (real \ 'a) stream \ bool" for S t where "t < t' \ s \ S \ trace_in S t s ((t', s')##\)" | "t \ t' \ trace_in S t s' \ \ trace_in S t s ((t', s')##\)" end lemma trace_in_simps[simp]: "trace_in ss t s (x##\) = (if t < fst x then s \ ss else trace_in ss t (snd x) \)" by (cases x) (subst trace_in.simps; auto) lemma trace_in_eq_lfp: "trace_in ss t = lfp (\F s. \(t', s')##\ \ if t < t' then s \ ss else F s' \)" unfolding trace_in_def by (intro arg_cong[where f=lfp] ext) (auto split: stream.splits) lemma trace_in_shiftD: "trace_in ss t s \ \ trace_in ss (t + t') s (smap (\(t, s'). (t + t', s')) \)" by (induction rule: trace_in.induct) auto lemma trace_in_shift[simp]: "trace_in ss t s (smap (\(t, s'). (t + t', s')) \) \ trace_in ss (t - t') s \" using trace_in_shiftD[of ss t s "smap (\(t, s'). (t + t', s')) \" "- t'"] trace_in_shiftD[of ss "t - t'" s \ t'] by (auto simp add: stream.map_comp prod.case_eq_if) lemma measurable_trace_in': "Measurable.pred (borel \\<^sub>M count_space UNIV \\<^sub>M T) (\(t, s, \). trace_in ss t s \)" (is "?M (\(t, s, \). trace_in ss t s \)") proof - let ?F = "\F. \(t, s, (t', s')##\) \ if t < t' then s \ ss else F (t, s', \)" have [measurable]: "Measurable.pred (count_space UNIV) (\x. x \ ss)" by simp have "trace_in ss = (\t s \. lfp ?F (t, s, \))" unfolding trace_in_def apply (subst lfp_arg) apply (subst lfp_rolling[where g="\F t s \. F (t, s, \)"]) subgoal by (auto simp: mono_def le_fun_def split: stream.splits) subgoal by (auto simp: mono_def le_fun_def split: stream.splits) subgoal by (intro arg_cong[where f=lfp]) (auto simp: mono_def le_fun_def split_beta' not_less fun_eq_iff split: stream.splits intro!: arg_cong[where f=lfp]) done then have eq: "(\(t, s, \). trace_in ss t s \) = lfp ?F" by simp have "sup_continuous ?F" by (auto simp: sup_continuous_def fun_eq_iff split: stream.splits) then show ?thesis unfolding eq proof (rule measurable_lfp) fix F assume "?M F" then show "?M (?F F)" by measurable qed qed lemma measurable_trace_in[measurable (raw)]: assumes [measurable]: "f \ M \\<^sub>M borel" "g \ M \\<^sub>M count_space UNIV" "h \ M \\<^sub>M T" shows "Measurable.pred M (\x. trace_in ss (f x) (g x) (h x))" using measurable_compose[of "\x. (f x, g x, h x)" M, OF _ measurable_trace_in'[of ss]] by simp definition p :: "'a \ 'a \ real \ real" where "p s s' t = \

(\ in K.lim_stream (0, s). trace_in {s'} t s \)" lemma p[measurable]: "(\(s, t). p s s' t) \ (count_space UNIV \\<^sub>M borel) \\<^sub>M borel" proof - have *: "(SIGMA x:space (count_space UNIV \\<^sub>M borel). {\ \ streams (space S). trace_in {s'} (snd x) (fst x) \}) = {x\space ((count_space UNIV \\<^sub>M borel) \\<^sub>M T). trace_in {s'} (snd (fst x)) (fst (fst x)) (snd x)}" by (auto simp: space_pair_measure) note measurable_trace_at'[measurable] show ?thesis unfolding p_def[abs_def] split_beta' by (rule measure_measurable_prob_algebra2[where N=T]) (auto simp: K.space_lim_stream * pred_def[symmetric] intro!: pred_count_space_const1 measurable_trace_at'[unfolded split_beta']) qed lemma p_nonpos: assumes "t \ 0" shows "p s s' t = of_bool (s = s')" proof - have "AE \ in K.lim_stream (0, s). trace_in {s'} t s \ = (s = s')" proof (subst K.AE_lim_stream) show "AE y in K (0, s). AE \ in K.lim_stream y. trace_in {s'} t s (y ## \) = (s = s')" using AE_K proof eventually_elim fix y :: "real \ 'a" assume "fst (0, s) < fst y \ snd y \ set_pmf (J (snd (0, s)))" with \t\0\ show "AE \ in K.lim_stream y. trace_in {s'} t s (y ## \) = (s = s')" by (cases y) auto qed qed auto then have "p s s' t = \

(\ in K.lim_stream (0, s). s = s')" unfolding p_def by (intro prob_space.prob_eq_AE K.prob_space_lim_stream) auto then show ?thesis using prob_space.prob_space[OF K.prob_space_lim_stream] by simp qed lemma p_0: "p s s' 0 = of_bool (s = s')" using p_nonpos[of 0] by simp lemma in_sets_T[measurable (raw)]: "Measurable.pred T P \ {\. P \} \ sets T" unfolding pred_def by simp lemma distr_id': "sets M = sets N \ distr M N (\x. x) = M" by (subst distr_cong[of M M N M _ "\x. x"] ) simp_all lemma p_nonneg[simp]: "0 \ p s s' t" by (simp add: p_def) lemma p_le_1[simp]: "p s s' t \ 1" unfolding p_def by (intro prob_space.prob_le_1 K.prob_space_lim_stream) simp lemma p_eq: assumes "0 \ t" shows "p s s'' t = (of_bool (s = s'') + (LINT u:{0..t}|lborel. escape_rate s * exp (escape_rate s * u) * (LINT s'|J s. p s' s'' u))) / exp (t * escape_rate s)" proof - have *: "(+) 0 = (\x::real. x)" by auto interpret L: prob_space "K.lim_stream x" for x by (rule K.prob_space_lim_stream) simp interpret E: prob_space "exponential (escape_rate s)" for s by (intro escape_rate_pos prob_space_exponential) have "p s s'' t = emeasure (K.lim_stream (0, s)) {\\space T. trace_in {s''} t s \}" by (simp add: p_def L.emeasure_eq_measure K.space_lim_stream space_stream_space del: in_space_T) also have "\ = (\\<^sup>+y. emeasure (K.lim_stream y) {\\space T. trace_in {s''} t s (y##\) } \K (0, s))" apply (subst K.lim_stream_eq[OF in_space_S]) apply (subst emeasure_bind_prob_algebra[OF K_in_space]) apply (measurable; fail) apply (measurable; fail) apply (subst bind_return_distr'[OF lim_stream_not_empty]) apply (measurable; fail) apply (simp add: emeasure_distr) done also have "\ = (\\<^sup>+y. indicator {t <..} (fst y) * of_bool (s = s'') + indicator {0<..t} (fst y) * p (snd y) s'' (t - fst y) \K (0, s))" apply (intro nn_integral_cong_AE) using AE_K apply eventually_elim subgoal for y using L.emeasure_space_1 apply (cases y) apply (auto split: split_indicator simp del: in_space_T) subgoal for t' s2 unfolding p_def L.emeasure_eq_measure[symmetric] K.space_lim_stream space_stream_space[symmetric] by (subst lim_0) (simp add: emeasure_distr) subgoal by (auto split: split_indicator cong: rev_conj_cong simp add: K.space_lim_stream space_stream_space simp del: in_space_T) done done also have "\ = (\\<^sup>+u. \\<^sup>+s'. indicator {t <..} u * of_bool (s = s'') + indicator {0<..t} u * p s' s'' (t - u) \J s \exponential (escape_rate s))" unfolding K_def by (simp add: K_def measure_pmf.nn_integral_fst[symmetric] * distr_id' sets_exponential) also have "\ = ennreal (exp (- t * escape_rate s) * of_bool (s = s'')) + (\\<^sup>+u. indicator {0<..t} u * \\<^sup>+s'. p s' s'' (t - u) \J s \exponential (escape_rate s))" using \0\t\ by (simp add: nn_integral_add nn_integral_cmult ennreal_indicator ennreal_mult emeasure_exponential_Ioi escape_rate_pos) also have "(\\<^sup>+u. indicator {0<..t} u * \\<^sup>+s'. p s' s'' (t - u) \J s \exponential (escape_rate s)) = (\\<^sup>+u. indicator {0<..t} u *\<^sub>R (LINT s'|J s. p s' s'' (t - u)) \exponential (escape_rate s))" by (simp add: measure_pmf.integrable_const_bound[of _ 1] nn_integral_eq_integral ennreal_mult ennreal_indicator) also have "\ = (LINT u:{0<..t}|exponential (escape_rate s). (LINT s'|J s. p s' s'' (t - u)))" unfolding set_lebesgue_integral_def by (intro nn_integral_eq_integral E.integrable_const_bound[of _ 1] AE_I2) (auto intro!: mult_le_one measure_pmf.integral_le_const measure_pmf.integrable_const_bound[of _ 1]) also have "\ = (LINT u:{0<..t}|lborel. escape_rate s * exp (- escape_rate s * u) * (LINT s'|J s. p s' s'' (t - u)))" unfolding exponential_def set_lebesgue_integral_def by (subst integral_density) (auto simp: ac_simps exponential_density_def fun_eq_iff split: split_indicator simp del: integral_mult_right integral_mult_right_zero intro!: arg_cong2[where f="integral\<^sup>L"]) also have "\ = (LINT u:{0..t}|lborel. escape_rate s * exp (- escape_rate s * (t - u)) * (LINT s'|J s. p s' s'' u))" using AE_lborel_singleton[of 0] AE_lborel_singleton[of t] unfolding set_lebesgue_integral_def by (subst lborel_integral_real_affine[where t=t and c="-1"]) (auto intro!: integral_cong_AE split: split_indicator) also have "\ = exp (- t * escape_rate s) * escape_rate s * (LINT u:{0..t}|lborel. exp (escape_rate s * u) * (LINT s'|J s. p s' s'' u))" by (simp add: field_simps exp_diff exp_minus) finally show "p s s'' t = (of_bool (s = s'') + (LBINT u:{0..t}. escape_rate s * exp (escape_rate s * u) * (LINT s'|J s. p s' s'' u))) / exp (t * escape_rate s)" unfolding set_lebesgue_integral_def by (simp del: ennreal_plus add: ennreal_plus[symmetric] exp_minus field_simps) qed lemma continuous_on_p: "continuous_on A (p s s')" proof - interpret E: prob_space "exponential (escape_rate s'')" for s'' by (intro escape_rate_pos prob_space_exponential) have "continuous_on {..0} (p s s')" by (simp add: p_nonpos continuous_on_const cong: continuous_on_cong_simp) moreover have "continuous_on {0..} (p s s')" proof (subst continuous_on_cong[OF refl p_eq]) let ?I = "\t. escape_rate s * exp (escape_rate s * t) * (LINT s''|J s. p s'' s' t)" show "continuous_on {0..} (\t. (of_bool (s = s') + (LBINT u:{0..t}. ?I u)) / exp (t * escape_rate s))" proof (intro continuous_intros continuous_on_LBINT[THEN continuous_on_subset]) fix t :: real assume t: "0 \ t" then have "0 \ x \ x \ t \ exp (x * escape_rate s) * (LINT s''|J s. p s'' s' x) \ exp (t * escape_rate s) * 1" for x by (intro mult_mono) (auto intro!: mult_mono measure_pmf.integral_le_const measure_pmf.integrable_const_bound[of _ 1]) with t show "set_integrable lborel {0..t} ?I" using escape_rate_pos[of s] unfolding set_integrable_def by (intro integrableI_bounded_set_indicator[where B="escape_rate s * exp (escape_rate s * t)"]) (auto simp: field_simps) qed auto qed simp ultimately have "continuous_on ({0..} \ {..0}) (p s s')" by (intro continuous_on_closed_Un) auto also have "{0..} \ {..0::real} = UNIV" by auto finally show ?thesis by (rule continuous_on_subset) simp qed lemma p_vector_derivative: \ \Backward equation\ assumes "0 \ t" shows "(p s s' has_vector_derivative (LINT s''|count_space UNIV. R s s'' * p s'' s' t) - escape_rate s * p s s' t) (at t within {0..})" (is "(_ has_vector_derivative ?A) _") proof - let ?I = "\t. escape_rate s * exp (escape_rate s * t) * (LINT s''|J s. p s'' s' t)" let ?p = "\t. (of_bool (s = s') + integral {0..t} ?I) * exp (t *\<^sub>R - escape_rate s)" { fix t :: real assume "0 \ t" have "p s s' t = (of_bool (s = s') + (LBINT u:{0..t}. ?I u)) * exp (- t * escape_rate s)" using p_eq[OF \0 \ t\, of s s'] by (simp add: exp_minus field_simps) also have "(LBINT u:{0..t}. ?I u) = integral {0..t} ?I" proof (intro set_borel_integral_eq_integral) have "0 \ x \ x \ t \ exp (x * escape_rate s) * (LINT s''|J s. p s'' s' x) \ exp (t * escape_rate s) * 1" for x by (intro mult_mono) (auto intro!: mult_mono measure_pmf.integral_le_const measure_pmf.integrable_const_bound[of _ 1]) with \0\t\ show "set_integrable lborel {0..t} ?I" using escape_rate_pos[of s] unfolding set_integrable_def by (intro integrableI_bounded_set_indicator[where B="escape_rate s * exp (escape_rate s * t)"]) (auto simp: field_simps) qed finally have "p s s' t = ?p t" by simp } note p_eq = this have at_eq: "at t within {0..} = at t within {0 .. t + 1}" by (intro at_within_nhd[where S="{..< t+1}"]) auto have c_I: "continuous_on {0..t + 1} ?I" by (intro continuous_intros continuous_on_LINT_pmf[where B=1] continuous_on_p) simp show ?thesis proof (subst has_vector_derivative_cong_ev) show "\\<^sub>F u in nhds t. u \ {0..} \ p s s' u = ?p u" "p s s' t = ?p t" using \0\t\ by (simp_all add: p_eq) have "(?p has_vector_derivative escape_rate s * ((LINT s''|J s. p s'' s' t) - p s s' t)) (at t within {0..})" unfolding at_eq apply (intro refl derivative_eq_intros) apply rule apply (rule integral_has_vector_derivative[OF c_I]) apply (simp add: \0 \ t\) apply rule apply (rule exp_scaleR_has_vector_derivative_right) apply (simp add: field_simps exp_minus p_eq \0\t\ split del: split_of_bool) done also have "escape_rate s * ((LINT s''|J s. p s'' s' t) - p s s' t) = (LINT s''|count_space UNIV. R s s'' * p s'' s' t) - escape_rate s * p s s' t" using escape_rate_pos[of s] by (simp add: measure_pmf_eq_density integral_density J.rep_eq field_simps) finally show "(?p has_vector_derivative ?A) (at t within {0..})" . qed qed coinductive wf_times :: "real \ (real \ 'a) stream \ bool" where "t < t' \ wf_times t' \ \ wf_times t ((t', s') ## \)" lemma wf_times_simp[simp]: "wf_times t (x ## \) \ t < fst x \ wf_times (fst x) \" by (cases x) (subst wf_times.simps; simp) lemma trace_in_merge_at: assumes \': "wf_times t' \'" shows "trace_in ss t x (merge_at \ t' \') \ (if t < t' then trace_in ss t x \ else \y. trace_in {y} t' x \ \ trace_in ss t y \')" (is "?merge \ ?cases") proof safe assume ?merge from this \' show ?cases proof (induction \\"merge_at \ t' \'" arbitrary: \ \') case (1 j s' y \'') then show ?case by (cases \) (auto split: if_splits) next case (2 j x \' s' \ \'') then show ?case by (cases \) (auto split: if_splits) qed next assume ?cases then show ?merge proof (split if_split_asm) assume "trace_in ss t x \" "t < t'" from this \' show ?thesis proof induction case 1 then show ?case by (cases \') auto qed auto next assume "\y. trace_in {y} t' x \ \ trace_in ss t y \'" "\ t < t'" then obtain y where "trace_in {y} t' x \" "trace_in ss t y \'" "t' \ t" by auto from this \' show ?thesis by induction auto qed qed lemma AE_lim_wf_times: "AE \ in K.lim_stream (t, s). wf_times t \" using AE_lim_stream proof eventually_elim fix \ assume *: "\i. snd (((t, s) ## \) !! i) \ DTMC.acc `` {snd (t, s)} \ snd (\ !! i) \ J (snd (((t, s) ## \) !! i)) \ fst (((t, s) ## \) !! i) < fst (\ !! i)" have "(t ## smap fst \) !! i < fst (\ !! i)" for i using *[THEN spec, of i] by (cases i) auto then show "wf_times t \" proof (coinduction arbitrary: t \) case wf_times from this[THEN spec, of 0] this[THEN spec, of "Suc i" for i] show ?case by (cases \) auto qed qed lemma wf_times_shiftD: "wf_times t' (smap (\(t', y). (t' + t, y)) \) \ wf_times (t' - t) \" apply (coinduction arbitrary: t' t \) subgoal for t' t \ apply (cases \; cases "shd \") apply (auto simp: ) subgoal for \' j x by (rule exI[of _ "j + t"]) auto done done lemma wf_times_shift[simp]: "wf_times t' (smap (\(t', y). (t' + t, y)) \) = wf_times (t' - t) \" using wf_times_shiftD[of "t' - t" "-t" "smap (\(t', y). (t' + t, y)) \"] by (auto simp: stream.map_comp stream.case_eq_if prod.case_eq_if wf_times_shiftD) lemma trace_in_unique: "trace_in {y1} t x \ \ trace_in {y2} t x \ \ y1 = y2" by (induction rule: trace_in.induct) auto lemma trace_at_eq: "trace_in {z} t x \ \ trace_at x \ t = z" by (induction rule: trace_in.induct) auto lemma AE_lim_acc: "AE \ in K.lim_stream (t, x). \t z. trace_in {z} t x \ \ (x, z) \ DTMC.acc" using AE_lim_stream proof (eventually_elim, safe) fix t' z \ assume *: "\i. snd (((t, x) ## \) !! i) \ DTMC.acc `` {snd (t, x)} \ snd (\ !! i) \ J (snd (((t, x) ## \) !! i)) \ fst (((t, x) ## \) !! i) < fst (\ !! i)" and t: "trace_in {z} t' x \" define X where "X = DTMC.acc `` {x}" have "(x ## smap snd \) !! i \ X" for i using *[THEN spec, of i] by (cases i) (auto simp: X_def) from t this have "z \ X" proof induction case (1 j y x \) with "1.prems"[of 0] show ?case by simp next case (2 j y \ x) with "2.prems"[of "Suc i" for i] show ?case by simp qed then show "(x, z) \ DTMC.acc" by (simp add: X_def) qed lemma p_add: assumes "0 \ t" "0 \ t'" shows "p x y (t + t') = (LINT z|count_space (DTMC.acc``{x}). p x z t * p z y t')" proof - interpret L: prob_space "K.lim_stream xy" for xy by (rule K.prob_space_lim_stream) simp interpret A: sigma_finite_measure "count_space (DTMC.acc``{x})" by (intro sigma_finite_measure_count_space_countable DTMC.countable_acc) simp interpret LA: pair_sigma_finite "count_space (DTMC.acc``{x})" "K.lim_stream xy" for xy by unfold_locales have "p x y (t + t') = (\\<^sup>+ \. \\<^sup>+\'. indicator {\\space T. trace_in {y} (t + t') x \} (merge_at \ t \') \K.lim_stream (t, trace_at x \ t) \K.lim_stream (0, x))" unfolding p_def L.emeasure_eq_measure[symmetric] apply (subst lim_time_split[OF \0 \ t\]) apply (subst emeasure_bind[OF lim_stream_not_empty measurable_prob_algebraD]) apply (measurable; fail) apply (measurable; fail) apply (intro nn_integral_cong) apply (subst emeasure_bind[OF lim_stream_not_empty measurable_prob_algebraD]) apply (measurable; fail) apply (measurable; fail) apply (simp add: in_space_lim_stream) done also have "\ = (\\<^sup>+ \. \\<^sup>+\'. indicator {\\space T. trace_in {y} (t + t') x \} (merge_at \ t (smap (\(t'', s). (t'' + t, s)) \')) \K.lim_stream (0, trace_at x \ t) \K.lim_stream (0, x))" unfolding lim_0[of t] by (subst nn_integral_distr) (measurable; fail)+ also have "\ = (\\<^sup>+ \. \\<^sup>+\'. of_bool (\z\DTMC.acc``{x}. trace_in {z} t x \ \ trace_in {y} t' z \') \K.lim_stream (0, trace_at x \ t) \K.lim_stream (0, x))" apply (rule nn_integral_cong_AE) using AE_lim_wf_times AE_lim_acc apply eventually_elim subgoal premises \ for \ apply (rule nn_integral_cong_AE) using AE_lim_wf_times AE_lim_acc apply eventually_elim using \ assms - apply (auto simp add: trace_in_merge_at indicator_def Bex_def) + apply (auto simp add: trace_in_merge_at indicator_eq_1_iff) done done also have "\ = (\\<^sup>+ \. \\<^sup>+\'. \\<^sup>+z. of_bool (trace_in {z} t x \ \ trace_in {y} t' z \') \count_space (DTMC.acc``{x}) \K.lim_stream (0, trace_at x \ t) \K.lim_stream (0, x))" by (intro nn_integral_cong of_bool_Bex_eq_nn_integral) (auto dest: trace_in_unique) also have "\ = (\\<^sup>+ \. \\<^sup>+z. \\<^sup>+\'. of_bool (trace_in {z} t x \ \ trace_in {y} t' z \') \K.lim_stream (0, trace_at x \ t) \count_space (DTMC.acc``{x}) \K.lim_stream (0, x))" apply (subst LA.Fubini') apply (subst measurable_split_conv) apply (rule measurable_compose_countable'[OF _ measurable_fst]) apply (auto simp: DTMC.countable_acc) done also have "\ = (\\<^sup>+z. \\<^sup>+ \. of_bool (trace_in {z} t x \) * \\<^sup>+\'. of_bool (trace_in {y} t' z \') \K.lim_stream (0, z) \K.lim_stream (0, x) \count_space (DTMC.acc``{x}))" apply (subst LA.Fubini') apply (subst measurable_split_conv) apply (rule measurable_compose_countable'[OF _ measurable_fst]) apply (rule nn_integral_measurable_subprob_algebra2) apply (measurable; fail) apply (rule measurable_prob_algebraD) apply (auto simp: DTMC.countable_acc trace_at_eq intro!: nn_integral_cong) done also have "\ = (\\<^sup>+z. (\\<^sup>+ \. of_bool (trace_in {z} t x \)\K.lim_stream (0, x)) * (\\<^sup>+\'. of_bool (trace_in {y} t' z \') \K.lim_stream (0, z)) \count_space (DTMC.acc``{x}))" by (auto intro!: nn_integral_cong simp: nn_integral_multc) also have "\ = (\\<^sup>+z. ennreal (p x z t) * ennreal (p z y t') \count_space (DTMC.acc``{x}))" unfolding p_def L.emeasure_eq_measure[symmetric] by (auto intro!: nn_integral_cong arg_cong2[where f="(*)"] simp: nn_integral_indicator[symmetric] simp del: nn_integral_indicator ) finally have "(\\<^sup>+z. p x z t * p z y t' \count_space (DTMC.acc``{x})) = p x y (t + t')" by (simp add: ennreal_mult) then show ?thesis by (subst (asm) nn_integral_eq_integrable) auto qed end end diff --git a/thys/Native_Word/Code_Target_Word_Base.thy b/thys/Native_Word/Code_Target_Word_Base.thy --- a/thys/Native_Word/Code_Target_Word_Base.thy +++ b/thys/Native_Word/Code_Target_Word_Base.thy @@ -1,393 +1,391 @@ (* Title: Code_Target_Word_Base.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Common base for target language implementations of word types\ theory Code_Target_Word_Base imports "HOL-Library.Word" "Word_Lib.Signed_Division_Word" Bits_Integer begin text \More lemmas\ lemma div_half_nat: fixes x y :: nat assumes "y \ 0" shows "(x div y, x mod y) = (let q = 2 * (x div 2 div y); r = x - q * y in if y \ r then (q + 1, r - y) else (q, r))" proof - let ?q = "2 * (x div 2 div y)" have q: "?q = x div y - x div y mod 2" by(metis div_mult2_eq mult.commute minus_mod_eq_mult_div [symmetric]) let ?r = "x - ?q * y" have r: "?r = x mod y + x div y mod 2 * y" by(simp add: q diff_mult_distrib minus_mod_eq_div_mult [symmetric])(metis diff_diff_cancel mod_less_eq_dividend mod_mult2_eq add.commute mult.commute) show ?thesis proof(cases "y \ x - ?q * y") case True with assms q have "x div y mod 2 \ 0" unfolding r by (metis Nat.add_0_right diff_0_eq_0 diff_Suc_1 le_div_geq mod2_gr_0 mod_div_trivial mult_0 neq0_conv numeral_1_eq_Suc_0 numerals(1)) hence "x div y = ?q + 1" unfolding q by simp moreover hence "x mod y = ?r - y" by simp(metis minus_div_mult_eq_mod [symmetric] diff_commute diff_diff_left mult_Suc) ultimately show ?thesis using True by(simp add: Let_def) next case False hence "x div y mod 2 = 0" unfolding r by(simp add: not_le)(metis Nat.add_0_right assms div_less div_mult_self2 mod_div_trivial mult.commute) hence "x div y = ?q" unfolding q by simp moreover hence "x mod y = ?r" by (metis minus_div_mult_eq_mod [symmetric]) ultimately show ?thesis using False by(simp add: Let_def) qed qed lemma div_half_word: fixes x y :: "'a :: len word" assumes "y \ 0" shows "(x div y, x mod y) = (let q = (x >> 1) div y << 1; r = x - q * y in if y \ r then (q + 1, r - y) else (q, r))" proof - obtain n where n: "x = of_nat n" "n < 2 ^ LENGTH('a)" by (rule that [of \unat x\]) simp_all moreover obtain m where m: "y = of_nat m" "m < 2 ^ LENGTH('a)" by (rule that [of \unat y\]) simp_all ultimately have [simp]: \unat (of_nat n :: 'a word) = n\ \unat (of_nat m :: 'a word) = m\ by (transfer, simp add: take_bit_of_nat take_bit_nat_eq_self_iff)+ let ?q = "(x >> 1) div y << 1" let ?q' = "2 * (n div 2 div m)" have "n div 2 div m < 2 ^ LENGTH('a)" using n by (metis of_nat_inverse unat_lt2p uno_simps(2)) hence q: "?q = of_nat ?q'" using n m by (auto simp add: shiftr_word_eq drop_bit_eq_div shiftl_t2n word_arith_nat_div uno_simps take_bit_nat_eq_self) from assms have "m \ 0" using m by -(rule notI, simp) from n have "2 * (n div 2 div m) < 2 ^ LENGTH('a)" by(metis mult.commute div_mult2_eq minus_mod_eq_mult_div [symmetric] less_imp_diff_less of_nat_inverse unat_lt2p uno_simps(2)) moreover have "2 * (n div 2 div m) * m < 2 ^ LENGTH('a)" using n unfolding div_mult2_eq[symmetric] by(subst (2) mult.commute)(simp add: minus_mod_eq_div_mult [symmetric] diff_mult_distrib minus_mod_eq_mult_div [symmetric] div_mult2_eq) moreover have "2 * (n div 2 div m) * m \ n" by (simp flip: div_mult2_eq ac_simps) ultimately have r: "x - ?q * y = of_nat (n - ?q' * m)" and "y \ x - ?q * y \ of_nat (n - ?q' * m) - y = of_nat (n - ?q' * m - m)" using n m unfolding q apply (simp_all add: of_nat_diff) apply (subst of_nat_diff) apply (simp_all add: word_le_nat_alt take_bit_nat_eq_self unat_sub_if' unat_word_ariths) - apply (cases \2 \ LENGTH('a)\) - apply (simp_all add: unat_word_ariths take_bit_nat_eq_self) done then show ?thesis using n m div_half_nat [OF \m \ 0\, of n] unfolding q by (simp add: word_le_nat_alt word_div_def word_mod_def Let_def take_bit_nat_eq_self flip: zdiv_int zmod_int split del: if_split split: if_split_asm) qed lemma word_test_bit_set_bits: "(BITS n. f n :: 'a :: len word) !! n \ n < LENGTH('a) \ f n" by (simp add: test_bit_eq_bit bit_set_bits_word_iff) lemma word_of_int_conv_set_bits: "word_of_int i = (BITS n. i !! n)" by (rule word_eqI) (auto simp add: word_test_bit_set_bits) lemma word_and_mask_or_conv_and_mask: "n !! index \ (n AND mask index) OR (1 << index) = n AND mask (index + 1)" for n :: \'a::len word\ by(rule word_eqI)(auto simp add: word_ao_nth word_size nth_shiftl simp del: shiftl_1) lemma uint_and_mask_or_full: fixes n :: "'a :: len word" assumes "n !! (LENGTH('a) - 1)" and "mask1 = mask (LENGTH('a) - 1)" and "mask2 = 1 << LENGTH('a) - 1" shows "uint (n AND mask1) OR mask2 = uint n" proof - have "mask2 = uint (1 << LENGTH('a) - 1 :: 'a word)" using assms by (simp add: uint_shiftl word_size bintrunc_shiftl del: shiftl_1) hence "uint (n AND mask1) OR mask2 = uint (n AND mask1 OR (1 << LENGTH('a) - 1 :: 'a word))" by(simp add: uint_or) also have "\ = uint (n AND mask (LENGTH('a) - 1 + 1))" using assms by(simp only: word_and_mask_or_conv_and_mask) also have "\ = uint n" by simp finally show ?thesis . qed text \Division on @{typ "'a word"} is unsigned, but Scala and OCaml only have signed division and modulus.\ lemmas word_sdiv_def = sdiv_word_def lemmas word_smod_def = smod_word_def lemma [code]: "x sdiv y = (let x' = sint x; y' = sint y; negative = (x' < 0) \ (y' < 0); result = abs x' div abs y' in word_of_int (if negative then -result else result))" for x y :: \'a::len word\ by (simp add: sdiv_word_def signed_divide_int_def sgn_if Let_def not_less not_le) lemma [code]: "x smod y = (let x' = sint x; y' = sint y; negative = (x' < 0); result = abs x' mod abs y' in word_of_int (if negative then -result else result))" for x y :: \'a::len word\ proof - have *: \k mod l = k - k div l * l\ for k l :: int by (simp add: minus_div_mult_eq_mod) show ?thesis by (simp add: smod_word_def signed_modulo_int_def signed_divide_int_def * sgn_if Let_def) qed text \ This algorithm implements unsigned division in terms of signed division. Taken from Hacker's Delight. \ lemma divmod_via_sdivmod: fixes x y :: "'a :: len word" assumes "y \ 0" shows "(x div y, x mod y) = (if 1 << (LENGTH('a) - 1) \ y then if x < y then (0, x) else (1, x - y) else let q = ((x >> 1) sdiv y) << 1; r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" proof(cases "1 << (LENGTH('a) - 1) \ y") case True note y = this show ?thesis proof(cases "x < y") case True then have "x mod y = x" by transfer simp thus ?thesis using True y by(simp add: word_div_lt_eq_0) next case False obtain n where n: "y = of_nat n" "n < 2 ^ LENGTH('a)" by (rule that [of \unat y\]) simp_all have "unat x < 2 ^ LENGTH('a)" by(rule unat_lt2p) also have "\ = 2 * 2 ^ (LENGTH('a) - 1)" by(metis Suc_pred len_gt_0 power_Suc One_nat_def) also have "\ \ 2 * n" using y n by transfer (simp add: push_bit_of_1 take_bit_eq_mod) finally have div: "x div of_nat n = 1" using False n by (simp add: word_div_eq_1_iff take_bit_nat_eq_self) moreover have "x mod y = x - x div y * y" by (simp add: minus_div_mult_eq_mod) with div n have "x mod y = x - y" by simp ultimately show ?thesis using False y n by simp qed next case False note y = this obtain n where n: "x = of_nat n" "n < 2 ^ LENGTH('a)" by (rule that [of \unat x\]) simp_all hence "int n div 2 + 2 ^ (LENGTH('a) - Suc 0) < 2 ^ LENGTH('a)" by (cases \LENGTH('a)\) (auto dest: less_imp_of_nat_less [where ?'a = int]) with y n have "sint (x >> 1) = uint (x >> 1)" by (simp add: sint_uint sbintrunc_mod2p shiftr_div_2n take_bit_nat_eq_self) moreover have "uint y + 2 ^ (LENGTH('a) - Suc 0) < 2 ^ LENGTH('a)" using y by (cases "LENGTH('a)") (simp_all add: not_le word_2p_lem word_size) then have "sint y = uint y" by (simp add: sint_uint sbintrunc_mod2p) ultimately show ?thesis using y apply (subst div_half_word [OF assms]) apply (simp add: sdiv_word_def signed_divide_int_def flip: uint_div) done qed text \More implementations tailored towards target-language implementations\ context includes integer.lifting begin lift_definition word_of_integer :: "integer \ 'a :: len word" is word_of_int . lemma word_of_integer_code [code]: "word_of_integer n = word_of_int (int_of_integer n)" by(simp add: word_of_integer.rep_eq) end lemma word_of_int_code: "uint (word_of_int x :: 'a word) = x AND mask (LENGTH('a :: len))" by (simp add: take_bit_eq_mask) context fixes f :: "nat \ bool" begin definition set_bits_aux :: \'a word \ nat \ 'a :: len word\ where \set_bits_aux w n = push_bit n w OR take_bit n (set_bits f)\ lemma set_bits_aux_conv: \set_bits_aux w n = (w << n) OR (set_bits f AND mask n)\ for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: set_bits_aux_def shiftl_word_eq bit_and_iff bit_or_iff bit_push_bit_iff bit_take_bit_iff bit_mask_iff bit_set_bits_word_iff) corollary set_bits_conv_set_bits_aux: \set_bits f = (set_bits_aux 0 (LENGTH('a)) :: 'a :: len word)\ by (simp add: set_bits_aux_conv) lemma set_bits_aux_0 [simp]: \set_bits_aux w 0 = w\ by (simp add: set_bits_aux_conv) lemma set_bits_aux_Suc [simp]: \set_bits_aux w (Suc n) = set_bits_aux ((w << 1) OR (if f n then 1 else 0)) n\ by (simp add: set_bits_aux_def shiftl_word_eq bit_eq_iff bit_or_iff bit_push_bit_iff bit_take_bit_iff bit_set_bits_word_iff) (auto simp add: bit_exp_iff not_less bit_1_iff less_Suc_eq_le) lemma set_bits_aux_simps [code]: \set_bits_aux w 0 = w\ \set_bits_aux w (Suc n) = set_bits_aux ((w << 1) OR (if f n then 1 else 0)) n\ by simp_all end lemma word_of_int_via_signed: fixes mask assumes mask_def: "mask = Bit_Operations.mask (LENGTH('a))" and shift_def: "shift = 1 << LENGTH('a)" and index_def: "index = LENGTH('a) - 1" and overflow_def:"overflow = 1 << (LENGTH('a) - 1)" and least_def: "least = - overflow" shows "(word_of_int i :: 'a :: len word) = (let i' = i AND mask in if i' !! index then if i' - shift < least \ overflow \ i' - shift then arbitrary1 i' else word_of_int (i' - shift) else if i' < least \ overflow \ i' then arbitrary2 i' else word_of_int i')" proof - define i' where "i' = i AND mask" have "shift = mask + 1" unfolding assms by(simp add: bin_mask_p1_conv_shift) hence "i' < shift" by(simp add: mask_def i'_def int_and_le) show ?thesis proof(cases "i' !! index") case True then have unf: "i' = overflow OR i'" apply (simp add: assms i'_def shiftl_eq_push_bit push_bit_of_1 flip: take_bit_eq_mask) apply (rule bit_eqI) apply (auto simp add: bit_take_bit_iff bit_or_iff bit_exp_iff) done have "overflow \ i'" by(subst unf)(rule le_int_or, simp add: bin_sign_and assms i'_def) hence "i' - shift < least \ False" unfolding assms by(cases "LENGTH('a)")(simp_all add: not_less) moreover have "overflow \ i' - shift \ False" using \i' < shift\ unfolding assms by(cases "LENGTH('a)")(auto simp add: not_le elim: less_le_trans) moreover have "word_of_int (i' - shift) = (word_of_int i :: 'a word)" using \i' < shift\ by (simp add: i'_def shift_def mask_def shiftl_eq_push_bit push_bit_of_1 word_of_int_eq_iff flip: take_bit_eq_mask) ultimately show ?thesis using True by(simp add: Let_def i'_def) next case False hence "i' = i AND Bit_Operations.mask (LENGTH('a) - 1)" unfolding assms i'_def by(clarsimp simp add: i'_def bin_nth_ops intro!: bin_eqI)(cases "LENGTH('a)", auto simp add: less_Suc_eq) also have "\ \ Bit_Operations.mask (LENGTH('a) - 1)" by(rule int_and_le) simp also have "\ < overflow" unfolding overflow_def by(simp add: bin_mask_p1_conv_shift[symmetric]) also have "least \ 0" unfolding least_def overflow_def by simp have "0 \ i'" by (simp add: i'_def mask_def) hence "least \ i'" using \least \ 0\ by simp moreover have "word_of_int i' = (word_of_int i :: 'a word)" by(rule word_eqI)(auto simp add: i'_def bin_nth_ops mask_def) ultimately show ?thesis using False by(simp add: Let_def i'_def) qed qed text \Quickcheck conversion functions\ context includes state_combinator_syntax begin definition qc_random_cnv :: "(natural \ 'a::term_of) \ natural \ Random.seed \ ('a \ (unit \ Code_Evaluation.term)) \ Random.seed" where "qc_random_cnv a_of_natural i = Random.range (i + 1) \\ (\k. Pair ( let n = a_of_natural k in (n, \_. Code_Evaluation.term_of n)))" end definition qc_exhaustive_cnv :: "(natural \ 'a) \ ('a \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "qc_exhaustive_cnv a_of_natural f d = Quickcheck_Exhaustive.exhaustive (%x. f (a_of_natural x)) d" definition qc_full_exhaustive_cnv :: "(natural \ ('a::term_of)) \ ('a \ (unit \ term) \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "qc_full_exhaustive_cnv a_of_natural f d = Quickcheck_Exhaustive.full_exhaustive (%(x, xt). f (a_of_natural x, %_. Code_Evaluation.term_of (a_of_natural x))) d" declare [[quickcheck_narrowing_ghc_options = "-XTypeSynonymInstances"]] definition qc_narrowing_drawn_from :: "'a list \ integer \ _" where "qc_narrowing_drawn_from xs = foldr Quickcheck_Narrowing.sum (map Quickcheck_Narrowing.cons (butlast xs)) (Quickcheck_Narrowing.cons (last xs))" locale quickcheck_narrowing_samples = fixes a_of_integer :: "integer \ 'a \ 'a :: {partial_term_of, term_of}" and zero :: "'a" and tr :: "typerep" begin function narrowing_samples :: "integer \ 'a list" where "narrowing_samples i = (if i > 0 then let (a, a') = a_of_integer i in narrowing_samples (i - 1) @ [a, a'] else [zero])" by pat_completeness auto termination including integer.lifting proof(relation "measure nat_of_integer") fix i :: integer assume "0 < i" thus "(i - 1, i) \ measure nat_of_integer" by simp(transfer, simp) qed simp definition partial_term_of_sample :: "integer \ 'a" where "partial_term_of_sample i = (if i < 0 then undefined else if i = 0 then zero else if i mod 2 = 0 then snd (a_of_integer (i div 2)) else fst (a_of_integer (i div 2 + 1)))" lemma partial_term_of_code: "partial_term_of (ty :: 'a itself) (Quickcheck_Narrowing.Narrowing_variable p t) \ Code_Evaluation.Free (STR ''_'') tr" "partial_term_of (ty :: 'a itself) (Quickcheck_Narrowing.Narrowing_constructor i []) \ Code_Evaluation.term_of (partial_term_of_sample i)" by (rule partial_term_of_anything)+ end lemmas [code] = quickcheck_narrowing_samples.narrowing_samples.simps quickcheck_narrowing_samples.partial_term_of_sample_def text \ The separate code target \SML_word\ collects setups for the code generator that PolyML does not provide. \ setup \Code_Target.add_derived_target ("SML_word", [(Code_ML.target_SML, I)])\ code_identifier code_module Code_Target_Word_Base \ (SML) Word and (Haskell) Word and (OCaml) Word and (Scala) Word end diff --git a/thys/Probabilistic_Prime_Tests/Fermat_Witness.thy b/thys/Probabilistic_Prime_Tests/Fermat_Witness.thy --- a/thys/Probabilistic_Prime_Tests/Fermat_Witness.thy +++ b/thys/Probabilistic_Prime_Tests/Fermat_Witness.thy @@ -1,752 +1,752 @@ (* File: Fermat_Witness.thy Authors: Daniel Stüwe Fermat witnesses as used in Fermat's test *) section \Fermat Witnesses\ theory Fermat_Witness imports Euler_Witness Carmichael_Numbers begin definition divide_out :: "'a :: factorial_semiring \ 'a \ 'a \ nat" where "divide_out p x = (x div p ^ multiplicity p x, multiplicity p x)" lemma fst_divide_out [simp]: "fst (divide_out p x) = x div p ^ multiplicity p x" and snd_divide_out [simp]: "snd (divide_out p x) = multiplicity p x" by (simp_all add: divide_out_def) function divide_out_aux :: "'a :: factorial_semiring \ 'a \ nat \ 'a \ nat" where "divide_out_aux p (x, acc) = (if x = 0 \ is_unit p \ \p dvd x then (x, acc) else divide_out_aux p (x div p, acc + 1))" by auto termination proof (relation "measure (\(p, x, _). multiplicity p x)") fix p x :: 'a and acc :: nat assume "\(x = 0 \ is_unit p \ \p dvd x)" thus "((p, x div p, acc + 1), p, x, acc) \ measure (\(p, x, _). multiplicity p x)" by (auto elim!: dvdE simp: multiplicity_times_same) qed auto lemmas [simp del] = divide_out_aux.simps lemma divide_out_aux_correct: "divide_out_aux p z = (fst z div p ^ multiplicity p (fst z), snd z + multiplicity p (fst z))" proof (induction p z rule: divide_out_aux.induct) case (1 p x acc) show ?case proof (cases "x = 0 \ is_unit p \ \p dvd x") case False have "x div p div p ^ multiplicity p (x div p) = x div p ^ multiplicity p x" using False by (subst dvd_div_mult2_eq [symmetric]) (auto elim!: dvdE simp: multiplicity_dvd multiplicity_times_same) with False show ?thesis using 1 by (subst divide_out_aux.simps) (auto elim: dvdE simp: multiplicity_times_same multiplicity_unit_left not_dvd_imp_multiplicity_0) qed (auto simp: divide_out_aux.simps multiplicity_unit_left not_dvd_imp_multiplicity_0) qed lemma divide_out_code [code]: "divide_out p x = divide_out_aux p (x, 0)" by (simp add: divide_out_aux_correct divide_out_def) lemma multiplicity_code [code]: "multiplicity p x = snd (divide_out_aux p (x, 0))" by (simp add: divide_out_aux_correct) (* TODO Move *) lemma multiplicity_times_same_power: assumes "x \ 0" "\is_unit p" "p \ 0" shows "multiplicity p (p ^ k * x) = multiplicity p x + k" using assms by (induction k) (simp_all add: multiplicity_times_same mult.assoc) lemma divide_out_unique_nat: fixes n :: nat assumes "\is_unit p" "p \ 0" "\p dvd m" and "n = p ^ k * m" shows "m = n div p ^ multiplicity p n" and "k = multiplicity p n" proof - from assms have "n \ 0" by (intro notI) auto thus *: "k = multiplicity p n" using assms by (auto simp: assms multiplicity_times_same_power not_dvd_imp_multiplicity_0) from assms show "m = n div p ^ multiplicity p n" unfolding * [symmetric] by auto qed context fixes a n :: nat begin definition "fermat_liar \ [a^(n-1) = 1] (mod n)" definition "fermat_witness \ [a^(n-1) \ 1] (mod n)" definition "strong_fermat_liar \ (\k m. odd m \ n - 1 = 2^k * m \ [a^m = 1](mod n) \ (\i \ {0..< k}. [a^(2 ^ i * m) = n-1] (mod n)))" definition "strong_fermat_witness \ \ strong_fermat_liar" lemma fermat_liar_witness_of_composition[iff]: "\fermat_liar = fermat_witness" "\fermat_witness = fermat_liar" unfolding fermat_liar_def and fermat_witness_def by simp_all lemma strong_fermat_liar_code [code]: "strong_fermat_liar \ (let (m, k) = divide_out 2 (n - 1) in [a^m = 1](mod n) \ (\i \ {0..< k}. [a^(2 ^ i * m) = n-1] (mod n)))" (is "?lhs = ?rhs") proof (cases "n > 1") case True define m where "m = (n - 1) div 2 ^ multiplicity 2 (n - 1)" define k where "k = multiplicity 2 (n - 1)" have mk: "odd m \ n - 1 = 2 ^ k * m" using True multiplicity_decompose[of "n - 1" 2] multiplicity_dvd[of 2 "n - 1"] by (auto simp: m_def k_def) show ?thesis proof assume ?lhs hence "[a^m = 1] (mod n) \ (\i \ {0..< k}. [a^(2 ^ i * m) = n-1] (mod n))" using True mk by (auto simp: divide_out_def strong_fermat_liar_def) thus ?rhs by (auto simp: Let_def divide_out_def m_def k_def) next assume ?rhs thus ?lhs using divide_out_unique_nat[of 2] by (auto simp: strong_fermat_liar_def divide_out_def) qed qed (auto simp: strong_fermat_liar_def divide_out_def) context assumes * : "a \ {1 ..< n}" begin lemma strong_fermat_witness_iff: "strong_fermat_witness = (\k m. odd m \ n - 1 = 2 ^ k * m \ [a ^ m \ 1] (mod n) \ (\i\{0.. n-1] (mod n)))" unfolding strong_fermat_witness_def strong_fermat_liar_def by blast lemma not_coprime_imp_witness: "\coprime a n \ fermat_witness" using * lucas_coprime_lemma[of "n-1" a n] by (auto simp: fermat_witness_def) corollary liar_imp_coprime: "fermat_liar \ coprime a n" using not_coprime_imp_witness fermat_liar_witness_of_composition by blast lemma fermat_witness_imp_strong_fermat_witness: assumes "1 < n" "fermat_witness" shows "strong_fermat_witness" proof - from \fermat_witness\ have "[a^(n-1) \ 1] (mod n)" unfolding fermat_witness_def . obtain k m where "odd m" and n: "n - 1 = 2 ^ k * m" using * by (auto intro: multiplicity_decompose'[of "(n-1)" 2]) moreover have "[a ^ m \ 1] (mod n)" using \[a^(n-1) \ 1] (mod n)\ n ord_divides by auto moreover { fix i :: nat assume "i \ {0.. k - 1" "0 < k" by auto then have "[a ^ (2 ^ i * m) \ n - 1] (mod n)" "[a ^ (2 ^ i * m) \ 1] (mod n)" proof(induction i rule: inc_induct) case base have * :"a ^ (n - 1) = (a ^ (2 ^ (k - 1) * m))\<^sup>2" using \0 < k\ n semiring_normalization_rules(27)[of "2 :: nat" "k - 1"] by (auto simp flip: power_even_eq simp add: algebra_simps) case 1 from * show ?case using \[a^(n-1) \ 1] (mod n)\ and square_minus_one_cong_one[OF \1 < n\] by auto case 2 from * show ?case using n \[a^(n-1) \ 1] (mod n)\ and square_one_cong_one by metis next case (step q) then have IH2: "[a ^ (2 ^ Suc q * m) \ 1] (mod n)" using \0 < k\ by blast+ have * : "a ^ (2 ^ Suc q * m) = (a ^ (2 ^ q * m))\<^sup>2" by (auto simp flip: power_even_eq simp add: algebra_simps) case 1 from * show ?case using IH2 and square_minus_one_cong_one[OF \1 < n\] by auto case 2 from * show ?case using IH2 and square_one_cong_one by metis qed } ultimately show ?thesis unfolding strong_fermat_witness_iff by blast qed corollary strong_fermat_liar_imp_fermat_liar: assumes "1 < n" "strong_fermat_liar" shows "fermat_liar" using fermat_witness_imp_strong_fermat_witness assms and fermat_liar_witness_of_composition strong_fermat_witness_def by blast lemma euler_liar_is_fermat_liar: assumes "1 < n" "euler_liar a n" "coprime a n" "odd n" shows "fermat_liar" proof - have "[Jacobi a n = a ^ ((n - 1) div 2)] (mod n)" using \euler_liar a n\ unfolding euler_witness_def by simp hence "[(Jacobi a n)^2 = (a ^ ((n - 1) div 2))^2] (mod n)" by (simp add: cong_pow) moreover have "Jacobi a n \ {1, -1}" using Jacobi_values Jacobi_eq_0_iff_not_coprime[of n] \coprime a n\ \1 < n\ by force ultimately have "[1 = (a ^ ((n - 1) div 2))^2] (mod n)" using cong_int_iff by force also have "(a ^ ((n - 1) div 2))^2 = a ^ (n - 1)" using \odd n\ by (simp flip: power_mult) finally show ?thesis using cong_sym fermat_liar_def by blast qed corollary fermat_witness_is_euler_witness: assumes "1 < n" "fermat_witness" "coprime a n" "odd n" shows "euler_witness a n" using assms euler_liar_is_fermat_liar fermat_liar_witness_of_composition by blast end end lemma one_is_fermat_liar[simp]: "1 < n \ fermat_liar 1 n" using fermat_liar_def by auto lemma one_is_strong_fermat_liar[simp]: "1 < n \ strong_fermat_liar 1 n" using strong_fermat_liar_def by auto lemma prime_imp_fermat_liar: "prime p \ a \ {1 ..< p} \ fermat_liar a p" unfolding fermat_liar_def using fermat_theorem and nat_dvd_not_less by simp lemma not_Carmichael_numberD: assumes "\Carmichael_number n" "\prime n" and "1 < n" shows "\ a \ {2 ..< n} . fermat_witness a n \ coprime a n" proof - obtain a where "coprime a n" "[a^(n-1) \ 1] (mod n)" using assms unfolding Carmichael_number_def by blast moreover from this and \1 < n\ have "a mod n \ {1.. {1 ..< n}" "coprime (a mod n) n" "[(a mod n)^(n-1) \ 1] (mod n)" using \1 < n\ by simp_all hence "fermat_witness (a mod n) n" using fermat_witness_def by blast hence "1 \ (a mod n)" using \1 < n\ \(a mod n) \ {1 ..< n}\ and one_is_fermat_liar fermat_liar_witness_of_composition(1) by metis thus ?thesis using \fermat_witness (a mod n) n\ \coprime (a mod n) n\ \(a mod n) \ {1.. by fastforce qed proposition prime_imp_strong_fermat_witness: fixes p :: nat assumes "prime p" "2 < p" "a \ {1 ..< p}" shows "strong_fermat_liar a p" proof - { fix k m :: nat define j where "j \ LEAST k. [a ^ (2^k * m) = 1] (mod p)" have "[a ^ (p - 1) = 1] (mod p)" using fermat_theorem[OF \prime p\, of a] \a \ {1 ..< p}\ by force moreover assume "odd m" and n: "p - 1 = 2 ^ k * m" ultimately have "[a ^ (2 ^ k * m) = 1] (mod p)" by simp moreover assume "[a ^ m \ 1] (mod p)" then have "0 < x" if "[a ^ (2 ^ x * m) = 1] (mod p)" for x using that by (auto intro: gr0I) ultimately have "0 < j" "j \ k" "[a ^ (2 ^ j * m) = 1] (mod p)" using LeastI2[of _ k "(<) 0"] Least_le[of _ k] LeastI[of _ k] by (simp_all add: j_def) moreover from this have "[a ^ (2^(j-1) * m) \ 1] (mod p)" using not_less_Least[of "j - 1" "\k. [a ^ (2^k * m) = 1] (mod p)"] unfolding j_def by simp moreover have "a ^ (2 ^ (j - 1) * m) * a ^ (2 ^ (j - 1) * m) = a ^ (2 ^ j * m)" using \0 < j\ by (simp add: algebra_simps semiring_normalization_rules(27) flip: power2_eq_square power_even_eq) ultimately have "(j-1)\{0..prime p\, of "a ^ (2 ^ (j-1) * m)"] by simp_all } then show ?thesis unfolding strong_fermat_liar_def by blast qed lemma ignore_one: fixes P :: "_ \ nat \ bool" assumes "P 1 n" "1 < n" shows "card {a \ {1.. a \ a < n \ P a n}" proof - have "insert 1 {a. 2 \ a \ a < n \ P a n} = {a. 1 \ a \ a < n \ P a n}" using assms by auto moreover have "card (insert 1 {a. 2 \ a \ a < n \ P a n}) = Suc (card {a. 2 \ a \ a < n \ P a n})" using card_insert_disjoint by auto ultimately show ?thesis by force qed text \Proofs in the following section are inspired by \cite{Cornwell, MillerRabinTest, lecture_notes}.\ proposition not_Carmichael_number_imp_card_fermat_witness_bound: fixes n :: nat assumes "\prime n" "\Carmichael_number n" "odd n" "1 < n" shows "(n-1) div 2 < card {a \ {1 ..< n}. fermat_witness a n}" and "(card {a. 2 \ a \ a < n \ strong_fermat_liar a n}) < real (n - 2) / 2" and "(card {a. 2 \ a \ a < n \ fermat_liar a n}) < real (n - 2) / 2" proof - define G where "G = Residues_Mult n" interpret residues_mult_nat n G by unfold_locales (use \n > 1\ in \simp_all only: G_def\) define h where "h \ \a. a ^ (n - 1) mod n" define ker where "ker = kernel G (G\carrier := h ` carrier G\) h" have "finite ker" by (auto simp: ker_def kernel_def) moreover have "1 \ ker" using \n > 1\ by (auto simp: ker_def kernel_def h_def) ultimately have [simp]: "card ker > 0" - by (subst card_gt_0_iff) (auto simp: ker_def kernel_def h_def) + by (subst card_gt_0_iff) auto have totatives_eq: "totatives n = {k\{1..n > 1\ by (force simp: totatives_def) have ker_altdef: "ker = {a \ {1..n > 1\ by (force simp: h_def cong_def intro: coprimeI_power_mod) have h_is_hom: "h \ hom G G" unfolding hom_def using nat_pow_closed by (auto simp: h_def power_mult_distrib mod_simps) then interpret h: group_hom G G h by unfold_locales obtain a where a: "a \ {2.. 1" using a by (auto simp: fermat_witness_def cong_def h_def) hence "2 \ card {1, h a}" by simp also have "\ \ card (h ` carrier G)" proof (intro card_mono; safe?) from \n > 1\ have "1 = h 1" by (simp add: h_def) also have "\ \ h ` carrier G" by (intro imageI) (use \n > 1\ in auto) finally show "1 \ h ` carrier G" . next show "h a \ h ` carrier G" using a by (intro imageI) (auto simp: totatives_def) qed auto also have "\ * card ker = order G" using homomorphism_thm_order[OF h.group_hom_axioms] by (simp add: ker_def order_def) also have "order G < n - 1" using totient_less_not_prime[of n] assms by (simp add: order_eq) finally have "card ker < (n - 1) div 2" using \odd n\ by (auto elim!: oddE) have "(n - 1) div 2 < (n - 1) - card ker" using \card ker < (n - 1) div 2\ by linarith also have "\ = card ({1.. {1.. {1.. a \ a < n \ fermat_liar a n} \ card (ker - {1})" by (intro card_mono) (auto simp: ker_altdef fermat_liar_def fermat_witness_def) also have "\ = card ker - 1" using \n > 1\ by (subst card_Diff_subset) (auto simp: ker_altdef fermat_liar_def) also have "\ < (n - 2) div 2" using \card ker < (n - 1) div 2\ \odd n\ \card ker > 0\ by linarith finally show *: "card {a. 2 \ a \ a < n \ fermat_liar a n} < real (n - 2) / 2" by simp have "card {a. 2 \ a \ a < n \ strong_fermat_liar a n} \ card {a. 2 \ a \ a < n \ fermat_liar a n}" by (intro card_mono) (auto intro!: strong_fermat_liar_imp_fermat_liar) moreover note * ultimately show "card {a. 2 \ a \ a < n \ strong_fermat_liar a n} < real (n - 2) / 2" by simp qed proposition Carmichael_number_imp_lower_bound_on_strong_fermat_witness: fixes n :: nat assumes Carmichael_number: "Carmichael_number n" shows "(n - 1) div 2 < card {a \ {1.. a \ a < n \ strong_fermat_liar a n}) < real (n - 2) / 2" proof - from assms have "n > 3" by (intro Carmichael_number_gt_3) hence "n - 1 \ 0" "\is_unit (2 :: nat)" by auto obtain k m where "odd m" and n_less: "n - 1 = 2 ^ k * m" using multiplicity_decompose'[OF \n - 1 \ 0\ \\is_unit (2::nat)\] by metis obtain p l where n: "n = p * l" and "prime p" "\ p dvd l" "2 < l" using Carmichael_number_imp_squarefree_alt[OF Carmichael_number] by blast then have "coprime p l" using prime_imp_coprime_nat by blast have "odd n" using Carmichael_number_odd Carmichael_number by simp have "2 < n" using \n > 3\ \odd n\ by presburger note prime_gt_1_nat[OF \prime p\] have "2 < p" using \odd n\ n \prime p\ prime_ge_2_nat and dvd_triv_left le_neq_implies_less by blast let ?P = "\ k. (\ a. coprime a p \ [a^(2^k * m) = 1] (mod p))" define j where "j \ LEAST k. ?P k" define H where "H \ {a \ {1.. ([a^(2^(j-1) * m) = 1] (mod n) \ [a^(2^(j-1) * m) = n - 1] (mod n))}" have k : "\a. coprime a n \ [a ^ (2 ^ k * m) = 1] (mod n)" using Carmichael_number unfolding Carmichael_number_def n_less by blast obtain k' m' where "odd m'" and p_less: "p - 1 = 2 ^ k' * m'" using \1 < p\ by (auto intro: multiplicity_decompose'[of "(p-1)" 2]) have "p - 1 dvd n - 1" using Carmichael_number_imp_dvd[OF Carmichael_number \prime p\] \n = p * l\ by fastforce then have "p - 1 dvd 2 ^ k' * m" unfolding n_less p_less using \odd m\ \odd m'\ and coprime_dvd_mult_left_iff[of "2^k'" m "2^k"] coprime_dvd_mult_right_iff[of m' "2^k" m] by auto have k': "\a. coprime a p \ [a ^ (2 ^ k' * m) = 1] (mod p)" proof safe fix a assume "coprime a p" hence "\ p dvd a" using p_coprime_right_nat[OF \prime p\] by simp have "[a ^ (2 ^ k' * m') = 1] (mod p)" unfolding p_less[symmetric] using fermat_theorem \prime p\ \\ p dvd a\ by blast then show "[a ^ (2 ^ k' * m) = 1] (mod p)" using \p - 1 dvd 2 ^ k' * m\ unfolding p_less n_less by (meson dvd_trans ord_divides) qed have j_prop: "[a^(2^j * m) = 1] (mod p)" if "coprime a p" for a using that LeastI[of ?P k', OF k', folded j_def] cong_modulus_mult coprime_mult_right_iff unfolding j_def n by blast have j_least: "[a^(2^i * m) = 1] (mod p)" if "coprime a p" "j \ i" for a i proof - obtain c where i: "i = j + c" using le_iff_add[of j i] \j \ i\ by blast then have "[a ^ (2 ^ i * m) = a ^ (2 ^ (j + c) * m)] (mod p)" by simp also have "[a ^ (2 ^ (j + c) * m) = (a ^ (2 ^ j * m)) ^ (2 ^ c)] (mod p)" by (simp flip: power_mult add: algebra_simps power_add) also note j_prop[OF \coprime a p\] also have "[1 ^ (2 ^ c) = 1] (mod p)" by simp finally show ?thesis . qed have neq_p: "[p - 1 \ 1](mod p)" using \2 < p\ and cong_less_modulus_unique_nat[of "p-1" 1 p] by linarith have "0 < j" proof (rule LeastI2[of ?P k', OF k', folded j_def], rule gr0I) fix x assume "\a. coprime a p \ [a ^ (2 ^ x * m) = 1] (mod p)" then have "[(p - 1) ^ (2 ^ x * m) = 1] (mod p)" using coprime_diff_one_left_nat[of p] prime_gt_1_nat[OF \prime p\] by simp moreover assume "x = 0" hence "[(p-1)^(2^x*m) = p - 1](mod p)" using \odd m\ odd_pow_cong[OF _ \odd m\, of p] prime_gt_1_nat[OF \prime p\] by auto ultimately show False using \[p - 1 \ 1] (mod p)\ by (simp add: cong_def) qed then have "j - 1 < j" by simp then obtain x where "coprime x p" "[x^(2^(j-1) * m) \ 1] (mod p)" using not_less_Least[of "j - 1" ?P, folded j_def] unfolding j_def by blast define G where "G = Residues_Mult n" interpret residues_mult_nat n G by unfold_locales (use \n > 3\ in \simp_all only: G_def\) have H_subset: "H \ carrier G" unfolding H_def by (auto simp: totatives_def) from \n > 3\ have \n > 1\ by simp interpret H: subgroup H G proof (rule subgroupI, goal_cases) case 1 then show ?case using H_subset . next case 2 then show ?case unfolding H_def using \1 < n\ by force next case (3 a) define y where "y = inv\<^bsub>G\<^esub> a" then have "y \ carrier G" using H_subset \a \ H\ by (auto simp del: carrier_eq) then have "1 \ y" "y < n" "coprime y n" using totatives_less[of y n] \n > 3\ by (auto simp: totatives_def) moreover have "[y ^ (2 ^ (j - Suc 0) * m) = Suc 0] (mod n)" if "[y ^ (2 ^ (j - Suc 0) * m) \ n - Suc 0] (mod n)" proof - from \a \ H\ have "[a * y = 1] (mod n)" using H_subset r_inv[of a] y_def by (auto simp: cong_def) hence "[(a * y) ^ (2 ^ (j - 1) * m) = 1 ^ (2 ^ (j - 1) * m)] (mod n)" by (intro cong_pow) hence "[(a * y) ^ (2 ^ (j - 1) * m) = 1] (mod n)" by simp hence * : "[a ^ (2 ^ (j - 1) * m) * y ^ (2 ^ (j - 1) * m) = 1] (mod n)" by (simp add: power_mult_distrib) from \a \ H\ have "1 \ a" "a < n" "coprime a n" unfolding H_def by auto have "[a ^ (2 ^ (j - 1) * m) = 1] (mod n) \ [a ^ (2 ^ (j - 1) * m) = n - 1] (mod n)" using \a \ H\ by (auto simp: H_def) thus ?thesis proof note * also assume "[a ^ (2 ^ (j - 1) * m) = 1] (mod n)" finally show ?thesis by simp next assume "[a ^ (2 ^ (j - 1) * m) = n - 1] (mod n)" then have "[y ^ (2 ^ (j - 1) * m) = n - 1] (mod n)" using minus_one_cong_solve[OF \1 < n\] * \coprime a n\ \coprime y n \coprime_power_left_iff by blast+ thus ?thesis using that by simp qed qed ultimately show ?case using \a \ H\ unfolding H_def y_def by auto next case (4 a b) hence "a \ totatives n" "b \ totatives n" by (auto simp: H_def totatives_def) hence "a * b mod n \ totatives n" using m_closed[of a b] by simp hence "a * b mod n \ {1..n > 3\ by (auto simp: totatives_def) moreover define x y where "x = a ^ (2 ^ (j - 1) * m)" and "y = b ^ (2 ^ (j - 1) * m)" have "[x * y = 1] (mod n) \ [x * y = n - 1] (mod n)" proof - have *: "x mod n \ {1, n - 1}" "y mod n \ {1, n - 1}" using 4 by (auto simp: H_def x_def y_def cong_def) have "[x * y = (x mod n) * (y mod n)] (mod n)" by (intro cong_mult) auto moreover have "((x mod n) * (y mod n)) mod n \ {1, n - 1}" using * square_minus_one_cong_one'[OF \1 < n\] \n > 1\ by (auto simp: cong_def) ultimately show ?thesis using \n > 1\ by (simp add: cong_def mod_simps) qed ultimately show ?case by (auto simp: H_def x_def y_def power_mult_distrib) qed { obtain a where "[a = x] (mod p)" "[a = 1] (mod l)" "a < p * l" using binary_chinese_remainder_unique_nat[of p l x 1] and \\ p dvd l\ \prime p\ prime_imp_coprime_nat by auto moreover have "coprime a p" using \coprime x p\ cong_imp_coprime[OF cong_sym[OF \[a = x] (mod p)\]] coprime_mult_right_iff unfolding n by blast moreover have "coprime a l" using coprime_1_left cong_imp_coprime[OF cong_sym[OF \[a = 1] (mod l)\]] by blast moreover from \prime p\ and \coprime a p\ have "a > 0" by (intro Nat.gr0I) auto ultimately have "a \ carrier G" using \2 < l\ by (auto intro: gre1I_nat simp: n totatives_def) have "[a ^ (2^(j-1) * m) \ 1] (mod p)" using \[x^(2^(j-1) * m) \ 1] (mod p)\ \[a = x] (mod p)\ and cong_trans cong_pow cong_sym by blast then have "[a ^ (2^(j-1) * m) \ 1] (mod n)" using cong_modulus_mult_nat n by fast moreover have "[a ^ (2 ^ (j - Suc 0) * m) \ n - 1] (mod n)" proof - have "[a ^ (2 ^ (j - 1) * m) = 1] (mod l)" using cong_pow[OF \[a = 1] (mod l)\] by auto moreover have "Suc 0 \ (n - Suc 0) mod l" using n \2 < l\ \odd n\ by (metis mod_Suc_eq mod_less mod_mult_self2_is_0 numeral_2_eq_2 odd_Suc_minus_one zero_neq_numeral) then have "[1 \ n - 1] (mod l)" using \2 < l\ \odd n\ unfolding cong_def by simp moreover have "l \ Suc 0" using \2 < l\ by simp ultimately have "[a ^ (2 ^ (j - Suc 0) * m) \ n - 1] (mod l)" by (auto simp add: cong_def n mod_simps dest: cong_modulus_mult_nat) then show ?thesis using cong_modulus_mult_nat mult.commute n by metis qed ultimately have "a \ H" unfolding H_def by auto hence "H \ carrier (G)" using H_subset subgroup.mem_carrier and \a \ carrier (G)\ by fast } have "card H \ order G div 2" by (intro proper_subgroup_imp_bound_on_card) (use \H \ carrier G\ H.is_subgroup in \auto\) also from assms have "\prime n" by (auto dest: Carmichael_number_not_prime) hence "order G div 2 < (n - 1) div 2" using totient_less_not_prime[OF \\ prime n\ \1 < n\] \odd n\ by (auto simp add: order_eq elim!: oddE) finally have "card H < (n - 1) div 2" . { { fix a assume "1 \ a" "a < n" hence "a \ {1.. 1] (mod n)" hence "[a ^ m \ 1] (mod n)" by (metis dvd_trans dvd_triv_right ord_divides) moreover assume "strong_fermat_liar a n" ultimately obtain i where "i \ {0 ..< k}" "[a^(2^i * m) = n-1](mod n)" unfolding strong_fermat_liar_def using \odd m\ n_less by blast then have "[a ^ (2 ^ i * m) = n - 1] (mod p)" unfolding n using cong_modulus_mult_nat by blast moreover have "[n - 1 \ 1] (mod p)" proof(subst cong_altdef_nat, goal_cases) case 1 then show ?case using \1 < n\ by linarith next case 2 have "\ p dvd 2" using \2 < p\ by (simp add: nat_dvd_not_less) moreover have "2 \ n" using \1 < n\ by linarith moreover have "p dvd n" using n by simp ultimately have "\ p dvd n - 2" using dvd_diffD1 by blast then show ?case by (simp add: numeral_2_eq_2) qed ultimately have "[a ^ (2 ^ i * m) \ Suc 0] (mod p)" using cong_sym by (simp add: cong_def) then have "i < j" using j_least[OF \coprime a p\, of i] by force have "[(a ^ (2 ^ Suc i * m)) = 1] (mod n)" using square_minus_one_cong_one[OF \1 < n\ \[a^(2^i * m) = n-1](mod n)\] by (simp add: power2_eq_square power_mult power_mult_distrib) { assume "i < j - 1" have "(2 :: nat) ^ (j - Suc 0) = ((2 ^ i) * 2 ^ (j - Suc i))" unfolding power_add[symmetric] using \i < j - 1\ by simp then have "[a ^ (2 ^ (j - 1) * m) = (a ^ (2 ^ i * m)) ^ (2^(j - 1 - i))] (mod n)" by (auto intro!: cong_pow_I simp flip: power_mult simp add: algebra_simps power_add) also note \[a ^ (2 ^ i * m) = n - 1] (mod n)\ also have "[(n - 1) ^ (2^(j - 1 - i)) = 1] (mod n) " using \1 < n\ \i < j - 1\ using even_pow_cong by auto finally have False using \[a ^ (2 ^ (j - 1) * m) \ 1] (mod n)\ by blast } hence "i = j - 1" using \i < j\ by fastforce hence "[a ^ (2 ^ (j - 1) * m) = n - 1] (mod n)" using \[a^(2^i * m) = n-1](mod n)\ by simp } hence "{a \ {1.. H" using strong_fermat_liar_imp_fermat_liar[of _ n, OF _ \1 < n\] liar_imp_coprime by (auto simp: H_def) } moreover have "finite H" unfolding H_def by auto ultimately have strong_fermat_liar_bounded: "card {a \ {1..card H < (n - 1) div 2\] by blast moreover { have "{1.. {1.. {1.. {1.. {1.. {1.. {1.. a \ a < n \ strong_fermat_liar a n}) < real (n - 2) / 2" using strong_fermat_liar_bounded ignore_one one_is_strong_fermat_liar \1 < n\ by simp qed corollary strong_fermat_witness_lower_bound: assumes "odd n" "n > 2" "\prime n" shows "card {a. 2 \ a \ a < n \ strong_fermat_liar a n} < real (n - 2) / 2" using Carmichael_number_imp_lower_bound_on_strong_fermat_witness(2)[of n] not_Carmichael_number_imp_card_fermat_witness_bound(2)[of n] assms by (cases "Carmichael_number n") auto end \ No newline at end of file diff --git a/thys/Probabilistic_Prime_Tests/Generalized_Primality_Test.thy b/thys/Probabilistic_Prime_Tests/Generalized_Primality_Test.thy --- a/thys/Probabilistic_Prime_Tests/Generalized_Primality_Test.thy +++ b/thys/Probabilistic_Prime_Tests/Generalized_Primality_Test.thy @@ -1,108 +1,108 @@ (* File: Generalized_Primality_Test.thy Authors: Daniel Stüwe, Manuel Eberl Generic probabilistic primality test *) section \A Generic View on Probabilistic Prime Tests\ theory Generalized_Primality_Test imports "HOL-Probability.Probability" Algebraic_Auxiliaries begin definition primality_test :: "(nat \ nat \ bool) \ nat \ bool pmf" where "primality_test P n = (if n < 3 \ even n then return_pmf (n = 2) else do { a \ pmf_of_set {2..< n}; return_pmf (P n a) })" (* TODO Move *) lemma expectation_of_bool_is_pmf: "measure_pmf.expectation M of_bool = pmf M True" by (simp add: integral_measure_pmf_real[where A=UNIV] UNIV_bool) lemma eq_bernoulli_pmfI: assumes "pmf p True = x" shows "p = bernoulli_pmf x" proof (intro pmf_eqI) fix b :: bool from assms have "x \ {0..1}" by (auto simp: pmf_le_1) thus "pmf p b = pmf (bernoulli_pmf x) b" using assms by (cases b) (auto simp: pmf_False_conv_True) qed text \ We require a probabilistic primality test to never classify a prime as composite. It may, however, mistakenly classify composites as primes. \ locale prob_primality_test = fixes P :: "nat \ nat \ bool" and n :: nat assumes P_works: "odd n \ 2 \ a \ a < n \ prime n \ P n a" begin lemma FalseD: assumes false: "False \ set_pmf (primality_test P n)" shows "\ prime n" proof - from false consider "n \ 2" "n < 3" | "n \ 2" "even n" | a where "\ P n a" "odd n" "2 \ a" "a < n" by (auto simp: primality_test_def not_less split: if_splits) then show ?thesis proof cases case 1 then show ?thesis by (cases rule: linorder_neqE_nat) (use prime_ge_2_nat[of n] in auto) next case 2 then show ?thesis using primes_dvd_imp_eq two_is_prime_nat by blast next case 3 then show ?thesis using P_works by blast qed qed theorem prime: assumes odd_prime: "prime n" shows "primality_test P n = return_pmf True" proof - have "set_pmf (primality_test P n) \ {True, False}" by auto moreover from assms have "False \ set_pmf (primality_test P n)" using FalseD by auto ultimately have "set_pmf (primality_test P n) \ {True}" by auto thus ?thesis by (subst (asm) set_pmf_subset_singleton) qed end text \ We call a primality test \q\-good for a fixed positive real number \q\ if the probability that it mistakenly classifies a composite as a prime is less than \q\. \ locale good_prob_primality_test = prob_primality_test + fixes q :: real assumes q_pos: "q > 0" assumes composite_witness_bound: "\prime n \ 2 < n \ odd n \ real (card {a . 2 \ a \ a < n \ P n a}) < q * real (n - 2)" begin lemma composite_aux: assumes "\prime n" shows "measure_pmf.expectation (primality_test P n) of_bool < q" unfolding primality_test_def using assms composite_witness_bound q_pos - by (auto simp: pmf_expectation_bind[where A = "{2..< n}"] sum_of_bool_eq_card field_simps - simp flip: sum_divide_distrib) + by (clarsimp simp add: pmf_expectation_bind[where A = "{2..< n}"] sum_of_bool_eq_card field_simps Int_def + simp flip: sum_divide_distrib) theorem composite: assumes "\prime n" shows "pmf (primality_test P n) True < q" using composite_aux[OF assms] by (simp add: expectation_of_bool_is_pmf) end 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,2196 +1,2194 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson, NICTA *) section \Bit values as reversed lists of bools\ theory Reversed_Bit_Lists imports "HOL-Library.Word" Typedef_Morphisms Least_significant_bit Most_significant_bit Even_More_List "HOL-Library.Sublist" Aligned 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) + by simp (auto simp add: bin_to_bl_aux_alt) qed lemma rbl_succ: "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))" apply (unfold bin_to_bl_def) apply (induction n arbitrary: bin) apply simp_all apply (case_tac bin rule: bin_exhaust) - apply simp - apply (simp add: bin_to_bl_aux_alt ac_simps) + apply (simp_all add: bin_to_bl_aux_alt ac_simps) done lemma rbl_add: "\bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = rev (bin_to_bl n (bina + binb))" apply (unfold bin_to_bl_def) apply (induct n) apply simp apply clarsimp apply (case_tac bina rule: bin_exhaust) apply (case_tac binb rule: bin_exhaust) apply (case_tac b) apply (case_tac [!] "ba") apply (auto simp: rbl_succ bin_to_bl_aux_alt Let_def ac_simps) done lemma rbl_add_long: "m \ n \ rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = rev (bin_to_bl n (bina + binb))" apply (rule box_equals [OF _ rbl_add_take2 rbl_add]) apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong) apply (rule rev_swap [THEN iffD1]) apply (simp add: rev_take drop_bin2bl) apply simp done lemma rbl_mult_gt1: "m \ length bl \ rbl_mult bl (rev (bin_to_bl m binb)) = rbl_mult bl (rev (bin_to_bl (length bl) binb))" apply (rule trans) apply (rule rbl_mult_take2 [symmetric]) apply simp_all apply (rule_tac f = "rbl_mult bl" in arg_cong) apply (rule rev_swap [THEN iffD1]) apply (simp add: rev_take drop_bin2bl) done lemma rbl_mult_gt: "m > n \ rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))" by (auto intro: trans [OF rbl_mult_gt1]) lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt] lemma rbbl_Cons: "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (of_bool b + 2 * x))" by (simp add: bin_to_bl_def) (simp add: bin_to_bl_aux_alt) lemma rbl_mult: "rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = rev (bin_to_bl n (bina * binb))" apply (induct n arbitrary: bina binb) apply simp_all apply (unfold bin_to_bl_def) apply clarsimp apply (case_tac bina rule: bin_exhaust) apply (case_tac binb rule: bin_exhaust) - apply simp - apply (simp add: bin_to_bl_aux_alt) - apply (simp add: rbbl_Cons rbl_mult_Suc rbl_add algebra_simps) + apply (simp_all add: bin_to_bl_aux_alt) + apply (simp_all add: rbbl_Cons rbl_mult_Suc rbl_add algebra_simps) done lemma sclem: "size (concat (map (bin_to_bl n) xs)) = length xs * n" - by (induct xs) auto + by (simp add: length_concat comp_def sum_list_triv) 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) 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]: \Word.the_int (of_bl bs :: 'a word) = horner_sum of_bool 2 (take LENGTH('a::len) (rev bs))\ apply (simp add: of_bl_eq flip: take_bit_horner_sum_bit_eq) apply transfer apply simp done lemma [code]: \to_bl w = map (bit w) (rev [0.. for w :: \'a::len word\ by (fact to_bl_eq_rev) lemma word_reverse_eq_of_bl_rev_to_bl: \word_reverse w = of_bl (rev (to_bl w))\ by (rule bit_word_eqI) (auto simp add: bit_word_reverse_iff bit_of_bl_iff nth_to_bl) lemmas word_reverse_no_def [simp] = word_reverse_eq_of_bl_rev_to_bl [of "numeral w"] for w lemma to_bl_word_rev: "to_bl (word_reverse w) = rev (to_bl w)" by (rule nth_equalityI) (simp_all add: nth_rev_to_bl word_reverse_def word_rep_drop flip: of_bl_eq) lemma to_bl_n1 [simp]: "to_bl (-1::'a::len word) = replicate (LENGTH('a)) True" apply (rule word_bl.Abs_inverse') apply simp apply (rule word_eqI) apply (clarsimp simp add: word_size) apply (auto simp add: word_bl.Abs_inverse test_bit_bl word_size) done lemma rbl_word_or: "rev (to_bl (x OR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_or rev_map) lemma rbl_word_and: "rev (to_bl (x AND y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_and rev_map) lemma rbl_word_xor: "rev (to_bl (x XOR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_xor rev_map) lemma rbl_word_not: "rev (to_bl (NOT x)) = map Not (rev (to_bl x))" by (simp add: bl_word_not rev_map) lemma bshiftr1_numeral [simp]: \bshiftr1 b (numeral w :: 'a word) = of_bl (b # butlast (bin_to_bl LENGTH('a::len) (numeral w)))\ by (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))" apply (auto simp add: to_bl_unfold) apply (rule bit_word_eqI) apply auto done lemma of_drop_to_bl: "of_bl (drop n (to_bl x)) = (x AND mask (size x - n))" by (simp add: of_bl_drop word_size_bl) lemma unat_of_bl_length: "unat (of_bl xs :: 'a::len word) < 2 ^ (length xs)" proof (cases "length xs < LENGTH('a)") case True then have "(of_bl xs::'a::len word) < 2 ^ length xs" by (simp add: of_bl_length_less) with True show ?thesis by (simp add: word_less_nat_alt unat_of_nat) next case False have "unat (of_bl xs::'a::len word) < 2 ^ LENGTH('a)" by (simp split: unat_split) also from False have "LENGTH('a) \ length xs" by simp then have "2 ^ LENGTH('a) \ (2::nat) ^ length xs" by (rule power_increasing) simp finally show ?thesis . qed lemma word_msb_alt: "msb w \ hd (to_bl w)" for w :: "'a::len word" apply (simp add: msb_word_eq) apply (subst hd_conv_nth) apply simp apply (subst nth_to_bl) apply simp apply simp done lemma word_lsb_last: \lsb w \ last (to_bl w)\ for w :: \'a::len word\ using nth_to_bl [of \LENGTH('a) - Suc 0\ w] by (simp add: lsb_odd last_conv_nth) lemma is_aligned_to_bl: "is_aligned (w :: 'a :: len word) n = (True \ set (drop (size w - n) (to_bl w)))" 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) lemma of_bl_length: "length xs < LENGTH('a) \ of_bl xs < (2 :: 'a::len word) ^ length xs" by (simp add: of_bl_length_less) lemma of_bl_mult_and_not_mask_eq: "\is_aligned (a :: 'a::len word) n; length b + m \ n\ \ a + of_bl b * (2^m) AND NOT(mask n) = a" apply (simp flip: push_bit_eq_mult subtract_mask(1) take_bit_eq_mask) apply (subst disjunctive_add) apply (auto simp add: bit_simps not_le not_less) apply (meson is_aligned_imp_not_bit is_aligned_weaken less_diff_conv2) apply (erule is_alignedE') apply (simp add: take_bit_push_bit) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma bin_to_bl_of_bl_eq: "\is_aligned (a::'a::len word) n; length b + c \ n; length b + c < LENGTH('a)\ \ bin_to_bl (length b) (uint ((a + of_bl b * 2^c) >> c)) = b" apply (simp flip: push_bit_eq_mult take_bit_eq_mask add: shiftr_eq_drop_bit) apply (subst disjunctive_add) apply (auto simp add: bit_simps not_le not_less unsigned_or_eq unsigned_drop_bit_eq unsigned_push_bit_eq bin_to_bl_or simp flip: bin_to_bl_def) apply (meson is_aligned_imp_not_bit is_aligned_weaken less_diff_conv2) apply (erule is_alignedE') apply (rule nth_equalityI) apply (auto simp add: nth_bin_to_bl bit_simps rev_nth simp flip: bin_to_bl_def) done (* FIXME: move to Word distribution *) lemma bin_nth_minus_Bit0[simp]: "0 < n \ bin_nth (numeral (num.Bit0 w)) n = bin_nth (numeral w) (n - 1)" by (cases n; simp) lemma bin_nth_minus_Bit1[simp]: "0 < n \ bin_nth (numeral (num.Bit1 w)) n = bin_nth (numeral w) (n - 1)" by (cases n; simp) (* casting a long word to a shorter word and casting back to the long word is equal to the original long word -- if the word is small enough. 'l is the longer word. 's is the shorter word. *) lemma bl_cast_long_short_long_ingoreLeadingZero_generic: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (of_bl :: _ \ 'l::len word) (to_bl ((of_bl::_ \ 's::len word) (to_bl w))) = w" by (rule word_uint_eqI) (simp add: uint_of_bl_is_bl_to_bin uint_of_bl_is_bl_to_bin_drop) (* Casting between longer and shorter word. 'l is the longer word. 's is the shorter word. For example: 'l::len word is 128 word (full ipv6 address) 's::len word is 16 word (address piece of ipv6 address in colon-text-representation) *) corollary ucast_short_ucast_long_ingoreLeadingZero: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (ucast:: 's::len word \ 'l::len word) ((ucast:: 'l::len word \ 's::len word) w) = w" apply (subst ucast_bl)+ apply (rule bl_cast_long_short_long_ingoreLeadingZero_generic; simp) done lemma length_drop_mask: fixes w::"'a::len word" shows "length (dropWhile Not (to_bl (w AND mask n))) \ n" proof - have "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)" for ls n by(subst takeWhile_append2) simp+ then show ?thesis unfolding bl_and_mask by (simp add: dropWhile_eq_drop) qed lemma map_bits_rev_to_bl: "map ((!!) x) [0.. of_bl xs * 2^c < (2::'a::len word) ^ (length xs + c)" by (simp add: of_bl_length word_less_power_trans2) lemma of_bl_max: "(of_bl xs :: 'a::len word) \ mask (length xs)" proof - define ys where \ys = rev xs\ have \take_bit (length ys) (horner_sum of_bool 2 ys :: 'a word) = horner_sum of_bool 2 ys\ by transfer (simp add: take_bit_horner_sum_bit_eq min_def) then have \(of_bl (rev ys) :: 'a word) \ mask (length ys)\ by (simp only: of_bl_rev_eq less_eq_mask_iff_take_bit_eq_self) with ys_def show ?thesis by simp qed end diff --git a/thys/Word_Lib/Strict_part_mono.thy b/thys/Word_Lib/Strict_part_mono.thy --- a/thys/Word_Lib/Strict_part_mono.thy +++ b/thys/Word_Lib/Strict_part_mono.thy @@ -1,56 +1,56 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) theory Strict_part_mono imports "HOL-Library.Word" More_Word begin definition strict_part_mono :: "'a set \ ('a :: order \ 'b :: order) \ bool" where "strict_part_mono S f \ \A\S. \B\S. A < B \ f A < f B" lemma strict_part_mono_by_steps: "strict_part_mono {..n :: nat} f = (n \ 0 \ f (n - 1) < f n \ strict_part_mono {.. n - 1} f)" apply (cases n; simp add: strict_part_mono_def) apply (safe; clarsimp) apply (case_tac "B = Suc nat"; simp) apply (case_tac "A = nat"; clarsimp) apply (erule order_less_trans [rotated]) apply simp done lemma strict_part_mono_singleton[simp]: "strict_part_mono {x} f" by (simp add: strict_part_mono_def) lemma strict_part_mono_lt: "\ x < f 0; strict_part_mono {.. n :: nat} f \ \ \m \ n. x < f m" - by (metis atMost_iff le_0_eq le_cases neq0_conv order.strict_trans strict_part_mono_def) + by (auto simp add: strict_part_mono_def Ball_def intro: order.strict_trans) lemma strict_part_mono_reverseE: "\ f n \ f m; strict_part_mono {.. N :: nat} f; n \ N \ \ n \ m" by (rule ccontr) (fastforce simp: linorder_not_le strict_part_mono_def) lemma two_power_strict_part_mono: "strict_part_mono {..LENGTH('a) - 1} (\x. (2 :: 'a :: len word) ^ x)" proof - { fix n have "n < LENGTH('a) \ strict_part_mono {..n} (\x. (2 :: 'a :: len word) ^ x)" proof (induct n) case 0 then show ?case by simp next case (Suc n) from Suc.prems have "2 ^ n < (2 :: 'a :: len word) ^ Suc n" using power_strict_increasing unat_power_lower word_less_nat_alt by fastforce with Suc show ?case by (subst strict_part_mono_by_steps) simp qed } then show ?thesis by simp qed end diff --git a/thys/Word_Lib/Traditional_Infix_Syntax.thy b/thys/Word_Lib/Traditional_Infix_Syntax.thy --- a/thys/Word_Lib/Traditional_Infix_Syntax.thy +++ b/thys/Word_Lib/Traditional_Infix_Syntax.thy @@ -1,1073 +1,1077 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson, NICTA *) section \Operation variants with traditional syntax\ theory Traditional_Infix_Syntax imports "HOL-Library.Word" More_Word Signed_Words begin class semiring_bit_syntax = semiring_bit_shifts begin definition test_bit :: \'a \ nat \ bool\ (infixl "!!" 100) where test_bit_eq_bit: \test_bit = bit\ definition shiftl :: \'a \ nat \ 'a\ (infixl "<<" 55) where shiftl_eq_push_bit: \a << n = push_bit n a\ definition shiftr :: \'a \ nat \ 'a\ (infixl ">>" 55) where shiftr_eq_drop_bit: \a >> n = drop_bit n a\ end instance word :: (len) semiring_bit_syntax .. context includes lifting_syntax begin lemma test_bit_word_transfer [transfer_rule]: \(pcr_word ===> (=)) (\k n. n < LENGTH('a) \ bit k n) (test_bit :: 'a::len word \ _)\ by (unfold test_bit_eq_bit) transfer_prover lemma shiftl_word_transfer [transfer_rule]: \(pcr_word ===> (=) ===> pcr_word) (\k n. push_bit n k) shiftl\ by (unfold shiftl_eq_push_bit) transfer_prover lemma shiftr_word_transfer [transfer_rule]: \(pcr_word ===> (=) ===> pcr_word) (\k n. (drop_bit n \ take_bit LENGTH('a)) k) (shiftr :: 'a::len word \ _)\ by (unfold shiftr_eq_drop_bit) transfer_prover end lemma test_bit_word_eq: \test_bit = (bit :: 'a::len word \ _)\ by (fact test_bit_eq_bit) lemma shiftl_word_eq: \w << n = push_bit n w\ for w :: \'a::len word\ by (fact shiftl_eq_push_bit) lemma shiftr_word_eq: \w >> n = drop_bit n w\ for w :: \'a::len word\ by (fact shiftr_eq_drop_bit) lemma test_bit_eq_iff: "test_bit u = test_bit v \ u = v" for u v :: "'a::len word" by (simp add: bit_eq_iff test_bit_eq_bit fun_eq_iff) lemma test_bit_size: "w !! n \ n < size w" for w :: "'a::len word" by transfer simp lemma word_eq_iff: "x = y \ (\n?P \ ?Q\) for x y :: "'a::len word" by transfer (auto simp add: bit_eq_iff bit_take_bit_iff) lemma word_eqI: "(\n. n < size u \ u !! n = v !! n) \ u = v" for u :: "'a::len word" by (simp add: word_size word_eq_iff) lemma word_eqD: "u = v \ u !! x = v !! x" for u v :: "'a::len word" by simp lemma test_bit_bin': "w !! n \ n < size w \ bit (uint w) n" by transfer (simp add: bit_take_bit_iff) lemmas test_bit_bin = test_bit_bin' [unfolded word_size] lemma word_test_bit_def: \test_bit a = bit (uint a)\ by transfer (simp add: fun_eq_iff bit_take_bit_iff) lemmas test_bit_def' = word_test_bit_def [THEN fun_cong] lemma word_test_bit_transfer [transfer_rule]: "(rel_fun pcr_word (rel_fun (=) (=))) (\x n. n < LENGTH('a) \ bit x n) (test_bit :: 'a::len word \ _)" by (simp only: test_bit_eq_bit) transfer_prover lemma test_bit_wi [simp]: "(word_of_int x :: 'a::len word) !! n \ n < LENGTH('a) \ bit x n" by transfer simp lemma word_ops_nth_size: "n < size x \ (x OR y) !! n = (x !! n | y !! n) \ (x AND y) !! n = (x !! n \ y !! n) \ (x XOR y) !! n = (x !! n \ y !! n) \ (NOT x) !! n = (\ x !! n)" for x :: "'a::len word" by transfer (simp add: bit_or_iff bit_and_iff bit_xor_iff bit_not_iff) lemma word_ao_nth: "(x OR y) !! n = (x !! n | y !! n) \ (x AND y) !! n = (x !! n \ y !! n)" for x :: "'a::len word" by transfer (auto simp add: bit_or_iff bit_and_iff) lemmas msb0 = len_gt_0 [THEN diff_Suc_less, THEN word_ops_nth_size [unfolded word_size]] lemmas msb1 = msb0 [where i = 0] lemma test_bit_numeral [simp]: "(numeral w :: 'a::len word) !! n \ n < LENGTH('a) \ bit (numeral w :: int) n" by transfer (rule refl) lemma test_bit_neg_numeral [simp]: "(- numeral w :: 'a::len word) !! n \ n < LENGTH('a) \ bit (- numeral w :: int) n" by transfer (rule refl) lemma test_bit_1 [iff]: "(1 :: 'a::len word) !! n \ n = 0" by transfer (auto simp add: bit_1_iff) lemma nth_0 [simp]: "\ (0 :: 'a::len word) !! n" by transfer simp lemma nth_minus1 [simp]: "(-1 :: 'a::len word) !! n \ n < LENGTH('a)" by transfer simp lemma shiftl1_code [code]: \shiftl1 w = push_bit 1 w\ by transfer (simp add: ac_simps) lemma uint_shiftr_eq: \uint (w >> n) = uint w div 2 ^ n\ by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit min_def le_less less_diff_conv) lemma shiftr1_code [code]: \shiftr1 w = drop_bit 1 w\ by transfer (simp add: drop_bit_Suc) lemma shiftl_def: \w << n = (shiftl1 ^^ n) w\ proof - have \push_bit n = (((*) 2 ^^ n) :: int \ int)\ for n by (induction n) (simp_all add: fun_eq_iff funpow_swap1, simp add: ac_simps) then show ?thesis by transfer simp qed lemma shiftr_def: \w >> n = (shiftr1 ^^ n) w\ proof - have \shiftr1 ^^ n = (drop_bit n :: 'a word \ 'a word)\ apply (induction n) apply simp apply (simp only: shiftr1_eq_div_2 [abs_def] drop_bit_eq_div [abs_def] funpow_Suc_right) apply (use div_exp_eq [of _ 1, where ?'a = \'a word\] in simp) done then show ?thesis by (simp add: shiftr_eq_drop_bit) qed lemma bit_shiftl_word_iff [bit_simps]: \bit (w << m) n \ m \ n \ n < LENGTH('a) \ bit w (n - m)\ for w :: \'a::len word\ by (simp add: shiftl_word_eq bit_push_bit_iff not_le) lemma bit_shiftr_word_iff [bit_simps]: \bit (w >> m) n \ bit w (m + n)\ for w :: \'a::len word\ by (simp add: shiftr_word_eq bit_drop_bit_eq) lift_definition sshiftr :: \'a::len word \ nat \ 'a word\ (infixl \>>>\ 55) is \\k n. take_bit LENGTH('a) (drop_bit n (signed_take_bit (LENGTH('a) - Suc 0) k))\ by (simp flip: signed_take_bit_decr_length_iff) lemma sshiftr_eq [code]: \w >>> n = signed_drop_bit n w\ by transfer simp lemma sshiftr_eq_funpow_sshiftr1: \w >>> n = (sshiftr1 ^^ n) w\ apply (rule sym) apply (simp add: sshiftr1_eq_signed_drop_bit_Suc_0 sshiftr_eq) apply (induction n) apply simp_all done lemma uint_sshiftr_eq: \uint (w >>> n) = take_bit LENGTH('a) (sint w div 2 ^ n)\ for w :: \'a::len word\ by transfer (simp flip: drop_bit_eq_div) lemma sshift1_code [code]: \sshiftr1 w = signed_drop_bit 1 w\ by transfer (simp add: drop_bit_Suc) lemma sshiftr_0 [simp]: "0 >>> n = 0" by transfer simp lemma sshiftr_n1 [simp]: "-1 >>> n = -1" by transfer simp lemma bit_sshiftr_word_iff [bit_simps]: \bit (w >>> m) n \ bit w (if LENGTH('a) - m \ n \ n < LENGTH('a) then LENGTH('a) - 1 else (m + n))\ for w :: \'a::len word\ apply transfer apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq bit_signed_take_bit_iff min_def not_le simp flip: bit_Suc) using le_less_Suc_eq apply fastforce using le_less_Suc_eq apply fastforce done lemma nth_sshiftr : "(w >>> m) !! n = (n < size w \ (if n + m \ size w then w !! (size w - 1) else w !! (n + m)))" apply transfer apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq bit_signed_take_bit_iff min_def not_le ac_simps) using le_less_Suc_eq apply fastforce using le_less_Suc_eq apply fastforce done lemma sshiftr_numeral [simp]: \(numeral k >>> numeral n :: 'a::len word) = word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (numeral k)))\ apply (rule word_eqI) apply (cases \LENGTH('a)\) apply (simp_all add: word_size bit_drop_bit_eq nth_sshiftr bit_signed_take_bit_iff min_def not_le not_less less_Suc_eq_le ac_simps) done setup \ Context.theory_map (fold SMT_Word.add_word_shift' [ (\<^term>\shiftl :: 'a::len word \ _\, "bvshl"), (\<^term>\shiftr :: 'a::len word \ _\, "bvlshr"), (\<^term>\sshiftr :: 'a::len word \ _\, "bvashr") ]) \ lemma revcast_down_us [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >>> n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_ucast_iff bit_sshiftr_word_iff ac_simps) done lemma revcast_down_ss [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >>> n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_word_scast_iff bit_sshiftr_word_iff ac_simps) done lemma sshiftr_div_2n: "sint (w >>> n) = sint w div 2 ^ n" using sint_signed_drop_bit_eq [of n w] by (simp add: drop_bit_eq_div sshiftr_eq) lemmas lsb0 = len_gt_0 [THEN word_ops_nth_size [unfolded word_size]] lemma nth_sint: fixes w :: "'a::len word" defines "l \ LENGTH('a)" shows "bit (sint w) n = (if n < l - 1 then w !! n else w !! (l - 1))" unfolding sint_uint l_def by (auto simp: bit_signed_take_bit_iff word_test_bit_def not_less min_def) lemma test_bit_2p: "(word_of_int (2 ^ n)::'a::len word) !! m \ m = n \ m < LENGTH('a)" by transfer (auto simp add: bit_exp_iff) lemma nth_w2p: "((2::'a::len word) ^ n) !! m \ m = n \ m < LENGTH('a::len)" by transfer (auto simp add: bit_exp_iff) lemma bang_is_le: "x !! m \ 2 ^ m \ x" for x :: "'a::len word" apply (rule xtrans(3)) apply (rule_tac [2] y = "x" in le_word_or2) apply (rule word_eqI) apply (auto simp add: word_ao_nth nth_w2p word_size) done lemma mask_eq: \mask n = (1 << n) - (1 :: 'a::len word)\ by transfer (simp add: mask_eq_exp_minus_1 push_bit_of_1) lemma nth_ucast: "(ucast w::'a::len word) !! n = (w !! n \ n < LENGTH('a))" by transfer (simp add: bit_take_bit_iff ac_simps) lemma shiftl_0 [simp]: "(0::'a::len word) << n = 0" by transfer simp lemma shiftr_0 [simp]: "(0::'a::len word) >> n = 0" by transfer simp lemma nth_shiftl1: "shiftl1 w !! n \ n < size w \ n > 0 \ w !! (n - 1)" by transfer (auto simp add: bit_double_iff) lemma nth_shiftl': "(w << m) !! n \ n < size w \ n >= m \ w !! (n - m)" for w :: "'a::len word" by transfer (auto simp add: bit_push_bit_iff) lemmas nth_shiftl = nth_shiftl' [unfolded word_size] lemma nth_shiftr1: "shiftr1 w !! n = w !! Suc n" by transfer (auto simp add: bit_take_bit_iff simp flip: bit_Suc) lemma nth_shiftr: "(w >> m) !! n = w !! (n + m)" for w :: "'a::len word" apply (unfold shiftr_def) apply (induct "m" arbitrary: n) apply (auto simp add: nth_shiftr1) done lemma nth_sshiftr1: "sshiftr1 w !! n = (if n = size w - 1 then w !! n else w !! Suc n)" apply transfer apply (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def simp flip: bit_Suc) using le_less_Suc_eq apply fastforce using le_less_Suc_eq apply fastforce done lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" by (fact uint_shiftr_eq) lemma shiftl_rev: "shiftl w n = word_reverse (shiftr (word_reverse w) n)" by (induct n) (auto simp add: shiftl_def shiftr_def shiftl1_rev) lemma rev_shiftl: "word_reverse w << n = word_reverse (w >> n)" by (simp add: shiftl_rev) lemma shiftr_rev: "w >> n = word_reverse (word_reverse w << n)" by (simp add: rev_shiftl) lemma rev_shiftr: "word_reverse w >> n = word_reverse (w << n)" by (simp add: shiftr_rev) lemma shiftl_numeral [simp]: \numeral k << numeral l = (push_bit (numeral l) (numeral k) :: 'a::len word)\ by (fact shiftl_word_eq) lemma shiftl_zero_size: "size x \ n \ x << n = 0" for x :: "'a::len word" apply transfer apply (simp add: take_bit_push_bit) done lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" for w :: "'a::len word" by (induct n) (auto simp: shiftl_def shiftl1_2t) lemma shiftr_numeral [simp]: \(numeral k >> numeral n :: 'a::len word) = drop_bit (numeral n) (numeral k)\ by (fact shiftr_word_eq) lemma shiftr_numeral_Suc [simp]: \(numeral k >> Suc 0 :: 'a::len word) = drop_bit (Suc 0) (numeral k)\ by (fact shiftr_word_eq) lemma drop_bit_numeral_bit0_1 [simp]: \drop_bit (Suc 0) (numeral k) = (word_of_int (drop_bit (Suc 0) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ by (metis Word_eq_word_of_int drop_bit_word.abs_eq of_int_numeral) lemma nth_mask [simp]: \(mask n :: 'a::len word) !! i \ i < n \ i < size (mask n :: 'a word)\ by (auto simp add: test_bit_word_eq word_size Word.bit_mask_iff) lemma slice_shiftr: "slice n w = ucast (w >> n)" apply (rule bit_word_eqI) apply (cases \n \ LENGTH('b)\) apply (auto simp add: bit_slice_iff bit_ucast_iff bit_shiftr_word_iff ac_simps dest: bit_imp_le_length) done lemma nth_slice: "(slice n w :: 'a::len word) !! m = (w !! (m + n) \ m < LENGTH('a))" by (simp add: slice_shiftr nth_ucast nth_shiftr) lemma revcast_down_uu [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >> n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_ucast_iff bit_shiftr_word_iff ac_simps) done lemma revcast_down_su [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >> n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_word_scast_iff bit_shiftr_word_iff ac_simps) done lemma cast_down_rev [OF refl]: "uc = ucast \ source_size uc = target_size uc + n \ uc w = revcast (w << n)" for w :: "'a::len word" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_word_ucast_iff bit_shiftl_word_iff) done lemma revcast_up [OF refl]: "rc = revcast \ source_size rc + n = target_size rc \ rc w = (ucast w :: 'a::len word) << n" apply (simp add: source_size_def target_size_def) apply (rule bit_word_eqI) apply (simp add: bit_revcast_iff bit_word_ucast_iff bit_shiftl_word_iff) apply auto apply (metis add.commute add_diff_cancel_right) apply (metis diff_add_inverse2 diff_diff_add) done lemmas rc1 = revcast_up [THEN revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] lemmas rc2 = revcast_down_uu [THEN revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] lemmas ucast_up = rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] lemmas ucast_down = rc2 [simplified rev_shiftr revcast_ucast [symmetric]] \ \problem posed by TPHOLs referee: criterion for overflow of addition of signed integers\ lemma sofl_test: \sint x + sint y = sint (x + y) \ (x + y XOR x) AND (x + y XOR y) >> (size x - 1) = 0\ for x y :: \'a::len word\ proof - obtain n where n: \LENGTH('a) = Suc n\ by (cases \LENGTH('a)\) simp_all have *: \sint x + sint y + 2 ^ Suc n > signed_take_bit n (sint x + sint y) \ sint x + sint y \ - (2 ^ n)\ \signed_take_bit n (sint x + sint y) > sint x + sint y - 2 ^ Suc n \ 2 ^ n > sint x + sint y\ using signed_take_bit_int_greater_eq [of \sint x + sint y\ n] signed_take_bit_int_less_eq [of n \sint x + sint y\] by (auto intro: ccontr) have \sint x + sint y = sint (x + y) \ (sint (x + y) < 0 \ sint x < 0) \ (sint (x + y) < 0 \ sint y < 0)\ using sint_less [of x] sint_greater_eq [of x] sint_less [of y] sint_greater_eq [of y] signed_take_bit_int_eq_self [of \LENGTH('a) - 1\ \sint x + sint y\] apply (auto simp add: not_less) apply (unfold sint_word_ariths) apply (subst signed_take_bit_int_eq_self) prefer 4 - apply (subst signed_take_bit_int_eq_self) - prefer 7 - apply (subst signed_take_bit_int_eq_self) - prefer 10 + apply (subst signed_take_bit_int_eq_self) + prefer 7 + apply (subst signed_take_bit_int_eq_self) + prefer 10 apply (subst signed_take_bit_int_eq_self) - apply (auto simp add: signed_take_bit_int_eq_self signed_take_bit_eq_take_bit_minus take_bit_Suc_from_most n not_less intro!: *) + apply (auto simp add: signed_take_bit_int_eq_self signed_take_bit_eq_take_bit_minus take_bit_Suc_from_most n not_less intro!: *) + apply (smt (z3) take_bit_nonnegative) + apply (smt (z3) take_bit_int_less_exp) + apply (smt (z3) take_bit_nonnegative) + apply (smt (z3) take_bit_int_less_exp) done then show ?thesis apply (simp only: One_nat_def word_size shiftr_word_eq drop_bit_eq_zero_iff_not_bit_last bit_and_iff bit_xor_iff) apply (simp add: bit_last_iff) done qed lemma shiftr_zero_size: "size x \ n \ x >> n = 0" for x :: "'a :: len word" by (rule word_eqI) (auto simp add: nth_shiftr dest: test_bit_size) lemma test_bit_cat [OF refl]: "wc = word_cat a b \ wc !! n = (n < size wc \ (if n < size b then b !! n else a !! (n - size b)))" apply (simp add: word_size not_less; transfer) apply (auto simp add: bit_concat_bit_iff bit_take_bit_iff) done \ \keep quantifiers for use in simplification\ lemma test_bit_split': "word_split c = (a, b) \ (\n m. b !! n = (n < size b \ c !! n) \ a !! m = (m < size a \ c !! (m + size b)))" by (auto simp add: word_split_bin' test_bit_bin bit_unsigned_iff word_size bit_drop_bit_eq ac_simps dest: bit_imp_le_length) lemma test_bit_split: "word_split c = (a, b) \ (\n::nat. b !! n \ n < size b \ c !! n) \ (\m::nat. a !! m \ m < size a \ c !! (m + size b))" by (simp add: test_bit_split') lemma test_bit_split_eq: "word_split c = (a, b) \ ((\n::nat. b !! n = (n < size b \ c !! n)) \ (\m::nat. a !! m = (m < size a \ c !! (m + size b))))" apply (rule_tac iffI) apply (rule_tac conjI) apply (erule test_bit_split [THEN conjunct1]) apply (erule test_bit_split [THEN conjunct2]) apply (case_tac "word_split c") apply (frule test_bit_split) apply (erule trans) apply (fastforce intro!: word_eqI simp add: word_size) done lemma test_bit_rcat: "sw = size (hd wl) \ rc = word_rcat wl \ rc !! n = (n < size rc \ n div sw < size wl \ (rev wl) ! (n div sw) !! (n mod sw))" for wl :: "'a::len word list" by (simp add: word_size word_rcat_def foldl_map rev_map bit_horner_sum_uint_exp_iff) (simp add: test_bit_eq_bit) lemmas test_bit_cong = arg_cong [where f = "test_bit", THEN fun_cong] lemma max_test_bit: "(max_word::'a::len word) !! n \ n < LENGTH('a)" by (fact nth_minus1) lemma shiftr_x_0 [iff]: "x >> 0 = x" for x :: "'a::len word" by transfer simp lemma shiftl_x_0 [simp]: "x << 0 = x" for x :: "'a::len word" by (simp add: shiftl_t2n) lemma shiftl_1 [simp]: "(1::'a::len word) << n = 2^n" by (simp add: shiftl_t2n) lemma shiftr_1[simp]: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" by (induct n) (auto simp: shiftr_def) lemma map_nth_0 [simp]: "map ((!!) (0::'a::len word)) xs = replicate (length xs) False" by (induct xs) auto lemma word_and_1: "n AND 1 = (if n !! 0 then 1 else 0)" for n :: "_ word" by (rule bit_word_eqI) (auto simp add: bit_and_iff test_bit_eq_bit bit_1_iff intro: gr0I) lemma test_bit_1' [simp]: "(1 :: 'a :: len word) !! n \ 0 < LENGTH('a) \ n = 0" by simp lemma shiftl0: "x << 0 = (x :: 'a :: len word)" by (fact shiftl_x_0) lemma word_ops_nth [simp]: fixes x y :: \'a::len word\ shows word_or_nth: "(x OR y) !! n = (x !! n \ y !! n)" and word_and_nth: "(x AND y) !! n = (x !! n \ y !! n)" and word_xor_nth: "(x XOR y) !! n = (x !! n \ y !! n)" by ((cases "n < size x", auto dest: test_bit_size simp: word_ops_nth_size word_size)[1])+ lemma and_not_mask: "w AND NOT (mask n) = (w >> n) << n" for w :: \'a::len word\ apply (rule word_eqI) apply (simp add : word_ops_nth_size word_size) apply (simp add : nth_shiftr nth_shiftl) by auto lemma and_mask: "w AND mask n = (w << (size w - n)) >> (size w - n)" for w :: \'a::len word\ apply (rule word_eqI) apply (simp add : word_ops_nth_size word_size) apply (simp add : nth_shiftr nth_shiftl) by auto lemma nth_w2p_same: "(2^n :: 'a :: len word) !! n = (n < LENGTH('a))" by (simp add : nth_w2p) lemma shiftr_div_2n_w: "n < size w \ w >> n = w div (2^n :: 'a :: len word)" apply (unfold word_div_def) apply (simp add: uint_2p_alt word_size) apply (metis uint_shiftr_eq word_of_int_uint) done lemma le_shiftr: "u \ v \ u >> (n :: nat) \ (v :: 'a :: len word) >> n" apply (unfold shiftr_def) apply (induct_tac "n") apply auto apply (erule le_shiftr1) done lemma shiftr_mask_le: "n <= m \ mask n >> m = (0 :: 'a::len word)" apply (rule word_eqI) apply (simp add: word_size nth_shiftr) done lemma shiftr_mask [simp]: \mask m >> m = (0::'a::len word)\ by (rule shiftr_mask_le) simp lemma word_leI: "(\n. \n < size (u::'a::len word); u !! n \ \ (v::'a::len word) !! n) \ u <= v" apply (rule xtrans(4)) apply (rule word_and_le2) apply (rule word_eqI) apply (simp add: word_ao_nth) apply safe apply assumption apply (erule_tac [2] asm_rl) apply (unfold word_size) by auto lemma le_mask_iff: "(w \ mask n) = (w >> n = 0)" for w :: \'a::len word\ apply safe apply (rule word_le_0_iff [THEN iffD1]) apply (rule xtrans(3)) apply (erule_tac [2] le_shiftr) apply simp apply (rule word_leI) apply (rename_tac n') apply (drule_tac x = "n' - n" in word_eqD) apply (simp add : nth_shiftr word_size) apply (case_tac "n <= n'") by auto lemma and_mask_eq_iff_shiftr_0: "(w AND mask n = w) = (w >> n = 0)" for w :: \'a::len word\ apply (unfold test_bit_eq_iff [THEN sym]) apply (rule iffI) apply (rule ext) apply (rule_tac [2] ext) apply (auto simp add : word_ao_nth nth_shiftr) apply (drule arg_cong) apply (drule iffD2) apply assumption apply (simp add : word_ao_nth) prefer 2 apply (simp add : word_size test_bit_bin) apply transfer apply (auto simp add: fun_eq_iff bit_simps) apply (metis add_diff_inverse_nat) done lemma mask_shiftl_decompose: "mask m << n = mask (m + n) AND NOT (mask n :: 'a::len word)" by (auto intro!: word_eqI simp: and_not_mask nth_shiftl nth_shiftr word_size) lemma bang_eq: fixes x :: "'a::len word" shows "(x = y) = (\n. x !! n = y !! n)" by (subst test_bit_eq_iff[symmetric]) fastforce lemma shiftl_over_and_dist: fixes a::"'a::len word" shows "(a AND b) << c = (a << c) AND (b << c)" apply(rule word_eqI) apply(simp add: word_ao_nth nth_shiftl, safe) done lemma shiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >> c = (a >> c) AND (b >> c)" apply(rule word_eqI) apply(simp add:nth_shiftr word_ao_nth) done lemma sshiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >>> c = (a >>> c) AND (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemma shiftl_over_or_dist: fixes a::"'a::len word" shows "a OR b << c = (a << c) OR (b << c)" apply(rule word_eqI) apply(simp add:nth_shiftl word_ao_nth, safe) done lemma shiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >> c = (a >> c) OR (b >> c)" apply(rule word_eqI) apply(simp add:nth_shiftr word_ao_nth) done lemma sshiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >>> c = (a >>> c) OR (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemmas shift_over_ao_dists = shiftl_over_or_dist shiftr_over_or_dist sshiftr_over_or_dist shiftl_over_and_dist shiftr_over_and_dist sshiftr_over_and_dist lemma shiftl_shiftl: fixes a::"'a::len word" shows "a << b << c = a << (b + c)" apply(rule word_eqI) apply(auto simp:word_size nth_shiftl add.commute add.left_commute) done lemma shiftr_shiftr: fixes a::"'a::len word" shows "a >> b >> c = a >> (b + c)" apply(rule word_eqI) apply(simp add:word_size nth_shiftr add.left_commute add.commute) done lemma shiftl_shiftr1: fixes a::"'a::len word" shows "c \ b \ a << b >> c = a AND (mask (size a - b)) << (b - c)" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth) done lemma shiftl_shiftr2: fixes a::"'a::len word" shows "b < c \ a << b >> c = (a >> (c - b)) AND (mask (size a - c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth) done lemma shiftr_shiftl1: fixes a::"'a::len word" shows "c \ b \ a >> b << c = (a >> (b - c)) AND (NOT (mask c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ops_nth_size) done lemma shiftr_shiftl2: fixes a::"'a::len word" shows "b < c \ a >> b << c = (a << (c - b)) AND (NOT (mask c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ops_nth_size) done lemmas multi_shift_simps = shiftl_shiftl shiftr_shiftr shiftl_shiftr1 shiftl_shiftr2 shiftr_shiftl1 shiftr_shiftl2 lemma shiftr_mask2: "n \ LENGTH('a) \ (mask n >> m :: ('a :: len) word) = mask (n - m)" apply (rule word_eqI) apply (simp add: nth_shiftr word_size) apply arith done lemma word_shiftl_add_distrib: fixes x :: "'a :: len word" shows "(x + y) << n = (x << n) + (y << n)" by (simp add: shiftl_t2n ring_distribs) lemma mask_shift: "(x AND NOT (mask y)) >> y = x >> y" for x :: \'a::len word\ apply (rule bit_eqI) apply (simp add: bit_and_iff bit_not_iff bit_shiftr_word_iff bit_mask_iff not_le) using bit_imp_le_length apply auto done lemma shiftr_div_2n': "unat (w >> n) = unat w div 2 ^ n" apply (unfold unat_eq_nat_uint) apply (subst shiftr_div_2n) apply (subst nat_div_distrib) apply simp apply (simp add: nat_power_eq) done lemma shiftl_shiftr_id: assumes nv: "n < LENGTH('a)" and xv: "x < 2 ^ (LENGTH('a) - n)" shows "x << n >> n = (x::'a::len word)" apply (simp add: shiftl_t2n) apply (rule word_eq_unatI) apply (subst shiftr_div_2n') apply (cases n) apply simp apply (subst iffD1 [OF unat_mult_lem])+ apply (subst unat_power_lower[OF nv]) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl]) apply (rule unat_power_lower) apply simp apply (subst unat_power_lower[OF nv]) apply simp done lemma ucast_shiftl_eq_0: fixes w :: "'a :: len word" shows "\ n \ LENGTH('b) \ \ ucast (w << n) = (0 :: 'b :: len word)" by transfer (simp add: take_bit_push_bit) lemma word_shift_nonzero: "\ (x::'a::len word) \ 2 ^ m; m + n < LENGTH('a::len); x \ 0\ \ x << n \ 0" apply (simp only: word_neq_0_conv word_less_nat_alt shiftl_t2n mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (subst mod_less) apply (rule order_le_less_trans) apply (erule mult_le_mono2) apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done lemma word_shiftr_lt: fixes w :: "'a::len word" shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))" apply (subst shiftr_div_2n') apply transfer apply (simp flip: drop_bit_eq_div add: drop_bit_nat_eq drop_bit_take_bit) done lemma neg_mask_test_bit: "(NOT(mask n) :: 'a :: len word) !! m = (n \ m \ m < LENGTH('a))" by (metis not_le nth_mask test_bit_bin word_ops_nth_size word_size) lemma upper_bits_unset_is_l2p: \(\n' \ n. n' < LENGTH('a) \ \ p !! n') \ (p < 2 ^ n)\ (is \?P \ ?Q\) if \n < LENGTH('a)\ for p :: "'a :: len word" proof assume ?Q then show ?P by (meson bang_is_le le_less_trans not_le word_power_increasing) next assume ?P have \take_bit n p = p\ proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ show \bit (take_bit n p) q \ bit p q\ proof (cases \q < n\) case True then show ?thesis by (auto simp add: bit_simps) next case False then have \n \ q\ by simp with \?P\ \q < LENGTH('a)\ have \\ bit p q\ by (simp add: test_bit_eq_bit) then show ?thesis by (simp add: bit_simps) qed qed with that show ?Q using take_bit_word_eq_self_iff [of n p] by auto qed lemma less_2p_is_upper_bits_unset: "p < 2 ^ n \ n < LENGTH('a) \ (\n' \ n. n' < LENGTH('a) \ \ p !! n')" for p :: "'a :: len word" by (meson le_less_trans le_mask_iff_lt_2n upper_bits_unset_is_l2p word_zero_le) lemma test_bit_over: "n \ size (x::'a::len word) \ (x !! n) = False" by transfer auto lemma le_mask_high_bits: "w \ mask n \ (\i \ {n ..< size w}. \ w !! i)" for w :: \'a::len word\ by (auto simp: word_size and_mask_eq_iff_le_mask[symmetric] word_eq_iff) lemma test_bit_conj_lt: "(x !! m \ m < LENGTH('a)) = x !! m" for x :: "'a :: len word" using test_bit_bin by blast lemma neg_test_bit: "(NOT x) !! n = (\ x !! n \ n < LENGTH('a))" for x :: "'a::len word" by (cases "n < LENGTH('a)") (auto simp add: test_bit_over word_ops_nth_size word_size) lemma shiftr_less_t2n': "\ x AND mask (n + m) = x; m < LENGTH('a) \ \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_drop_bit ac_simps) done lemma shiftr_less_t2n: "x < 2 ^ (n + m) \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (rule shiftr_less_t2n') apply (erule less_mask_eq) apply (rule ccontr) apply (simp add: not_less) apply (subst (asm) p2_eq_0[symmetric]) apply (simp add: power_add) done lemma shiftr_eq_0: "n \ LENGTH('a) \ ((w::'a::len word) >> n) = 0" apply (cut_tac shiftr_less_t2n'[of w n 0], simp) apply (simp add: mask_eq_iff) apply (simp add: lt2p_lem) apply simp done lemma shiftr_not_mask_0: "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) AND NOT (mask m) = 0" by (rule bit_word_eqI) (auto simp add: bit_simps dest: bit_imp_le_length) lemma shiftl_less_t2n: fixes x :: "'a :: len word" shows "\ x < (2 ^ (m - n)); m < LENGTH('a) \ \ (x << n) < 2 ^ m" apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_push_bit) done lemma shiftl_less_t2n': "(x::'a::len word) < 2 ^ m \ m+n < LENGTH('a) \ x << n < 2 ^ (m + n)" by (rule shiftl_less_t2n) simp_all lemma nth_w2p_scast [simp]: "((scast ((2::'a::len signed word) ^ n) :: 'a word) !! m) \ ((((2::'a::len word) ^ n) :: 'a word) !! m)" by transfer (auto simp add: bit_simps) lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (clarsimp simp: word_eq_iff) lemma signed_shift_guard_to_word: "\ n < len_of TYPE ('a); n > 0 \ \ (unat (x :: 'a :: len word) * 2 ^ y < 2 ^ n) = (x = 0 \ x < (1 << n >> y))" apply (simp only: nat_mult_power_less_eq) apply (cases "y \ n") apply (simp only: shiftl_shiftr1) apply (subst less_mask_eq) apply (simp add: word_less_nat_alt word_size) apply (rule order_less_le_trans[rotated], rule power_increasing[where n=1]) apply simp apply simp apply simp apply (simp add: nat_mult_power_less_eq word_less_nat_alt word_size) apply auto[1] apply (simp only: shiftl_shiftr2, simp add: unat_eq_0) done lemma nth_bounded: "\(x :: 'a :: len word) !! n; x < 2 ^ m; m \ len_of TYPE ('a)\ \ n < m" apply (rule ccontr) apply (auto simp add: not_less test_bit_word_eq) apply (meson bit_imp_le_length bit_uint_iff less_2p_is_upper_bits_unset test_bit_bin) done lemma shiftl_mask_is_0[simp]: "(x << n) AND mask n = 0" for x :: \'a::len word\ by (simp flip: take_bit_eq_mask add: shiftl_eq_push_bit take_bit_push_bit) lemma rshift_sub_mask_eq: "(a >> (size a - b)) AND mask b = a >> (size a - b)" for a :: \'a::len word\ using shiftl_shiftr2[where a=a and b=0 and c="size a - b"] apply (cases "b < size a") apply simp apply (simp add: linorder_not_less mask_eq_decr_exp word_size p2_eq_0[THEN iffD2]) done lemma shiftl_shiftr3: "b \ c \ a << b >> c = (a >> c - b) AND mask (size a - c)" for a :: \'a::len word\ apply (cases "b = c") apply (simp add: shiftl_shiftr1) apply (simp add: shiftl_shiftr2) done lemma and_mask_shiftr_comm: "m \ size w \ (w AND mask m) >> n = (w >> n) AND mask (m-n)" for w :: \'a::len word\ by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) lemma and_mask_shiftl_comm: "m+n \ size w \ (w AND mask m) << n = (w << n) AND mask (m+n)" for w :: \'a::len word\ by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1) lemma le_mask_shiftl_le_mask: "s = m + n \ x \ mask n \ x << m \ mask s" for x :: \'a::len word\ by (simp add: le_mask_iff shiftl_shiftr3) lemma word_and_1_shiftl: "x AND (1 << n) = (if x !! n then (1 << n) else 0)" for x :: "'a :: len word" apply (rule bit_word_eqI; transfer) apply (auto simp add: bit_simps not_le ac_simps) done lemmas word_and_1_shiftls' = word_and_1_shiftl[where n=0] word_and_1_shiftl[where n=1] word_and_1_shiftl[where n=2] lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified] lemma word_and_mask_shiftl: "x AND (mask n << m) = ((x >> m) AND mask n) << m" for x :: \'a::len word\ apply (rule bit_word_eqI; transfer) apply (auto simp add: bit_simps not_le ac_simps) done lemma shift_times_fold: "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)" by (simp add: shiftl_t2n ac_simps power_add) lemma of_bool_nth: "of_bool (x !! v) = (x >> v) AND 1" for x :: \'a::len word\ by (simp add: test_bit_word_eq shiftr_word_eq bit_eq_iff) (auto simp add: bit_1_iff bit_and_iff bit_drop_bit_eq intro: ccontr) lemma shiftr_mask_eq: "(x >> n) AND mask (size x - n) = x >> n" for x :: "'a :: len word" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_drop_bit) done lemma shiftr_mask_eq': "m = (size x - n) \ (x >> n) AND mask m = x >> n" for x :: "'a :: len word" by (simp add: shiftr_mask_eq) lemma and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x AND y) = 0) = (\ (x !! n))" apply safe apply (drule_tac u="(x AND (1 << n))" and x=n in word_eqD) apply (simp add: nth_w2p) apply (simp add: test_bit_bin) apply (rule bit_word_eqI) apply (auto simp add: bit_simps test_bit_eq_bit) done lemma and_neq_0_is_nth: \x AND y \ 0 \ x !! n\ if \y = 2 ^ n\ for x y :: \'a::len word\ apply (simp add: bit_eq_iff bit_simps) using that apply (simp add: bit_simps not_le) apply transfer apply auto done lemma nth_is_and_neq_0: "(x::'a::len word) !! n = (x AND 2 ^ n \ 0)" by (subst and_neq_0_is_nth; rule refl) lemma word_shift_zero: "\ x << n = 0; x \ 2^m; m + n < LENGTH('a)\ \ (x::'a::len word) = 0" apply (rule ccontr) apply (drule (2) word_shift_nonzero) apply simp done lemma mask_shift_and_negate[simp]:"(w AND mask n << m) AND NOT (mask n << m) = 0" for w :: \'a::len word\ by (clarsimp simp add: mask_eq_decr_exp Parity.bit_eq_iff bit_and_iff bit_not_iff shiftl_word_eq bit_push_bit_iff) end