diff --git a/src/HOL/Computational_Algebra/Polynomial_Factorial.thy b/src/HOL/Computational_Algebra/Polynomial_Factorial.thy --- a/src/HOL/Computational_Algebra/Polynomial_Factorial.thy +++ b/src/HOL/Computational_Algebra/Polynomial_Factorial.thy @@ -1,869 +1,869 @@ (* Title: HOL/Computational_Algebra/Polynomial_Factorial.thy Author: Manuel Eberl *) section \Polynomials, fractions and rings\ theory Polynomial_Factorial imports Complex_Main Polynomial Normalized_Fraction begin subsection \Lifting elements into the field of fractions\ definition to_fract :: "'a :: idom \ 'a fract" where "to_fract x = Fract x 1" \ \FIXME: more idiomatic name, abbreviation\ lemma to_fract_0 [simp]: "to_fract 0 = 0" by (simp add: to_fract_def eq_fract Zero_fract_def) lemma to_fract_1 [simp]: "to_fract 1 = 1" by (simp add: to_fract_def eq_fract One_fract_def) lemma to_fract_add [simp]: "to_fract (x + y) = to_fract x + to_fract y" by (simp add: to_fract_def) lemma to_fract_diff [simp]: "to_fract (x - y) = to_fract x - to_fract y" by (simp add: to_fract_def) lemma to_fract_uminus [simp]: "to_fract (-x) = -to_fract x" by (simp add: to_fract_def) lemma to_fract_mult [simp]: "to_fract (x * y) = to_fract x * to_fract y" by (simp add: to_fract_def) lemma to_fract_eq_iff [simp]: "to_fract x = to_fract y \ x = y" by (simp add: to_fract_def eq_fract) lemma to_fract_eq_0_iff [simp]: "to_fract x = 0 \ x = 0" by (simp add: to_fract_def Zero_fract_def eq_fract) lemma snd_quot_of_fract_nonzero [simp]: "snd (quot_of_fract x) \ 0" by transfer simp lemma Fract_quot_of_fract [simp]: "Fract (fst (quot_of_fract x)) (snd (quot_of_fract x)) = x" by transfer (simp del: fractrel_iff, subst fractrel_normalize_quot_left, simp) lemma to_fract_quot_of_fract: assumes "snd (quot_of_fract x) = 1" shows "to_fract (fst (quot_of_fract x)) = x" proof - have "x = Fract (fst (quot_of_fract x)) (snd (quot_of_fract x))" by simp also note assms finally show ?thesis by (simp add: to_fract_def) qed lemma snd_quot_of_fract_Fract_whole: assumes "y dvd x" shows "snd (quot_of_fract (Fract x y)) = 1" using assms by transfer (auto simp: normalize_quot_def Let_def gcd_proj2_if_dvd) lemma Fract_conv_to_fract: "Fract a b = to_fract a / to_fract b" by (simp add: to_fract_def) lemma quot_of_fract_to_fract [simp]: "quot_of_fract (to_fract x) = (x, 1)" unfolding to_fract_def by transfer (simp add: normalize_quot_def) lemma fst_quot_of_fract_eq_0_iff [simp]: "fst (quot_of_fract x) = 0 \ x = 0" by transfer simp lemma snd_quot_of_fract_to_fract [simp]: "snd (quot_of_fract (to_fract x)) = 1" unfolding to_fract_def by (rule snd_quot_of_fract_Fract_whole) simp_all lemma coprime_quot_of_fract: "coprime (fst (quot_of_fract x)) (snd (quot_of_fract x))" by transfer (simp add: coprime_normalize_quot) lemma unit_factor_snd_quot_of_fract: "unit_factor (snd (quot_of_fract x)) = 1" using quot_of_fract_in_normalized_fracts[of x] by (simp add: normalized_fracts_def case_prod_unfold) lemma unit_factor_1_imp_normalized: "unit_factor x = 1 \ normalize x = x" by (subst (2) normalize_mult_unit_factor [symmetric, of x]) (simp del: normalize_mult_unit_factor) lemma normalize_snd_quot_of_fract: "normalize (snd (quot_of_fract x)) = snd (quot_of_fract x)" by (intro unit_factor_1_imp_normalized unit_factor_snd_quot_of_fract) subsection \Lifting polynomial coefficients to the field of fractions\ abbreviation (input) fract_poly where "fract_poly \ map_poly to_fract" abbreviation (input) unfract_poly where "unfract_poly \ map_poly (fst \ quot_of_fract)" lemma fract_poly_smult [simp]: "fract_poly (smult c p) = smult (to_fract c) (fract_poly p)" by (simp add: smult_conv_map_poly map_poly_map_poly o_def) lemma fract_poly_0 [simp]: "fract_poly 0 = 0" by (simp add: poly_eqI coeff_map_poly) lemma fract_poly_1 [simp]: "fract_poly 1 = 1" by (simp add: map_poly_pCons) lemma fract_poly_add [simp]: "fract_poly (p + q) = fract_poly p + fract_poly q" by (intro poly_eqI) (simp_all add: coeff_map_poly) lemma fract_poly_diff [simp]: "fract_poly (p - q) = fract_poly p - fract_poly q" by (intro poly_eqI) (simp_all add: coeff_map_poly) lemma to_fract_sum [simp]: "to_fract (sum f A) = sum (\x. to_fract (f x)) A" by (cases "finite A", induction A rule: finite_induct) simp_all lemma fract_poly_mult [simp]: "fract_poly (p * q) = fract_poly p * fract_poly q" by (intro poly_eqI) (simp_all add: coeff_map_poly coeff_mult) lemma fract_poly_eq_iff [simp]: "fract_poly p = fract_poly q \ p = q" by (auto simp: poly_eq_iff coeff_map_poly) lemma fract_poly_eq_0_iff [simp]: "fract_poly p = 0 \ p = 0" using fract_poly_eq_iff[of p 0] by (simp del: fract_poly_eq_iff) lemma fract_poly_dvd: "p dvd q \ fract_poly p dvd fract_poly q" by (auto elim!: dvdE) lemma prod_mset_fract_poly: "(\x\#A. map_poly to_fract (f x)) = fract_poly (prod_mset (image_mset f A))" by (induct A) (simp_all add: ac_simps) lemma is_unit_fract_poly_iff: "p dvd 1 \ fract_poly p dvd 1 \ content p = 1" proof safe assume A: "p dvd 1" with fract_poly_dvd [of p 1] show "is_unit (fract_poly p)" by simp from A show "content p = 1" by (auto simp: is_unit_poly_iff normalize_1_iff) next assume A: "fract_poly p dvd 1" and B: "content p = 1" from A obtain c where c: "fract_poly p = [:c:]" by (auto simp: is_unit_poly_iff) { fix n :: nat assume "n > 0" have "to_fract (coeff p n) = coeff (fract_poly p) n" by (simp add: coeff_map_poly) also note c also from \n > 0\ have "coeff [:c:] n = 0" by (simp add: coeff_pCons split: nat.splits) finally have "coeff p n = 0" by simp } hence "degree p \ 0" by (intro degree_le) simp_all with B show "p dvd 1" by (auto simp: is_unit_poly_iff normalize_1_iff elim!: degree_eq_zeroE) qed lemma fract_poly_is_unit: "p dvd 1 \ fract_poly p dvd 1" using fract_poly_dvd[of p 1] by simp lemma fract_poly_smult_eqE: fixes c :: "'a :: {idom_divide,ring_gcd,semiring_gcd_mult_normalize} fract" assumes "fract_poly p = smult c (fract_poly q)" obtains a b where "c = to_fract b / to_fract a" "smult a p = smult b q" "coprime a b" "normalize a = a" proof - define a b where "a = fst (quot_of_fract c)" and "b = snd (quot_of_fract c)" have "smult (to_fract a) (fract_poly q) = smult (to_fract b) (fract_poly p)" by (subst smult_eq_iff) (simp_all add: a_def b_def Fract_conv_to_fract [symmetric] assms) hence "fract_poly (smult a q) = fract_poly (smult b p)" by (simp del: fract_poly_eq_iff) hence "smult b p = smult a q" by (simp only: fract_poly_eq_iff) moreover have "c = to_fract a / to_fract b" "coprime b a" "normalize b = b" by (simp_all add: a_def b_def coprime_quot_of_fract [of c] ac_simps normalize_snd_quot_of_fract Fract_conv_to_fract [symmetric]) ultimately show ?thesis by (intro that[of a b]) qed subsection \Fractional content\ abbreviation (input) Lcm_coeff_denoms :: "'a :: {semiring_Gcd,idom_divide,ring_gcd,semiring_gcd_mult_normalize} fract poly \ 'a" where "Lcm_coeff_denoms p \ Lcm (snd ` quot_of_fract ` set (coeffs p))" definition fract_content :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} fract poly \ 'a fract" where "fract_content p = (let d = Lcm_coeff_denoms p in Fract (content (unfract_poly (smult (to_fract d) p))) d)" definition primitive_part_fract :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} fract poly \ 'a poly" where "primitive_part_fract p = primitive_part (unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p))" lemma primitive_part_fract_0 [simp]: "primitive_part_fract 0 = 0" by (simp add: primitive_part_fract_def) lemma fract_content_eq_0_iff [simp]: "fract_content p = 0 \ p = 0" unfolding fract_content_def Let_def Zero_fract_def by (subst eq_fract) (auto simp: Lcm_0_iff map_poly_eq_0_iff) lemma content_primitive_part_fract [simp]: fixes p :: "'a :: {semiring_gcd_mult_normalize, factorial_semiring, ring_gcd, semiring_Gcd,idom_divide} fract poly" shows "p \ 0 \ content (primitive_part_fract p) = 1" unfolding primitive_part_fract_def by (rule content_primitive_part) (auto simp: primitive_part_fract_def map_poly_eq_0_iff Lcm_0_iff) lemma content_times_primitive_part_fract: "smult (fract_content p) (fract_poly (primitive_part_fract p)) = p" proof - define p' where "p' = unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p)" have "fract_poly p' = map_poly (to_fract \ fst \ quot_of_fract) (smult (to_fract (Lcm_coeff_denoms p)) p)" unfolding primitive_part_fract_def p'_def by (subst map_poly_map_poly) (simp_all add: o_assoc) also have "\ = smult (to_fract (Lcm_coeff_denoms p)) p" proof (intro map_poly_idI, unfold o_apply) fix c assume "c \ set (coeffs (smult (to_fract (Lcm_coeff_denoms p)) p))" then obtain c' where c: "c' \ set (coeffs p)" "c = to_fract (Lcm_coeff_denoms p) * c'" by (auto simp add: Lcm_0_iff coeffs_smult split: if_splits) note c(2) also have "c' = Fract (fst (quot_of_fract c')) (snd (quot_of_fract c'))" by simp also have "to_fract (Lcm_coeff_denoms p) * \ = Fract (Lcm_coeff_denoms p * fst (quot_of_fract c')) (snd (quot_of_fract c'))" unfolding to_fract_def by (subst mult_fract) simp_all also have "snd (quot_of_fract \) = 1" by (intro snd_quot_of_fract_Fract_whole dvd_mult2 dvd_Lcm) (insert c(1), auto) finally show "to_fract (fst (quot_of_fract c)) = c" by (rule to_fract_quot_of_fract) qed also have "p' = smult (content p') (primitive_part p')" by (rule content_times_primitive_part [symmetric]) also have "primitive_part p' = primitive_part_fract p" by (simp add: primitive_part_fract_def p'_def) also have "fract_poly (smult (content p') (primitive_part_fract p)) = smult (to_fract (content p')) (fract_poly (primitive_part_fract p))" by simp finally have "smult (to_fract (content p')) (fract_poly (primitive_part_fract p)) = smult (to_fract (Lcm_coeff_denoms p)) p" . thus ?thesis by (subst (asm) smult_eq_iff) (auto simp add: Let_def p'_def Fract_conv_to_fract field_simps Lcm_0_iff fract_content_def) qed lemma fract_content_fract_poly [simp]: "fract_content (fract_poly p) = to_fract (content p)" proof - have "Lcm_coeff_denoms (fract_poly p) = 1" by (auto simp: set_coeffs_map_poly) hence "fract_content (fract_poly p) = to_fract (content (map_poly (fst \ quot_of_fract \ to_fract) p))" by (simp add: fract_content_def to_fract_def fract_collapse map_poly_map_poly del: Lcm_1_iff) also have "map_poly (fst \ quot_of_fract \ to_fract) p = p" by (intro map_poly_idI) simp_all finally show ?thesis . qed lemma content_decompose_fract: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide, semiring_gcd_mult_normalize} fract poly" obtains c p' where "p = smult c (map_poly to_fract p')" "content p' = 1" proof (cases "p = 0") case True hence "p = smult 0 (map_poly to_fract 1)" "content 1 = 1" by simp_all thus ?thesis .. next case False thus ?thesis by (rule that[OF content_times_primitive_part_fract [symmetric] content_primitive_part_fract]) qed subsection \More properties of content and primitive part\ lemma lift_prime_elem_poly: assumes "prime_elem (c :: 'a :: semidom)" shows "prime_elem [:c:]" proof (rule prime_elemI) fix a b assume *: "[:c:] dvd a * b" from * have dvd: "c dvd coeff (a * b) n" for n by (subst (asm) const_poly_dvd_iff) blast { define m where "m = (GREATEST m. \c dvd coeff b m)" assume "\[:c:] dvd b" hence A: "\i. \c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast have B: "\i. \c dvd coeff b i \ i \ degree b" by (auto intro: le_degree) have coeff_m: "\c dvd coeff b m" unfolding m_def by (rule GreatestI_ex_nat[OF A B]) have "i \ m" if "\c dvd coeff b i" for i unfolding m_def by (blast intro!: Greatest_le_nat that B) hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force have "c dvd coeff a i" for i proof (induction i rule: nat_descend_induct[of "degree a"]) case (base i) thus ?case by (simp add: coeff_eq_0) next case (descend i) let ?A = "{..i+m} - {i}" have "c dvd coeff (a * b) (i + m)" by (rule dvd) also have "coeff (a * b) (i + m) = (\k\i + m. coeff a k * coeff b (i + m - k))" by (simp add: coeff_mult) also have "{..i+m} = insert i ?A" by auto also have "(\k\\. coeff a k * coeff b (i + m - k)) = coeff a i * coeff b m + (\k\?A. coeff a k * coeff b (i + m - k))" (is "_ = _ + ?S") by (subst sum.insert) simp_all finally have eq: "c dvd coeff a i * coeff b m + ?S" . moreover have "c dvd ?S" proof (rule dvd_sum) fix k assume k: "k \ {..i+m} - {i}" show "c dvd coeff a k * coeff b (i + m - k)" proof (cases "k < i") case False with k have "c dvd coeff a k" by (intro descend.IH) simp thus ?thesis by simp next case True hence "c dvd coeff b (i + m - k)" by (intro dvd_b) simp thus ?thesis by simp qed qed ultimately have "c dvd coeff a i * coeff b m" by (simp add: dvd_add_left_iff) with assms coeff_m show "c dvd coeff a i" by (simp add: prime_elem_dvd_mult_iff) qed hence "[:c:] dvd a" by (subst const_poly_dvd_iff) blast } then show "[:c:] dvd a \ [:c:] dvd b" by blast next from assms show "[:c:] \ 0" and "\ [:c:] dvd 1" by (simp_all add: prime_elem_def is_unit_const_poly_iff) qed lemma prime_elem_const_poly_iff: fixes c :: "'a :: semidom" shows "prime_elem [:c:] \ prime_elem c" proof assume A: "prime_elem [:c:]" show "prime_elem c" proof (rule prime_elemI) fix a b assume "c dvd a * b" hence "[:c:] dvd [:a:] * [:b:]" by (simp add: mult_ac) from A and this have "[:c:] dvd [:a:] \ [:c:] dvd [:b:]" by (rule prime_elem_dvd_multD) thus "c dvd a \ c dvd b" by simp qed (insert A, auto simp: prime_elem_def is_unit_poly_iff) qed (auto intro: lift_prime_elem_poly) lemma fract_poly_dvdD: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide, semiring_gcd_mult_normalize} poly" assumes "fract_poly p dvd fract_poly q" "content p = 1" shows "p dvd q" proof - from assms(1) obtain r where r: "fract_poly q = fract_poly p * r" by (erule dvdE) from content_decompose_fract[of r] guess c r' . note r' = this from r r' have eq: "fract_poly q = smult c (fract_poly (p * r'))" by simp from fract_poly_smult_eqE[OF this] guess a b . note ab = this have "content (smult a q) = content (smult b (p * r'))" by (simp only: ab(2)) hence eq': "normalize b = a * content q" by (simp add: assms content_mult r' ab(4)) have "1 = gcd a (normalize b)" by (simp add: ab) also note eq' also have "gcd a (a * content q) = a" by (simp add: gcd_proj1_if_dvd ab(4)) finally have [simp]: "a = 1" by simp from eq ab have "q = p * ([:b:] * r')" by simp thus ?thesis by (rule dvdI) qed subsection \Polynomials over a field are a Euclidean ring\ context begin interpretation field_poly: normalization_euclidean_semiring_multiplicative where zero = "0 :: 'a :: field poly" and one = 1 and plus = plus and minus = minus and times = times and normalize = "\p. smult (inverse (lead_coeff p)) p" and unit_factor = "\p. [:lead_coeff p:]" and euclidean_size = "\p. if p = 0 then 0 else 2 ^ degree p" and divide = divide and modulo = modulo rewrites "dvd.dvd (times :: 'a poly \ _) = Rings.dvd" and "comm_monoid_mult.prod_mset times 1 = prod_mset" and "comm_semiring_1.irreducible times 1 0 = irreducible" and "comm_semiring_1.prime_elem times 1 0 = prime_elem" proof - show "dvd.dvd (times :: 'a poly \ _) = Rings.dvd" by (simp add: dvd_dict) show "comm_monoid_mult.prod_mset times 1 = prod_mset" by (simp add: prod_mset_dict) show "comm_semiring_1.irreducible times 1 0 = irreducible" by (simp add: irreducible_dict) show "comm_semiring_1.prime_elem times 1 0 = prime_elem" by (simp add: prime_elem_dict) show "class.normalization_euclidean_semiring_multiplicative divide plus minus (0 :: 'a poly) times 1 modulo (\p. if p = 0 then 0 else 2 ^ degree p) (\p. [:lead_coeff p:]) (\p. smult (inverse (lead_coeff p)) p)" proof (standard, fold dvd_dict) fix p :: "'a poly" show "[:lead_coeff p:] * smult (inverse (lead_coeff p)) p = p" by (cases "p = 0") simp_all next fix p :: "'a poly" assume "is_unit p" then show "[:lead_coeff p:] = p" by (elim is_unit_polyE) (auto simp: monom_0 one_poly_def field_simps) next fix p :: "'a poly" assume "p \ 0" then show "is_unit [:lead_coeff p:]" by (simp add: is_unit_pCons_iff) next fix a b :: "'a poly" assume "is_unit a" thus "[:lead_coeff (a * b):] = a * [:lead_coeff b:]" by (auto elim!: is_unit_polyE) qed (auto simp: lead_coeff_mult Rings.div_mult_mod_eq intro!: degree_mod_less' degree_mult_right_le) qed lemma field_poly_irreducible_imp_prime: "prime_elem p" if "irreducible p" for p :: "'a :: field poly" using that by (fact field_poly.irreducible_imp_prime_elem) -find_theorems name:prod_mset_prime_factorization + lemma field_poly_prod_mset_prime_factorization: "prod_mset (field_poly.prime_factorization p) = smult (inverse (lead_coeff p)) p" if "p \ 0" for p :: "'a :: field poly" using that by (fact field_poly.prod_mset_prime_factorization) lemma field_poly_in_prime_factorization_imp_prime: "prime_elem p" if "p \# field_poly.prime_factorization x" for p :: "'a :: field poly" by (rule field_poly.prime_imp_prime_elem, rule field_poly.in_prime_factors_imp_prime) (fact that) subsection \Primality and irreducibility in polynomial rings\ lemma nonconst_poly_irreducible_iff: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "degree p \ 0" shows "irreducible p \ irreducible (fract_poly p) \ content p = 1" proof safe assume p: "irreducible p" from content_decompose[of p] guess p' . note p' = this hence "p = [:content p:] * p'" by simp from p this have "[:content p:] dvd 1 \ p' dvd 1" by (rule irreducibleD) moreover have "\p' dvd 1" proof assume "p' dvd 1" hence "degree p = 0" by (subst p') (auto simp: is_unit_poly_iff) with assms show False by contradiction qed ultimately show [simp]: "content p = 1" by (simp add: is_unit_const_poly_iff) show "irreducible (map_poly to_fract p)" proof (rule irreducibleI) have "fract_poly p = 0 \ p = 0" by (intro map_poly_eq_0_iff) auto with assms show "map_poly to_fract p \ 0" by auto next show "\is_unit (fract_poly p)" proof assume "is_unit (map_poly to_fract p)" hence "degree (map_poly to_fract p) = 0" by (auto simp: is_unit_poly_iff) hence "degree p = 0" by (simp add: degree_map_poly) with assms show False by contradiction qed next fix q r assume qr: "fract_poly p = q * r" from content_decompose_fract[of q] guess cg q' . note q = this from content_decompose_fract[of r] guess cr r' . note r = this from qr q r p have nz: "cg \ 0" "cr \ 0" by auto from qr have eq: "fract_poly p = smult (cr * cg) (fract_poly (q' * r'))" by (simp add: q r) from fract_poly_smult_eqE[OF this] guess a b . note ab = this hence "content (smult a p) = content (smult b (q' * r'))" by (simp only:) with ab(4) have a: "a = normalize b" by (simp add: content_mult q r) then have "normalize b = gcd a b" by simp with \coprime a b\ have "normalize b = 1" by simp then have "a = 1" "is_unit b" by (simp_all add: a normalize_1_iff) note eq also from ab(1) \a = 1\ have "cr * cg = to_fract b" by simp also have "smult \ (fract_poly (q' * r')) = fract_poly (smult b (q' * r'))" by simp finally have "p = ([:b:] * q') * r'" by (simp del: fract_poly_smult) from p and this have "([:b:] * q') dvd 1 \ r' dvd 1" by (rule irreducibleD) hence "q' dvd 1 \ r' dvd 1" by (auto dest: dvd_mult_right simp del: mult_pCons_left) hence "fract_poly q' dvd 1 \ fract_poly r' dvd 1" by (auto simp: fract_poly_is_unit) with q r show "is_unit q \ is_unit r" by (auto simp add: is_unit_smult_iff dvd_field_iff nz) qed next assume irred: "irreducible (fract_poly p)" and primitive: "content p = 1" show "irreducible p" proof (rule irreducibleI) from irred show "p \ 0" by auto next from irred show "\p dvd 1" by (auto simp: irreducible_def dest: fract_poly_is_unit) next fix q r assume qr: "p = q * r" hence "fract_poly p = fract_poly q * fract_poly r" by simp from irred and this have "fract_poly q dvd 1 \ fract_poly r dvd 1" by (rule irreducibleD) with primitive qr show "q dvd 1 \ r dvd 1" by (auto simp: content_prod_eq_1_iff is_unit_fract_poly_iff) qed qed private lemma irreducible_imp_prime_poly: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "irreducible p" shows "prime_elem p" proof (cases "degree p = 0") case True with assms show ?thesis by (auto simp: prime_elem_const_poly_iff irreducible_const_poly_iff intro!: irreducible_imp_prime_elem elim!: degree_eq_zeroE) next case False from assms False have irred: "irreducible (fract_poly p)" and primitive: "content p = 1" by (simp_all add: nonconst_poly_irreducible_iff) from irred have prime: "prime_elem (fract_poly p)" by (rule field_poly_irreducible_imp_prime) show ?thesis proof (rule prime_elemI) fix q r assume "p dvd q * r" hence "fract_poly p dvd fract_poly (q * r)" by (rule fract_poly_dvd) hence "fract_poly p dvd fract_poly q * fract_poly r" by simp from prime and this have "fract_poly p dvd fract_poly q \ fract_poly p dvd fract_poly r" by (rule prime_elem_dvd_multD) with primitive show "p dvd q \ p dvd r" by (auto dest: fract_poly_dvdD) qed (insert assms, auto simp: irreducible_def) qed lemma degree_primitive_part_fract [simp]: "degree (primitive_part_fract p) = degree p" proof - have "p = smult (fract_content p) (fract_poly (primitive_part_fract p))" by (simp add: content_times_primitive_part_fract) also have "degree \ = degree (primitive_part_fract p)" by (auto simp: degree_map_poly) finally show ?thesis .. qed lemma irreducible_primitive_part_fract: fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd,semiring_gcd_mult_normalize} fract poly" assumes "irreducible p" shows "irreducible (primitive_part_fract p)" proof - from assms have deg: "degree (primitive_part_fract p) \ 0" by (intro notI) (auto elim!: degree_eq_zeroE simp: irreducible_def is_unit_poly_iff dvd_field_iff) hence [simp]: "p \ 0" by auto note \irreducible p\ also have "p = [:fract_content p:] * fract_poly (primitive_part_fract p)" by (simp add: content_times_primitive_part_fract) also have "irreducible \ \ irreducible (fract_poly (primitive_part_fract p))" by (intro irreducible_mult_unit_left) (simp_all add: is_unit_poly_iff dvd_field_iff) finally show ?thesis using deg by (simp add: nonconst_poly_irreducible_iff) qed lemma prime_elem_primitive_part_fract: fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd,semiring_gcd_mult_normalize} fract poly" shows "irreducible p \ prime_elem (primitive_part_fract p)" by (intro irreducible_imp_prime_poly irreducible_primitive_part_fract) lemma irreducible_linear_field_poly: fixes a b :: "'a::field" assumes "b \ 0" shows "irreducible [:a,b:]" proof (rule irreducibleI) fix p q assume pq: "[:a,b:] = p * q" also from pq assms have "degree \ = degree p + degree q" by (intro degree_mult_eq) auto finally have "degree p = 0 \ degree q = 0" using assms by auto with assms pq show "is_unit p \ is_unit q" by (auto simp: is_unit_const_poly_iff dvd_field_iff elim!: degree_eq_zeroE) qed (insert assms, auto simp: is_unit_poly_iff) lemma prime_elem_linear_field_poly: "(b :: 'a :: field) \ 0 \ prime_elem [:a,b:]" by (rule field_poly_irreducible_imp_prime, rule irreducible_linear_field_poly) lemma irreducible_linear_poly: fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd,semiring_gcd_mult_normalize}" shows "b \ 0 \ coprime a b \ irreducible [:a,b:]" by (auto intro!: irreducible_linear_field_poly simp: nonconst_poly_irreducible_iff content_def map_poly_pCons) lemma prime_elem_linear_poly: fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd,semiring_gcd_mult_normalize}" shows "b \ 0 \ coprime a b \ prime_elem [:a,b:]" by (rule irreducible_imp_prime_poly, rule irreducible_linear_poly) subsection \Prime factorisation of polynomials\ private lemma poly_prime_factorization_exists_content_1: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "p \ 0" "content p = 1" shows "\A. (\p. p \# A \ prime_elem p) \ prod_mset A = normalize p" proof - let ?P = "field_poly.prime_factorization (fract_poly p)" define c where "c = prod_mset (image_mset fract_content ?P)" define c' where "c' = c * to_fract (lead_coeff p)" define e where "e = prod_mset (image_mset primitive_part_fract ?P)" define A where "A = image_mset (normalize \ primitive_part_fract) ?P" have "content e = (\x\#field_poly.prime_factorization (map_poly to_fract p). content (primitive_part_fract x))" by (simp add: e_def content_prod_mset multiset.map_comp o_def) also have "image_mset (\x. content (primitive_part_fract x)) ?P = image_mset (\_. 1) ?P" by (intro image_mset_cong content_primitive_part_fract) auto finally have content_e: "content e = 1" by simp from \p \ 0\ have "fract_poly p = [:lead_coeff (fract_poly p):] * smult (inverse (lead_coeff (fract_poly p))) (fract_poly p)" by simp also have "[:lead_coeff (fract_poly p):] = [:to_fract (lead_coeff p):]" by (simp add: monom_0 degree_map_poly coeff_map_poly) also from assms have "smult (inverse (lead_coeff (fract_poly p))) (fract_poly p) = prod_mset ?P" by (subst field_poly_prod_mset_prime_factorization) simp_all also have "\ = prod_mset (image_mset id ?P)" by simp also have "image_mset id ?P = image_mset (\x. [:fract_content x:] * fract_poly (primitive_part_fract x)) ?P" by (intro image_mset_cong) (auto simp: content_times_primitive_part_fract) also have "prod_mset \ = smult c (fract_poly e)" by (subst prod_mset.distrib) (simp_all add: prod_mset_fract_poly prod_mset_const_poly c_def e_def) also have "[:to_fract (lead_coeff p):] * \ = smult c' (fract_poly e)" by (simp add: c'_def) finally have eq: "fract_poly p = smult c' (fract_poly e)" . also obtain b where b: "c' = to_fract b" "is_unit b" proof - from fract_poly_smult_eqE[OF eq] guess a b . note ab = this from ab(2) have "content (smult a p) = content (smult b e)" by (simp only: ) with assms content_e have "a = normalize b" by (simp add: ab(4)) with ab have ab': "a = 1" "is_unit b" by (simp_all add: normalize_1_iff) with ab ab' have "c' = to_fract b" by auto from this and \is_unit b\ show ?thesis by (rule that) qed hence "smult c' (fract_poly e) = fract_poly (smult b e)" by simp finally have "p = smult b e" by (simp only: fract_poly_eq_iff) hence "p = [:b:] * e" by simp with b have "normalize p = normalize e" by (simp only: normalize_mult) (simp add: is_unit_normalize is_unit_poly_iff) also have "normalize e = prod_mset A" by (simp add: multiset.map_comp e_def A_def normalize_prod_mset) finally have "prod_mset A = normalize p" .. have "prime_elem p" if "p \# A" for p using that by (auto simp: A_def prime_elem_primitive_part_fract prime_elem_imp_irreducible dest!: field_poly_in_prime_factorization_imp_prime ) from this and \prod_mset A = normalize p\ show ?thesis by (intro exI[of _ A]) blast qed lemma poly_prime_factorization_exists: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "p \ 0" shows "\A. (\p. p \# A \ prime_elem p) \ normalize (prod_mset A) = normalize p" proof - define B where "B = image_mset (\x. [:x:]) (prime_factorization (content p))" have "\A. (\p. p \# A \ prime_elem p) \ prod_mset A = normalize (primitive_part p)" by (rule poly_prime_factorization_exists_content_1) (insert assms, simp_all) then guess A by (elim exE conjE) note A = this have "normalize (prod_mset (A + B)) = normalize (prod_mset A * normalize (prod_mset B))" by simp also from assms have "normalize (prod_mset B) = normalize [:content p:]" by (simp add: prod_mset_const_poly normalize_const_poly prod_mset_prime_factorization_weak B_def) also have "prod_mset A = normalize (primitive_part p)" using A by simp finally have "normalize (prod_mset (A + B)) = normalize (primitive_part p * [:content p:])" by simp moreover have "\p. p \# B \ prime_elem p" by (auto simp: B_def intro!: lift_prime_elem_poly dest: in_prime_factors_imp_prime) ultimately show ?thesis using A by (intro exI[of _ "A + B"]) (auto) qed end subsection \Typeclass instances\ instance poly :: ("{factorial_ring_gcd,semiring_gcd_mult_normalize}") factorial_semiring by standard (rule poly_prime_factorization_exists) instantiation poly :: ("{factorial_ring_gcd, semiring_gcd_mult_normalize}") factorial_ring_gcd begin definition gcd_poly :: "'a poly \ 'a poly \ 'a poly" where [code del]: "gcd_poly = gcd_factorial" definition lcm_poly :: "'a poly \ 'a poly \ 'a poly" where [code del]: "lcm_poly = lcm_factorial" definition Gcd_poly :: "'a poly set \ 'a poly" where [code del]: "Gcd_poly = Gcd_factorial" definition Lcm_poly :: "'a poly set \ 'a poly" where [code del]: "Lcm_poly = Lcm_factorial" instance by standard (simp_all add: gcd_poly_def lcm_poly_def Gcd_poly_def Lcm_poly_def) end instance poly :: ("{factorial_ring_gcd, semiring_gcd_mult_normalize}") semiring_gcd_mult_normalize .. instantiation poly :: ("{field,factorial_ring_gcd,semiring_gcd_mult_normalize}") "{unique_euclidean_ring, normalization_euclidean_semiring}" begin definition euclidean_size_poly :: "'a poly \ nat" where "euclidean_size_poly p = (if p = 0 then 0 else 2 ^ degree p)" definition division_segment_poly :: "'a poly \ 'a poly" where [simp]: "division_segment_poly p = 1" instance proof show "(q * p + r) div p = q" if "p \ 0" and "euclidean_size r < euclidean_size p" for q p r :: "'a poly" proof - from \p \ 0\ eucl_rel_poly [of r p] have "eucl_rel_poly (r + q * p) p (q + r div p, r mod p)" by (simp add: eucl_rel_poly_iff distrib_right) then have "(r + q * p) div p = q + r div p" by (rule div_poly_eq) with that show ?thesis by (cases "r = 0") (simp_all add: euclidean_size_poly_def div_poly_less ac_simps) qed qed (auto simp: euclidean_size_poly_def Rings.div_mult_mod_eq div_poly_less degree_mult_eq power_add intro!: degree_mod_less' split: if_splits) end instance poly :: ("{field, normalization_euclidean_semiring, factorial_ring_gcd, semiring_gcd_mult_normalize}") euclidean_ring_gcd by (rule euclidean_ring_gcd_class.intro, rule factorial_euclidean_semiring_gcdI) standard subsection \Polynomial GCD\ lemma gcd_poly_decompose: fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" shows "gcd p q = smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))" proof (rule sym, rule gcdI) have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd [:content p:] * primitive_part p" by (intro mult_dvd_mono) simp_all thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd p" by simp next have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd [:content q:] * primitive_part q" by (intro mult_dvd_mono) simp_all thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd q" by simp next fix d assume "d dvd p" "d dvd q" hence "[:content d:] * primitive_part d dvd [:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q)" by (intro mult_dvd_mono) auto thus "d dvd smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))" by simp qed (auto simp: normalize_smult) lemma gcd_poly_pseudo_mod: fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" assumes nz: "q \ 0" and prim: "content p = 1" "content q = 1" shows "gcd p q = gcd q (primitive_part (pseudo_mod p q))" proof - define r s where "r = fst (pseudo_divmod p q)" and "s = snd (pseudo_divmod p q)" define a where "a = [:coeff q (degree q) ^ (Suc (degree p) - degree q):]" have [simp]: "primitive_part a = unit_factor a" by (simp add: a_def unit_factor_poly_def unit_factor_power monom_0) from nz have [simp]: "a \ 0" by (auto simp: a_def) have rs: "pseudo_divmod p q = (r, s)" by (simp add: r_def s_def) have "gcd (q * r + s) q = gcd q s" using gcd_add_mult[of q r s] by (simp add: gcd.commute add_ac mult_ac) with pseudo_divmod(1)[OF nz rs] have "gcd (p * a) q = gcd q s" by (simp add: a_def) also from prim have "gcd (p * a) q = gcd p q" by (subst gcd_poly_decompose) (auto simp: primitive_part_mult gcd_mult_unit1 primitive_part_prim simp del: mult_pCons_right ) also from prim have "gcd q s = gcd q (primitive_part s)" by (subst gcd_poly_decompose) (simp_all add: primitive_part_prim) also have "s = pseudo_mod p q" by (simp add: s_def pseudo_mod_def) finally show ?thesis . qed lemma degree_pseudo_mod_less: assumes "q \ 0" "pseudo_mod p q \ 0" shows "degree (pseudo_mod p q) < degree q" using pseudo_mod(2)[of q p] assms by auto function gcd_poly_code_aux :: "'a :: factorial_ring_gcd poly \ 'a poly \ 'a poly" where "gcd_poly_code_aux p q = (if q = 0 then normalize p else gcd_poly_code_aux q (primitive_part (pseudo_mod p q)))" by auto termination by (relation "measure ((\p. if p = 0 then 0 else Suc (degree p)) \ snd)") (auto simp: degree_pseudo_mod_less) declare gcd_poly_code_aux.simps [simp del] lemma gcd_poly_code_aux_correct: assumes "content p = 1" "q = 0 \ content q = 1" shows "gcd_poly_code_aux p q = gcd p q" using assms proof (induction p q rule: gcd_poly_code_aux.induct) case (1 p q) show ?case proof (cases "q = 0") case True thus ?thesis by (subst gcd_poly_code_aux.simps) auto next case False hence "gcd_poly_code_aux p q = gcd_poly_code_aux q (primitive_part (pseudo_mod p q))" by (subst gcd_poly_code_aux.simps) simp_all also from "1.prems" False have "primitive_part (pseudo_mod p q) = 0 \ content (primitive_part (pseudo_mod p q)) = 1" by (cases "pseudo_mod p q = 0") auto with "1.prems" False have "gcd_poly_code_aux q (primitive_part (pseudo_mod p q)) = gcd q (primitive_part (pseudo_mod p q))" by (intro 1) simp_all also from "1.prems" False have "\ = gcd p q" by (intro gcd_poly_pseudo_mod [symmetric]) auto finally show ?thesis . qed qed definition gcd_poly_code :: "'a :: factorial_ring_gcd poly \ 'a poly \ 'a poly" where "gcd_poly_code p q = (if p = 0 then normalize q else if q = 0 then normalize p else smult (gcd (content p) (content q)) (gcd_poly_code_aux (primitive_part p) (primitive_part q)))" lemma gcd_poly_code [code]: "gcd p q = gcd_poly_code p q" by (simp add: gcd_poly_code_def gcd_poly_code_aux_correct gcd_poly_decompose [symmetric]) lemma lcm_poly_code [code]: fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" shows "lcm p q = normalize (p * q div gcd p q)" by (fact lcm_gcd) lemmas Gcd_poly_set_eq_fold [code] = Gcd_set_eq_fold [where ?'a = "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"] lemmas Lcm_poly_set_eq_fold [code] = Lcm_set_eq_fold [where ?'a = "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"] text \Example: @{lemma "Lcm {[:1, 2, 3:], [:2, 3, 4:]} = [:[:2:], [:7:], [:16:], [:17:], [:12 :: int:]:]" by eval} \ end diff --git a/src/HOL/Real_Asymp/Multiseries_Expansion.thy b/src/HOL/Real_Asymp/Multiseries_Expansion.thy --- a/src/HOL/Real_Asymp/Multiseries_Expansion.thy +++ b/src/HOL/Real_Asymp/Multiseries_Expansion.thy @@ -1,5390 +1,5389 @@ (* File: Multiseries_Expansion.thy Author: Manuel Eberl, TU München Asymptotic bases and Multiseries expansions of real-valued functions. Contains automation to prove limits and asymptotic relationships between such functions. *) section \Multiseries expansions\ theory Multiseries_Expansion imports "HOL-Library.BNF_Corec" "HOL-Library.Landau_Symbols" Lazy_Eval Inst_Existentials Eventuallize begin (* TODO Move *) lemma real_powr_at_top: assumes "(p::real) > 0" shows "filterlim (\x. x powr p) at_top at_top" proof (subst filterlim_cong[OF refl refl]) show "LIM x at_top. exp (p * ln x) :> at_top" by (rule filterlim_compose[OF exp_at_top filterlim_tendsto_pos_mult_at_top[OF tendsto_const]]) (simp_all add: ln_at_top assms) show "eventually (\x. x powr p = exp (p * ln x)) at_top" using eventually_gt_at_top[of 0] by eventually_elim (simp add: powr_def) qed subsection \Streams and lazy lists\ codatatype 'a msstream = MSSCons 'a "'a msstream" primrec mssnth :: "'a msstream \ nat \ 'a" where "mssnth xs 0 = (case xs of MSSCons x xs \ x)" | "mssnth xs (Suc n) = (case xs of MSSCons x xs \ mssnth xs n)" codatatype 'a msllist = MSLNil | MSLCons 'a "'a msllist" for map: mslmap fun lcoeff where "lcoeff MSLNil n = 0" | "lcoeff (MSLCons x xs) 0 = x" | "lcoeff (MSLCons x xs) (Suc n) = lcoeff xs n" primcorec msllist_of_msstream :: "'a msstream \ 'a msllist" where "msllist_of_msstream xs = (case xs of MSSCons x xs \ MSLCons x (msllist_of_msstream xs))" lemma msllist_of_msstream_MSSCons [simp]: "msllist_of_msstream (MSSCons x xs) = MSLCons x (msllist_of_msstream xs)" by (subst msllist_of_msstream.code) simp lemma lcoeff_llist_of_stream [simp]: "lcoeff (msllist_of_msstream xs) n = mssnth xs n" by (induction "msllist_of_msstream xs" n arbitrary: xs rule: lcoeff.induct; subst msllist_of_msstream.code) (auto split: msstream.splits) subsection \Convergent power series\ subsubsection \Definition\ definition convergent_powser :: "real msllist \ bool" where "convergent_powser cs \ (\R>0. \x. abs x < R \ summable (\n. lcoeff cs n * x ^ n))" lemma convergent_powser_stl: assumes "convergent_powser (MSLCons c cs)" shows "convergent_powser cs" proof - from assms obtain R where "R > 0" "\x. abs x < R \ summable (\n. lcoeff (MSLCons c cs) n * x ^ n)" unfolding convergent_powser_def by blast hence "summable (\n. lcoeff cs n * x ^ n)" if "abs x < R" for x using that by (subst (asm) summable_powser_split_head [symmetric]) simp thus ?thesis using \R > 0\ by (auto simp: convergent_powser_def) qed definition powser :: "real msllist \ real \ real" where "powser cs x = suminf (\n. lcoeff cs n * x ^ n)" lemma isCont_powser: assumes "convergent_powser cs" shows "isCont (powser cs) 0" proof - from assms obtain R where R: "R > 0" "\x. abs x < R \ summable (\n. lcoeff cs n * x ^ n)" unfolding convergent_powser_def by blast hence *: "summable (\n. lcoeff cs n * (R/2) ^ n)" by (intro R) simp_all from \R > 0\ show ?thesis unfolding powser_def by (intro isCont_powser[OF *]) simp_all qed lemma powser_MSLNil [simp]: "powser MSLNil = (\_. 0)" by (simp add: fun_eq_iff powser_def) lemma powser_MSLCons: assumes "convergent_powser (MSLCons c cs)" shows "eventually (\x. powser (MSLCons c cs) x = x * powser cs x + c) (nhds 0)" proof - from assms obtain R where R: "R > 0" "\x. abs x < R \ summable (\n. lcoeff (MSLCons c cs) n * x ^ n)" unfolding convergent_powser_def by blast from R have "powser (MSLCons c cs) x = x * powser cs x + c" if "abs x < R" for x using that unfolding powser_def by (subst powser_split_head) simp_all moreover have "eventually (\x. abs x < R) (nhds 0)" using \R > 0\ filterlim_ident[of "nhds (0::real)"] by (subst (asm) tendsto_iff) (simp add: dist_real_def) ultimately show ?thesis by (auto elim: eventually_mono) qed definition convergent_powser' :: "real msllist \ (real \ real) \ bool" where "convergent_powser' cs f \ (\R>0. \x. abs x < R \ (\n. lcoeff cs n * x ^ n) sums f x)" lemma convergent_powser'_imp_convergent_powser: "convergent_powser' cs f \ convergent_powser cs" unfolding convergent_powser_def convergent_powser'_def by (auto simp: sums_iff) lemma convergent_powser'_eventually: assumes "convergent_powser' cs f" shows "eventually (\x. powser cs x = f x) (nhds 0)" proof - from assms obtain R where "R > 0" "\x. abs x < R \ (\n. lcoeff cs n * x ^ n) sums f x" unfolding convergent_powser'_def by blast hence "powser cs x = f x" if "abs x < R" for x using that by (simp add: powser_def sums_iff) moreover have "eventually (\x. abs x < R) (nhds 0)" using \R > 0\ filterlim_ident[of "nhds (0::real)"] by (subst (asm) tendsto_iff) (simp add: dist_real_def) ultimately show ?thesis by (auto elim: eventually_mono) qed subsubsection \Geometric series\ primcorec mssalternate :: "'a \ 'a \ 'a msstream" where "mssalternate a b = MSSCons a (mssalternate b a)" lemma case_mssalternate [simp]: "(case mssalternate a b of MSSCons c d \ f c d) = f a (mssalternate b a)" by (subst mssalternate.code) simp lemma mssnth_mssalternate: "mssnth (mssalternate a b) n = (if even n then a else b)" by (induction n arbitrary: a b; subst mssalternate.code) simp_all lemma convergent_powser'_geometric: "convergent_powser' (msllist_of_msstream (mssalternate 1 (-1))) (\x. inverse (1 + x))" proof - have "(\n. lcoeff (msllist_of_msstream (mssalternate 1 (-1))) n * x ^ n) sums (inverse (1 + x))" if "abs x < 1" for x :: real proof - have "(\n. (-1) ^ n * x ^ n) sums (inverse (1 + x))" using that geometric_sums[of "-x"] by (simp add: field_simps power_mult_distrib [symmetric]) also have "(\n. (-1) ^ n * x ^ n) = (\n. lcoeff (msllist_of_msstream (mssalternate 1 (-1))) n * x ^ n)" by (auto simp add: mssnth_mssalternate fun_eq_iff) finally show ?thesis . qed thus ?thesis unfolding convergent_powser'_def by (force intro!: exI[of _ 1]) qed subsubsection \Exponential series\ primcorec exp_series_stream_aux :: "real \ real \ real msstream" where "exp_series_stream_aux acc n = MSSCons acc (exp_series_stream_aux (acc / n) (n + 1))" lemma mssnth_exp_series_stream_aux: "mssnth (exp_series_stream_aux (1 / fact n) (n + 1)) m = 1 / fact (n + m)" proof (induction m arbitrary: n) case (0 n) thus ?case by (subst exp_series_stream_aux.code) simp next case (Suc m n) from Suc.IH[of "n + 1"] show ?case by (subst exp_series_stream_aux.code) (simp add: algebra_simps) qed definition exp_series_stream :: "real msstream" where "exp_series_stream = exp_series_stream_aux 1 1" lemma mssnth_exp_series_stream: "mssnth exp_series_stream n = 1 / fact n" unfolding exp_series_stream_def using mssnth_exp_series_stream_aux[of 0 n] by simp lemma convergent_powser'_exp: "convergent_powser' (msllist_of_msstream exp_series_stream) exp" proof - have "(\n. lcoeff (msllist_of_msstream exp_series_stream) n * x ^ n) sums exp x" for x :: real using exp_converges[of x] by (simp add: mssnth_exp_series_stream field_split_simps) thus ?thesis by (auto intro: exI[of _ "1::real"] simp: convergent_powser'_def) qed subsubsection \Logarithm series\ primcorec ln_series_stream_aux :: "bool \ real \ real msstream" where "ln_series_stream_aux b n = MSSCons (if b then -inverse n else inverse n) (ln_series_stream_aux (\b) (n+1))" lemma mssnth_ln_series_stream_aux: "mssnth (ln_series_stream_aux b x) n = (if b then - ((-1)^n) * inverse (x + real n) else ((-1)^n) * inverse (x + real n))" by (induction n arbitrary: b x; subst ln_series_stream_aux.code) simp_all definition ln_series_stream :: "real msstream" where "ln_series_stream = MSSCons 0 (ln_series_stream_aux False 1)" lemma mssnth_ln_series_stream: "mssnth ln_series_stream n = (-1) ^ Suc n / real n" unfolding ln_series_stream_def by (cases n) (simp_all add: mssnth_ln_series_stream_aux field_simps) lemma ln_series': assumes "abs (x::real) < 1" shows "(\n. - ((-x)^n) / of_nat n) sums ln (1 + x)" proof - have "\n\1. norm (-((-x)^n)) / of_nat n \ norm x ^ n / 1" by (intro exI[of _ "1::nat"] allI impI frac_le) (simp_all add: power_abs) hence "\N. \n\N. norm (-((-x)^n) / of_nat n) \ norm x ^ n" by (intro exI[of _ 1]) simp_all moreover from assms have "summable (\n. norm x ^ n)" by (intro summable_geometric) simp_all ultimately have *: "summable (\n. - ((-x)^n) / of_nat n)" by (rule summable_comparison_test) moreover from assms have "0 < 1 + x" "1 + x < 2" by simp_all from ln_series[OF this] have "ln (1 + x) = (\n. - ((-x) ^ Suc n) / real (Suc n))" by (simp_all add: power_minus' mult_ac) hence "ln (1 + x) = (\n. - ((-x) ^ n / real n))" by (subst (asm) suminf_split_head[OF *]) simp_all ultimately show ?thesis by (simp add: sums_iff) qed lemma convergent_powser'_ln: "convergent_powser' (msllist_of_msstream ln_series_stream) (\x. ln (1 + x))" proof - have "(\n. lcoeff (msllist_of_msstream ln_series_stream)n * x ^ n) sums ln (1 + x)" if "abs x < 1" for x proof - from that have "(\n. - ((- x) ^ n) / real n) sums ln (1 + x)" by (rule ln_series') also have "(\n. - ((- x) ^ n) / real n) = (\n. lcoeff (msllist_of_msstream ln_series_stream) n * x ^ n)" by (auto simp: fun_eq_iff mssnth_ln_series_stream power_mult_distrib [symmetric]) finally show ?thesis . qed thus ?thesis unfolding convergent_powser'_def by (force intro!: exI[of _ 1]) qed subsubsection \Generalized binomial series\ primcorec gbinomial_series_aux :: "bool \ real \ real \ real \ real msllist" where "gbinomial_series_aux abort x n acc = (if abort \ acc = 0 then MSLNil else MSLCons acc (gbinomial_series_aux abort x (n + 1) ((x - n) / (n + 1) * acc)))" lemma gbinomial_series_abort [simp]: "gbinomial_series_aux True x n 0 = MSLNil" by (subst gbinomial_series_aux.code) simp definition gbinomial_series :: "bool \ real \ real msllist" where "gbinomial_series abort x = gbinomial_series_aux abort x 0 1" (* TODO Move *) lemma gbinomial_Suc_rec: "(x gchoose (Suc k)) = (x gchoose k) * ((x - of_nat k) / (of_nat k + 1))" proof - have "((x - 1) + 1) gchoose (Suc k) = x * (x - 1 gchoose k) / (1 + of_nat k)" by (subst gbinomial_factors) simp also have "x * (x - 1 gchoose k) = (x - of_nat k) * (x gchoose k)" by (rule gbinomial_absorb_comp [symmetric]) finally show ?thesis by (simp add: algebra_simps) qed lemma lcoeff_gbinomial_series_aux: "lcoeff (gbinomial_series_aux abort x m (x gchoose m)) n = (x gchoose (n + m))" proof (induction n arbitrary: m) case 0 show ?case by (subst gbinomial_series_aux.code) simp next case (Suc n m) show ?case proof (cases "abort \ (x gchoose m) = 0") case True with gbinomial_mult_fact[of m x] obtain k where "x = real k" "k < m" by auto hence "(x gchoose Suc (n + m)) = 0" using gbinomial_mult_fact[of "Suc (n + m)" x] by (simp add: gbinomial_altdef_of_nat) with True show ?thesis by (subst gbinomial_series_aux.code) simp next case False hence "lcoeff (gbinomial_series_aux abort x m (x gchoose m)) (Suc n) = lcoeff (gbinomial_series_aux abort x (Suc m) ((x gchoose m) * ((x - real m) / (real m + 1)))) n" by (subst gbinomial_series_aux.code) (auto simp: algebra_simps) also have "((x gchoose m) * ((x - real m) / (real m + 1))) = x gchoose (Suc m)" by (rule gbinomial_Suc_rec [symmetric]) also have "lcoeff (gbinomial_series_aux abort x (Suc m) \) n = x gchoose (n + Suc m)" by (rule Suc.IH) finally show ?thesis by simp qed qed lemma lcoeff_gbinomial_series [simp]: "lcoeff (gbinomial_series abort x) n = (x gchoose n)" using lcoeff_gbinomial_series_aux[of abort x 0 n] by (simp add: gbinomial_series_def) lemma gbinomial_ratio_limit: fixes a :: "'a :: real_normed_field" assumes "a \ \" shows "(\n. (a gchoose n) / (a gchoose Suc n)) \ -1" proof (rule Lim_transform_eventually) let ?f = "\n. inverse (a / of_nat (Suc n) - of_nat n / of_nat (Suc n))" from eventually_gt_at_top[of "0::nat"] show "eventually (\n. ?f n = (a gchoose n) /(a gchoose Suc n)) sequentially" proof eventually_elim fix n :: nat assume n: "n > 0" then obtain q where q: "n = Suc q" by (cases n) blast let ?P = "\i=0..i=0..n. a - of_nat i))" by (simp add: gbinomial_prod_rev atLeastLessThanSuc_atLeastAtMost) also from q have "(\i=0..n. a - of_nat i) = ?P * (a - of_nat n)" by (simp add: prod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost) also have "?P / \ = (?P / ?P) / (a - of_nat n)" by (rule divide_divide_eq_left[symmetric]) also from assms have "?P / ?P = 1" by auto also have "of_nat (Suc n) * (1 / (a - of_nat n)) = inverse (inverse (of_nat (Suc n)) * (a - of_nat n))" by (simp add: field_simps) also have "inverse (of_nat (Suc n)) * (a - of_nat n) = a / of_nat (Suc n) - of_nat n / of_nat (Suc n)" by (simp add: field_simps del: of_nat_Suc) finally show "?f n = (a gchoose n) / (a gchoose Suc n)" by simp qed have "(\n. norm a / (of_nat (Suc n))) \ 0" unfolding divide_inverse by (intro tendsto_mult_right_zero LIMSEQ_inverse_real_of_nat) hence "(\n. a / of_nat (Suc n)) \ 0" by (subst tendsto_norm_zero_iff[symmetric]) (simp add: norm_divide del: of_nat_Suc) hence "?f \ inverse (0 - 1)" by (intro tendsto_inverse tendsto_diff LIMSEQ_n_over_Suc_n) simp_all thus "?f \ -1" by simp qed lemma summable_gbinomial: fixes a z :: real assumes "norm z < 1" shows "summable (\n. (a gchoose n) * z ^ n)" proof (cases "z = 0 \ a \ \") case False define r where "r = (norm z + 1) / 2" from assms have r: "r > norm z" "r < 1" by (simp_all add: r_def field_simps) from False have "abs z > 0" by auto from False have "a \ \" by (auto elim!: Nats_cases) hence *: "(\n. norm (z / ((a gchoose n) / (a gchoose (Suc n))))) \ norm (z / (-1))" by (intro tendsto_norm tendsto_divide tendsto_const gbinomial_ratio_limit) simp_all hence "\\<^sub>F x in at_top. norm (z / ((a gchoose x) / (a gchoose Suc x))) > 0" using assms False by (intro order_tendstoD) auto hence nz: "\\<^sub>F x in at_top. (a gchoose x) \ 0" by eventually_elim auto have "\\<^sub>F x in at_top. norm (z / ((a gchoose x) / (a gchoose Suc x))) < r" using assms r by (intro order_tendstoD(2)[OF *]) simp_all with nz have "\\<^sub>F n in at_top. norm ((a gchoose (Suc n)) * z) \ r * norm (a gchoose n)" by eventually_elim (simp add: field_simps abs_mult split: if_splits) hence "\\<^sub>F n in at_top. norm (z ^ n) * norm ((a gchoose (Suc n)) * z) \ norm (z ^ n) * (r * norm (a gchoose n))" by eventually_elim (insert False, intro mult_left_mono, auto) hence "\\<^sub>F n in at_top. norm ((a gchoose (Suc n)) * z ^ (Suc n)) \ r * norm ((a gchoose n) * z ^ n)" by (simp add: abs_mult mult_ac) then obtain N where N: "\n. n \ N \ norm ((a gchoose (Suc n)) * z ^ (Suc n)) \ r * norm ((a gchoose n) * z ^ n)" unfolding eventually_at_top_linorder by blast show "summable (\n. (a gchoose n) * z ^ n)" by (intro summable_ratio_test[OF r(2), of N] N) next case True thus ?thesis proof assume "a \ \" then obtain n where [simp]: "a = of_nat n" by (auto elim: Nats_cases) from eventually_gt_at_top[of n] have "eventually (\n. (a gchoose n) * z ^ n = 0) at_top" by eventually_elim (simp add: binomial_gbinomial [symmetric]) from summable_cong[OF this] show ?thesis by simp qed auto qed lemma gen_binomial_real: fixes z :: real assumes "norm z < 1" shows "(\n. (a gchoose n) * z^n) sums (1 + z) powr a" proof - define K where "K = 1 - (1 - norm z) / 2" from assms have K: "K > 0" "K < 1" "norm z < K" unfolding K_def by (auto simp: field_simps intro!: add_pos_nonneg) let ?f = "\n. a gchoose n" and ?f' = "diffs (\n. a gchoose n)" have summable_strong: "summable (\n. ?f n * z ^ n)" if "norm z < 1" for z using that by (intro summable_gbinomial) with K have summable: "summable (\n. ?f n * z ^ n)" if "norm z < K" for z using that by auto hence summable': "summable (\n. ?f' n * z ^ n)" if "norm z < K" for z using that by (intro termdiff_converges[of _ K]) simp_all define f f' where [abs_def]: "f z = (\n. ?f n * z ^ n)" "f' z = (\n. ?f' n * z ^ n)" for z { fix z :: real assume z: "norm z < K" from summable_mult2[OF summable'[OF z], of z] have summable1: "summable (\n. ?f' n * z ^ Suc n)" by (simp add: mult_ac) hence summable2: "summable (\n. of_nat n * ?f n * z^n)" unfolding diffs_def by (subst (asm) summable_Suc_iff) have "(1 + z) * f' z = (\n. ?f' n * z^n) + (\n. ?f' n * z^Suc n)" unfolding f_f'_def using summable' z by (simp add: algebra_simps suminf_mult) also have "(\n. ?f' n * z^n) = (\n. of_nat (Suc n) * ?f (Suc n) * z^n)" by (intro suminf_cong) (simp add: diffs_def) also have "(\n. ?f' n * z^Suc n) = (\n. of_nat n * ?f n * z ^ n)" using summable1 suminf_split_initial_segment[OF summable1] unfolding diffs_def by (subst suminf_split_head, subst (asm) summable_Suc_iff) simp_all also have "(\n. of_nat (Suc n) * ?f (Suc n) * z^n) + (\n. of_nat n * ?f n * z^n) = (\n. a * ?f n * z^n)" by (subst gbinomial_mult_1, subst suminf_add) (insert summable'[OF z] summable2, simp_all add: summable_powser_split_head algebra_simps diffs_def) also have "\ = a * f z" unfolding f_f'_def by (subst suminf_mult[symmetric]) (simp_all add: summable[OF z] mult_ac) finally have "a * f z = (1 + z) * f' z" by simp } note deriv = this have [derivative_intros]: "(f has_field_derivative f' z) (at z)" if "norm z < of_real K" for z unfolding f_f'_def using K that by (intro termdiffs_strong[of "?f" K z] summable_strong) simp_all have "f 0 = (\n. if n = 0 then 1 else 0)" unfolding f_f'_def by (intro suminf_cong) simp also have "\ = 1" using sums_single[of 0 "\_. 1::real"] unfolding sums_iff by simp finally have [simp]: "f 0 = 1" . have "f z * (1 + z) powr (-a) = f 0 * (1 + 0) powr (-a)" proof (rule DERIV_isconst3[where ?f = "\x. f x * (1 + x) powr (-a)"]) fix z :: real assume z': "z \ {-K<.. 0" by (auto dest!: minus_unique) from z K have "norm z < 1" by simp hence "((\z. f z * (1 + z) powr (-a)) has_field_derivative f' z * (1 + z) powr (-a) - a * f z * (1 + z) powr (-a-1)) (at z)" using z by (auto intro!: derivative_eq_intros) also from z have "a * f z = (1 + z) * f' z" by (rule deriv) also have "f' z * (1 + z) powr - a - (1 + z) * f' z * (1 + z) powr (- a - 1) = 0" using \norm z < 1\ by (auto simp add: field_simps powr_diff) finally show "((\z. f z * (1 + z) powr (-a)) has_field_derivative 0) (at z)" . qed (insert K, auto) also have "f 0 * (1 + 0) powr -a = 1" by simp finally have "f z = (1 + z) powr a" using K by (simp add: field_simps dist_real_def powr_minus) with summable K show ?thesis unfolding f_f'_def by (simp add: sums_iff) qed lemma convergent_powser'_gbinomial: "convergent_powser' (gbinomial_series abort p) (\x. (1 + x) powr p)" proof - have "(\n. lcoeff (gbinomial_series abort p) n * x ^ n) sums (1 + x) powr p" if "abs x < 1" for x using that gen_binomial_real[of x p] by simp thus ?thesis unfolding convergent_powser'_def by (force intro!: exI[of _ 1]) qed lemma convergent_powser'_gbinomial': "convergent_powser' (gbinomial_series abort (real n)) (\x. (1 + x) ^ n)" proof - { fix x :: real have "(\k. if k \ {0..n} then real (n choose k) * x ^ k else 0) sums (x + 1) ^ n" using sums_If_finite_set[of "{..n}" "\k. real (n choose k) * x ^ k"] by (subst binomial_ring) simp also have "(\k. if k \ {0..n} then real (n choose k) * x ^ k else 0) = (\k. lcoeff (gbinomial_series abort (real n)) k * x ^ k)" by (simp add: fun_eq_iff binomial_gbinomial [symmetric]) finally have "\ sums (1 + x) ^ n" by (simp add: add_ac) } thus ?thesis unfolding convergent_powser'_def by (auto intro!: exI[of _ 1]) qed subsubsection \Sine and cosine\ primcorec sin_series_stream_aux :: "bool \ real \ real \ real msstream" where "sin_series_stream_aux b acc m = MSSCons (if b then -inverse acc else inverse acc) (sin_series_stream_aux (\b) (acc * (m + 1) * (m + 2)) (m + 2))" lemma mssnth_sin_series_stream_aux: "mssnth (sin_series_stream_aux b (fact m) m) n = (if b then -1 else 1) * (-1) ^ n / (fact (2 * n + m))" proof (induction n arbitrary: b m) case (0 b m) thus ?case by (subst sin_series_stream_aux.code) (simp add: field_simps) next case (Suc n b m) show ?case using Suc.IH[of "\b" "m + 2"] by (subst sin_series_stream_aux.code) (auto simp: field_simps) qed definition sin_series_stream where "sin_series_stream = sin_series_stream_aux False 1 1" lemma mssnth_sin_series_stream: "mssnth sin_series_stream n = (-1) ^ n / fact (2 * n + 1)" using mssnth_sin_series_stream_aux[of False 1 n] unfolding sin_series_stream_def by simp definition cos_series_stream where "cos_series_stream = sin_series_stream_aux False 1 0" lemma mssnth_cos_series_stream: "mssnth cos_series_stream n = (-1) ^ n / fact (2 * n)" using mssnth_sin_series_stream_aux[of False 0 n] unfolding cos_series_stream_def by simp lemma summable_pre_sin: "summable (\n. mssnth sin_series_stream n * (x::real) ^ n)" proof - have *: "summable (\n. 1 / fact n * abs x ^ n)" using exp_converges[of "abs x"] by (simp add: sums_iff field_simps) { fix n :: nat have "\x\ ^ n / fact (2 * n + 1) \ \x\ ^ n / fact n" by (intro divide_left_mono fact_mono) auto hence "norm (mssnth sin_series_stream n * x ^ n) \ 1 / fact n * abs x ^ n" by (simp add: mssnth_sin_series_stream abs_mult power_abs field_simps) } thus ?thesis by (intro summable_comparison_test[OF _ *]) (auto intro!: exI[of _ 0]) qed lemma summable_pre_cos: "summable (\n. mssnth cos_series_stream n * (x::real) ^ n)" proof - have *: "summable (\n. 1 / fact n * abs x ^ n)" using exp_converges[of "abs x"] by (simp add: sums_iff field_simps) { fix n :: nat have "\x\ ^ n / fact (2 * n) \ \x\ ^ n / fact n" by (intro divide_left_mono fact_mono) auto hence "norm (mssnth cos_series_stream n * x ^ n) \ 1 / fact n * abs x ^ n" by (simp add: mssnth_cos_series_stream abs_mult power_abs field_simps) } thus ?thesis by (intro summable_comparison_test[OF _ *]) (auto intro!: exI[of _ 0]) qed lemma cos_conv_pre_cos: "cos x = powser (msllist_of_msstream cos_series_stream) (x ^ 2)" proof - have "(\n. cos_coeff (2 * n) * x ^ (2 * n)) sums cos x" using cos_converges[of x] by (subst sums_mono_reindex[of "\n. 2 * n"]) (auto simp: strict_mono_def cos_coeff_def elim!: evenE) also have "(\n. cos_coeff (2 * n) * x ^ (2 * n)) = (\n. mssnth cos_series_stream n * (x ^ 2) ^ n)" by (simp add: fun_eq_iff mssnth_cos_series_stream cos_coeff_def power_mult) finally have sums: "(\n. mssnth cos_series_stream n * x\<^sup>2 ^ n) sums cos x" . thus ?thesis by (simp add: powser_def sums_iff) qed lemma sin_conv_pre_sin: "sin x = x * powser (msllist_of_msstream sin_series_stream) (x ^ 2)" proof - have "(\n. sin_coeff (2 * n + 1) * x ^ (2 * n + 1)) sums sin x" using sin_converges[of x] by (subst sums_mono_reindex[of "\n. 2 * n + 1"]) (auto simp: strict_mono_def sin_coeff_def elim!: oddE) also have "(\n. sin_coeff (2 * n + 1) * x ^ (2 * n + 1)) = (\n. x * mssnth sin_series_stream n * (x ^ 2) ^ n)" by (simp add: fun_eq_iff mssnth_sin_series_stream sin_coeff_def power_mult [symmetric] algebra_simps) finally have sums: "(\n. x * mssnth sin_series_stream n * x\<^sup>2 ^ n) sums sin x" . thus ?thesis using summable_pre_sin[of "x^2"] by (simp add: powser_def sums_iff suminf_mult [symmetric] mult.assoc) qed lemma convergent_powser_sin_series_stream: "convergent_powser (msllist_of_msstream sin_series_stream)" (is "convergent_powser ?cs") proof - show ?thesis using summable_pre_sin unfolding convergent_powser_def by (intro exI[of _ 1]) auto qed lemma convergent_powser_cos_series_stream: "convergent_powser (msllist_of_msstream cos_series_stream)" (is "convergent_powser ?cs") proof - show ?thesis using summable_pre_cos unfolding convergent_powser_def by (intro exI[of _ 1]) auto qed subsubsection \Arc tangent\ definition arctan_coeffs :: "nat \ 'a :: {real_normed_div_algebra, banach}" where "arctan_coeffs n = (if odd n then (-1) ^ (n div 2) / of_nat n else 0)" lemma arctan_converges: assumes "norm x < 1" shows "(\n. arctan_coeffs n * x ^ n) sums arctan x" proof - have A: "(\n. diffs arctan_coeffs n * x ^ n) sums (1 / (1 + x^2))" if "norm x < 1" for x :: real proof - let ?f = "\n. (if odd n then 0 else (-1) ^ (n div 2)) * x ^ n" from that have "norm x ^ 2 < 1 ^ 2" by (intro power_strict_mono) simp_all hence "(\n. (-(x^2))^n) sums (1 / (1 - (-(x^2))))" by (intro geometric_sums) simp_all also have "1 - (-(x^2)) = 1 + x^2" by simp also have "(\n. (-(x^2))^n) = (\n. ?f (2*n))" by (auto simp: fun_eq_iff power_minus' power_mult) also have "range ((*) (2::nat)) = {n. even n}" by (auto elim!: evenE) hence "(\n. ?f (2*n)) sums (1 / (1 + x^2)) \ ?f sums (1 / (1 + x^2))" by (intro sums_mono_reindex) (simp_all add: strict_mono_Suc_iff) also have "?f = (\n. diffs arctan_coeffs n * x ^ n)" by (simp add: fun_eq_iff arctan_coeffs_def diffs_def) finally show "(\n. diffs arctan_coeffs n * x ^ n) sums (1 / (1 + x^2))" . qed have B: "summable (\n. arctan_coeffs n * x ^ n)" if "norm x < 1" for x :: real proof (rule summable_comparison_test) show "\N. \n\N. norm (arctan_coeffs n * x ^ n) \ 1 * norm x ^ n" unfolding norm_mult norm_power by (intro exI[of _ "0::nat"] allI impI mult_right_mono) (simp_all add: arctan_coeffs_def field_split_simps) from that show "summable (\n. 1 * norm x ^ n)" by (intro summable_mult summable_geometric) simp_all qed define F :: "real \ real" where "F = (\x. \n. arctan_coeffs n * x ^ n)" have [derivative_intros]: "(F has_field_derivative (1 / (1 + x^2))) (at x)" if "norm x < 1" for x :: real proof - define K :: real where "K = (1 + norm x) / 2" from that have K: "norm x < K" "K < 1" by (simp_all add: K_def field_simps) have "(F has_field_derivative (\n. diffs arctan_coeffs n * x ^ n)) (at x)" unfolding F_def using K by (intro termdiffs_strong[where K = K] B) simp_all also from A[OF that] have "(\n. diffs arctan_coeffs n * x ^ n) = 1 / (1 + x^2)" by (simp add: sums_iff) finally show ?thesis . qed from assms have "arctan x - F x = arctan 0 - F 0" by (intro DERIV_isconst3[of "-1" 1 _ _ "\x. arctan x - F x"]) (auto intro!: derivative_eq_intros simp: field_simps) hence "F x = arctan x" by (simp add: F_def arctan_coeffs_def) with B[OF assms] show ?thesis by (simp add: sums_iff F_def) qed primcorec arctan_series_stream_aux :: "bool \ real \ real msstream" where "arctan_series_stream_aux b n = MSSCons (if b then -inverse n else inverse n) (arctan_series_stream_aux (\b) (n + 2))" lemma mssnth_arctan_series_stream_aux: "mssnth (arctan_series_stream_aux b n) m = (-1) ^ (if b then Suc m else m) / (2*m + n)" by (induction m arbitrary: b n; subst arctan_series_stream_aux.code) (auto simp: field_simps split: if_splits) definition arctan_series_stream where "arctan_series_stream = arctan_series_stream_aux False 1" lemma mssnth_arctan_series_stream: "mssnth arctan_series_stream n = (-1) ^ n / (2 * n + 1)" by (simp add: arctan_series_stream_def mssnth_arctan_series_stream_aux) lemma summable_pre_arctan: assumes "norm x < 1" shows "summable (\n. mssnth arctan_series_stream n * x ^ n)" (is "summable ?f") proof (rule summable_comparison_test) show "summable (\n. abs x ^ n)" using assms by (intro summable_geometric) auto show "\N. \n\N. norm (?f n) \ abs x ^ n" proof (intro exI[of _ 0] allI impI) fix n :: nat have "norm (?f n) = \x\ ^ n / (2 * real n + 1)" by (simp add: mssnth_arctan_series_stream abs_mult power_abs) also have "\ \ \x\ ^ n / 1" by (intro divide_left_mono) auto finally show "norm (?f n) \ abs x ^ n" by simp qed qed lemma arctan_conv_pre_arctan: assumes "norm x < 1" shows "arctan x = x * powser (msllist_of_msstream arctan_series_stream) (x ^ 2)" proof - from assms have "norm x * norm x < 1 * 1" by (intro mult_strict_mono) auto hence norm_less: "norm (x ^ 2) < 1" by (simp add: power2_eq_square) have "(\n. arctan_coeffs n * x ^ n) sums arctan x" by (intro arctan_converges assms) also have "?this \ (\n. arctan_coeffs (2 * n + 1) * x ^ (2 * n + 1)) sums arctan x" by (intro sums_mono_reindex [symmetric]) (auto simp: arctan_coeffs_def strict_mono_def elim!: oddE) also have "(\n. arctan_coeffs (2 * n + 1) * x ^ (2 * n + 1)) = (\n. x * mssnth arctan_series_stream n * (x ^ 2) ^ n)" by (simp add: fun_eq_iff mssnth_arctan_series_stream power_mult [symmetric] arctan_coeffs_def mult_ac) finally have "(\n. x * mssnth arctan_series_stream n * x\<^sup>2 ^ n) sums arctan x" . thus ?thesis using summable_pre_arctan[OF norm_less] assms by (simp add: powser_def sums_iff suminf_mult [symmetric] mult.assoc) qed lemma convergent_powser_arctan: "convergent_powser (msllist_of_msstream arctan_series_stream)" unfolding convergent_powser_def using summable_pre_arctan by (intro exI[of _ 1]) auto lemma arctan_inverse_pos: "x > 0 \ arctan x = pi / 2 - arctan (inverse x)" using arctan_inverse[of x] by (simp add: field_simps) lemma arctan_inverse_neg: "x < 0 \ arctan x = -pi / 2 - arctan (inverse x)" using arctan_inverse[of x] by (simp add: field_simps) subsection \Multiseries\ subsubsection \Asymptotic bases\ type_synonym basis = "(real \ real) list" lemma transp_ln_smallo_ln: "transp (\f g. (\x::real. ln (g x)) \ o(\x. ln (f x)))" by (rule transpI, erule landau_o.small.trans) definition basis_wf :: "basis \ bool" where "basis_wf basis \ (\f\set basis. filterlim f at_top at_top) \ sorted_wrt (\f g. (\x. ln (g x)) \ o(\x. ln (f x))) basis" lemma basis_wf_Nil [simp]: "basis_wf []" by (simp add: basis_wf_def) lemma basis_wf_Cons: "basis_wf (f # basis) \ filterlim f at_top at_top \ (\g\set basis. (\x. ln (g x)) \ o(\x. ln (f x))) \ basis_wf basis" unfolding basis_wf_def by auto (* TODO: Move *) lemma sorted_wrt_snoc: assumes trans: "transp P" and "sorted_wrt P xs" "P (last xs) y" shows "sorted_wrt P (xs @ [y])" proof (subst sorted_wrt_append, intro conjI ballI) fix x y' assume x: "x \ set xs" and y': "y' \ set [y]" from y' have [simp]: "y' = y" by simp show "P x y'" proof (cases "x = last xs") case False from x have eq: "xs = butlast xs @ [last xs]" by (subst append_butlast_last_id) auto from x and False have x': "x \ set (butlast xs)" by (subst (asm) eq) auto have "sorted_wrt P xs" by fact also note eq finally have "P x (last xs)" using x' by (subst (asm) sorted_wrt_append) auto with \P (last xs) y\ have "P x y" using transpD[OF trans] by blast thus ?thesis by simp qed (insert assms, auto) qed (insert assms, auto) lemma basis_wf_snoc: assumes "bs \ []" assumes "basis_wf bs" "filterlim b at_top at_top" assumes "(\x. ln (b x)) \ o(\x. ln (last bs x))" shows "basis_wf (bs @ [b])" proof - have "transp (\f g. (\x. ln (g x)) \ o(\x. ln (f x)))" by (auto simp: transp_def intro: landau_o.small.trans) thus ?thesis using assms by (auto simp add: basis_wf_def sorted_wrt_snoc[OF transp_ln_smallo_ln]) qed lemma basis_wf_singleton: "basis_wf [b] \ filterlim b at_top at_top" by (simp add: basis_wf_Cons) lemma basis_wf_many: "basis_wf (b # b' # bs) \ filterlim b at_top at_top \ (\x. ln (b' x)) \ o(\x. ln (b x)) \ basis_wf (b' # bs)" unfolding basis_wf_def by (subst sorted_wrt2[OF transp_ln_smallo_ln]) auto subsubsection \Monomials\ type_synonym monom = "real \ real list" definition eval_monom :: "monom \ basis \ real \ real" where "eval_monom = (\(c,es) basis x. c * prod_list (map (\(b,e). b x powr e) (zip basis es)))" lemma eval_monom_Nil1 [simp]: "eval_monom (c, []) basis = (\_. c)" by (simp add: eval_monom_def split: prod.split) lemma eval_monom_Nil2 [simp]: "eval_monom monom [] = (\_. fst monom)" by (simp add: eval_monom_def split: prod.split) lemma eval_monom_Cons: "eval_monom (c, e # es) (b # basis) = (\x. eval_monom (c, es) basis x * b x powr e)" by (simp add: eval_monom_def fun_eq_iff split: prod.split) definition inverse_monom :: "monom \ monom" where "inverse_monom monom = (case monom of (c, es) \ (inverse c, map uminus es))" lemma length_snd_inverse_monom [simp]: "length (snd (inverse_monom monom)) = length (snd monom)" by (simp add: inverse_monom_def split: prod.split) lemma eval_monom_pos: assumes "basis_wf basis" "fst monom > 0" shows "eventually (\x. eval_monom monom basis x > 0) at_top" proof (cases monom) case (Pair c es) with assms have "c > 0" by simp with assms(1) show ?thesis unfolding Pair proof (induction es arbitrary: basis) case (Cons e es) note A = this thus ?case proof (cases basis) case (Cons b basis') with A(1)[of basis'] A(2,3) have A: "\\<^sub>F x in at_top. eval_monom (c, es) basis' x > 0" and B: "eventually (\x. b x > 0) at_top" by (simp_all add: basis_wf_Cons filterlim_at_top_dense) thus ?thesis by eventually_elim (simp add: Cons eval_monom_Cons) qed simp qed simp qed lemma eval_monom_uminus: "eval_monom (-c, es) basis x = -eval_monom (c, es) basis x" by (simp add: eval_monom_def) lemma eval_monom_neg: assumes "basis_wf basis" "fst monom < 0" shows "eventually (\x. eval_monom monom basis x < 0) at_top" proof - from assms have "eventually (\x. eval_monom (-fst monom, snd monom) basis x > 0) at_top" by (intro eval_monom_pos) simp_all thus ?thesis by (simp add: eval_monom_uminus) qed lemma eval_monom_nonzero: assumes "basis_wf basis" "fst monom \ 0" shows "eventually (\x. eval_monom monom basis x \ 0) at_top" proof (cases "fst monom" "0 :: real" rule: linorder_cases) case greater with assms have "eventually (\x. eval_monom monom basis x > 0) at_top" by (simp add: eval_monom_pos) thus ?thesis by eventually_elim simp next case less with assms have "eventually (\x. eval_monom (-fst monom, snd monom) basis x > 0) at_top" by (simp add: eval_monom_pos) thus ?thesis by eventually_elim (simp add: eval_monom_uminus) qed (insert assms, simp_all) subsubsection \Typeclass for multiseries\ class multiseries = plus + minus + times + uminus + inverse + fixes is_expansion :: "'a \ basis \ bool" and expansion_level :: "'a itself \ nat" and eval :: "'a \ real \ real" and zero_expansion :: 'a and const_expansion :: "real \ 'a" and powr_expansion :: "bool \ 'a \ real \ 'a" and power_expansion :: "bool \ 'a \ nat \ 'a" and trimmed :: "'a \ bool" and dominant_term :: "'a \ monom" assumes is_expansion_length: "is_expansion F basis \ length basis = expansion_level (TYPE('a))" assumes is_expansion_zero: "basis_wf basis \ length basis = expansion_level (TYPE('a)) \ is_expansion zero_expansion basis" assumes is_expansion_const: "basis_wf basis \ length basis = expansion_level (TYPE('a)) \ is_expansion (const_expansion c) basis" assumes is_expansion_uminus: "basis_wf basis \ is_expansion F basis \ is_expansion (-F) basis" assumes is_expansion_add: "basis_wf basis \ is_expansion F basis \ is_expansion G basis \ is_expansion (F + G) basis" assumes is_expansion_minus: "basis_wf basis \ is_expansion F basis \ is_expansion G basis \ is_expansion (F - G) basis" assumes is_expansion_mult: "basis_wf basis \ is_expansion F basis \ is_expansion G basis \ is_expansion (F * G) basis" assumes is_expansion_inverse: "basis_wf basis \ trimmed F \ is_expansion F basis \ is_expansion (inverse F) basis" assumes is_expansion_divide: "basis_wf basis \ trimmed G \ is_expansion F basis \ is_expansion G basis \ is_expansion (F / G) basis" assumes is_expansion_powr: "basis_wf basis \ trimmed F \ fst (dominant_term F) > 0 \ is_expansion F basis \ is_expansion (powr_expansion abort F p) basis" assumes is_expansion_power: "basis_wf basis \ trimmed F \ is_expansion F basis \ is_expansion (power_expansion abort F n) basis" assumes is_expansion_imp_smallo: "basis_wf basis \ is_expansion F basis \ filterlim b at_top at_top \ (\g\set basis. (\x. ln (g x)) \ o(\x. ln (b x))) \ e > 0 \ eval F \ o(\x. b x powr e)" assumes is_expansion_imp_smallomega: "basis_wf basis \ is_expansion F basis \ filterlim b at_top at_top \ trimmed F \ (\g\set basis. (\x. ln (g x)) \ o(\x. ln (b x))) \ e < 0 \ eval F \ \(\x. b x powr e)" assumes trimmed_imp_eventually_sgn: "basis_wf basis \ is_expansion F basis \ trimmed F \ eventually (\x. sgn (eval F x) = sgn (fst (dominant_term F))) at_top" assumes trimmed_imp_eventually_nz: "basis_wf basis \ is_expansion F basis \ trimmed F \ eventually (\x. eval F x \ 0) at_top" assumes trimmed_imp_dominant_term_nz: "trimmed F \ fst (dominant_term F) \ 0" assumes dominant_term: "basis_wf basis \ is_expansion F basis \ trimmed F \ eval F \[at_top] eval_monom (dominant_term F) basis" assumes dominant_term_bigo: "basis_wf basis \ is_expansion F basis \ eval F \ O(eval_monom (1, snd (dominant_term F)) basis)" assumes length_dominant_term: "length (snd (dominant_term F)) = expansion_level TYPE('a)" assumes fst_dominant_term_uminus [simp]: "fst (dominant_term (- F)) = -fst (dominant_term F)" assumes trimmed_uminus_iff [simp]: "trimmed (-F) \ trimmed F" assumes add_zero_expansion_left [simp]: "zero_expansion + F = F" assumes add_zero_expansion_right [simp]: "F + zero_expansion = F" assumes eval_zero [simp]: "eval zero_expansion x = 0" assumes eval_const [simp]: "eval (const_expansion c) x = c" assumes eval_uminus [simp]: "eval (-F) = (\x. -eval F x)" assumes eval_plus [simp]: "eval (F + G) = (\x. eval F x + eval G x)" assumes eval_minus [simp]: "eval (F - G) = (\x. eval F x - eval G x)" assumes eval_times [simp]: "eval (F * G) = (\x. eval F x * eval G x)" assumes eval_inverse [simp]: "eval (inverse F) = (\x. inverse (eval F x))" assumes eval_divide [simp]: "eval (F / G) = (\x. eval F x / eval G x)" assumes eval_powr [simp]: "eval (powr_expansion abort F p) = (\x. eval F x powr p)" assumes eval_power [simp]: "eval (power_expansion abort F n) = (\x. eval F x ^ n)" assumes minus_eq_plus_uminus: "F - G = F + (-G)" assumes times_const_expansion_1: "const_expansion 1 * F = F" assumes trimmed_const_expansion: "trimmed (const_expansion c) \ c \ 0" begin definition trimmed_pos where "trimmed_pos F \ trimmed F \ fst (dominant_term F) > 0" definition trimmed_neg where "trimmed_neg F \ trimmed F \ fst (dominant_term F) < 0" lemma trimmed_pos_uminus: "trimmed_neg F \ trimmed_pos (-F)" by (simp add: trimmed_neg_def trimmed_pos_def) lemma trimmed_pos_imp_trimmed: "trimmed_pos x \ trimmed x" by (simp add: trimmed_pos_def) lemma trimmed_neg_imp_trimmed: "trimmed_neg x \ trimmed x" by (simp add: trimmed_neg_def) end subsubsection \Zero-rank expansions\ instantiation real :: multiseries begin inductive is_expansion_real :: "real \ basis \ bool" where "is_expansion_real c []" definition expansion_level_real :: "real itself \ nat" where expansion_level_real_def [simp]: "expansion_level_real _ = 0" definition eval_real :: "real \ real \ real" where eval_real_def [simp]: "eval_real c = (\_. c)" definition zero_expansion_real :: "real" where zero_expansion_real_def [simp]: "zero_expansion_real = 0" definition const_expansion_real :: "real \ real" where const_expansion_real_def [simp]: "const_expansion_real x = x" definition powr_expansion_real :: "bool \ real \ real \ real" where powr_expansion_real_def [simp]: "powr_expansion_real _ x p = x powr p" definition power_expansion_real :: "bool \ real \ nat \ real" where power_expansion_real_def [simp]: "power_expansion_real _ x n = x ^ n" definition trimmed_real :: "real \ bool" where trimmed_real_def [simp]: "trimmed_real x \ x \ 0" definition dominant_term_real :: "real \ monom" where dominant_term_real_def [simp]: "dominant_term_real c = (c, [])" instance proof fix basis :: basis and b :: "real \ real" and e F :: real assume *: "basis_wf basis" "filterlim b at_top at_top" "is_expansion F basis" "0 < e" have "(\x. b x powr e) \ \(\_. 1)" by (intro smallomegaI_filterlim_at_infinity filterlim_at_top_imp_at_infinity) (auto intro!: filterlim_compose[OF real_powr_at_top] * ) with * show "eval F \ o(\x. b x powr e)" by (cases "F = 0") (auto elim!: is_expansion_real.cases simp: smallomega_iff_smallo) next fix basis :: basis and b :: "real \ real" and e F :: real assume *: "basis_wf basis" "filterlim b at_top at_top" "is_expansion F basis" "e < 0" "trimmed F" from * have pos: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) have "(\x. b x powr -e) \ \(\_. 1)" by (intro smallomegaI_filterlim_at_infinity filterlim_at_top_imp_at_infinity) (auto intro!: filterlim_compose[OF real_powr_at_top] * ) with pos have "(\x. b x powr e) \ o(\_. 1)" unfolding powr_minus by (subst (asm) landau_omega.small.inverse_eq2) (auto elim: eventually_mono simp: smallomega_iff_smallo) with * show "eval F \ \(\x. b x powr e)" by (auto elim!: is_expansion_real.cases simp: smallomega_iff_smallo) qed (auto intro!: is_expansion_real.intros elim!: is_expansion_real.cases) end lemma eval_real: "eval (c :: real) x = c" by simp subsubsection \Operations on multiseries\ lemma ms_aux_cases [case_names MSLNil MSLCons]: fixes xs :: "('a \ real) msllist" obtains "xs = MSLNil" | c e xs' where "xs = MSLCons (c, e) xs'" proof (cases xs) case (MSLCons x xs') with that(2)[of "fst x" "snd x" xs'] show ?thesis by auto qed auto definition ms_aux_hd_exp :: "('a \ real) msllist \ real option" where "ms_aux_hd_exp xs = (case xs of MSLNil \ None | MSLCons (_, e) _ \ Some e)" primrec ms_exp_gt :: "real \ real option \ bool" where "ms_exp_gt x None = True" | "ms_exp_gt x (Some y) \ x > y" lemma ms_aux_hd_exp_MSLNil [simp]: "ms_aux_hd_exp MSLNil = None" by (simp add: ms_aux_hd_exp_def split: prod.split) lemma ms_aux_hd_exp_MSLCons [simp]: "ms_aux_hd_exp (MSLCons x xs) = Some (snd x)" by (simp add: ms_aux_hd_exp_def split: prod.split) coinductive is_expansion_aux :: "('a :: multiseries \ real) msllist \ (real \ real) \ basis \ bool" where is_expansion_MSLNil: "eventually (\x. f x = 0) at_top \ length basis = Suc (expansion_level TYPE('a)) \ is_expansion_aux MSLNil f basis" | is_expansion_MSLCons: "is_expansion_aux xs (\x. f x - eval C x * b x powr e) (b # basis) \ is_expansion C basis \ (\e'. e' > e \ f \ o(\x. b x powr e')) \ ms_exp_gt e (ms_aux_hd_exp xs) \ is_expansion_aux (MSLCons (C, e) xs) f (b # basis)" inductive_cases is_expansion_aux_MSLCons: "is_expansion_aux (MSLCons (c, e) xs) f basis" lemma is_expansion_aux_basis_nonempty: "is_expansion_aux F f basis \ basis \ []" by (erule is_expansion_aux.cases) auto lemma is_expansion_aux_expansion_level: assumes "is_expansion_aux (G :: ('a::multiseries \ real) msllist) g basis" shows "length basis = Suc (expansion_level TYPE('a))" using assms by (cases rule: is_expansion_aux.cases) (auto dest: is_expansion_length) lemma is_expansion_aux_imp_smallo: assumes "is_expansion_aux xs f basis" "ms_exp_gt p (ms_aux_hd_exp xs)" shows "f \ o(\x. hd basis x powr p)" using assms proof (cases rule: is_expansion_aux.cases) case is_expansion_MSLNil show ?thesis by (simp add: landau_o.small.in_cong[OF is_expansion_MSLNil(2)]) next case (is_expansion_MSLCons xs C b e basis) with assms have "f \ o(\x. b x powr p)" by (intro is_expansion_MSLCons) (simp_all add: ms_aux_hd_exp_def) thus ?thesis by (simp add: is_expansion_MSLCons) qed lemma is_expansion_aux_imp_smallo': assumes "is_expansion_aux xs f basis" obtains e where "f \ o(\x. hd basis x powr e)" proof - define e where "e = (case ms_aux_hd_exp xs of None \ 0 | Some e \ e + 1)" have "ms_exp_gt e (ms_aux_hd_exp xs)" by (auto simp add: e_def ms_aux_hd_exp_def split: msllist.splits) from assms this have "f \ o(\x. hd basis x powr e)" by (rule is_expansion_aux_imp_smallo) from that[OF this] show ?thesis . qed lemma is_expansion_aux_imp_smallo'': assumes "is_expansion_aux xs f basis" "ms_exp_gt e' (ms_aux_hd_exp xs)" obtains e where "e < e'" "f \ o(\x. hd basis x powr e)" proof - define e where "e = (case ms_aux_hd_exp xs of None \ e' - 1 | Some e \ (e' + e) / 2)" from assms have "ms_exp_gt e (ms_aux_hd_exp xs)" "e < e'" by (cases xs; simp add: e_def field_simps)+ from assms(1) this(1) have "f \ o(\x. hd basis x powr e)" by (rule is_expansion_aux_imp_smallo) from that[OF \e < e'\ this] show ?thesis . qed definition trimmed_ms_aux :: "('a :: multiseries \ real) msllist \ bool" where "trimmed_ms_aux xs = (case xs of MSLNil \ False | MSLCons (C, _) _ \ trimmed C)" lemma not_trimmed_ms_aux_MSLNil [simp]: "\trimmed_ms_aux MSLNil" by (simp add: trimmed_ms_aux_def) lemma trimmed_ms_aux_MSLCons: "trimmed_ms_aux (MSLCons x xs) = trimmed (fst x)" by (simp add: trimmed_ms_aux_def split: prod.split) lemma trimmed_ms_aux_imp_nz: assumes "basis_wf basis" "is_expansion_aux xs f basis" "trimmed_ms_aux xs" shows "eventually (\x. f x \ 0) at_top" proof (cases xs rule: ms_aux_cases) case (MSLCons C e xs') from assms this obtain b basis' where *: "basis = b # basis'" "is_expansion_aux xs' (\x. f x - eval C x * b x powr e) (b # basis')" "ms_exp_gt e (ms_aux_hd_exp xs')" "is_expansion C basis'" "\e'. e' > e \ f \ o(\x. b x powr e')" by (auto elim!: is_expansion_aux_MSLCons) from assms(1,3) * have nz: "eventually (\x. eval C x \ 0) at_top" by (auto simp: trimmed_ms_aux_def MSLCons basis_wf_Cons intro: trimmed_imp_eventually_nz[of basis']) from assms(1) * have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_nz: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) from is_expansion_aux_imp_smallo''[OF *(2,3)] obtain e' where e': "e' < e" "(\x. f x - eval C x * b x powr e) \ o(\x. b x powr e')" unfolding list.sel by blast note this(2) also have "\ = o(\x. b x powr (e' - e) * b x powr e)" by (simp add: powr_add [symmetric]) also from assms * e' have "eval C \ \(\x. b x powr (e' - e))" by (intro is_expansion_imp_smallomega[OF _ *(4)]) (simp_all add: MSLCons basis_wf_Cons trimmed_ms_aux_def) hence "(\x. b x powr (e' - e) * b x powr e) \ o(\x. eval C x * b x powr e)" by (intro landau_o.small_big_mult is_expansion_imp_smallomega) (simp_all add: smallomega_iff_smallo) finally have "(\x. f x - eval C x * b x powr e + eval C x * b x powr e) \ \(\x. eval C x * b x powr e)" by (subst bigtheta_sym, subst landau_theta.plus_absorb1) simp_all hence theta: "f \ \(\x. eval C x * b x powr e)" by simp have "eventually (\x. eval C x * b x powr e \ 0) at_top" using b_nz nz by eventually_elim simp thus ?thesis by (subst eventually_nonzero_bigtheta [OF theta]) qed (insert assms, simp_all add: trimmed_ms_aux_def) lemma is_expansion_aux_imp_smallomega: assumes "basis_wf basis" "is_expansion_aux xs f basis" "filterlim b' at_top at_top" "trimmed_ms_aux xs" "\g\set basis. (\x. ln (g x)) \ o(\x. ln (b' x))" "r < 0" shows "f \ \(\x. b' x powr r)" proof (cases xs rule: ms_aux_cases) case (MSLCons C e xs') from assms this obtain b basis' where *: "basis = b # basis'" "trimmed C" "is_expansion_aux xs' (\x. f x - eval C x * b x powr e) (b # basis')" "ms_exp_gt e (ms_aux_hd_exp xs')" "is_expansion C basis'" "\e'. e' > e \ f \ o(\x. b x powr e')" by (auto elim!: is_expansion_aux_MSLCons simp: trimmed_ms_aux_def) from assms * have nz: "eventually (\x. eval C x \ 0) at_top" by (intro trimmed_imp_eventually_nz[of basis']) (simp_all add: basis_wf_Cons) from assms(1) * have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_nz: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) from is_expansion_aux_imp_smallo''[OF *(3,4)] obtain e' where e': "e' < e" "(\x. f x - eval C x * b x powr e) \ o(\x. b x powr e')" unfolding list.sel by blast note this(2) also have "\ = o(\x. b x powr (e' - e) * b x powr e)" by (simp add: powr_add [symmetric]) also from assms * e' have "eval C \ \(\x. b x powr (e' - e))" by (intro is_expansion_imp_smallomega[OF _ *(5)]) (simp_all add: MSLCons basis_wf_Cons trimmed_ms_aux_def) hence "(\x. b x powr (e' - e) * b x powr e) \ o(\x. eval C x * b x powr e)" by (intro landau_o.small_big_mult is_expansion_imp_smallomega) (simp_all add: smallomega_iff_smallo) finally have "(\x. f x - eval C x * b x powr e + eval C x * b x powr e) \ \(\x. eval C x * b x powr e)" by (subst bigtheta_sym, subst landau_theta.plus_absorb1) simp_all hence theta: "f \ \(\x. eval C x * b x powr e)" by simp also from * assms e' have "(\x. eval C x * b x powr e) \ \(\x. b x powr (e' - e) * b x powr e)" by (intro landau_omega.small_big_mult is_expansion_imp_smallomega[of basis']) (simp_all add: basis_wf_Cons) also have "\ = \(\x. b x powr e')" by (simp add: powr_add [symmetric]) also from *(1) assms have "(\x. b x powr e') \ \(\x. b' x powr r)" unfolding smallomega_iff_smallo by (intro ln_smallo_imp_flat') (simp_all add: basis_wf_Cons) finally show ?thesis . qed (insert assms, simp_all add: trimmed_ms_aux_def) definition dominant_term_ms_aux :: "('a :: multiseries \ real) msllist \ monom" where "dominant_term_ms_aux xs = (case xs of MSLNil \ (0, replicate (Suc (expansion_level TYPE('a))) 0) | MSLCons (C, e) _ \ case dominant_term C of (c, es) \ (c, e # es))" lemma dominant_term_ms_aux_MSLCons: "dominant_term_ms_aux (MSLCons (C, e) xs) = map_prod id (\es. e # es) (dominant_term C)" by (simp add: dominant_term_ms_aux_def case_prod_unfold map_prod_def) lemma length_dominant_term_ms_aux [simp]: "length (snd (dominant_term_ms_aux (xs :: ('a :: multiseries \ real) msllist))) = Suc (expansion_level TYPE('a))" proof (cases xs rule: ms_aux_cases) case (MSLCons C e xs') hence "length (snd (dominant_term_ms_aux xs)) = Suc (length (snd (dominant_term C)))" by (simp add: dominant_term_ms_aux_def split: prod.splits) also note length_dominant_term finally show ?thesis . qed (simp_all add: dominant_term_ms_aux_def split: msllist.splits prod.splits) definition const_ms_aux :: "real \ ('a :: multiseries \ real) msllist" where "const_ms_aux c = MSLCons (const_expansion c, 0) MSLNil" definition uminus_ms_aux :: "('a :: uminus \ real) msllist \ ('a \ real) msllist" where "uminus_ms_aux xs = mslmap (map_prod uminus id) xs" corec (friend) plus_ms_aux :: "('a :: plus \ real) msllist \ ('a \ real) msllist \ ('a \ real) msllist" where "plus_ms_aux xs ys = (case (xs, ys) of (MSLNil, MSLNil) \ MSLNil | (MSLNil, MSLCons y ys) \ MSLCons y ys | (MSLCons x xs, MSLNil) \ MSLCons x xs | (MSLCons x xs, MSLCons y ys) \ if snd x > snd y then MSLCons x (plus_ms_aux xs (MSLCons y ys)) else if snd x < snd y then MSLCons y (plus_ms_aux (MSLCons x xs) ys) else MSLCons (fst x + fst y, snd x) (plus_ms_aux xs ys))" context begin friend_of_corec mslmap where "mslmap (f :: 'a \ 'a) xs = (case xs of MSLNil \ MSLNil | MSLCons x xs \ MSLCons (f x) (mslmap f xs))" by (simp split: msllist.splits, transfer_prover) corec (friend) times_ms_aux :: "('a :: {plus,times} \ real) msllist \ ('a \ real) msllist \ ('a \ real) msllist" where "times_ms_aux xs ys = (case (xs, ys) of (MSLNil, ys) \ MSLNil | (xs, MSLNil) \ MSLNil | (MSLCons x xs, MSLCons y ys) \ MSLCons (fst x * fst y, snd x + snd y) (plus_ms_aux (mslmap (map_prod ((*) (fst x)) ((+) (snd x))) ys) (times_ms_aux xs (MSLCons y ys))))" definition scale_shift_ms_aux :: "('a :: times \ real) \ ('a \ real) msllist \ ('a \ real) msllist" where "scale_shift_ms_aux = (\(a,b) xs. mslmap (map_prod ((*) a) ((+) b)) xs)" lemma times_ms_aux_altdef: "times_ms_aux xs ys = (case (xs, ys) of (MSLNil, ys) \ MSLNil | (xs, MSLNil) \ MSLNil | (MSLCons x xs, MSLCons y ys) \ MSLCons (fst x * fst y, snd x + snd y) (plus_ms_aux (scale_shift_ms_aux x ys) (times_ms_aux xs (MSLCons y ys))))" by (subst times_ms_aux.code) (simp_all add: scale_shift_ms_aux_def split: msllist.splits) end corec powser_ms_aux :: "real msllist \ ('a :: multiseries \ real) msllist \ ('a \ real) msllist" where "powser_ms_aux cs xs = (case cs of MSLNil \ MSLNil | MSLCons c cs \ MSLCons (const_expansion c, 0) (times_ms_aux xs (powser_ms_aux cs xs)))" definition minus_ms_aux :: "('a :: multiseries \ real) msllist \ _" where "minus_ms_aux xs ys = plus_ms_aux xs (uminus_ms_aux ys)" lemma is_expansion_aux_coinduct [consumes 1, case_names is_expansion_aux]: fixes xs :: "('a :: multiseries \ real) msllist" assumes "X xs f basis" "(\xs f basis. X xs f basis \ (xs = MSLNil \ (\\<^sub>F x in at_top. f x = 0) \ length basis = Suc (expansion_level TYPE('a))) \ (\xs' b e basis' C. xs = MSLCons (C, e) xs' \ basis = b # basis' \ (X xs' (\x. f x - eval C x * b x powr e) (b # basis')) \ is_expansion C basis' \ (\x>e. f \ o(\xa. b xa powr x)) \ ms_exp_gt e (ms_aux_hd_exp xs')))" shows "is_expansion_aux xs f basis" using assms(1) proof (coinduction arbitrary: xs f) case (is_expansion_aux xs f) from assms(2)[OF is_expansion_aux] show ?case by blast qed lemma is_expansion_aux_cong: assumes "is_expansion_aux F f basis" assumes "eventually (\x. f x = g x) at_top" shows "is_expansion_aux F g basis" using assms proof (coinduction arbitrary: F f g rule: is_expansion_aux_coinduct) case (is_expansion_aux F f g) from this have ev: "eventually (\x. g x = f x) at_top" by (simp add: eq_commute) from ev have [simp]: "eventually (\x. g x = 0) at_top \ eventually (\x. f x = 0) at_top" by (rule eventually_subst') from ev have [simp]: "(\\<^sub>F x in at_top. h x = g x - h' x) \ (\\<^sub>F x in at_top. h x = f x - h' x)" for h h' by (rule eventually_subst') note [simp] = landau_o.small.in_cong[OF ev] from is_expansion_aux(1) show ?case by (cases rule: is_expansion_aux.cases) force+ qed lemma scale_shift_ms_aux_conv_mslmap: "scale_shift_ms_aux x = mslmap (map_prod ((*) (fst x)) ((+) (snd x)))" by (rule ext) (simp add: scale_shift_ms_aux_def map_prod_def case_prod_unfold) fun inverse_ms_aux :: "('a :: multiseries \ real) msllist \ ('a \ real) msllist" where "inverse_ms_aux (MSLCons x xs) = (let c' = inverse (fst x) in scale_shift_ms_aux (c', -snd x) (powser_ms_aux (msllist_of_msstream (mssalternate 1 (-1))) (scale_shift_ms_aux (c', -snd x) xs)))" (* TODO: Move? *) lemma inverse_prod_list: "inverse (prod_list xs :: 'a :: field) = prod_list (map inverse xs)" by (induction xs) simp_all lemma eval_inverse_monom: "eval_monom (inverse_monom monom) basis = (\x. inverse (eval_monom monom basis x))" by (rule ext) (simp add: eval_monom_def inverse_monom_def zip_map2 o_def case_prod_unfold inverse_prod_list powr_minus) fun powr_ms_aux :: "bool \ ('a :: multiseries \ real) msllist \ real \ ('a \ real) msllist" where "powr_ms_aux abort (MSLCons x xs) p = scale_shift_ms_aux (powr_expansion abort (fst x) p, snd x * p) (powser_ms_aux (gbinomial_series abort p) (scale_shift_ms_aux (inverse (fst x), -snd x) xs))" fun power_ms_aux :: "bool \ ('a :: multiseries \ real) msllist \ nat \ ('a \ real) msllist" where "power_ms_aux abort (MSLCons x xs) n = scale_shift_ms_aux (power_expansion abort (fst x) n, snd x * real n) (powser_ms_aux (gbinomial_series abort (real n)) (scale_shift_ms_aux (inverse (fst x), -snd x) xs))" definition sin_ms_aux' :: "('a :: multiseries \ real) msllist \ ('a \ real) msllist" where "sin_ms_aux' xs = times_ms_aux xs (powser_ms_aux (msllist_of_msstream sin_series_stream) (times_ms_aux xs xs))" definition cos_ms_aux' :: "('a :: multiseries \ real) msllist \ ('a \ real) msllist" where "cos_ms_aux' xs = powser_ms_aux (msllist_of_msstream cos_series_stream) (times_ms_aux xs xs)" subsubsection \Basic properties of multiseries operations\ lemma uminus_ms_aux_MSLNil [simp]: "uminus_ms_aux MSLNil = MSLNil" by (simp add: uminus_ms_aux_def) lemma uminus_ms_aux_MSLCons: "uminus_ms_aux (MSLCons x xs) = MSLCons (-fst x, snd x) (uminus_ms_aux xs)" by (simp add: uminus_ms_aux_def map_prod_def split: prod.splits) lemma uminus_ms_aux_MSLNil_iff [simp]: "uminus_ms_aux xs = MSLNil \ xs = MSLNil" by (simp add: uminus_ms_aux_def) lemma hd_exp_uminus [simp]: "ms_aux_hd_exp (uminus_ms_aux xs) = ms_aux_hd_exp xs" by (simp add: uminus_ms_aux_def ms_aux_hd_exp_def split: msllist.splits prod.split) lemma scale_shift_ms_aux_MSLNil_iff [simp]: "scale_shift_ms_aux x F = MSLNil \ F = MSLNil" by (auto simp: scale_shift_ms_aux_def split: prod.split) lemma scale_shift_ms_aux_MSLNil [simp]: "scale_shift_ms_aux x MSLNil = MSLNil" by (simp add: scale_shift_ms_aux_def split: prod.split) lemma scale_shift_ms_aux_1 [simp]: "scale_shift_ms_aux (1 :: real, 0) xs = xs" by (simp add: scale_shift_ms_aux_def map_prod_def msllist.map_id) lemma scale_shift_ms_aux_MSLCons: "scale_shift_ms_aux x (MSLCons y F) = MSLCons (fst x * fst y, snd x + snd y) (scale_shift_ms_aux x F)" by (auto simp: scale_shift_ms_aux_def map_prod_def split: prod.split) lemma hd_exp_basis: "ms_aux_hd_exp (scale_shift_ms_aux x xs) = map_option ((+) (snd x)) (ms_aux_hd_exp xs)" by (auto simp: ms_aux_hd_exp_def scale_shift_ms_aux_MSLCons split: msllist.split) lemma plus_ms_aux_MSLNil_iff: "plus_ms_aux F G = MSLNil \ F = MSLNil \ G = MSLNil" by (subst plus_ms_aux.code) (simp split: msllist.splits) lemma plus_ms_aux_MSLNil [simp]: "plus_ms_aux F MSLNil = F" "plus_ms_aux MSLNil G = G" by (subst plus_ms_aux.code, simp split: msllist.splits)+ lemma plus_ms_aux_MSLCons: "plus_ms_aux (MSLCons x xs) (MSLCons y ys) = (if snd x > snd y then MSLCons x (plus_ms_aux xs (MSLCons y ys)) else if snd x < snd y then MSLCons y (plus_ms_aux (MSLCons x xs) ys) else MSLCons (fst x + fst y, snd x) (plus_ms_aux xs ys))" by (subst plus_ms_aux.code) simp lemma hd_exp_plus: "ms_aux_hd_exp (plus_ms_aux xs ys) = combine_options max (ms_aux_hd_exp xs) (ms_aux_hd_exp ys)" by (cases xs; cases ys) (simp_all add: plus_ms_aux_MSLCons ms_aux_hd_exp_def max_def split: prod.split) lemma minus_ms_aux_MSLNil_left [simp]: "minus_ms_aux MSLNil ys = uminus_ms_aux ys" by (simp add: minus_ms_aux_def) lemma minus_ms_aux_MSLNil_right [simp]: "minus_ms_aux xs MSLNil = xs" by (simp add: minus_ms_aux_def) lemma times_ms_aux_MSLNil_iff: "times_ms_aux F G = MSLNil \ F = MSLNil \ G = MSLNil" by (subst times_ms_aux.code) (simp split: msllist.splits) lemma times_ms_aux_MSLNil [simp]: "times_ms_aux MSLNil G = MSLNil" "times_ms_aux F MSLNil = MSLNil" by (subst times_ms_aux.code, simp split: msllist.splits)+ lemma times_ms_aux_MSLCons: "times_ms_aux (MSLCons x xs) (MSLCons y ys) = MSLCons (fst x * fst y, snd x + snd y) (plus_ms_aux (scale_shift_ms_aux x ys) (times_ms_aux xs (MSLCons y ys)))" by (subst times_ms_aux_altdef) simp lemma hd_exp_times: "ms_aux_hd_exp (times_ms_aux xs ys) = (case (ms_aux_hd_exp xs, ms_aux_hd_exp ys) of (Some x, Some y) \ Some (x + y) | _ \ None)" by (cases xs; cases ys) (simp_all add: times_ms_aux_MSLCons ms_aux_hd_exp_def split: prod.split) subsubsection \Induction upto friends for multiseries\ inductive ms_closure :: "(('a :: multiseries \ real) msllist \ (real \ real) \ basis \ bool) \ ('a :: multiseries \ real) msllist \ (real \ real) \ basis \ bool" for P where ms_closure_embed: "P xs f basis \ ms_closure P xs f basis" | ms_closure_cong: "ms_closure P xs f basis \ eventually (\x. f x = g x) at_top \ ms_closure P xs g basis" | ms_closure_MSLCons: "ms_closure P xs (\x. f x - eval C x * b x powr e) (b # basis) \ is_expansion C basis \ (\e'. e' > e \ f \ o(\x. b x powr e')) \ ms_exp_gt e (ms_aux_hd_exp xs) \ ms_closure P (MSLCons (C, e) xs) f (b # basis)" | ms_closure_add: "ms_closure P xs f basis \ ms_closure P ys g basis \ ms_closure P (plus_ms_aux xs ys) (\x. f x + g x) basis" | ms_closure_mult: "ms_closure P xs f basis \ ms_closure P ys g basis \ ms_closure P (times_ms_aux xs ys) (\x. f x * g x) basis" | ms_closure_basis: "ms_closure P xs f (b # basis) \ is_expansion C basis \ ms_closure P (scale_shift_ms_aux (C,e) xs) (\x. eval C x * b x powr e * f x) (b # basis)" | ms_closure_embed': "is_expansion_aux xs f basis \ ms_closure P xs f basis" lemma is_expansion_aux_coinduct_upto [consumes 2, case_names A B]: fixes xs :: "('a :: multiseries \ real) msllist" assumes *: "X xs f basis" and basis: "basis_wf basis" and step: "\xs f basis. X xs f basis \ (xs = MSLNil \ eventually (\x. f x = 0) at_top \ length basis = Suc (expansion_level TYPE('a))) \ (\xs' b e basis' C. xs = MSLCons (C, e) xs' \ basis = b # basis' \ ms_closure X xs' (\x. f x - eval C x * b x powr e) (b # basis') \ is_expansion C basis' \ (\e'>e. f \ o(\x. b x powr e')) \ ms_exp_gt e (ms_aux_hd_exp xs'))" shows "is_expansion_aux xs f basis" proof - from * have "ms_closure X xs f basis" by (rule ms_closure_embed[of X xs f basis]) thus ?thesis proof (coinduction arbitrary: xs f rule: is_expansion_aux_coinduct) case (is_expansion_aux xs f) from this and basis show ?case proof (induction rule: ms_closure.induct) case (ms_closure_embed xs f basis) from step[OF ms_closure_embed(1)] show ?case by auto next case (ms_closure_MSLCons xs f C b e basis) thus ?case by auto next case (ms_closure_cong xs f basis g) note ev = \\\<^sub>F x in at_top. f x = g x\ hence ev_zero_iff: "eventually (\x. f x = 0) at_top \ eventually (\x. g x = 0) at_top" by (rule eventually_subst') have *: "ms_closure X xs' (\x. f x - c x * b x powr e) basis \ ms_closure X xs' (\x. g x - c x * b x powr e) basis" for xs' c b e by (rule iffI; erule ms_closure.intros(2)) (insert ev, auto elim: eventually_mono) from ms_closure_cong show ?case by (simp add: ev_zero_iff * landau_o.small.in_cong[OF ev]) next case (ms_closure_embed' xs f basis) from this show ?case by (cases rule: is_expansion_aux.cases) (force intro: ms_closure.intros(7))+ next case (ms_closure_basis xs f b basis C e) show ?case proof (cases xs rule: ms_aux_cases) case MSLNil with ms_closure_basis show ?thesis by (auto simp: scale_shift_ms_aux_def split: prod.split elim: eventually_mono) next case (MSLCons C' e' xs') with ms_closure_basis have IH: "ms_closure X (MSLCons (C', e') xs') f (b # basis)" "is_expansion C basis" "xs = MSLCons (C', e') xs'" "ms_closure X xs' (\x. f x - eval C' x * b x powr e') (b # basis)" "ms_exp_gt e' (ms_aux_hd_exp xs')" "is_expansion C' basis" "\x. x > e' \ f \ o(\xa. b xa powr x)" by auto { fix e'' :: real assume *: "e + e' < e''" define d where "d = (e'' - e - e') / 2" from * have "d > 0" by (simp add: d_def) hence "(\x. eval C x * b x powr e * f x) \ o(\x. b x powr d * b x powr e * b x powr (e'+d))" using ms_closure_basis(4) IH by (intro landau_o.small.mult[OF landau_o.small_big_mult] IH is_expansion_imp_smallo[OF _ IH(2)]) (simp_all add: basis_wf_Cons) also have "\ = o(\x. b x powr e'')" by (simp add: d_def powr_add [symmetric]) finally have "(\x. eval C x * b x powr e * f x) \ \" . } moreover have "ms_closure X xs' (\x. f x - eval C' x * b x powr e') (b # basis)" using ms_closure_basis IH by auto note ms_closure.intros(6)[OF this IH(2), of e] moreover have "ms_exp_gt (e + e') (ms_aux_hd_exp (scale_shift_ms_aux (C, e) xs'))" using \ms_exp_gt e' (ms_aux_hd_exp xs')\ by (cases xs') (simp_all add: hd_exp_basis scale_shift_ms_aux_def ) ultimately show ?thesis using IH(2,6) MSLCons ms_closure_basis(4) by (auto simp: scale_shift_ms_aux_MSLCons algebra_simps powr_add basis_wf_Cons intro: is_expansion_mult) qed next case (ms_closure_add xs f basis ys g) show ?case proof (cases xs ys rule: ms_aux_cases[case_product ms_aux_cases]) case MSLNil_MSLNil with ms_closure_add have "eventually (\x. f x = 0) at_top" "eventually (\x. g x = 0) at_top" by simp_all hence "eventually (\x. f x + g x = 0) at_top" by eventually_elim simp with MSLNil_MSLNil ms_closure_add show ?thesis by simp next case (MSLNil_MSLCons c e xs') with ms_closure_add have "eventually (\x. f x = 0) at_top" by simp hence ev: "eventually (\x. f x + g x = g x) at_top" by eventually_elim simp have *: "ms_closure X xs (\x. f x + g x - y x) basis \ ms_closure X xs (\x. g x - y x) basis" for y basis xs by (rule iffI; erule ms_closure_cong) (insert ev, simp_all) from MSLNil_MSLCons ms_closure_add show ?thesis by (simp add: * landau_o.small.in_cong[OF ev]) next case (MSLCons_MSLNil c e xs') with ms_closure_add have "eventually (\x. g x = 0) at_top" by simp hence ev: "eventually (\x. f x + g x = f x) at_top" by eventually_elim simp have *: "ms_closure X xs (\x. f x + g x - y x) basis \ ms_closure X xs (\x. f x - y x) basis" for y basis xs by (rule iffI; erule ms_closure_cong) (insert ev, simp_all) from MSLCons_MSLNil ms_closure_add show ?thesis by (simp add: * landau_o.small.in_cong[OF ev]) next case (MSLCons_MSLCons C1 e1 xs' C2 e2 ys') with ms_closure_add obtain b basis' where IH: "basis_wf (b # basis')" "basis = b # basis'" "ms_closure X (MSLCons (C1, e1) xs') f (b # basis')" "ms_closure X (MSLCons (C2, e2) ys') g (b # basis')" "xs = MSLCons (C1, e1) xs'" "ys = MSLCons (C2, e2) ys'" "ms_closure X xs' (\x. f x - eval C1 x * b x powr e1) (b # basis')" "ms_closure X ys' (\x. g x - eval C2 x * b x powr e2) (b # basis')" "is_expansion C1 basis'" "\e. e > e1 \ f \ o(\x. b x powr e)" "is_expansion C2 basis'" "\e. e > e2 \ g \ o(\x. b x powr e)" "ms_exp_gt e1 (ms_aux_hd_exp xs')" "ms_exp_gt e2 (ms_aux_hd_exp ys')" by blast have o: "(\x. f x + g x) \ o(\x. b x powr e)" if "e > max e1 e2" for e using that by (intro sum_in_smallo IH) simp_all show ?thesis proof (cases e1 e2 rule: linorder_cases) case less have gt: "ms_exp_gt e2 (combine_options max (Some e1) (ms_aux_hd_exp ys'))" using \ms_exp_gt e2 _\ less by (cases "ms_aux_hd_exp ys'") auto have "ms_closure X (plus_ms_aux xs ys') (\x. f x + (g x - eval C2 x * b x powr e2)) (b # basis')" by (rule ms_closure.intros(4)) (insert IH, simp_all) with less IH(2,11,14) o gt show ?thesis by (intro disjI2) (simp add: MSLCons_MSLCons plus_ms_aux_MSLCons algebra_simps hd_exp_plus) next case greater have gt: "ms_exp_gt e1 (combine_options max (ms_aux_hd_exp xs') (Some e2))" using \ms_exp_gt e1 _\ greater by (cases "ms_aux_hd_exp xs'") auto have "ms_closure X (plus_ms_aux xs' ys) (\x. (f x - eval C1 x * b x powr e1) + g x) (b # basis')" by (rule ms_closure.intros(4)) (insert IH, simp_all) with greater IH(2,9,13) o gt show ?thesis by (intro disjI2) (simp add: MSLCons_MSLCons plus_ms_aux_MSLCons algebra_simps hd_exp_plus) next case equal have gt: "ms_exp_gt e2 (combine_options max (ms_aux_hd_exp xs') (ms_aux_hd_exp ys'))" using \ms_exp_gt e1 _\ \ms_exp_gt e2 _\ equal by (cases "ms_aux_hd_exp xs'"; cases "ms_aux_hd_exp ys'") auto have "ms_closure X (plus_ms_aux xs' ys') (\x. (f x - eval C1 x * b x powr e1) + (g x - eval C2 x * b x powr e2)) (b # basis')" by (rule ms_closure.intros(4)) (insert IH, simp_all) with equal IH(1,2,9,11,13,14) o gt show ?thesis by (intro disjI2) (auto intro: is_expansion_add simp: plus_ms_aux_MSLCons basis_wf_Cons algebra_simps hd_exp_plus MSLCons_MSLCons) qed qed next case (ms_closure_mult xs f basis ys g) show ?case proof (cases "xs = MSLNil \ ys = MSLNil") case True with ms_closure_mult have "eventually (\x. f x = 0) at_top \ eventually (\x. g x = 0) at_top" by blast hence "eventually (\x. f x * g x = 0) at_top" by (auto elim: eventually_mono) with ms_closure_mult True show ?thesis by auto next case False with ms_closure_mult obtain C1 e1 xs' C2 e2 ys' b basis' where IH: "xs = MSLCons (C1, e1) xs'" "ys = MSLCons (C2, e2) ys'" "basis_wf (b # basis')" "basis = b # basis'" "ms_closure X (MSLCons (C1, e1) xs') f (b # basis')" "ms_closure X (MSLCons (C2, e2) ys') g (b # basis')" "ms_closure X xs' (\x. f x - eval C1 x * b x powr e1) (b # basis')" "ms_closure X ys' (\x. g x - eval C2 x * b x powr e2) (b # basis')" "is_expansion C1 basis'" "\e. e > e1 \ f \ o(\x. b x powr e)" "is_expansion C2 basis'" "\e. e > e2 \ g \ o(\x. b x powr e)" "ms_exp_gt e1 (ms_aux_hd_exp xs')" "ms_exp_gt e2 (ms_aux_hd_exp ys')" by blast have o: "(\x. f x * g x) \ o(\x. b x powr e)" if "e > e1 + e2" for e proof - define d where "d = (e - e1 - e2) / 2" from that have d: "d > 0" by (simp add: d_def) hence "(\x. f x * g x) \ o(\x. b x powr (e1 + d) * b x powr (e2 + d))" by (intro landau_o.small.mult IH) simp_all also have "\ = o(\x. b x powr e)" by (simp add: d_def powr_add [symmetric]) finally show ?thesis . qed have "ms_closure X (plus_ms_aux (scale_shift_ms_aux (C1, e1) ys') (times_ms_aux xs' ys)) (\x. eval C1 x * b x powr e1 * (g x - eval C2 x * b x powr e2) + ((f x - eval C1 x * b x powr e1) * g x)) (b # basis')" (is "ms_closure _ ?zs ?f _") using ms_closure_mult IH(4) by (intro ms_closure.intros(4-6) IH) simp_all also have "?f = (\x. f x * g x - eval C1 x * eval C2 x * b x powr (e1 + e2))" by (intro ext) (simp add: algebra_simps powr_add) finally have "ms_closure X ?zs \ (b # basis')" . moreover from \ms_exp_gt e1 _\ \ms_exp_gt e2 _\ have "ms_exp_gt (e1 + e2) (combine_options max (map_option ((+) e1) (ms_aux_hd_exp ys')) (case ms_aux_hd_exp xs' of None \ None | Some x \ (case Some e2 of None \ None | Some y \ Some (x + y))))" by (cases "ms_aux_hd_exp xs'"; cases "ms_aux_hd_exp ys'") (simp_all add: hd_exp_times hd_exp_plus hd_exp_basis IH) ultimately show ?thesis using IH(1,2,3,4,9,11) o by (auto simp: times_ms_aux_MSLCons basis_wf_Cons hd_exp_times hd_exp_plus hd_exp_basis intro: is_expansion_mult) qed qed qed qed subsubsection \Dominant terms\ lemma one_smallo_powr: "e > (0::real) \ (\_. 1) \ o(\x. x powr e)" by (auto simp: smallomega_iff_smallo [symmetric] real_powr_at_top smallomegaI_filterlim_at_top_norm) lemma powr_smallo_one: "e < (0::real) \ (\x. x powr e) \ o(\_. 1)" by (intro smalloI_tendsto) (auto intro!: tendsto_neg_powr filterlim_ident) lemma eval_monom_smallo': assumes "basis_wf (b # basis)" "e > 0" shows "eval_monom x basis \ o(\x. b x powr e)" proof (cases x) case (Pair c es) from assms show ?thesis unfolding Pair proof (induction es arbitrary: b basis e) case (Nil b basis e) hence b: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) have "eval_monom (c, []) basis \ O(\_. 1)" by simp also have "(\_. 1) \ o(\x. b x powr e)" by (rule landau_o.small.compose[OF _ b]) (insert Nil, simp add: one_smallo_powr) finally show ?case . next case (Cons e' es b basis e) show ?case proof (cases basis) case Nil from Cons have b: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) from Nil have "eval_monom (c, e' # es) basis \ O(\_. 1)" by simp also have "(\_. 1) \ o(\x. b x powr e)" by (rule landau_o.small.compose[OF _ b]) (insert Cons.prems, simp add: one_smallo_powr) finally show ?thesis . next case (Cons b' basis') from Cons.prems have "eval_monom (c, es) basis' \ o(\x. b' x powr 1)" by (intro Cons.IH) (simp_all add: basis_wf_Cons Cons) hence "(\x. eval_monom (c, es) basis' x * b' x powr e') \ o(\x. b' x powr 1 * b' x powr e')" by (rule landau_o.small_big_mult) simp_all also have "\ = o(\x. b' x powr (1 + e'))" by (simp add: powr_add) also from Cons.prems have "(\x. b' x powr (1 + e')) \ o(\x. b x powr e)" by (intro ln_smallo_imp_flat) (simp_all add: basis_wf_Cons Cons) finally show ?thesis by (simp add: powr_add [symmetric] Cons eval_monom_Cons) qed qed qed lemma eval_monom_smallomega': assumes "basis_wf (b # basis)" "e < 0" shows "eval_monom (1, es) basis \ \(\x. b x powr e)" using assms proof (induction es arbitrary: b basis e) case (Nil b basis e) hence b: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) have "eval_monom (1, []) basis \ \(\_. 1)" by simp also have "(\_. 1) \ \(\x. b x powr e)" unfolding smallomega_iff_smallo by (rule landau_o.small.compose[OF _ b]) (insert Nil, simp add: powr_smallo_one) finally show ?case . next case (Cons e' es b basis e) show ?case proof (cases basis) case Nil from Cons have b: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) from Nil have "eval_monom (1, e' # es) basis \ \(\_. 1)" by simp also have "(\_. 1) \ \(\x. b x powr e)" unfolding smallomega_iff_smallo by (rule landau_o.small.compose[OF _ b]) (insert Cons.prems, simp add: powr_smallo_one) finally show ?thesis . next case (Cons b' basis') from Cons.prems have "eval_monom (1, es) basis' \ \(\x. b' x powr -1)" by (intro Cons.IH) (simp_all add: basis_wf_Cons Cons) hence "(\x. eval_monom (1, es) basis' x * b' x powr e') \ \(\x. b' x powr -1 * b' x powr e')" by (rule landau_omega.small_big_mult) simp_all also have "\ = \(\x. b' x powr (e' - 1))" by (simp add: powr_diff powr_minus field_simps) also from Cons.prems have "(\x. b' x powr (e' - 1)) \ \(\x. b x powr e)" unfolding smallomega_iff_smallo by (intro ln_smallo_imp_flat') (simp_all add: basis_wf_Cons Cons) finally show ?thesis by (simp add: powr_add [symmetric] Cons eval_monom_Cons) qed qed lemma dominant_term_ms_aux: assumes basis: "basis_wf basis" and xs: "trimmed_ms_aux xs" "is_expansion_aux xs f basis" shows "f \[at_top] eval_monom (dominant_term_ms_aux xs) basis" (is ?thesis1) and "eventually (\x. sgn (f x) = sgn (fst (dominant_term_ms_aux xs))) at_top" (is ?thesis2) proof - include asymp_equiv_notation from xs(2,1) obtain xs' C b e basis' where xs': "trimmed C" "xs = MSLCons (C, e) xs'" "basis = b # basis'" "is_expansion_aux xs' (\x. f x - eval C x * b x powr e) (b # basis')" "is_expansion C basis'" "(\e'. e < e' \ f \ o(\x. b x powr e'))" "ms_exp_gt e (ms_aux_hd_exp xs')" by (cases rule: is_expansion_aux.cases) (auto simp: trimmed_ms_aux_MSLCons) from is_expansion_aux_imp_smallo''[OF xs'(4,7)] obtain e' where e': "e' < e" "(\x. f x - eval C x * b x powr e) \ o(\x. b x powr e')" unfolding list.sel by blast from basis xs' have "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) note e'(2) also have "o(\x. b x powr e') = o(\x. b x powr (e' - e) * b x powr e)" by (simp add: powr_add [symmetric]) also from xs' basis e'(1) have "eval C \ \(\x. b x powr (e' - e))" by (intro is_expansion_imp_smallomega[of basis']) (simp_all add: basis_wf_Cons) hence "(\x. b x powr (e' - e) * b x powr e) \ o(\x. eval C x * b x powr e)" by (intro landau_o.small_big_mult) (simp_all add: smallomega_iff_smallo) finally have *: "(\x. f x - eval C x * b x powr e) \ o(\x. eval C x * b x powr e)" . hence "f \ (\x. eval C x * b x powr e)" by (intro smallo_imp_asymp_equiv) simp also from basis xs' have "eval C \ (\x. eval_monom (dominant_term C) basis' x)" by (intro dominant_term) (simp_all add: basis_wf_Cons) also have "(\a. eval_monom (dominant_term C) basis' a * b a powr e) = eval_monom (dominant_term_ms_aux xs) basis" by (auto simp add: dominant_term_ms_aux_def xs' eval_monom_Cons fun_eq_iff split: prod.split) finally show ?thesis1 by - (simp_all add: asymp_equiv_intros) have "eventually (\x. sgn (eval C x * b x powr e + (f x - eval C x * b x powr e)) = sgn (eval C x * b x powr e)) at_top" by (intro smallo_imp_eventually_sgn *) moreover have "eventually (\x. sgn (eval C x) = sgn (fst (dominant_term C))) at_top" using xs' basis by (intro trimmed_imp_eventually_sgn[of basis']) (simp_all add: basis_wf_Cons) ultimately have "eventually (\x. sgn (f x) = sgn (fst (dominant_term C))) at_top" using b_pos by eventually_elim (simp_all add: sgn_mult) thus ?thesis2 using xs'(2) by (auto simp: dominant_term_ms_aux_def split: prod.split) qed lemma eval_monom_smallo: assumes "basis \ []" "basis_wf basis" "length es = length basis" "e > hd es" shows "eval_monom (c, es) basis \ o(\x. hd basis x powr e)" proof (cases es; cases basis) fix b basis' e' es' assume [simp]: "basis = b # basis'" "es = e' # es'" from assms have "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" using filterlim_at_top_dense by blast have "(\x. eval_monom (1, es') basis' x * b x powr e') \ o(\x. b x powr (e - e') * b x powr e')" using assms by (intro landau_o.small_big_mult eval_monom_smallo') auto also have "(\x. b x powr (e - e') * b x powr e') \ \(\x. b x powr e)" by (intro bigthetaI_cong eventually_mono[OF b_pos]) (auto simp: powr_diff) finally show ?thesis by (simp add: eval_monom_def algebra_simps) qed (insert assms, auto) lemma eval_monom_smallomega: assumes "basis \ []" "basis_wf basis" "length es = length basis" "e < hd es" shows "eval_monom (1, es) basis \ \(\x. hd basis x powr e)" proof (cases es; cases basis) fix b basis' e' es' assume [simp]: "basis = b # basis'" "es = e' # es'" from assms have "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" using filterlim_at_top_dense by blast have "(\x. eval_monom (1, es') basis' x * b x powr e') \ \(\x. b x powr (e - e') * b x powr e')" using assms by (intro landau_omega.small_big_mult eval_monom_smallomega') auto also have "(\x. b x powr (e - e') * b x powr e') \ \(\x. b x powr e)" by (intro bigthetaI_cong eventually_mono[OF b_pos]) (auto simp: powr_diff) finally show ?thesis by (simp add: eval_monom_Cons) qed (insert assms, auto) subsubsection \Correctness of multiseries operations\ lemma drop_zero_ms_aux: assumes "eventually (\x. eval C x = 0) at_top" assumes "is_expansion_aux (MSLCons (C, e) xs) f basis" shows "is_expansion_aux xs f basis" proof (rule is_expansion_aux_cong) from assms(2) show "is_expansion_aux xs (\x. f x - eval C x * hd basis x powr e) basis" by (auto elim: is_expansion_aux_MSLCons) from assms(1) show "eventually (\x. f x - eval C x * hd basis x powr e = f x) at_top" by eventually_elim simp qed lemma dominant_term_ms_aux_bigo: assumes basis: "basis_wf basis" and xs: "is_expansion_aux xs f basis" shows "f \ O(eval_monom (1, snd (dominant_term_ms_aux xs)) basis)" (is ?thesis1) proof (cases xs rule: ms_aux_cases) case MSLNil with assms have "eventually (\x. f x = 0) at_top" by (auto elim: is_expansion_aux.cases) hence "f \ \(\_. 0)" by (intro bigthetaI_cong) auto also have "(\_. 0) \ O(eval_monom (1, snd (dominant_term_ms_aux xs)) basis)" by simp finally show ?thesis . next case [simp]: (MSLCons C e xs') from xs obtain b basis' where xs': "basis = b # basis'" "is_expansion_aux xs' (\x. f x - eval C x * b x powr e) (b # basis')" "is_expansion C basis'" "(\e'. e < e' \ f \ o(\x. b x powr e'))" "ms_exp_gt e (ms_aux_hd_exp xs')" by (cases rule: is_expansion_aux.cases) auto from is_expansion_aux_imp_smallo''[OF xs'(2,5)] obtain e' where e': "e' < e" "(\x. f x - eval C x * b x powr e) \ o(\x. b x powr e')" unfolding list.sel by blast from basis xs' have "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) let ?h = "(\x. eval_monom (1, snd (dominant_term C)) basis' x * b x powr e)" note e'(2) also have "(\x. b x powr e') \ \(\x. b x powr (e' - e) * b x powr e)" by (intro bigthetaI_cong eventually_mono[OF b_pos]) (auto simp: powr_diff) also have "(\x. b x powr (e' - e) * b x powr e) \ o(?h)" unfolding smallomega_iff_smallo [symmetric] using e'(1) basis xs' by (intro landau_omega.small_big_mult landau_omega.big_refl eval_monom_smallomega') (simp_all add: basis_wf_Cons) finally have "(\x. f x - eval C x * b x powr e) \ O(?h)" by (rule landau_o.small_imp_big) moreover have "(\x. eval C x * b x powr e) \ O(?h)" using basis xs' by (intro landau_o.big.mult dominant_term_bigo) (auto simp: basis_wf_Cons) ultimately have "(\x. f x - eval C x * b x powr e + eval C x * b x powr e) \ O(?h)" by (rule sum_in_bigo) hence "f \ O(?h)" by simp also have "?h = eval_monom (1, snd (dominant_term_ms_aux xs)) basis" using xs' by (auto simp: dominant_term_ms_aux_def case_prod_unfold eval_monom_Cons) finally show ?thesis . qed lemma const_ms_aux: assumes basis: "basis_wf basis" and "length basis = Suc (expansion_level TYPE('a::multiseries))" shows "is_expansion_aux (const_ms_aux c :: ('a \ real) msllist) (\_. c) basis" proof - from assms(2) obtain b basis' where [simp]: "basis = b # basis'" by (cases basis) simp_all from basis have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) have "(\_. c) \ o(\x. b x powr e)" if "e > 0" for e proof - have "(\_. c) \ O(\_. 1)" by (cases "c = 0") simp_all also from b_pos have "(\_. 1) \ \(\x. b x powr 0)" by (intro bigthetaI_cong) (auto elim: eventually_mono) also from that b_limit have "(\x. b x powr 0) \ o(\x. b x powr e)" by (subst powr_smallo_iff) auto finally show ?thesis . qed with b_pos assms show ?thesis by (auto intro!: is_expansion_aux.intros simp: const_ms_aux_def is_expansion_const basis_wf_Cons elim: eventually_mono) qed lemma uminus_ms_aux: assumes basis: "basis_wf basis" assumes F: "is_expansion_aux F f basis" shows "is_expansion_aux (uminus_ms_aux F) (\x. -f x) basis" using F proof (coinduction arbitrary: F f rule: is_expansion_aux_coinduct) case (is_expansion_aux F f) note IH = is_expansion_aux show ?case proof (cases F rule: ms_aux_cases) case MSLNil with IH show ?thesis by (auto simp: uminus_ms_aux_def elim: is_expansion_aux.cases) next case (MSLCons C e xs) with IH obtain b basis' where IH: "basis = b # basis'" "is_expansion_aux xs (\x. f x - eval C x * b x powr e) (b # basis')" "is_expansion C basis'" "\e'. e < e' \ f \ o(\x. b x powr e')" "ms_exp_gt e (ms_aux_hd_exp xs)" by (auto elim: is_expansion_aux_MSLCons) from basis IH have basis': "basis_wf basis'" by (simp add: basis_wf_Cons) with IH basis show ?thesis by (force simp: uminus_ms_aux_MSLCons MSLCons is_expansion_uminus) qed qed lemma plus_ms_aux: assumes "basis_wf basis" "is_expansion_aux F f basis" "is_expansion_aux G g basis" shows "is_expansion_aux (plus_ms_aux F G) (\x. f x + g x) basis" using assms proof (coinduction arbitrary: F f G g rule: is_expansion_aux_coinduct) case (is_expansion_aux F f G g) note IH = this show ?case proof (cases F G rule: ms_aux_cases[case_product ms_aux_cases]) assume [simp]: "F = MSLNil" "G = MSLNil" with IH have *: "eventually (\x. f x = 0) at_top" "eventually (\x. g x = 0) at_top" and length: "length basis = Suc (expansion_level TYPE('a))" by (auto elim: is_expansion_aux.cases) from * have "eventually (\x. f x + g x = 0) at_top" by eventually_elim simp with length show ?case by simp next assume [simp]: "F = MSLNil" fix C e G' assume G: "G = MSLCons (C, e) G'" from IH have f: "eventually (\x. f x = 0) at_top" by (auto elim: is_expansion_aux.cases) from f have "eventually (\x. f x + g x = g x) at_top" by eventually_elim simp note eq = landau_o.small.in_cong[OF this] from IH(3) G obtain b basis' where G': "basis = b # basis'" "is_expansion_aux G' (\x. g x - eval C x * b x powr e) (b # basis')" "is_expansion C basis'" "\e'>e. g \ o(\x. b x powr e')" "ms_exp_gt e (ms_aux_hd_exp G')" by (auto elim!: is_expansion_aux_MSLCons) show ?thesis by (rule disjI2, inst_existentials G' b e basis' C F f G' "\x. g x - eval C x * b x powr e") (insert G' IH(1,2), simp_all add: G eq algebra_simps) next assume [simp]: "G = MSLNil" fix C e F' assume F: "F = MSLCons (C, e) F'" from IH have g: "eventually (\x. g x = 0) at_top" by (auto elim: is_expansion_aux.cases) hence "eventually (\x. f x + g x = f x) at_top" by eventually_elim simp note eq = landau_o.small.in_cong[OF this] from IH F obtain b basis' where F': "basis = b # basis'" "is_expansion_aux F' (\x. f x - eval C x * b x powr e) (b # basis')" "is_expansion C basis'" "\e'>e. f \ o(\x. b x powr e')" "ms_exp_gt e (ms_aux_hd_exp F')" by (auto elim!: is_expansion_aux_MSLCons) show ?thesis by (rule disjI2, inst_existentials F' b e basis' C F' "\x. f x - eval C x * b x powr e" G g) (insert F g F' IH, simp_all add: eq algebra_simps) next fix C1 e1 F' C2 e2 G' assume F: "F = MSLCons (C1, e1) F'" and G: "G = MSLCons (C2, e2) G'" from IH F obtain b basis' where basis': "basis = b # basis'" and F': "is_expansion_aux F' (\x. f x - eval C1 x * b x powr e1) (b # basis')" "is_expansion C1 basis'" "\e'. e' > e1 \ f \ o(\x. b x powr e')" "ms_exp_gt e1 (ms_aux_hd_exp F')" by (auto elim!: is_expansion_aux_MSLCons) from IH G basis' have G': "is_expansion_aux G' (\x. g x - eval C2 x * b x powr e2) (b # basis')" "is_expansion C2 basis'" "\e'. e' > e2 \ g \ o(\x. b x powr e')" "ms_exp_gt e2 (ms_aux_hd_exp G')" by (auto elim!: is_expansion_aux_MSLCons) hence *: "\x>max e1 e2. (\x. f x + g x) \ o(\xa. b xa powr x)" by (intro allI impI sum_in_smallo F' G') simp_all consider "e1 > e2" | "e1 < e2" | "e1 = e2" by force thus ?thesis proof cases assume e: "e1 > e2" have gt: "ms_exp_gt e1 (combine_options max (ms_aux_hd_exp F') (Some e2))" using \ms_exp_gt e1 _\ e by (cases "ms_aux_hd_exp F'") simp_all show ?thesis by (rule disjI2, inst_existentials "plus_ms_aux F' G" b e1 basis' C1 F' "\x. f x - eval C1 x * b x powr e1" G g) (insert e F'(1,2,4) IH(1,3) basis'(1) * gt, simp_all add: F G plus_ms_aux_MSLCons algebra_simps hd_exp_plus) next assume e: "e1 < e2" have gt: "ms_exp_gt e2 (combine_options max (Some e1) (ms_aux_hd_exp G'))" using \ms_exp_gt e2 _\ e by (cases "ms_aux_hd_exp G'") simp_all show ?thesis by (rule disjI2, inst_existentials "plus_ms_aux F G'" b e2 basis' C2 F f G' "\x. g x - eval C2 x * b x powr e2") (insert e G'(1,2,4) IH(1,2) basis'(1) * gt, simp_all add: F G plus_ms_aux_MSLCons algebra_simps hd_exp_plus) next assume e: "e1 = e2" have gt: "ms_exp_gt e2 (combine_options max (ms_aux_hd_exp F') (ms_aux_hd_exp G'))" using \ms_exp_gt e1 _\ \ms_exp_gt e2 _\ e by (cases "ms_aux_hd_exp F'"; cases "ms_aux_hd_exp G'") simp_all show ?thesis by (rule disjI2, inst_existentials "plus_ms_aux F' G'" b e1 basis' "C1 + C2" F' "\x. f x - eval C1 x * b x powr e1" G' "\x. g x - eval C2 x * b x powr e2") (insert e F'(1,2,4) G'(1,2,4) IH(1) basis'(1) * gt, simp_all add: F G plus_ms_aux_MSLCons algebra_simps hd_exp_plus is_expansion_add basis_wf_Cons) qed qed qed lemma minus_ms_aux: assumes "basis_wf basis" "is_expansion_aux F f basis" "is_expansion_aux G g basis" shows "is_expansion_aux (minus_ms_aux F G) (\x. f x - g x) basis" proof - have "is_expansion_aux (minus_ms_aux F G) (\x. f x + (- g x)) basis" unfolding minus_ms_aux_def by (intro plus_ms_aux uminus_ms_aux assms) thus ?thesis by simp qed lemma scale_shift_ms_aux: assumes basis: "basis_wf (b # basis)" assumes F: "is_expansion_aux F f (b # basis)" assumes C: "is_expansion C basis" shows "is_expansion_aux (scale_shift_ms_aux (C, e) F) (\x. eval C x * b x powr e * f x) (b # basis)" using F proof (coinduction arbitrary: F f rule: is_expansion_aux_coinduct) case (is_expansion_aux F f) note IH = is_expansion_aux show ?case proof (cases F rule: ms_aux_cases) case MSLNil with IH show ?thesis by (auto simp: scale_shift_ms_aux_def elim!: is_expansion_aux.cases eventually_mono) next case (MSLCons C' e' xs) with IH have IH: "is_expansion_aux xs (\x. f x - eval C' x * b x powr e') (b # basis)" "is_expansion C' basis" "ms_exp_gt e' (ms_aux_hd_exp xs)" "\e''. e' < e'' \ f \ o(\x. b x powr e'')" by (auto elim: is_expansion_aux_MSLCons) have "(\x. eval C x * b x powr e * f x) \ o(\x. b x powr e'')" if "e'' > e + e'" for e'' proof - define d where "d = (e'' - e - e') / 2" from that have d: "d > 0" by (simp add: d_def) have "(\x. eval C x * b x powr e * f x) \ o(\x. b x powr d * b x powr e * b x powr (e' + d))" using d basis by (intro landau_o.small.mult[OF landau_o.small_big_mult] is_expansion_imp_smallo[OF _ C] IH) (simp_all add: basis_wf_Cons) also have "\ = o(\x. b x powr e'')" by (simp add: d_def powr_add [symmetric]) finally show ?thesis . qed moreover have "ms_exp_gt (e + e') (ms_aux_hd_exp (scale_shift_ms_aux (C, e) xs))" using IH(3) by (cases "ms_aux_hd_exp xs") (simp_all add: hd_exp_basis) ultimately show ?thesis using basis IH(1,2) by (intro disjI2, inst_existentials "scale_shift_ms_aux (C, e) xs" b "e + e'" basis "C * C'" xs "\x. f x - eval C' x * b x powr e'") (simp_all add: scale_shift_ms_aux_MSLCons MSLCons is_expansion_mult basis_wf_Cons algebra_simps powr_add C) qed qed lemma times_ms_aux: assumes basis: "basis_wf basis" assumes F: "is_expansion_aux F f basis" and G: "is_expansion_aux G g basis" shows "is_expansion_aux (times_ms_aux F G) (\x. f x * g x) basis" using F G proof (coinduction arbitrary: F f G g rule: is_expansion_aux_coinduct_upto) case (B F f G g) note IH = this show ?case proof (cases "F = MSLNil \ G = MSLNil") case True with IH have "eventually (\x. f x = 0) at_top \ eventually (\x. g x = 0) at_top" and length: "length basis = Suc (expansion_level TYPE('a))" by (auto elim: is_expansion_aux.cases) from this(1) have "eventually (\x. f x * g x = 0) at_top" by (auto elim!: eventually_mono) with length True show ?thesis by auto next case False from False obtain C1 e1 F' where F: "F = MSLCons (C1, e1) F'" by (cases F rule: ms_aux_cases) simp_all from False obtain C2 e2 G' where G: "G = MSLCons (C2, e2) G'" by (cases G rule: ms_aux_cases) simp_all from IH(1) F obtain b basis' where basis': "basis = b # basis'" and F': "is_expansion_aux F' (\x. f x - eval C1 x * b x powr e1) (b # basis')" "is_expansion C1 basis'" "\e'. e' > e1 \ f \ o(\x. b x powr e')" "ms_exp_gt e1 (ms_aux_hd_exp F')" by (auto elim!: is_expansion_aux_MSLCons) from IH(2) G basis' have G': "is_expansion_aux G' (\x. g x - eval C2 x * b x powr e2) (b # basis')" "is_expansion C2 basis'" "\e'. e' > e2 \ g \ o(\x. b x powr e')" "ms_exp_gt e2 (ms_aux_hd_exp G')" by (auto elim!: is_expansion_aux_MSLCons) let ?P = "(\xs'' fa'' basisa''. \F f G. xs'' = times_ms_aux F G \ (\g. fa'' = (\x. f x * g x) \ basisa'' = b # basis' \ is_expansion_aux F f (b # basis') \ is_expansion_aux G g (b # basis')))" have "ms_closure ?P (plus_ms_aux (scale_shift_ms_aux (C1, e1) G') (times_ms_aux F' G)) (\x. eval C1 x * b x powr e1 * (g x - eval C2 x * b x powr e2) + (f x - eval C1 x * b x powr e1) * g x) (b # basis')" (is "ms_closure _ ?zs _ _") using IH(2) basis' basis by (intro ms_closure_add ms_closure_embed'[OF scale_shift_ms_aux] ms_closure_mult[OF ms_closure_embed' ms_closure_embed'] F' G') simp_all hence "ms_closure ?P (plus_ms_aux (scale_shift_ms_aux (C1, e1) G') (times_ms_aux F' G)) (\x. f x * g x - eval C1 x * eval C2 x * b x powr (e1 + e2)) (b # basis')" by (simp add: algebra_simps powr_add) moreover have "(\x. f x * g x) \ o(\x. b x powr e)" if "e > e1 + e2" for e proof - define d where "d = (e - e1 - e2) / 2" from that have "d > 0" by (simp add: d_def) hence "(\x. f x * g x) \ o(\x. b x powr (e1 + d) * b x powr (e2 + d))" by (intro landau_o.small.mult F' G') simp_all also have "\ = o(\x. b x powr e)" by (simp add: d_def powr_add [symmetric]) finally show ?thesis . qed moreover from \ms_exp_gt e1 _\ \ms_exp_gt e2 _\ have "ms_exp_gt (e1 + e2) (ms_aux_hd_exp ?zs)" by (cases "ms_aux_hd_exp F'"; cases "ms_aux_hd_exp G'") (simp_all add: hd_exp_times hd_exp_plus hd_exp_basis G) ultimately show ?thesis using F'(2) G'(2) basis' basis by (simp add: times_ms_aux_MSLCons basis_wf_Cons is_expansion_mult F G) qed qed (insert basis, simp_all) lemma powser_ms_aux_MSLNil_iff [simp]: "powser_ms_aux cs f = MSLNil \ cs = MSLNil" by (subst powser_ms_aux.code) (simp split: msllist.splits) lemma powser_ms_aux_MSLNil [simp]: "powser_ms_aux MSLNil f = MSLNil" by (subst powser_ms_aux.code) simp lemma powser_ms_aux_MSLNil' [simp]: "powser_ms_aux (MSLCons c cs) MSLNil = MSLCons (const_expansion c, 0) MSLNil" by (subst powser_ms_aux.code) simp lemma powser_ms_aux_MSLCons: "powser_ms_aux (MSLCons c cs) f = MSLCons (const_expansion c, 0) (times_ms_aux f (powser_ms_aux cs f))" by (subst powser_ms_aux.code) simp lemma hd_exp_powser: "ms_aux_hd_exp (powser_ms_aux cs f) = (if cs = MSLNil then None else Some 0)" by (subst powser_ms_aux.code) (simp split: msllist.splits) lemma powser_ms_aux: assumes "convergent_powser cs" and basis: "basis_wf basis" assumes G: "is_expansion_aux G g basis" "ms_exp_gt 0 (ms_aux_hd_exp G)" shows "is_expansion_aux (powser_ms_aux cs G) (powser cs \ g) basis" using assms(1) proof (coinduction arbitrary: cs rule: is_expansion_aux_coinduct_upto) case (B cs) show ?case proof (cases cs) case MSLNil with G show ?thesis by (auto simp: is_expansion_aux_expansion_level) next case (MSLCons c cs') from is_expansion_aux_basis_nonempty[OF G(1)] obtain b basis' where basis': "basis = b # basis'" by (cases basis) simp_all from B have conv: "convergent_powser cs'" by (auto simp: MSLCons dest: convergent_powser_stl) from basis basis' have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) from G(2) have "g \ o(\x. hd basis x powr 0)" by (intro is_expansion_aux_imp_smallo[OF G(1)]) simp with basis' have "g \ o(\x. b x powr 0)" by simp also have "(\x. b x powr 0) \ \(\x. 1)" using b_pos by (intro bigthetaI_cong) (auto elim!: eventually_mono) finally have g_limit: "(g \ 0) at_top" by (auto dest: smalloD_tendsto) let ?P = "(\xs'' fa'' basisa''. \cs. xs'' = powser_ms_aux cs G \ fa'' = powser cs \ g \ basisa'' = b # basis' \ convergent_powser cs)" have "ms_closure ?P (times_ms_aux G (powser_ms_aux cs' G)) (\x. g x * (powser cs' \ g) x) basis" using conv by (intro ms_closure_mult [OF ms_closure_embed' ms_closure_embed] G) (auto simp add: basis' o_def) also have "(\x. g x * (powser cs' \ g) x) = (\x. x * powser cs' x) \ g" by (simp add: o_def) finally have "ms_closure ?P (times_ms_aux G (powser_ms_aux cs' G)) (\x. (powser cs \ g) x - c * b x powr 0) basis" unfolding o_def proof (rule ms_closure_cong) note b_pos moreover have "eventually (\x. powser cs (g x) = g x * powser cs' (g x) + c) at_top" proof - from B have "eventually (\x. powser cs x = x * powser cs' x + c) (nhds 0)" by (simp add: MSLCons powser_MSLCons) from this and g_limit show ?thesis by (rule eventually_compose_filterlim) qed ultimately show "eventually (\x. g x * powser cs' (g x) = powser cs (g x) - c * b x powr 0) at_top" by eventually_elim simp qed moreover from basis G have "is_expansion (const_expansion c :: 'a) basis'" by (intro is_expansion_const) (auto dest: is_expansion_aux_expansion_level simp: basis' basis_wf_Cons) moreover have "powser cs \ g \ o(\x. b x powr e)" if "e > 0" for e proof - have "((powser cs \ g) \ powser cs 0) at_top" unfolding o_def by (intro isCont_tendsto_compose[OF _ g_limit] isCont_powser B) hence "powser cs \ g \ O(\_. 1)" by (intro bigoI_tendsto[where c = "powser cs 0"]) (simp_all add: o_def) also from b_pos have "O(\_. 1) = O(\x. b x powr 0)" by (intro landau_o.big.cong) (auto elim: eventually_mono) also from that b_limit have "(\x. b x powr 0) \ o(\x. b x powr e)" by (subst powr_smallo_iff) simp_all finally show ?thesis . qed moreover from G have "ms_exp_gt 0 (ms_aux_hd_exp (times_ms_aux G (powser_ms_aux cs' G)))" by (cases "ms_aux_hd_exp G") (simp_all add: hd_exp_times MSLCons hd_exp_powser) ultimately show ?thesis by (simp add: MSLCons powser_ms_aux_MSLCons basis') qed qed (insert assms, simp_all) lemma powser_ms_aux': assumes powser: "convergent_powser' cs f" and basis: "basis_wf basis" assumes G: "is_expansion_aux G g basis" "ms_exp_gt 0 (ms_aux_hd_exp G)" shows "is_expansion_aux (powser_ms_aux cs G) (f \ g) basis" proof (rule is_expansion_aux_cong) from is_expansion_aux_basis_nonempty[OF G(1)] basis have "filterlim (hd basis) at_top at_top" by (cases basis) (simp_all add: basis_wf_Cons) hence pos: "eventually (\x. hd basis x > 0) at_top" by (simp add: filterlim_at_top_dense) from G(2) have "g \ o(\x. hd basis x powr 0)" by (intro is_expansion_aux_imp_smallo[OF G(1)]) simp also have "(\x. hd basis x powr 0) \ \(\x. 1)" using pos by (intro bigthetaI_cong) (auto elim!: eventually_mono) finally have g_limit: "(g \ 0) at_top" by (auto dest: smalloD_tendsto) from powser have "eventually (\x. powser cs x = f x) (nhds 0)" by (rule convergent_powser'_eventually) from this and g_limit show "eventually (\x. (powser cs \ g) x = (f \ g) x) at_top" unfolding o_def by (rule eventually_compose_filterlim) qed (rule assms powser_ms_aux convergent_powser'_imp_convergent_powser)+ lemma inverse_ms_aux: assumes basis: "basis_wf basis" assumes F: "is_expansion_aux F f basis" "trimmed_ms_aux F" shows "is_expansion_aux (inverse_ms_aux F) (\x. inverse (f x)) basis" proof (cases F rule: ms_aux_cases) case (MSLCons C e F') from F MSLCons obtain b basis' where F': "basis = b # basis'" "trimmed C" "is_expansion_aux F' (\x. f x - eval C x * b x powr e) (b # basis')" "is_expansion C basis'" "(\e'. e < e' \ f \ o(\x. b x powr e'))" "ms_exp_gt e (ms_aux_hd_exp F')" by (auto elim!: is_expansion_aux_MSLCons simp: trimmed_ms_aux_def) define f' where "f' = (\x. f x - eval C x * b x powr e)" from basis F' have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) moreover from F'(6) have "ms_exp_gt 0 (ms_aux_hd_exp (scale_shift_ms_aux (inverse C, - e) F'))" by (cases "ms_aux_hd_exp F'") (simp_all add: hd_exp_basis) ultimately have "is_expansion_aux (inverse_ms_aux F) (\x. eval (inverse C) x * b x powr (-e) * ((\x. inverse (1 + x)) \ (\x. eval (inverse C) x * b x powr (-e) * f' x)) x) (b # basis')" (is "is_expansion_aux _ ?h _") unfolding MSLCons inverse_ms_aux.simps Let_def fst_conv snd_conv f'_def eval_inverse [symmetric] using basis F(2) F'(1,3,4) by (intro scale_shift_ms_aux powser_ms_aux' is_expansion_inverse) (simp_all add: convergent_powser'_geometric basis_wf_Cons trimmed_ms_aux_def MSLCons) also have "b # basis' = basis" by (simp add: F') finally show ?thesis proof (rule is_expansion_aux_cong, goal_cases) case 1 from basis F' have "eventually (\x. eval C x \ 0) at_top" by (intro trimmed_imp_eventually_nz[of basis']) (simp_all add: basis_wf_Cons) with b_pos show ?case by eventually_elim (simp add: o_def field_simps powr_minus f'_def) qed qed (insert assms, auto simp: trimmed_ms_aux_def) lemma hd_exp_inverse: "xs \ MSLNil \ ms_aux_hd_exp (inverse_ms_aux xs) = map_option uminus (ms_aux_hd_exp xs)" by (cases xs) (auto simp: Let_def hd_exp_basis hd_exp_powser inverse_ms_aux.simps) lemma eval_pos_if_dominant_term_pos: assumes "basis_wf basis" "is_expansion F basis" "trimmed F" "fst (dominant_term F) > 0" shows "eventually (\x. eval F x > 0) at_top" proof - have "eval F \[at_top] eval_monom (dominant_term F) basis" by (intro dominant_term assms) from trimmed_imp_eventually_sgn[OF assms(1-3)] have "\\<^sub>F x in at_top. sgn (eval F x) = sgn (fst (dominant_term F))" . moreover from assms have "eventually (\x. eval_monom (dominant_term F) basis x > 0) at_top" by (intro eval_monom_pos) ultimately show ?thesis by eventually_elim (insert assms, simp add: sgn_if split: if_splits) qed lemma eval_pos_if_dominant_term_pos': assumes "basis_wf basis" "trimmed_ms_aux F" "is_expansion_aux F f basis" "fst (dominant_term_ms_aux F) > 0" shows "eventually (\x. f x > 0) at_top" proof - have "f \[at_top] eval_monom (dominant_term_ms_aux F) basis" by (intro dominant_term_ms_aux assms) from dominant_term_ms_aux(2)[OF assms(1-3)] have "\\<^sub>F x in at_top. sgn (f x) = sgn (fst (dominant_term_ms_aux F))" . moreover from assms have "eventually (\x. eval_monom (dominant_term_ms_aux F) basis x > 0) at_top" by (intro eval_monom_pos) ultimately show ?thesis by eventually_elim (insert assms, simp add: sgn_if split: if_splits) qed lemma eval_neg_if_dominant_term_neg': assumes "basis_wf basis" "trimmed_ms_aux F" "is_expansion_aux F f basis" "fst (dominant_term_ms_aux F) < 0" shows "eventually (\x. f x < 0) at_top" proof - have "f \[at_top] eval_monom (dominant_term_ms_aux F) basis" by (intro dominant_term_ms_aux assms) from dominant_term_ms_aux(2)[OF assms(1-3)] have "\\<^sub>F x in at_top. sgn (f x) = sgn (fst (dominant_term_ms_aux F))" . moreover from assms have "eventually (\x. eval_monom (dominant_term_ms_aux F) basis x < 0) at_top" by (intro eval_monom_neg) ultimately show ?thesis by eventually_elim (insert assms, simp add: sgn_if split: if_splits) qed lemma fst_dominant_term_ms_aux_MSLCons: "fst (dominant_term_ms_aux (MSLCons x xs)) = fst (dominant_term (fst x))" by (auto simp: dominant_term_ms_aux_def split: prod.splits) lemma powr_ms_aux: assumes basis: "basis_wf basis" assumes F: "is_expansion_aux F f basis" "trimmed_ms_aux F" "fst (dominant_term_ms_aux F) > 0" shows "is_expansion_aux (powr_ms_aux abort F p) (\x. f x powr p) basis" proof (cases F rule: ms_aux_cases) case (MSLCons C e F') from F MSLCons obtain b basis' where F': "basis = b # basis'" "trimmed C" "fst (dominant_term C) > 0" "is_expansion_aux F' (\x. f x - eval C x * b x powr e) (b # basis')" "is_expansion C basis'" "(\e'. e < e' \ f \ o(\x. b x powr e'))" "ms_exp_gt e (ms_aux_hd_exp F')" by (auto elim!: is_expansion_aux_MSLCons simp: trimmed_ms_aux_def dominant_term_ms_aux_def split: prod.splits) define f' where "f' = (\x. f x - eval C x * b x powr e)" from basis F' have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) moreover from F' have "ms_exp_gt 0 (ms_aux_hd_exp (scale_shift_ms_aux (inverse C, - e) F'))" by (cases "ms_aux_hd_exp F'") (simp_all add: hd_exp_basis) ultimately have "is_expansion_aux (powr_ms_aux abort F p) (\x. eval (powr_expansion abort C p) x * b x powr (e * p) * ((\x. (1 + x) powr p) \ (\x. eval (inverse C) x * b x powr (-e) * f' x)) x) (b # basis')" (is "is_expansion_aux _ ?h _") unfolding MSLCons powr_ms_aux.simps Let_def fst_conv snd_conv f'_def eval_inverse [symmetric] using basis F'(1-5) by (intro scale_shift_ms_aux powser_ms_aux' is_expansion_inverse is_expansion_powr) (simp_all add: MSLCons basis_wf_Cons convergent_powser'_gbinomial) also have "b # basis' = basis" by (simp add: F') finally show ?thesis proof (rule is_expansion_aux_cong, goal_cases) case 1 from basis F' have "eventually (\x. eval C x > 0) at_top" by (intro eval_pos_if_dominant_term_pos[of basis']) (simp_all add: basis_wf_Cons) moreover from basis F have "eventually (\x. f x > 0) at_top" by (intro eval_pos_if_dominant_term_pos'[of basis F]) ultimately show ?case using b_pos by eventually_elim (simp add: powr_powr [symmetric] powr_minus powr_mult powr_divide f'_def field_simps) qed qed (insert assms, auto simp: trimmed_ms_aux_def) lemma power_ms_aux: assumes basis: "basis_wf basis" assumes F: "is_expansion_aux F f basis" "trimmed_ms_aux F" shows "is_expansion_aux (power_ms_aux abort F n) (\x. f x ^ n) basis" proof (cases F rule: ms_aux_cases) case (MSLCons C e F') from F MSLCons obtain b basis' where F': "basis = b # basis'" "trimmed C" "is_expansion_aux F' (\x. f x - eval C x * b x powr e) (b # basis')" "is_expansion C basis'" "(\e'. e < e' \ f \ o(\x. b x powr e'))" "ms_exp_gt e (ms_aux_hd_exp F')" by (auto elim!: is_expansion_aux_MSLCons simp: trimmed_ms_aux_def dominant_term_ms_aux_def split: prod.splits) define f' where "f' = (\x. f x - eval C x * b x powr e)" from basis F' have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) moreover have "ms_exp_gt 0 (ms_aux_hd_exp (scale_shift_ms_aux (inverse C, - e) F'))" using F'(6) by (cases "ms_aux_hd_exp F'") (simp_all add: hd_exp_basis) ultimately have "is_expansion_aux (power_ms_aux abort F n) (\x. eval (power_expansion abort C n) x * b x powr (e * real n) * ((\x. (1 + x) ^ n) \ (\x. eval (inverse C) x * b x powr (-e) * f' x)) x) (b # basis')" (is "is_expansion_aux _ ?h _") unfolding MSLCons power_ms_aux.simps Let_def fst_conv snd_conv f'_def eval_inverse [symmetric] using basis F'(1-4) by (intro scale_shift_ms_aux powser_ms_aux' is_expansion_inverse is_expansion_power) (simp_all add: MSLCons basis_wf_Cons convergent_powser'_gbinomial') also have "b # basis' = basis" by (simp add: F') finally show ?thesis proof (rule is_expansion_aux_cong, goal_cases) case 1 from F'(1,2,4) basis have "eventually (\x. eval C x \ 0) at_top" using trimmed_imp_eventually_nz[of basis' C] by (simp add: basis_wf_Cons) thus ?case using b_pos by eventually_elim (simp add: field_simps f'_def powr_minus powr_powr [symmetric] powr_realpow power_mult_distrib [symmetric] power_divide [symmetric] del: power_mult_distrib power_divide) qed qed (insert assms, auto simp: trimmed_ms_aux_def) lemma snd_dominant_term_ms_aux_MSLCons: "snd (dominant_term_ms_aux (MSLCons (C, e) xs)) = e # snd (dominant_term C)" by (simp add: dominant_term_ms_aux_def case_prod_unfold) subsubsection \Type-class instantiations\ datatype 'a ms = MS "('a \ real) msllist" "real \ real" instantiation ms :: (uminus) uminus begin primrec uminus_ms where "-(MS xs f) = MS (uminus_ms_aux xs) (\x. -f x)" instance .. end instantiation ms :: (plus) plus begin fun plus_ms :: "'a ms \ 'a ms \ 'a ms" where "MS xs f + MS ys g = MS (plus_ms_aux xs ys) (\x. f x + g x)" instance .. end instantiation ms :: ("{plus,uminus}") minus begin definition minus_ms :: "'a ms \ 'a ms \ 'a ms" where "F - G = F + -(G :: 'a ms)" instance .. end instantiation ms :: ("{plus,times}") times begin fun times_ms :: "'a ms \ 'a ms \ 'a ms" where "MS xs f * MS ys g = MS (times_ms_aux xs ys) (\x. f x * g x)" instance .. end instantiation ms :: (multiseries) inverse begin fun inverse_ms :: "'a ms \ 'a ms" where "inverse_ms (MS xs f) = MS (inverse_ms_aux xs) (\x. inverse (f x))" fun divide_ms :: "'a ms \ 'a ms \ 'a ms" where "divide_ms (MS xs f) (MS ys g) = MS (times_ms_aux xs (inverse_ms_aux ys)) (\x. f x / g x)" instance .. end instantiation ms :: (multiseries) multiseries begin definition expansion_level_ms :: "'a ms itself \ nat" where expansion_level_ms_def [simp]: "expansion_level_ms _ = Suc (expansion_level (TYPE('a)))" definition zero_expansion_ms :: "'a ms" where "zero_expansion = MS MSLNil (\_. 0)" definition const_expansion_ms :: "real \ 'a ms" where "const_expansion c = MS (const_ms_aux c) (\_. c)" primrec is_expansion_ms :: "'a ms \ basis \ bool" where "is_expansion (MS xs f) = is_expansion_aux xs f" lemma is_expansion_ms_def': "is_expansion F = (case F of MS xs f \ is_expansion_aux xs f)" by (simp add: is_expansion_ms_def split: ms.splits) primrec eval_ms :: "'a ms \ real \ real" where "eval_ms (MS _ f) = f" lemma eval_ms_def': "eval F = (case F of MS _ f \ f)" by (cases F) simp_all primrec powr_expansion_ms :: "bool \ 'a ms \ real \ 'a ms" where "powr_expansion_ms abort (MS xs f) p = MS (powr_ms_aux abort xs p) (\x. f x powr p)" primrec power_expansion_ms :: "bool \ 'a ms \ nat \ 'a ms" where "power_expansion_ms abort (MS xs f) n = MS (power_ms_aux abort xs n) (\x. f x ^ n)" primrec trimmed_ms :: "'a ms \ bool" where "trimmed_ms (MS xs _) \ trimmed_ms_aux xs" primrec dominant_term_ms :: "'a ms \ monom" where "dominant_term_ms (MS xs _) = dominant_term_ms_aux xs" lemma length_dominant_term_ms: "length (snd (dominant_term (F :: 'a ms))) = Suc (expansion_level TYPE('a))" by (cases F) (simp_all add: length_dominant_term) instance proof (standard, goal_cases length_basis zero const uminus plus minus times inverse divide powr power smallo smallomega trimmed dominant dominant_bigo) case (smallo basis F b e) from \is_expansion F basis\ obtain xs f where F: "F = MS xs f" "is_expansion_aux xs f basis" by (simp add: is_expansion_ms_def' split: ms.splits) from F(2) have nonempty: "basis \ []" by (rule is_expansion_aux_basis_nonempty) with smallo have filterlim_hd_basis: "filterlim (hd basis) at_top at_top" by (cases basis) (simp_all add: basis_wf_Cons) from F(2) obtain e' where "f \ o(\x. hd basis x powr e')" by (erule is_expansion_aux_imp_smallo') also from smallo nonempty filterlim_hd_basis have "(\x. hd basis x powr e') \ o(\x. b x powr e)" by (intro ln_smallo_imp_flat) auto finally show ?case by (simp add: F) next case (smallomega basis F b e) obtain xs f where F: "F = MS xs f" by (cases F) simp_all from this smallomega have *: "is_expansion_aux xs f basis" by simp with smallomega F have "f \ \(\x. b x powr e)" by (intro is_expansion_aux_imp_smallomega [OF _ *]) (simp_all add: is_expansion_ms_def' split: ms.splits) with F show ?case by simp next case (minus basis F G) thus ?case by (simp add: minus_ms_def is_expansion_ms_def' add_uminus_conv_diff [symmetric] plus_ms_aux uminus_ms_aux del: add_uminus_conv_diff split: ms.splits) next case (divide basis F G) have "G / F = G * inverse F" by (cases F; cases G) (simp add: divide_inverse) with divide show ?case by (simp add: is_expansion_ms_def' divide_inverse times_ms_aux inverse_ms_aux split: ms.splits) next fix c :: real show "trimmed (const_expansion c :: 'a ms) \ c \ 0" by (simp add: const_expansion_ms_def trimmed_ms_aux_def const_ms_aux_def trimmed_const_expansion split: msllist.splits) next fix F :: "'a ms" assume "trimmed F" thus "fst (dominant_term F) \ 0" by (cases F) (auto simp: dominant_term_ms_aux_def trimmed_ms_aux_MSLCons case_prod_unfold trimmed_imp_dominant_term_nz split: msllist.splits) next fix F :: "'a ms" have "times_ms_aux (MSLCons (const_expansion 1, 0) MSLNil) xs = xs" for xs :: "('a \ real) msllist" proof (coinduction arbitrary: xs rule: msllist.coinduct_upto) case Eq_real_prod_msllist have "map_prod ((*) (const_expansion 1 :: 'a)) ((+) (0::real)) = (\x. x)" by (rule ext) (simp add: map_prod_def times_const_expansion_1 case_prod_unfold) moreover have "mslmap \ = (\x. x)" by (rule ext) (simp add: msllist.map_ident) ultimately have "scale_shift_ms_aux (const_expansion 1 :: 'a, 0) = (\x. x)" by (simp add: scale_shift_ms_aux_conv_mslmap) thus ?case by (cases xs rule: ms_aux_cases) (auto simp: times_ms_aux_MSLCons times_const_expansion_1 times_ms_aux.corec.cong_refl) qed thus "const_expansion 1 * F = F" by (cases F) (simp add: const_expansion_ms_def const_ms_aux_def) next fix F :: "'a ms" show "fst (dominant_term (- F)) = - fst (dominant_term F)" "trimmed (-F) \ trimmed F" by (cases F; simp add: dominant_term_ms_aux_def case_prod_unfold uminus_ms_aux_MSLCons trimmed_ms_aux_def split: msllist.splits)+ next fix F :: "'a ms" show "zero_expansion + F = F" "F + zero_expansion = F" by (cases F; simp add: zero_expansion_ms_def)+ qed (auto simp: eval_ms_def' const_expansion_ms_def trimmed_ms_aux_imp_nz is_expansion_ms_def' const_ms_aux uminus_ms_aux plus_ms_aux times_ms_aux inverse_ms_aux is_expansion_aux_expansion_level dominant_term_ms_aux length_dominant_term_ms minus_ms_def powr_ms_aux power_ms_aux zero_expansion_ms_def is_expansion_aux.intros(1) dominant_term_ms_aux_bigo split: ms.splits) end subsubsection \Changing the evaluation function of a multiseries\ definition modify_eval :: "(real \ real) \ 'a ms \ 'a ms" where "modify_eval f F = (case F of MS xs _ \ MS xs f)" lemma eval_modify_eval [simp]: "eval (modify_eval f F) = f" by (cases F) (simp add: modify_eval_def) subsubsection \Scaling expansions\ definition scale_ms :: "real \ 'a :: multiseries \ 'a" where "scale_ms c F = const_expansion c * F" lemma scale_ms_real [simp]: "scale_ms c (c' :: real) = c * c'" by (simp add: scale_ms_def) definition scale_ms_aux :: "real \ ('a :: multiseries \ real) msllist \ ('a \ real) msllist" where "scale_ms_aux c xs = scale_shift_ms_aux (const_expansion c, 0) xs" lemma scale_ms_aux_eq_MSLNil_iff [simp]: "scale_ms_aux x xs = MSLNil \ xs = MSLNil" unfolding scale_ms_aux_def by simp lemma times_ms_aux_singleton: "times_ms_aux (MSLCons (c, e) MSLNil) xs = scale_shift_ms_aux (c, e) xs" proof (coinduction arbitrary: xs rule: msllist.coinduct_strong) case (Eq_msllist xs) thus ?case by (cases xs rule: ms_aux_cases) (simp_all add: scale_shift_ms_aux_def times_ms_aux_MSLCons) qed lemma scale_ms [simp]: "scale_ms c (MS xs f) = MS (scale_ms_aux c xs) (\x. c * f x)" by (simp add: scale_ms_def scale_ms_aux_def const_expansion_ms_def const_ms_aux_def times_ms_aux_singleton) lemma scale_ms_aux_MSLNil [simp]: "scale_ms_aux c MSLNil = MSLNil" by (simp add: scale_ms_aux_def) lemma scale_ms_aux_MSLCons: "scale_ms_aux c (MSLCons (c', e) xs) = MSLCons (scale_ms c c', e) (scale_ms_aux c xs)" by (simp add: scale_ms_aux_def scale_shift_ms_aux_MSLCons scale_ms_def) subsubsection \Additional convenience rules\ lemma trimmed_pos_real_iff [simp]: "trimmed_pos (x :: real) \ x > 0" by (auto simp: trimmed_pos_def) lemma trimmed_pos_ms_iff: "trimmed_pos (MS xs f) \ (case xs of MSLNil \ False | MSLCons x xs \ trimmed_pos (fst x))" by (auto simp add: trimmed_pos_def dominant_term_ms_aux_def trimmed_ms_aux_def split: msllist.splits prod.splits) lemma not_trimmed_pos_MSLNil [simp]: "\trimmed_pos (MS MSLNil f)" by (simp add: trimmed_pos_ms_iff) lemma trimmed_pos_MSLCons [simp]: "trimmed_pos (MS (MSLCons x xs) f) = trimmed_pos (fst x)" by (simp add: trimmed_pos_ms_iff) lemma drop_zero_ms: assumes "eventually (\x. eval C x = 0) at_top" assumes "is_expansion (MS (MSLCons (C, e) xs) f) basis" shows "is_expansion (MS xs f) basis" using assms by (auto simp: is_expansion_ms_def intro: drop_zero_ms_aux) lemma expansion_level_eq_Nil: "length [] = expansion_level TYPE(real)" by simp lemma expansion_level_eq_Cons: "length xs = expansion_level TYPE('a::multiseries) \ length (x # xs) = expansion_level TYPE('a ms)" by simp lemma basis_wf_insert_ln: assumes "basis_wf [b]" shows "basis_wf [\x. ln (b x)]" (is ?thesis1) "basis_wf [b, \x. ln (b x)]" (is ?thesis2) "is_expansion (MS (MSLCons (1::real,1) MSLNil) (\x. ln (b x))) [\x. ln (b x)]" (is ?thesis3) proof - have "ln \ o(\x. x :: real)" using ln_x_over_x_tendsto_0 by (intro smalloI_tendsto) auto with assms show ?thesis1 ?thesis2 by (auto simp: basis_wf_Cons intro!: landau_o.small.compose[of _ _ _ "\x. ln (b x)"] filterlim_compose[OF ln_at_top]) from assms have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence ev: "eventually (\x. b x > 1) at_top" by (simp add: filterlim_at_top_dense) have "(\x::real. x) \ \(\x. x powr 1)" by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto also have "(\x::real. x powr 1) \ o(\a. a powr e)" if "e > 1" for e by (subst powr_smallo_iff) (auto simp: that filterlim_ident) finally show ?thesis3 using assms ev by (auto intro!: is_expansion_aux.intros landau_o.small.compose[of _ at_top _ "\x. ln (b x)"] filterlim_compose[OF ln_at_top b_limit] is_expansion_real.intros elim!: eventually_mono) qed lemma basis_wf_lift_modification: assumes "basis_wf (b' # b # bs)" "basis_wf (b # bs')" shows "basis_wf (b' # b # bs')" using assms by (simp add: basis_wf_many) lemma default_basis_wf: "basis_wf [\x. x]" by (simp add: basis_wf_Cons filterlim_ident) subsubsection \Lifting expansions\ definition is_lifting where "is_lifting f basis basis' \ (\C. eval (f C) = eval C \ (is_expansion C basis \ is_expansion (f C) basis') \ trimmed (f C) = trimmed C \ fst (dominant_term (f C)) = fst (dominant_term C))" lemma is_liftingD: assumes "is_lifting f basis basis'" shows "eval (f C) = eval C" "is_expansion C basis \ is_expansion (f C) basis'" "trimmed (f C) \ trimmed C" "fst (dominant_term (f C)) = fst (dominant_term C)" using assms [unfolded is_lifting_def] unfolding is_lifting_def by blast+ definition lift_ms :: "'a :: multiseries \ 'a ms" where "lift_ms C = MS (MSLCons (C, 0) MSLNil) (eval C)" lemma is_expansion_lift: assumes "basis_wf (b # basis)" "is_expansion C basis" shows "is_expansion (lift_ms C) (b # basis)" proof - from assms have "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) moreover from assms have "eval C \ o(\x. b x powr e)" if "e > 0" for e using that by (intro is_expansion_imp_smallo[OF _ assms(2)]) (simp_all add: basis_wf_Cons) ultimately show ?thesis using assms by (auto intro!: is_expansion_aux.intros simp: lift_ms_def is_expansion_length elim: eventually_mono) qed lemma eval_lift_ms [simp]: "eval (lift_ms C) = eval C" by (simp add: lift_ms_def) lemma is_lifting_lift: assumes "basis_wf (b # basis)" shows "is_lifting lift_ms basis (b # basis)" using is_expansion_lift[OF assms] unfolding is_lifting_def by (auto simp: lift_ms_def trimmed_ms_aux_MSLCons dominant_term_ms_aux_def case_prod_unfold) definition map_ms_aux :: "('a :: multiseries \ 'b :: multiseries) \ ('a \ real) msllist \ ('b \ real) msllist" where "map_ms_aux f xs = mslmap (\(c,e). (f c, e)) xs" lemma map_ms_aux_MSLNil [simp]: "map_ms_aux f MSLNil = MSLNil" by (simp add: map_ms_aux_def) lemma map_ms_aux_MSLCons: "map_ms_aux f (MSLCons (C, e) xs) = MSLCons (f C, e) (map_ms_aux f xs)" by (simp add: map_ms_aux_def) lemma hd_exp_map [simp]: "ms_aux_hd_exp (map_ms_aux f xs) = ms_aux_hd_exp xs" by (simp add: ms_aux_hd_exp_def map_ms_aux_def split: msllist.splits) lemma map_ms_altdef: "map_ms f (MS xs g) = MS (map_ms_aux f xs) g" by (simp add: map_ms_aux_def map_prod_def) lemma eval_map_ms [simp]: "eval (map_ms f F) = eval F" by (cases F) simp_all lemma is_expansion_map_aux: fixes f :: "'a :: multiseries \ 'b :: multiseries" assumes "is_expansion_aux xs g (b # basis)" assumes "\C. is_expansion C basis \ is_expansion (f C) basis'" assumes "length basis' = expansion_level TYPE('b)" assumes "\C. eval (f C) = eval C" shows "is_expansion_aux (map_ms_aux f xs) g (b # basis')" using assms(1) proof (coinduction arbitrary: xs g rule: is_expansion_aux_coinduct) case (is_expansion_aux xs g) show ?case proof (cases xs rule: ms_aux_cases) case MSLNil with is_expansion_aux show ?thesis by (auto simp: assms elim: is_expansion_aux.cases) next case (MSLCons c e xs') with is_expansion_aux show ?thesis by (auto elim!: is_expansion_aux_MSLCons simp: map_ms_aux_MSLCons assms) qed qed lemma is_expansion_map: fixes f :: "'a :: multiseries \ 'b :: multiseries" and F :: "'a ms" assumes "is_expansion G (b # basis)" assumes "\C. is_expansion C basis \ is_expansion (f C) basis'" assumes "\C. eval (f C) = eval C" assumes "length basis' = expansion_level TYPE('b)" shows "is_expansion (map_ms f G) (b # basis')" using assms by (cases G, simp only: map_ms_altdef) (auto intro!: is_expansion_map_aux) lemma is_expansion_map_final: fixes f :: "'a :: multiseries \ 'b :: multiseries" and F :: "'a ms" assumes "is_lifting f basis basis'" assumes "is_expansion G (b # basis)" assumes "length basis' = expansion_level TYPE('b)" shows "is_expansion (map_ms f G) (b # basis')" by (rule is_expansion_map[OF assms(2)]) (insert assms, simp_all add: is_lifting_def) lemma fst_dominant_term_map_ms: "is_lifting f basis basis' \ fst (dominant_term (map_ms f C)) = fst (dominant_term C)" by (cases C) (simp add: dominant_term_ms_aux_def case_prod_unfold is_lifting_def split: msllist.splits) lemma trimmed_map_ms: "is_lifting f basis basis' \ trimmed (map_ms f C) \ trimmed C" by (cases C) (simp add: trimmed_ms_aux_def is_lifting_def split: msllist.splits) lemma is_lifting_map: fixes f :: "'a :: multiseries \ 'b :: multiseries" and F :: "'a ms" assumes "is_lifting f basis basis'" assumes "length basis' = expansion_level TYPE('b)" shows "is_lifting (map_ms f) (b # basis) (b # basis')" unfolding is_lifting_def by (intro allI conjI impI is_expansion_map_final[OF assms(1)]) (insert assms, simp_all add: is_lifting_def fst_dominant_term_map_ms[OF assms(1)] trimmed_map_ms[OF assms(1)]) lemma is_lifting_id: "is_lifting (\x. x) basis basis" by (simp add: is_lifting_def) lemma is_lifting_trans: "is_lifting f basis basis' \ is_lifting g basis' basis'' \ is_lifting (\x. g (f x)) basis basis''" by (auto simp: is_lifting_def) lemma is_expansion_X: "is_expansion (MS (MSLCons (1 :: real, 1) MSLNil) (\x. x)) [\x. x]" proof - have "(\x::real. x) \ \(\x. x powr 1)" by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto also have "(\x::real. x powr 1) \ o(\a. a powr e)" if "e > 1" for e by (subst powr_smallo_iff) (auto simp: that filterlim_ident) finally show ?thesis by (auto intro!: is_expansion_aux.intros is_expansion_real.intros eventually_mono[OF eventually_gt_at_top[of "0::real"]]) qed inductive expands_to :: "(real \ real) \ 'a :: multiseries \ basis \ bool" (infix "(expands'_to)" 50) where "is_expansion F basis \ eventually (\x. eval F x = f x) at_top \ (f expands_to F) basis" lemma dominant_term_expands_to: assumes "basis_wf basis" "(f expands_to F) basis" "trimmed F" shows "f \[at_top] eval_monom (dominant_term F) basis" proof - from assms have "eval F \[at_top] f" by (intro asymp_equiv_refl_ev) (simp add: expands_to.simps) also from dominant_term[OF assms(1) _ assms(3)] assms(2) have "eval F \[at_top] eval_monom (dominant_term F) basis" by (simp add: expands_to.simps) finally show ?thesis by (subst asymp_equiv_sym) simp_all qed lemma expands_to_cong: "(f expands_to F) basis \ eventually (\x. f x = g x) at_top \ (g expands_to F) basis" by (auto simp: expands_to.simps elim: eventually_elim2) lemma expands_to_cong': assumes "(f expands_to MS xs g) basis" "eventually (\x. g x = g' x) at_top" shows "(f expands_to MS xs g') basis" proof - have "is_expansion_aux xs g' basis" by (rule is_expansion_aux_cong [OF _ assms(2)]) (insert assms(1), simp add: expands_to.simps) with assms show ?thesis by (auto simp: expands_to.simps elim: eventually_elim2) qed lemma eval_expands_to: "(f expands_to F) basis \ eventually (\x. eval F x = f x) at_top" by (simp add: expands_to.simps) lemma expands_to_real_compose: assumes "(f expands_to (c::real)) bs" shows "((\x. g (f x)) expands_to g c) bs" using assms by (auto simp: expands_to.simps is_expansion_real.simps elim: eventually_mono) lemma expands_to_lift_compose: assumes "(f expands_to MS (MSLCons (C, e) xs) f') bs'" assumes "eventually (\x. f' x - h x = 0) at_top" "e = 0" assumes "((\x. g (h x)) expands_to G) bs" "basis_wf (b # bs)" shows "((\x. g (f x)) expands_to lift_ms G) (b # bs)" proof - from assms have "\\<^sub>F x in at_top. f' x = f x" "\\<^sub>F x in at_top. eval G x = g (h x)" by (auto simp: expands_to.simps) with assms(2) have "\\<^sub>F x in at_top. eval G x = g (f x)" by eventually_elim simp with assms show ?thesis by (intro expands_to.intros is_expansion_lift) (auto simp: expands_to.simps) qed lemma expands_to_zero: "basis_wf basis \ length basis = expansion_level TYPE('a) \ ((\_. 0) expands_to (zero_expansion :: 'a :: multiseries)) basis" by (auto simp add: expands_to.simps is_expansion_zero) lemma expands_to_const: "basis_wf basis \ length basis = expansion_level TYPE('a) \ ((\_. c) expands_to (const_expansion c :: 'a :: multiseries)) basis" by (auto simp add: expands_to.simps is_expansion_const) lemma expands_to_X: "((\x. x) expands_to MS (MSLCons (1 :: real, 1) MSLNil) (\x. x)) [\x. x]" using is_expansion_X by (simp add: expands_to.simps) lemma expands_to_uminus: "basis_wf basis \ (f expands_to F) basis \ ((\x. -f x) expands_to -F) basis" by (auto simp: expands_to.simps is_expansion_uminus) lemma expands_to_add: "basis_wf basis \ (f expands_to F) basis \ (g expands_to G) basis \ ((\x. f x + g x) expands_to F + G) basis" by (auto simp: expands_to.simps is_expansion_add elim: eventually_elim2) lemma expands_to_minus: "basis_wf basis \ (f expands_to F) basis \ (g expands_to G) basis \ ((\x. f x - g x) expands_to F - G) basis" by (auto simp: expands_to.simps is_expansion_minus elim: eventually_elim2) lemma expands_to_mult: "basis_wf basis \ (f expands_to F) basis \ (g expands_to G) basis \ ((\x. f x * g x) expands_to F * G) basis" by (auto simp: expands_to.simps is_expansion_mult elim: eventually_elim2) lemma expands_to_inverse: "trimmed F \ basis_wf basis \ (f expands_to F) basis \ ((\x. inverse (f x)) expands_to inverse F) basis" by (auto simp: expands_to.simps is_expansion_inverse) lemma expands_to_divide: "trimmed G \ basis_wf basis \ (f expands_to F) basis \ (g expands_to G) basis \ ((\x. f x / g x) expands_to F / G) basis" by (auto simp: expands_to.simps is_expansion_divide elim: eventually_elim2) lemma expands_to_powr_0: "eventually (\x. f x = 0) at_top \ g \ g \ basis_wf bs \ length bs = expansion_level TYPE('a) \ ((\x. f x powr g x) expands_to (zero_expansion :: 'a :: multiseries)) bs" by (erule (1) expands_to_cong[OF expands_to_zero]) simp_all lemma expands_to_powr_const: "trimmed_pos F \ basis_wf basis \ (f expands_to F) basis \ p \ p \ ((\x. f x powr p) expands_to powr_expansion abort F p) basis" by (auto simp: expands_to.simps is_expansion_powr trimmed_pos_def elim: eventually_mono) lemma expands_to_powr_const_0: "eventually (\x. f x = 0) at_top \ basis_wf bs \ length bs = expansion_level TYPE('a :: multiseries) \ p \ p \ ((\x. f x powr p) expands_to (zero_expansion :: 'a)) bs" by (rule expands_to_cong[OF expands_to_zero]) auto lemma expands_to_powr: assumes "trimmed_pos F" "basis_wf basis'" "(f expands_to F) basis'" assumes "((\x. exp (ln (f x) * g x)) expands_to E) basis" shows "((\x. f x powr g x) expands_to E) basis" using assms(4) proof (rule expands_to_cong) from eval_pos_if_dominant_term_pos[of basis' F] assms(1-4) show "eventually (\x. exp (ln (f x) * g x) = f x powr g x) at_top" by (auto simp: expands_to.simps trimmed_pos_def powr_def elim: eventually_elim2) qed lemma expands_to_ln_powr: assumes "trimmed_pos F" "basis_wf basis'" "(f expands_to F) basis'" assumes "((\x. ln (f x) * g x) expands_to E) basis" shows "((\x. ln (f x powr g x)) expands_to E) basis" using assms(4) proof (rule expands_to_cong) from eval_pos_if_dominant_term_pos[of basis' F] assms(1-4) show "eventually (\x. ln (f x) * g x = ln (f x powr g x)) at_top" by (auto simp: expands_to.simps trimmed_pos_def powr_def elim: eventually_elim2) qed lemma expands_to_exp_ln: assumes "trimmed_pos F" "basis_wf basis" "(f expands_to F) basis" shows "((\x. exp (ln (f x))) expands_to F) basis" using assms(3) proof (rule expands_to_cong) from eval_pos_if_dominant_term_pos[of basis F] assms show "eventually (\x. f x = exp (ln (f x))) at_top" by (auto simp: expands_to.simps trimmed_pos_def powr_def elim: eventually_elim2) qed lemma expands_to_power: assumes "trimmed F" "basis_wf basis" "(f expands_to F) basis" shows "((\x. f x ^ n) expands_to power_expansion abort F n) basis" using assms by (auto simp: expands_to.simps is_expansion_power elim: eventually_mono) lemma expands_to_power_0: assumes "eventually (\x. f x = 0) at_top" "basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)" "n \ n" shows "((\x. f x ^ n) expands_to (const_expansion (0 ^ n) :: 'a)) basis" by (rule expands_to_cong[OF expands_to_const]) (insert assms, auto elim: eventually_mono) lemma expands_to_0th_root: assumes "n = 0" "basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)" "f \ f" shows "((\x. root n (f x)) expands_to (zero_expansion :: 'a)) basis" by (rule expands_to_cong[OF expands_to_zero]) (insert assms, auto) lemma expands_to_root_0: assumes "n > 0" "eventually (\x. f x = 0) at_top" "basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)" shows "((\x. root n (f x)) expands_to (zero_expansion :: 'a)) basis" by (rule expands_to_cong[OF expands_to_zero]) (insert assms, auto elim: eventually_mono) lemma expands_to_root: assumes "n > 0" "trimmed_pos F" "basis_wf basis" "(f expands_to F) basis" shows "((\x. root n (f x)) expands_to powr_expansion False F (inverse (real n))) basis" proof - have "((\x. f x powr (inverse (real n))) expands_to powr_expansion False F (inverse (real n))) basis" using assms(2-) by (rule expands_to_powr_const) moreover have "eventually (\x. f x powr (inverse (real n)) = root n (f x)) at_top" using eval_pos_if_dominant_term_pos[of basis F] assms by (auto simp: trimmed_pos_def expands_to.simps root_powr_inverse field_simps elim: eventually_elim2) ultimately show ?thesis by (rule expands_to_cong) qed lemma expands_to_root_neg: assumes "n > 0" "trimmed_neg F" "basis_wf basis" "(f expands_to F) basis" shows "((\x. root n (f x)) expands_to -powr_expansion False (-F) (inverse (real n))) basis" proof (rule expands_to_cong) show "((\x. -root n (-f x)) expands_to -powr_expansion False (-F) (inverse (real n))) basis" using assms by (intro expands_to_uminus expands_to_root trimmed_pos_uminus) auto qed (simp_all add: real_root_minus) lemma expands_to_sqrt: assumes "trimmed_pos F" "basis_wf basis" "(f expands_to F) basis" shows "((\x. sqrt (f x)) expands_to powr_expansion False F (1/2)) basis" using expands_to_root[of 2 F basis f] assms by (simp add: sqrt_def) lemma expands_to_exp_real: "(f expands_to (F :: real)) basis \ ((\x. exp (f x)) expands_to exp F) basis" by (auto simp: expands_to.simps is_expansion_real.simps elim: eventually_mono) lemma expands_to_exp_MSLNil: assumes "(f expands_to (MS MSLNil f' :: 'a :: multiseries ms)) bs" "basis_wf bs" shows "((\x. exp (f x)) expands_to (const_expansion 1 :: 'a ms)) bs" proof - from assms have "eventually (\x. f' x = 0) at_top" "eventually (\x. f' x = f x) at_top" and [simp]: "length bs = Suc (expansion_level(TYPE('a)))" by (auto simp: expands_to.simps elim: is_expansion_aux.cases) from this(1,2) have "eventually (\x. 1 = exp (f x)) at_top" by eventually_elim simp thus ?thesis by (auto simp: expands_to.simps intro!: is_expansion_const assms) qed lemma expands_to_exp_zero: "(g expands_to MS xs f) basis \ eventually (\x. f x = 0) at_top \ basis_wf basis \ length basis = expansion_level TYPE('a :: multiseries) \ ((\x. exp (g x)) expands_to (const_expansion 1 :: 'a)) basis" by (auto simp: expands_to.simps intro!: is_expansion_const elim: eventually_elim2) lemma expands_to_sin_real: "(f expands_to (F :: real)) basis \ ((\x. sin (f x)) expands_to sin F) basis" by (auto simp: expands_to.simps is_expansion_real.simps elim: eventually_mono) lemma expands_to_cos_real: "(f expands_to (F :: real)) basis \ ((\x. cos (f x)) expands_to cos F) basis" by (auto simp: expands_to.simps is_expansion_real.simps elim: eventually_mono) lemma expands_to_sin_MSLNil: "(f expands_to MS (MSLNil:: ('a \ real) msllist) g) basis \ basis_wf basis \ ((\x. sin (f x)) expands_to MS (MSLNil :: ('a :: multiseries \ real) msllist) (\x. 0)) basis" by (auto simp: expands_to.simps dominant_term_ms_aux_def intro!: is_expansion_aux.intros elim: eventually_elim2 is_expansion_aux.cases) lemma expands_to_cos_MSLNil: "(f expands_to MS (MSLNil:: ('a :: multiseries \ real) msllist) g) basis \ basis_wf basis \ ((\x. cos (f x)) expands_to (const_expansion 1 :: 'a ms)) basis" by (auto simp: expands_to.simps dominant_term_ms_aux_def const_expansion_ms_def intro!: const_ms_aux elim: is_expansion_aux.cases eventually_elim2) lemma sin_ms_aux': assumes "ms_exp_gt 0 (ms_aux_hd_exp xs)" "basis_wf basis" "is_expansion_aux xs f basis" shows "is_expansion_aux (sin_ms_aux' xs) (\x. sin (f x)) basis" unfolding sin_ms_aux'_def sin_conv_pre_sin power2_eq_square using assms by (intro times_ms_aux powser_ms_aux[unfolded o_def] convergent_powser_sin_series_stream) (auto simp: hd_exp_times add_neg_neg split: option.splits) lemma cos_ms_aux': assumes "ms_exp_gt 0 (ms_aux_hd_exp xs)" "basis_wf basis" "is_expansion_aux xs f basis" shows "is_expansion_aux (cos_ms_aux' xs) (\x. cos (f x)) basis" unfolding cos_ms_aux'_def cos_conv_pre_cos power2_eq_square using assms by (intro times_ms_aux powser_ms_aux[unfolded o_def] convergent_powser_cos_series_stream) (auto simp: hd_exp_times add_neg_neg split: option.splits) lemma expands_to_sin_ms_neg_exp: assumes "e < 0" "basis_wf basis" "(f expands_to MS (MSLCons (C, e) xs) g) basis" shows "((\x. sin (f x)) expands_to MS (sin_ms_aux' (MSLCons (C, e) xs)) (\x. sin (g x))) basis" using assms by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def intro!: sin_ms_aux' elim: eventually_mono) lemma expands_to_cos_ms_neg_exp: assumes "e < 0" "basis_wf basis" "(f expands_to MS (MSLCons (C, e) xs) g) basis" shows "((\x. cos (f x)) expands_to MS (cos_ms_aux' (MSLCons (C, e) xs)) (\x. cos (g x))) basis" using assms by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def intro!: cos_ms_aux' elim: eventually_mono) lemma is_expansion_sin_ms_zero_exp: fixes F :: "('a :: multiseries \ real) msllist" assumes basis: "basis_wf (b # basis)" assumes F: "is_expansion (MS (MSLCons (C, 0) xs) f) (b # basis)" assumes Sin: "((\x. sin (eval C x)) expands_to (Sin :: 'a)) basis" assumes Cos: "((\x. cos (eval C x)) expands_to (Cos :: 'a)) basis" shows "is_expansion (MS (plus_ms_aux (scale_shift_ms_aux (Sin, 0) (cos_ms_aux' xs)) (scale_shift_ms_aux (Cos, 0) (sin_ms_aux' xs))) (\x. sin (f x))) (b # basis)" (is "is_expansion (MS ?A _) _") proof - let ?g = "\x. f x - eval C x * b x powr 0" let ?h = "\x. eval Sin x * b x powr 0 * cos (?g x) + eval Cos x * b x powr 0 * sin (?g x)" from basis have "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) from F have F': "is_expansion_aux xs ?g (b # basis)" "ms_exp_gt 0 (ms_aux_hd_exp xs)" "is_expansion C basis" by (auto elim!: is_expansion_aux_MSLCons) have "is_expansion_aux ?A ?h (b # basis)" using basis Sin Cos unfolding F'(1) by (intro plus_ms_aux scale_shift_ms_aux cos_ms_aux' sin_ms_aux' F') (simp_all add: expands_to.simps) moreover from b_pos eval_expands_to[OF Sin] eval_expands_to[OF Cos] have "eventually (\x. ?h x = sin (f x)) at_top" by eventually_elim (simp add: sin_add [symmetric]) ultimately have "is_expansion_aux ?A (\x. sin (f x)) (b # basis)" by (rule is_expansion_aux_cong) thus ?thesis by simp qed lemma is_expansion_cos_ms_zero_exp: fixes F :: "('a :: multiseries \ real) msllist" assumes basis: "basis_wf (b # basis)" assumes F: "is_expansion (MS (MSLCons (C, 0) xs) f) (b # basis)" assumes Sin: "((\x. sin (eval C x)) expands_to (Sin :: 'a)) basis" assumes Cos: "((\x. cos (eval C x)) expands_to (Cos :: 'a)) basis" shows "is_expansion (MS (minus_ms_aux (scale_shift_ms_aux (Cos, 0) (cos_ms_aux' xs)) (scale_shift_ms_aux (Sin, 0) (sin_ms_aux' xs))) (\x. cos (f x))) (b # basis)" (is "is_expansion (MS ?A _) _") proof - let ?g = "\x. f x - eval C x * b x powr 0" let ?h = "\x. eval Cos x * b x powr 0 * cos (?g x) - eval Sin x * b x powr 0 * sin (?g x)" from basis have "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) from F have F': "is_expansion_aux xs ?g (b # basis)" "ms_exp_gt 0 (ms_aux_hd_exp xs)" "is_expansion C basis" by (auto elim!: is_expansion_aux_MSLCons) have "is_expansion_aux ?A ?h (b # basis)" using basis Sin Cos unfolding F'(1) by (intro minus_ms_aux scale_shift_ms_aux cos_ms_aux' sin_ms_aux' F') (simp_all add: expands_to.simps) moreover from b_pos eval_expands_to[OF Sin] eval_expands_to[OF Cos] have "eventually (\x. ?h x = cos (f x)) at_top" by eventually_elim (simp add: cos_add [symmetric]) ultimately have "is_expansion_aux ?A (\x. cos (f x)) (b # basis)" by (rule is_expansion_aux_cong) thus ?thesis by simp qed lemma expands_to_sin_ms_zero_exp: assumes e: "e = 0" and basis: "basis_wf (b # basis)" assumes F: "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)" assumes Sin: "((\x. sin (c x)) expands_to Sin) basis" assumes Cos: "((\x. cos (c x)) expands_to Cos) basis" assumes C: "eval C = c" shows "((\x. sin (f x)) expands_to MS (plus_ms_aux (scale_shift_ms_aux (Sin, 0) (cos_ms_aux' xs)) (scale_shift_ms_aux (Cos, 0) (sin_ms_aux' xs))) (\x. sin (g x))) (b # basis)" using is_expansion_sin_ms_zero_exp[of b basis C xs g Sin Cos] assms by (auto simp: expands_to.simps elim: eventually_mono) lemma expands_to_cos_ms_zero_exp: assumes e: "e = 0" and basis: "basis_wf (b # basis)" assumes F: "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)" assumes Sin: "((\x. sin (c x)) expands_to Sin) basis" assumes Cos: "((\x. cos (c x)) expands_to Cos) basis" assumes C: "eval C = c" shows "((\x. cos (f x)) expands_to MS (minus_ms_aux (scale_shift_ms_aux (Cos, 0) (cos_ms_aux' xs)) (scale_shift_ms_aux (Sin, 0) (sin_ms_aux' xs))) (\x. cos (g x))) (b # basis)" using is_expansion_cos_ms_zero_exp[of b basis C xs g Sin Cos] assms by (auto simp: expands_to.simps elim: eventually_mono) lemma expands_to_sgn_zero: assumes "eventually (\x. f x = 0) at_top" "basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)" shows "((\x. sgn (f x)) expands_to (zero_expansion :: 'a)) basis" by (rule expands_to_cong[OF expands_to_zero]) (insert assms, auto simp: sgn_eq_0_iff) lemma expands_to_sgn_pos: assumes "trimmed_pos F" "(f expands_to F) basis" "basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)" shows "((\x. sgn (f x)) expands_to (const_expansion 1 :: 'a)) basis" proof (rule expands_to_cong[OF expands_to_const]) from trimmed_imp_eventually_sgn[of basis F] assms have "eventually (\x. sgn (eval F x) = 1) at_top" by (simp add: expands_to.simps trimmed_pos_def) moreover from assms have "eventually (\x. eval F x = f x) at_top" by (simp add: expands_to.simps) ultimately show "eventually (\x. 1 = sgn (f x)) at_top" by eventually_elim simp qed (insert assms, auto) lemma expands_to_sgn_neg: assumes "trimmed_neg F" "(f expands_to F) basis" "basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)" shows "((\x. sgn (f x)) expands_to (const_expansion (-1) :: 'a)) basis" proof (rule expands_to_cong[OF expands_to_const]) from trimmed_imp_eventually_sgn[of basis F] assms have "eventually (\x. sgn (eval F x) = -1) at_top" by (simp add: expands_to.simps trimmed_neg_def) moreover from assms have "eventually (\x. eval F x = f x) at_top" by (simp add: expands_to.simps) ultimately show "eventually (\x. -1 = sgn (f x)) at_top" by eventually_elim simp qed (insert assms, auto) lemma expands_to_abs_pos: assumes "trimmed_pos F" "basis_wf basis" "(f expands_to F) basis" shows "((\x. abs (f x)) expands_to F) basis" using assms(3) proof (rule expands_to_cong) from trimmed_imp_eventually_sgn[of basis F] assms have "eventually (\x. sgn (eval F x) = 1) at_top" by (simp add: expands_to.simps trimmed_pos_def) with assms(3) show "eventually (\x. f x = abs (f x)) at_top" by (auto simp: expands_to.simps sgn_if split: if_splits elim: eventually_elim2) qed lemma expands_to_abs_neg: assumes "trimmed_neg F" "basis_wf basis" "(f expands_to F) basis" shows "((\x. abs (f x)) expands_to -F) basis" using expands_to_uminus[OF assms(2,3)] proof (rule expands_to_cong) from trimmed_imp_eventually_sgn[of basis F] assms have "eventually (\x. sgn (eval F x) = -1) at_top" by (simp add: expands_to.simps trimmed_neg_def) with assms(3) show "eventually (\x. -f x = abs (f x)) at_top" by (auto simp: expands_to.simps sgn_if split: if_splits elim: eventually_elim2) qed lemma expands_to_imp_eventually_nz: assumes "basis_wf basis" "(f expands_to F) basis" "trimmed F" shows "eventually (\x. f x \ 0) at_top" using trimmed_imp_eventually_nz[OF assms(1), of F] assms(2,3) by (auto simp: expands_to.simps elim: eventually_elim2) lemma expands_to_imp_eventually_pos: assumes "basis_wf basis" "(f expands_to F) basis" "trimmed_pos F" shows "eventually (\x. f x > 0) at_top" using eval_pos_if_dominant_term_pos[of basis F] assms by (auto simp: expands_to.simps trimmed_pos_def elim: eventually_elim2) lemma expands_to_imp_eventually_neg: assumes "basis_wf basis" "(f expands_to F) basis" "trimmed_neg F" shows "eventually (\x. f x < 0) at_top" using eval_pos_if_dominant_term_pos[of basis "-F"] assms by (auto simp: expands_to.simps trimmed_neg_def is_expansion_uminus elim: eventually_elim2) lemma expands_to_imp_eventually_gt: assumes "basis_wf basis" "((\x. f x - g x) expands_to F) basis" "trimmed_pos F" shows "eventually (\x. f x > g x) at_top" using expands_to_imp_eventually_pos[OF assms] by simp lemma expands_to_imp_eventually_lt: assumes "basis_wf basis" "((\x. f x - g x) expands_to F) basis" "trimmed_neg F" shows "eventually (\x. f x < g x) at_top" using expands_to_imp_eventually_neg[OF assms] by simp lemma eventually_diff_zero_imp_eq: fixes f g :: "real \ real" assumes "eventually (\x. f x - g x = 0) at_top" shows "eventually (\x. f x = g x) at_top" using assms by eventually_elim simp lemma smallo_trimmed_imp_eventually_less: fixes f g :: "real \ real" assumes "f \ o(g)" "(g expands_to G) bs" "basis_wf bs" "trimmed_pos G" shows "eventually (\x. f x < g x) at_top" proof - from assms(2-4) have pos: "eventually (\x. g x > 0) at_top" using expands_to_imp_eventually_pos by blast have "(1 / 2 :: real) > 0" by simp from landau_o.smallD[OF assms(1) this] and pos show ?thesis by eventually_elim (simp add: abs_if split: if_splits) qed lemma smallo_trimmed_imp_eventually_greater: fixes f g :: "real \ real" assumes "f \ o(g)" "(g expands_to G) bs" "basis_wf bs" "trimmed_neg G" shows "eventually (\x. f x > g x) at_top" proof - from assms(2-4) have pos: "eventually (\x. g x < 0) at_top" using expands_to_imp_eventually_neg by blast have "(1 / 2 :: real) > 0" by simp from landau_o.smallD[OF assms(1) this] and pos show ?thesis by eventually_elim (simp add: abs_if split: if_splits) qed lemma expands_to_min_lt: assumes "(f expands_to F) basis" "eventually (\x. f x < g x) at_top" shows "((\x. min (f x) (g x)) expands_to F) basis" using assms(1) by (rule expands_to_cong) (insert assms(2), auto simp: min_def elim: eventually_mono) lemma expands_to_min_eq: assumes "(f expands_to F) basis" "eventually (\x. f x = g x) at_top" shows "((\x. min (f x) (g x)) expands_to F) basis" using assms(1) by (rule expands_to_cong) (insert assms(2), auto simp: min_def elim: eventually_mono) lemma expands_to_min_gt: assumes "(g expands_to G) basis" "eventually (\x. f x > g x) at_top" shows "((\x. min (f x) (g x)) expands_to G) basis" using assms(1) by (rule expands_to_cong) (insert assms(2), auto simp: min_def elim: eventually_mono) lemma expands_to_max_gt: assumes "(f expands_to F) basis" "eventually (\x. f x > g x) at_top" shows "((\x. max (f x) (g x)) expands_to F) basis" using assms(1) by (rule expands_to_cong) (insert assms(2), auto simp: max_def elim: eventually_mono) lemma expands_to_max_eq: assumes "(f expands_to F) basis" "eventually (\x. f x = g x) at_top" shows "((\x. max (f x) (g x)) expands_to F) basis" using assms(1) by (rule expands_to_cong) (insert assms(2), auto simp: max_def elim: eventually_mono) lemma expands_to_max_lt: assumes "(g expands_to G) basis" "eventually (\x. f x < g x) at_top" shows "((\x. max (f x) (g x)) expands_to G) basis" using assms(1) by (rule expands_to_cong) (insert assms(2), auto simp: max_def elim: eventually_mono) lemma expands_to_lift: "is_lifting f basis basis' \ (c expands_to C) basis \ (c expands_to (f C)) basis'" by (simp add: is_lifting_def expands_to.simps) lemma expands_to_basic_powr: assumes "basis_wf (b # basis)" "length basis = expansion_level TYPE('a::multiseries)" shows "((\x. b x powr e) expands_to MS (MSLCons (const_expansion 1 :: 'a, e) MSLNil) (\x. b x powr e)) (b # basis)" using assms by (auto simp: expands_to.simps basis_wf_Cons powr_smallo_iff intro!: is_expansion_aux.intros is_expansion_const powr_smallo_iff) lemma expands_to_basic_inverse: assumes "basis_wf (b # basis)" "length basis = expansion_level TYPE('a::multiseries)" shows "((\x. inverse (b x)) expands_to MS (MSLCons (const_expansion 1 :: 'a, -1) MSLNil) (\x. b x powr -1)) (b # basis)" proof - from assms have "eventually (\x. b x > 0) at_top" by (simp add: basis_wf_Cons filterlim_at_top_dense) moreover have "(\a. inverse (a powr 1)) \ o(\a. a powr e')" if "e' > -1" for e' :: real using that by (simp add: landau_o.small.inverse_eq2 powr_add [symmetric] one_smallo_powr) ultimately show ?thesis using assms by (auto simp: expands_to.simps basis_wf_Cons powr_minus elim: eventually_mono intro!: is_expansion_aux.intros is_expansion_const landau_o.small.compose[of _ at_top _ b]) qed lemma expands_to_div': assumes "basis_wf basis" "(f expands_to F) basis" "((\x. inverse (g x)) expands_to G) basis" shows "((\x. f x / g x) expands_to F * G) basis" using expands_to_mult[OF assms] by (simp add: field_split_simps) lemma expands_to_basic: assumes "basis_wf (b # basis)" "length basis = expansion_level TYPE('a::multiseries)" shows "(b expands_to MS (MSLCons (const_expansion 1 :: 'a, 1) MSLNil) b) (b # basis)" proof - from assms have "eventually (\x. b x > 0) at_top" by (simp add: basis_wf_Cons filterlim_at_top_dense) moreover { fix e' :: real assume e': "e' > 1" have "(\x::real. x) \ \(\x. x powr 1)" by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto also have "(\x. x powr 1) \ o(\x. x powr e')" using e' by (subst powr_smallo_iff) (auto simp: filterlim_ident) finally have "(\x. x) \ o(\x. x powr e')" . } ultimately show ?thesis using assms by (auto simp: expands_to.simps basis_wf_Cons elim: eventually_mono intro!: is_expansion_aux.intros is_expansion_const landau_o.small.compose[of _ at_top _ b]) qed lemma expands_to_lift': assumes "basis_wf (b # basis)" "(f expands_to MS xs g) basis" shows "(f expands_to (MS (MSLCons (MS xs g, 0) MSLNil) g)) (b # basis)" proof - have "is_lifting lift_ms basis (b # basis)" by (rule is_lifting_lift) fact+ from expands_to_lift[OF this assms(2)] show ?thesis by (simp add: lift_ms_def) qed lemma expands_to_lift'': assumes "basis_wf (b # basis)" "(f expands_to F) basis" shows "(f expands_to (MS (MSLCons (F, 0) MSLNil) (eval F))) (b # basis)" proof - have "is_lifting lift_ms basis (b # basis)" by (rule is_lifting_lift) fact+ from expands_to_lift[OF this assms(2)] show ?thesis by (simp add: lift_ms_def) qed lemma expands_to_drop_zero: assumes "eventually (\x. eval C x = 0) at_top" assumes "(f expands_to (MS (MSLCons (C, e) xs) g)) basis" shows "(f expands_to (MS xs g)) basis" using assms drop_zero_ms[of C e xs g basis] by (simp add: expands_to.simps) lemma expands_to_hd'': assumes "(f expands_to (MS (MSLCons (C, e) xs) g)) (b # basis)" shows "(eval C expands_to C) basis" using assms by (auto simp: expands_to.simps elim: is_expansion_aux_MSLCons) lemma expands_to_hd: assumes "(f expands_to (MS (MSLCons (MS ys h, e) xs) g)) (b # basis)" shows "(h expands_to MS ys h) basis" using assms by (auto simp: expands_to.simps elim: is_expansion_aux_MSLCons) lemma expands_to_hd': assumes "(f expands_to (MS (MSLCons (c :: real, e) xs) g)) (b # basis)" shows "((\_. c) expands_to c) basis" using assms by (auto simp: expands_to.simps elim: is_expansion_aux_MSLCons) lemma expands_to_trim_cong: assumes "(f expands_to (MS (MSLCons (C, e) xs) g)) (b # basis)" assumes "(eval C expands_to C') basis" shows "(f expands_to (MS (MSLCons (C', e) xs) g)) (b # basis)" proof - from assms(1) have "is_expansion_aux xs (\x. g x - eval C x * b x powr e) (b # basis)" by (auto simp: expands_to.simps elim!: is_expansion_aux_MSLCons) hence "is_expansion_aux xs (\x. g x - eval C' x * b x powr e) (b # basis)" by (rule is_expansion_aux_cong) (insert assms(2), auto simp: expands_to.simps elim: eventually_mono) with assms show ?thesis by (auto simp: expands_to.simps elim!: is_expansion_aux_MSLCons intro!: is_expansion_aux.intros) qed lemma is_expansion_aux_expands_to: assumes "(f expands_to MS xs g) basis" shows "is_expansion_aux xs f basis" proof - from assms have "is_expansion_aux xs g basis" "eventually (\x. g x = f x) at_top" by (simp_all add: expands_to.simps) thus ?thesis by (rule is_expansion_aux_cong) qed lemma ln_series_stream_aux_code: "ln_series_stream_aux True c = MSSCons (- inverse c) (ln_series_stream_aux False (c + 1))" "ln_series_stream_aux False c = MSSCons (inverse c) (ln_series_stream_aux True (c + 1))" by (subst ln_series_stream_aux.code, simp)+ definition powser_ms_aux' where "powser_ms_aux' cs xs = powser_ms_aux (msllist_of_msstream cs) xs" lemma ln_ms_aux: fixes C L :: "'a :: multiseries" assumes trimmed: "trimmed_pos C" and basis: "basis_wf (b # basis)" assumes F: "is_expansion_aux (MSLCons (C, e) xs) f (b # basis)" assumes L: "((\x. ln (eval C x) + e * ln (b x)) expands_to L) basis" shows "is_expansion_aux (MSLCons (L, 0) (times_ms_aux (scale_shift_ms_aux (inverse C, - e) xs) (powser_ms_aux' (ln_series_stream_aux False 1) (scale_shift_ms_aux (inverse C, - e) xs)))) (\x. ln (f x)) (b # basis)" (is "is_expansion_aux ?G _ _") proof - from F have F': "is_expansion_aux xs (\x. f x - eval C x * b x powr e) (b # basis)" "is_expansion C basis" "\e'>e. f \ o(\x. b x powr e')" "ms_exp_gt e (ms_aux_hd_exp xs)" by (auto elim!: is_expansion_aux_MSLCons) from basis have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 1) at_top" by (simp add: filterlim_at_top_dense) from eval_pos_if_dominant_term_pos[of basis C] trimmed F' basis have C_pos: "eventually (\x. eval C x > 0) at_top" by (auto simp: basis_wf_Cons trimmed_pos_def) from eval_pos_if_dominant_term_pos'[OF basis _ F] trimmed have f_pos: "eventually (\x. f x > 0) at_top" by (simp add: trimmed_pos_def trimmed_ms_aux_def dominant_term_ms_aux_def case_prod_unfold) from F' have "ms_exp_gt 0 (map_option ((+) (-e)) (ms_aux_hd_exp xs))" by (cases "ms_aux_hd_exp xs") simp_all hence "is_expansion_aux (powser_ms_aux' ln_series_stream (scale_shift_ms_aux (inverse C, -e) xs)) ((\x. ln (1 + x)) \ (\x. eval (inverse C) x * b x powr -e * (f x - eval C x * b x powr e))) (b # basis)" (is "is_expansion_aux _ ?g _") unfolding powser_ms_aux'_def using powser_ms_aux' basis F' trimmed by (intro powser_ms_aux' convergent_powser'_ln scale_shift_ms_aux is_expansion_inverse F') (simp_all add: trimmed_pos_def hd_exp_basis basis_wf_Cons) moreover have "eventually (\x. ?g x = ln (f x) - eval L x * b x powr 0) at_top" using b_pos C_pos f_pos eval_expands_to[OF L] by eventually_elim (simp add: powr_minus algebra_simps ln_mult ln_inverse expands_to.simps ln_powr) ultimately have "is_expansion_aux (powser_ms_aux' ln_series_stream (scale_shift_ms_aux (inverse C, -e) xs)) (\x. ln (f x) - eval L x * b x powr 0) (b # basis)" by (rule is_expansion_aux_cong) hence *: "is_expansion_aux (times_ms_aux (scale_shift_ms_aux (inverse C, - e) xs) (powser_ms_aux' (ln_series_stream_aux False 1) (scale_shift_ms_aux (inverse C, - e) xs))) (\x. ln (f x) - eval L x * b x powr 0) (b # basis)" unfolding ln_series_stream_def powser_ms_aux'_def powser_ms_aux_MSLCons msllist_of_msstream_MSSCons by (rule drop_zero_ms_aux [rotated]) simp_all moreover from F' have exp: "ms_exp_gt 0 (map_option ((+) (-e)) (ms_aux_hd_exp xs))" by (cases "ms_aux_hd_exp xs") simp_all moreover have "(\x. ln (f x)) \ o(\x. b x powr e')" if "e' > 0" for e' proof - from is_expansion_aux_imp_smallo[OF *, of e'] that exp have "(\x. ln (f x) - eval L x * b x powr 0) \ o(\x. b x powr e')" by (cases "ms_aux_hd_exp xs") (simp_all add: hd_exp_times hd_exp_powser hd_exp_basis powser_ms_aux'_def) moreover { from L F' basis that have "eval L \ o(\x. b x powr e')" by (intro is_expansion_imp_smallo[of basis]) (simp_all add: basis_wf_Cons expands_to.simps) also have "eval L \ o(\x. b x powr e') \ (\x. eval L x * b x powr 0) \ o(\x. b x powr e')" using b_pos by (intro landau_o.small.in_cong) (auto elim: eventually_mono) finally have \ . } ultimately have "(\x. ln (f x) - eval L x * b x powr 0 + eval L x * b x powr 0) \ o(\x. b x powr e')" by (rule sum_in_smallo) thus ?thesis by simp qed ultimately show "is_expansion_aux ?G (\x. ln (f x)) (b # basis)" using L by (intro is_expansion_aux.intros) (auto simp: expands_to.simps hd_exp_times hd_exp_powser hd_exp_basis powser_ms_aux'_def split: option.splits) qed lemma expands_to_ln: fixes C L :: "'a :: multiseries" assumes trimmed: "trimmed_pos (MS (MSLCons (C, e) xs) g)" and basis: "basis_wf (b # basis)" assumes F: "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)" assumes L: "((\x. ln (h x) + e * ln (b x)) expands_to L) basis" and h: "eval C = h" shows "((\x. ln (f x)) expands_to MS (MSLCons (L, 0) (times_ms_aux (scale_shift_ms_aux (inverse C, - e) xs) (powser_ms_aux' (ln_series_stream_aux False 1) (scale_shift_ms_aux (inverse C, - e) xs)))) (\x. ln (f x))) (b # basis)" using is_expansion_aux_expands_to[OF F] assms by (auto simp: expands_to.simps intro!: ln_ms_aux) lemma trimmed_lifting: assumes "is_lifting f basis basis'" assumes "trimmed F" shows "trimmed (f F)" using assms by (simp add: is_lifting_def) lemma trimmed_pos_lifting: assumes "is_lifting f basis basis'" assumes "trimmed_pos F" shows "trimmed_pos (f F)" using assms by (simp add: is_lifting_def trimmed_pos_def) lemma expands_to_ln_aux_0: assumes e: "e = 0" assumes L1: "((\x. ln (g x)) expands_to L) basis" shows "((\x. ln (g x) + e * ln (b x)) expands_to L) basis" using assms by simp lemma expands_to_ln_aux_1: assumes e: "e = 1" assumes basis: "basis_wf (b # basis)" assumes L1: "((\x. ln (g x)) expands_to L1) basis" assumes L2: "((\x. ln (b x)) expands_to L2) basis" shows "((\x. ln (g x) + e * ln (b x)) expands_to L1 + L2) basis" using assms by (intro expands_to_add) (simp_all add: basis_wf_Cons) lemma expands_to_ln_eventually_1: assumes "basis_wf basis" "length basis = expansion_level TYPE('a)" assumes "eventually (\x. f x - 1 = 0) at_top" shows "((\x. ln (f x)) expands_to (zero_expansion :: 'a :: multiseries)) basis" by (rule expands_to_cong [OF expands_to_zero]) (insert assms, auto elim: eventually_mono) lemma expands_to_ln_aux: assumes basis: "basis_wf (b # basis)" assumes L1: "((\x. ln (g x)) expands_to L1) basis" assumes L2: "((\x. ln (b x)) expands_to L2) basis" shows "((\x. ln (g x) + e * ln (b x)) expands_to L1 + scale_ms e L2) basis" proof - from L1 have "length basis = expansion_level TYPE('a)" by (auto simp: expands_to.simps is_expansion_length) with assms show ?thesis unfolding scale_ms_def by (intro expands_to_add assms expands_to_mult expands_to_const) (simp_all add: basis_wf_Cons) qed lemma expands_to_ln_to_expands_to_ln_eval: assumes "((\x. ln (f x) + F x) expands_to L) basis" shows "((\x. ln (eval (MS xs f) x) + F x) expands_to L) basis" using assms by simp lemma expands_to_ln_const: "((\x. ln (eval (C :: real) x)) expands_to ln C) []" by (simp add: expands_to.simps is_expansion_real.intros) lemma expands_to_meta_eq_cong: assumes "(f expands_to F) basis" "F \ G" shows "(f expands_to G) basis" using assms by simp lemma expands_to_meta_eq_cong': assumes "(g expands_to F) basis" "f \ g" shows "(f expands_to F) basis" using assms by simp lemma uminus_ms_aux_MSLCons_eval: "uminus_ms_aux (MSLCons (c, e) xs) = MSLCons (-c, e) (uminus_ms_aux xs)" by (simp add: uminus_ms_aux_MSLCons) lemma scale_shift_ms_aux_MSLCons_eval: "scale_shift_ms_aux (c, e) (MSLCons (c', e') xs) = MSLCons (c * c', e + e') (scale_shift_ms_aux (c, e) xs)" by (simp add: scale_shift_ms_aux_MSLCons) lemma times_ms_aux_MSLCons_eval: "times_ms_aux (MSLCons (c1, e1) xs) (MSLCons (c2, e2) ys) = MSLCons (c1 * c2, e1 + e2) (plus_ms_aux (scale_shift_ms_aux (c1, e1) ys) (times_ms_aux xs (MSLCons (c2, e2) ys)))" by (simp add: times_ms_aux_MSLCons) lemma plus_ms_aux_MSLCons_eval: "plus_ms_aux (MSLCons (c1, e1) xs) (MSLCons (c2, e2) ys) = CMP_BRANCH (COMPARE e1 e2) (MSLCons (c2, e2) (plus_ms_aux (MSLCons (c1, e1) xs) ys)) (MSLCons (c1 + c2, e1) (plus_ms_aux xs ys)) (MSLCons (c1, e1) (plus_ms_aux xs (MSLCons (c2, e2) ys)))" by (simp add: CMP_BRANCH_def COMPARE_def plus_ms_aux_MSLCons) lemma inverse_ms_aux_MSLCons: "inverse_ms_aux (MSLCons (C, e) xs) = scale_shift_ms_aux (inverse C, - e) (powser_ms_aux' (mssalternate 1 (- 1)) (scale_shift_ms_aux (inverse C, - e) xs))" by (simp add: Let_def inverse_ms_aux.simps powser_ms_aux'_def) lemma powr_ms_aux_MSLCons: "powr_ms_aux abort (MSLCons (C, e) xs) p = scale_shift_ms_aux (powr_expansion abort C p, e * p) (powser_ms_aux (gbinomial_series abort p) (scale_shift_ms_aux (inverse C, - e) xs))" by simp lemma power_ms_aux_MSLCons: "power_ms_aux abort (MSLCons (C, e) xs) n = scale_shift_ms_aux (power_expansion abort C n, e * real n) (powser_ms_aux (gbinomial_series abort (real n)) (scale_shift_ms_aux (inverse C, - e) xs))" by simp lemma minus_ms_aux_MSLCons_eval: "minus_ms_aux (MSLCons (c1, e1) xs) (MSLCons (c2, e2) ys) = CMP_BRANCH (COMPARE e1 e2) (MSLCons (-c2, e2) (minus_ms_aux (MSLCons (c1, e1) xs) ys)) (MSLCons (c1 - c2, e1) (minus_ms_aux xs ys)) (MSLCons (c1, e1) (minus_ms_aux xs (MSLCons (c2, e2) ys)))" by (simp add: minus_ms_aux_def plus_ms_aux_MSLCons_eval uminus_ms_aux_MSLCons minus_eq_plus_uminus) lemma minus_ms_altdef: "MS xs f - MS ys g = MS (minus_ms_aux xs ys) (\x. f x - g x)" by (simp add: minus_ms_def minus_ms_aux_def) lemma const_expansion_ms_eval: "const_expansion c = MS (MSLCons (const_expansion c, 0) MSLNil) (\_. c)" by (simp add: const_expansion_ms_def const_ms_aux_def) lemma powser_ms_aux'_MSSCons: "powser_ms_aux' (MSSCons c cs) xs = MSLCons (const_expansion c, 0) (times_ms_aux xs (powser_ms_aux' cs xs))" by (simp add: powser_ms_aux'_def powser_ms_aux_MSLCons) lemma sin_ms_aux'_altdef: "sin_ms_aux' xs = times_ms_aux xs (powser_ms_aux' sin_series_stream (times_ms_aux xs xs))" by (simp add: sin_ms_aux'_def powser_ms_aux'_def) lemma cos_ms_aux'_altdef: "cos_ms_aux' xs = powser_ms_aux' cos_series_stream (times_ms_aux xs xs)" by (simp add: cos_ms_aux'_def powser_ms_aux'_def) lemma sin_series_stream_aux_code: "sin_series_stream_aux True acc m = MSSCons (-inverse acc) (sin_series_stream_aux False (acc * (m + 1) * (m + 2)) (m + 2))" "sin_series_stream_aux False acc m = MSSCons (inverse acc) (sin_series_stream_aux True (acc * (m + 1) * (m + 2)) (m + 2))" by (subst sin_series_stream_aux.code; simp)+ lemma arctan_series_stream_aux_code: "arctan_series_stream_aux True n = MSSCons (-inverse n) (arctan_series_stream_aux False (n + 2))" "arctan_series_stream_aux False n = MSSCons (inverse n) (arctan_series_stream_aux True (n + 2))" by (subst arctan_series_stream_aux.code; simp)+ subsubsection \Composition with an asymptotic power series\ context fixes h :: "real \ real" and F :: "real filter" begin coinductive asymp_powser :: "(real \ real) \ real msstream \ bool" where "asymp_powser f cs \ f' \ O[F](\_. 1) \ eventually (\x. f' x = c + f x * h x) F \ asymp_powser f' (MSSCons c cs)" lemma asymp_powser_coinduct [case_names asymp_powser]: "P x1 x2 \ (\x1 x2. P x1 x2 \ \f cs c. x2 = MSSCons c cs \ asymp_powser f cs \ x1 \ O[F](\_. 1) \ (\\<^sub>F x in F. x1 x = c + f x * h x)) \ asymp_powser x1 x2" using asymp_powser.coinduct[of P x1 x2] by blast lemma asymp_powser_coinduct' [case_names asymp_powser]: "P x1 x2 \ (\x1 x2. P x1 x2 \ \f cs c. x2 = MSSCons c cs \ P f cs \ x1 \ O[F](\_. 1) \ (\\<^sub>F x in F. x1 x = c + f x * h x)) \ asymp_powser x1 x2" using asymp_powser.coinduct[of P x1 x2] by blast lemma asymp_powser_transfer: assumes "asymp_powser f cs" "eventually (\x. f x = f' x) F" shows "asymp_powser f' cs" using assms(1) proof (cases rule: asymp_powser.cases) case (1 f cs' c) have "asymp_powser f' (MSSCons c cs')" by (rule asymp_powser.intros) (insert 1 assms(2), auto elim: eventually_elim2 dest: landau_o.big.in_cong) thus ?thesis by (simp add: 1) qed lemma sum_bigo_imp_asymp_powser: assumes filterlim_h: "filterlim h (at 0) F" assumes "(\n. (\x. f x - (\k O[F](\x. h x ^ n))" shows "asymp_powser f cs" using assms(2) proof (coinduction arbitrary: f cs rule: asymp_powser_coinduct') case (asymp_powser f cs) from filterlim_h have h_nz:"eventually (\x. h x \ 0) F" by (auto simp: filterlim_at) show ?case proof (cases cs) case (MSSCons c cs') define f' where "f' = (\x. (f x - c) / h x)" have "(\x. f' x - (\k O[F](\x. h x ^ n)" for n proof - have "(\x. h x * (f' x - (\i \[F](\x. f x - c - h x * (\ix. f x - c - h x * (\i O[F](\x. h x * h x ^ n)" unfolding sum.lessThan_Suc_shift by (simp add: MSSCons algebra_simps sum_distrib_left sum_distrib_right) finally show ?thesis by (subst (asm) landau_o.big.mult_cancel_left) (insert h_nz, auto) qed moreover from h_nz have "\\<^sub>F x in F. f x = c + f' x * h x" by eventually_elim (simp add: f'_def) ultimately show ?thesis using spec[OF asymp_powser, of 0] by (auto simp: MSSCons intro!: exI[of _ f']) qed qed end lemma asymp_powser_o: assumes "asymp_powser h F f cs" assumes "filterlim g F G" shows "asymp_powser (h \ g) G (f \ g) cs" using assms(1) proof (coinduction arbitrary: f cs rule: asymp_powser_coinduct') case (asymp_powser f cs) thus ?case proof (cases rule: asymp_powser.cases) case (1 f' cs' c) from assms(2) have "filtermap g G \ F" by (simp add: filterlim_def) moreover have "eventually (\x. f x = c + f' x * h x) F" by fact ultimately have "eventually (\x. f x = c + f' x * h x) (filtermap g G)" by (rule filter_leD) hence "eventually (\x. f (g x) = c + f' (g x) * h (g x)) G" by (simp add: eventually_filtermap) moreover from 1 have "f \ g \ O[G](\_. 1)" unfolding o_def by (intro landau_o.big.compose[OF _ assms(2)]) auto ultimately show ?thesis using 1 by force qed qed lemma asymp_powser_compose: assumes "asymp_powser h F f cs" assumes "filterlim g F G" shows "asymp_powser (\x. h (g x)) G (\x. f (g x)) cs" using asymp_powser_o[OF assms] by (simp add: o_def) lemma hd_exp_powser': "ms_aux_hd_exp (powser_ms_aux' cs f) = Some 0" by (simp add: powser_ms_aux'_def hd_exp_powser) lemma powser_ms_aux_asymp_powser: assumes "asymp_powser h at_top f cs" and basis: "basis_wf bs" assumes "is_expansion_aux xs h bs" "ms_exp_gt 0 (ms_aux_hd_exp xs)" shows "is_expansion_aux (powser_ms_aux' cs xs) f bs" using assms(1) proof (coinduction arbitrary: cs f rule: is_expansion_aux_coinduct_upto) show "basis_wf bs" by fact next case (B cs f) thus ?case proof (cases rule: asymp_powser.cases) case (1 f' cs' c) from assms(3) obtain b bs' where [simp]: "bs = b # bs'" by (cases bs) (auto dest: is_expansion_aux_basis_nonempty) from basis have b_at_top: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" by (auto simp: filterlim_at_top_dense) have A: "f \ o(\x. b x powr e)" if "e > 0" for e :: real proof - have "f \ O(\_. 1)" by fact also have "(\_::real. 1) \ o(\x. b x powr e)" by (rule landau_o.small.compose[OF _ b_at_top]) (insert that, simp_all add: one_smallo_powr) finally show ?thesis . qed have "ms_closure (\xsa'' fa'' basis''. \cs. xsa'' = powser_ms_aux' cs xs \ basis'' = b # bs' \ asymp_powser h at_top fa'' cs) (times_ms_aux xs (powser_ms_aux' cs' xs)) (\x. h x * f' x) bs" (is "ms_closure ?Q ?ys _ _") by (rule ms_closure_mult[OF ms_closure_embed' ms_closure_embed]) (insert assms 1, auto intro!: exI[of _ cs']) moreover have "eventually (\x. h x * f' x = f x - c * b x powr 0) at_top" using b_pos and \\\<^sub>F x in at_top. f x = c + f' x * h x\ by eventually_elim simp ultimately have "ms_closure ?Q ?ys (\x. f x - c * b x powr 0) bs" by (rule ms_closure_cong) with 1 assms A show ?thesis using is_expansion_aux_expansion_level[OF assms(3)] by (auto simp: powser_ms_aux'_MSSCons basis_wf_Cons hd_exp_times hd_exp_powser' intro!: is_expansion_const ms_closure_add ms_closure_mult split: option.splits) qed qed subsubsection \Arc tangent\ definition arctan_ms_aux where "arctan_ms_aux xs = times_ms_aux xs (powser_ms_aux' arctan_series_stream (times_ms_aux xs xs))" lemma arctan_ms_aux_0: assumes "ms_exp_gt 0 (ms_aux_hd_exp xs)" "basis_wf basis" "is_expansion_aux xs f basis" shows "is_expansion_aux (arctan_ms_aux xs) (\x. arctan (f x)) basis" proof (rule is_expansion_aux_cong) let ?g = "powser (msllist_of_msstream arctan_series_stream)" show "is_expansion_aux (arctan_ms_aux xs) (\x. f x * (?g \ (\x. f x * f x)) x) basis" unfolding arctan_ms_aux_def powser_ms_aux'_def using assms by (intro times_ms_aux powser_ms_aux powser_ms_aux convergent_powser_arctan) (auto simp: hd_exp_times add_neg_neg split: option.splits) have "f \ o(\x. hd basis x powr 0)" using assms by (intro is_expansion_aux_imp_smallo[OF assms(3)]) auto also have "(\x. hd basis x powr 0) \ O(\_. 1)" by (intro bigoI[of _ 1]) auto finally have "filterlim f (nhds 0) at_top" by (auto dest!: smalloD_tendsto) from order_tendstoD(2)[OF tendsto_norm [OF this], of 1] have "eventually (\x. norm (f x) < 1) at_top" by simp thus "eventually (\x. f x * (?g \ (\x. f x * f x)) x = arctan (f x)) at_top" by eventually_elim (simp add: arctan_conv_pre_arctan power2_eq_square) qed lemma arctan_ms_aux_inf: assumes "\e>0. ms_aux_hd_exp xs = Some e" "fst (dominant_term_ms_aux xs) > 0" "trimmed_ms_aux xs" "basis_wf basis" "is_expansion_aux xs f basis" shows "is_expansion_aux (minus_ms_aux (const_ms_aux (pi / 2)) (arctan_ms_aux (inverse_ms_aux xs))) (\x. arctan (f x)) basis" (is "is_expansion_aux ?xs' _ _") proof (rule is_expansion_aux_cong) from \trimmed_ms_aux xs\ have [simp]: "xs \ MSLNil" by auto thus "is_expansion_aux ?xs' (\x. pi / 2 - arctan (inverse (f x))) basis" using assms by (intro minus_ms_aux arctan_ms_aux_0 inverse_ms_aux const_ms_aux) (auto simp: hd_exp_inverse is_expansion_aux_expansion_level) next from assms(2-5) have "eventually (\x. f x > 0) at_top" by (intro eval_pos_if_dominant_term_pos') simp_all thus "eventually (\x. ((\x. pi/2 - arctan (inverse (f x)))) x = arctan (f x)) at_top" unfolding o_def by eventually_elim (subst arctan_inverse_pos, simp_all) qed lemma arctan_ms_aux_neg_inf: assumes "\e>0. ms_aux_hd_exp xs = Some e" "fst (dominant_term_ms_aux xs) < 0" "trimmed_ms_aux xs" "basis_wf basis" "is_expansion_aux xs f basis" shows "is_expansion_aux (minus_ms_aux (const_ms_aux (-pi / 2)) (arctan_ms_aux (inverse_ms_aux xs))) (\x. arctan (f x)) basis" (is "is_expansion_aux ?xs' _ _") proof (rule is_expansion_aux_cong) from \trimmed_ms_aux xs\ have [simp]: "xs \ MSLNil" by auto thus "is_expansion_aux ?xs' (\x. -pi / 2 - arctan (inverse (f x))) basis" using assms by (intro minus_ms_aux arctan_ms_aux_0 inverse_ms_aux const_ms_aux) (auto simp: hd_exp_inverse is_expansion_aux_expansion_level) next from assms(2-5) have "eventually (\x. f x < 0) at_top" by (intro eval_neg_if_dominant_term_neg') simp_all thus "eventually (\x. ((\x. -pi/2 - arctan (inverse (f x)))) x = arctan (f x)) at_top" unfolding o_def by eventually_elim (subst arctan_inverse_neg, simp_all) qed lemma expands_to_arctan_zero: assumes "((\_::real. 0::real) expands_to C) bs" "eventually (\x. f x = 0) at_top" shows "((\x::real. arctan (f x)) expands_to C) bs" using assms by (auto simp: expands_to.simps elim: eventually_elim2) lemma expands_to_arctan_ms_neg_exp: assumes "e < 0" "basis_wf basis" "(f expands_to MS (MSLCons (C, e) xs) g) basis" shows "((\x. arctan (f x)) expands_to MS (arctan_ms_aux (MSLCons (C, e) xs)) (\x. arctan (g x))) basis" using assms by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def intro!: arctan_ms_aux_0 elim: eventually_mono) lemma expands_to_arctan_ms_pos_exp_pos: assumes "e > 0" "trimmed_pos (MS (MSLCons (C, e) xs) g)" "basis_wf basis" "(f expands_to MS (MSLCons (C, e) xs) g) basis" shows "((\x. arctan (f x)) expands_to MS (minus_ms_aux (const_ms_aux (pi / 2)) (arctan_ms_aux (inverse_ms_aux (MSLCons (C, e) xs)))) (\x. arctan (g x))) basis" using arctan_ms_aux_inf[of "MSLCons (C, e) xs" basis g] assms by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def dominant_term_ms_aux_MSLCons trimmed_pos_def elim: eventually_mono) lemma expands_to_arctan_ms_pos_exp_neg: assumes "e > 0" "trimmed_neg (MS (MSLCons (C, e) xs) g)" "basis_wf basis" "(f expands_to MS (MSLCons (C, e) xs) g) basis" shows "((\x. arctan (f x)) expands_to MS (minus_ms_aux (const_ms_aux (-pi/2)) (arctan_ms_aux (inverse_ms_aux (MSLCons (C, e) xs)))) (\x. arctan (g x))) basis" using arctan_ms_aux_neg_inf[of "MSLCons (C, e) xs" basis g] assms by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def dominant_term_ms_aux_MSLCons trimmed_neg_def elim: eventually_mono) subsubsection \Exponential function\ (* TODO: Move *) lemma ln_smallo_powr: assumes "(e::real) > 0" shows "(\x. ln x) \ o(\x. x powr e)" proof - have *: "ln \ o(\x::real. x)" using ln_x_over_x_tendsto_0 by (intro smalloI_tendsto) auto have "(\x. ln x) \ \(\x::real. ln x powr 1)" by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 1]]) auto also have "(\x::real. ln x powr 1) \ o(\x. x powr e)" by (intro ln_smallo_imp_flat filterlim_ident ln_at_top assms landau_o.small.compose[of _ at_top _ ln] *) finally show ?thesis . qed lemma basis_wf_insert_exp_pos: assumes "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)" "basis_wf (b # basis)" "trimmed (MS (MSLCons (C, e) xs) g)" "e - 0 > 0" shows "(\x. ln (b x)) \ o(f)" proof - from assms have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) from assms have "is_expansion_aux xs (\x. g x - eval C x * b x powr e) (b # basis)" "ms_exp_gt e (ms_aux_hd_exp xs)" by (auto elim!: is_expansion_aux_MSLCons simp: expands_to.simps) from is_expansion_aux_imp_smallo''[OF this] obtain e' where e': "e' < e" "(\x. g x - eval C x * b x powr e) \ o(\x. b x powr e')" unfolding list.sel . define eps where "eps = e/2" have "(\x. ln (b x)) \ o(\x. b x powr (e - eps))" unfolding eps_def by (rule landau_o.small.compose[OF _ b_limit]) (insert e'(1) \e - 0 > 0\, simp add: ln_smallo_powr) also from assms e'(1) have "eval C \ \(\x. b x powr -eps)" unfolding eps_def by (intro is_expansion_imp_smallomega[of basis]) (auto simp: basis_wf_Cons expands_to.simps trimmed_pos_def trimmed_ms_aux_MSLCons elim!: is_expansion_aux_MSLCons) hence "(\x. b x powr -eps * b x powr e) \ o(\x. eval C x * b x powr e)" by (intro landau_o.small_big_mult) (simp_all add: smallomega_iff_smallo) hence "(\x. b x powr (e - eps)) \ o(\x. eval C x * b x powr e)" by (simp add: powr_add [symmetric] field_simps) finally have "(\x. ln (b x)) \ o(\x. eval C x * b x powr e)" . also { from e' have "(\x. g x - eval C x * b x powr e) \ o(\x. b x powr (e' - e) * b x powr e)" by (simp add: powr_add [symmetric]) also from assms e'(1) have "eval C \ \(\x. b x powr (e' - e))" unfolding eps_def by (intro is_expansion_imp_smallomega[of basis]) (auto simp: basis_wf_Cons expands_to.simps trimmed_pos_def trimmed_ms_aux_MSLCons elim!: is_expansion_aux_MSLCons) hence "(\x. b x powr (e' - e) * b x powr e) \ o(\x. eval C x * b x powr e)" by (intro landau_o.small_big_mult is_expansion_imp_smallo) (simp_all add: smallomega_iff_smallo) finally have "o(\x. eval C x * b x powr e) = o(\x. g x - eval C x * b x powr e + eval C x * b x powr e)" by (intro landau_o.small.plus_absorb1 [symmetric]) simp_all } also have "o(\x. g x - eval C x * b x powr e + eval C x * b x powr e) = o(g)" by simp also from assms have "g \ \(f)" by (intro bigthetaI_cong) (simp add: expands_to.simps) finally show "(\x. ln (b x)) \ o(f)" . qed lemma basis_wf_insert_exp_uminus: "(\x. ln (b x)) \ o(f) \ (\x. ln (b x)) \ o(\x. -f x :: real)" by simp lemma basis_wf_insert_exp_uminus': "f \ o(\x. ln (b x)) \ (\x. -f x) \ o(\x. ln (b x))" by simp lemma expands_to_exp_insert_pos: fixes C :: "'a :: multiseries" assumes "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)" "basis_wf (b # basis)" "trimmed_pos (MS (MSLCons (C, e) xs) g)" "(\x. ln (b x)) \ o(f)" shows "filterlim (\x. exp (f x)) at_top at_top" "((\x. ln (exp (f x))) expands_to MS (MSLCons (C, e) xs) g) (b # basis)" "((\x. exp (f x)) expands_to MS (MSLCons (const_expansion 1 :: 'a ms, 1) MSLNil) (\x. exp (g x))) ((\x. exp (f x)) # b # basis)" (is ?thesis3) proof - have "ln \ \(\x::real. 1)" by (intro smallomegaI_filterlim_at_top_norm) (auto intro!: filterlim_compose[OF filterlim_abs_real] ln_at_top) with assms have "(\x. 1) \ o(\x. ln (b x))" by (intro landau_o.small.compose[of _ at_top _ b]) (simp_all add: basis_wf_Cons smallomega_iff_smallo) also note \(\x. ln (b x)) \ o(f)\ finally have f: "f \ \(\_. 1)" by (simp add: smallomega_iff_smallo) hence f: "filterlim (\x. abs (f x)) at_top at_top" by (auto dest: smallomegaD_filterlim_at_top_norm) define c where "c = fst (dominant_term C)" define es where "es = snd (dominant_term C)" from trimmed_imp_eventually_sgn[OF assms(2), of "MS (MSLCons (C, e) xs) g"] assms have "\\<^sub>F x in at_top. abs (f x) = f x" by (auto simp: dominant_term_ms_aux_MSLCons trimmed_pos_def expands_to.simps sgn_if split: if_splits elim: eventually_elim2) from filterlim_cong[OF refl refl this] f show *: "filterlim (\x. exp (f x)) at_top at_top" by (auto intro: filterlim_compose[OF exp_at_top]) { fix e' :: real assume e': "e' > 1" from assms have "(\x. exp (g x)) \ \(\x. exp (f x) powr 1)" by (intro bigthetaI_cong) (auto elim: eventually_mono simp: expands_to.simps) also from e' * have "(\x. exp (f x) powr 1) \ o(\x. exp (f x) powr e')" by (subst powr_smallo_iff) (insert *, simp_all) finally have "(\x. exp (g x)) \ o(\x. exp (f x) powr e')" . } with assms show ?thesis3 by (auto intro!: is_expansion_aux.intros is_expansion_const simp: expands_to.simps is_expansion_aux_expansion_level dest: is_expansion_aux_expansion_level) qed (insert assms, simp_all) lemma expands_to_exp_insert_neg: fixes C :: "'a :: multiseries" assumes "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)" "basis_wf (b # basis)" "trimmed_neg (MS (MSLCons (C, e) xs) g)" "(\x. ln (b x)) \ o(f)" shows "filterlim (\x. exp (-f x)) at_top at_top" (is ?thesis1) "((\x. ln (exp (-f x))) expands_to MS (MSLCons (-C, e) (uminus_ms_aux xs)) (\x. - g x)) (b # basis)" "trimmed_pos (MS (MSLCons (-C, e) (uminus_ms_aux xs)) (\x. - g x))" "((\x. exp (f x)) expands_to MS (MSLCons (const_expansion 1 :: 'a ms, -1) MSLNil) (\x. exp (g x))) ((\x. exp (-f x)) # b # basis)" (is ?thesis3) proof - have "ln \ \(\x::real. 1)" by (intro smallomegaI_filterlim_at_top_norm) (auto intro!: filterlim_compose[OF filterlim_abs_real] ln_at_top) with assms have "(\x. 1) \ o(\x. ln (b x))" by (intro landau_o.small.compose[of _ at_top _ b]) (simp_all add: basis_wf_Cons smallomega_iff_smallo) also note \(\x. ln (b x)) \ o(f)\ finally have f: "f \ \(\_. 1)" by (simp add: smallomega_iff_smallo) hence f: "filterlim (\x. abs (f x)) at_top at_top" by (auto dest: smallomegaD_filterlim_at_top_norm) from trimmed_imp_eventually_sgn[OF assms(2), of "MS (MSLCons (C, e) xs) g"] assms have "\\<^sub>F x in at_top. abs (f x) = -f x" by (auto simp: dominant_term_ms_aux_MSLCons trimmed_neg_def expands_to.simps sgn_if split: if_splits elim: eventually_elim2) from filterlim_cong[OF refl refl this] f show *: "filterlim (\x. exp (-f x)) at_top at_top" by (auto intro: filterlim_compose[OF exp_at_top]) have "((\x. -f x) expands_to -MS (MSLCons (C, e) xs) g) (b # basis)" by (intro expands_to_uminus assms) thus "((\x. ln (exp (-f x))) expands_to MS (MSLCons (-C, e) (uminus_ms_aux xs)) (\x. - g x)) (b # basis)" by (simp add: uminus_ms_aux_MSLCons) { fix e' :: real assume e': "e' > -1" from assms have "(\x. exp (g x)) \ \(\x. exp (-f x) powr -1)" by (intro bigthetaI_cong) (auto elim: eventually_mono simp: expands_to.simps powr_minus exp_minus) also from e' * have "(\x. exp (-f x) powr -1) \ o(\x. exp (-f x) powr e')" by (subst powr_smallo_iff) (insert *, simp_all) finally have "(\x. exp (g x)) \ o(\x. exp (-f x) powr e')" . } with assms show ?thesis3 by (auto intro!: is_expansion_aux.intros is_expansion_const simp: expands_to.simps is_expansion_aux_expansion_level powr_minus exp_minus dest: is_expansion_aux_expansion_level) qed (insert assms, simp_all add: trimmed_pos_def trimmed_neg_def trimmed_ms_aux_MSLCons dominant_term_ms_aux_MSLCons) lemma is_expansion_aux_exp_neg: assumes "is_expansion_aux F f basis" "basis_wf basis" "ms_exp_gt 0 (ms_aux_hd_exp F)" shows "is_expansion_aux (powser_ms_aux' exp_series_stream F) (\x. exp (f x)) basis" using powser_ms_aux'[OF convergent_powser'_exp assms(2,1,3)] by (simp add: o_def powser_ms_aux'_def) lemma is_expansion_exp_neg: assumes "is_expansion (MS (MSLCons (C, e) xs) f) basis" "basis_wf basis" "e < 0" shows "is_expansion (MS (powser_ms_aux' exp_series_stream (MSLCons (C, e) xs)) (\x. exp (f x))) basis" using is_expansion_aux_exp_neg[of "MSLCons (C, e) xs" f basis] assms by simp lemma expands_to_exp_neg: assumes "(f expands_to MS (MSLCons (C, e) xs) g) basis" "basis_wf basis" "e - 0 < 0" shows "((\x. exp (f x)) expands_to MS (powser_ms_aux' exp_series_stream (MSLCons (C, e) xs)) (\x. exp (g x))) basis" using assms is_expansion_exp_neg[of C e xs g basis] by (auto simp: expands_to.simps) lemma basis_wf_insert_exp_near: assumes "basis_wf (b # basis)" "basis_wf ((\x. exp (f x)) # basis)" "f \ o(\x. ln (b x))" shows "basis_wf (b # (\x. exp (f x)) # basis)" using assms by (simp_all add: basis_wf_Cons) definition scale_ms_aux' where "scale_ms_aux' c F = scale_shift_ms_aux (c, 0) F" definition shift_ms_aux where "shift_ms_aux e F = scale_shift_ms_aux (const_expansion 1, e) F" lemma scale_ms_1 [simp]: "scale_ms 1 F = F" by (simp add: scale_ms_def times_const_expansion_1) lemma scale_ms_aux_1 [simp]: "scale_ms_aux 1 xs = xs" by (simp add: scale_ms_aux_def scale_shift_ms_aux_def map_prod_def msllist.map_ident case_prod_unfold times_const_expansion_1) lemma scale_ms_aux'_1 [simp]: "scale_ms_aux' (const_expansion 1) xs = xs" by (simp add: scale_ms_aux'_def scale_shift_ms_aux_def map_prod_def msllist.map_ident case_prod_unfold times_const_expansion_1) lemma shift_ms_aux_0 [simp]: "shift_ms_aux 0 xs = xs" by (simp add: shift_ms_aux_def scale_shift_ms_aux_def map_prod_def case_prod_unfold times_const_expansion_1 msllist.map_ident) lemma scale_ms_aux'_MSLNil [simp]: "scale_ms_aux' C MSLNil = MSLNil" by (simp add: scale_ms_aux'_def) lemma scale_ms_aux'_MSLCons [simp]: "scale_ms_aux' C (MSLCons (C', e) xs) = MSLCons (C * C', e) (scale_ms_aux' C xs)" by (simp add: scale_ms_aux'_def scale_shift_ms_aux_MSLCons) lemma shift_ms_aux_MSLNil [simp]: "shift_ms_aux e MSLNil = MSLNil" by (simp add: shift_ms_aux_def) lemma shift_ms_aux_MSLCons [simp]: "shift_ms_aux e (MSLCons (C, e') xs) = MSLCons (C, e' + e) (shift_ms_aux e xs)" by (simp add: shift_ms_aux_def scale_shift_ms_aux_MSLCons times_const_expansion_1) lemma scale_ms_aux: fixes xs :: "('a :: multiseries \ real) msllist" assumes "is_expansion_aux xs f basis" "basis_wf basis" shows "is_expansion_aux (scale_ms_aux c xs) (\x. c * f x) basis" proof (cases basis) case (Cons b basis') show ?thesis proof (rule is_expansion_aux_cong) show "is_expansion_aux (scale_ms_aux c xs) (\x. eval (const_expansion c :: 'a) x * b x powr 0 * f x) basis" using assms unfolding scale_ms_aux_def Cons by (intro scale_shift_ms_aux is_expansion_const) (auto simp: basis_wf_Cons dest: is_expansion_aux_expansion_level) next from assms have "eventually (\x. b x > 0) at_top" by (simp add: basis_wf_Cons Cons filterlim_at_top_dense) thus "eventually (\x. eval (const_expansion c :: 'a) x * b x powr 0 * f x = c * f x) at_top" by eventually_elim simp qed qed (insert assms, auto dest: is_expansion_aux_basis_nonempty) lemma scale_ms_aux': assumes "is_expansion_aux xs f (b # basis)" "is_expansion C basis" "basis_wf (b # basis)" shows "is_expansion_aux (scale_ms_aux' C xs) (\x. eval C x * f x) (b # basis)" proof (rule is_expansion_aux_cong) show "is_expansion_aux (scale_ms_aux' C xs) (\x. eval C x * b x powr 0 * f x) (b # basis)" using assms unfolding scale_ms_aux'_def Cons by (intro scale_shift_ms_aux) simp_all next from assms have "eventually (\x. b x > 0) at_top" by (simp add: basis_wf_Cons filterlim_at_top_dense) thus "eventually (\x. eval C x * b x powr 0 * f x = eval C x * f x) at_top" by eventually_elim simp qed lemma shift_ms_aux: fixes xs :: "('a :: multiseries \ real) msllist" assumes "is_expansion_aux xs f (b # basis)" "basis_wf (b # basis)" shows "is_expansion_aux (shift_ms_aux e xs) (\x. b x powr e * f x) (b # basis)" unfolding shift_ms_aux_def using scale_shift_ms_aux[OF assms(2,1) is_expansion_const[of _ 1], of e] assms by (auto dest!: is_expansion_aux_expansion_level simp: basis_wf_Cons) lemma expands_to_exp_0: assumes "(f expands_to MS (MSLCons (MS ys c, e) xs) g) (b # basis)" "basis_wf (b # basis)" "e - 0 = 0" "((\x. exp (c x)) expands_to E) basis" shows "((\x. exp (f x)) expands_to MS (scale_ms_aux' E (powser_ms_aux' exp_series_stream xs)) (\x. exp (g x))) (b # basis)" (is "(_ expands_to MS ?H _) _") proof - from assms have "is_expansion_aux ?H (\x. eval E x * exp (g x - c x * b x powr 0)) (b # basis)" by (intro scale_ms_aux' is_expansion_aux_exp_neg) (auto elim!: is_expansion_aux_MSLCons simp: expands_to.simps) moreover { from \basis_wf (b # basis)\ have "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense basis_wf_Cons) moreover from assms(4) have "eventually (\x. eval E x = exp (c x)) at_top" by (simp add: expands_to.simps) ultimately have "eventually (\x. eval E x * exp (g x - c x * b x powr 0) = exp (g x)) at_top" by eventually_elim (simp add: exp_diff) } ultimately have "is_expansion_aux ?H (\x. exp (g x)) (b # basis)" by (rule is_expansion_aux_cong) with assms(1) show ?thesis by (simp add: expands_to.simps) qed lemma expands_to_exp_0_real: assumes "(f expands_to MS (MSLCons (c::real, e) xs) g) [b]" "basis_wf [b]" "e - 0 = 0" shows "((\x. exp (f x)) expands_to MS (scale_ms_aux (exp c) (powser_ms_aux' exp_series_stream xs)) (\x. exp (g x))) [b]" (is "(_ expands_to MS ?H _) _") proof - from assms have "is_expansion_aux ?H (\x. exp c * exp (g x - c * b x powr 0)) [b]" by (intro scale_ms_aux is_expansion_aux_exp_neg) (auto elim!: is_expansion_aux_MSLCons simp: expands_to.simps) moreover from \basis_wf [b]\ have "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense basis_wf_Cons) hence "eventually (\x. exp c * exp (g x - c * b x powr 0) = exp (g x)) at_top" by eventually_elim (simp add: exp_diff) ultimately have "is_expansion_aux ?H (\x. exp (g x)) [b]" by (rule is_expansion_aux_cong) with assms(1) show ?thesis by (simp add: expands_to.simps) qed lemma expands_to_exp_0_pull_out1: assumes "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)" assumes "((\x. ln (b x)) expands_to L) basis" "basis_wf (b # basis)" "e - 0 = 0" "c \ c" shows "((\x. f x - c * ln (b x)) expands_to MS (MSLCons (C - scale_ms c L, e) xs) (\x. g x - c * ln (b x))) (b # basis)" proof - from \basis_wf (b # basis)\ have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) have "(\x. c * ln (b x)) \ o(\x. b x powr e')" if "e' > 0" for e' using that by (intro landau_o.small.compose[OF _ b_limit]) (simp add: ln_smallo_powr) hence *: "(\x. g x - c * ln (b x)) \ o(\x. b x powr e')" if "e' > 0" for e' using assms(1,4) that by (intro sum_in_smallo) (auto simp: expands_to.simps elim!: is_expansion_aux_MSLCons) have "is_expansion_aux xs (\x. (g x - c * b x powr 0 * eval L x) - eval (C - scale_ms c L) x * b x powr e) (b # basis)" (is "is_expansion_aux _ ?h _") using assms by (auto simp: expands_to.simps algebra_simps scale_ms_def elim!: is_expansion_aux_MSLCons) moreover from b_limit have "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense) hence "eventually (\x. ?h x = g x - c * ln (b x) - eval (C - const_expansion c * L) x * b x powr e) at_top" (is "eventually (\x. _ = ?h x) _") using assms(2) by (auto simp: expands_to.simps scale_ms_def elim: eventually_elim2) ultimately have **: "is_expansion_aux xs ?h (b # basis)" by (rule is_expansion_aux_cong) from assms(1,4) have "ms_exp_gt 0 (ms_aux_hd_exp xs)" "is_expansion C basis" by (auto simp: expands_to.simps elim!: is_expansion_aux_MSLCons) moreover from assms(1) have "length basis = expansion_level TYPE('a)" by (auto simp: expands_to.simps dest: is_expansion_aux_expansion_level) ultimately have "is_expansion_aux (MSLCons (C - scale_ms c L, e) xs) (\x. g x - c * ln (b x)) (b # basis)" using assms unfolding scale_ms_def by (intro is_expansion_aux.intros is_expansion_minus is_expansion_mult is_expansion_const * **) (simp_all add: basis_wf_Cons expands_to.simps) with assms(1) show ?thesis by (auto simp: expands_to.simps) qed lemma expands_to_exp_0_pull_out2: assumes "((\x. exp (f x - c * ln (b x))) expands_to MS xs g) (b # basis)" assumes "basis_wf (b # basis)" shows "((\x. exp (f x)) expands_to MS (shift_ms_aux c xs) (\x. b x powr c * g x)) (b # basis)" proof (rule expands_to.intros) show "is_expansion (MS (shift_ms_aux c xs) (\x. b x powr c * g x)) (b # basis)" using assms unfolding expands_to.simps by (auto intro!: shift_ms_aux) from assms have "eventually (\x. b x > 0) at_top" by (simp add: basis_wf_Cons filterlim_at_top_dense) with assms(1) show "\\<^sub>F x in at_top. eval (MS (shift_ms_aux c xs) (\x. b x powr c * g x)) x = exp (f x)" by (auto simp: expands_to.simps exp_diff powr_def elim: eventually_elim2) qed lemma expands_to_exp_0_pull_out2_nat: assumes "((\x. exp (f x - c * ln (b x))) expands_to MS xs g) (b # basis)" assumes "basis_wf (b # basis)" "c = real n" shows "((\x. exp (f x)) expands_to MS (shift_ms_aux c xs) (\x. b x ^ n * g x)) (b # basis)" using expands_to_exp_0_pull_out2[OF assms(1-2)] proof (rule expands_to_cong') from assms(2) have "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense basis_wf_Cons) with assms(3) show "eventually (\x. b x powr c * g x = b x ^ n * g x) at_top" by (auto elim!: eventually_mono simp: powr_realpow) qed lemma expands_to_exp_0_pull_out2_neg_nat: assumes "((\x. exp (f x - c * ln (b x))) expands_to MS xs g) (b # basis)" assumes "basis_wf (b # basis)" "c = -real n" shows "((\x. exp (f x)) expands_to MS (shift_ms_aux c xs) (\x. g x / b x ^ n)) (b # basis)" using expands_to_exp_0_pull_out2[OF assms(1-2)] proof (rule expands_to_cong') from assms(2) have "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense basis_wf_Cons) with assms(3) show "eventually (\x. b x powr c * g x = g x / b x ^ n) at_top" by (auto elim!: eventually_mono simp: powr_realpow powr_minus field_split_simps) qed lemma eval_monom_collapse: "c * eval_monom (c', es) basis x = eval_monom (c * c', es) basis x" by (simp add: eval_monom_def) lemma eval_monom_1_conv: "eval_monom a basis = (\x. fst a * eval_monom (1, snd a) basis x)" by (simp add: eval_monom_def case_prod_unfold) subsubsection \Comparing exponent lists\ primrec flip_cmp_result where "flip_cmp_result LT = GT" | "flip_cmp_result GT = LT" | "flip_cmp_result EQ = EQ" fun compare_lists :: "real list \ real list \ cmp_result" where "compare_lists [] [] = EQ" | "compare_lists (a # as) (b # bs) = (if a < b then LT else if b < a then GT else compare_lists as bs)" | "compare_lists _ _ = undefined" declare compare_lists.simps [simp del] lemma compare_lists_Nil [simp]: "compare_lists [] [] = EQ" by (fact compare_lists.simps) lemma compare_lists_Cons [simp]: "a < b \ compare_lists (a # as) (b # bs) = LT" "a > b \ compare_lists (a # as) (b # bs) = GT" "a = b \ compare_lists (a # as) (b # bs) = compare_lists as bs" by (simp_all add: compare_lists.simps(2)) lemma flip_compare_lists: "length as = length bs \ flip_cmp_result (compare_lists as bs) = compare_lists bs as" by (induction as bs rule: compare_lists.induct) (auto simp: compare_lists.simps(2)) lemma compare_lists_induct [consumes 1, case_names Nil Eq Neq]: fixes as bs :: "real list" assumes "length as = length bs" assumes "P [] []" assumes "\a as bs. P as bs \ P (a # as) (a # bs)" assumes "\a as b bs. a \ b \ P (a # as) (b # bs)" shows "P as bs" using assms(1) proof (induction as bs rule: list_induct2) case (Cons a as b bs) thus ?case by (cases "a < b") (auto intro: assms) qed (simp_all add: assms) lemma eval_monom_smallo_eval_monom: assumes "length es1 = length es2" "length es2 = length basis" "basis_wf basis" assumes "compare_lists es1 es2 = LT" shows "eval_monom (1, es1) basis \ o(eval_monom (1, es2) basis)" using assms proof (induction es1 es2 basis rule: list_induct3) case (Cons e1 es1 e2 es2 b basis) show ?case proof (cases "e1 = e2") case True from Cons.prems have "eventually (\x. b x > 0) at_top" by (simp add: basis_wf_Cons filterlim_at_top_dense) with Cons True show ?thesis by (auto intro: landau_o.small_big_mult simp: eval_monom_Cons basis_wf_Cons) next case False with Cons.prems have "e1 < e2" by (cases e1 e2 rule: linorder_cases) simp_all with Cons.prems have "(\x. eval_monom (1, es1) basis x * eval_monom (inverse_monom (1, es2)) basis x * b x powr e1) \ o(\x. b x powr ((e2 - e1)/2) * b x powr ((e2 - e1)/2) * b x powr e1)" (is "?f \ _") by (intro landau_o.small_big_mult[OF _ landau_o.big_refl] landau_o.small.mult eval_monom_smallo') simp_all also have "\ = o(\x. b x powr e2)" by (simp add: powr_add [symmetric]) also have "?f = (\x. eval_monom (1, e1 # es1) (b # basis) x / eval_monom (1, es2) basis x)" by (simp add: eval_inverse_monom field_simps eval_monom_Cons) also have "\ \ o(\x. b x powr e2) \ eval_monom (1, e1 # es1) (b # basis) \ o(eval_monom (1, e2 # es2) (b # basis))" using Cons.prems by (subst landau_o.small.divide_eq2) (simp_all add: eval_monom_Cons mult_ac eval_monom_nonzero basis_wf_Cons) finally show ?thesis . qed qed simp_all lemma eval_monom_eq: assumes "length es1 = length es2" "length es2 = length basis" "basis_wf basis" assumes "compare_lists es1 es2 = EQ" shows "eval_monom (1, es1) basis = eval_monom (1, es2) basis" using assms by (induction es1 es2 basis rule: list_induct3) (auto simp: basis_wf_Cons eval_monom_Cons compare_lists.simps(2) split: if_splits) definition compare_expansions :: "'a :: multiseries \ 'a \ cmp_result \ real \ real" where "compare_expansions F G = (case compare_lists (snd (dominant_term F)) (snd (dominant_term G)) of LT \ (LT, 0, 0) | GT \ (GT, 0, 0) | EQ \ (EQ, fst (dominant_term F), fst (dominant_term G)))" lemma compare_expansions_real: "compare_expansions (c1 :: real) c2 = (EQ, c1, c2)" by (simp add: compare_expansions_def) lemma compare_expansions_MSLCons_eval: "compare_expansions (MS (MSLCons (C1, e1) xs) f) (MS (MSLCons (C2, e2) ys) g) = CMP_BRANCH (COMPARE e1 e2) (LT, 0, 0) (compare_expansions C1 C2) (GT, 0, 0)" by (simp add: compare_expansions_def dominant_term_ms_aux_def case_prod_unfold COMPARE_def CMP_BRANCH_def split: cmp_result.splits) lemma compare_expansions_LT_I: assumes "e1 - e2 < 0" shows "compare_expansions (MS (MSLCons (C1, e1) xs) f) (MS (MSLCons (C2, e2) ys) g) = (LT, 0, 0)" using assms by (simp add: compare_expansions_MSLCons_eval CMP_BRANCH_def COMPARE_def) lemma compare_expansions_GT_I: assumes "e1 - e2 > 0" shows "compare_expansions (MS (MSLCons (C1, e1) xs) f) (MS (MSLCons (C2, e2) ys) g) = (GT, 0, 0)" using assms by (simp add: compare_expansions_MSLCons_eval CMP_BRANCH_def COMPARE_def) lemma compare_expansions_same_exp: assumes "e1 - e2 = 0" "compare_expansions C1 C2 = res" shows "compare_expansions (MS (MSLCons (C1, e1) xs) f) (MS (MSLCons (C2, e2) ys) g) = res" using assms by (simp add: compare_expansions_MSLCons_eval CMP_BRANCH_def COMPARE_def) lemma compare_expansions_GT_flip: "length (snd (dominant_term F)) = length (snd (dominant_term G)) \ compare_expansions F G = (GT, c) \ compare_expansions G F = (LT, c)" using flip_compare_lists[of "snd (dominant_term F)" "snd (dominant_term G)"] by (auto simp: compare_expansions_def split: cmp_result.splits) lemma compare_expansions_LT: assumes "compare_expansions F G = (LT, c, c')" "trimmed G" "(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis" shows "f \ o(g)" proof - from assms have "f \ \(eval F)" by (auto simp: expands_to.simps eq_commute intro!: bigthetaI_cong) also have "eval F \ O(eval_monom (1, snd (dominant_term F)) basis)" using assms by (intro dominant_term_bigo) (auto simp: expands_to.simps) also have "eval_monom (1, snd (dominant_term F)) basis \ o(eval_monom (1, snd (dominant_term G)) basis)" using assms by (intro eval_monom_smallo_eval_monom) (auto simp: length_dominant_term expands_to.simps is_expansion_length compare_expansions_def split: cmp_result.splits) also have "eval_monom (1, snd (dominant_term G)) basis \ \(eval_monom (dominant_term G) basis)" by (subst (2) eval_monom_1_conv, subst landau_theta.cmult) (insert assms, simp_all add: trimmed_imp_dominant_term_nz) also have "eval_monom (dominant_term G) basis \ \(g)" by (intro asymp_equiv_imp_bigtheta[OF asymp_equiv_symI] dominant_term_expands_to asymp_equiv_imp_bigtheta assms) finally show ?thesis . qed lemma compare_expansions_GT: assumes "compare_expansions F G = (GT, c, c')" "trimmed F" "(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis" shows "g \ o(f)" by (rule compare_expansions_LT[OF _ assms(2,4,3)]) (insert assms, simp_all add: compare_expansions_GT_flip length_dominant_term) lemma compare_expansions_EQ: assumes "compare_expansions F G = (EQ, c, c')" "trimmed F" "trimmed G" "(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis" shows "(\x. c' * f x) \[at_top] (\x. c * g x)" proof - from assms(1) have c: "c = fst (dominant_term F)" "c' = fst (dominant_term G)" by (auto simp: compare_expansions_def split: cmp_result.splits) have "(\x. c' * f x) \[at_top] (\x. c' * (c * eval_monom (1, snd (dominant_term F)) basis x))" unfolding c by (rule asymp_equiv_mult, rule asymp_equiv_refl, subst eval_monom_collapse) (auto intro: dominant_term_expands_to assms) also have "eval_monom (1, snd (dominant_term F)) basis = eval_monom (1, snd (dominant_term G)) basis" using assms by (intro eval_monom_eq) (simp_all add: compare_expansions_def length_dominant_term is_expansion_length expands_to.simps split: cmp_result.splits) also have "(\x. c' * (c * \ x)) = (\x. c * (c' * \ x))" by (simp add: mult_ac) also have "\ \[at_top] (\x. c * g x)" unfolding c by (rule asymp_equiv_mult, rule asymp_equiv_refl, subst eval_monom_collapse, rule asymp_equiv_symI) (auto intro: dominant_term_expands_to assms) finally show ?thesis by (simp add: asymp_equiv_sym) qed lemma compare_expansions_EQ_imp_bigo: assumes "compare_expansions F G = (EQ, c, c')" "trimmed G" "(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis" shows "f \ O(g)" proof - from assms(1) have c: "c = fst (dominant_term F)" "c' = fst (dominant_term G)" by (auto simp: compare_expansions_def split: cmp_result.splits) from assms(3) have [simp]: "expansion_level TYPE('a) = length basis" by (auto simp: expands_to.simps is_expansion_length) have "f \ \(eval F)" using assms by (intro bigthetaI_cong) (auto simp: expands_to.simps eq_commute) also have "eval F \ O(eval_monom (1, snd (dominant_term F)) basis)" using assms by (intro dominant_term_bigo assms) (auto simp: expands_to.simps) also have "eval_monom (1, snd (dominant_term F)) basis = eval_monom (1, snd (dominant_term G)) basis" using assms by (intro eval_monom_eq) (auto simp: compare_expansions_def case_prod_unfold length_dominant_term split: cmp_result.splits) also have "\ \ O(eval_monom (dominant_term G) basis)" using assms(2) by (auto simp add: eval_monom_def case_prod_unfold dest: trimmed_imp_dominant_term_nz) also have "eval_monom (dominant_term G) basis \ \(eval G)" by (rule asymp_equiv_imp_bigtheta, rule asymp_equiv_symI, rule dominant_term) (insert assms, auto simp: expands_to.simps) also have "eval G \ \(g)" using assms by (intro bigthetaI_cong) (auto simp: expands_to.simps) finally show ?thesis . qed lemma compare_expansions_EQ_same: assumes "compare_expansions F G = (EQ, c, c')" "trimmed F" "trimmed G" "(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis" "c = c'" shows "f \[at_top] g" proof - from assms have [simp]: "c' \ 0" by (auto simp: compare_expansions_def trimmed_imp_dominant_term_nz split: cmp_result.splits) have "(\x. inverse c * (c' * f x)) \[at_top] (\x. inverse c * (c * g x))" by (rule asymp_equiv_mult[OF asymp_equiv_refl]) (rule compare_expansions_EQ[OF assms(1-6)]) with assms(7) show ?thesis by (simp add: field_split_simps) qed lemma compare_expansions_EQ_imp_bigtheta: assumes "compare_expansions F G = (EQ, c, c')" "trimmed F" "trimmed G" "(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis" shows "f \ \(g)" proof - from assms have "(\x. c' * f x) \ \(\x. c * g x)" by (intro asymp_equiv_imp_bigtheta compare_expansions_EQ) moreover from assms have "c \ 0" "c' \ 0" by (auto simp: compare_expansions_def trimmed_imp_dominant_term_nz split: cmp_result.splits) ultimately show ?thesis by simp qed lemma expands_to_MSLCons_0_asymp_equiv_hd: assumes "(f expands_to (MS (MSLCons (C, 0) xs) g)) basis" "trimmed (MS (MSLCons (C, 0) xs) f)" "basis_wf basis" shows "f \[at_top] eval C" proof - from assms(1) is_expansion_aux_basis_nonempty obtain b basis' where [simp]: "basis = b # basis'" by (cases basis) (auto simp: expands_to.simps) from assms have b_pos: "eventually (\x. b x > 0) at_top" by (simp add: filterlim_at_top_dense basis_wf_Cons) from assms have "f \[at_top] eval_monom (dominant_term (MS (MSLCons (C, 0) xs) g)) basis" by (intro dominant_term_expands_to) simp_all also have "\ = (\x. eval_monom (dominant_term C) basis' x * b x powr 0)" by (simp_all add: dominant_term_ms_aux_def case_prod_unfold eval_monom_Cons) also have "eventually (\x. \ x = eval_monom (dominant_term C) basis' x) at_top" using b_pos by eventually_elim simp also from assms have "eval_monom (dominant_term C) basis' \[at_top] eval C" by (intro asymp_equiv_symI [OF dominant_term_expands_to] expands_to_hd''[of f C 0 xs g b basis']) (auto simp: trimmed_ms_aux_MSLCons basis_wf_Cons) finally show ?thesis . qed lemma compare_expansions_LT': assumes "compare_expansions (MS ys h) G \ (LT, c, c')" "trimmed G" "(f expands_to (MS (MSLCons (MS ys h, e) xs) f')) (b # basis)" "(g expands_to G) basis" "e = 0" "basis_wf (b # basis)" shows "h \ o(g)" proof - from assms show ?thesis by (intro compare_expansions_LT[OF _ assms(2) _ assms(4)]) (auto intro: expands_to_hd'' simp: trimmed_ms_aux_MSLCons basis_wf_Cons expands_to.simps elim!: is_expansion_aux_MSLCons) qed lemma compare_expansions_GT': assumes "compare_expansions C G \ (GT, c, c')" "trimmed (MS (MSLCons (C, e) xs) f')" "(f expands_to (MS (MSLCons (C, e) xs) f')) (b # basis)" "(g expands_to G) basis" "e = 0" "basis_wf (b # basis)" shows "g \ o(f)" proof - from assms have "g \ o(eval C)" by (intro compare_expansions_GT[of C G c c' _ basis]) (auto intro: expands_to_hd'' simp: trimmed_ms_aux_MSLCons basis_wf_Cons) also from assms have "f \ \(eval C)" by (intro asymp_equiv_imp_bigtheta expands_to_MSLCons_0_asymp_equiv_hd[of f C xs f' "b#basis"]) auto finally show ?thesis . qed lemma compare_expansions_EQ': assumes "compare_expansions C G = (EQ, c, c')" "trimmed (MS (MSLCons (C, e) xs) f')" "trimmed G" "(f expands_to (MS (MSLCons (C, e) xs) f')) (b # basis)" "(g expands_to G) basis" "e = 0" "basis_wf (b # basis)" shows "f \[at_top] (\x. c / c' * g x)" proof - from assms have [simp]: "c' \ 0" by (auto simp: compare_expansions_def trimmed_imp_dominant_term_nz split: cmp_result.splits) from assms have "f \[at_top] eval C" by (intro expands_to_MSLCons_0_asymp_equiv_hd[of f C xs f' "b#basis"]) auto also from assms(2,4,7) have *: "(\x. c' * eval C x) \[at_top] (\x. c * g x)" by (intro compare_expansions_EQ[OF assms(1) _ assms(3) _ assms(5)]) (auto intro: expands_to_hd'' simp: trimmed_ms_aux_MSLCons basis_wf_Cons) have "(\x. inverse c' * (c' * eval C x)) \[at_top] (\x. inverse c' * (c * g x))" by (rule asymp_equiv_mult) (insert *, simp_all) hence "eval C \[at_top] (\x. c / c' * g x)" by (simp add: field_split_simps) finally show ?thesis . qed lemma expands_to_insert_ln: assumes "basis_wf [b]" shows "((\x. ln (b x)) expands_to MS (MSLCons (1::real,1) MSLNil) (\x. ln (b x))) [\x. ln (b x)]" proof - from assms have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b: "eventually (\x. b x > 1) at_top" by (simp add: filterlim_at_top_dense) have "(\x::real. x) \ \(\x. x powr 1)" by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto also have "(\x::real. x powr 1) \ o(\a. a powr e)" if "e > 1" for e by (subst powr_smallo_iff) (auto simp: that filterlim_ident) finally show ?thesis using b assms by (auto intro!: is_expansion_aux.intros landau_o.small.compose[of _ at_top _ "\x. ln (b x)"] filterlim_compose[OF ln_at_top b_limit] is_expansion_real.intros elim!: eventually_mono simp: expands_to.simps) qed lemma trimmed_pos_insert_ln: assumes "basis_wf [b]" shows "trimmed_pos (MS (MSLCons (1::real, 1) MSLNil) (\x. ln (b x)))" by (simp_all add: trimmed_ms_aux_def) primrec compare_list_0 where "compare_list_0 [] = EQ" | "compare_list_0 (c # cs) = CMP_BRANCH (COMPARE c 0) LT (compare_list_0 cs) GT" lemma compare_reals_diff_sgnD: "a - (b :: real) < 0 \ a < b" "a - b = 0 \ a = b" "a - b > 0 \ a > b" by simp_all lemma expands_to_real_imp_filterlim: assumes "(f expands_to (c :: real)) basis" shows "(f \ c) at_top" using assms by (auto elim!: expands_to.cases simp: eq_commute[of c] tendsto_eventually) lemma expands_to_MSLNil_imp_filterlim: assumes "(f expands_to MS MSLNil f') basis" shows "(f \ 0) at_top" proof - from assms have "eventually (\x. f' x = 0) at_top" "eventually (\x. f' x = f x) at_top" by (auto elim!: expands_to.cases is_expansion_aux.cases[of MSLNil]) hence "eventually (\x. f x = 0) at_top" by eventually_elim auto thus ?thesis by (simp add: tendsto_eventually) qed lemma expands_to_neg_exponent_imp_filterlim: assumes "(f expands_to MS (MSLCons (C, e) xs) f') basis" "basis_wf basis" "e < 0" shows "(f \ 0) at_top" proof - let ?F = "MS (MSLCons (C, e) xs) f'" let ?es = "snd (dominant_term ?F)" from assms have exp: "is_expansion ?F basis" by (simp add: expands_to.simps) from assms have "f \ \(f')" by (intro bigthetaI_cong) (simp add: expands_to.simps eq_commute) also from dominant_term_bigo[OF assms(2) exp] have "f' \ O(eval_monom (1, ?es) basis)" by simp also have "(eval_monom (1, ?es) basis) \ o(\x. hd basis x powr 0)" using exp assms by (intro eval_monom_smallo) (auto simp: is_expansion_aux_basis_nonempty dominant_term_ms_aux_def case_prod_unfold length_dominant_term is_expansion_aux_expansion_level) also have "(\x. hd basis x powr 0) \ O(\_. 1)" by (intro landau_o.bigI[of 1]) auto finally show ?thesis by (auto dest!: smalloD_tendsto) qed lemma expands_to_pos_exponent_imp_filterlim: assumes "(f expands_to MS (MSLCons (C, e) xs) f') basis" "trimmed (MS (MSLCons (C, e) xs) f')" "basis_wf basis" "e > 0" shows "filterlim f at_infinity at_top" proof - let ?F = "MS (MSLCons (C, e) xs) f'" let ?es = "snd (dominant_term ?F)" from assms have exp: "is_expansion ?F basis" by (simp add: expands_to.simps) with assms have "filterlim (hd basis) at_top at_top" using is_expansion_aux_basis_nonempty[of "MSLCons (C, e) xs" f' basis] by (cases basis) (simp_all add: basis_wf_Cons) hence b_pos: "eventually (\x. hd basis x > 0) at_top" using filterlim_at_top_dense by blast from assms have "f \ \(f')" by (intro bigthetaI_cong) (simp add: expands_to.simps eq_commute) also from dominant_term[OF assms(3) exp assms(2)] have "f' \ \(eval_monom (dominant_term ?F) basis)" by (simp add: asymp_equiv_imp_bigtheta) also have "(eval_monom (dominant_term ?F) basis) \ \(eval_monom (1, ?es) basis)" using assms by (simp add: eval_monom_def case_prod_unfold dominant_term_ms_aux_def trimmed_imp_dominant_term_nz trimmed_ms_aux_def) also have "eval_monom (1, ?es) basis \ \(\x. hd basis x powr 0)" using exp assms by (intro eval_monom_smallomega) (auto simp: is_expansion_aux_basis_nonempty dominant_term_ms_aux_def case_prod_unfold length_dominant_term is_expansion_aux_expansion_level) also have "(\x. hd basis x powr 0) \ \(\_. 1)" by (intro bigthetaI_cong eventually_mono[OF b_pos]) auto finally show ?thesis by (auto dest!: smallomegaD_filterlim_at_infinity) qed lemma expands_to_zero_exponent_imp_filterlim: assumes "(f expands_to MS (MSLCons (C, e) xs) f') basis" "basis_wf basis" "e = 0" shows "filterlim (eval C) at_infinity at_top \ filterlim f at_infinity at_top" and "filterlim (eval C) (nhds L) at_top \ filterlim f (nhds L) at_top" proof - from assms obtain b basis' where *: "basis = b # basis'" "is_expansion C basis'" "ms_exp_gt 0 (ms_aux_hd_exp xs)" "is_expansion_aux xs (\x. f' x - eval C x * b x powr 0) basis" by (auto elim!: expands_to.cases is_expansion_aux_MSLCons) from *(1) assms have "filterlim b at_top at_top" by (simp add: basis_wf_Cons) hence b_pos: "eventually (\x. b x > 0) at_top" using filterlim_at_top_dense by blast from assms(1) have "eventually (\x. f' x = f x) at_top" by (simp add: expands_to.simps) hence "eventually (\x. f' x - eval C x * b x powr 0 = f x - eval C x) at_top" using b_pos by eventually_elim auto from *(4) and this have "is_expansion_aux xs (\x. f x - eval C x) basis" by (rule is_expansion_aux_cong) hence "(\x. f x - eval C x) \ o(\x. hd basis x powr 0)" by (rule is_expansion_aux_imp_smallo) fact also have "(\x. hd basis x powr 0) \ \(\_. 1)" by (intro bigthetaI_cong eventually_mono[OF b_pos]) (auto simp: *(1)) finally have lim: "filterlim (\x. f x - eval C x) (nhds 0) at_top" by (auto dest: smalloD_tendsto) show "filterlim f at_infinity at_top" if "filterlim (eval C) at_infinity at_top" using tendsto_add_filterlim_at_infinity[OF lim that] by simp show "filterlim f (nhds L) at_top" if "filterlim (eval C) (nhds L) at_top" using tendsto_add[OF lim that] by simp qed lemma expands_to_lift_function: assumes "eventually (\x. f x - eval c x = 0) at_top" "((\x. g (eval (c :: 'a :: multiseries) x)) expands_to G) bs'" shows "((\x. g (f x)) expands_to G) bs'" by (rule expands_to_cong[OF assms(2)]) (insert assms, auto elim: eventually_mono) lemma drop_zero_ms': fixes c f assumes "c = (0::real)" "(f expands_to MS (MSLCons (c, e) xs) g) basis" shows "(f expands_to MS xs g) basis" using assms drop_zero_ms[of c e xs g basis] by (simp add: expands_to.simps) lemma trimmed_realI: "c \ (0::real) \ trimmed c" by simp lemma trimmed_pos_realI: "c > (0::real) \ trimmed_pos c" by simp lemma trimmed_neg_realI: "c < (0::real) \ trimmed_neg c" by (simp add: trimmed_neg_def) lemma trimmed_pos_hd_coeff: "trimmed_pos (MS (MSLCons (C, e) xs) f) \ trimmed_pos C" by simp lemma lift_trimmed: "trimmed C \ trimmed (MS (MSLCons (C, e) xs) f)" by (auto simp: trimmed_ms_aux_def) lemma lift_trimmed_pos: "trimmed_pos C \ trimmed_pos (MS (MSLCons (C, e) xs) f)" by simp lemma lift_trimmed_neg: "trimmed_neg C \ trimmed_neg (MS (MSLCons (C, e) xs) f)" by (simp add: trimmed_neg_def fst_dominant_term_ms_aux_MSLCons trimmed_ms_aux_MSLCons) lemma lift_trimmed_pos': "trimmed_pos C \ MS (MSLCons (C, e) xs) f \ MS (MSLCons (C, e) xs) f \ trimmed_pos (MS (MSLCons (C, e) xs) f)" by simp lemma lift_trimmed_neg': "trimmed_neg C \ MS (MSLCons (C, e) xs) f \ MS (MSLCons (C, e) xs) f \ trimmed_neg (MS (MSLCons (C, e) xs) f)" by (simp add: lift_trimmed_neg) lemma trimmed_eq_cong: "trimmed C \ C \ C' \ trimmed C'" and trimmed_pos_eq_cong: "trimmed_pos C \ C \ C' \ trimmed_pos C'" and trimmed_neg_eq_cong: "trimmed_neg C \ C \ C' \ trimmed_neg C'" by simp_all lemma trimmed_hd: "trimmed (MS (MSLCons (C, e) xs) f) \ trimmed C" by (simp add: trimmed_ms_aux_MSLCons) lemma trim_lift_eq: assumes "A \ MS (MSLCons (C, e) xs) f" "C \ D" shows "A \ MS (MSLCons (D, e) xs) f" using assms by simp lemma basis_wf_manyI: "filterlim b' at_top at_top \ (\x. ln (b x)) \ o(\x. ln (b' x)) \ basis_wf (b # basis) \ basis_wf (b' # b # basis)" by (simp_all add: basis_wf_many) lemma ln_smallo_ln_exp: "(\x. ln (b x)) \ o(f) \ (\x. ln (b x)) \ o(\x. ln (exp (f x :: real)))" by simp subsection \Reification and other technical details\ text \ The following is used by the automation in order to avoid writing terms like $x^2$ or $x^{-2}$ as \<^term>\\x::real. x powr 2\ etc.\ but as the more agreeable \<^term>\\x::real. x ^ 2\ or \<^term>\\x::real. inverse (x ^ 2)\. \ lemma intyness_0: "0 \ real 0" and intyness_1: "1 \ real 1" and intyness_numeral: "num \ num \ numeral num \ real (numeral num)" and intyness_uminus: "x \ real n \ -x \ -real n" and intyness_of_nat: "n \ n \ real n \ real n" by simp_all lemma intyness_simps: "real a + real b = real (a + b)" "real a * real b = real (a * b)" "real a ^ n = real (a ^ n)" "1 = real 1" "0 = real 0" "numeral num = real (numeral num)" by simp_all lemma odd_Numeral1: "odd (Numeral1)" by simp lemma even_addI: "even a \ even b \ even (a + b)" "odd a \ odd b \ even (a + b)" by auto lemma odd_addI: "even a \ odd b \ odd (a + b)" "odd a \ even b \ odd (a + b)" by auto lemma even_diffI: "even a \ even b \ even (a - b :: 'a :: unique_euclidean_ring_with_nat)" "odd a \ odd b \ even (a - b)" by auto lemma odd_diffI: "even a \ odd b \ odd (a - b :: 'a :: unique_euclidean_ring_with_nat)" "odd a \ even b \ odd (a - b)" by auto lemma even_multI: "even a \ even (a * b)" "even b \ even (a * b)" by auto lemma odd_multI: "odd a \ odd b \ odd (a * b)" by auto lemma even_uminusI: "even a \ even (-a :: 'a :: unique_euclidean_ring_with_nat)" and odd_uminusI: "odd a \ odd (-a :: 'a :: unique_euclidean_ring_with_nat)" by auto lemma odd_powerI: "odd a \ odd (a ^ n)" by auto text \ The following theorem collection is used to pre-process functions for multiseries expansion. \ named_theorems real_asymp_reify_simps lemmas [real_asymp_reify_simps] = sinh_field_def cosh_field_def tanh_real_altdef arsinh_def arcosh_def artanh_def text \ This is needed in order to handle things like \<^term>\\n. f n ^ n\. \ definition powr_nat :: "real \ real \ real" where "powr_nat x y = (if y = 0 then 1 else if x < 0 then cos (pi * y) * (-x) powr y else x powr y)" lemma powr_nat_of_nat: "powr_nat x (of_nat n) = x ^ n" by (cases "x > 0") (auto simp: powr_nat_def powr_realpow not_less power_mult_distrib [symmetric]) lemma powr_nat_conv_powr: "x > 0 \ powr_nat x y = x powr y" by (simp_all add: powr_nat_def) lemma reify_power: "x ^ n \ powr_nat x (of_nat n)" by (simp add: powr_nat_of_nat) lemma sqrt_conv_root [real_asymp_reify_simps]: "sqrt x = root 2 x" by (simp add: sqrt_def) lemma tan_conv_sin_cos [real_asymp_reify_simps]: "tan x = sin x / cos x" by (simp add: tan_def) definition rfloor :: "real \ real" where "rfloor x = real_of_int (floor x)" definition rceil :: "real \ real" where "rceil x = real_of_int (ceiling x)" definition rnatmod :: "real \ real \ real" where "rnatmod x y = real (nat \x\ mod nat \y\)" definition rintmod :: "real \ real \ real" where "rintmod x y = real_of_int (\x\ mod \y\)" lemmas [real_asymp_reify_simps] = ln_exp log_def rfloor_def [symmetric] rceil_def [symmetric] lemma expands_to_powr_nat_0_0: assumes "eventually (\x. f x = 0) at_top" "eventually (\x. g x = 0) at_top" "basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)" shows "((\x. powr_nat (f x) (g x)) expands_to (const_expansion 1 :: 'a)) basis" proof (rule expands_to_cong [OF expands_to_const]) from assms(1,2) show "eventually (\x. 1 = powr_nat (f x) (g x)) at_top" by eventually_elim (simp add: powr_nat_def) qed fact+ lemma expands_to_powr_nat_0: assumes "eventually (\x. f x = 0) at_top" "(g expands_to G) basis" "trimmed G" "basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)" shows "((\x. powr_nat (f x) (g x)) expands_to (zero_expansion :: 'a)) basis" proof (rule expands_to_cong [OF expands_to_zero]) from assms have "eventually (\x. g x \ 0) at_top" using expands_to_imp_eventually_nz[of basis g G] by auto with assms(1) show "eventually (\x. 0 = powr_nat (f x) (g x)) at_top" by eventually_elim (simp add: powr_nat_def) qed (insert assms, auto elim!: eventually_mono simp: powr_nat_def) lemma expands_to_powr_nat: assumes "trimmed_pos F" "basis_wf basis'" "(f expands_to F) basis'" assumes "((\x. exp (ln (f x) * g x)) expands_to E) basis" shows "((\x. powr_nat (f x) (g x)) expands_to E) basis" using assms(4) proof (rule expands_to_cong) from eval_pos_if_dominant_term_pos[of basis' F] assms(1-4) show "eventually (\x. exp (ln (f x) * g x) = powr_nat (f x) (g x)) at_top" by (auto simp: expands_to.simps trimmed_pos_def powr_def powr_nat_def elim: eventually_elim2) qed subsection \Some technical lemmas\ lemma landau_meta_eq_cong: "f \ L(g) \ f' \ f \ g' \ g \ f' \ L(g')" and asymp_equiv_meta_eq_cong: "f \[at_top] g \ f' \ f \ g' \ g \ f' \[at_top] g'" by simp_all lemma filterlim_mono': "filterlim f F G \ F \ F' \ filterlim f F' G" by (erule (1) filterlim_mono) simp_all lemma at_within_le_nhds: "at x within A \ nhds x" by (simp add: at_within_def) lemma at_within_le_at: "at x within A \ at x" by (rule at_le) simp_all lemma at_right_to_top': "at_right c = filtermap (\x::real. c + inverse x) at_top" by (subst at_right_to_0, subst at_right_to_top) (simp add: filtermap_filtermap add_ac) lemma at_left_to_top': "at_left c = filtermap (\x::real. c - inverse x) at_top" by (subst at_left_minus, subst at_right_to_0, subst at_right_to_top) (simp add: filtermap_filtermap add_ac) lemma at_left_to_top: "at_left 0 = filtermap (\x::real. - inverse x) at_top" by (simp add: at_left_to_top') lemma filterlim_conv_filtermap: "G = filtermap g G' \ PROP (Trueprop (filterlim f F G)) \ PROP (Trueprop (filterlim (\x. f (g x)) F G'))" by (simp add: filterlim_filtermap) lemma eventually_conv_filtermap: "G = filtermap g G' \ PROP (Trueprop (eventually P G)) \ PROP (Trueprop (eventually (\x. P (g x)) G'))" by (simp add: eventually_filtermap) lemma eventually_lt_imp_eventually_le: "eventually (\x. f x < (g x :: real)) F \ eventually (\x. f x \ g x) F" by (erule eventually_mono) simp lemma smallo_imp_smallomega: "f \ o[F](g) \ g \ \[F](f)" by (simp add: smallomega_iff_smallo) lemma bigo_imp_bigomega: "f \ O[F](g) \ g \ \[F](f)" by (simp add: bigomega_iff_bigo) context fixes L L' :: "real filter \ (real \ _) \ _" and Lr :: "real filter \ (real \ real) \ _" assumes LS: "landau_symbol L L' Lr" begin interpretation landau_symbol L L' Lr by (fact LS) lemma landau_symbol_at_top_imp_at_bot: "(\x. f (-x)) \ L' at_top (\x. g (-x)) \ f \ L at_bot g" by (simp add: in_filtermap_iff at_bot_mirror) lemma landau_symbol_at_top_imp_at_right_0: "(\x. f (inverse x)) \ L' at_top (\x. g (inverse x)) \ f \ L (at_right 0) g" by (simp add: in_filtermap_iff at_right_to_top') lemma landau_symbol_at_top_imp_at_left_0: "(\x. f ( -inverse x)) \ L' at_top (\x. g (-inverse x)) \ f \ L (at_left 0) g" by (simp add: in_filtermap_iff at_left_to_top') lemma landau_symbol_at_top_imp_at_right: "(\x. f (a + inverse x)) \ L' at_top (\x. g (a + inverse x)) \ f \ L (at_right a) g" by (simp add: in_filtermap_iff at_right_to_top') lemma landau_symbol_at_top_imp_at_left: "(\x. f (a - inverse x)) \ L' at_top (\x. g (a - inverse x)) \ f \ L (at_left a) g" by (simp add: in_filtermap_iff at_left_to_top') lemma landau_symbol_at_top_imp_at: "(\x. f (a - inverse x)) \ L' at_top (\x. g (a - inverse x)) \ (\x. f (a + inverse x)) \ L' at_top (\x. g (a + inverse x)) \ f \ L (at a) g" by (simp add: sup at_eq_sup_left_right landau_symbol_at_top_imp_at_left landau_symbol_at_top_imp_at_right) lemma landau_symbol_at_top_imp_at_0: "(\x. f (-inverse x)) \ L' at_top (\x. g (-inverse x)) \ (\x. f (inverse x)) \ L' at_top (\x. g (inverse x)) \ f \ L (at 0) g" by (rule landau_symbol_at_top_imp_at) simp_all end context fixes f g :: "real \ real" begin lemma asymp_equiv_at_top_imp_at_bot: "(\x. f (- x)) \[at_top] (\x. g (-x)) \ f \[at_bot] g" by (simp add: asymp_equiv_def filterlim_at_bot_mirror) lemma asymp_equiv_at_top_imp_at_right_0: "(\x. f (inverse x)) \[at_top] (\x. g (inverse x)) \ f \[at_right 0] g" by (simp add: at_right_to_top asymp_equiv_filtermap_iff) lemma asymp_equiv_at_top_imp_at_left_0: "(\x. f (-inverse x)) \[at_top] (\x. g (-inverse x)) \ f \[at_left 0] g" by (simp add: at_left_to_top asymp_equiv_filtermap_iff) lemma asymp_equiv_at_top_imp_at_right: "(\x. f (a + inverse x)) \[at_top] (\x. g (a + inverse x)) \ f \[at_right a] g" by (simp add: at_right_to_top' asymp_equiv_filtermap_iff) lemma asymp_equiv_at_top_imp_at_left: "(\x. f (a - inverse x)) \[at_top] (\x. g (a - inverse x)) \ f \[at_left a] g" by (simp add: at_left_to_top' asymp_equiv_filtermap_iff) lemma asymp_equiv_at_top_imp_at: "(\x. f (a - inverse x)) \[at_top] (\x. g (a - inverse x)) \ (\x. f (a + inverse x)) \[at_top] (\x. g (a + inverse x)) \ f \[at a] g" unfolding asymp_equiv_def using asymp_equiv_at_top_imp_at_left[of a] asymp_equiv_at_top_imp_at_right[of a] by (intro filterlim_split_at) (auto simp: asymp_equiv_def) lemma asymp_equiv_at_top_imp_at_0: "(\x. f (-inverse x)) \[at_top] (\x. g (-inverse x)) \ (\x. f (inverse x)) \[at_top] (\x. g (inverse x)) \ f \[at 0] g" by (rule asymp_equiv_at_top_imp_at) auto end lemmas eval_simps = eval_const eval_plus eval_minus eval_uminus eval_times eval_inverse eval_divide eval_map_ms eval_lift_ms eval_real_def eval_ms.simps lemma real_eqI: "a - b = 0 \ a = (b::real)" by simp lemma lift_ms_simps: "lift_ms (MS xs f) = MS (MSLCons (MS xs f, 0) MSLNil) f" "lift_ms (c :: real) = MS (MSLCons (c, 0) MSLNil) (\_. c)" by (simp_all add: lift_ms_def) lemmas landau_reduce_to_top = landau_symbols [THEN landau_symbol_at_top_imp_at_bot] landau_symbols [THEN landau_symbol_at_top_imp_at_left_0] landau_symbols [THEN landau_symbol_at_top_imp_at_right_0] landau_symbols [THEN landau_symbol_at_top_imp_at_left] landau_symbols [THEN landau_symbol_at_top_imp_at_right] asymp_equiv_at_top_imp_at_bot asymp_equiv_at_top_imp_at_left_0 asymp_equiv_at_top_imp_at_right_0 asymp_equiv_at_top_imp_at_left asymp_equiv_at_top_imp_at_right lemmas landau_at_top_imp_at = landau_symbols [THEN landau_symbol_at_top_imp_at_0] landau_symbols [THEN landau_symbol_at_top_imp_at] asymp_equiv_at_top_imp_at_0 asymp_equiv_at_top_imp_at lemma nhds_leI: "x = y \ nhds x \ nhds y" by simp lemma of_nat_diff_real: "of_nat (a - b) = max (0::real) (of_nat a - of_nat b)" and of_nat_max_real: "of_nat (max a b) = max (of_nat a) (of_nat b :: real)" and of_nat_min_real: "of_nat (min a b) = min (of_nat a) (of_nat b :: real)" and of_int_max_real: "of_int (max c d) = max (of_int c) (of_int d :: real)" and of_int_min_real: "of_int (min c d) = min (of_int c) (of_int d :: real)" by simp_all named_theorems real_asymp_nat_reify lemmas [real_asymp_nat_reify] = of_nat_add of_nat_mult of_nat_diff_real of_nat_max_real of_nat_min_real of_nat_power of_nat_Suc of_nat_numeral lemma of_nat_div_real [real_asymp_nat_reify]: "of_nat (a div b) = max 0 (rfloor (of_nat a / of_nat b))" by (simp add: rfloor_def floor_divide_of_nat_eq) lemma of_nat_mod_real [real_asymp_nat_reify]: "of_nat (a mod b) = rnatmod (of_nat a) (of_nat b)" by (simp add: rnatmod_def) lemma real_nat [real_asymp_nat_reify]: "real (nat a) = max 0 (of_int a)" by simp named_theorems real_asymp_int_reify lemmas [real_asymp_int_reify] = of_int_add of_int_mult of_int_diff of_int_minus of_int_max_real of_int_min_real of_int_power of_int_of_nat_eq of_int_numeral lemma of_int_div_real [real_asymp_int_reify]: "of_int (a div b) = rfloor (of_int a / of_int b)" by (simp add: rfloor_def floor_divide_of_int_eq) named_theorems real_asymp_preproc lemmas [real_asymp_preproc] = of_nat_add of_nat_mult of_nat_diff_real of_nat_max_real of_nat_min_real of_nat_power of_nat_Suc of_nat_numeral of_int_add of_int_mult of_int_diff of_int_minus of_int_max_real of_int_min_real of_int_power of_int_of_nat_eq of_int_numeral real_nat lemma of_int_mod_real [real_asymp_int_reify]: "of_int (a mod b) = rintmod (of_int a) (of_int b)" by (simp add: rintmod_def) lemma filterlim_of_nat_at_top: "filterlim f F at_top \ filterlim (\x. f (of_nat x :: real)) F at_top" by (erule filterlim_compose[OF _filterlim_real_sequentially]) lemma asymp_equiv_real_nat_transfer: "f \[at_top] g \ (\x. f (of_nat x :: real)) \[at_top] (\x. g (of_nat x))" unfolding asymp_equiv_def by (rule filterlim_of_nat_at_top) lemma eventually_nat_real: assumes "eventually P (at_top :: real filter)" shows "eventually (\x. P (real x)) (at_top :: nat filter)" using assms filterlim_real_sequentially unfolding filterlim_def le_filter_def eventually_filtermap by auto lemmas real_asymp_nat_intros = filterlim_of_nat_at_top eventually_nat_real smallo_real_nat_transfer bigo_real_nat_transfer smallomega_real_nat_transfer bigomega_real_nat_transfer bigtheta_real_nat_transfer asymp_equiv_real_nat_transfer lemma filterlim_of_int_at_top: "filterlim f F at_top \ filterlim (\x. f (of_int x :: real)) F at_top" by (erule filterlim_compose[OF _ filterlim_real_of_int_at_top]) (* TODO Move *) lemma filterlim_real_of_int_at_bot [tendsto_intros]: "filterlim real_of_int at_bot at_bot" unfolding filterlim_at_bot proof fix C :: real show "eventually (\n. real_of_int n \ C) at_bot" using eventually_le_at_bot[of "\C\"] by eventually_elim linarith qed lemma filterlim_of_int_at_bot: "filterlim f F at_bot \ filterlim (\x. f (of_int x :: real)) F at_bot" by (erule filterlim_compose[OF _ filterlim_real_of_int_at_bot]) -find_theorems of_int at_bot lemma asymp_equiv_real_int_transfer_at_top: "f \[at_top] g \ (\x. f (of_int x :: real)) \[at_top] (\x. g (of_int x))" unfolding asymp_equiv_def by (rule filterlim_of_int_at_top) lemma asymp_equiv_real_int_transfer_at_bot: "f \[at_bot] g \ (\x. f (of_int x :: real)) \[at_bot] (\x. g (of_int x))" unfolding asymp_equiv_def by (rule filterlim_of_int_at_bot) lemma eventually_int_real_at_top: assumes "eventually P (at_top :: real filter)" shows "eventually (\x. P (of_int x)) (at_top :: int filter)" using assms filterlim_of_int_at_top filterlim_iff filterlim_real_of_int_at_top by blast lemma eventually_int_real_at_bot: assumes "eventually P (at_bot :: real filter)" shows "eventually (\x. P (of_int x)) (at_bot :: int filter)" using assms filterlim_of_int_at_bot filterlim_iff filterlim_real_of_int_at_bot by blast lemmas real_asymp_int_intros = filterlim_of_int_at_bot filterlim_of_int_at_top landau_symbols[THEN landau_symbol.compose[OF _ _ filterlim_real_of_int_at_top]] landau_symbols[THEN landau_symbol.compose[OF _ _ filterlim_real_of_int_at_bot]] asymp_equiv_real_int_transfer_at_top asymp_equiv_real_int_transfer_at_bot (* TODO: Move? *) lemma tendsto_discrete_iff: "filterlim f (nhds (c :: 'a :: discrete_topology)) F \ eventually (\x. f x = c) F" by (auto simp: nhds_discrete filterlim_principal) lemma tendsto_of_nat_iff: "filterlim (\n. of_nat (f n)) (nhds (of_nat c :: 'a :: real_normed_div_algebra)) F \ eventually (\n. f n = c) F" proof assume "((\n. of_nat (f n)) \ (of_nat c :: 'a)) F" hence "eventually (\n. \real (f n) - real c\ < 1/2) F" by (force simp: tendsto_iff dist_of_nat dest: spec[of _ "1/2"]) thus "eventually (\n. f n = c) F" by eventually_elim (simp add: abs_if split: if_splits) next assume "eventually (\n. f n = c) F" hence "eventually (\n. of_nat (f n) = (of_nat c :: 'a)) F" by eventually_elim simp thus "filterlim (\n. of_nat (f n)) (nhds (of_nat c :: 'a)) F" by (rule tendsto_eventually) qed lemma tendsto_of_int_iff: "filterlim (\n. of_int (f n)) (nhds (of_int c :: 'a :: real_normed_div_algebra)) F \ eventually (\n. f n = c) F" proof assume "((\n. of_int (f n)) \ (of_int c :: 'a)) F" hence "eventually (\n. \real_of_int (f n) - of_int c\ < 1/2) F" by (force simp: tendsto_iff dist_of_int dest: spec[of _ "1/2"]) thus "eventually (\n. f n = c) F" by eventually_elim (simp add: abs_if split: if_splits) next assume "eventually (\n. f n = c) F" hence "eventually (\n. of_int (f n) = (of_int c :: 'a)) F" by eventually_elim simp thus "filterlim (\n. of_int (f n)) (nhds (of_int c :: 'a)) F" by (rule tendsto_eventually) qed lemma filterlim_at_top_int_iff_filterlim_real: "filterlim f at_top F \ filterlim (\x. real_of_int (f x)) at_top F" (is "?lhs = ?rhs") proof assume ?rhs hence "filterlim (\x. floor (real_of_int (f x))) at_top F" by (intro filterlim_compose[OF filterlim_floor_sequentially]) also have "?this \ ?lhs" by (intro filterlim_cong refl) auto finally show ?lhs . qed (auto intro: filterlim_compose[OF filterlim_real_of_int_at_top]) lemma filterlim_floor_at_bot: "filterlim (floor :: real \ int) at_bot at_bot" proof (subst filterlim_at_bot, rule allI) fix C :: int show "eventually (\x::real. \x\ \ C) at_bot" using eventually_le_at_bot[of "real_of_int C"] by eventually_elim linarith qed lemma filterlim_at_bot_int_iff_filterlim_real: "filterlim f at_bot F \ filterlim (\x. real_of_int (f x)) at_bot F" (is "?lhs = ?rhs") proof assume ?rhs hence "filterlim (\x. floor (real_of_int (f x))) at_bot F" by (intro filterlim_compose[OF filterlim_floor_at_bot]) also have "?this \ ?lhs" by (intro filterlim_cong refl) auto finally show ?lhs . qed (auto intro: filterlim_compose[OF filterlim_real_of_int_at_bot]) (* END TODO *) lemma real_asymp_real_nat_transfer: "filterlim (\n. real (f n)) at_top F \ filterlim f at_top F" "filterlim (\n. real (f n)) (nhds (real c)) F \ eventually (\n. f n = c) F" "eventually (\n. real (f n) \ real (g n)) F \ eventually (\n. f n \ g n) F" "eventually (\n. real (f n) < real (g n)) F \ eventually (\n. f n < g n) F" by (simp_all add: filterlim_sequentially_iff_filterlim_real tendsto_of_nat_iff) lemma real_asymp_real_int_transfer: "filterlim (\n. real_of_int (f n)) at_top F \ filterlim f at_top F" "filterlim (\n. real_of_int (f n)) at_bot F \ filterlim f at_bot F" "filterlim (\n. real_of_int (f n)) (nhds (of_int c)) F \ eventually (\n. f n = c) F" "eventually (\n. real_of_int (f n) \ real_of_int (g n)) F \ eventually (\n. f n \ g n) F" "eventually (\n. real_of_int (f n) < real_of_int (g n)) F \ eventually (\n. f n < g n) F" by (simp_all add: tendsto_of_int_iff filterlim_at_top_int_iff_filterlim_real filterlim_at_bot_int_iff_filterlim_real) lemma eventually_at_left_at_right_imp_at: "eventually P (at_left a) \ eventually P (at_right a) \ eventually P (at (a::real))" by (simp add: eventually_at_split) lemma filtermap_at_left_shift: "filtermap (\x. x - d) (at_left a) = at_left (a - d)" for a d :: "real" by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric]) lemma at_left_to_0: "at_left a = filtermap (\x. x + a) (at_left 0)" for a :: real using filtermap_at_left_shift[of "-a" 0] by simp lemma filterlim_at_leftI: assumes "filterlim (\x. f x - c) (at_left 0) F" shows "filterlim f (at_left (c::real)) F" proof - from assms have "filtermap (\x. f x - c) F \ at_left 0" by (simp add: filterlim_def) hence "filtermap (\x. x + c) (filtermap (\x. f x - c) F) \ filtermap (\x. x + c) (at_left 0)" by (rule filtermap_mono) thus ?thesis by (subst at_left_to_0) (simp_all add: filterlim_def filtermap_filtermap) qed lemma filterlim_at_rightI: assumes "filterlim (\x. f x - c) (at_right 0) F" shows "filterlim f (at_right (c::real)) F" proof - from assms have "filtermap (\x. f x - c) F \ at_right 0" by (simp add: filterlim_def) hence "filtermap (\x. x + c) (filtermap (\x. f x - c) F) \ filtermap (\x. x + c) (at_right 0)" by (rule filtermap_mono) thus ?thesis by (subst at_right_to_0) (simp_all add: filterlim_def filtermap_filtermap) qed lemma filterlim_atI': assumes "filterlim (\x. f x - c) (at 0) F" shows "filterlim f (at (c::real)) F" proof - from assms have "filtermap (\x. f x - c) F \ at 0" by (simp add: filterlim_def) hence "filtermap (\x. x + c) (filtermap (\x. f x - c) F) \ filtermap (\x. x + c) (at 0)" by (rule filtermap_mono) thus ?thesis by (subst at_to_0) (simp_all add: filterlim_def filtermap_filtermap) qed lemma gbinomial_series_aux_altdef: "gbinomial_series_aux False x n acc = MSLCons acc (gbinomial_series_aux False x (n + 1) ((x - n) / (n + 1) * acc))" "gbinomial_series_aux True x n acc = (if acc = 0 then MSLNil else MSLCons acc (gbinomial_series_aux True x (n + 1) ((x - n) / (n + 1) * acc)))" by (subst gbinomial_series_aux.code, simp)+ subsection \Proof method setup\ text \ The following theorem collection is the list of rewrite equations to be used by the lazy evaluation package. If something is missing here, normalisation of terms into head normal form will fail. \ named_theorems real_asymp_eval_eqs lemmas [real_asymp_eval_eqs] = msllist.sel fst_conv snd_conv CMP_BRANCH.simps plus_ms.simps minus_ms_altdef minus_ms_aux_MSLNil_left minus_ms_aux_MSLNil_right minus_ms_aux_MSLCons_eval times_ms.simps uminus_ms.simps divide_ms.simps inverse_ms.simps inverse_ms_aux_MSLCons const_expansion_ms_eval const_expansion_real_def times_ms_aux_MSLNil times_ms_aux_MSLCons_eval scale_shift_ms_aux_MSLCons_eval scale_shift_ms_aux_MSLNil uminus_ms_aux_MSLCons_eval uminus_ms_aux_MSLNil scale_ms_aux_MSLNil scale_ms_aux_MSLCons scale_ms_def shift_ms_aux_MSLNil shift_ms_aux_MSLCons scale_ms_aux'_MSLNil scale_ms_aux'_MSLCons exp_series_stream_def exp_series_stream_aux.code plus_ms_aux_MSLNil plus_ms_aux_MSLCons_eval map_ms_altdef map_ms_aux_MSLCons map_ms_aux_MSLNil lift_ms_simps powser_ms_aux_MSLNil powser_ms_aux_MSLCons ln_series_stream_aux_code gbinomial_series_def gbinomial_series_aux_altdef mssalternate.code powr_expansion_ms.simps powr_expansion_real_def powr_ms_aux_MSLCons power_expansion_ms.simps power_expansion_real_def power_ms_aux_MSLCons powser_ms_aux'_MSSCons sin_ms_aux'_altdef cos_ms_aux'_altdef const_ms_aux_def compare_expansions_MSLCons_eval compare_expansions_real zero_expansion_ms_def zero_expansion_real_def sin_series_stream_def cos_series_stream_def arctan_series_stream_def arctan_ms_aux_def sin_series_stream_aux_code arctan_series_stream_aux_code if_True if_False text \ The following constant and theorem collection exist in order to register constructors for lazy evaluation. All constants marked in such a way will be passed to the lazy evaluation package as free constructors upon which pattern-matching can be done. \ definition REAL_ASYMP_EVAL_CONSTRUCTOR :: "'a \ bool" where [simp]: "REAL_ASYMP_EVAL_CONSTRUCTOR x = True" named_theorems exp_log_eval_constructor lemma exp_log_eval_constructors [exp_log_eval_constructor]: "REAL_ASYMP_EVAL_CONSTRUCTOR MSLNil" "REAL_ASYMP_EVAL_CONSTRUCTOR MSLCons" "REAL_ASYMP_EVAL_CONSTRUCTOR MSSCons" "REAL_ASYMP_EVAL_CONSTRUCTOR LT" "REAL_ASYMP_EVAL_CONSTRUCTOR EQ" "REAL_ASYMP_EVAL_CONSTRUCTOR GT" "REAL_ASYMP_EVAL_CONSTRUCTOR Pair" "REAL_ASYMP_EVAL_CONSTRUCTOR True" "REAL_ASYMP_EVAL_CONSTRUCTOR False" "REAL_ASYMP_EVAL_CONSTRUCTOR MS" by simp_all text \ The following constant exists in order to mark functions as having plug-in support for multiseries expansions (i.\,e.\ this can be used to add support for arbitrary functions later on) \ definition REAL_ASYMP_CUSTOM :: "'a \ bool" where [simp]: "REAL_ASYMP_CUSTOM x = True" lemmas [simp del] = ms.map inverse_ms_aux.simps divide_ms.simps definition expansion_with_remainder_term :: "(real \ real) \ (real \ real) set \ bool" where "expansion_with_remainder_term _ _ = True" notation (output) expansion_with_remainder_term (infixl "+" 10) ML_file \asymptotic_basis.ML\ ML_file \exp_log_expression.ML\ ML_file \expansion_interface.ML\ ML_file \multiseries_expansion.ML\ ML_file \real_asymp.ML\ method_setup real_asymp = \(Scan.lift (Scan.optional (Args.parens (Args.$$$ "verbose") >> K true) false) --| Method.sections Simplifier.simp_modifiers) >> (fn verbose => fn ctxt => SIMPLE_METHOD' (Real_Asymp_Basic.tac verbose ctxt))\ "Semi-automatic decision procedure for asymptotics of exp-log functions" end diff --git a/src/HOL/Word/Word.thy b/src/HOL/Word/Word.thy --- a/src/HOL/Word/Word.thy +++ b/src/HOL/Word/Word.thy @@ -1,5321 +1,5321 @@ (* Title: HOL/Word/Word.thy Author: Jeremy Dawson and Gerwin Klein, NICTA *) section \A type of finite bit strings\ theory Word imports "HOL-Library.Type_Length" "HOL-Library.Boolean_Algebra" "HOL-Library.Bit_Operations" Bits_Int Traditional_Syntax Bit_Comprehension Misc_Typedef begin subsection \Preliminaries\ lemma signed_take_bit_decr_length_iff: \signed_take_bit (LENGTH('a::len) - Suc 0) k = signed_take_bit (LENGTH('a) - Suc 0) l \ take_bit LENGTH('a) k = take_bit LENGTH('a) l\ by (cases \LENGTH('a)\) (simp_all add: signed_take_bit_eq_iff_take_bit_eq) subsection \Fundamentals\ subsubsection \Type definition\ quotient_type (overloaded) 'a word = int / \\k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len) l\ morphisms rep Word by (auto intro!: equivpI reflpI sympI transpI) hide_const (open) rep \ \only for foundational purpose\ hide_const (open) Word \ \only for code generation\ subsubsection \Basic arithmetic\ instantiation word :: (len) comm_ring_1 begin lift_definition zero_word :: \'a word\ is 0 . lift_definition one_word :: \'a word\ is 1 . lift_definition plus_word :: \'a word \ 'a word \ 'a word\ is \(+)\ by (auto simp add: take_bit_eq_mod intro: mod_add_cong) lift_definition minus_word :: \'a word \ 'a word \ 'a word\ is \(-)\ by (auto simp add: take_bit_eq_mod intro: mod_diff_cong) lift_definition uminus_word :: \'a word \ 'a word\ is uminus by (auto simp add: take_bit_eq_mod intro: mod_minus_cong) lift_definition times_word :: \'a word \ 'a word \ 'a word\ is \(*)\ by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) instance by (standard; transfer) (simp_all add: algebra_simps) end context includes lifting_syntax notes power_transfer [transfer_rule] transfer_rule_of_bool [transfer_rule] transfer_rule_numeral [transfer_rule] transfer_rule_of_nat [transfer_rule] transfer_rule_of_int [transfer_rule] begin lemma power_transfer_word [transfer_rule]: \(pcr_word ===> (=) ===> pcr_word) (^) (^)\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) of_bool of_bool\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) numeral numeral\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) int of_nat\ by transfer_prover lemma [transfer_rule]: \((=) ===> pcr_word) (\k. k) of_int\ proof - have \((=) ===> pcr_word) of_int of_int\ by transfer_prover then show ?thesis by (simp add: id_def) qed lemma [transfer_rule]: \(pcr_word ===> (\)) even ((dvd) 2 :: 'a::len word \ bool)\ proof - have even_word_unfold: "even k \ (\l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \ ?Q") for k :: int proof assume ?P then show ?Q by auto next assume ?Q then obtain l where "take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l)" .. then have "even (take_bit LENGTH('a) k)" by simp then show ?P by simp qed show ?thesis by (simp only: even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def]) transfer_prover qed end lemma word_exp_length_eq_0 [simp]: \(2 :: 'a::len word) ^ LENGTH('a) = 0\ by transfer simp lemma exp_eq_zero_iff: \2 ^ n = (0 :: 'a::len word) \ n \ LENGTH('a)\ by transfer simp subsubsection \Basic code generation setup\ context begin qualified lift_definition the_int :: \'a::len word \ int\ is \take_bit LENGTH('a)\ . end lemma [code abstype]: \Word.Word (Word.the_int w) = w\ by transfer simp lemma Word_eq_word_of_int [code_post, simp]: \Word.Word = of_int\ by (rule; transfer) simp quickcheck_generator word constructors: \0 :: 'a::len word\, \numeral :: num \ 'a::len word\ instantiation word :: (len) equal begin lift_definition equal_word :: \'a word \ 'a word \ bool\ is \\k l. take_bit LENGTH('a) k = take_bit LENGTH('a) l\ by simp instance by (standard; transfer) rule end lemma [code]: \HOL.equal v w \ HOL.equal (Word.the_int v) (Word.the_int w)\ by transfer (simp add: equal) lemma [code]: \Word.the_int 0 = 0\ by transfer simp lemma [code]: \Word.the_int 1 = 1\ by transfer simp lemma [code]: \Word.the_int (v + w) = take_bit LENGTH('a) (Word.the_int v + Word.the_int w)\ for v w :: \'a::len word\ by transfer (simp add: take_bit_add) lemma [code]: \Word.the_int (- w) = (let k = Word.the_int w in if w = 0 then 0 else 2 ^ LENGTH('a) - k)\ for w :: \'a::len word\ by transfer (auto simp add: take_bit_eq_mod zmod_zminus1_eq_if) lemma [code]: \Word.the_int (v - w) = take_bit LENGTH('a) (Word.the_int v - Word.the_int w)\ for v w :: \'a::len word\ by transfer (simp add: take_bit_diff) lemma [code]: \Word.the_int (v * w) = take_bit LENGTH('a) (Word.the_int v * Word.the_int w)\ for v w :: \'a::len word\ by transfer (simp add: take_bit_mult) subsubsection \Basic conversions\ abbreviation word_of_nat :: \nat \ 'a::len word\ where \word_of_nat \ of_nat\ abbreviation word_of_int :: \int \ 'a::len word\ where \word_of_int \ of_int\ lemma word_of_nat_eq_iff: \word_of_nat m = (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m = take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma word_of_int_eq_iff: \word_of_int k = (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k = take_bit LENGTH('a) l\ by transfer rule lemma word_of_nat_eq_0_iff [simp]: \word_of_nat n = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd n\ using word_of_nat_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff) lemma word_of_int_eq_0_iff [simp]: \word_of_int k = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd k\ using word_of_int_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff) context semiring_1 begin lift_definition unsigned :: \'b::len word \ 'a\ is \of_nat \ nat \ take_bit LENGTH('b)\ by simp lemma unsigned_0 [simp]: \unsigned 0 = 0\ by transfer simp lemma unsigned_1 [simp]: \unsigned 1 = 1\ by transfer simp lemma unsigned_numeral [simp]: \unsigned (numeral n :: 'b::len word) = of_nat (take_bit LENGTH('b) (numeral n))\ by transfer (simp add: nat_take_bit_eq) lemma unsigned_neg_numeral [simp]: \unsigned (- numeral n :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) (- numeral n)))\ by transfer simp end context semiring_1 begin lemma unsigned_of_nat [simp]: \unsigned (word_of_nat n :: 'b::len word) = of_nat (take_bit LENGTH('b) n)\ by transfer (simp add: nat_eq_iff take_bit_of_nat) lemma unsigned_of_int [simp]: \unsigned (word_of_int k :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) k))\ by transfer simp end context semiring_char_0 begin lemma unsigned_word_eqI: \v = w\ if \unsigned v = unsigned w\ using that by transfer (simp add: eq_nat_nat_iff) lemma word_eq_iff_unsigned: \v = w \ unsigned v = unsigned w\ by (auto intro: unsigned_word_eqI) end context ring_1 begin lift_definition signed :: \'b::len word \ 'a\ is \of_int \ signed_take_bit (LENGTH('b) - Suc 0)\ by (simp flip: signed_take_bit_decr_length_iff) lemma signed_0 [simp]: \signed 0 = 0\ by transfer simp lemma signed_1 [simp]: \signed (1 :: 'b::len word) = (if LENGTH('b) = 1 then - 1 else 1)\ by (transfer fixing: uminus; cases \LENGTH('b)\) (simp_all add: sbintrunc_minus_simps) lemma signed_minus_1 [simp]: \signed (- 1 :: 'b::len word) = - 1\ by (transfer fixing: uminus) simp lemma signed_numeral [simp]: \signed (numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (numeral n))\ by transfer simp lemma signed_neg_numeral [simp]: \signed (- numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (- numeral n))\ by transfer simp lemma signed_of_nat [simp]: \signed (word_of_nat n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) (int n))\ by transfer simp lemma signed_of_int [simp]: \signed (word_of_int n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) n)\ by transfer simp end context ring_char_0 begin lemma signed_word_eqI: \v = w\ if \signed v = signed w\ using that by transfer (simp flip: signed_take_bit_decr_length_iff) lemma word_eq_iff_signed: \v = w \ signed v = signed w\ by (auto intro: signed_word_eqI) end abbreviation unat :: \'a::len word \ nat\ where \unat \ unsigned\ abbreviation uint :: \'a::len word \ int\ where \uint \ unsigned\ abbreviation sint :: \'a::len word \ int\ where \sint \ signed\ abbreviation ucast :: \'a::len word \ 'b::len word\ where \ucast \ unsigned\ abbreviation scast :: \'a::len word \ 'b::len word\ where \scast \ signed\ context includes lifting_syntax begin lemma [transfer_rule]: \(pcr_word ===> (=)) (nat \ take_bit LENGTH('a)) (unat :: 'a::len word \ nat)\ using unsigned.transfer [where ?'a = nat] by simp lemma [transfer_rule]: \(pcr_word ===> (=)) (take_bit LENGTH('a)) (uint :: 'a::len word \ int)\ using unsigned.transfer [where ?'a = int] by (simp add: comp_def) lemma [transfer_rule]: \(pcr_word ===> (=)) (signed_take_bit (LENGTH('a) - Suc 0)) (sint :: 'a::len word \ int)\ using signed.transfer [where ?'a = int] by simp lemma [transfer_rule]: \(pcr_word ===> pcr_word) (take_bit LENGTH('a)) (ucast :: 'a::len word \ 'b::len word)\ proof (rule rel_funI) fix k :: int and w :: \'a word\ assume \pcr_word k w\ then have \w = word_of_int k\ by (simp add: pcr_word_def cr_word_def relcompp_apply) moreover have \pcr_word (take_bit LENGTH('a) k) (ucast (word_of_int k :: 'a word))\ by transfer (simp add: pcr_word_def cr_word_def relcompp_apply) ultimately show \pcr_word (take_bit LENGTH('a) k) (ucast w)\ by simp qed lemma [transfer_rule]: \(pcr_word ===> pcr_word) (signed_take_bit (LENGTH('a) - Suc 0)) (scast :: 'a::len word \ 'b::len word)\ proof (rule rel_funI) fix k :: int and w :: \'a word\ assume \pcr_word k w\ then have \w = word_of_int k\ by (simp add: pcr_word_def cr_word_def relcompp_apply) moreover have \pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast (word_of_int k :: 'a word))\ by transfer (simp add: pcr_word_def cr_word_def relcompp_apply) ultimately show \pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast w)\ by simp qed end lemma of_nat_unat [simp]: \of_nat (unat w) = unsigned w\ by transfer simp lemma of_int_uint [simp]: \of_int (uint w) = unsigned w\ by transfer simp lemma of_int_sint [simp]: \of_int (sint a) = signed a\ by transfer (simp_all add: take_bit_signed_take_bit) lemma nat_uint_eq [simp]: \nat (uint w) = unat w\ by transfer simp text \Aliasses only for code generation\ context begin qualified lift_definition of_int :: \int \ 'a::len word\ is \take_bit LENGTH('a)\ . qualified lift_definition of_nat :: \nat \ 'a::len word\ is \int \ take_bit LENGTH('a)\ . qualified lift_definition the_nat :: \'a::len word \ nat\ is \nat \ take_bit LENGTH('a)\ by simp qualified lift_definition the_signed_int :: \'a::len word \ int\ is \signed_take_bit (LENGTH('a) - Suc 0)\ by (simp add: signed_take_bit_decr_length_iff) qualified lift_definition cast :: \'a::len word \ 'b::len word\ is \take_bit LENGTH('a)\ by simp qualified lift_definition signed_cast :: \'a::len word \ 'b::len word\ is \signed_take_bit (LENGTH('a) - Suc 0)\ by (metis signed_take_bit_decr_length_iff) end lemma [code_abbrev, simp]: \Word.the_int = uint\ by transfer rule lemma [code]: \Word.the_int (Word.of_int k :: 'a::len word) = take_bit LENGTH('a) k\ by transfer simp lemma [code_abbrev, simp]: \Word.of_int = word_of_int\ by (rule; transfer) simp lemma [code]: \Word.the_int (Word.of_nat n :: 'a::len word) = take_bit LENGTH('a) (int n)\ by transfer (simp add: take_bit_of_nat) lemma [code_abbrev, simp]: \Word.of_nat = word_of_nat\ by (rule; transfer) (simp add: take_bit_of_nat) lemma [code]: \Word.the_nat w = nat (Word.the_int w)\ by transfer simp lemma [code_abbrev, simp]: \Word.the_nat = unat\ by (rule; transfer) simp lemma [code]: \Word.the_signed_int w = signed_take_bit (LENGTH('a) - Suc 0) (Word.the_int w)\ for w :: \'a::len word\ by transfer simp lemma [code_abbrev, simp]: \Word.the_signed_int = sint\ by (rule; transfer) simp lemma [code]: \Word.the_int (Word.cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_int w)\ for w :: \'a::len word\ by transfer simp lemma [code_abbrev, simp]: \Word.cast = ucast\ by (rule; transfer) simp lemma [code]: \Word.the_int (Word.signed_cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_signed_int w)\ for w :: \'a::len word\ by transfer simp lemma [code_abbrev, simp]: \Word.signed_cast = scast\ by (rule; transfer) simp lemma [code]: \unsigned w = of_nat (nat (Word.the_int w))\ by transfer simp lemma [code]: \signed w = of_int (Word.the_signed_int w)\ by transfer simp subsubsection \Basic ordering\ instantiation word :: (len) linorder begin lift_definition less_eq_word :: "'a word \ 'a word \ bool" is "\a b. take_bit LENGTH('a) a \ take_bit LENGTH('a) b" by simp lift_definition less_word :: "'a word \ 'a word \ bool" is "\a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b" by simp instance by (standard; transfer) auto end interpretation word_order: ordering_top \(\)\ \(<)\ \- 1 :: 'a::len word\ by (standard; transfer) (simp add: take_bit_eq_mod zmod_minus1) interpretation word_coorder: ordering_top \(\)\ \(>)\ \0 :: 'a::len word\ by (standard; transfer) simp lemma word_of_nat_less_eq_iff: \word_of_nat m \ (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m \ take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma word_of_int_less_eq_iff: \word_of_int k \ (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ by transfer rule lemma word_of_nat_less_iff: \word_of_nat m < (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m < take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma word_of_int_less_iff: \word_of_int k < (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k < take_bit LENGTH('a) l\ by transfer rule lemma word_le_def [code]: "a \ b \ uint a \ uint b" by transfer rule lemma word_less_def [code]: "a < b \ uint a < uint b" by transfer rule lemma word_greater_zero_iff: \a > 0 \ a \ 0\ for a :: \'a::len word\ by transfer (simp add: less_le) lemma of_nat_word_less_eq_iff: \of_nat m \ (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m \ take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma of_nat_word_less_iff: \of_nat m < (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m < take_bit LENGTH('a) n\ by transfer (simp add: take_bit_of_nat) lemma of_int_word_less_eq_iff: \of_int k \ (of_int l :: 'a::len word) \ take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ by transfer rule lemma of_int_word_less_iff: \of_int k < (of_int l :: 'a::len word) \ take_bit LENGTH('a) k < take_bit LENGTH('a) l\ by transfer rule subsection \Bit-wise operations\ instantiation word :: (len) semiring_modulo begin lift_definition divide_word :: \'a word \ 'a word \ 'a word\ is \\a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b\ by simp lift_definition modulo_word :: \'a word \ 'a word \ 'a word\ is \\a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b\ by simp instance proof show "a div b * b + a mod b = a" for a b :: "'a word" proof transfer fix k l :: int define r :: int where "r = 2 ^ LENGTH('a)" then have r: "take_bit LENGTH('a) k = k mod r" for k by (simp add: take_bit_eq_mod) have "k mod r = ((k mod r) div (l mod r) * (l mod r) + (k mod r) mod (l mod r)) mod r" by (simp add: div_mult_mod_eq) also have "... = (((k mod r) div (l mod r) * (l mod r)) mod r + (k mod r) mod (l mod r)) mod r" by (simp add: mod_add_left_eq) also have "... = (((k mod r) div (l mod r) * l) mod r + (k mod r) mod (l mod r)) mod r" by (simp add: mod_mult_right_eq) finally have "k mod r = ((k mod r) div (l mod r) * l + (k mod r) mod (l mod r)) mod r" by (simp add: mod_simps) with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l + take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k" by simp qed qed end instance word :: (len) semiring_parity proof show "\ 2 dvd (1::'a word)" by transfer simp show even_iff_mod_2_eq_0: "2 dvd a \ a mod 2 = 0" for a :: "'a word" by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) show "\ 2 dvd a \ a mod 2 = 1" for a :: "'a word" by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) qed lemma word_bit_induct [case_names zero even odd]: \P a\ if word_zero: \P 0\ and word_even: \\a. P a \ 0 < a \ a < 2 ^ (LENGTH('a) - Suc 0) \ P (2 * a)\ and word_odd: \\a. P a \ a < 2 ^ (LENGTH('a) - Suc 0) \ P (1 + 2 * a)\ for P and a :: \'a::len word\ proof - define m :: nat where \m = LENGTH('a) - Suc 0\ then have l: \LENGTH('a) = Suc m\ by simp define n :: nat where \n = unat a\ then have \n < 2 ^ LENGTH('a)\ by transfer (simp add: take_bit_eq_mod) then have \n < 2 * 2 ^ m\ by (simp add: l) then have \P (of_nat n)\ proof (induction n rule: nat_bit_induct) case zero show ?case by simp (rule word_zero) next case (even n) then have \n < 2 ^ m\ by simp with even.IH have \P (of_nat n)\ by simp moreover from \n < 2 ^ m\ even.hyps have \0 < (of_nat n :: 'a word)\ by (auto simp add: word_greater_zero_iff l) moreover from \n < 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\ using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] by (simp add: l take_bit_eq_mod) ultimately have \P (2 * of_nat n)\ by (rule word_even) then show ?case by simp next case (odd n) then have \Suc n \ 2 ^ m\ by simp with odd.IH have \P (of_nat n)\ by simp moreover from \Suc n \ 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\ using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] by (simp add: l take_bit_eq_mod) ultimately have \P (1 + 2 * of_nat n)\ by (rule word_odd) then show ?case by simp qed moreover have \of_nat (nat (uint a)) = a\ by transfer simp ultimately show ?thesis by (simp add: n_def) qed lemma bit_word_half_eq: \(of_bool b + a * 2) div 2 = a\ if \a < 2 ^ (LENGTH('a) - Suc 0)\ for a :: \'a::len word\ proof (cases \2 \ LENGTH('a::len)\) case False have \of_bool (odd k) < (1 :: int) \ even k\ for k :: int by auto with False that show ?thesis by transfer (simp add: eq_iff) next case True obtain n where length: \LENGTH('a) = Suc n\ by (cases \LENGTH('a)\) simp_all show ?thesis proof (cases b) case False moreover have \a * 2 div 2 = a\ using that proof transfer fix k :: int from length have \k * 2 mod 2 ^ LENGTH('a) = (k mod 2 ^ n) * 2\ by simp moreover assume \take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\ with \LENGTH('a) = Suc n\ have \k mod 2 ^ LENGTH('a) = k mod 2 ^ n\ by (simp add: take_bit_eq_mod divmod_digit_0) ultimately have \take_bit LENGTH('a) (k * 2) = take_bit LENGTH('a) k * 2\ by (simp add: take_bit_eq_mod) with True show \take_bit LENGTH('a) (take_bit LENGTH('a) (k * 2) div take_bit LENGTH('a) 2) = take_bit LENGTH('a) k\ by simp qed ultimately show ?thesis by simp next case True moreover have \(1 + a * 2) div 2 = a\ using that proof transfer fix k :: int from length have \(1 + k * 2) mod 2 ^ LENGTH('a) = 1 + (k mod 2 ^ n) * 2\ using pos_zmod_mult_2 [of \2 ^ n\ k] by (simp add: ac_simps) moreover assume \take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\ with \LENGTH('a) = Suc n\ have \k mod 2 ^ LENGTH('a) = k mod 2 ^ n\ by (simp add: take_bit_eq_mod divmod_digit_0) ultimately have \take_bit LENGTH('a) (1 + k * 2) = 1 + take_bit LENGTH('a) k * 2\ by (simp add: take_bit_eq_mod) with True show \take_bit LENGTH('a) (take_bit LENGTH('a) (1 + k * 2) div take_bit LENGTH('a) 2) = take_bit LENGTH('a) k\ by (auto simp add: take_bit_Suc) qed ultimately show ?thesis by simp qed qed lemma even_mult_exp_div_word_iff: \even (a * 2 ^ m div 2 ^ n) \ \ ( m \ n \ n < LENGTH('a) \ odd (a div 2 ^ (n - m)))\ for a :: \'a::len word\ by transfer (auto simp flip: drop_bit_eq_div simp add: even_drop_bit_iff_not_bit bit_take_bit_iff, simp_all flip: push_bit_eq_mult add: bit_push_bit_iff_int) instantiation word :: (len) semiring_bits begin lift_definition bit_word :: \'a word \ nat \ bool\ is \\k n. n < LENGTH('a) \ bit k n\ proof fix k l :: int and n :: nat assume *: \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ show \n < LENGTH('a) \ bit k n \ n < LENGTH('a) \ bit l n\ proof (cases \n < LENGTH('a)\) case True from * have \bit (take_bit LENGTH('a) k) n \ bit (take_bit LENGTH('a) l) n\ by simp then show ?thesis by (simp add: bit_take_bit_iff) next case False then show ?thesis by simp qed qed instance proof show \P a\ if stable: \\a. a div 2 = a \ P a\ and rec: \\a b. P a \ (of_bool b + 2 * a) div 2 = a \ P (of_bool b + 2 * a)\ for P and a :: \'a word\ proof (induction a rule: word_bit_induct) case zero have \0 div 2 = (0::'a word)\ by transfer simp with stable [of 0] show ?case by simp next case (even a) with rec [of a False] show ?case using bit_word_half_eq [of a False] by (simp add: ac_simps) next case (odd a) with rec [of a True] show ?case using bit_word_half_eq [of a True] by (simp add: ac_simps) qed show \bit a n \ odd (a div 2 ^ n)\ for a :: \'a word\ and n by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit bit_iff_odd_drop_bit) show \0 div a = 0\ for a :: \'a word\ by transfer simp show \a div 1 = a\ for a :: \'a word\ by transfer simp show \a mod b div b = 0\ for a b :: \'a word\ apply transfer apply (simp add: take_bit_eq_mod) apply (subst (3) mod_pos_pos_trivial [of _ \2 ^ LENGTH('a)\]) apply simp_all apply (metis le_less mod_by_0 pos_mod_conj zero_less_numeral zero_less_power) using pos_mod_bound [of \2 ^ LENGTH('a)\] apply simp proof - fix aa :: int and ba :: int have f1: "\i n. (i::int) mod 2 ^ n = 0 \ 0 < i mod 2 ^ n" by (metis le_less take_bit_eq_mod take_bit_nonnegative) have "(0::int) < 2 ^ len_of (TYPE('a)::'a itself) \ ba mod 2 ^ len_of (TYPE('a)::'a itself) \ 0 \ aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)" by (metis (no_types) mod_by_0 unique_euclidean_semiring_numeral_class.pos_mod_bound zero_less_numeral zero_less_power) then show "aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)" using f1 by (meson le_less less_le_trans unique_euclidean_semiring_numeral_class.pos_mod_bound) qed show \(1 + a) div 2 = a div 2\ if \even a\ for a :: \'a word\ using that by transfer (auto dest: le_Suc_ex simp add: mod_2_eq_odd take_bit_Suc elim!: evenE) show \(2 :: 'a word) ^ m div 2 ^ n = of_bool ((2 :: 'a word) ^ m \ 0 \ n \ m) * 2 ^ (m - n)\ for m n :: nat by transfer (simp, simp add: exp_div_exp_eq) show "a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)" for a :: "'a word" and m n :: nat apply transfer apply (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: drop_bit_eq_div) apply (simp add: drop_bit_take_bit) done show "a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n" for a :: "'a word" and m n :: nat by transfer (auto simp flip: take_bit_eq_mod simp add: ac_simps) show \a * 2 ^ m mod 2 ^ n = a mod 2 ^ (n - m) * 2 ^ m\ if \m \ n\ for a :: "'a word" and m n :: nat using that apply transfer apply (auto simp flip: take_bit_eq_mod) apply (auto simp flip: push_bit_eq_mult simp add: push_bit_take_bit split: split_min_lin) done show \a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\ for a :: "'a word" and m n :: nat by transfer (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: take_bit_eq_mod drop_bit_eq_div split: split_min_lin) show \even ((2 ^ m - 1) div (2::'a word) ^ n) \ 2 ^ n = (0::'a word) \ m \ n\ for m n :: nat by transfer (auto simp add: take_bit_of_mask even_mask_div_iff) show \even (a * 2 ^ m div 2 ^ n) \ n < m \ (2::'a word) ^ n = 0 \ m \ n \ even (a div 2 ^ (n - m))\ for a :: \'a word\ and m n :: nat proof transfer show \even (take_bit LENGTH('a) (k * 2 ^ m) div take_bit LENGTH('a) (2 ^ n)) \ n < m \ take_bit LENGTH('a) ((2::int) ^ n) = take_bit LENGTH('a) 0 \ (m \ n \ even (take_bit LENGTH('a) k div take_bit LENGTH('a) (2 ^ (n - m))))\ for m n :: nat and k l :: int by (auto simp flip: take_bit_eq_mod drop_bit_eq_div push_bit_eq_mult simp add: div_push_bit_of_1_eq_drop_bit drop_bit_take_bit drop_bit_push_bit_int [of n m]) qed qed end lemma bit_word_eqI: \a = b\ if \\n. n < LENGTH('a) \ bit a n \ bit b n\ for a b :: \'a::len word\ using that by transfer (auto simp add: nat_less_le bit_eq_iff bit_take_bit_iff) lemma bit_imp_le_length: \n < LENGTH('a)\ if \bit w n\ for w :: \'a::len word\ using that by transfer simp lemma not_bit_length [simp]: \\ bit w LENGTH('a)\ for w :: \'a::len word\ by transfer simp instantiation word :: (len) semiring_bit_shifts begin lift_definition push_bit_word :: \nat \ 'a word \ 'a word\ is push_bit proof - show \take_bit LENGTH('a) (push_bit n k) = take_bit LENGTH('a) (push_bit n l)\ if \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ for k l :: int and n :: nat proof - from that have \take_bit (LENGTH('a) - n) (take_bit LENGTH('a) k) = take_bit (LENGTH('a) - n) (take_bit LENGTH('a) l)\ by simp moreover have \min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n\ by simp ultimately show ?thesis by (simp add: take_bit_push_bit) qed qed lift_definition drop_bit_word :: \nat \ 'a word \ 'a word\ is \\n. drop_bit n \ take_bit LENGTH('a)\ by (simp add: take_bit_eq_mod) lift_definition take_bit_word :: \nat \ 'a word \ 'a word\ is \\n. take_bit (min LENGTH('a) n)\ by (simp add: ac_simps) (simp only: flip: take_bit_take_bit) instance proof show \push_bit n a = a * 2 ^ n\ for n :: nat and a :: \'a word\ by transfer (simp add: push_bit_eq_mult) show \drop_bit n a = a div 2 ^ n\ for n :: nat and a :: \'a word\ by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit) show \take_bit n a = a mod 2 ^ n\ for n :: nat and a :: \'a word\ by transfer (auto simp flip: take_bit_eq_mod) qed end lemma [code]: \Word.the_int (take_bit n w) = (if n < LENGTH('a::len) then take_bit n (Word.the_int w) else Word.the_int w)\ for w :: \'a::len word\ by transfer (simp add: not_le not_less ac_simps min_absorb2) instantiation word :: (len) ring_bit_operations begin lift_definition not_word :: \'a word \ 'a word\ is not by (simp add: take_bit_not_iff) lift_definition and_word :: \'a word \ 'a word \ 'a word\ is \and\ by simp lift_definition or_word :: \'a word \ 'a word \ 'a word\ is or by simp lift_definition xor_word :: \'a word \ 'a word \ 'a word\ is xor by simp lift_definition mask_word :: \nat \ 'a word\ is mask . instance by (standard; transfer) (auto simp add: minus_eq_not_minus_1 mask_eq_exp_minus_1 bit_not_iff bit_and_iff bit_or_iff bit_xor_iff) end lemma [code_abbrev]: \push_bit n 1 = (2 :: 'a::len word) ^ n\ by (fact push_bit_of_1) lemma [code]: \NOT w = Word.of_int (NOT (Word.the_int w))\ for w :: \'a::len word\ by transfer (simp add: take_bit_not_take_bit) lemma [code]: \Word.the_int (v AND w) = Word.the_int v AND Word.the_int w\ by transfer simp lemma [code]: \Word.the_int (v OR w) = Word.the_int v OR Word.the_int w\ by transfer simp lemma [code]: \Word.the_int (v XOR w) = Word.the_int v XOR Word.the_int w\ by transfer simp lemma [code]: \Word.the_int (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ by transfer simp context includes lifting_syntax begin lemma set_bit_word_transfer [transfer_rule]: \((=) ===> pcr_word ===> pcr_word) set_bit set_bit\ by (unfold set_bit_def) transfer_prover lemma unset_bit_word_transfer [transfer_rule]: \((=) ===> pcr_word ===> pcr_word) unset_bit unset_bit\ by (unfold unset_bit_def) transfer_prover lemma flip_bit_word_transfer [transfer_rule]: \((=) ===> pcr_word ===> pcr_word) flip_bit flip_bit\ by (unfold flip_bit_def) transfer_prover lemma signed_take_bit_word_transfer [transfer_rule]: \((=) ===> pcr_word ===> pcr_word) (\n k. signed_take_bit n (take_bit LENGTH('a::len) k)) (signed_take_bit :: nat \ 'a word \ 'a word)\ proof - let ?K = \\n (k :: int). take_bit (min LENGTH('a) n) k OR of_bool (n < LENGTH('a) \ bit k n) * NOT (mask n)\ let ?W = \\n (w :: 'a word). take_bit n w OR of_bool (bit w n) * NOT (mask n)\ have \((=) ===> pcr_word ===> pcr_word) ?K ?W\ by transfer_prover also have \?K = (\n k. signed_take_bit n (take_bit LENGTH('a::len) k))\ by (simp add: fun_eq_iff signed_take_bit_def bit_take_bit_iff ac_simps) also have \?W = signed_take_bit\ by (simp add: fun_eq_iff signed_take_bit_def) finally show ?thesis . qed end subsection \Conversions including casts\ subsubsection \Generic unsigned conversion\ context semiring_bits begin lemma bit_unsigned_iff: \bit (unsigned w) n \ 2 ^ n \ 0 \ bit w n\ for w :: \'b::len word\ by (transfer fixing: bit) (simp add: bit_of_nat_iff bit_nat_iff bit_take_bit_iff) end context semiring_bit_shifts begin lemma unsigned_push_bit_eq: \unsigned (push_bit n w) = take_bit LENGTH('b) (push_bit n (unsigned w))\ for w :: \'b::len word\ proof (rule bit_eqI) fix m assume \2 ^ m \ 0\ show \bit (unsigned (push_bit n w)) m = bit (take_bit LENGTH('b) (push_bit n (unsigned w))) m\ proof (cases \n \ m\) case True with \2 ^ m \ 0\ have \2 ^ (m - n) \ 0\ by (metis (full_types) diff_add exp_add_not_zero_imp) with True show ?thesis by (simp add: bit_unsigned_iff bit_push_bit_iff Parity.bit_push_bit_iff bit_take_bit_iff not_le exp_eq_zero_iff ac_simps) next case False then show ?thesis by (simp add: not_le bit_unsigned_iff bit_push_bit_iff Parity.bit_push_bit_iff bit_take_bit_iff) qed qed lemma unsigned_take_bit_eq: \unsigned (take_bit n w) = take_bit n (unsigned w)\ for w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_take_bit_iff Parity.bit_take_bit_iff) end context semiring_bit_operations begin lemma unsigned_and_eq: \unsigned (v AND w) = unsigned v AND unsigned w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_and_iff Bit_Operations.bit_and_iff) lemma unsigned_or_eq: \unsigned (v OR w) = unsigned v OR unsigned w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_or_iff Bit_Operations.bit_or_iff) lemma unsigned_xor_eq: \unsigned (v XOR w) = unsigned v XOR unsigned w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_xor_iff Bit_Operations.bit_xor_iff) end context ring_bit_operations begin lemma unsigned_not_eq: \unsigned (NOT w) = take_bit LENGTH('b) (NOT (unsigned w))\ for w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_take_bit_iff bit_not_iff Bit_Operations.bit_not_iff exp_eq_zero_iff not_le) end context unique_euclidean_semiring_numeral begin lemma unsigned_greater_eq: \0 \ unsigned w\ for w :: \'b::len word\ by (transfer fixing: less_eq) simp lemma unsigned_less: \unsigned w < 2 ^ LENGTH('b)\ for w :: \'b::len word\ by (transfer fixing: less) simp end context linordered_semidom begin lemma word_less_eq_iff_unsigned: "a \ b \ unsigned a \ unsigned b" by (transfer fixing: less_eq) (simp add: nat_le_eq_zle) lemma word_less_iff_unsigned: "a < b \ unsigned a < unsigned b" by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF take_bit_nonnegative]) end subsubsection \Generic signed conversion\ context ring_bit_operations begin lemma bit_signed_iff: \bit (signed w) n \ 2 ^ n \ 0 \ bit w (min (LENGTH('b) - Suc 0) n)\ for w :: \'b::len word\ by (transfer fixing: bit) (auto simp add: bit_of_int_iff Bit_Operations.bit_signed_take_bit_iff min_def) lemma signed_push_bit_eq: \signed (push_bit n w) = signed_take_bit (LENGTH('b) - Suc 0) (push_bit n (signed w :: 'a))\ for w :: \'b::len word\ proof (rule bit_eqI) fix m assume \2 ^ m \ 0\ define q where \q = LENGTH('b) - Suc 0\ then have *: \LENGTH('b) = Suc q\ by simp show \bit (signed (push_bit n w)) m \ bit (signed_take_bit (LENGTH('b) - Suc 0) (push_bit n (signed w :: 'a))) m\ proof (cases \q \ m\) case True moreover define r where \r = m - q\ ultimately have \m = q + r\ by simp moreover from \m = q + r\ \2 ^ m \ 0\ have \2 ^ q \ 0\ \2 ^ r \ 0\ using exp_add_not_zero_imp_left [of q r] exp_add_not_zero_imp_right [of q r] by simp_all moreover from \2 ^ q \ 0\ have \2 ^ (q - n) \ 0\ by (rule exp_not_zero_imp_exp_diff_not_zero) ultimately show ?thesis by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff min_def * exp_eq_zero_iff le_diff_conv2) next case False then show ?thesis using exp_not_zero_imp_exp_diff_not_zero [of m n] by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff min_def not_le not_less * le_diff_conv2 less_diff_conv2 Parity.exp_eq_0_imp_not_bit exp_eq_0_imp_not_bit exp_eq_zero_iff) qed qed lemma signed_take_bit_eq: \signed (take_bit n w) = (if n < LENGTH('b) then take_bit n (signed w) else signed w)\ for w :: \'b::len word\ by (transfer fixing: take_bit; cases \LENGTH('b)\) (auto simp add: Bit_Operations.signed_take_bit_take_bit Bit_Operations.take_bit_signed_take_bit take_bit_of_int min_def less_Suc_eq) lemma signed_not_eq: \signed (NOT w) = signed_take_bit LENGTH('b) (NOT (signed w))\ for w :: \'b::len word\ proof (rule bit_eqI) fix n assume \2 ^ n \ 0\ define q where \q = LENGTH('b) - Suc 0\ then have *: \LENGTH('b) = Suc q\ by simp show \bit (signed (NOT w)) n \ bit (signed_take_bit LENGTH('b) (NOT (signed w))) n\ proof (cases \q < n\) case True moreover define r where \r = n - Suc q\ ultimately have \n = r + Suc q\ by simp moreover from \2 ^ n \ 0\ \n = r + Suc q\ have \2 ^ Suc q \ 0\ using exp_add_not_zero_imp_right by blast ultimately show ?thesis by (simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def exp_eq_zero_iff) next case False then show ?thesis by (auto simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def exp_eq_zero_iff) qed qed lemma signed_and_eq: \signed (v AND w) = signed v AND signed w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_signed_iff bit_and_iff Bit_Operations.bit_and_iff) lemma signed_or_eq: \signed (v OR w) = signed v OR signed w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_signed_iff bit_or_iff Bit_Operations.bit_or_iff) lemma signed_xor_eq: \signed (v XOR w) = signed v XOR signed w\ for v w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_signed_iff bit_xor_iff Bit_Operations.bit_xor_iff) end subsubsection \More\ lemma sint_greater_eq: \- (2 ^ (LENGTH('a) - Suc 0)) \ sint w\ for w :: \'a::len word\ proof (cases \bit w (LENGTH('a) - Suc 0)\) case True then show ?thesis by transfer (simp add: signed_take_bit_eq_if_negative minus_exp_eq_not_mask or_greater_eq ac_simps) next have *: \- (2 ^ (LENGTH('a) - Suc 0)) \ (0::int)\ by simp case False then show ?thesis by transfer (auto simp add: signed_take_bit_eq intro: order_trans *) qed lemma sint_less: \sint w < 2 ^ (LENGTH('a) - Suc 0)\ for w :: \'a::len word\ by (cases \bit w (LENGTH('a) - Suc 0)\; transfer) (simp_all add: signed_take_bit_eq signed_take_bit_def not_eq_complement mask_eq_exp_minus_1 OR_upper) lemma unat_div_distrib: \unat (v div w) = unat v div unat w\ proof transfer fix k l have \nat (take_bit LENGTH('a) k) div nat (take_bit LENGTH('a) l) \ nat (take_bit LENGTH('a) k)\ by (rule div_le_dividend) also have \nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\ by (simp add: nat_less_iff) finally show \(nat \ take_bit LENGTH('a)) (take_bit LENGTH('a) k div take_bit LENGTH('a) l) = (nat \ take_bit LENGTH('a)) k div (nat \ take_bit LENGTH('a)) l\ by (simp add: nat_take_bit_eq div_int_pos_iff nat_div_distrib take_bit_nat_eq_self_iff) qed lemma unat_mod_distrib: \unat (v mod w) = unat v mod unat w\ proof transfer fix k l have \nat (take_bit LENGTH('a) k) mod nat (take_bit LENGTH('a) l) \ nat (take_bit LENGTH('a) k)\ by (rule mod_less_eq_dividend) also have \nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\ by (simp add: nat_less_iff) finally show \(nat \ take_bit LENGTH('a)) (take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = (nat \ take_bit LENGTH('a)) k mod (nat \ take_bit LENGTH('a)) l\ by (simp add: nat_take_bit_eq mod_int_pos_iff less_le nat_mod_distrib take_bit_nat_eq_self_iff) qed lemma uint_div_distrib: \uint (v div w) = uint v div uint w\ proof - have \int (unat (v div w)) = int (unat v div unat w)\ by (simp add: unat_div_distrib) then show ?thesis by (simp add: of_nat_div) qed lemma uint_mod_distrib: \uint (v mod w) = uint v mod uint w\ proof - have \int (unat (v mod w)) = int (unat v mod unat w)\ by (simp add: unat_mod_distrib) then show ?thesis by (simp add: of_nat_mod) qed context semiring_bit_shifts begin lemma unsigned_ucast_eq: \unsigned (ucast w :: 'c::len word) = take_bit LENGTH('c) (unsigned w)\ for w :: \'b::len word\ by (rule bit_eqI) (simp add: bit_unsigned_iff Word.bit_unsigned_iff bit_take_bit_iff exp_eq_zero_iff not_le) end context ring_bit_operations begin lemma signed_ucast_eq: \signed (ucast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)\ for w :: \'b::len word\ proof (rule bit_eqI) fix n assume \2 ^ n \ 0\ then have \2 ^ (min (LENGTH('c) - Suc 0) n) \ 0\ by (simp add: min_def) (metis (mono_tags) diff_diff_cancel exp_not_zero_imp_exp_diff_not_zero) then show \bit (signed (ucast w :: 'c::len word)) n \ bit (signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)) n\ by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_unsigned_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le) qed lemma signed_scast_eq: \signed (scast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (signed w)\ for w :: \'b::len word\ proof (rule bit_eqI) fix n assume \2 ^ n \ 0\ then have \2 ^ (min (LENGTH('c) - Suc 0) n) \ 0\ by (simp add: min_def) (metis (mono_tags) diff_diff_cancel exp_not_zero_imp_exp_diff_not_zero) then show \bit (signed (scast w :: 'c::len word)) n \ bit (signed_take_bit (LENGTH('c) - Suc 0) (signed w)) n\ by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_signed_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le) qed end lemma uint_nonnegative: "0 \ uint w" by (fact unsigned_greater_eq) lemma uint_bounded: "uint w < 2 ^ LENGTH('a)" for w :: "'a::len word" by (fact unsigned_less) lemma uint_idem: "uint w mod 2 ^ LENGTH('a) = uint w" for w :: "'a::len word" by transfer (simp add: take_bit_eq_mod) lemma word_uint_eqI: "uint a = uint b \ a = b" by (fact unsigned_word_eqI) lemma word_uint_eq_iff: "a = b \ uint a = uint b" by (fact word_eq_iff_unsigned) lemma uint_word_of_int_eq: \uint (word_of_int k :: 'a::len word) = take_bit LENGTH('a) k\ by transfer rule lemma uint_word_of_int: "uint (word_of_int k :: 'a::len word) = k mod 2 ^ LENGTH('a)" by (simp add: uint_word_of_int_eq take_bit_eq_mod) lemma word_of_int_uint: "word_of_int (uint w) = w" by transfer simp lemma word_div_def [code]: "a div b = word_of_int (uint a div uint b)" by transfer rule lemma word_mod_def [code]: "a mod b = word_of_int (uint a mod uint b)" by transfer rule lemma split_word_all: "(\x::'a::len word. PROP P x) \ (\x. PROP P (word_of_int x))" proof fix x :: "'a word" assume "\x. PROP P (word_of_int x)" then have "PROP P (word_of_int (uint x))" . then show "PROP P x" by (simp only: word_of_int_uint) qed lemma sint_uint: \sint w = signed_take_bit (LENGTH('a) - Suc 0) (uint w)\ for w :: \'a::len word\ by (cases \LENGTH('a)\; transfer) (simp_all add: signed_take_bit_take_bit) lemma unat_eq_nat_uint: \unat w = nat (uint w)\ by simp lemma ucast_eq: \ucast w = word_of_int (uint w)\ by transfer simp lemma scast_eq: \scast w = word_of_int (sint w)\ by transfer simp lemma uint_0_eq: \uint 0 = 0\ by (fact unsigned_0) lemma uint_1_eq: \uint 1 = 1\ by (fact unsigned_1) lemma word_m1_wi: "- 1 = word_of_int (- 1)" by simp lemma uint_0_iff: "uint x = 0 \ x = 0" by (auto simp add: unsigned_word_eqI) lemma unat_0_iff: "unat x = 0 \ x = 0" by (auto simp add: unsigned_word_eqI) lemma unat_0: "unat 0 = 0" by (fact unsigned_0) lemma unat_gt_0: "0 < unat x \ x \ 0" by (auto simp: unat_0_iff [symmetric]) lemma ucast_0: "ucast 0 = 0" by (fact unsigned_0) lemma sint_0: "sint 0 = 0" by (fact signed_0) lemma scast_0: "scast 0 = 0" by (fact signed_0) lemma sint_n1: "sint (- 1) = - 1" by (fact signed_minus_1) lemma scast_n1: "scast (- 1) = - 1" by (fact signed_minus_1) lemma uint_1: "uint (1::'a::len word) = 1" by (fact uint_1_eq) lemma unat_1: "unat (1::'a::len word) = 1" by (fact unsigned_1) lemma ucast_1: "ucast (1::'a::len word) = 1" by (fact unsigned_1) instantiation word :: (len) size begin lift_definition size_word :: \'a word \ nat\ is \\_. LENGTH('a)\ .. instance .. end lemma word_size [code]: \size w = LENGTH('a)\ for w :: \'a::len word\ by (fact size_word.rep_eq) lemma word_size_gt_0 [iff]: "0 < size w" for w :: "'a::len word" by (simp add: word_size) lemmas lens_gt_0 = word_size_gt_0 len_gt_0 lemma lens_not_0 [iff]: \size w \ 0\ for w :: \'a::len word\ by auto lift_definition source_size :: \('a::len word \ 'b) \ nat\ is \\_. LENGTH('a)\ . lift_definition target_size :: \('a \ 'b::len word) \ nat\ is \\_. LENGTH('b)\ .. lift_definition is_up :: \('a::len word \ 'b::len word) \ bool\ is \\_. LENGTH('a) \ LENGTH('b)\ .. lift_definition is_down :: \('a::len word \ 'b::len word) \ bool\ is \\_. LENGTH('a) \ LENGTH('b)\ .. lemma is_up_eq: \is_up f \ source_size f \ target_size f\ for f :: \'a::len word \ 'b::len word\ by (simp add: source_size.rep_eq target_size.rep_eq is_up.rep_eq) lemma is_down_eq: \is_down f \ target_size f \ source_size f\ for f :: \'a::len word \ 'b::len word\ by (simp add: source_size.rep_eq target_size.rep_eq is_down.rep_eq) lift_definition word_int_case :: \(int \ 'b) \ 'a::len word \ 'b\ is \\f. f \ take_bit LENGTH('a)\ by simp lemma word_int_case_eq_uint [code]: \word_int_case f w = f (uint w)\ by transfer simp translations "case x of XCONST of_int y \ b" \ "CONST word_int_case (\y. b) x" "case x of (XCONST of_int :: 'a) y \ b" \ "CONST word_int_case (\y. b) x" subsection \Arithmetic operations\ text \Legacy theorems:\ lemma word_add_def [code]: "a + b = word_of_int (uint a + uint b)" by transfer (simp add: take_bit_add) lemma word_sub_wi [code]: "a - b = word_of_int (uint a - uint b)" by transfer (simp add: take_bit_diff) lemma word_mult_def [code]: "a * b = word_of_int (uint a * uint b)" by transfer (simp add: take_bit_eq_mod mod_simps) lemma word_minus_def [code]: "- a = word_of_int (- uint a)" by transfer (simp add: take_bit_minus) lemma word_0_wi: "0 = word_of_int 0" by transfer simp lemma word_1_wi: "1 = word_of_int 1" by transfer simp lift_definition word_succ :: "'a::len word \ 'a word" is "\x. x + 1" by (auto simp add: take_bit_eq_mod intro: mod_add_cong) lift_definition word_pred :: "'a::len word \ 'a word" is "\x. x - 1" by (auto simp add: take_bit_eq_mod intro: mod_diff_cong) lemma word_succ_alt [code]: "word_succ a = word_of_int (uint a + 1)" by transfer (simp add: take_bit_eq_mod mod_simps) lemma word_pred_alt [code]: "word_pred a = word_of_int (uint a - 1)" by transfer (simp add: take_bit_eq_mod mod_simps) lemmas word_arith_wis = word_add_def word_sub_wi word_mult_def word_minus_def word_succ_alt word_pred_alt word_0_wi word_1_wi lemma wi_homs: shows wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)" and wi_hom_sub: "word_of_int a - word_of_int b = word_of_int (a - b)" and wi_hom_mult: "word_of_int a * word_of_int b = word_of_int (a * b)" and wi_hom_neg: "- word_of_int a = word_of_int (- a)" and wi_hom_succ: "word_succ (word_of_int a) = word_of_int (a + 1)" and wi_hom_pred: "word_pred (word_of_int a) = word_of_int (a - 1)" by (transfer, simp)+ lemmas wi_hom_syms = wi_homs [symmetric] lemmas word_of_int_homs = wi_homs word_0_wi word_1_wi lemmas word_of_int_hom_syms = word_of_int_homs [symmetric] lemma double_eq_zero_iff: \2 * a = 0 \ a = 0 \ a = 2 ^ (LENGTH('a) - Suc 0)\ for a :: \'a::len word\ proof - define n where \n = LENGTH('a) - Suc 0\ then have *: \LENGTH('a) = Suc n\ by simp have \a = 0\ if \2 * a = 0\ and \a \ 2 ^ (LENGTH('a) - Suc 0)\ using that by transfer (auto simp add: take_bit_eq_0_iff take_bit_eq_mod *) moreover have \2 ^ LENGTH('a) = (0 :: 'a word)\ by transfer simp then have \2 * 2 ^ (LENGTH('a) - Suc 0) = (0 :: 'a word)\ by (simp add: *) ultimately show ?thesis by auto qed subsection \Ordering\ lift_definition word_sle :: \'a::len word \ 'a word \ bool\ (\(_/ <=s _)\ [50, 51] 50) is \\k l. signed_take_bit (LENGTH('a) - 1) k \ signed_take_bit (LENGTH('a) - 1) l\ by (simp flip: signed_take_bit_decr_length_iff) lemma word_sle_eq [code]: \a <=s b \ sint a \ sint b\ by transfer simp lift_definition word_sless :: \'a::len word \ 'a word \ bool\ (\(_/ [50, 51] 50) is \\k l. signed_take_bit (LENGTH('a) - 1) k < signed_take_bit (LENGTH('a) - 1) l\ by (simp flip: signed_take_bit_decr_length_iff) lemma word_sless_eq: \x x <=s y \ x \ y\ by transfer (simp add: signed_take_bit_decr_length_iff less_le) lemma [code]: \a sint a < sint b\ by transfer simp lemma word_less_alt: "a < b \ uint a < uint b" by (fact word_less_def) lemma signed_linorder: "class.linorder word_sle word_sless" by (standard; transfer) (auto simp add: signed_take_bit_decr_length_iff) interpretation signed: linorder "word_sle" "word_sless" by (rule signed_linorder) lemma word_zero_le [simp]: "0 \ y" for y :: "'a::len word" by transfer simp lemma word_m1_ge [simp] : "word_pred 0 \ y" (* FIXME: delete *) by transfer (simp add: take_bit_minus_one_eq_mask mask_eq_exp_minus_1 bintr_lt2p) lemma word_n1_ge [simp]: "y \ -1" for y :: "'a::len word" by (fact word_order.extremum) lemmas word_not_simps [simp] = word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD] lemma word_gt_0: "0 < y \ 0 \ y" for y :: "'a::len word" by (simp add: less_le) lemmas word_gt_0_no [simp] = word_gt_0 [of "numeral y"] for y lemma word_sless_alt: "a sint a < sint b" by transfer simp lemma word_le_nat_alt: "a \ b \ unat a \ unat b" by transfer (simp add: nat_le_eq_zle) lemma word_less_nat_alt: "a < b \ unat a < unat b" by transfer (auto simp add: less_le [of 0]) lemmas unat_mono = word_less_nat_alt [THEN iffD1] instance word :: (len) wellorder proof fix P :: "'a word \ bool" and a assume *: "(\b. (\a. a < b \ P a) \ P b)" have "wf (measure unat)" .. moreover have "{(a, b :: ('a::len) word). a < b} \ measure unat" by (auto simp add: word_less_nat_alt) ultimately have "wf {(a, b :: ('a::len) word). a < b}" by (rule wf_subset) then show "P a" using * by induction blast qed lemma wi_less: "(word_of_int n < (word_of_int m :: 'a::len word)) = (n mod 2 ^ LENGTH('a) < m mod 2 ^ LENGTH('a))" by transfer (simp add: take_bit_eq_mod) lemma wi_le: "(word_of_int n \ (word_of_int m :: 'a::len word)) = (n mod 2 ^ LENGTH('a) \ m mod 2 ^ LENGTH('a))" by transfer (simp add: take_bit_eq_mod) subsection \Bit-wise operations\ lemma uint_take_bit_eq: \uint (take_bit n w) = take_bit n (uint w)\ by transfer (simp add: ac_simps) lemma take_bit_word_eq_self: \take_bit n w = w\ if \LENGTH('a) \ n\ for w :: \'a::len word\ using that by transfer simp lemma take_bit_length_eq [simp]: \take_bit LENGTH('a) w = w\ for w :: \'a::len word\ by (rule take_bit_word_eq_self) simp lemma bit_word_of_int_iff: \bit (word_of_int k :: 'a::len word) n \ n < LENGTH('a) \ bit k n\ by transfer rule lemma bit_uint_iff: \bit (uint w) n \ n < LENGTH('a) \ bit w n\ for w :: \'a::len word\ by transfer (simp add: bit_take_bit_iff) lemma bit_sint_iff: \bit (sint w) n \ n \ LENGTH('a) \ bit w (LENGTH('a) - 1) \ bit w n\ for w :: \'a::len word\ by transfer (auto simp add: bit_signed_take_bit_iff min_def le_less not_less) lemma bit_word_ucast_iff: \bit (ucast w :: 'b::len word) n \ n < LENGTH('a) \ n < LENGTH('b) \ bit w n\ for w :: \'a::len word\ by transfer (simp add: bit_take_bit_iff ac_simps) lemma bit_word_scast_iff: \bit (scast w :: 'b::len word) n \ n < LENGTH('b) \ (bit w n \ LENGTH('a) \ n \ bit w (LENGTH('a) - Suc 0))\ for w :: \'a::len word\ by transfer (auto simp add: bit_signed_take_bit_iff le_less min_def) lift_definition shiftl1 :: \'a::len word \ 'a word\ is \(*) 2\ by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) lemma shiftl1_eq: \shiftl1 w = word_of_int (2 * uint w)\ by transfer (simp add: take_bit_eq_mod mod_simps) lemma shiftl1_eq_mult_2: \shiftl1 = (*) 2\ by (rule ext, transfer) simp lemma bit_shiftl1_iff: \bit (shiftl1 w) n \ 0 < n \ n < LENGTH('a) \ bit w (n - 1)\ for w :: \'a::len word\ by (simp add: shiftl1_eq_mult_2 bit_double_iff exp_eq_zero_iff not_le) (simp add: ac_simps) lift_definition shiftr1 :: \'a::len word \ 'a word\ \ \shift right as unsigned or as signed, ie logical or arithmetic\ is \\k. take_bit LENGTH('a) k div 2\ by simp lemma shiftr1_eq_div_2: \shiftr1 w = w div 2\ by transfer simp lemma bit_shiftr1_iff: \bit (shiftr1 w) n \ bit w (Suc n)\ by transfer (auto simp flip: bit_Suc simp add: bit_take_bit_iff) lemma shiftr1_eq: \shiftr1 w = word_of_int (uint w div 2)\ by transfer simp instantiation word :: (len) semiring_bit_syntax begin lift_definition test_bit_word :: \'a::len word \ nat \ bool\ is \\k n. n < LENGTH('a) \ bit k n\ proof fix k l :: int and n :: nat assume *: \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ show \n < LENGTH('a) \ bit k n \ n < LENGTH('a) \ bit l n\ proof (cases \n < LENGTH('a)\) case True from * have \bit (take_bit LENGTH('a) k) n \ bit (take_bit LENGTH('a) l) n\ by simp then show ?thesis by (simp add: bit_take_bit_iff) next case False then show ?thesis by simp qed qed lemma test_bit_word_eq: \test_bit = (bit :: 'a word \ _)\ by transfer simp lemma bit_word_iff_drop_bit_and [code]: \bit a n \ drop_bit n a AND 1 = 1\ for a :: \'a::len word\ by (simp add: bit_iff_odd_drop_bit odd_iff_mod_2_eq_one and_one_eq) lemma [code]: \test_bit a n \ drop_bit n a AND 1 = 1\ for a :: \'a::len word\ by (simp add: test_bit_word_eq bit_word_iff_drop_bit_and) lift_definition shiftl_word :: \'a::len word \ nat \ 'a word\ is \\k n. push_bit n k\ proof - show \take_bit LENGTH('a) (push_bit n k) = take_bit LENGTH('a) (push_bit n l)\ if \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ for k l :: int and n :: nat proof - from that have \take_bit (LENGTH('a) - n) (take_bit LENGTH('a) k) = take_bit (LENGTH('a) - n) (take_bit LENGTH('a) l)\ by simp moreover have \min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n\ by simp ultimately show ?thesis by (simp add: take_bit_push_bit) qed qed lemma shiftl_word_eq: \w << n = push_bit n w\ for w :: \'a::len word\ by transfer rule lift_definition shiftr_word :: \'a::len word \ nat \ 'a word\ is \\k n. drop_bit n (take_bit LENGTH('a) k)\ by simp lemma shiftr_word_eq: \w >> n = drop_bit n w\ for w :: \'a::len word\ by transfer simp instance by (standard; transfer) simp_all end lemma shiftl_code [code]: \w << n = w * 2 ^ n\ for w :: \'a::len word\ by transfer (simp add: push_bit_eq_mult) lemma shiftl1_code [code]: \shiftl1 w = w << 1\ by transfer (simp add: push_bit_eq_mult 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 [code]: \Word.the_int (w >> n) = drop_bit n (Word.the_int w)\ by transfer (simp add: drop_bit_take_bit min_def le_less less_diff_conv) lemma shiftr1_code [code]: \shiftr1 w = w >> 1\ by transfer (simp add: drop_bit_Suc) lemma word_test_bit_def: \test_bit a = bit (uint a)\ by transfer (simp add: fun_eq_iff bit_take_bit_iff) 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 \drop_bit n = (((\k::int. k div 2) ^^ n))\ for n by (rule sym, induction n) (simp_all add: fun_eq_iff drop_bit_Suc flip: drop_bit_half) then show ?thesis apply transfer apply simp apply (metis bintrunc_bintrunc rco_bintr) done qed lemma bit_shiftl_word_iff: \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 exp_eq_zero_iff not_le) lemma [code]: \push_bit n w = w << n\ for w :: \'a::len word\ by (simp add: shiftl_word_eq) lemma [code]: \drop_bit n w = w >> n\ for w :: \'a::len word\ by (simp add: shiftr_word_eq) lemma bit_shiftr_word_iff: \bit (w >> m) n \ bit w (m + n)\ for w :: \'a::len word\ by (simp add: shiftr_word_eq bit_drop_bit_eq) lemma word_not_def: "NOT (a::'a::len word) = word_of_int (NOT (uint a))" and word_and_def: "(a::'a word) AND b = word_of_int (uint a AND uint b)" and word_or_def: "(a::'a word) OR b = word_of_int (uint a OR uint b)" and word_xor_def: "(a::'a word) XOR b = word_of_int (uint a XOR uint b)" by (transfer, simp add: take_bit_not_take_bit)+ lift_definition setBit :: \'a::len word \ nat \ 'a word\ is \\k n. set_bit n k\ by (simp add: take_bit_set_bit_eq) lemma set_Bit_eq: \setBit w n = set_bit n w\ by transfer simp lemma bit_setBit_iff: \bit (setBit w m) n \ (m = n \ n < LENGTH('a) \ bit w n)\ for w :: \'a::len word\ by transfer (auto simp add: bit_set_bit_iff) lift_definition clearBit :: \'a::len word \ nat \ 'a word\ is \\k n. unset_bit n k\ by (simp add: take_bit_unset_bit_eq) lemma clear_Bit_eq: \clearBit w n = unset_bit n w\ by transfer simp lemma bit_clearBit_iff: \bit (clearBit w m) n \ m \ n \ bit w n\ for w :: \'a::len word\ by transfer (auto simp add: bit_unset_bit_iff) definition even_word :: \'a::len word \ bool\ where [code_abbrev]: \even_word = even\ lemma even_word_iff [code]: \even_word a \ a AND 1 = 0\ by (simp add: and_one_eq even_iff_mod_2_eq_zero even_word_def) lemma map_bit_range_eq_if_take_bit_eq: \map (bit k) [0.. if \take_bit n k = take_bit n l\ for k l :: int using that proof (induction n arbitrary: k l) case 0 then show ?case by simp next case (Suc n) from Suc.prems have \take_bit n (k div 2) = take_bit n (l div 2)\ by (simp add: take_bit_Suc) then have \map (bit (k div 2)) [0.. by (rule Suc.IH) moreover have \bit (r div 2) = bit r \ Suc\ for r :: int by (simp add: fun_eq_iff bit_Suc) moreover from Suc.prems have \even k \ even l\ by (auto simp add: take_bit_Suc elim!: evenE oddE) arith+ ultimately show ?case by (simp only: map_Suc_upt upt_conv_Cons flip: list.map_comp) simp qed lemma take_bit_word_Bit0_eq [simp]: \take_bit (numeral n) (numeral (num.Bit0 m) :: 'a::len word) = 2 * take_bit (pred_numeral n) (numeral m)\ (is ?P) and take_bit_word_Bit1_eq [simp]: \take_bit (numeral n) (numeral (num.Bit1 m) :: 'a::len word) = 1 + 2 * take_bit (pred_numeral n) (numeral m)\ (is ?Q) and take_bit_word_minus_Bit0_eq [simp]: \take_bit (numeral n) (- numeral (num.Bit0 m) :: 'a::len word) = 2 * take_bit (pred_numeral n) (- numeral m)\ (is ?R) and take_bit_word_minus_Bit1_eq [simp]: \take_bit (numeral n) (- numeral (num.Bit1 m) :: 'a::len word) = 1 + 2 * take_bit (pred_numeral n) (- numeral (Num.inc m))\ (is ?S) proof - define w :: \'a::len word\ where \w = numeral m\ moreover define q :: nat where \q = pred_numeral n\ ultimately have num: \numeral m = w\ \numeral (num.Bit0 m) = 2 * w\ \numeral (num.Bit1 m) = 1 + 2 * w\ \numeral (Num.inc m) = 1 + w\ \pred_numeral n = q\ \numeral n = Suc q\ by (simp_all only: w_def q_def numeral_Bit0 [of m] numeral_Bit1 [of m] ac_simps numeral_inc numeral_eq_Suc flip: mult_2) have even: \take_bit (Suc q) (2 * w) = 2 * take_bit q w\ for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_take_bit_iff bit_double_iff) have odd: \take_bit (Suc q) (1 + 2 * w) = 1 + 2 * take_bit q w\ for w :: \'a::len word\ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_double_iff even_bit_succ_iff) show ?P using even [of w] by (simp add: num) show ?Q using odd [of w] by (simp add: num) show ?R using even [of \- w\] by (simp add: num) show ?S using odd [of \- (1 + w)\] by (simp add: num) qed subsection \Type-definition locale instantiations\ definition uints :: "nat \ int set" \ \the sets of integers representing the words\ where "uints n = range (take_bit n)" definition sints :: "nat \ int set" where "sints n = range (signed_take_bit (n - 1))" lemma uints_num: "uints n = {i. 0 \ i \ i < 2 ^ n}" by (simp add: uints_def range_bintrunc) lemma sints_num: "sints n = {i. - (2 ^ (n - 1)) \ i \ i < 2 ^ (n - 1)}" by (simp add: sints_def range_sbintrunc) definition unats :: "nat \ nat set" where "unats n = {i. i < 2 ^ n}" \ \naturals\ lemma uints_unats: "uints n = int ` unats n" apply (unfold unats_def uints_num) apply safe apply (rule_tac image_eqI) apply (erule_tac nat_0_le [symmetric]) by auto lemma unats_uints: "unats n = nat ` uints n" by (auto simp: uints_unats image_iff) lemma td_ext_uint: "td_ext (uint :: 'a word \ int) word_of_int (uints (LENGTH('a::len))) (\w::int. w mod 2 ^ LENGTH('a))" apply (unfold td_ext_def') apply transfer apply (simp add: uints_num take_bit_eq_mod) done interpretation word_uint: td_ext "uint::'a::len word \ int" word_of_int "uints (LENGTH('a::len))" "\w. w mod 2 ^ LENGTH('a::len)" by (fact td_ext_uint) lemmas td_uint = word_uint.td_thm lemmas int_word_uint = word_uint.eq_norm lemma td_ext_ubin: "td_ext (uint :: 'a word \ int) word_of_int (uints (LENGTH('a::len))) (take_bit (LENGTH('a)))" apply standard apply transfer apply simp done interpretation word_ubin: td_ext "uint::'a::len word \ int" word_of_int "uints (LENGTH('a::len))" "take_bit (LENGTH('a::len))" by (fact td_ext_ubin) lemma td_ext_unat [OF refl]: "n = LENGTH('a::len) \ td_ext (unat :: 'a word \ nat) of_nat (unats n) (\i. i mod 2 ^ n)" apply (standard; transfer) apply (simp_all add: unats_def take_bit_of_nat take_bit_nat_eq_self_iff flip: take_bit_eq_mod) done lemmas unat_of_nat = td_ext_unat [THEN td_ext.eq_norm] interpretation word_unat: td_ext "unat::'a::len word \ nat" of_nat "unats (LENGTH('a::len))" "\i. i mod 2 ^ LENGTH('a::len)" by (rule td_ext_unat) lemmas td_unat = word_unat.td_thm lemmas unat_lt2p [iff] = word_unat.Rep [unfolded unats_def mem_Collect_eq] lemma unat_le: "y \ unat z \ y \ unats (LENGTH('a))" for z :: "'a::len word" apply (unfold unats_def) apply clarsimp apply (rule xtrans, rule unat_lt2p, assumption) done lemma td_ext_sbin: "td_ext (sint :: 'a word \ int) word_of_int (sints (LENGTH('a::len))) (signed_take_bit (LENGTH('a) - 1))" by (standard; transfer) (auto simp add: sints_def) lemma td_ext_sint: "td_ext (sint :: 'a word \ int) word_of_int (sints (LENGTH('a::len))) (\w. (w + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1))" using td_ext_sbin [where ?'a = 'a] by (simp add: no_sbintr_alt2) text \ We do \sint\ before \sbin\, before \sint\ is the user version and interpretations do not produce thm duplicates. I.e. we get the name \word_sint.Rep_eqD\, but not \word_sbin.Req_eqD\, because the latter is the same thm as the former. \ interpretation word_sint: td_ext "sint ::'a::len word \ int" word_of_int "sints (LENGTH('a::len))" "\w. (w + 2^(LENGTH('a::len) - 1)) mod 2^LENGTH('a::len) - 2 ^ (LENGTH('a::len) - 1)" by (rule td_ext_sint) interpretation word_sbin: td_ext "sint ::'a::len word \ int" word_of_int "sints (LENGTH('a::len))" "signed_take_bit (LENGTH('a::len) - 1)" by (rule td_ext_sbin) lemmas int_word_sint = td_ext_sint [THEN td_ext.eq_norm] lemmas td_sint = word_sint.td subsection \More shift operations\ lift_definition sshiftr1 :: \'a::len word \ 'a word\ is \\k. take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - 1) k div 2)\ by (simp flip: signed_take_bit_decr_length_iff) 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) - 1) k))\ by (simp flip: signed_take_bit_decr_length_iff) lift_definition bshiftr1 :: \bool \ 'a::len word \ 'a word\ is \\b k. take_bit LENGTH('a) k div 2 + of_bool b * 2 ^ (LENGTH('a) - Suc 0)\ by (fact arg_cong) lemma sshiftr1_eq: \sshiftr1 w = word_of_int (sint w div 2)\ by transfer simp lemma sshiftr_eq: \w >>> n = (sshiftr1 ^^ n) w\ proof - have *: \(\k::int. take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - Suc 0) k div 2)) ^^ Suc n = take_bit LENGTH('a) \ drop_bit (Suc n) \ signed_take_bit (LENGTH('a) - Suc 0)\ for n apply (induction n) apply (auto simp add: fun_eq_iff drop_bit_Suc) apply (metis (no_types, lifting) Suc_pred funpow_swap1 len_gt_0 sbintrunc_bintrunc sbintrunc_rest) done show ?thesis apply transfer apply simp subgoal for k n apply (cases n) apply (simp_all only: *) apply simp_all done done qed 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 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 [code]: \Word.the_int (w >>> n) = take_bit LENGTH('a) (drop_bit n (Word.the_signed_int w))\ for w :: \'a::len word\ by transfer simp lemma sshift1_code [code]: \sshiftr1 w = w >>> 1\ by transfer (simp add: drop_bit_Suc) subsection \Rotation\ lift_definition word_rotr :: \nat \ 'a::len word \ 'a::len word\ is \\n k. concat_bit (LENGTH('a) - n mod LENGTH('a)) (drop_bit (n mod LENGTH('a)) (take_bit LENGTH('a) k)) (take_bit (n mod LENGTH('a)) k)\ subgoal for n k l apply (simp add: concat_bit_def nat_le_iff less_imp_le take_bit_tightened [of \LENGTH('a)\ k l \n mod LENGTH('a::len)\]) done done lift_definition word_rotl :: \nat \ 'a::len word \ 'a::len word\ is \\n k. concat_bit (n mod LENGTH('a)) (drop_bit (LENGTH('a) - n mod LENGTH('a)) (take_bit LENGTH('a) k)) (take_bit (LENGTH('a) - n mod LENGTH('a)) k)\ subgoal for n k l apply (simp add: concat_bit_def nat_le_iff less_imp_le take_bit_tightened [of \LENGTH('a)\ k l \LENGTH('a) - n mod LENGTH('a::len)\]) done done lift_definition word_roti :: \int \ 'a::len word \ 'a::len word\ is \\r k. concat_bit (LENGTH('a) - nat (r mod int LENGTH('a))) (drop_bit (nat (r mod int LENGTH('a))) (take_bit LENGTH('a) k)) (take_bit (nat (r mod int LENGTH('a))) k)\ subgoal for r k l apply (simp add: concat_bit_def nat_le_iff less_imp_le take_bit_tightened [of \LENGTH('a)\ k l \nat (r mod int LENGTH('a::len))\]) done done lemma word_rotl_eq_word_rotr [code]: \word_rotl n = (word_rotr (LENGTH('a) - n mod LENGTH('a)) :: 'a::len word \ 'a word)\ by (rule ext, cases \n mod LENGTH('a) = 0\; transfer) simp_all lemma word_roti_eq_word_rotr_word_rotl [code]: \word_roti i w = (if i \ 0 then word_rotr (nat i) w else word_rotl (nat (- i)) w)\ proof (cases \i \ 0\) case True moreover define n where \n = nat i\ ultimately have \i = int n\ by simp moreover have \word_roti (int n) = (word_rotr n :: _ \ 'a word)\ by (rule ext, transfer) (simp add: nat_mod_distrib) ultimately show ?thesis by simp next case False moreover define n where \n = nat (- i)\ ultimately have \i = - int n\ \n > 0\ by simp_all moreover have \word_roti (- int n) = (word_rotl n :: _ \ 'a word)\ by (rule ext, transfer) (simp add: zmod_zminus1_eq_if flip: of_nat_mod of_nat_diff) ultimately show ?thesis by simp qed lemma bit_word_rotr_iff: \bit (word_rotr m w) n \ n < LENGTH('a) \ bit w ((n + m) mod LENGTH('a))\ for w :: \'a::len word\ proof transfer fix k :: int and m n :: nat define q where \q = m mod LENGTH('a)\ have \q < LENGTH('a)\ by (simp add: q_def) then have \q \ LENGTH('a)\ by simp have \m mod LENGTH('a) = q\ by (simp add: q_def) moreover have \(n + m) mod LENGTH('a) = (n + q) mod LENGTH('a)\ by (subst mod_add_right_eq [symmetric]) (simp add: \m mod LENGTH('a) = q\) moreover have \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - q) (drop_bit q (take_bit LENGTH('a) k)) (take_bit q k)) n \ n < LENGTH('a) \ bit k ((n + q) mod LENGTH('a))\ using \q < LENGTH('a)\ by (cases \q + n \ LENGTH('a)\) (auto simp add: bit_concat_bit_iff bit_drop_bit_eq bit_take_bit_iff le_mod_geq ac_simps) ultimately show \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - m mod LENGTH('a)) (drop_bit (m mod LENGTH('a)) (take_bit LENGTH('a) k)) (take_bit (m mod LENGTH('a)) k)) n \ n < LENGTH('a) \ (n + m) mod LENGTH('a) < LENGTH('a) \ bit k ((n + m) mod LENGTH('a))\ by simp qed lemma bit_word_rotl_iff: \bit (word_rotl m w) n \ n < LENGTH('a) \ bit w ((n + (LENGTH('a) - m mod LENGTH('a))) mod LENGTH('a))\ for w :: \'a::len word\ by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff) lemma bit_word_roti_iff: \bit (word_roti k w) n \ n < LENGTH('a) \ bit w (nat ((int n + k) mod int LENGTH('a)))\ for w :: \'a::len word\ proof transfer fix k l :: int and n :: nat define m where \m = nat (k mod int LENGTH('a))\ have \m < LENGTH('a)\ by (simp add: nat_less_iff m_def) then have \m \ LENGTH('a)\ by simp have \k mod int LENGTH('a) = int m\ by (simp add: nat_less_iff m_def) moreover have \(int n + k) mod int LENGTH('a) = int ((n + m) mod LENGTH('a))\ by (subst mod_add_right_eq [symmetric]) (simp add: of_nat_mod \k mod int LENGTH('a) = int m\) moreover have \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - m) (drop_bit m (take_bit LENGTH('a) l)) (take_bit m l)) n \ n < LENGTH('a) \ bit l ((n + m) mod LENGTH('a))\ using \m < LENGTH('a)\ by (cases \m + n \ LENGTH('a)\) (auto simp add: bit_concat_bit_iff bit_drop_bit_eq bit_take_bit_iff nat_less_iff not_le not_less ac_simps le_diff_conv le_mod_geq) ultimately show \n < LENGTH('a) \ bit (concat_bit (LENGTH('a) - nat (k mod int LENGTH('a))) (drop_bit (nat (k mod int LENGTH('a))) (take_bit LENGTH('a) l)) (take_bit (nat (k mod int LENGTH('a))) l)) n \ n < LENGTH('a) \ nat ((int n + k) mod int LENGTH('a)) < LENGTH('a) \ bit l (nat ((int n + k) mod int LENGTH('a)))\ by simp qed lemma uint_word_rotr_eq: \uint (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a)) (drop_bit (n mod LENGTH('a)) (uint w)) (uint (take_bit (n mod LENGTH('a)) w))\ for w :: \'a::len word\ apply transfer apply (simp add: concat_bit_def take_bit_drop_bit push_bit_take_bit min_def) using mod_less_divisor not_less apply blast done lemma [code]: \Word.the_int (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a)) (drop_bit (n mod LENGTH('a)) (Word.the_int w)) (Word.the_int (take_bit (n mod LENGTH('a)) w))\ for w :: \'a::len word\ using uint_word_rotr_eq [of n w] by simp subsection \Split and cat operations\ lift_definition word_cat :: \'a::len word \ 'b::len word \ 'c::len word\ is \\k l. concat_bit LENGTH('b) l (take_bit LENGTH('a) k)\ by (simp add: bit_eq_iff bit_concat_bit_iff bit_take_bit_iff) lemma word_cat_eq: \(word_cat v w :: 'c::len word) = push_bit LENGTH('b) (ucast v) + ucast w\ for v :: \'a::len word\ and w :: \'b::len word\ by transfer (simp add: concat_bit_eq ac_simps) lemma word_cat_eq' [code]: \word_cat a b = word_of_int (concat_bit LENGTH('b) (uint b) (uint a))\ for a :: \'a::len word\ and b :: \'b::len word\ by transfer simp lemma bit_word_cat_iff: \bit (word_cat v w :: 'c::len word) n \ n < LENGTH('c) \ (if n < LENGTH('b) then bit w n else bit v (n - LENGTH('b)))\ for v :: \'a::len word\ and w :: \'b::len word\ by transfer (simp add: bit_concat_bit_iff bit_take_bit_iff) definition word_split :: "'a::len word \ 'b::len word \ 'c::len word" where "word_split a = (case bin_split (LENGTH('c)) (uint a) of (u, v) \ (word_of_int u, word_of_int v))" definition word_rcat :: \'a::len word list \ 'b::len word\ where \word_rcat = word_of_int \ horner_sum uint (2 ^ LENGTH('a)) \ rev\ lemma word_rcat_eq: \word_rcat ws = word_of_int (bin_rcat (LENGTH('a::len)) (map uint ws))\ for ws :: \'a::len word list\ apply (simp add: word_rcat_def bin_rcat_def rev_map) apply transfer apply (simp add: horner_sum_foldr foldr_map comp_def) done definition word_rsplit :: "'a::len word \ 'b::len word list" where "word_rsplit w = map word_of_int (bin_rsplit (LENGTH('b)) (LENGTH('a), uint w))" abbreviation (input) max_word :: \'a::len word\ \ \Largest representable machine integer.\ where "max_word \ - 1" subsection \Theorems about typedefs\ lemma sint_sbintrunc': "sint (word_of_int bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) bin" by (auto simp: sint_uint word_ubin.eq_norm sbintrunc_bintrunc_lt) lemma uint_sint: "uint w = take_bit (LENGTH('a)) (sint w)" for w :: "'a::len word" by (auto simp: sint_uint bintrunc_sbintrunc_le) lemma bintr_uint: "LENGTH('a) \ n \ take_bit n (uint w) = uint w" for w :: "'a::len word" apply (subst word_ubin.norm_Rep [symmetric]) apply (simp only: bintrunc_bintrunc_min word_size) apply (simp add: min.absorb2) done lemma wi_bintr: "LENGTH('a::len) \ n \ word_of_int (take_bit n w) = (word_of_int w :: 'a word)" by (auto simp: word_ubin.norm_eq_iff [symmetric] min.absorb1) lemma uints_mod: "uints n = range (\w. w mod 2 ^ n)" by (fact uints_def [unfolded no_bintr_alt1]) lemma word_numeral_alt: "numeral b = word_of_int (numeral b)" by (induct b, simp_all only: numeral.simps word_of_int_homs) declare word_numeral_alt [symmetric, code_abbrev] lemma word_neg_numeral_alt: "- numeral b = word_of_int (- numeral b)" by (simp only: word_numeral_alt wi_hom_neg) declare word_neg_numeral_alt [symmetric, code_abbrev] lemma uint_bintrunc [simp]: "uint (numeral bin :: 'a word) = take_bit (LENGTH('a::len)) (numeral bin)" unfolding word_numeral_alt by (rule word_ubin.eq_norm) lemma uint_bintrunc_neg [simp]: "uint (- numeral bin :: 'a word) = take_bit (LENGTH('a::len)) (- numeral bin)" by (simp only: word_neg_numeral_alt word_ubin.eq_norm) lemma sint_sbintrunc [simp]: "sint (numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (numeral bin)" by (simp only: word_numeral_alt word_sbin.eq_norm) lemma sint_sbintrunc_neg [simp]: "sint (- numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (- numeral bin)" by (simp only: word_neg_numeral_alt word_sbin.eq_norm) lemma unat_bintrunc [simp]: "unat (numeral bin :: 'a::len word) = nat (take_bit (LENGTH('a)) (numeral bin))" by transfer simp lemma unat_bintrunc_neg [simp]: "unat (- numeral bin :: 'a::len word) = nat (take_bit (LENGTH('a)) (- numeral bin))" by transfer simp lemma size_0_eq: "size w = 0 \ v = w" for v w :: "'a::len word" apply (unfold word_size) apply (rule word_uint.Rep_eqD) apply (rule box_equals) defer apply (rule word_ubin.norm_Rep)+ apply simp done lemma uint_ge_0 [iff]: "0 \ uint x" for x :: "'a::len word" using word_uint.Rep [of x] by (simp add: uints_num) lemma uint_lt2p [iff]: "uint x < 2 ^ LENGTH('a)" for x :: "'a::len word" using word_uint.Rep [of x] by (simp add: uints_num) lemma sint_ge: "- (2 ^ (LENGTH('a) - 1)) \ sint x" for x :: "'a::len word" using word_sint.Rep [of x] by (simp add: sints_num) lemma sint_lt: "sint x < 2 ^ (LENGTH('a) - 1)" for x :: "'a::len word" using word_sint.Rep [of x] by (simp add: sints_num) lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" by (simp add: sign_Pls_ge_0) lemma uint_m2p_neg: "uint x - 2 ^ LENGTH('a) < 0" for x :: "'a::len word" by (simp only: diff_less_0_iff_less uint_lt2p) lemma uint_m2p_not_non_neg: "\ 0 \ uint x - 2 ^ LENGTH('a)" for x :: "'a::len word" by (simp only: not_le uint_m2p_neg) lemma lt2p_lem: "LENGTH('a) \ n \ uint w < 2 ^ n" for w :: "'a::len word" by (metis bintr_lt2p bintr_uint) lemma uint_le_0_iff [simp]: "uint x \ 0 \ uint x = 0" by (fact uint_ge_0 [THEN leD, THEN antisym_conv1]) lemma uint_nat: "uint w = int (unat w)" by transfer simp lemma uint_numeral: "uint (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" by (simp only: word_numeral_alt int_word_uint) lemma uint_neg_numeral: "uint (- numeral b :: 'a::len word) = - numeral b mod 2 ^ LENGTH('a)" by (simp only: word_neg_numeral_alt int_word_uint) lemma unat_numeral: "unat (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" by transfer (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) lemma sint_numeral: "sint (numeral b :: 'a::len word) = (numeral b + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)" unfolding word_numeral_alt by (rule int_word_sint) lemma word_of_int_0 [simp, code_post]: "word_of_int 0 = 0" unfolding word_0_wi .. lemma word_of_int_1 [simp, code_post]: "word_of_int 1 = 1" unfolding word_1_wi .. lemma word_of_int_neg_1 [simp]: "word_of_int (- 1) = - 1" by (simp add: wi_hom_syms) lemma word_of_int_numeral [simp] : "(word_of_int (numeral bin) :: 'a::len word) = numeral bin" by (simp only: word_numeral_alt) lemma word_of_int_neg_numeral [simp]: "(word_of_int (- numeral bin) :: 'a::len word) = - numeral bin" by (simp only: word_numeral_alt wi_hom_syms) lemma word_int_case_wi: "word_int_case f (word_of_int i :: 'b word) = f (i mod 2 ^ LENGTH('b::len))" by transfer (simp add: take_bit_eq_mod) lemma word_int_split: "P (word_int_case f x) = (\i. x = (word_of_int i :: 'b::len word) \ 0 \ i \ i < 2 ^ LENGTH('b) \ P (f i))" by transfer (auto simp add: take_bit_eq_mod) lemma word_int_split_asm: "P (word_int_case f x) = (\n. x = (word_of_int n :: 'b::len word) \ 0 \ n \ n < 2 ^ LENGTH('b::len) \ \ P (f n))" by transfer (auto simp add: take_bit_eq_mod) lemmas uint_range' = word_uint.Rep [unfolded uints_num mem_Collect_eq] lemmas sint_range' = word_sint.Rep [unfolded One_nat_def sints_num mem_Collect_eq] lemma uint_range_size: "0 \ uint w \ uint w < 2 ^ size w" unfolding word_size by (rule uint_range') lemma sint_range_size: "- (2 ^ (size w - Suc 0)) \ sint w \ sint w < 2 ^ (size w - Suc 0)" unfolding word_size by (rule sint_range') lemma sint_above_size: "2 ^ (size w - 1) \ x \ sint w < x" for w :: "'a::len word" unfolding word_size by (rule less_le_trans [OF sint_lt]) lemma sint_below_size: "x \ - (2 ^ (size w - 1)) \ x \ sint w" for w :: "'a::len word" unfolding word_size by (rule order_trans [OF _ sint_ge]) subsection \Testing bits\ lemma test_bit_eq_iff: "test_bit u = test_bit v \ u = v" for u v :: "'a::len word" unfolding word_test_bit_def by (simp add: bin_nth_eq_iff) lemma test_bit_size [rule_format] : "w !! n \ n < size w" for w :: "'a::len word" apply (unfold word_test_bit_def) apply (subst word_ubin.norm_Rep [symmetric]) apply (simp only: nth_bintr word_size) apply fast done 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 \ bin_nth (uint w) n" by (simp add: word_test_bit_def word_size nth_bintr [symmetric]) lemmas test_bit_bin = test_bit_bin' [unfolded word_size] lemma bin_nth_uint_imp: "bin_nth (uint w) n \ n < LENGTH('a)" for w :: "'a::len word" apply (rule nth_bintr [THEN iffD1, THEN conjunct1]) apply (subst word_ubin.norm_Rep) apply assumption done lemma bin_nth_sint: "LENGTH('a) \ n \ bin_nth (sint w) n = bin_nth (sint w) (LENGTH('a) - 1)" for w :: "'a::len word" apply (subst word_sbin.norm_Rep [symmetric]) apply (auto simp add: nth_sbintr) done lemmas bintr_num = word_ubin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b lemmas sbintr_num = word_sbin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b lemma num_of_bintr': "take_bit (LENGTH('a::len)) (numeral a :: int) = (numeral b) \ numeral a = (numeral b :: 'a word)" unfolding bintr_num by (erule subst, simp) lemma num_of_sbintr': "signed_take_bit (LENGTH('a::len) - 1) (numeral a :: int) = (numeral b) \ numeral a = (numeral b :: 'a word)" unfolding sbintr_num by (erule subst, simp) lemma num_abs_bintr: "(numeral x :: 'a word) = word_of_int (take_bit (LENGTH('a::len)) (numeral x))" by (simp only: word_ubin.Abs_norm word_numeral_alt) lemma num_abs_sbintr: "(numeral x :: 'a word) = word_of_int (signed_take_bit (LENGTH('a::len) - 1) (numeral x))" by (simp only: word_sbin.Abs_norm word_numeral_alt) text \ \cast\ -- note, no arg for new length, as it's determined by type of result, thus in \cast w = w\, the type means cast to length of \w\! \ lemma bit_ucast_iff: \bit (ucast a :: 'a::len word) n \ n < LENGTH('a::len) \ Parity.bit a n\ by transfer (simp add: bit_take_bit_iff) lemma ucast_id [simp]: "ucast w = w" by transfer simp lemma scast_id [simp]: "scast w = w" by transfer simp 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 ucast_mask_eq: \ucast (mask n :: 'b word) = mask (min LENGTH('b::len) n)\ by (simp add: bit_eq_iff) (auto simp add: bit_mask_iff bit_ucast_iff exp_eq_zero_iff) \ \literal u(s)cast\ lemma ucast_bintr [simp]: "ucast (numeral w :: 'a::len word) = word_of_int (take_bit (LENGTH('a)) (numeral w))" by transfer simp (* TODO: neg_numeral *) lemma scast_sbintr [simp]: "scast (numeral w ::'a::len word) = word_of_int (signed_take_bit (LENGTH('a) - Suc 0) (numeral w))" by transfer simp lemma source_size: "source_size (c::'a::len word \ _) = LENGTH('a)" by transfer simp lemma target_size: "target_size (c::_ \ 'b::len word) = LENGTH('b)" by transfer simp lemma is_down: "is_down c \ LENGTH('b) \ LENGTH('a)" for c :: "'a::len word \ 'b::len word" by transfer simp lemma is_up: "is_up c \ LENGTH('a) \ LENGTH('b)" for c :: "'a::len word \ 'b::len word" by transfer simp lemma is_up_down: \is_up c \ is_down d\ for c :: \'a::len word \ 'b::len word\ and d :: \'b::len word \ 'a::len word\ by transfer simp context fixes dummy_types :: \'a::len \ 'b::len\ begin private abbreviation (input) UCAST :: \'a::len word \ 'b::len word\ where \UCAST == ucast\ private abbreviation (input) SCAST :: \'a::len word \ 'b::len word\ where \SCAST == scast\ lemma down_cast_same: \UCAST = scast\ if \is_down UCAST\ by (rule ext, use that in transfer) (simp add: take_bit_signed_take_bit) lemma sint_up_scast: \sint (SCAST w) = sint w\ if \is_up SCAST\ using that by transfer (simp add: min_def Suc_leI le_diff_iff) lemma uint_up_ucast: \uint (UCAST w) = uint w\ if \is_up UCAST\ using that by transfer (simp add: min_def) lemma ucast_up_ucast: \ucast (UCAST w) = ucast w\ if \is_up UCAST\ using that by transfer (simp add: ac_simps) lemma ucast_up_ucast_id: \ucast (UCAST w) = w\ if \is_up UCAST\ using that by (simp add: ucast_up_ucast) lemma scast_up_scast: \scast (SCAST w) = scast w\ if \is_up SCAST\ using that by transfer (simp add: ac_simps) lemma scast_up_scast_id: \scast (SCAST w) = w\ if \is_up SCAST\ using that by (simp add: scast_up_scast) lemma isduu: \is_up UCAST\ if \is_down d\ for d :: \'b word \ 'a word\ using that is_up_down [of UCAST d] by simp lemma isdus: \is_up SCAST\ if \is_down d\ for d :: \'b word \ 'a word\ using that is_up_down [of SCAST d] by simp lemmas ucast_down_ucast_id = isduu [THEN ucast_up_ucast_id] lemmas scast_down_scast_id = isdus [THEN scast_up_scast_id] lemma up_ucast_surj: \surj (ucast :: 'b word \ 'a word)\ if \is_up UCAST\ by (rule surjI) (use that in \rule ucast_up_ucast_id\) lemma up_scast_surj: \surj (scast :: 'b word \ 'a word)\ if \is_up SCAST\ by (rule surjI) (use that in \rule scast_up_scast_id\) lemma down_ucast_inj: \inj_on UCAST A\ if \is_down (ucast :: 'b word \ 'a word)\ by (rule inj_on_inverseI) (use that in \rule ucast_down_ucast_id\) lemma down_scast_inj: \inj_on SCAST A\ if \is_down (scast :: 'b word \ 'a word)\ by (rule inj_on_inverseI) (use that in \rule scast_down_scast_id\) lemma ucast_down_wi: \UCAST (word_of_int x) = word_of_int x\ if \is_down UCAST\ using that by transfer simp lemma ucast_down_no: \UCAST (numeral bin) = numeral bin\ if \is_down UCAST\ using that by transfer simp end lemmas test_bit_def' = word_test_bit_def [THEN fun_cong] lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def lemma bit_last_iff: \bit w (LENGTH('a) - Suc 0) \ sint w < 0\ (is \?P \ ?Q\) for w :: \'a::len word\ proof - have \?P \ bit (uint w) (LENGTH('a) - Suc 0)\ by (simp add: bit_uint_iff) also have \\ \ ?Q\ by (simp add: sint_uint) finally show ?thesis . qed lemma drop_bit_eq_zero_iff_not_bit_last: \drop_bit (LENGTH('a) - Suc 0) w = 0 \ \ bit w (LENGTH('a) - Suc 0)\ for w :: "'a::len word" apply (cases \LENGTH('a)\) apply simp_all apply (simp add: bit_iff_odd_drop_bit) apply transfer apply (simp add: take_bit_drop_bit) apply (auto simp add: drop_bit_eq_div take_bit_eq_mod min_def) apply (auto elim!: evenE) apply (metis div_exp_eq mod_div_trivial mult.commute nonzero_mult_div_cancel_left power_Suc0_right power_add zero_neq_numeral) done subsection \Word Arithmetic\ lemmas word_div_no [simp] = word_div_def [of "numeral a" "numeral b"] for a b lemmas word_mod_no [simp] = word_mod_def [of "numeral a" "numeral b"] for a b lemmas word_less_no [simp] = word_less_def [of "numeral a" "numeral b"] for a b lemmas word_le_no [simp] = word_le_def [of "numeral a" "numeral b"] for a b lemmas word_sless_no [simp] = word_sless_eq [of "numeral a" "numeral b"] for a b lemmas word_sle_no [simp] = word_sle_eq [of "numeral a" "numeral b"] for a b lemma size_0_same': "size w = 0 \ w = v" for v w :: "'a::len word" by (unfold word_size) simp lemmas size_0_same = size_0_same' [unfolded word_size] lemmas unat_eq_0 = unat_0_iff lemmas unat_eq_zero = unat_0_iff subsection \Transferring goals from words to ints\ lemma word_ths: shows word_succ_p1: "word_succ a = a + 1" and word_pred_m1: "word_pred a = a - 1" and word_pred_succ: "word_pred (word_succ a) = a" and word_succ_pred: "word_succ (word_pred a) = a" and word_mult_succ: "word_succ a * b = b + a * b" by (transfer, simp add: algebra_simps)+ lemma uint_cong: "x = y \ uint x = uint y" by simp lemma uint_word_ariths: fixes a b :: "'a::len word" shows "uint (a + b) = (uint a + uint b) mod 2 ^ LENGTH('a::len)" and "uint (a - b) = (uint a - uint b) mod 2 ^ LENGTH('a)" and "uint (a * b) = uint a * uint b mod 2 ^ LENGTH('a)" and "uint (- a) = - uint a mod 2 ^ LENGTH('a)" and "uint (word_succ a) = (uint a + 1) mod 2 ^ LENGTH('a)" and "uint (word_pred a) = (uint a - 1) mod 2 ^ LENGTH('a)" and "uint (0 :: 'a word) = 0 mod 2 ^ LENGTH('a)" and "uint (1 :: 'a word) = 1 mod 2 ^ LENGTH('a)" by (simp_all only: word_arith_wis uint_word_of_int_eq flip: take_bit_eq_mod) lemma uint_word_arith_bintrs: fixes a b :: "'a::len word" shows "uint (a + b) = take_bit (LENGTH('a)) (uint a + uint b)" and "uint (a - b) = take_bit (LENGTH('a)) (uint a - uint b)" and "uint (a * b) = take_bit (LENGTH('a)) (uint a * uint b)" and "uint (- a) = take_bit (LENGTH('a)) (- uint a)" and "uint (word_succ a) = take_bit (LENGTH('a)) (uint a + 1)" and "uint (word_pred a) = take_bit (LENGTH('a)) (uint a - 1)" and "uint (0 :: 'a word) = take_bit (LENGTH('a)) 0" and "uint (1 :: 'a word) = take_bit (LENGTH('a)) 1" by (simp_all add: uint_word_ariths take_bit_eq_mod) lemma sint_word_ariths: fixes a b :: "'a::len word" shows "sint (a + b) = signed_take_bit (LENGTH('a) - 1) (sint a + sint b)" and "sint (a - b) = signed_take_bit (LENGTH('a) - 1) (sint a - sint b)" and "sint (a * b) = signed_take_bit (LENGTH('a) - 1) (sint a * sint b)" and "sint (- a) = signed_take_bit (LENGTH('a) - 1) (- sint a)" and "sint (word_succ a) = signed_take_bit (LENGTH('a) - 1) (sint a + 1)" and "sint (word_pred a) = signed_take_bit (LENGTH('a) - 1) (sint a - 1)" and "sint (0 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 0" and "sint (1 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 1" apply (simp_all only: word_sbin.inverse_norm [symmetric]) apply (simp_all add: wi_hom_syms) apply transfer apply simp apply transfer apply simp done lemmas uint_div_alt = word_div_def [THEN trans [OF uint_cong int_word_uint]] lemmas uint_mod_alt = word_mod_def [THEN trans [OF uint_cong int_word_uint]] lemma word_pred_0_n1: "word_pred 0 = word_of_int (- 1)" unfolding word_pred_m1 by simp lemma succ_pred_no [simp]: "word_succ (numeral w) = numeral w + 1" "word_pred (numeral w) = numeral w - 1" "word_succ (- numeral w) = - numeral w + 1" "word_pred (- numeral w) = - numeral w - 1" by (simp_all add: word_succ_p1 word_pred_m1) lemma word_sp_01 [simp]: "word_succ (- 1) = 0 \ word_succ 0 = 1 \ word_pred 0 = - 1 \ word_pred 1 = 0" by (simp_all add: word_succ_p1 word_pred_m1) \ \alternative approach to lifting arithmetic equalities\ lemma word_of_int_Ex: "\y. x = word_of_int y" by (rule_tac x="uint x" in exI) simp subsection \Order on fixed-length words\ lift_definition udvd :: \'a::len word \ 'a::len word \ bool\ (infixl \udvd\ 50) is \\k l. take_bit LENGTH('a) k dvd take_bit LENGTH('a) l\ by simp lemma udvd_iff_dvd: \x udvd y \ unat x dvd unat y\ by transfer (simp add: nat_dvd_iff) lemma udvd_iff_dvd_int: \v udvd w \ uint v dvd uint w\ by transfer rule lemma udvdI [intro]: \v udvd w\ if \unat w = unat v * unat u\ proof - from that have \unat v dvd unat w\ .. then show ?thesis by (simp add: udvd_iff_dvd) qed lemma udvdE [elim]: fixes v w :: \'a::len word\ assumes \v udvd w\ obtains u :: \'a word\ where \unat w = unat v * unat u\ proof (cases \v = 0\) case True moreover from True \v udvd w\ have \w = 0\ by transfer simp ultimately show thesis using that by simp next case False then have \unat v > 0\ by (simp add: unat_gt_0) from \v udvd w\ have \unat v dvd unat w\ by (simp add: udvd_iff_dvd) then obtain n where \unat w = unat v * n\ .. moreover have \n < 2 ^ LENGTH('a)\ proof (rule ccontr) assume \\ n < 2 ^ LENGTH('a)\ then have \n \ 2 ^ LENGTH('a)\ by (simp add: not_le) then have \unat v * n \ 2 ^ LENGTH('a)\ using \unat v > 0\ mult_le_mono [of 1 \unat v\ \2 ^ LENGTH('a)\ n] by simp with \unat w = unat v * n\ unat_lt2p [of w] show False by simp qed ultimately have \unat w = unat v * unat (word_of_nat n :: 'a word)\ by (auto simp add: take_bit_nat_eq_self_iff intro: sym) with that show thesis . qed lemma udvd_imp_mod_eq_0: \w mod v = 0\ if \v udvd w\ using that by transfer simp lemma mod_eq_0_imp_udvd [intro?]: \v udvd w\ if \w mod v = 0\ proof - from that have \unat (w mod v) = unat 0\ by simp then have \unat w mod unat v = 0\ by (simp add: unat_mod_distrib) then have \unat v dvd unat w\ .. then show ?thesis by (simp add: udvd_iff_dvd) qed lemma udvd_nat_alt: \a udvd b \ (\n. unat b = n * unat a)\ by (auto simp add: udvd_iff_dvd) lemma udvd_unfold_int: \a udvd b \ (\n\0. uint b = n * uint a)\ apply (auto elim!: dvdE simp add: udvd_iff_dvd) apply (simp only: uint_nat) apply auto using of_nat_0_le_iff apply blast apply (simp only: unat_eq_nat_uint) apply (simp add: nat_mult_distrib) done lemma unat_minus_one: \unat (w - 1) = unat w - 1\ if \w \ 0\ proof - have "0 \ uint w" by (fact uint_nonnegative) moreover from that have "0 \ uint w" by (simp add: uint_0_iff) ultimately have "1 \ uint w" by arith from uint_lt2p [of w] have "uint w - 1 < 2 ^ LENGTH('a)" by arith with \1 \ uint w\ have "(uint w - 1) mod 2 ^ LENGTH('a) = uint w - 1" by (auto intro: mod_pos_pos_trivial) with \1 \ uint w\ have "nat ((uint w - 1) mod 2 ^ LENGTH('a)) = nat (uint w) - 1" by (auto simp del: nat_uint_eq) then show ?thesis by (simp only: unat_eq_nat_uint int_word_uint word_arith_wis mod_diff_right_eq) qed lemma measure_unat: "p \ 0 \ unat (p - 1) < unat p" by (simp add: unat_minus_one) (simp add: unat_0_iff [symmetric]) lemmas uint_add_ge0 [simp] = add_nonneg_nonneg [OF uint_ge_0 uint_ge_0] lemmas uint_mult_ge0 [simp] = mult_nonneg_nonneg [OF uint_ge_0 uint_ge_0] lemma uint_sub_lt2p [simp]: "uint x - uint y < 2 ^ LENGTH('a)" for x :: "'a::len word" and y :: "'b::len word" using uint_ge_0 [of y] uint_lt2p [of x] by arith subsection \Conditions for the addition (etc) of two words to overflow\ lemma uint_add_lem: "(uint x + uint y < 2 ^ LENGTH('a)) = (uint (x + y) = uint x + uint y)" for x y :: "'a::len word" by (metis add.right_neutral add_mono_thms_linordered_semiring(1) mod_pos_pos_trivial of_nat_0_le_iff uint_lt2p uint_nat uint_word_ariths(1)) lemma uint_mult_lem: "(uint x * uint y < 2 ^ LENGTH('a)) = (uint (x * y) = uint x * uint y)" for x y :: "'a::len word" by (metis mod_pos_pos_trivial uint_lt2p uint_mult_ge0 uint_word_ariths(3)) lemma uint_sub_lem: "uint x \ uint y \ uint (x - y) = uint x - uint y" - by (metis (mono_tags, hide_lams) diff_ge_0_iff_ge mod_pos_pos_trivial of_nat_0_le_iff take_bit_eq_mod uint_nat uint_sub_lt2p word_sub_wi word_ubin.eq_norm) find_theorems uint \- _\ + by (metis (mono_tags, hide_lams) diff_ge_0_iff_ge mod_pos_pos_trivial of_nat_0_le_iff take_bit_eq_mod uint_nat uint_sub_lt2p word_sub_wi word_ubin.eq_norm) lemma uint_add_le: "uint (x + y) \ uint x + uint y" unfolding uint_word_ariths by (simp add: zmod_le_nonneg_dividend) lemma uint_sub_ge: "uint (x - y) \ uint x - uint y" unfolding uint_word_ariths by (simp add: int_mod_ge) lemma mod_add_if_z: "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ (x + y) mod z = (if x + y < z then x + y else x + y - z)" for x y z :: int apply (auto simp add: not_less) apply (rule antisym) apply (metis diff_ge_0_iff_ge minus_mod_self2 zmod_le_nonneg_dividend) apply (simp only: flip: minus_mod_self2 [of \x + y\ z]) apply (rule int_mod_ge) apply simp_all done lemma uint_plus_if': "uint (a + b) = (if uint a + uint b < 2 ^ LENGTH('a) then uint a + uint b else uint a + uint b - 2 ^ LENGTH('a))" for a b :: "'a::len word" using mod_add_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) lemma mod_sub_if_z: "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ (x - y) mod z = (if y \ x then x - y else x - y + z)" for x y z :: int apply (auto simp add: not_le) apply (rule antisym) apply (simp only: flip: mod_add_self2 [of \x - y\ z]) apply (rule zmod_le_nonneg_dividend) apply simp apply (metis add.commute add.right_neutral add_le_cancel_left diff_ge_0_iff_ge int_mod_ge le_less le_less_trans mod_add_self1 not_less) done lemma uint_sub_if': "uint (a - b) = (if uint b \ uint a then uint a - uint b else uint a - uint b + 2 ^ LENGTH('a))" for a b :: "'a::len word" using mod_sub_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) subsection \Definition of \uint_arith\\ lemma word_of_int_inverse: "word_of_int r = a \ 0 \ r \ r < 2 ^ LENGTH('a) \ uint a = r" for a :: "'a::len word" apply (erule word_uint.Abs_inverse' [rotated]) apply (simp add: uints_num) done lemma uint_split: "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ P i)" for x :: "'a::len word" by transfer (auto simp add: take_bit_eq_mod) lemma uint_split_asm: "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ \ P i)" for x :: "'a::len word" by auto (metis take_bit_int_eq_self_iff) lemmas uint_splits = uint_split uint_split_asm lemmas uint_arith_simps = word_le_def word_less_alt word_uint.Rep_inject [symmetric] uint_sub_if' uint_plus_if' \ \use this to stop, eg. \2 ^ LENGTH(32)\ being simplified\ lemma power_False_cong: "False \ a ^ b = c ^ d" by auto \ \\uint_arith_tac\: reduce to arithmetic on int, try to solve by arith\ ML \ fun uint_arith_simpset ctxt = ctxt addsimps @{thms uint_arith_simps} delsimps @{thms word_uint.Rep_inject} |> fold Splitter.add_split @{thms if_split_asm} |> fold Simplifier.add_cong @{thms power_False_cong} fun uint_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.arith_tac ctxt n t handle Cooper.COOPER _ => Seq.empty; in [ clarify_tac ctxt 1, full_simp_tac (uint_arith_simpset ctxt) 1, ALLGOALS (full_simp_tac (put_simpset HOL_ss ctxt |> fold Splitter.add_split @{thms uint_splits} |> fold Simplifier.add_cong @{thms power_False_cong})), rewrite_goals_tac ctxt @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN REPEAT (eresolve_tac ctxt [conjE] n) THEN REPEAT (dresolve_tac ctxt @{thms word_of_int_inverse} n THEN assume_tac ctxt n THEN assume_tac ctxt n)), TRYALL arith_tac' ] end fun uint_arith_tac ctxt = SELECT_GOAL (EVERY (uint_arith_tacs ctxt)) \ method_setup uint_arith = \Scan.succeed (SIMPLE_METHOD' o uint_arith_tac)\ "solving word arithmetic via integers and arith" subsection \More on overflows and monotonicity\ lemma no_plus_overflow_uint_size: "x \ x + y \ uint x + uint y < 2 ^ size x" for x y :: "'a::len word" unfolding word_size by uint_arith lemmas no_olen_add = no_plus_overflow_uint_size [unfolded word_size] lemma no_ulen_sub: "x \ x - y \ uint y \ uint x" for x y :: "'a::len word" by uint_arith lemma no_olen_add': "x \ y + x \ uint y + uint x < 2 ^ LENGTH('a)" for x y :: "'a::len word" by (simp add: ac_simps no_olen_add) lemmas olen_add_eqv = trans [OF no_olen_add no_olen_add' [symmetric]] lemmas uint_plus_simple_iff = trans [OF no_olen_add uint_add_lem] lemmas uint_plus_simple = uint_plus_simple_iff [THEN iffD1] lemmas uint_minus_simple_iff = trans [OF no_ulen_sub uint_sub_lem] lemmas uint_minus_simple_alt = uint_sub_lem [folded word_le_def] lemmas word_sub_le_iff = no_ulen_sub [folded word_le_def] lemmas word_sub_le = word_sub_le_iff [THEN iffD2] lemma word_less_sub1: "x \ 0 \ 1 < x \ 0 < x - 1" for x :: "'a::len word" by uint_arith lemma word_le_sub1: "x \ 0 \ 1 \ x \ 0 \ x - 1" for x :: "'a::len word" by uint_arith lemma sub_wrap_lt: "x < x - z \ x < z" for x z :: "'a::len word" by uint_arith lemma sub_wrap: "x \ x - z \ z = 0 \ x < z" for x z :: "'a::len word" by uint_arith lemma plus_minus_not_NULL_ab: "x \ ab - c \ c \ ab \ c \ 0 \ x + c \ 0" for x ab c :: "'a::len word" by uint_arith lemma plus_minus_no_overflow_ab: "x \ ab - c \ c \ ab \ x \ x + c" for x ab c :: "'a::len word" by uint_arith lemma le_minus': "a + c \ b \ a \ a + c \ c \ b - a" for a b c :: "'a::len word" by uint_arith lemma le_plus': "a \ b \ c \ b - a \ a + c \ b" for a b c :: "'a::len word" by uint_arith lemmas le_plus = le_plus' [rotated] lemmas le_minus = leD [THEN thin_rl, THEN le_minus'] (* FIXME *) lemma word_plus_mono_right: "y \ z \ x \ x + z \ x + y \ x + z" for x y z :: "'a::len word" by uint_arith lemma word_less_minus_cancel: "y - x < z - x \ x \ z \ y < z" for x y z :: "'a::len word" by uint_arith lemma word_less_minus_mono_left: "y < z \ x \ y \ y - x < z - x" for x y z :: "'a::len word" by uint_arith lemma word_less_minus_mono: "a < c \ d < b \ a - b < a \ c - d < c \ a - b < c - d" for a b c d :: "'a::len word" by uint_arith lemma word_le_minus_cancel: "y - x \ z - x \ x \ z \ y \ z" for x y z :: "'a::len word" by uint_arith lemma word_le_minus_mono_left: "y \ z \ x \ y \ y - x \ z - x" for x y z :: "'a::len word" by uint_arith lemma word_le_minus_mono: "a \ c \ d \ b \ a - b \ a \ c - d \ c \ a - b \ c - d" for a b c d :: "'a::len word" by uint_arith lemma plus_le_left_cancel_wrap: "x + y' < x \ x + y < x \ x + y' < x + y \ y' < y" for x y y' :: "'a::len word" by uint_arith lemma plus_le_left_cancel_nowrap: "x \ x + y' \ x \ x + y \ x + y' < x + y \ y' < y" for x y y' :: "'a::len word" by uint_arith lemma word_plus_mono_right2: "a \ a + b \ c \ b \ a \ a + c" for a b c :: "'a::len word" by uint_arith lemma word_less_add_right: "x < y - z \ z \ y \ x + z < y" for x y z :: "'a::len word" by uint_arith lemma word_less_sub_right: "x < y + z \ y \ x \ x - y < z" for x y z :: "'a::len word" by uint_arith lemma word_le_plus_either: "x \ y \ x \ z \ y \ y + z \ x \ y + z" for x y z :: "'a::len word" by uint_arith lemma word_less_nowrapI: "x < z - k \ k \ z \ 0 < k \ x < x + k" for x z k :: "'a::len word" by uint_arith lemma inc_le: "i < m \ i + 1 \ m" for i m :: "'a::len word" by uint_arith lemma inc_i: "1 \ i \ i < m \ 1 \ i + 1 \ i + 1 \ m" for i m :: "'a::len word" by uint_arith lemma udvd_incr_lem: "up < uq \ up = ua + n * uint K \ uq = ua + n' * uint K \ up + uint K \ uq" by auto (metis int_distrib(1) linorder_not_less mult.left_neutral mult_right_mono uint_nonnegative zless_imp_add1_zle) lemma udvd_incr': "p < q \ uint p = ua + n * uint K \ uint q = ua + n' * uint K \ p + K \ q" apply (unfold word_less_alt word_le_def) apply (drule (2) udvd_incr_lem) apply (erule uint_add_le [THEN order_trans]) done lemma udvd_decr': "p < q \ uint p = ua + n * uint K \ uint q = ua + n' * uint K \ p \ q - K" apply (unfold word_less_alt word_le_def) apply (drule (2) udvd_incr_lem) apply (drule le_diff_eq [THEN iffD2]) apply (erule order_trans) apply (rule uint_sub_ge) done lemmas udvd_incr_lem0 = udvd_incr_lem [where ua=0, unfolded add_0_left] lemmas udvd_incr0 = udvd_incr' [where ua=0, unfolded add_0_left] lemmas udvd_decr0 = udvd_decr' [where ua=0, unfolded add_0_left] lemma udvd_minus_le': "xy < k \ z udvd xy \ z udvd k \ xy \ k - z" apply (unfold udvd_unfold_int) apply clarify apply (erule (2) udvd_decr0) done lemma udvd_incr2_K: "p < a + s \ a \ a + s \ K udvd s \ K udvd p - a \ a \ p \ 0 < K \ p \ p + K \ p + K \ a + s" supply [[simproc del: linordered_ring_less_cancel_factor]] apply (unfold udvd_unfold_int) apply clarify apply (simp add: uint_arith_simps split: if_split_asm) prefer 2 apply (insert uint_range' [of s])[1] apply arith apply (drule add.commute [THEN xtrans(1)]) apply (simp flip: diff_less_eq) apply (subst (asm) mult_less_cancel_right) apply simp apply (simp add: diff_eq_eq not_less) apply (subst (asm) (3) zless_iff_Suc_zadd) apply auto apply (auto simp add: algebra_simps) apply (drule less_le_trans [of _ \2 ^ LENGTH('a)\]) apply assumption apply (simp add: mult_less_0_iff) done subsection \Arithmetic type class instantiations\ lemmas word_le_0_iff [simp] = word_zero_le [THEN leD, THEN antisym_conv1] lemma word_of_int_nat: "0 \ x \ word_of_int x = of_nat (nat x)" by simp text \ note that \iszero_def\ is only for class \comm_semiring_1_cancel\, which requires word length \\ 1\, ie \'a::len word\ \ lemma iszero_word_no [simp]: "iszero (numeral bin :: 'a::len word) = iszero (take_bit LENGTH('a) (numeral bin :: int))" using word_ubin.norm_eq_iff [where 'a='a, of "numeral bin" 0] by (simp add: iszero_def [symmetric]) text \Use \iszero\ to simplify equalities between word numerals.\ lemmas word_eq_numeral_iff_iszero [simp] = eq_numeral_iff_iszero [where 'a="'a::len word"] subsection \Word and nat\ lemma word_nchotomy: "\w :: 'a::len word. \n. w = of_nat n \ n < 2 ^ LENGTH('a)" apply (rule allI) apply (rule word_unat.Abs_cases) apply (unfold unats_def) apply auto done lemma of_nat_eq: "of_nat n = w \ (\q. n = unat w + q * 2 ^ LENGTH('a))" for w :: "'a::len word" using mod_div_mult_eq [of n "2 ^ LENGTH('a)", symmetric] by (auto simp add: word_unat.inverse_norm) lemma of_nat_eq_size: "of_nat n = w \ (\q. n = unat w + q * 2 ^ size w)" unfolding word_size by (rule of_nat_eq) lemma of_nat_0: "of_nat m = (0::'a::len word) \ (\q. m = q * 2 ^ LENGTH('a))" by (simp add: of_nat_eq) lemma of_nat_2p [simp]: "of_nat (2 ^ LENGTH('a)) = (0::'a::len word)" by (fact mult_1 [symmetric, THEN iffD2 [OF of_nat_0 exI]]) lemma of_nat_gt_0: "of_nat k \ 0 \ 0 < k" by (cases k) auto lemma of_nat_neq_0: "0 < k \ k < 2 ^ LENGTH('a::len) \ of_nat k \ (0 :: 'a word)" by (auto simp add : of_nat_0) lemma Abs_fnat_hom_add: "of_nat a + of_nat b = of_nat (a + b)" by simp lemma Abs_fnat_hom_mult: "of_nat a * of_nat b = (of_nat (a * b) :: 'a::len word)" by (simp add: wi_hom_mult) lemma Abs_fnat_hom_Suc: "word_succ (of_nat a) = of_nat (Suc a)" by transfer (simp add: ac_simps) lemma Abs_fnat_hom_0: "(0::'a::len word) = of_nat 0" by simp lemma Abs_fnat_hom_1: "(1::'a::len word) = of_nat (Suc 0)" by simp lemmas Abs_fnat_homs = Abs_fnat_hom_add Abs_fnat_hom_mult Abs_fnat_hom_Suc Abs_fnat_hom_0 Abs_fnat_hom_1 lemma word_arith_nat_add: "a + b = of_nat (unat a + unat b)" by simp lemma word_arith_nat_mult: "a * b = of_nat (unat a * unat b)" by simp lemma word_arith_nat_Suc: "word_succ a = of_nat (Suc (unat a))" by (subst Abs_fnat_hom_Suc [symmetric]) simp lemma word_arith_nat_div: "a div b = of_nat (unat a div unat b)" by (metis of_int_of_nat_eq of_nat_unat of_nat_div word_div_def) lemma word_arith_nat_mod: "a mod b = of_nat (unat a mod unat b)" by (metis of_int_of_nat_eq of_nat_mod of_nat_unat word_mod_def) lemmas word_arith_nat_defs = word_arith_nat_add word_arith_nat_mult word_arith_nat_Suc Abs_fnat_hom_0 Abs_fnat_hom_1 word_arith_nat_div word_arith_nat_mod lemma unat_cong: "x = y \ unat x = unat y" by simp lemmas unat_word_ariths = word_arith_nat_defs [THEN trans [OF unat_cong unat_of_nat]] lemmas word_sub_less_iff = word_sub_le_iff [unfolded linorder_not_less [symmetric] Not_eq_iff] lemma unat_add_lem: "unat x + unat y < 2 ^ LENGTH('a) \ unat (x + y) = unat x + unat y" for x y :: "'a::len word" apply (auto simp: unat_word_ariths) apply (metis unat_lt2p word_unat.eq_norm) done lemma unat_mult_lem: "unat x * unat y < 2 ^ LENGTH('a) \ unat (x * y) = unat x * unat y" for x y :: "'a::len word" apply (auto simp: unat_word_ariths) apply (metis unat_lt2p word_unat.eq_norm) done lemma unat_plus_if': \unat (a + b) = (if unat a + unat b < 2 ^ LENGTH('a) then unat a + unat b else unat a + unat b - 2 ^ LENGTH('a))\ for a b :: \'a::len word\ apply (auto simp: unat_word_ariths not_less) apply (subst (asm) le_iff_add) apply auto apply (metis add_less_cancel_left add_less_cancel_right le_less_trans less_imp_le mod_less unat_lt2p) done lemma le_no_overflow: "x \ b \ a \ a + b \ x \ a + b" for a b x :: "'a::len word" apply (erule order_trans) apply (erule olen_add_eqv [THEN iffD1]) done lemmas un_ui_le = trans [OF word_le_nat_alt [symmetric] word_le_def] lemma unat_sub_if_size: "unat (x - y) = (if unat y \ unat x then unat x - unat y else unat x + 2 ^ size x - unat y)" supply nat_uint_eq [simp del] apply (unfold word_size) apply (simp add: un_ui_le) apply (auto simp add: unat_eq_nat_uint uint_sub_if') apply (rule nat_diff_distrib) prefer 3 apply (simp add: algebra_simps) apply (rule nat_diff_distrib [THEN trans]) prefer 3 apply (subst nat_add_distrib) prefer 3 apply (simp add: nat_power_eq) apply auto apply uint_arith done lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size] lemma uint_div: \uint (x div y) = uint x div uint y\ by (fact uint_div_distrib) lemma unat_div: \unat (x div y) = unat x div unat y\ by (fact unat_div_distrib) lemma uint_mod: \uint (x mod y) = uint x mod uint y\ by (fact uint_mod_distrib) lemma unat_mod: \unat (x mod y) = unat x mod unat y\ by (fact unat_mod_distrib) text \Definition of \unat_arith\ tactic\ lemma unat_split: "P (unat x) \ (\n. of_nat n = x \ n < 2^LENGTH('a) \ P n)" for x :: "'a::len word" by auto (metis take_bit_nat_eq_self_iff) lemma unat_split_asm: "P (unat x) \ (\n. of_nat n = x \ n < 2^LENGTH('a) \ \ P n)" for x :: "'a::len word" by auto (metis take_bit_nat_eq_self_iff) lemmas of_nat_inverse = word_unat.Abs_inverse' [rotated, unfolded unats_def, simplified] lemmas unat_splits = unat_split unat_split_asm lemmas unat_arith_simps = word_le_nat_alt word_less_nat_alt word_unat.Rep_inject [symmetric] unat_sub_if' unat_plus_if' unat_div unat_mod \ \\unat_arith_tac\: tactic to reduce word arithmetic to \nat\, try to solve via \arith\\ ML \ fun unat_arith_simpset ctxt = ctxt addsimps @{thms unat_arith_simps} delsimps @{thms word_unat.Rep_inject} |> fold Splitter.add_split @{thms if_split_asm} |> fold Simplifier.add_cong @{thms power_False_cong} fun unat_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.arith_tac ctxt n t handle Cooper.COOPER _ => Seq.empty; in [ clarify_tac ctxt 1, full_simp_tac (unat_arith_simpset ctxt) 1, ALLGOALS (full_simp_tac (put_simpset HOL_ss ctxt |> fold Splitter.add_split @{thms unat_splits} |> fold Simplifier.add_cong @{thms power_False_cong})), rewrite_goals_tac ctxt @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN REPEAT (eresolve_tac ctxt [conjE] n) THEN REPEAT (dresolve_tac ctxt @{thms of_nat_inverse} n THEN assume_tac ctxt n)), TRYALL arith_tac' ] end fun unat_arith_tac ctxt = SELECT_GOAL (EVERY (unat_arith_tacs ctxt)) \ method_setup unat_arith = \Scan.succeed (SIMPLE_METHOD' o unat_arith_tac)\ "solving word arithmetic via natural numbers and arith" lemma no_plus_overflow_unat_size: "x \ x + y \ unat x + unat y < 2 ^ size x" for x y :: "'a::len word" unfolding word_size by unat_arith lemmas no_olen_add_nat = no_plus_overflow_unat_size [unfolded word_size] lemmas unat_plus_simple = trans [OF no_olen_add_nat unat_add_lem] lemma word_div_mult: "0 < y \ unat x * unat y < 2 ^ LENGTH('a) \ x * y div y = x" for x y :: "'a::len word" apply unat_arith apply clarsimp apply (subst unat_mult_lem [THEN iffD1]) apply auto done lemma div_lt': "i \ k div x \ unat i * unat x < 2 ^ LENGTH('a)" for i k x :: "'a::len word" apply unat_arith apply clarsimp apply (drule mult_le_mono1) apply (erule order_le_less_trans) apply (metis add_lessD1 div_mult_mod_eq unat_lt2p) done lemmas div_lt'' = order_less_imp_le [THEN div_lt'] lemma div_lt_mult: "i < k div x \ 0 < x \ i * x < k" for i k x :: "'a::len word" apply (frule div_lt'' [THEN unat_mult_lem [THEN iffD1]]) apply (simp add: unat_arith_simps) apply (drule (1) mult_less_mono1) apply (erule order_less_le_trans) apply auto done lemma div_le_mult: "i \ k div x \ 0 < x \ i * x \ k" for i k x :: "'a::len word" apply (frule div_lt' [THEN unat_mult_lem [THEN iffD1]]) apply (simp add: unat_arith_simps) apply (drule mult_le_mono1) apply (erule order_trans) apply auto done lemma div_lt_uint': "i \ k div x \ uint i * uint x < 2 ^ LENGTH('a)" for i k x :: "'a::len word" apply (unfold uint_nat) apply (drule div_lt') apply (metis of_nat_less_iff of_nat_mult of_nat_numeral of_nat_power) done lemmas div_lt_uint'' = order_less_imp_le [THEN div_lt_uint'] lemma word_le_exists': "x \ y \ \z. y = x + z \ uint x + uint z < 2 ^ LENGTH('a)" for x y z :: "'a::len word" by (metis add_diff_cancel_left' add_diff_eq uint_add_lem uint_plus_simple) lemmas plus_minus_not_NULL = order_less_imp_le [THEN plus_minus_not_NULL_ab] lemmas plus_minus_no_overflow = order_less_imp_le [THEN plus_minus_no_overflow_ab] lemmas mcs = word_less_minus_cancel word_less_minus_mono_left word_le_minus_cancel word_le_minus_mono_left lemmas word_l_diffs = mcs [where y = "w + x", unfolded add_diff_cancel] for w x lemmas word_diff_ls = mcs [where z = "w + x", unfolded add_diff_cancel] for w x lemmas word_plus_mcs = word_diff_ls [where y = "v + x", unfolded add_diff_cancel] for v x lemmas le_unat_uoi = unat_le [THEN word_unat.Abs_inverse] lemmas thd = times_div_less_eq_dividend lemmas uno_simps [THEN le_unat_uoi] = mod_le_divisor div_le_dividend lemma word_mod_div_equality: "(n div b) * b + (n mod b) = n" for n b :: "'a::len word" by (fact div_mult_mod_eq) lemma word_div_mult_le: "a div b * b \ a" for a b :: "'a::len word" by (metis div_le_mult mult_not_zero order.not_eq_order_implies_strict order_refl word_zero_le) lemma word_mod_less_divisor: "0 < n \ m mod n < n" for m n :: "'a::len word" by (simp add: unat_arith_simps) lemma word_of_int_power_hom: "word_of_int a ^ n = (word_of_int (a ^ n) :: 'a::len word)" by (induct n) (simp_all add: wi_hom_mult [symmetric]) lemma word_arith_power_alt: "a ^ n = (word_of_int (uint a ^ n) :: 'a::len word)" by (simp add : word_of_int_power_hom [symmetric]) lemma unatSuc: "1 + n \ 0 \ unat (1 + n) = Suc (unat n)" for n :: "'a::len word" by unat_arith subsection \Cardinality, finiteness of set of words\ lemma inj_on_word_of_int: \inj_on (word_of_int :: int \ 'a word) {0..<2 ^ LENGTH('a::len)}\ apply (rule inj_onI) apply transfer apply (simp add: take_bit_eq_mod) done lemma inj_uint: \inj uint\ by (rule injI) simp lemma range_uint: \range (uint :: 'a word \ int) = {0..<2 ^ LENGTH('a::len)}\ by transfer (auto simp add: bintr_lt2p range_bintrunc) lemma UNIV_eq: \(UNIV :: 'a word set) = word_of_int ` {0..<2 ^ LENGTH('a::len)}\ proof - have \uint ` (UNIV :: 'a word set) = uint ` (word_of_int :: int \ 'a word) ` {0..<2 ^ LENGTH('a::len)}\ by transfer (simp add: range_bintrunc, auto simp add: take_bit_eq_mod) then show ?thesis using inj_image_eq_iff [of \uint :: 'a word \ int\ \UNIV :: 'a word set\ \word_of_int ` {0..<2 ^ LENGTH('a)} :: 'a word set\, OF inj_uint] by simp qed lemma card_word: "CARD('a word) = 2 ^ LENGTH('a::len)" by (simp add: UNIV_eq card_image inj_on_word_of_int) lemma card_word_size: "CARD('a word) = 2 ^ size x" for x :: "'a::len word" unfolding word_size by (rule card_word) instance word :: (len) finite by standard (simp add: UNIV_eq) subsection \Bitwise Operations on Words\ lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or \ \following definitions require both arithmetic and bit-wise word operations\ \ \to get \word_no_log_defs\ from \word_log_defs\, using \bin_log_bintrs\\ lemmas wils1 = bin_log_bintrs [THEN word_ubin.norm_eq_iff [THEN iffD1], folded word_ubin.eq_norm, THEN eq_reflection] \ \the binary operations only\ (* BH: why is this needed? *) lemmas word_log_binary_defs = word_and_def word_or_def word_xor_def lemma word_wi_log_defs: "NOT (word_of_int a) = word_of_int (NOT a)" "word_of_int a AND word_of_int b = word_of_int (a AND b)" "word_of_int a OR word_of_int b = word_of_int (a OR b)" "word_of_int a XOR word_of_int b = word_of_int (a XOR b)" by (transfer, rule refl)+ lemma word_no_log_defs [simp]: "NOT (numeral a) = word_of_int (NOT (numeral a))" "NOT (- numeral a) = word_of_int (NOT (- numeral a))" "numeral a AND numeral b = word_of_int (numeral a AND numeral b)" "numeral a AND - numeral b = word_of_int (numeral a AND - numeral b)" "- numeral a AND numeral b = word_of_int (- numeral a AND numeral b)" "- numeral a AND - numeral b = word_of_int (- numeral a AND - numeral b)" "numeral a OR numeral b = word_of_int (numeral a OR numeral b)" "numeral a OR - numeral b = word_of_int (numeral a OR - numeral b)" "- numeral a OR numeral b = word_of_int (- numeral a OR numeral b)" "- numeral a OR - numeral b = word_of_int (- numeral a OR - numeral b)" "numeral a XOR numeral b = word_of_int (numeral a XOR numeral b)" "numeral a XOR - numeral b = word_of_int (numeral a XOR - numeral b)" "- numeral a XOR numeral b = word_of_int (- numeral a XOR numeral b)" "- numeral a XOR - numeral b = word_of_int (- numeral a XOR - numeral b)" by (transfer, rule refl)+ text \Special cases for when one of the arguments equals 1.\ lemma word_bitwise_1_simps [simp]: "NOT (1::'a::len word) = -2" "1 AND numeral b = word_of_int (1 AND numeral b)" "1 AND - numeral b = word_of_int (1 AND - numeral b)" "numeral a AND 1 = word_of_int (numeral a AND 1)" "- numeral a AND 1 = word_of_int (- numeral a AND 1)" "1 OR numeral b = word_of_int (1 OR numeral b)" "1 OR - numeral b = word_of_int (1 OR - numeral b)" "numeral a OR 1 = word_of_int (numeral a OR 1)" "- numeral a OR 1 = word_of_int (- numeral a OR 1)" "1 XOR numeral b = word_of_int (1 XOR numeral b)" "1 XOR - numeral b = word_of_int (1 XOR - numeral b)" "numeral a XOR 1 = word_of_int (numeral a XOR 1)" "- numeral a XOR 1 = word_of_int (- numeral a XOR 1)" by (transfer, simp)+ text \Special cases for when one of the arguments equals -1.\ lemma word_bitwise_m1_simps [simp]: "NOT (-1::'a::len word) = 0" "(-1::'a::len word) AND x = x" "x AND (-1::'a::len word) = x" "(-1::'a::len word) OR x = -1" "x OR (-1::'a::len word) = -1" " (-1::'a::len word) XOR x = NOT x" "x XOR (-1::'a::len word) = NOT x" by (transfer, simp)+ lemma uint_and: \uint (x AND y) = uint x AND uint y\ by transfer simp lemma uint_or: \uint (x OR y) = uint x OR uint y\ by transfer simp lemma uint_xor: \uint (x XOR y) = uint x XOR uint y\ by transfer simp lemma test_bit_wi [simp]: "(word_of_int x :: 'a::len word) !! n \ n < LENGTH('a) \ bin_nth x n" by (simp add: word_test_bit_def word_ubin.eq_norm nth_bintr) 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 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" unfolding word_size by transfer (simp add: bin_nth_ops) 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: bin_nth_ops) 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) \ bin_nth (numeral w) n" by transfer (rule refl) lemma test_bit_neg_numeral [simp]: "(- numeral w :: 'a::len word) !! n \ n < LENGTH('a) \ bin_nth (- numeral w) n" by transfer (rule refl) lemma test_bit_1 [simp]: "(1 :: 'a::len word) !! n \ n = 0" by transfer auto 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 \ \get from commutativity, associativity etc of \int_and\ etc to same for \word_and etc\\ lemmas bwsimps = wi_hom_add word_wi_log_defs lemma word_bw_assocs: "(x AND y) AND z = x AND y AND z" "(x OR y) OR z = x OR y OR z" "(x XOR y) XOR z = x XOR y XOR z" for x :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_bw_comms: "x AND y = y AND x" "x OR y = y OR x" "x XOR y = y XOR x" for x :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_bw_lcs: "y AND x AND z = x AND y AND z" "y OR x OR z = x OR y OR z" "y XOR x XOR z = x XOR y XOR z" for x :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_log_esimps: "x AND 0 = 0" "x AND -1 = x" "x OR 0 = x" "x OR -1 = -1" "x XOR 0 = x" "x XOR -1 = NOT x" "0 AND x = 0" "-1 AND x = x" "0 OR x = x" "-1 OR x = -1" "0 XOR x = x" "-1 XOR x = NOT x" for x :: "'a::len word" by simp_all lemma word_not_dist: "NOT (x OR y) = NOT x AND NOT y" "NOT (x AND y) = NOT x OR NOT y" for x :: "'a::len word" by simp_all lemma word_bw_same: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: "'a::len word" by simp_all lemma word_ao_absorbs [simp]: "x AND (y OR x) = x" "x OR y AND x = x" "x AND (x OR y) = x" "y AND x OR x = x" "(y OR x) AND x = x" "x OR x AND y = x" "(x OR y) AND x = x" "x AND y OR x = x" for x :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_not_not [simp]: "NOT (NOT x) = x" for x :: "'a::len word" by simp lemma word_ao_dist: "(x OR y) AND z = x AND z OR y AND z" for x :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_oa_dist: "x AND y OR z = (x OR z) AND (y OR z)" for x :: "'a::len word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_add_not [simp]: "x + NOT x = -1" for x :: "'a::len word" by transfer (simp add: bin_add_not) lemma word_plus_and_or [simp]: "(x AND y) + (x OR y) = x + y" for x :: "'a::len word" by transfer (simp add: plus_and_or) lemma leoa: "w = x OR y \ y = w AND y" for x :: "'a::len word" by auto lemma leao: "w' = x' AND y' \ x' = x' OR w'" for x' :: "'a::len word" by auto lemma word_ao_equiv: "w = w OR w' \ w' = w AND w'" for w w' :: "'a::len word" by (auto intro: leoa leao) lemma le_word_or2: "x \ x OR y" for x y :: "'a::len word" by (auto simp: word_le_def uint_or intro: le_int_or) lemmas le_word_or1 = xtrans(3) [OF word_bw_comms (2) le_word_or2] lemmas word_and_le1 = xtrans(3) [OF word_ao_absorbs (4) [symmetric] le_word_or2] lemmas word_and_le2 = xtrans(3) [OF word_ao_absorbs (8) [symmetric] le_word_or2] lemma bit_horner_sum_bit_word_iff: \bit (horner_sum of_bool (2 :: 'a::len word) bs) n \ n < min LENGTH('a) (length bs) \ bs ! n\ by transfer (simp add: bit_horner_sum_bit_iff) definition word_reverse :: \'a::len word \ 'a word\ where \word_reverse w = horner_sum of_bool 2 (rev (map (bit w) [0.. lemma bit_word_reverse_iff: \bit (word_reverse w) n \ n < LENGTH('a) \ bit w (LENGTH('a) - Suc n)\ for w :: \'a::len word\ by (cases \n < LENGTH('a)\) (simp_all add: word_reverse_def bit_horner_sum_bit_word_iff rev_nth) lemma word_rev_rev [simp] : "word_reverse (word_reverse w) = w" by (rule bit_word_eqI) (auto simp add: bit_word_reverse_iff bit_imp_le_length Suc_diff_Suc) lemma word_rev_gal: "word_reverse w = u \ word_reverse u = w" by (metis word_rev_rev) lemma word_rev_gal': "u = word_reverse w \ w = word_reverse u" by simp 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 "bin_nth (sint w) n = (if n < l - 1 then w !! n else w !! (l - 1))" unfolding sint_uint l_def by (auto simp: nth_sbintr word_test_bit_def [symmetric]) lemma setBit_no [simp]: "setBit (numeral bin) n = word_of_int (bin_sc n True (numeral bin))" by transfer (simp add: bin_sc_eq) lemma clearBit_no [simp]: "clearBit (numeral bin) n = word_of_int (bin_sc n False (numeral bin))" by transfer (simp add: bin_sc_eq) 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 uint_2p: "(0::'a::len word) < 2 ^ n \ uint (2 ^ n::'a::len word) = 2 ^ n" apply (cases \n < LENGTH('a)\; transfer) apply auto done lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a::len word) = 2 ^ n" by (induct n) (simp_all add: wi_hom_syms) 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 subsection \Bit comprehension\ instantiation word :: (len) bit_comprehension begin definition word_set_bits_def: \(BITS n. P n) = (horner_sum of_bool 2 (map P [0.. instance .. end lemma bit_set_bits_word_iff: \bit (set_bits P :: 'a::len word) n \ n < LENGTH('a) \ P n\ by (auto simp add: word_set_bits_def bit_horner_sum_bit_word_iff) lemma set_bits_bit_eq: \set_bits (bit w) = w\ for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_set_bits_word_iff bit_imp_le_length) lemma set_bits_K_False [simp]: \set_bits (\_. False) = (0 :: 'a :: len word)\ by (rule bit_word_eqI) (simp add: bit_set_bits_word_iff) interpretation test_bit: td_ext "(!!) :: 'a::len word \ nat \ bool" set_bits "{f. \i. f i \ i < LENGTH('a::len)}" "(\h i. h i \ i < LENGTH('a::len))" by standard (auto simp add: test_bit_word_eq bit_imp_le_length bit_set_bits_word_iff set_bits_bit_eq) lemmas td_nth = test_bit.td_thm subsection \Shifting, Rotating, and Splitting Words\ lemma shiftl1_wi [simp]: "shiftl1 (word_of_int w) = word_of_int (2 * w)" by transfer simp lemma shiftl1_numeral [simp]: "shiftl1 (numeral w) = numeral (Num.Bit0 w)" unfolding word_numeral_alt shiftl1_wi by simp lemma shiftl1_neg_numeral [simp]: "shiftl1 (- numeral w) = - numeral (Num.Bit0 w)" unfolding word_neg_numeral_alt shiftl1_wi by simp lemma shiftl1_0 [simp] : "shiftl1 0 = 0" by transfer simp lemma shiftl1_def_u: "shiftl1 w = word_of_int (2 * uint w)" by (fact shiftl1_eq) lemma shiftl1_def_s: "shiftl1 w = word_of_int (2 * sint w)" by (simp add: shiftl1_def_u wi_hom_syms) lemma shiftr1_0 [simp]: "shiftr1 0 = 0" by transfer simp lemma sshiftr1_0 [simp]: "sshiftr1 0 = 0" by transfer simp lemma sshiftr1_n1 [simp]: "sshiftr1 (- 1) = - 1" by transfer simp 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 sshiftr_0 [simp]: "0 >>> n = 0" by transfer simp lemma sshiftr_n1 [simp]: "-1 >>> n = -1" 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 text \ see paper page 10, (1), (2), \shiftr1_def\ is of the form of (1), where \f\ (ie \bin_rest\) takes normal arguments to normal results, thus we get (2) from (1) \ lemma uint_shiftr1: "uint (shiftr1 w) = bin_rest (uint w)" by transfer simp lemma bit_sshiftr1_iff: \bit (sshiftr1 w) n \ bit w (if n = LENGTH('a) - 1 then LENGTH('a) - 1 else Suc n)\ for w :: \'a::len word\ 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 bit_sshiftr_word_iff: \bit (w >>> m) n \ bit w (if LENGTH('a) - m \ n \ n < LENGTH('a) then LENGTH('a) - 1 else (m + n))\ for w :: \'a::len word\ 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_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 nth_sshiftr : "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 shiftr1_div_2: "uint (shiftr1 w) = uint w div 2" by (fact uint_shiftr1) lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2" by transfer simp lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" apply (unfold shiftr_def) apply (induct n) apply simp apply (simp add: shiftr1_div_2 mult.commute zdiv_zmult2_eq [symmetric]) done lemma sshiftr_div_2n: "sint (sshiftr w n) = sint w div 2 ^ n" apply transfer apply (rule bit_eqI) apply (simp add: bit_signed_take_bit_iff bit_drop_bit_eq min_def flip: drop_bit_eq_div) done lemma bit_bshiftr1_iff: \bit (bshiftr1 b w) n \ b \ n = LENGTH('a) - 1 \ bit w (Suc n)\ for w :: \'a::len word\ apply transfer apply (simp add: bit_take_bit_iff flip: bit_Suc) apply (subst disjunctive_add) apply (auto simp add: bit_take_bit_iff bit_or_iff bit_exp_iff simp flip: bit_Suc) done subsubsection \shift functions in terms of lists of bools\ lemma shiftl1_rev: "shiftl1 w = word_reverse (shiftr1 (word_reverse w))" apply (rule bit_word_eqI) apply (auto simp add: bit_shiftl1_iff bit_word_reverse_iff bit_shiftr1_iff Suc_diff_Suc) done 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 \ \note -- the following results use \'a::len word < number_ring\\ lemma shiftl1_2t: "shiftl1 w = 2 * w" for w :: "'a::len word" by (simp add: shiftl1_eq wi_hom_mult [symmetric]) lemma shiftl1_p: "shiftl1 w = w + w" for w :: "'a::len word" by (simp add: shiftl1_2t) lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" for w :: "'a::len word" by (induct n) (auto simp: shiftl_def shiftl1_2t) lemma shiftr1_bintr [simp]: "(shiftr1 (numeral w) :: 'a::len word) = word_of_int (bin_rest (take_bit (LENGTH('a)) (numeral w)))" by transfer simp lemma sshiftr1_sbintr [simp]: "(sshiftr1 (numeral w) :: 'a::len word) = word_of_int (bin_rest (signed_take_bit (LENGTH('a) - 1) (numeral w)))" unfolding sshiftr1_eq word_numeral_alt by (simp add: word_sbin.eq_norm) text \TODO: rules for \<^term>\- (numeral n)\\ lemma drop_bit_word_numeral [simp]: \drop_bit (numeral n) (numeral k) = (word_of_int (drop_bit (numeral n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ by transfer simp lemma shiftr_numeral [simp]: \(numeral k >> numeral n :: 'a::len word) = drop_bit (numeral n) (numeral k)\ by (fact shiftr_word_eq) 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 nth_sbintr not_le not_less less_Suc_eq_le ac_simps) done lemma zip_replicate: "n \ length ys \ zip (replicate n x) ys = map (\y. (x, y)) ys" apply (induct ys arbitrary: n) apply simp_all apply (case_tac n) apply simp_all done lemma align_lem_or [rule_format] : "\x m. length x = n + m \ length y = n + m \ drop m x = replicate n False \ take m y = replicate m False \ map2 (|) x y = take m x @ drop m y" apply (induct y) apply force apply clarsimp apply (case_tac x) apply force apply (case_tac m) apply auto apply (drule_tac t="length xs" for xs in sym) apply (auto simp: zip_replicate o_def) done lemma align_lem_and [rule_format] : "\x m. length x = n + m \ length y = n + m \ drop m x = replicate n False \ take m y = replicate m False \ map2 (\) x y = replicate (n + m) False" apply (induct y) apply force apply clarsimp apply (case_tac x) apply force apply (case_tac m) apply auto apply (drule_tac t="length xs" for xs in sym) apply (auto simp: zip_replicate o_def map_replicate_const) done subsubsection \Mask\ lemma minus_1_eq_mask: \- 1 = (mask LENGTH('a) :: 'a::len word)\ by (rule bit_eqI) (simp add: bit_exp_iff bit_mask_iff exp_eq_zero_iff) lemma mask_eq_decr_exp: \mask n = 2 ^ n - (1 :: 'a::len word)\ by (fact mask_eq_exp_minus_1) lemma mask_Suc_rec: \mask (Suc n) = 2 * mask n + (1 :: 'a::len word)\ by (simp add: mask_eq_exp_minus_1) context begin qualified lemma bit_mask_iff: \bit (mask m :: 'a::len word) n \ n < min LENGTH('a) m\ by (simp add: bit_mask_iff exp_eq_zero_iff not_le) end 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 mask_bin: "mask n = word_of_int (take_bit n (- 1))" by (auto simp add: nth_bintr word_size intro: word_eqI) lemma and_mask_bintr: "w AND mask n = word_of_int (take_bit n (uint w))" apply (rule word_eqI) apply (simp add: nth_bintr word_size word_ops_nth_size) apply (auto simp add: test_bit_bin) done lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (take_bit n i)" by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff) lemma and_mask_wi': "word_of_int i AND mask n = (word_of_int (take_bit (min LENGTH('a) n) i) :: 'a::len word)" by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff) lemma and_mask_no: "numeral i AND mask n = word_of_int (take_bit n (numeral i))" unfolding word_numeral_alt by (rule and_mask_wi) lemma and_mask_mod_2p: "w AND mask n = word_of_int (uint w mod 2 ^ n)" by (simp only: and_mask_bintr take_bit_eq_mod) lemma uint_mask_eq: \uint (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ by transfer simp lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n" apply (simp add: uint_and uint_mask_eq) apply (rule AND_upper2'') apply simp apply (auto simp add: mask_eq_exp_minus_1 min_def power_add algebra_simps dest!: le_Suc_ex) apply (metis add_minus_cancel le_add2 one_le_numeral power_add power_increasing uminus_add_conv_diff zle_diff1_eq) done lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" for b n :: int by auto (metis pos_mod_conj)+ lemma mask_eq_iff: "w AND mask n = w \ uint w < 2 ^ n" apply (simp add: and_mask_bintr) apply (simp add: word_ubin.inverse_norm) apply (simp add: eq_mod_iff take_bit_eq_mod min_def) apply (fast intro!: lt2p_lem) done lemma and_mask_dvd: "2 ^ n dvd uint w \ w AND mask n = 0" by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 uint_0_iff) lemma and_mask_dvd_nat: "2 ^ n dvd unat w \ w AND mask n = 0" by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 unat_0_iff uint_0_iff) lemma word_2p_lem: "n < size w \ w < 2 ^ n = (uint w < 2 ^ n)" for w :: "'a::len word" by transfer simp lemma less_mask_eq: "x < 2 ^ n \ x AND mask n = x" for x :: "'a::len word" apply (simp add: and_mask_bintr) apply transfer apply (simp add: ac_simps) apply (auto simp add: min_def) apply (metis bintrunc_bintrunc_ge mod_pos_pos_trivial mult.commute mult.left_neutral mult_zero_left not_le of_bool_def take_bit_eq_mod take_bit_nonnegative) done lemmas mask_eq_iff_w2p = trans [OF mask_eq_iff word_2p_lem [symmetric]] lemmas and_mask_less' = iffD2 [OF word_2p_lem and_mask_lt_2p, simplified word_size] lemma and_mask_less_size: "n < size x \ x AND mask n < 2 ^ n" for x :: \'a::len word\ unfolding word_size by (erule and_mask_less') lemma word_mod_2p_is_mask [OF refl]: "c = 2 ^ n \ c > 0 \ x mod c = x AND mask n" for c x :: "'a::len word" by (auto simp: word_mod_def uint_2p and_mask_mod_2p) lemma mask_eqs: "(a AND mask n) + b AND mask n = a + b AND mask n" "a + (b AND mask n) AND mask n = a + b AND mask n" "(a AND mask n) - b AND mask n = a - b AND mask n" "a - (b AND mask n) AND mask n = a - b AND mask n" "a * (b AND mask n) AND mask n = a * b AND mask n" "(b AND mask n) * a AND mask n = b * a AND mask n" "(a AND mask n) + (b AND mask n) AND mask n = a + b AND mask n" "(a AND mask n) - (b AND mask n) AND mask n = a - b AND mask n" "(a AND mask n) * (b AND mask n) AND mask n = a * b AND mask n" "- (a AND mask n) AND mask n = - a AND mask n" "word_succ (a AND mask n) AND mask n = word_succ a AND mask n" "word_pred (a AND mask n) AND mask n = word_pred a AND mask n" using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b] apply (auto simp flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) apply transfer apply (simp add: take_bit_eq_mod mod_simps) done lemma mask_power_eq: "(x AND mask n) ^ k AND mask n = x ^ k AND mask n" for x :: \'a::len word\ using word_of_int_Ex [where x=x] apply (auto simp flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_eq_mod mod_simps) done lemma mask_full [simp]: "mask LENGTH('a) = (- 1 :: 'a::len word)" by transfer (simp add: take_bit_minus_one_eq_mask) subsubsection \Slices\ definition slice1 :: \nat \ 'a::len word \ 'b::len word\ where \slice1 n w = (if n < LENGTH('a) then ucast (drop_bit (LENGTH('a) - n) w) else push_bit (n - LENGTH('a)) (ucast w))\ lemma bit_slice1_iff: \bit (slice1 m w :: 'b::len word) n \ m - LENGTH('a) \ n \ n < min LENGTH('b) m \ bit w (n + (LENGTH('a) - m) - (m - LENGTH('a)))\ for w :: \'a::len word\ by (auto simp add: slice1_def bit_ucast_iff bit_drop_bit_eq bit_push_bit_iff exp_eq_zero_iff not_less not_le ac_simps dest: bit_imp_le_length) definition slice :: \nat \ 'a::len word \ 'b::len word\ where \slice n = slice1 (LENGTH('a) - n)\ lemma bit_slice_iff: \bit (slice m w :: 'b::len word) n \ n < min LENGTH('b) (LENGTH('a) - m) \ bit w (n + LENGTH('a) - (LENGTH('a) - m))\ for w :: \'a::len word\ by (simp add: slice_def word_size bit_slice1_iff) lemma slice1_0 [simp] : "slice1 n 0 = 0" unfolding slice1_def by simp lemma slice_0 [simp] : "slice n 0 = 0" unfolding slice_def by auto 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 ucast_slice1: "ucast w = slice1 (size w) w" apply (simp add: slice1_def) apply transfer apply simp done lemma ucast_slice: "ucast w = slice 0 w" by (simp add: slice_def slice1_def) lemma slice_id: "slice 0 t = t" by (simp only: ucast_slice [symmetric] ucast_id) lemma rev_slice1: \slice1 n (word_reverse w :: 'b::len word) = word_reverse (slice1 k w :: 'a::len word)\ if \n + k = LENGTH('a) + LENGTH('b)\ proof (rule bit_word_eqI) fix m assume *: \m < LENGTH('a)\ from that have **: \LENGTH('b) = n + k - LENGTH('a)\ by simp show \bit (slice1 n (word_reverse w :: 'b word) :: 'a word) m \ bit (word_reverse (slice1 k w :: 'a word)) m\ apply (simp add: bit_slice1_iff bit_word_reverse_iff) using * ** apply (cases \n \ LENGTH('a)\; cases \k \ LENGTH('a)\) apply auto done qed lemma rev_slice: "n + k + LENGTH('a::len) = LENGTH('b::len) \ slice n (word_reverse (w::'b word)) = word_reverse (slice k w :: 'a word)" apply (unfold slice_def word_size) apply (rule rev_slice1) apply arith done subsubsection \Revcast\ definition revcast :: \'a::len word \ 'b::len word\ where \revcast = slice1 LENGTH('b)\ lemma bit_revcast_iff: \bit (revcast w :: 'b::len word) n \ LENGTH('b) - LENGTH('a) \ n \ n < LENGTH('b) \ bit w (n + (LENGTH('a) - LENGTH('b)) - (LENGTH('b) - LENGTH('a)))\ for w :: \'a::len word\ by (simp add: revcast_def bit_slice1_iff) lemma revcast_slice1 [OF refl]: "rc = revcast w \ slice1 (size rc) w = rc" by (simp add: revcast_def word_size) lemma revcast_rev_ucast [OF refl refl refl]: "cs = [rc, uc] \ rc = revcast (word_reverse w) \ uc = ucast w \ rc = word_reverse uc" apply auto apply (rule bit_word_eqI) apply (cases \LENGTH('a) \ LENGTH('b)\) apply (simp_all add: bit_revcast_iff bit_word_reverse_iff bit_ucast_iff not_le bit_imp_le_length) using bit_imp_le_length apply fastforce using bit_imp_le_length apply fastforce done lemma revcast_ucast: "revcast w = word_reverse (ucast (word_reverse w))" using revcast_rev_ucast [of "word_reverse w"] by simp lemma ucast_revcast: "ucast w = word_reverse (revcast (word_reverse w))" by (fact revcast_rev_ucast [THEN word_rev_gal']) lemma ucast_rev_revcast: "ucast (word_reverse w) = word_reverse (revcast w)" by (fact revcast_ucast [THEN word_rev_gal']) text "linking revcast and cast via shift" lemmas wsst_TYs = source_size target_size word_size 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_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_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 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 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]] lemmas sym_notr = not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]] \ \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_range' [of x] sint_range' [of y] apply (auto simp add: not_less) apply (unfold sint_word_ariths word_sbin.set_iff_norm [symmetric] sints_num) apply (auto simp add: signed_take_bit_eq_take_bit_minus take_bit_Suc_from_most n not_less intro!: *) 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) subsection \Split and cat\ lemmas word_split_bin' = word_split_def lemmas word_cat_bin' = word_cat_eq lemma word_rsplit_no: "(word_rsplit (numeral bin :: 'b::len word) :: 'a word list) = map word_of_int (bin_rsplit (LENGTH('a::len)) (LENGTH('b), take_bit (LENGTH('b)) (numeral bin)))" by (simp add: word_rsplit_def of_nat_take_bit) lemmas word_rsplit_no_cl [simp] = word_rsplit_no [unfolded bin_rsplitl_def bin_rsplit_l [symmetric]] 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 lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ a = take_bit (LENGTH('a) - n) a \ b = take_bit (LENGTH('a)) b" for w :: "'a::len word" apply (frule word_ubin.norm_Rep [THEN ssubst]) apply (drule bin_split_trunc1) apply (drule sym [THEN trans]) apply assumption apply safe 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)))" apply (unfold word_split_bin' test_bit_bin) apply (clarify) apply (clarsimp simp: word_ubin.eq_norm nth_bintr word_size split: prod.splits) apply (auto simp add: bit_take_bit_iff bit_drop_bit_eq bit_unsigned_iff ac_simps exp_eq_zero_iff dest: bit_imp_le_length) done 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 \ \this odd result is analogous to \ucast_id\, result to the length given by the result type\ lemma word_cat_id: "word_cat a b = b" by transfer simp \ \limited hom result\ lemma word_cat_hom: "LENGTH('a::len) \ LENGTH('b::len) + LENGTH('c::len) \ (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = word_of_int (bin_cat w (size b) (uint b))" apply transfer using bintr_cat by auto lemma word_cat_split_alt: "size w \ size u + size v \ word_split w = (u, v) \ word_cat u v = w" apply (rule word_eqI) apply (drule test_bit_split) apply (clarsimp simp add : test_bit_cat word_size) apply safe apply arith done lemmas word_cat_split_size = sym [THEN [2] word_cat_split_alt [symmetric]] subsubsection \Split and slice\ lemma split_slices: "word_split w = (u, v) \ u = slice (size v) w \ v = slice 0 w" apply (drule test_bit_split) apply (rule conjI) apply (rule word_eqI, clarsimp simp: nth_slice word_size)+ done lemma slice_cat1 [OF refl]: "wc = word_cat a b \ size wc >= size a + size b \ slice (size b) wc = a" apply safe apply (rule word_eqI) apply (simp add: nth_slice test_bit_cat word_size) done lemmas slice_cat2 = trans [OF slice_id word_cat_id] lemma cat_slices: "a = slice n c \ b = slice 0 c \ n = size b \ size a + size b >= size c \ word_cat a b = c" apply safe apply (rule word_eqI) apply (simp add: nth_slice test_bit_cat word_size) apply safe apply arith done lemma word_split_cat_alt: "w = word_cat u v \ size u + size v \ size w \ word_split w = (u, v)" apply (case_tac "word_split w") apply (rule trans, assumption) apply (drule test_bit_split) apply safe apply (rule word_eqI, clarsimp simp: test_bit_cat word_size)+ done text \ This odd result arises from the fact that the statement of the result implies that the decoded words are of the same type, and therefore of the same length, as the original word.\ lemma word_rsplit_same: "word_rsplit w = [w]" by (simp add: word_rsplit_def bin_rsplit_all) lemma word_rsplit_empty_iff_size: "word_rsplit w = [] \ size w = 0" by (simp add: word_rsplit_def bin_rsplit_def word_size bin_rsplit_aux_simp_alt Let_def split: prod.split) lemma test_bit_rsplit: "sw = word_rsplit w \ m < size (hd sw) \ k < length sw \ (rev sw ! k) !! m = w !! (k * size (hd sw) + m)" for sw :: "'a::len word list" apply (unfold word_rsplit_def word_test_bit_def) apply (rule trans) apply (rule_tac f = "\x. bin_nth x m" in arg_cong) apply (rule nth_map [symmetric]) apply simp apply (rule bin_nth_rsplit) apply simp_all apply (simp add : word_size rev_map) apply (rule trans) defer apply (rule map_ident [THEN fun_cong]) apply (rule refl [THEN map_cong]) apply (simp add : word_ubin.eq_norm) apply (erule bin_rsplit_size_sign [OF len_gt_0 refl]) done lemma horner_sum_uint_exp_Cons_eq: \horner_sum uint (2 ^ LENGTH('a)) (w # ws) = concat_bit LENGTH('a) (uint w) (horner_sum uint (2 ^ LENGTH('a)) ws)\ for ws :: \'a::len word list\ by (simp add: concat_bit_eq push_bit_eq_mult) lemma bit_horner_sum_uint_exp_iff: \bit (horner_sum uint (2 ^ LENGTH('a)) ws) n \ n div LENGTH('a) < length ws \ bit (ws ! (n div LENGTH('a))) (n mod LENGTH('a))\ for ws :: \'a::len word list\ proof (induction ws arbitrary: n) case Nil then show ?case by simp next case (Cons w ws) then show ?case by (cases \n \ LENGTH('a)\) (simp_all only: horner_sum_uint_exp_Cons_eq, simp_all add: bit_concat_bit_iff le_div_geq le_mod_geq bit_uint_iff Cons) qed 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 bin_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 test_bit_rsplit_alt: \(word_rsplit w :: 'b::len word list) ! i !! m \ w !! ((length (word_rsplit w :: 'b::len word list) - Suc i) * size (hd (word_rsplit w :: 'b::len word list)) + m)\ if \i < length (word_rsplit w :: 'b::len word list)\ \m < size (hd (word_rsplit w :: 'b::len word list))\ \0 < length (word_rsplit w :: 'b::len word list)\ for w :: \'a::len word\ apply (rule trans) apply (rule test_bit_cong) apply (rule rev_nth [of _ \rev (word_rsplit w)\, simplified rev_rev_ident]) apply simp apply (rule that(1)) apply simp apply (rule test_bit_rsplit) apply (rule refl) apply (rule asm_rl) apply (rule that(2)) apply (rule diff_Suc_less) apply (rule that(3)) done lemma word_rsplit_len_indep [OF refl refl refl refl]: "[u,v] = p \ [su,sv] = q \ word_rsplit u = su \ word_rsplit v = sv \ length su = length sv" by (auto simp: word_rsplit_def bin_rsplit_len_indep) lemma length_word_rsplit_size: "n = LENGTH('a::len) \ length (word_rsplit w :: 'a word list) \ m \ size w \ m * n" by (auto simp: word_rsplit_def word_size bin_rsplit_len_le) lemmas length_word_rsplit_lt_size = length_word_rsplit_size [unfolded Not_eq_iff linorder_not_less [symmetric]] lemma length_word_rsplit_exp_size: "n = LENGTH('a::len) \ length (word_rsplit w :: 'a word list) = (size w + n - 1) div n" by (auto simp: word_rsplit_def word_size bin_rsplit_len) lemma length_word_rsplit_even_size: "n = LENGTH('a::len) \ size w = m * n \ length (word_rsplit w :: 'a word list) = m" by (cases \LENGTH('a)\) (simp_all add: length_word_rsplit_exp_size div_nat_eqI) lemmas length_word_rsplit_exp_size' = refl [THEN length_word_rsplit_exp_size] \ \alternative proof of \word_rcat_rsplit\\ lemmas tdle = times_div_less_eq_dividend lemmas dtle = xtrans(4) [OF tdle mult.commute] lemma word_rcat_rsplit: "word_rcat (word_rsplit w) = w" apply (rule word_eqI) apply (clarsimp simp: test_bit_rcat word_size) apply (subst refl [THEN test_bit_rsplit]) apply (simp_all add: word_size refl [THEN length_word_rsplit_size [simplified not_less [symmetric], simplified]]) apply safe apply (erule xtrans(7), rule dtle)+ done lemma size_word_rsplit_rcat_size: "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) \ length (word_rsplit frcw::'a word list) = length ws" for ws :: "'a::len word list" and frcw :: "'b::len word" by (cases \LENGTH('a)\) (simp_all add: word_size length_word_rsplit_exp_size' div_nat_eqI) lemma msrevs: "0 < n \ (k * n + m) div n = m div n + k" "(k * n + m) mod n = m mod n" for n :: nat by (auto simp: add.commute) lemma word_rsplit_rcat_size [OF refl]: "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) \ word_rsplit frcw = ws" for ws :: "'a::len word list" apply (frule size_word_rsplit_rcat_size, assumption) apply (clarsimp simp add : word_size) apply (rule nth_equalityI, assumption) apply clarsimp apply (rule word_eqI [rule_format]) apply (rule trans) apply (rule test_bit_rsplit_alt) apply (clarsimp simp: word_size)+ apply (rule trans) apply (rule test_bit_rcat [OF refl refl]) apply (simp add: word_size) apply (subst rev_nth) apply arith apply (simp add: le0 [THEN [2] xtrans(7), THEN diff_Suc_less]) apply safe apply (simp add: diff_mult_distrib) apply (cases "size ws") apply simp_all done subsection \Rotation\ lemma word_rotr_word_rotr_eq: \word_rotr m (word_rotr n w) = word_rotr (m + n) w\ by (rule bit_word_eqI) (simp add: bit_word_rotr_iff ac_simps mod_add_right_eq) lemma word_rot_rl [simp]: \word_rotl k (word_rotr k v) = v\ apply (rule bit_word_eqI) apply (simp add: word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps) apply (auto dest: bit_imp_le_length) apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_if mod_mult_self2_is_0) apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_less mod_mult_self2_is_0) done lemma word_rot_lr [simp]: \word_rotr k (word_rotl k v) = v\ apply (rule bit_word_eqI) apply (simp add: word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps) apply (auto dest: bit_imp_le_length) apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_if mod_mult_self2_is_0) apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_less mod_mult_self2_is_0) done lemma word_rot_gal: \word_rotr n v = w \ word_rotl n w = v\ by auto lemma word_rot_gal': \w = word_rotr n v \ v = word_rotl n w\ by auto lemma word_rotr_rev: \word_rotr n w = word_reverse (word_rotl n (word_reverse w))\ proof (rule bit_word_eqI) fix m assume \m < LENGTH('a)\ moreover have \1 + ((int m + int n mod int LENGTH('a)) mod int LENGTH('a) + ((int LENGTH('a) * 2) mod int LENGTH('a) - (1 + (int m + int n mod int LENGTH('a)))) mod int LENGTH('a)) = int LENGTH('a)\ apply (cases \(1 + (int m + int n mod int LENGTH('a))) mod int LENGTH('a) = 0\) using zmod_zminus1_eq_if [of \1 + (int m + int n mod int LENGTH('a))\ \int LENGTH('a)\] apply simp_all apply (auto simp add: algebra_simps) apply (simp add: minus_equation_iff [of \int m\]) apply (drule sym [of _ \int m\]) apply simp apply (metis add.commute add_minus_cancel diff_minus_eq_add len_gt_0 less_imp_of_nat_less less_nat_zero_code mod_mult_self3 of_nat_gt_0 zmod_minus1) apply (metis (no_types, hide_lams) Abs_fnat_hom_add less_not_refl mod_Suc of_nat_Suc of_nat_gt_0 of_nat_mod) done then have \int ((m + n) mod LENGTH('a)) = int (LENGTH('a) - Suc ((LENGTH('a) - Suc m + LENGTH('a) - n mod LENGTH('a)) mod LENGTH('a)))\ using \m < LENGTH('a)\ by (simp only: of_nat_mod mod_simps) (simp add: of_nat_diff of_nat_mod Suc_le_eq add_less_mono algebra_simps mod_simps) then have \(m + n) mod LENGTH('a) = LENGTH('a) - Suc ((LENGTH('a) - Suc m + LENGTH('a) - n mod LENGTH('a)) mod LENGTH('a))\ by simp ultimately show \bit (word_rotr n w) m \ bit (word_reverse (word_rotl n (word_reverse w))) m\ by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff bit_word_reverse_iff) qed lemma word_roti_0 [simp]: "word_roti 0 w = w" by transfer simp lemma word_roti_add: "word_roti (m + n) w = word_roti m (word_roti n w)" by (rule bit_word_eqI) (simp add: bit_word_roti_iff nat_less_iff mod_simps ac_simps) lemma word_roti_conv_mod': "word_roti n w = word_roti (n mod int (size w)) w" by transfer simp lemmas word_roti_conv_mod = word_roti_conv_mod' [unfolded word_size] subsubsection \"Word rotation commutes with bit-wise operations\ \ \using locale to not pollute lemma namespace\ locale word_rotate begin lemma word_rot_logs: "word_rotl n (NOT v) = NOT (word_rotl n v)" "word_rotr n (NOT v) = NOT (word_rotr n v)" "word_rotl n (x AND y) = word_rotl n x AND word_rotl n y" "word_rotr n (x AND y) = word_rotr n x AND word_rotr n y" "word_rotl n (x OR y) = word_rotl n x OR word_rotl n y" "word_rotr n (x OR y) = word_rotr n x OR word_rotr n y" "word_rotl n (x XOR y) = word_rotl n x XOR word_rotl n y" "word_rotr n (x XOR y) = word_rotr n x XOR word_rotr n y" apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotl_iff bit_not_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotr_iff bit_not_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotl_iff bit_and_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotr_iff bit_and_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotl_iff bit_or_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotr_iff bit_or_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotl_iff bit_xor_iff algebra_simps exp_eq_zero_iff not_le) apply (rule bit_word_eqI) apply (auto simp add: bit_word_rotr_iff bit_xor_iff algebra_simps exp_eq_zero_iff not_le) done end lemmas word_rot_logs = word_rotate.word_rot_logs lemma word_rotx_0 [simp] : "word_rotr i 0 = 0 \ word_rotl i 0 = 0" by transfer simp_all lemma word_roti_0' [simp] : "word_roti n 0 = 0" by transfer simp declare word_roti_eq_word_rotr_word_rotl [simp] subsection \Maximum machine word\ lemma word_int_cases: fixes x :: "'a::len word" obtains n where "x = word_of_int n" and "0 \ n" and "n < 2^LENGTH('a)" by (cases x rule: word_uint.Abs_cases) (simp add: uints_num) lemma word_nat_cases [cases type: word]: fixes x :: "'a::len word" obtains n where "x = of_nat n" and "n < 2^LENGTH('a)" by (cases x rule: word_unat.Abs_cases) (simp add: unats_def) lemma max_word_max [intro!]: "n \ max_word" by (fact word_order.extremum) lemma word_of_int_2p_len: "word_of_int (2 ^ LENGTH('a)) = (0::'a::len word)" by (subst word_uint.Abs_norm [symmetric]) simp lemma word_pow_0: "(2::'a::len word) ^ LENGTH('a) = 0" by (fact word_exp_length_eq_0) lemma max_word_wrap: "x + 1 = 0 \ x = max_word" by (simp add: eq_neg_iff_add_eq_0) lemma max_test_bit: "(max_word::'a::len word) !! n \ n < LENGTH('a)" by (fact nth_minus1) lemma word_and_max: "x AND max_word = x" by (fact word_log_esimps) lemma word_or_max: "x OR max_word = max_word" by (fact word_log_esimps) lemma word_ao_dist2: "x AND (y OR z) = x AND y OR x AND z" for x y z :: "'a::len word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_oa_dist2: "x OR y AND z = (x OR y) AND (x OR z)" for x y z :: "'a::len word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_and_not [simp]: "x AND NOT x = 0" for x :: "'a::len word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_or_not [simp]: "x OR NOT x = max_word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_xor_and_or: "x XOR y = x AND NOT y OR NOT x AND y" for x y :: "'a::len word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) 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 uint_lt_0 [simp]: "uint x < 0 = False" by (simp add: linorder_not_less) lemma shiftr1_1 [simp]: "shiftr1 (1::'a::len word) = 0" by transfer simp lemma shiftr_1[simp]: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" by (induct n) (auto simp: shiftr_def) lemma word_less_1 [simp]: "x < 1 \ x = 0" for x :: "'a::len word" by (simp add: word_less_nat_alt unat_0_iff) lemma map_nth_0 [simp]: "map ((!!) (0::'a::len word)) xs = replicate (length xs) False" by (induct xs) auto lemma uint_plus_if_size: "uint (x + y) = (if uint x + uint y < 2^size x then uint x + uint y else uint x + uint y - 2^size x)" apply (simp only: word_arith_wis word_size uint_word_of_int_eq) apply (auto simp add: not_less take_bit_int_eq_self_iff) apply (metis not_less uint_plus_if' word_add_def word_ubin.eq_norm) done lemma unat_plus_if_size: "unat (x + y) = (if unat x + unat y < 2^size x then unat x + unat y else unat x + unat y - 2^size x)" for x y :: "'a::len word" apply (subst word_arith_nat_defs) apply (subst unat_of_nat) apply (auto simp add: not_less word_size) apply (metis not_le unat_plus_if' unat_word_ariths(1)) done lemma word_neq_0_conv: "w \ 0 \ 0 < w" for w :: "'a::len word" by (fact word_coorder.not_eq_extremum) lemma max_lt: "unat (max a b div c) = unat (max a b) div unat c" for c :: "'a::len word" by (fact unat_div) lemma uint_sub_if_size: "uint (x - y) = (if uint y \ uint x then uint x - uint y else uint x - uint y + 2^size x)" apply (simp only: word_arith_wis word_size uint_word_of_int_eq) apply (auto simp add: take_bit_int_eq_self_iff not_le) apply (metis not_le uint_sub_if' word_sub_wi word_ubin.eq_norm) done lemma unat_sub: \unat (a - b) = unat a - unat b\ if \b \ a\ proof - from that have \unat b \ unat a\ by transfer simp with that show ?thesis apply transfer apply simp apply (subst take_bit_diff [symmetric]) apply (subst nat_take_bit_eq) apply (simp add: nat_le_eq_zle) apply (simp add: nat_diff_distrib take_bit_nat_eq_self_iff less_imp_diff_less bintr_lt2p) done qed lemmas word_less_sub1_numberof [simp] = word_less_sub1 [of "numeral w"] for w lemmas word_le_sub1_numberof [simp] = word_le_sub1 [of "numeral w"] for w lemma word_of_int_minus: "word_of_int (2^LENGTH('a) - i) = (word_of_int (-i)::'a::len word)" proof - have *: "2^LENGTH('a) - i = -i + 2^LENGTH('a)" by simp show ?thesis apply (subst *) apply (subst word_uint.Abs_norm [symmetric], subst mod_add_self2) apply simp done qed lemmas word_of_int_inj = word_uint.Abs_inject [unfolded uints_num, simplified] lemma word_le_less_eq: "x \ y \ x = y \ x < y" for x y :: "'z::len word" by (auto simp add: order_class.le_less) lemma mod_plus_cong: fixes b b' :: int assumes 1: "b = b'" and 2: "x mod b' = x' mod b'" and 3: "y mod b' = y' mod b'" and 4: "x' + y' = z'" shows "(x + y) mod b = z' mod b'" proof - from 1 2[symmetric] 3[symmetric] have "(x + y) mod b = (x' mod b' + y' mod b') mod b'" by (simp add: mod_add_eq) also have "\ = (x' + y') mod b'" by (simp add: mod_add_eq) finally show ?thesis by (simp add: 4) qed lemma mod_minus_cong: fixes b b' :: int assumes "b = b'" and "x mod b' = x' mod b'" and "y mod b' = y' mod b'" and "x' - y' = z'" shows "(x - y) mod b = z' mod b'" using assms [symmetric] by (auto intro: mod_diff_cong) lemma word_induct_less: \P m\ if zero: \P 0\ and less: \\n. n < m \ P n \ P (1 + n)\ for m :: \'a::len word\ proof - define q where \q = unat m\ with less have \\n. n < word_of_nat q \ P n \ P (1 + n)\ by simp then have \P (word_of_nat q :: 'a word)\ proof (induction q) case 0 show ?case by (simp add: zero) next case (Suc q) show ?case proof (cases \1 + word_of_nat q = (0 :: 'a word)\) case True then show ?thesis by (simp add: zero) next case False then have *: \word_of_nat q < (word_of_nat (Suc q) :: 'a word)\ by (simp add: unatSuc word_less_nat_alt) then have **: \n < (1 + word_of_nat q :: 'a word) \ n \ (word_of_nat q :: 'a word)\ for n by (metis (no_types, lifting) add.commute inc_le le_less_trans not_less of_nat_Suc) have \P (word_of_nat q)\ apply (rule Suc.IH) apply (rule Suc.prems) apply (erule less_trans) apply (rule *) apply assumption done with * have \P (1 + word_of_nat q)\ by (rule Suc.prems) then show ?thesis by simp qed qed with \q = unat m\ show ?thesis by simp qed lemma word_induct: "P 0 \ (\n. P n \ P (1 + n)) \ P m" for P :: "'a::len word \ bool" by (rule word_induct_less) lemma word_induct2 [induct type]: "P 0 \ (\n. 1 + n \ 0 \ P n \ P (1 + n)) \ P n" for P :: "'b::len word \ bool" apply (rule word_induct_less) apply simp_all apply (case_tac "1 + na = 0") apply auto done subsection \Recursion combinator for words\ definition word_rec :: "'a \ ('b::len word \ 'a \ 'a) \ 'b word \ 'a" where "word_rec forZero forSuc n = rec_nat forZero (forSuc \ of_nat) (unat n)" lemma word_rec_0: "word_rec z s 0 = z" by (simp add: word_rec_def) lemma word_rec_Suc: "1 + n \ 0 \ word_rec z s (1 + n) = s n (word_rec z s n)" for n :: "'a::len word" apply (auto simp add: word_rec_def unat_word_ariths) apply (metis (mono_tags, lifting) old.nat.simps(7) unatSuc word_unat.Rep_inverse word_unat.eq_norm word_unat.td_th) done lemma word_rec_Pred: "n \ 0 \ word_rec z s n = s (n - 1) (word_rec z s (n - 1))" apply (rule subst[where t="n" and s="1 + (n - 1)"]) apply simp apply (subst word_rec_Suc) apply simp apply simp done lemma word_rec_in: "f (word_rec z (\_. f) n) = word_rec (f z) (\_. f) n" by (induct n) (simp_all add: word_rec_0 word_rec_Suc) lemma word_rec_in2: "f n (word_rec z f n) = word_rec (f 0 z) (f \ (+) 1) n" by (induct n) (simp_all add: word_rec_0 word_rec_Suc) lemma word_rec_twice: "m \ n \ word_rec z f n = word_rec (word_rec z f (n - m)) (f \ (+) (n - m)) m" apply (erule rev_mp) apply (rule_tac x=z in spec) apply (rule_tac x=f in spec) apply (induct n) apply (simp add: word_rec_0) apply clarsimp apply (rule_tac t="1 + n - m" and s="1 + (n - m)" in subst) apply simp apply (case_tac "1 + (n - m) = 0") apply (simp add: word_rec_0) apply (rule_tac f = "word_rec a b" for a b in arg_cong) apply (rule_tac t="m" and s="m + (1 + (n - m))" in subst) apply simp apply (simp (no_asm_use)) apply (simp add: word_rec_Suc word_rec_in2) apply (erule impE) apply uint_arith apply (drule_tac x="x \ (+) 1" in spec) apply (drule_tac x="x 0 xa" in spec) apply simp apply (rule_tac t="\a. x (1 + (n - m + a))" and s="\a. x (1 + (n - m) + a)" in subst) apply (clarsimp simp add: fun_eq_iff) apply (rule_tac t="(1 + (n - m + xb))" and s="1 + (n - m) + xb" in subst) apply simp apply (rule refl) apply (rule refl) done lemma word_rec_id: "word_rec z (\_. id) n = z" by (induct n) (auto simp add: word_rec_0 word_rec_Suc) lemma word_rec_id_eq: "\m < n. f m = id \ word_rec z f n = z" apply (erule rev_mp) apply (induct n) apply (auto simp add: word_rec_0 word_rec_Suc) apply (drule spec, erule mp) apply uint_arith apply (drule_tac x=n in spec, erule impE) apply uint_arith apply simp done lemma word_rec_max: "\m\n. m \ - 1 \ f m = id \ word_rec z f (- 1) = word_rec z f n" apply (subst word_rec_twice[where n="-1" and m="-1 - n"]) apply simp apply simp apply (rule word_rec_id_eq) apply clarsimp apply (drule spec, rule mp, erule mp) apply (rule word_plus_mono_right2[OF _ order_less_imp_le]) prefer 2 apply assumption apply simp apply (erule contrapos_pn) apply simp apply (drule arg_cong[where f="\x. x - n"]) apply simp done subsection \More\ 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 mask_1: "mask 1 = 1" by transfer (simp add: min_def mask_Suc_exp) lemma mask_Suc_0: "mask (Suc 0) = 1" using mask_1 by simp lemma mask_numeral: "mask (numeral n) = 2 * mask (pred_numeral n) + (1 :: 'a::len word)" by (simp add: mask_Suc_rec numeral_eq_Suc) lemma bin_last_bintrunc: "bin_last (take_bit l n) = (l > 0 \ bin_last n)" by simp 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 bintrunc_shiftl: "take_bit n (m << i) = take_bit (n - i) m << i" for m :: int by (rule bit_eqI) (auto simp add: bit_take_bit_iff) lemma uint_shiftl: "uint (n << i) = take_bit (size n) (uint n << i)" by transfer (simp add: push_bit_take_bit shiftl_eq_push_bit) subsection \Misc\ ML_file \Tools/word_lib.ML\ ML_file \Tools/smt_word.ML\ end