Theory Landau_Symbols
section ‹Definition of Landau symbols›
theory Landau_Symbols
imports
Complex_Main
begin
lemma eventually_subst':
"eventually (λx. f x = g x) F ⟹ eventually (λx. P x (f x)) F = eventually (λx. P x (g x)) F"
by (rule eventually_subst, erule eventually_rev_mp) simp
subsection ‹Definition of Landau symbols›
text ‹
Our Landau symbols are sign-oblivious, i.e. any function always has the same growth
as its absolute. This has the advantage of making some cancelling rules for sums nicer, but
introduces some problems in other places. Nevertheless, we found this definition more convenient
to work with.
›
definition bigo :: "'a filter ⇒ ('a ⇒ ('b :: real_normed_field)) ⇒ ('a ⇒ 'b) set"
(‹(1O[_]'(_'))›)
where "bigo F g = {f. (∃c>0. eventually (λx. norm (f x) ≤ c * norm (g x)) F)}"
definition smallo :: "'a filter ⇒ ('a ⇒ ('b :: real_normed_field)) ⇒ ('a ⇒ 'b) set"
(‹(1o[_]'(_'))›)
where "smallo F g = {f. (∀c>0. eventually (λx. norm (f x) ≤ c * norm (g x)) F)}"
definition bigomega :: "'a filter ⇒ ('a ⇒ ('b :: real_normed_field)) ⇒ ('a ⇒ 'b) set"
(‹(1Ω[_]'(_'))›)
where "bigomega F g = {f. (∃c>0. eventually (λx. norm (f x) ≥ c * norm (g x)) F)}"
definition smallomega :: "'a filter ⇒ ('a ⇒ ('b :: real_normed_field)) ⇒ ('a ⇒ 'b) set"
(‹(1ω[_]'(_'))›)
where "smallomega F g = {f. (∀c>0. eventually (λx. norm (f x) ≥ c * norm (g x)) F)}"
definition bigtheta :: "'a filter ⇒ ('a ⇒ ('b :: real_normed_field)) ⇒ ('a ⇒ 'b) set"
(‹(1Θ[_]'(_'))›)
where "bigtheta F g = bigo F g ∩ bigomega F g"
abbreviation bigo_at_top (‹(2O'(_'))›) where
"O(g) ≡ bigo at_top g"
abbreviation smallo_at_top (‹(2o'(_'))›) where
"o(g) ≡ smallo at_top g"
abbreviation bigomega_at_top (‹(2Ω'(_'))›) where
"Ω(g) ≡ bigomega at_top g"
abbreviation smallomega_at_top (‹(2ω'(_'))›) where
"ω(g) ≡ smallomega at_top g"
abbreviation bigtheta_at_top (‹(2Θ'(_'))›) where
"Θ(g) ≡ bigtheta at_top g"
text ‹The following is a set of properties that all Landau symbols satisfy.›
named_theorems landau_divide_simps
locale landau_symbol =
fixes L :: "'a filter ⇒ ('a ⇒ ('b :: real_normed_field)) ⇒ ('a ⇒ 'b) set"
and L' :: "'c filter ⇒ ('c ⇒ ('b :: real_normed_field)) ⇒ ('c ⇒ 'b) set"
and Lr :: "'a filter ⇒ ('a ⇒ real) ⇒ ('a ⇒ real) set"
assumes bot': "L bot f = UNIV"
assumes filter_mono': "F1 ≤ F2 ⟹ L F2 f ⊆ L F1 f"
assumes in_filtermap_iff:
"f' ∈ L (filtermap h' F') g' ⟷ (λx. f' (h' x)) ∈ L' F' (λx. g' (h' x))"
assumes filtercomap:
"f' ∈ L F'' g' ⟹ (λx. f' (h' x)) ∈ L' (filtercomap h' F'') (λx. g' (h' x))"
assumes sup: "f ∈ L F1 g ⟹ f ∈ L F2 g ⟹ f ∈ L (sup F1 F2) g"
assumes in_cong: "eventually (λx. f x = g x) F ⟹ f ∈ L F (h) ⟷ g ∈ L F (h)"
assumes cong: "eventually (λx. f x = g x) F ⟹ L F (f) = L F (g)"
assumes cong_bigtheta: "f ∈ Θ[F](g) ⟹ L F (f) = L F (g)"
assumes in_cong_bigtheta: "f ∈ Θ[F](g) ⟹ f ∈ L F (h) ⟷ g ∈ L F (h)"
assumes cmult [simp]: "c ≠ 0 ⟹ L F (λx. c * f x) = L F (f)"
assumes cmult_in_iff [simp]: "c ≠ 0 ⟹ (λx. c * f x) ∈ L F (g) ⟷ f ∈ L F (g)"
assumes mult_left [simp]: "f ∈ L F (g) ⟹ (λx. h x * f x) ∈ L F (λx. h x * g x)"
assumes inverse: "eventually (λx. f x ≠ 0) F ⟹ eventually (λx. g x ≠ 0) F ⟹
f ∈ L F (g) ⟹ (λx. inverse (g x)) ∈ L F (λx. inverse (f x))"
assumes subsetI: "f ∈ L F (g) ⟹ L F (f) ⊆ L F (g)"
assumes plus_subset1: "f ∈ o[F](g) ⟹ L F (g) ⊆ L F (λx. f x + g x)"
assumes trans: "f ∈ L F (g) ⟹ g ∈ L F (h) ⟹ f ∈ L F (h)"
assumes compose: "f ∈ L F (g) ⟹ filterlim h' F G ⟹ (λx. f (h' x)) ∈ L' G (λx. g (h' x))"
assumes norm_iff [simp]: "(λx. norm (f x)) ∈ Lr F (λx. norm (g x)) ⟷ f ∈ L F (g)"
assumes abs [simp]: "Lr Fr (λx. ¦fr x¦) = Lr Fr fr"
assumes abs_in_iff [simp]: "(λx. ¦fr x¦) ∈ Lr Fr gr ⟷ fr ∈ Lr Fr gr"
begin
lemma bot [simp]: "f ∈ L bot g" by (simp add: bot')
lemma filter_mono: "F1 ≤ F2 ⟹ f ∈ L F2 g ⟹ f ∈ L F1 g"
using filter_mono'[of F1 F2] by blast
lemma cong_ex:
"eventually (λx. f1 x = f2 x) F ⟹ eventually (λx. g1 x = g2 x) F ⟹
f1 ∈ L F (g1) ⟷ f2 ∈ L F (g2)"
by (subst cong, assumption, subst in_cong, assumption, rule refl)
lemma cong_ex_bigtheta:
"f1 ∈ Θ[F](f2) ⟹ g1 ∈ Θ[F](g2) ⟹ f1 ∈ L F (g1) ⟷ f2 ∈ L F (g2)"
by (subst cong_bigtheta, assumption, subst in_cong_bigtheta, assumption, rule refl)
lemma bigtheta_trans1:
"f ∈ L F (g) ⟹ g ∈ Θ[F](h) ⟹ f ∈ L F (h)"
by (subst cong_bigtheta[symmetric])
lemma bigtheta_trans2:
"f ∈ Θ[F](g) ⟹ g ∈ L F (h) ⟹ f ∈ L F (h)"
by (subst in_cong_bigtheta)
lemma cmult' [simp]: "c ≠ 0 ⟹ L F (λx. f x * c) = L F (f)"
by (subst mult.commute) (rule cmult)
lemma cmult_in_iff' [simp]: "c ≠ 0 ⟹ (λx. f x * c) ∈ L F (g) ⟷ f ∈ L F (g)"
by (subst mult.commute) (rule cmult_in_iff)
lemma cdiv [simp]: "c ≠ 0 ⟹ L F (λx. f x / c) = L F (f)"
using cmult'[of "inverse c" F f] by (simp add: field_simps)
lemma cdiv_in_iff' [simp]: "c ≠ 0 ⟹ (λx. f x / c) ∈ L F (g) ⟷ f ∈ L F (g)"
using cmult_in_iff'[of "inverse c" f] by (simp add: field_simps)
lemma uminus [simp]: "L F (λx. -g x) = L F (g)" using cmult[of "-1"] by simp
lemma uminus_in_iff [simp]: "(λx. -f x) ∈ L F (g) ⟷ f ∈ L F (g)"
using cmult_in_iff[of "-1"] by simp
lemma const: "c ≠ 0 ⟹ L F (λ_. c) = L F (λ_. 1)"
by (subst (2) cmult[symmetric]) simp_all
lemma const' [simp]: "NO_MATCH 1 c ⟹ c ≠ 0 ⟹ L F (λ_. c) = L F (λ_. 1)"
by (rule const)
lemma const_in_iff: "c ≠ 0 ⟹ (λ_. c) ∈ L F (f) ⟷ (λ_. 1) ∈ L F (f)"
using cmult_in_iff'[of c "λ_. 1"] by simp
lemma const_in_iff' [simp]: "NO_MATCH 1 c ⟹ c ≠ 0 ⟹ (λ_. c) ∈ L F (f) ⟷ (λ_. 1) ∈ L F (f)"
by (rule const_in_iff)
lemma plus_subset2: "g ∈ o[F](f) ⟹ L F (f) ⊆ L F (λx. f x + g x)"
by (subst add.commute) (rule plus_subset1)
lemma mult_right [simp]: "f ∈ L F (g) ⟹ (λx. f x * h x) ∈ L F (λx. g x * h x)"
using mult_left by (simp add: mult.commute)
lemma mult: "f1 ∈ L F (g1) ⟹ f2 ∈ L F (g2) ⟹ (λx. f1 x * f2 x) ∈ L F (λx. g1 x * g2 x)"
by (rule trans, erule mult_left, erule mult_right)
lemma inverse_cancel:
assumes "eventually (λx. f x ≠ 0) F"
assumes "eventually (λx. g x ≠ 0) F"
shows "(λx. inverse (f x)) ∈ L F (λx. inverse (g x)) ⟷ g ∈ L F (f)"
proof
assume "(λx. inverse (f x)) ∈ L F (λx. inverse (g x))"
from inverse[OF _ _ this] assms show "g ∈ L F (f)" by simp
qed (intro inverse assms)
lemma divide_right:
assumes "eventually (λx. h x ≠ 0) F"
assumes "f ∈ L F (g)"
shows "(λx. f x / h x) ∈ L F (λx. g x / h x)"
by (subst (1 2) divide_inverse) (intro mult_right inverse assms)
lemma divide_right_iff:
assumes "eventually (λx. h x ≠ 0) F"
shows "(λx. f x / h x) ∈ L F (λx. g x / h x) ⟷ f ∈ L F (g)"
proof
assume "(λx. f x / h x) ∈ L F (λx. g x / h x)"
from mult_right[OF this, of h] assms show "f ∈ L F (g)"
by (subst (asm) cong_ex[of _ f F _ g]) (auto elim!: eventually_mono)
qed (simp add: divide_right assms)
lemma divide_left:
assumes "eventually (λx. f x ≠ 0) F"
assumes "eventually (λx. g x ≠ 0) F"
assumes "g ∈ L F(f)"
shows "(λx. h x / f x) ∈ L F (λx. h x / g x)"
by (subst (1 2) divide_inverse) (intro mult_left inverse assms)
lemma divide_left_iff:
assumes "eventually (λx. f x ≠ 0) F"
assumes "eventually (λx. g x ≠ 0) F"
assumes "eventually (λx. h x ≠ 0) F"
shows "(λx. h x / f x) ∈ L F (λx. h x / g x) ⟷ g ∈ L F (f)"
proof
assume A: "(λx. h x / f x) ∈ L F (λx. h x / g x)"
from assms have B: "eventually (λx. h x / f x / h x = inverse (f x)) F"
by eventually_elim (simp add: divide_inverse)
from assms have C: "eventually (λx. h x / g x / h x = inverse (g x)) F"
by eventually_elim (simp add: divide_inverse)
from divide_right[OF assms(3) A] assms show "g ∈ L F (f)"
by (subst (asm) cong_ex[OF B C]) (simp add: inverse_cancel)
qed (simp add: divide_left assms)
lemma divide:
assumes "eventually (λx. g1 x ≠ 0) F"
assumes "eventually (λx. g2 x ≠ 0) F"
assumes "f1 ∈ L F (f2)" "g2 ∈ L F (g1)"
shows "(λx. f1 x / g1 x) ∈ L F (λx. f2 x / g2 x)"
by (subst (1 2) divide_inverse) (intro mult inverse assms)
lemma divide_eq1:
assumes "eventually (λx. h x ≠ 0) F"
shows "f ∈ L F (λx. g x / h x) ⟷ (λx. f x * h x) ∈ L F (g)"
proof-
have "f ∈ L F (λx. g x / h x) ⟷ (λx. f x * h x / h x) ∈ L F (λx. g x / h x)"
using assms by (intro in_cong) (auto elim: eventually_mono)
thus ?thesis by (simp only: divide_right_iff assms)
qed
lemma divide_eq2:
assumes "eventually (λx. h x ≠ 0) F"
shows "(λx. f x / h x) ∈ L F (λx. g x) ⟷ f ∈ L F (λx. g x * h x)"
proof-
have "L F (λx. g x) = L F (λx. g x * h x / h x)"
using assms by (intro cong) (auto elim: eventually_mono)
thus ?thesis by (simp only: divide_right_iff assms)
qed
lemma inverse_eq1:
assumes "eventually (λx. g x ≠ 0) F"
shows "f ∈ L F (λx. inverse (g x)) ⟷ (λx. f x * g x) ∈ L F (λ_. 1)"
using divide_eq1[of g F f "λ_. 1"] by (simp add: divide_inverse assms)
lemma inverse_eq2:
assumes "eventually (λx. f x ≠ 0) F"
shows "(λx. inverse (f x)) ∈ L F (g) ⟷ (λx. 1) ∈ L F (λx. f x * g x)"
using divide_eq2[of f F "λ_. 1" g] by (simp add: divide_inverse assms mult_ac)
lemma inverse_flip:
assumes "eventually (λx. g x ≠ 0) F"
assumes "eventually (λx. h x ≠ 0) F"
assumes "(λx. inverse (g x)) ∈ L F (h)"
shows "(λx. inverse (h x)) ∈ L F (g)"
using assms by (simp add: divide_eq1 divide_eq2 inverse_eq_divide mult.commute)
lemma lift_trans:
assumes "f ∈ L F (g)"
assumes "(λx. t x (g x)) ∈ L F (h)"
assumes "⋀f g. f ∈ L F (g) ⟹ (λx. t x (f x)) ∈ L F (λx. t x (g x))"
shows "(λx. t x (f x)) ∈ L F (h)"
by (rule trans[OF assms(3)[OF assms(1)] assms(2)])
lemma lift_trans':
assumes "f ∈ L F (λx. t x (g x))"
assumes "g ∈ L F (h)"
assumes "⋀g h. g ∈ L F (h) ⟹ (λx. t x (g x)) ∈ L F (λx. t x (h x))"
shows "f ∈ L F (λx. t x (h x))"
by (rule trans[OF assms(1) assms(3)[OF assms(2)]])
lemma lift_trans_bigtheta:
assumes "f ∈ L F (g)"
assumes "(λx. t x (g x)) ∈ Θ[F](h)"
assumes "⋀f g. f ∈ L F (g) ⟹ (λx. t x (f x)) ∈ L F (λx. t x (g x))"
shows "(λx. t x (f x)) ∈ L F (h)"
using cong_bigtheta[OF assms(2)] assms(3)[OF assms(1)] by simp
lemma lift_trans_bigtheta':
assumes "f ∈ L F (λx. t x (g x))"
assumes "g ∈ Θ[F](h)"
assumes "⋀g h. g ∈ Θ[F](h) ⟹ (λx. t x (g x)) ∈ Θ[F](λx. t x (h x))"
shows "f ∈ L F (λx. t x (h x))"
using cong_bigtheta[OF assms(3)[OF assms(2)]] assms(1) by simp
lemma (in landau_symbol) mult_in_1:
assumes "f ∈ L F (λ_. 1)" "g ∈ L F (λ_. 1)"
shows "(λx. f x * g x) ∈ L F (λ_. 1)"
using mult[OF assms] by simp
lemma (in landau_symbol) of_real_cancel:
"(λx. of_real (f x)) ∈ L F (λx. of_real (g x)) ⟹ f ∈ Lr F g"
by (subst (asm) norm_iff [symmetric], subst (asm) (1 2) norm_of_real) simp_all
lemma (in landau_symbol) of_real_iff:
"(λx. of_real (f x)) ∈ L F (λx. of_real (g x)) ⟷ f ∈ Lr F g"
by (subst norm_iff [symmetric], subst (1 2) norm_of_real) simp_all
lemmas [landau_divide_simps] =
inverse_cancel divide_left_iff divide_eq1 divide_eq2 inverse_eq1 inverse_eq2
end
text ‹
The symbols $O$ and $o$ and $\Omega$ and $\omega$ are dual, so for many rules, replacing $O$ with
$\Omega$, $o$ with $\omega$, and $\leq$ with $\geq$ in a theorem yields another valid theorem.
The following locale captures this fact.
›
locale landau_pair =
fixes L l :: "'a filter ⇒ ('a ⇒ ('b :: real_normed_field)) ⇒ ('a ⇒ 'b) set"
fixes L' l' :: "'c filter ⇒ ('c ⇒ ('b :: real_normed_field)) ⇒ ('c ⇒ 'b) set"
fixes Lr lr :: "'a filter ⇒ ('a ⇒ real) ⇒ ('a ⇒ real) set"
and R :: "real ⇒ real ⇒ bool"
assumes L_def: "L F g = {f. ∃c>0. eventually (λx. R (norm (f x)) (c * norm (g x))) F}"
and l_def: "l F g = {f. ∀c>0. eventually (λx. R (norm (f x)) (c * norm (g x))) F}"
and L'_def: "L' F' g' = {f. ∃c>0. eventually (λx. R (norm (f x)) (c * norm (g' x))) F'}"
and l'_def: "l' F' g' = {f. ∀c>0. eventually (λx. R (norm (f x)) (c * norm (g' x))) F'}"
and Lr_def: "Lr F'' g'' = {f. ∃c>0. eventually (λx. R (norm (f x)) (c * norm (g'' x))) F''}"
and lr_def: "lr F'' g'' = {f. ∀c>0. eventually (λx. R (norm (f x)) (c * norm (g'' x))) F''}"
and R: "R = (≤) ∨ R = (≥)"
interpretation landau_o:
landau_pair bigo smallo bigo smallo bigo smallo "(≤)"
by unfold_locales (auto simp: bigo_def smallo_def intro!: ext)
interpretation landau_omega:
landau_pair bigomega smallomega bigomega smallomega bigomega smallomega "(≥)"
by unfold_locales (auto simp: bigomega_def smallomega_def intro!: ext)
context landau_pair
begin
lemmas R_E = disjE [OF R, case_names le ge]
lemma bigI:
"c > 0 ⟹ eventually (λx. R (norm (f x)) (c * norm (g x))) F ⟹ f ∈ L F (g)"
unfolding L_def by blast
lemma bigE:
assumes "f ∈ L F (g)"
obtains c where "c > 0" "eventually (λx. R (norm (f x)) (c * (norm (g x)))) F"
using assms unfolding L_def by blast
lemma smallI:
"(⋀c. c > 0 ⟹ eventually (λx. R (norm (f x)) (c * (norm (g x)))) F) ⟹ f ∈ l F (g)"
unfolding l_def by blast
lemma smallD:
"f ∈ l F (g) ⟹ c > 0 ⟹ eventually (λx. R (norm (f x)) (c * (norm (g x)))) F"
unfolding l_def by blast
lemma bigE_nonneg_real:
assumes "f ∈ Lr F (g)" "eventually (λx. f x ≥ 0) F"
obtains c where "c > 0" "eventually (λx. R (f x) (c * ¦g x¦)) F"
proof-
from assms(1) obtain c where c: "c > 0" "eventually (λx. R (norm (f x)) (c * norm (g x))) F"
by (auto simp: Lr_def)
from c(2) assms(2) have "eventually (λx. R (f x) (c * ¦g x¦)) F"
by eventually_elim simp
from c(1) and this show ?thesis by (rule that)
qed
lemma smallD_nonneg_real:
assumes "f ∈ lr F (g)" "eventually (λx. f x ≥ 0) F" "c > 0"
shows "eventually (λx. R (f x) (c * ¦g x¦)) F"
using assms by (auto simp: lr_def dest!: spec[of _ c] elim: eventually_elim2)
lemma small_imp_big: "f ∈ l F (g) ⟹ f ∈ L F (g)"
by (rule bigI[OF _ smallD, of 1]) simp_all
lemma small_subset_big: "l F (g) ⊆ L F (g)"
using small_imp_big by blast
lemma R_refl [simp]: "R x x" using R by auto
lemma R_linear: "¬R x y ⟹ R y x"
using R by auto
lemma R_trans [trans]: "R a b ⟹ R b c ⟹ R a c"
using R by auto
lemma R_mult_left_mono: "R a b ⟹ c ≥ 0 ⟹ R (c*a) (c*b)"
using R by (auto simp: mult_left_mono)
lemma R_mult_right_mono: "R a b ⟹ c ≥ 0 ⟹ R (a*c) (b*c)"
using R by (auto simp: mult_right_mono)
lemma big_trans:
assumes "f ∈ L F (g)" "g ∈ L F (h)"
shows "f ∈ L F (h)"
proof-
from assms obtain c d where *: "0 < c" "0 < d"
and **: "∀⇩F x in F. R (norm (f x)) (c * norm (g x))"
"∀⇩F x in F. R (norm (g x)) (d * norm (h x))"
by (elim bigE)
from ** have "eventually (λx. R (norm (f x)) (c * d * (norm (h x)))) F"
proof eventually_elim
fix x assume "R (norm (f x)) (c * (norm (g x)))"
also assume "R (norm (g x)) (d * (norm (h x)))"
with ‹0 < c› have "R (c * (norm (g x))) (c * (d * (norm (h x))))"
by (intro R_mult_left_mono) simp_all
finally show "R (norm (f x)) (c * d * (norm (h x)))"
by (simp add: algebra_simps)
qed
with * show ?thesis by (intro bigI[of "c*d"]) simp_all
qed
lemma big_small_trans:
assumes "f ∈ L F (g)" "g ∈ l F (h)"
shows "f ∈ l F (h)"
proof (rule smallI)
fix c :: real assume c: "c > 0"
from assms(1) obtain d where d: "d > 0" and *: "∀⇩F x in F. R (norm (f x)) (d * norm (g x))"
by (elim bigE)
from assms(2) c d have "eventually (λx. R (norm (g x)) (c * inverse d * norm (h x))) F"
by (intro smallD) simp_all
with * show "eventually (λx. R (norm (f x)) (c * (norm (h x)))) F"
proof eventually_elim
case (elim x)
show ?case
by (use elim(1) in ‹rule R_trans›) (use elim(2) R d in ‹auto simp: field_simps›)
qed
qed
lemma small_big_trans:
assumes "f ∈ l F (g)" "g ∈ L F (h)"
shows "f ∈ l F (h)"
proof (rule smallI)
fix c :: real assume c: "c > 0"
from assms(2) obtain d where d: "d > 0" and *: "∀⇩F x in F. R (norm (g x)) (d * norm (h x))"
by (elim bigE)
from assms(1) c d have "eventually (λx. R (norm (f x)) (c * inverse d * norm (g x))) F"
by (intro smallD) simp_all
with * show "eventually (λx. R (norm (f x)) (c * norm (h x))) F"
by eventually_elim (rotate_tac 2, erule R_trans, insert R c d, auto simp: field_simps)
qed
lemma small_trans:
"f ∈ l F (g) ⟹ g ∈ l F (h) ⟹ f ∈ l F (h)"
by (rule big_small_trans[OF small_imp_big])
lemma small_big_trans':
"f ∈ l F (g) ⟹ g ∈ L F (h) ⟹ f ∈ L F (h)"
by (rule small_imp_big[OF small_big_trans])
lemma big_small_trans':
"f ∈ L F (g) ⟹ g ∈ l F (h) ⟹ f ∈ L F (h)"
by (rule small_imp_big[OF big_small_trans])
lemma big_subsetI [intro]: "f ∈ L F (g) ⟹ L F (f) ⊆ L F (g)"
by (intro subsetI) (drule (1) big_trans)
lemma small_subsetI [intro]: "f ∈ L F (g) ⟹ l F (f) ⊆ l F (g)"
by (intro subsetI) (drule (1) small_big_trans)
lemma big_refl [simp]: "f ∈ L F (f)"
by (rule bigI[of 1]) simp_all
lemma small_refl_iff: "f ∈ l F (f) ⟷ eventually (λx. f x = 0) F"
proof (rule iffI[OF _ smallI])
assume f: "f ∈ l F f"
have "(1/2::real) > 0" "(2::real) > 0" by simp_all
from smallD[OF f this(1)] smallD[OF f this(2)]
show "eventually (λx. f x = 0) F" by eventually_elim (insert R, auto)
next
fix c :: real assume "c > 0" "eventually (λx. f x = 0) F"
from this(2) show "eventually (λx. R (norm (f x)) (c * norm (f x))) F"
by eventually_elim simp_all
qed
lemma big_small_asymmetric: "f ∈ L F (g) ⟹ g ∈ l F (f) ⟹ eventually (λx. f x = 0) F"
by (drule (1) big_small_trans) (simp add: small_refl_iff)
lemma small_big_asymmetric: "f ∈ l F (g) ⟹ g ∈ L F (f) ⟹ eventually (λx. f x = 0) F"
by (drule (1) small_big_trans) (simp add: small_refl_iff)
lemma small_asymmetric: "f ∈ l F (g) ⟹ g ∈ l F (f) ⟹ eventually (λx. f x = 0) F"
by (drule (1) small_trans) (simp add: small_refl_iff)
lemma plus_aux:
assumes "f ∈ o[F](g)"
shows "g ∈ L F (λx. f x + g x)"
proof (rule R_E)
assume R: "R = (≤)"
have A: "1/2 > (0::real)" by simp
have B: "1/2 * (norm (g x)) ≤ norm (f x + g x)"
if "norm (f x) ≤ 1/2 * norm (g x)" for x
proof -
from that have "1/2 * (norm (g x)) ≤ (norm (g x)) - (norm (f x))"
by simp
also have "norm (g x) - norm (f x) ≤ norm (f x + g x)"
by (subst add.commute) (rule norm_diff_ineq)
finally show ?thesis by simp
qed
show "g ∈ L F (λx. f x + g x)"
apply (rule bigI[of "2"])
apply simp
apply (use landau_o.smallD[OF assms A] in eventually_elim)
apply (use B in ‹simp add: R algebra_simps›)
done
next
assume R: "R = (λx y. x ≥ y)"
show "g ∈ L F (λx. f x + g x)"
proof (rule bigI[of "1/2"])
show "eventually (λx. R (norm (g x)) (1/2 * norm (f x + g x))) F"
using landau_o.smallD[OF assms zero_less_one]
proof eventually_elim
case (elim x)
have "norm (f x + g x) ≤ norm (f x) + norm (g x)"
by (rule norm_triangle_ineq)
also note elim
finally show ?case by (simp add: R)
qed
qed simp_all
qed
end
lemma summable_comparison_test_bigo:
fixes f :: "nat ⇒ real"
assumes "summable (λn. norm (g n))" "f ∈ O(g)"
shows "summable f"
proof -
from ‹f ∈ O(g)› obtain C where C: "eventually (λx. norm (f x) ≤ C * norm (g x)) at_top"
by (auto elim: landau_o.bigE)
thus ?thesis
by (rule summable_comparison_test_ev) (insert assms, auto intro: summable_mult)
qed
lemma bigomega_iff_bigo: "g ∈ Ω[F](f) ⟷ f ∈ O[F](g)"
proof
assume "f ∈ O[F](g)"
then obtain c where "0 < c" "∀⇩F x in F. norm (f x) ≤ c * norm (g x)"
by (rule landau_o.bigE)
then show "g ∈ Ω[F](f)"
by (intro landau_omega.bigI[of "inverse c"]) (simp_all add: field_simps)
next
assume "g ∈ Ω[F](f)"
then obtain c where "0 < c" "∀⇩F x in F. c * norm (f x) ≤ norm (g x)"
by (rule landau_omega.bigE)
then show "f ∈ O[F](g)"
by (intro landau_o.bigI[of "inverse c"]) (simp_all add: field_simps)
qed
lemma smallomega_iff_smallo: "g ∈ ω[F](f) ⟷ f ∈ o[F](g)"
proof
assume "f ∈ o[F](g)"
from landau_o.smallD[OF this, of "inverse c" for c]
show "g ∈ ω[F](f)" by (intro landau_omega.smallI) (simp_all add: field_simps)
next
assume "g ∈ ω[F](f)"
from landau_omega.smallD[OF this, of "inverse c" for c]
show "f ∈ o[F](g)" by (intro landau_o.smallI) (simp_all add: field_simps)
qed
context landau_pair
begin
lemma big_mono:
"eventually (λx. R (norm (f x)) (norm (g x))) F ⟹ f ∈ L F (g)"
by (rule bigI[OF zero_less_one]) simp
lemma big_mult:
assumes "f1 ∈ L F (g1)" "f2 ∈ L F (g2)"
shows "(λx. f1 x * f2 x) ∈ L F (λx. g1 x * g2 x)"
proof-
from assms obtain c1 c2 where *: "c1 > 0" "c2 > 0"
and **: "∀⇩F x in F. R (norm (f1 x)) (c1 * norm (g1 x))"
"∀⇩F x in F. R (norm (f2 x)) (c2 * norm (g2 x))"
by (elim bigE)
from * have "c1 * c2 > 0" by simp
moreover have "eventually (λx. R (norm (f1 x * f2 x)) (c1 * c2 * norm (g1 x * g2 x))) F"
using **
proof eventually_elim
case (elim x)
show ?case
proof (cases rule: R_E)
case le
have "norm (f1 x) * norm (f2 x) ≤ (c1 * norm (g1 x)) * (c2 * norm (g2 x))"
using elim le * by (intro mult_mono mult_nonneg_nonneg) auto
with le show ?thesis by (simp add: le norm_mult mult_ac)
next
case ge
have "(c1 * norm (g1 x)) * (c2 * norm (g2 x)) ≤ norm (f1 x) * norm (f2 x)"
using elim ge * by (intro mult_mono mult_nonneg_nonneg) auto
with ge show ?thesis by (simp_all add: norm_mult mult_ac)
qed
qed
ultimately show ?thesis by (rule bigI)
qed
lemma small_big_mult:
assumes "f1 ∈ l F (g1)" "f2 ∈ L F (g2)"
shows "(λx. f1 x * f2 x) ∈ l F (λx. g1 x * g2 x)"
proof (rule smallI)
fix c1 :: real assume c1: "c1 > 0"
from assms(2) obtain c2 where c2: "c2 > 0"
and *: "∀⇩F x in F. R (norm (f2 x)) (c2 * norm (g2 x))" by (elim bigE)
from assms(1) c1 c2 have "eventually (λx. R (norm (f1 x)) (c1 * inverse c2 * norm (g1 x))) F"
by (auto intro!: smallD)
with * show "eventually (λx. R (norm (f1 x * f2 x)) (c1 * norm (g1 x * g2 x))) F"
proof eventually_elim
case (elim x)
show ?case
proof (cases rule: R_E)
case le
have "norm (f1 x) * norm (f2 x) ≤ (c1 * inverse c2 * norm (g1 x)) * (c2 * norm (g2 x))"
using elim le c1 c2 by (intro mult_mono mult_nonneg_nonneg) auto
with le c2 show ?thesis by (simp add: le norm_mult field_simps)
next
case ge
have "norm (f1 x) * norm (f2 x) ≥ (c1 * inverse c2 * norm (g1 x)) * (c2 * norm (g2 x))"
using elim ge c1 c2 by (intro mult_mono mult_nonneg_nonneg) auto
with ge c2 show ?thesis by (simp add: ge norm_mult field_simps)
qed
qed
qed
lemma big_small_mult:
"f1 ∈ L F (g1) ⟹ f2 ∈ l F (g2) ⟹ (λx. f1 x * f2 x) ∈ l F (λx. g1 x * g2 x)"
by (subst (1 2) mult.commute) (rule small_big_mult)
lemma small_mult: "f1 ∈ l F (g1) ⟹ f2 ∈ l F (g2) ⟹ (λx. f1 x * f2 x) ∈ l F (λx. g1 x * g2 x)"
by (rule small_big_mult, assumption, rule small_imp_big)
lemmas mult = big_mult small_big_mult big_small_mult small_mult
lemma big_power:
assumes "f ∈ L F (g)"
shows "(λx. f x ^ m) ∈ L F (λx. g x ^ m)"
using assms by (induction m) (auto intro: mult)
lemma (in landau_pair) small_power:
assumes "f ∈ l F (g)" "m > 0"
shows "(λx. f x ^ m) ∈ l F (λx. g x ^ m)"
proof -
have "(λx. f x * f x ^ (m - 1)) ∈ l F (λx. g x * g x ^ (m - 1))"
by (intro small_big_mult assms big_power[OF small_imp_big])
thus ?thesis
using assms by (cases m) (simp_all add: mult_ac)
qed
lemma big_power_increasing:
assumes "(λ_. 1) ∈ L F f" "m ≤ n"
shows "(λx. f x ^ m) ∈ L F (λx. f x ^ n)"
proof -
have "(λx. f x ^ m * 1 ^ (n - m)) ∈ L F (λx. f x ^ m * f x ^ (n - m))"
using assms by (intro mult big_power) auto
also have "(λx. f x ^ m * f x ^ (n - m)) = (λx. f x ^ (m + (n - m)))"
by (subst power_add [symmetric]) (rule refl)
also have "m + (n - m) = n"
using assms by simp
finally show ?thesis by simp
qed
lemma small_power_increasing:
assumes "(λ_. 1) ∈ l F f" "m < n"
shows "(λx. f x ^ m) ∈ l F (λx. f x ^ n)"
proof -
note [trans] = small_big_trans
have "(λx. f x ^ m * 1) ∈ l F (λx. f x ^ m * f x)"
using assms by (intro big_small_mult) auto
also have "(λx. f x ^ m * f x) = (λx. f x ^ Suc m)"
by (simp add: mult_ac)
also have "… ∈ L F (λx. f x ^ n)"
using assms by (intro big_power_increasing[OF small_imp_big]) auto
finally show ?thesis by simp
qed
sublocale big: landau_symbol L L' Lr
proof
have L: "L = bigo ∨ L = bigomega"
by (rule R_E) (auto simp: bigo_def L_def bigomega_def fun_eq_iff)
have A: "(λx. c * f x) ∈ L F f" if "c ≠ 0" for c :: 'b and F and f :: "'a ⇒ 'b"
using that by (intro bigI[of "norm c"]) (simp_all add: norm_mult)
show "L F (λx. c * f x) = L F f" if "c ≠ 0" for c :: 'b and F and f :: "'a ⇒ 'b"
using ‹c ≠ 0› and A[of c f] and A[of "inverse c" "λx. c * f x"]
by (intro equalityI big_subsetI) (simp_all add: field_simps)
show "((λx. c * f x) ∈ L F g) = (f ∈ L F g)" if "c ≠ 0"
for c :: 'b and F and f g :: "'a ⇒ 'b"
proof -
from ‹c ≠ 0› and A[of c f] and A[of "inverse c" "λx. c * f x"]
have "(λx. c * f x) ∈ L F f" "f ∈ L F (λx. c * f x)"
by (simp_all add: field_simps)
then show ?thesis by (intro iffI) (erule (1) big_trans)+
qed
show "(λx. inverse (g x)) ∈ L F (λx. inverse (f x))"
if *: "f ∈ L F (g)" and **: "eventually (λx. f x ≠ 0) F" "eventually (λx. g x ≠ 0) F"
for f g :: "'a ⇒ 'b" and F
proof -
from * obtain c where c: "c > 0" and ***: "∀⇩F x in F. R (norm (f x)) (c * norm (g x))"
by (elim bigE)
from ** *** have "eventually (λx. R (norm (inverse (g x))) (c * norm (inverse (f x)))) F"
by eventually_elim (rule R_E, simp_all add: field_simps norm_inverse norm_divide c)
with c show ?thesis by (rule bigI)
qed
show "L F g ⊆ L F (λx. f x + g x)" if "f ∈ o[F](g)" for f g :: "'a ⇒ 'b" and F
using plus_aux that by (blast intro!: big_subsetI)
show "L F (f) = L F (g)" if "eventually (λx. f x = g x) F" for f g :: "'a ⇒ 'b" and F
unfolding L_def by (subst eventually_subst'[OF that]) (rule refl)
show "f ∈ L F (h) ⟷ g ∈ L F (h)" if "eventually (λx. f x = g x) F"
for f g h :: "'a ⇒ 'b" and F
unfolding L_def mem_Collect_eq
by (subst (1) eventually_subst'[OF that]) (rule refl)
show "L F f ⊆ L F g" if "f ∈ L F g" for f g :: "'a ⇒ 'b" and F
using that by (rule big_subsetI)
show "L F (f) = L F (g)" if "f ∈ Θ[F](g)" for f g :: "'a ⇒ 'b" and F
using L that unfolding bigtheta_def
by (intro equalityI big_subsetI) (auto simp: bigomega_iff_bigo)
show "f ∈ L F (h) ⟷ g ∈ L F (h)" if "f ∈ Θ[F](g)" for f g h :: "'a ⇒ 'b" and F
by (rule disjE[OF L])
(use that in ‹auto simp: bigtheta_def bigomega_iff_bigo intro: landau_o.big_trans›)
show "(λx. h x * f x) ∈ L F (λx. h x * g x)" if "f ∈ L F g" for f g h :: "'a ⇒ 'b" and F
using that by (intro big_mult) simp
show "f ∈ L F (h)" if "f ∈ L F g" "g ∈ L F h" for f g h :: "'a ⇒ 'b" and F
using that by (rule big_trans)
show "(λx. f (h x)) ∈ L' G (λx. g (h x))"
if "f ∈ L F g" and "filterlim h F G"
for f g :: "'a ⇒ 'b" and h :: "'c ⇒ 'a" and F G
using that by (auto simp: L_def L'_def filterlim_iff)
show "f ∈ L (sup F G) g" if "f ∈ L F g" "f ∈ L G g"
for f g :: "'a ⇒ 'b" and F G :: "'a filter"
proof -
from that [THEN bigE] obtain c1 c2
where *: "c1 > 0" "c2 > 0"
and **: "∀⇩F x in F. R (norm (f x)) (c1 * norm (g x))"
"∀⇩F x in G. R (norm (f x)) (c2 * norm (g x))" .
define c where "c = (if R c1 c2 then c2 else c1)"
from * have c: "R c1 c" "R c2 c" "c > 0"
by (auto simp: c_def dest: R_linear)
with ** have "eventually (λx. R (norm (f x)) (c * norm (g x))) F"
"eventually (λx. R (norm (f x)) (c * norm (g x))) G"
by (force elim: eventually_mono intro: R_trans[OF _ R_mult_right_mono])+
with c show "f ∈ L (sup F G) g"
by (auto simp: L_def eventually_sup)
qed
show "((λx. f (h x)) ∈ L' (filtercomap h F) (λx. g (h x)))" if "(f ∈ L F g)"
for f g :: "'a ⇒ 'b" and h :: "'c ⇒ 'a" and F G :: "'a filter"
using that unfolding L_def L'_def by auto
qed (auto simp: L_def Lr_def eventually_filtermap L'_def
intro: filter_leD exI[of _ "1::real"])
sublocale small: landau_symbol l l' lr
proof
have A: "(λx. c * f x) ∈ L F f" if "c ≠ 0" for c :: 'b and f :: "'a ⇒ 'b" and F
using that by (intro bigI[of "norm c"]) (simp_all add: norm_mult)
show "l F (λx. c * f x) = l F f" if "c ≠ 0" for c :: 'b and f :: "'a ⇒ 'b" and F
using that and A[of c f] and A[of "inverse c" "λx. c * f x"]
by (intro equalityI small_subsetI) (simp_all add: field_simps)
show "((λx. c * f x) ∈ l F g) = (f ∈ l F g)" if "c ≠ 0" for c :: 'b and f g :: "'a ⇒ 'b" and F
proof -
from that and A[of c f] and A[of "inverse c" "λx. c * f x"]
have "(λx. c * f x) ∈ L F f" "f ∈ L F (λx. c * f x)"
by (simp_all add: field_simps)
then show ?thesis
by (intro iffI) (erule (1) big_small_trans)+
qed
show "l F g ⊆ l F (λx. f x + g x)" if "f ∈ o[F](g)" for f g :: "'a ⇒ 'b" and F
using plus_aux that by (blast intro!: small_subsetI)
show "(λx. inverse (g x)) ∈ l F (λx. inverse (f x))"
if A: "f ∈ l F (g)" and B: "eventually (λx. f x ≠ 0) F" "eventually (λx. g x ≠ 0) F"
for f g :: "'a ⇒ 'b" and F
proof (rule smallI)
fix c :: real assume c: "c > 0"
from B smallD[OF A c]
show "eventually (λx. R (norm (inverse (g x))) (c * norm (inverse (f x)))) F"
by eventually_elim (rule R_E, simp_all add: field_simps norm_inverse norm_divide)
qed
show "l F (f) = l F (g)" if "eventually (λx. f x = g x) F" for f g :: "'a ⇒ 'b" and F
unfolding l_def by (subst eventually_subst'[OF that]) (rule refl)
show "f ∈ l F (h) ⟷ g ∈ l F (h)" if "eventually (λx. f x = g x) F" for f g h :: "'a ⇒ 'b" and F
unfolding l_def mem_Collect_eq by (subst (1) eventually_subst'[OF that]) (rule refl)
show "l F f ⊆ l F g" if "f ∈ l F g" for f g :: "'a ⇒ 'b" and F
using that by (intro small_subsetI small_imp_big)
show "l F (f) = l F (g)" if "f ∈ Θ[F](g)" for f g :: "'a ⇒ 'b" and F
proof -
have L: "L = bigo ∨ L = bigomega"
by (rule R_E) (auto simp: bigo_def L_def bigomega_def fun_eq_iff)
with that show ?thesis unfolding bigtheta_def
by (intro equalityI small_subsetI) (auto simp: bigomega_iff_bigo)
qed
show "f ∈ l F (h) ⟷ g ∈ l F (h)" if "f ∈ Θ[F](g)" for f g h :: "'a ⇒ 'b" and F
proof -
have l: "l = smallo ∨ l = smallomega"
by (rule R_E) (auto simp: smallo_def l_def smallomega_def fun_eq_iff)
show ?thesis
by (rule disjE[OF l])
(use that in ‹auto simp: bigtheta_def bigomega_iff_bigo smallomega_iff_smallo
intro: landau_o.big_small_trans landau_o.small_big_trans›)
qed
show "(λx. h x * f x) ∈ l F (λx. h x * g x)" if "f ∈ l F g" for f g h :: "'a ⇒ 'b" and F
using that by (intro big_small_mult) simp
show "f ∈ l F (h)" if "f ∈ l F g" "g ∈ l F h" for f g h :: "'a ⇒ 'b" and F
using that by (rule small_trans)
show "(λx. f (h x)) ∈ l' G (λx. g (h x))" if "f ∈ l F g" and "filterlim h F G"
for f g :: "'a ⇒ 'b" and h :: "'c ⇒ 'a" and F G
using that by (auto simp: l_def l'_def filterlim_iff)
show "((λx. f (h x)) ∈ l' (filtercomap h F) (λx. g (h x)))" if "f ∈ l F g"
for f g :: "'a ⇒ 'b" and h :: "'c ⇒ 'a" and F G :: "'a filter"
using that unfolding l_def l'_def by auto
qed (auto simp: l_def lr_def eventually_filtermap l'_def eventually_sup intro: filter_leD)
text ‹These rules allow chaining of Landau symbol propositions in Isar with "also".›
lemma big_mult_1: "f ∈ L F (g) ⟹ (λ_. 1) ∈ L F (h) ⟹ f ∈ L F (λx. g x * h x)"
and big_mult_1': "(λ_. 1) ∈ L F (g) ⟹ f ∈ L F (h) ⟹ f ∈ L F (λx. g x * h x)"
and small_mult_1: "f ∈ l F (g) ⟹ (λ_. 1) ∈ L F (h) ⟹ f ∈ l F (λx. g x * h x)"
and small_mult_1': "(λ_. 1) ∈ L F (g) ⟹ f ∈ l F (h) ⟹ f ∈ l F (λx. g x * h x)"
and small_mult_1'': "f ∈ L F (g) ⟹ (λ_. 1) ∈ l F (h) ⟹ f ∈ l F (λx. g x * h x)"
and small_mult_1''': "(λ_. 1) ∈ l F (g) ⟹ f ∈ L F (h) ⟹ f ∈ l F (λx. g x * h x)"
by (drule (1) big.mult big_small_mult small_big_mult, simp)+
lemma big_1_mult: "f ∈ L F (g) ⟹ h ∈ L F (λ_. 1) ⟹ (λx. f x * h x) ∈ L F (g)"
and big_1_mult': "h ∈ L F (λ_. 1) ⟹ f ∈ L F (g) ⟹ (λx. f x * h x) ∈ L F (g)"
and small_1_mult: "f ∈ l F (g) ⟹ h ∈ L F (λ_. 1) ⟹ (λx. f x * h x) ∈ l F (g)"
and small_1_mult': "h ∈ L F (λ_. 1) ⟹ f ∈ l F (g) ⟹ (λx. f x * h x) ∈ l F (g)"
and small_1_mult'': "f ∈ L F (g) ⟹ h ∈ l F (λ_. 1) ⟹ (λx. f x * h x) ∈ l F (g)"
and small_1_mult''': "h ∈ l F (λ_. 1) ⟹ f ∈ L F (g) ⟹ (λx. f x * h x) ∈ l F (g)"
by (drule (1) big.mult big_small_mult small_big_mult, simp)+
lemmas mult_1_trans =
big_mult_1 big_mult_1' small_mult_1 small_mult_1' small_mult_1'' small_mult_1'''
big_1_mult big_1_mult' small_1_mult small_1_mult' small_1_mult'' small_1_mult'''
lemma big_equal_iff_bigtheta: "L F (f) = L F (g) ⟷ f ∈ Θ[F](g)"
proof
have L: "L = bigo ∨ L = bigomega"
by (rule R_E) (auto simp: fun_eq_iff L_def bigo_def bigomega_def)
fix f g :: "'a ⇒ 'b" assume "L F (f) = L F (g)"
with big_refl[of f F] big_refl[of g F] have "f ∈ L F (g)" "g ∈ L F (f)" by simp_all
thus "f ∈ Θ[F](g)" using L unfolding bigtheta_def by (auto simp: bigomega_iff_bigo)
qed (rule big.cong_bigtheta)
lemma big_prod:
assumes "⋀x. x ∈ A ⟹ f x ∈ L F (g x)"
shows "(λy. ∏x∈A. f x y) ∈ L F (λy. ∏x∈A. g x y)"
using assms by (induction A rule: infinite_finite_induct) (auto intro!: big.mult)
lemma big_prod_in_1:
assumes "⋀x. x ∈ A ⟹ f x ∈ L F (λ_. 1)"
shows "(λy. ∏x∈A. f x y) ∈ L F (λ_. 1)"
using assms by (induction A rule: infinite_finite_induct) (auto intro!: big.mult_in_1)
end
context landau_symbol
begin
lemma plus_absorb1:
assumes "f ∈ o[F](g)"
shows "L F (λx. f x + g x) = L F (g)"
proof (intro equalityI)
from plus_subset1 and assms show "L F g ⊆ L F (λx. f x + g x)" .
from landau_o.small.plus_subset1[OF assms] and assms have "(λx. -f x) ∈ o[F](λx. f x + g x)"
by (auto simp: landau_o.small.uminus_in_iff)
from plus_subset1[OF this] show "L F (λx. f x + g x) ⊆ L F (g)" by simp
qed
lemma plus_absorb2: "g ∈ o[F](f) ⟹ L F (λx. f x + g x) = L F (f)"
using plus_absorb1[of g F f] by (simp add: add.commute)
lemma diff_absorb1: "f ∈ o[F](g) ⟹ L F (λx. f x - g x) = L F (g)"
by (simp only: diff_conv_add_uminus plus_absorb1 landau_o.small.uminus uminus)
lemma diff_absorb2: "g ∈ o[F](f) ⟹ L F (λx. f x - g x) = L F (f)"
by (simp only: diff_conv_add_uminus plus_absorb2 landau_o.small.uminus_in_iff)
lemmas absorb = plus_absorb1 plus_absorb2 diff_absorb1 diff_absorb2
end
lemma bigthetaI [intro]: "f ∈ O[F](g) ⟹ f ∈ Ω[F](g) ⟹ f ∈ Θ[F](g)"
unfolding bigtheta_def bigomega_def by blast
lemma bigthetaD1 [dest]: "f ∈ Θ[F](g) ⟹ f ∈ O[F](g)"
and bigthetaD2 [dest]: "f ∈ Θ[F](g) ⟹ f ∈ Ω[F](g)"
unfolding bigtheta_def bigo_def bigomega_def by blast+
lemma bigtheta_refl [simp]: "f ∈ Θ[F](f)"
unfolding bigtheta_def by simp
lemma bigtheta_sym: "f ∈ Θ[F](g) ⟷ g ∈ Θ[F](f)"
unfolding bigtheta_def by (auto simp: bigomega_iff_bigo)
lemmas landau_flip =
bigomega_iff_bigo[symmetric] smallomega_iff_smallo[symmetric]
bigomega_iff_bigo smallomega_iff_smallo bigtheta_sym
interpretation landau_theta: landau_symbol bigtheta bigtheta bigtheta
proof
fix f g :: "'a ⇒ 'b" and F
assume "f ∈ o[F](g)"
hence "O[F](g) ⊆ O[F](λx. f x + g x)" "Ω[F](g) ⊆ Ω[F](λx. f x + g x)"
by (rule landau_o.big.plus_subset1 landau_omega.big.plus_subset1)+
thus "Θ[F](g) ⊆ Θ[F](λx. f x + g x)" unfolding bigtheta_def by blast
next
fix f g :: "'a ⇒ 'b" and F
assume "f ∈ Θ[F](g)"
thus A: "Θ[F](f) = Θ[F](g)"
apply (subst (1 2) bigtheta_def)
apply (subst landau_o.big.cong_bigtheta landau_omega.big.cong_bigtheta, assumption)+
apply (rule refl)
done
thus "Θ[F](f) ⊆ Θ[F](g)" by simp
fix h :: "'a ⇒ 'b"
show "f ∈ Θ[F](h) ⟷ g ∈ Θ[F](h)" by (subst (1 2) bigtheta_sym) (simp add: A)
next
fix f g h :: "'a ⇒ 'b" and F
assume "f ∈ Θ[F](g)" "g ∈ Θ[F](h)"
thus "f ∈ Θ[F](h)" unfolding bigtheta_def
by (blast intro: landau_o.big.trans landau_omega.big.trans)
next
fix f :: "'a ⇒ 'b" and F1 F2 :: "'a filter"
assume "F1 ≤ F2"
thus "Θ[F2](f) ⊆ Θ[F1](f)"
by (auto simp: bigtheta_def intro: landau_o.big.filter_mono landau_omega.big.filter_mono)
qed (auto simp: bigtheta_def landau_o.big.norm_iff
landau_o.big.cmult landau_omega.big.cmult
landau_o.big.cmult_in_iff landau_omega.big.cmult_in_iff
landau_o.big.in_cong landau_omega.big.in_cong
landau_o.big.mult landau_omega.big.mult
landau_o.big.inverse landau_omega.big.inverse
landau_o.big.compose landau_omega.big.compose
landau_o.big.bot' landau_omega.big.bot'
landau_o.big.in_filtermap_iff landau_omega.big.in_filtermap_iff
landau_o.big.sup landau_omega.big.sup
landau_o.big.filtercomap landau_omega.big.filtercomap
dest: landau_o.big.cong landau_omega.big.cong)
lemmas landau_symbols =
landau_o.big.landau_symbol_axioms landau_o.small.landau_symbol_axioms
landau_omega.big.landau_symbol_axioms landau_omega.small.landau_symbol_axioms
landau_theta.landau_symbol_axioms
lemma bigoI [intro]:
assumes "eventually (λx. (norm (f x)) ≤ c * (norm (g x))) F"
shows "f ∈ O[F](g)"
proof (rule landau_o.bigI)
show "max 1 c > 0" by simp
have "c * (norm (g x)) ≤ max 1 c * (norm (g x))" for x
by (simp add: mult_right_mono)
with assms show "eventually (λx. (norm (f x)) ≤ max 1 c * (norm (g x))) F"
by (auto elim!: eventually_mono dest: order.trans)
qed
lemma smallomegaD [dest]:
assumes "f ∈ ω[F](g)"
shows "eventually (λx. (norm (f x)) ≥ c * (norm (g x))) F"
proof (cases "c > 0")
case False
show ?thesis
by (intro always_eventually allI, rule order.trans[of _ 0])
(insert False, auto intro!: mult_nonpos_nonneg)
qed (blast dest: landau_omega.smallD[OF assms, of c])
lemma bigthetaI':
assumes "c1 > 0" "c2 > 0"
assumes "eventually (λx. c1 * (norm (g x)) ≤ (norm (f x)) ∧ (norm (f x)) ≤ c2 * (norm (g x))) F"
shows "f ∈ Θ[F](g)"
apply (rule bigthetaI)
apply (rule landau_o.bigI[OF assms(2)]) using assms(3) apply (eventually_elim, simp)
apply (rule landau_omega.bigI[OF assms(1)]) using assms(3) apply (eventually_elim, simp)
done
lemma bigthetaI_cong: "eventually (λx. f x = g x) F ⟹ f ∈ Θ[F](g)"
by (intro bigthetaI'[of 1 1]) (auto elim!: eventually_mono)
lemma (in landau_symbol) ev_eq_trans1:
"f ∈ L F (λx. g x (h x)) ⟹ eventually (λx. h x = h' x) F ⟹ f ∈ L F (λx. g x (h' x))"
by (rule bigtheta_trans1[OF _ bigthetaI_cong]) (auto elim!: eventually_mono)
lemma (in landau_symbol) ev_eq_trans2:
"eventually (λx. f x = f' x) F ⟹ (λx. g x (f' x)) ∈ L F (h) ⟹ (λx. g x (f x)) ∈ L F (h)"
by (rule bigtheta_trans2[OF bigthetaI_cong]) (auto elim!: eventually_mono)
declare landau_o.smallI landau_omega.bigI landau_omega.smallI [intro]
declare landau_o.bigE landau_omega.bigE [elim]
declare landau_o.smallD
lemma (in landau_symbol) bigtheta_trans1':
"f ∈ L F (g) ⟹ h ∈ Θ[F](g) ⟹ f ∈ L F (h)"
by (subst cong_bigtheta[symmetric]) (simp add: bigtheta_sym)
lemma (in landau_symbol) bigtheta_trans2':
"g ∈ Θ[F](f) ⟹ g ∈ L F (h) ⟹ f ∈ L F (h)"
by (rule bigtheta_trans2, subst bigtheta_sym)
lemma bigo_bigomega_trans: "f ∈ O[F](g) ⟹ h ∈ Ω[F](g) ⟹ f ∈ O[F](h)"
and bigo_smallomega_trans: "f ∈ O[F](g) ⟹ h ∈ ω[F](g) ⟹ f ∈ o[F](h)"
and smallo_bigomega_trans: "f ∈ o[F](g) ⟹ h ∈ Ω[F](g) ⟹ f ∈ o[F](h)"
and smallo_smallomega_trans: "f ∈ o[F](g) ⟹ h ∈ ω[F](g) ⟹ f ∈ o[F](h)"
and bigomega_bigo_trans: "f ∈ Ω[F](g) ⟹ h ∈ O[F](g) ⟹ f ∈ Ω[F](h)"
and bigomega_smallo_trans: "f ∈ Ω[F](g) ⟹ h ∈ o[F](g) ⟹ f ∈ ω[F](h)"
and smallomega_bigo_trans: "f ∈ ω[F](g) ⟹ h ∈ O[F](g) ⟹ f ∈ ω[F](h)"
and smallomega_smallo_trans: "f ∈ ω[F](g) ⟹ h ∈ o[F](g) ⟹ f ∈ ω[F](h)"
by (unfold bigomega_iff_bigo smallomega_iff_smallo)
(erule (1) landau_o.big_trans landau_o.big_small_trans landau_o.small_big_trans
landau_o.big_trans landau_o.small_trans)+
lemmas landau_trans_lift [trans] =
landau_symbols[THEN landau_symbol.lift_trans]
landau_symbols[THEN landau_symbol.lift_trans']
landau_symbols[THEN landau_symbol.lift_trans_bigtheta]
landau_symbols[THEN landau_symbol.lift_trans_bigtheta']
lemmas landau_mult_1_trans [trans] =
landau_o.mult_1_trans landau_omega.mult_1_trans
lemmas landau_trans [trans] =
landau_symbols[THEN landau_symbol.bigtheta_trans1]
landau_symbols[THEN landau_symbol.bigtheta_trans2]
landau_symbols[THEN landau_symbol.bigtheta_trans1']
landau_symbols[THEN landau_symbol.bigtheta_trans2']
landau_symbols[THEN landau_symbol.ev_eq_trans1]
landau_symbols[THEN landau_symbol.ev_eq_trans2]
landau_o.big_trans landau_o.small_trans landau_o.small_big_trans landau_o.big_small_trans
landau_omega.big_trans landau_omega.small_trans
landau_omega.small_big_trans landau_omega.big_small_trans
bigo_bigomega_trans bigo_smallomega_trans smallo_bigomega_trans smallo_smallomega_trans
bigomega_bigo_trans bigomega_smallo_trans smallomega_bigo_trans smallomega_smallo_trans
lemma bigtheta_inverse [simp]:
shows "(λx. inverse (f x)) ∈ Θ[F](λx. inverse (g x)) ⟷ f ∈ Θ[F](g)"
proof -
have "(λx. inverse (f x)) ∈ O[F](λx. inverse (g x))"
if A: "f ∈ Θ[F](g)"
for f g :: "'a ⇒ 'b" and F
proof -
from A obtain c1 c2 :: real where *: "c1 > 0" "c2 > 0"
and **: "∀⇩F x in F. norm (f x) ≤ c1 * norm (g x)"
"∀⇩F x in F. c2 * norm (g x) ≤ norm (f x)"
unfolding bigtheta_def by (elim landau_o.bigE landau_omega.bigE IntE)
from ‹c2 > 0› have c2: "inverse c2 > 0" by simp
from ** have "eventually (λx. norm (inverse (f x)) ≤ inverse c2 * norm (inverse (g x))) F"
proof eventually_elim
fix x assume A: "norm (f x) ≤ c1 * norm (g x)" "c2 * norm (g x) ≤ norm (f x)"
from A * have "f x = 0 ⟷ g x = 0"
by (auto simp: field_simps mult_le_0_iff)
with A * show "norm (inverse (f x)) ≤ inverse c2 * norm (inverse (g x))"
by (force simp: field_simps norm_inverse norm_divide)
qed
with c2 show ?thesis by (rule landau_o.bigI)
qed
then show ?thesis
unfolding bigtheta_def
by (force simp: bigomega_iff_bigo bigtheta_sym)
qed
lemma bigtheta_divide:
assumes "f1 ∈ Θ(f2)" "g1 ∈ Θ(g2)"
shows "(λx. f1 x / g1 x) ∈ Θ(λx. f2 x / g2 x)"
by (subst (1 2) divide_inverse, intro landau_theta.mult) (simp_all add: bigtheta_inverse assms)
lemma eventually_nonzero_bigtheta:
assumes "f ∈ Θ[F](g)"
shows "eventually (λx. f x ≠ 0) F ⟷ eventually (λx. g x ≠ 0) F"
proof -
have "eventually (λx. g x ≠ 0) F"
if A: "f ∈ Θ[F](g)" and B: "eventually (λx. f x ≠ 0) F"
for f g :: "'a ⇒ 'b"
proof -
from A obtain c1 c2 where
"∀⇩F x in F. norm (f x) ≤ c1 * norm (g x)"
"∀⇩F x in F. c2 * norm (g x) ≤ norm (f x)"
unfolding bigtheta_def by (elim landau_o.bigE landau_omega.bigE IntE)
with B show ?thesis by eventually_elim auto
qed
with assms show ?thesis by (force simp: bigtheta_sym)
qed
subsection ‹Landau symbols and limits›
lemma bigoI_tendsto_norm:
fixes f g
assumes "((λx. norm (f x / g x)) ⤏ c) F"
assumes "eventually (λx. g x ≠ 0) F"
shows "f ∈ O[F](g)"
proof (rule bigoI)
from assms have "eventually (λx. dist (norm (f x / g x)) c < 1) F"
using tendstoD by force
thus "eventually (λx. (norm (f x)) ≤ (norm c + 1) * (norm (g x))) F"
unfolding dist_real_def using assms(2)
proof eventually_elim
case (elim x)
have "(norm (f x)) - norm c * (norm (g x)) ≤ norm ((norm (f x)) - c * (norm (g x)))"
unfolding norm_mult [symmetric] using norm_triangle_ineq2[of "norm (f x)" "c * norm (g x)"]
by (simp add: norm_mult abs_mult)
also from elim have "… = norm (norm (g x)) * norm (norm (f x / g x) - c)"
unfolding norm_mult [symmetric] by (simp add: algebra_simps norm_divide)
also from elim have "norm (norm (f x / g x) - c) ≤ 1" by simp
hence "norm (norm (g x)) * norm (norm (f x / g x) - c) ≤ norm (norm (g x)) * 1"
by (rule mult_left_mono) simp_all
finally show ?case by (simp add: algebra_simps)
qed
qed
lemma bigoI_tendsto:
assumes "((λx. f x / g x) ⤏ c) F"
assumes "eventually (λx. g x ≠ 0) F"
shows "f ∈ O[F](g)"
using assms by (rule bigoI_tendsto_norm[OF tendsto_norm])
lemma bigomegaI_tendsto_norm:
assumes c_not_0: "(c::real) ≠ 0"
assumes lim: "((λx. norm (f x / g x)) ⤏ c) F"
shows "f ∈ Ω[F](g)"
proof (cases "F = bot")
case False
show ?thesis
proof (rule landau_omega.bigI)
from lim have "c ≥ 0" by (rule tendsto_lowerbound) (insert False, simp_all)
with c_not_0 have "c > 0" by simp
with c_not_0 show "c/2 > 0" by simp
from lim have ev: "⋀ε. ε > 0 ⟹ eventually (λx. norm (norm (f x / g x) - c) < ε) F"
by (subst (asm) tendsto_iff) (simp add: dist_real_def)
from ev[OF ‹c/2 > 0›] show "eventually (λx. (norm (f x)) ≥ c/2 * (norm (g x))) F"
proof (eventually_elim)
fix x assume B: "norm (norm (f x / g x) - c) < c / 2"
from B have g: "g x ≠ 0" by auto
from B have "-c/2 < -norm (norm (f x / g x) - c)" by simp
also have "... ≤ norm (f x / g x) - c" by simp
finally show "(norm (f x)) ≥ c/2 * (norm (g x))" using g
by (simp add: field_simps norm_mult norm_divide)
qed
qed
qed simp
lemma bigomegaI_tendsto:
assumes c_not_0: "(c::real) ≠ 0"
assumes lim: "((λx. f x / g x) ⤏ c) F"
shows "f ∈ Ω[F](g)"
by (rule bigomegaI_tendsto_norm[OF _ tendsto_norm, of c]) (insert assms, simp_all)
lemma smallomegaI_filterlim_at_top_norm:
assumes lim: "filterlim (λx. norm (f x / g x)) at_top F"
shows "f ∈ ω[F](g)"
proof (rule landau_omega.smallI)
fix c :: real assume c_pos: "c > 0"
from lim have ev: "eventually (λx. norm (f x / g x) ≥ c) F"
by (subst (asm) filterlim_at_top) simp
thus "eventually (λx. (norm (f x)) ≥ c * (norm (g x))) F"
proof eventually_elim
fix x assume A: "norm (f x / g x) ≥ c"
from A c_pos have "g x ≠ 0" by auto
with A show "(norm (f x)) ≥ c * (norm (g x))" by (simp add: field_simps norm_divide)
qed
qed
lemma smallomegaI_filterlim_at_infinity:
assumes lim: "filterlim (λx. f x / g x) at_infinity F"
shows "f ∈ ω[F](g)"
proof (rule smallomegaI_filterlim_at_top_norm)
from lim show "filterlim (λx. norm (f x / g x)) at_top F"
by (rule filterlim_at_infinity_imp_norm_at_top)
qed
lemma smallomegaD_filterlim_at_top_norm:
assumes "f ∈ ω[F](g)"
assumes "eventually (λx. g x ≠ 0) F"
shows "LIM x F. norm (f x / g x) :> at_top"
proof (subst filterlim_at_top_gt, clarify)
fix c :: real assume c: "c > 0"
from landau_omega.smallD[OF assms(1) this] assms(2)
show "eventually (λx. norm (f x / g x) ≥ c) F"
by eventually_elim (simp add: field_simps norm_divide)
qed
lemma smallomegaD_filterlim_at_infinity:
assumes "f ∈ ω[F](g)"
assumes "eventually (λx. g x ≠ 0) F"
shows "LIM x F. f x / g x :> at_infinity"
using assms by (intro filterlim_norm_at_top_imp_at_infinity smallomegaD_filterlim_at_top_norm)
lemma smallomega_1_conv_filterlim: "f ∈ ω[F](λ_. 1) ⟷ filterlim f at_infinity F"
by (auto intro: smallomegaI_filterlim_at_infinity dest: smallomegaD_filterlim_at_infinity)
lemma smalloI_tendsto:
assumes lim: "((λx. f x / g x) ⤏ 0) F"
assumes "eventually (λx. g x ≠ 0) F"
shows "f ∈ o[F](g)"
proof (rule landau_o.smallI)
fix c :: real assume c_pos: "c > 0"
from c_pos and lim have ev: "eventually (λx. norm (f x / g x) < c) F"
by (subst (asm) tendsto_iff) (simp add: dist_real_def)
with assms(2) show "eventually (λx. (norm (f x)) ≤ c * (norm (g x))) F"
by eventually_elim (simp add: field_simps norm_divide)
qed
lemma smalloD_tendsto:
assumes "f ∈ o[F](g)"
shows "((λx. f x / g x) ⤏ 0) F"
unfolding tendsto_iff
proof clarify
fix e :: real assume e: "e > 0"
hence "e/2 > 0" by simp
from landau_o.smallD[OF assms this] show "eventually (λx. dist (f x / g x) 0 < e) F"
proof eventually_elim
fix x assume "(norm (f x)) ≤ e/2 * (norm (g x))"
with e have "dist (f x / g x) 0 ≤ e/2"
by (cases "g x = 0") (simp_all add: dist_real_def norm_divide field_simps)
also from e have "... < e" by simp
finally show "dist (f x / g x) 0 < e" by simp
qed
qed
lemma bigthetaI_tendsto_norm:
assumes c_not_0: "(c::real) ≠ 0"
assumes lim: "((λx. norm (f x / g x)) ⤏ c) F"
shows "f ∈ Θ[F](g)"
proof (rule bigthetaI)
from c_not_0 have "¦c¦ > 0" by simp
with lim have "eventually (λx. norm (norm (f x / g x) - c) < ¦c¦) F"
by (subst (asm) tendsto_iff) (simp add: dist_real_def)
hence g: "eventually (λx. g x ≠ 0) F" by eventually_elim (auto simp add: field_simps)
from lim g show "f ∈ O[F](g)" by (rule bigoI_tendsto_norm)
from c_not_0 and lim show "f ∈ Ω[F](g)" by (rule bigomegaI_tendsto_norm)
qed
lemma bigthetaI_tendsto:
assumes c_not_0: "(c::real) ≠ 0"
assumes lim: "((λx. f x / g x) ⤏ c) F"
shows "f ∈ Θ[F](g)"
using assms by (intro bigthetaI_tendsto_norm[OF _ tendsto_norm, of "c"]) simp_all
lemma tendsto_add_smallo:
assumes "(f1 ⤏ a) F"
assumes "f2 ∈ o[F](f1)"
shows "((λx. f1 x + f2 x) ⤏ a) F"
proof (subst filterlim_cong[OF refl refl])
from landau_o.smallD[OF assms(2) zero_less_one]
have "eventually (λx. norm (f2 x) ≤ norm (f1 x)) F" by simp
thus "eventually (λx. f1 x + f2 x = f1 x * (1 + f2 x / f1 x)) F"
by eventually_elim (auto simp: field_simps)
next
from assms(1) show "((λx. f1 x * (1 + f2 x / f1 x)) ⤏ a) F"
by (force intro: tendsto_eq_intros smalloD_tendsto[OF assms(2)])
qed
lemma tendsto_diff_smallo:
shows "(f1 ⤏ a) F ⟹ f2 ∈ o[F](f1) ⟹ ((λx. f1 x - f2 x) ⤏ a) F"
using tendsto_add_smallo[of f1 a F "λx. -f2 x"] by simp
lemma tendsto_add_smallo_iff:
assumes "f2 ∈ o[F](f1)"
shows "(f1 ⤏ a) F ⟷ ((λx. f1 x + f2 x) ⤏ a) F"
proof
assume "((λx. f1 x + f2 x) ⤏ a) F"
hence "((λx. f1 x + f2 x - f2 x) ⤏ a) F"
by (rule tendsto_diff_smallo) (simp add: landau_o.small.plus_absorb2 assms)
thus "(f1 ⤏ a) F" by simp
qed (rule tendsto_add_smallo[OF _ assms])
lemma tendsto_diff_smallo_iff:
shows "f2 ∈ o[F](f1) ⟹ (f1 ⤏ a) F ⟷ ((λx. f1 x - f2 x) ⤏ a) F"
using tendsto_add_smallo_iff[of "λx. -f2 x" F f1 a] by simp
lemma tendsto_divide_smallo:
assumes "((λx. f1 x / g1 x) ⤏ a) F"
assumes "f2 ∈ o[F](f1)" "g2 ∈ o[F](g1)"
assumes "eventually (λx. g1 x ≠ 0) F"
shows "((λx. (f1 x + f2 x) / (g1 x + g2 x)) ⤏ a) F" (is "(?f ⤏ _) _")
proof (subst tendsto_cong)
let ?f' = "λx. (f1 x / g1 x) * (1 + f2 x / f1 x) / (1 + g2 x / g1 x)"
have "(?f' ⤏ a * (1 + 0) / (1 + 0)) F"
by (rule tendsto_mult tendsto_divide tendsto_add assms tendsto_const
smalloD_tendsto[OF assms(2)] smalloD_tendsto[OF assms(3)])+ simp_all
thus "(?f' ⤏ a) F" by simp
have "(1/2::real) > 0" by simp
from landau_o.smallD[OF assms(2) this] landau_o.smallD[OF assms(3) this]
have "eventually (λx. norm (f2 x) ≤ norm (f1 x)/2) F"
"eventually (λx. norm (g2 x) ≤ norm (g1 x)/2) F" by simp_all
with assms(4) show "eventually (λx. ?f x = ?f' x) F"
proof eventually_elim
fix x assume A: "norm (f2 x) ≤ norm (f1 x)/2" and
B: "norm (g2 x) ≤ norm (g1 x)/2" and C: "g1 x ≠ 0"
show "?f x = ?f' x"
proof (cases "f1 x = 0")
assume D: "f1 x ≠ 0"
from D have "f1 x + f2 x = f1 x * (1 + f2 x/f1 x)" by (simp add: field_simps)
moreover from C have "g1 x + g2 x = g1 x * (1 + g2 x/g1 x)" by (simp add: field_simps)
ultimately have "?f x = (f1 x * (1 + f2 x/f1 x)) / (g1 x * (1 + g2 x/g1 x))" by (simp only:)
also have "... = ?f' x" by simp
finally show ?thesis .
qed (insert A, simp)
qed
qed
lemma bigo_powr:
fixes f :: "'a ⇒ real"
assumes "f ∈ O[F](g)" "p ≥ 0"
shows "(λx. ¦f x¦ powr p) ∈ O[F](λx. ¦g x¦ powr p)"
proof-
from assms(1) obtain c where c: "c > 0" and *: "∀⇩F x in F. norm (f x) ≤ c * norm (g x)"
by (elim landau_o.bigE landau_omega.bigE IntE)
from assms(2) * have "eventually (λx. (norm (f x)) powr p ≤ (c * norm (g x)) powr p) F"
by (auto elim!: eventually_mono intro!: powr_mono2)
with c show "(λx. ¦f x¦ powr p) ∈ O[F](λx. ¦g x¦ powr p)"
by (intro bigoI[of _ "c powr p"]) (simp_all add: powr_mult)
qed
lemma smallo_powr:
fixes f :: "'a ⇒ real"
assumes "f ∈ o[F](g)" "p > 0"
shows "(λx. ¦f x¦ powr p) ∈ o[F](λx. ¦g x¦ powr p)"
proof (rule landau_o.smallI)
fix c :: real assume c: "c > 0"
hence "c powr (1/p) > 0" by simp
from landau_o.smallD[OF assms(1) this]
show "eventually (λx. norm (¦f x¦ powr p) ≤ c * norm (¦g x¦ powr p)) F"
proof eventually_elim
fix x assume "(norm (f x)) ≤ c powr (1 / p) * (norm (g x))"
with assms(2) have "(norm (f x)) powr p ≤ (c powr (1 / p) * (norm (g x))) powr p"
by (intro powr_mono2) simp_all
also from assms(2) c have "... = c * (norm (g x)) powr p"
by (simp add: field_simps powr_mult powr_powr)
finally show "norm (¦f x¦ powr p) ≤ c * norm (¦g x¦ powr p)" by simp
qed
qed
lemma smallo_powr_nonneg:
fixes f :: "'a ⇒ real"
assumes "f ∈ o[F](g)" "p > 0" "eventually (λx. f x ≥ 0) F" "eventually (λx. g x ≥ 0) F"
shows "(λx. f x powr p) ∈ o[F](λx. g x powr p)"
proof -
from assms(3) have "(λx. f x powr p) ∈ Θ[F](λx. ¦f x¦ powr p)"
by (intro bigthetaI_cong) (auto elim!: eventually_mono)
also have "(λx. ¦f x¦ powr p) ∈ o[F](λx. ¦g x¦ powr p)" by (intro smallo_powr) fact+
also from assms(4) have "(λx. ¦g x¦ powr p) ∈ Θ[F](λx. g x powr p)"
by (intro bigthetaI_cong) (auto elim!: eventually_mono)
finally show ?thesis .
qed
lemma bigtheta_powr:
fixes f :: "'a ⇒ real"
shows "f ∈ Θ[F](g) ⟹ (λx. ¦f x¦ powr p) ∈ Θ[F](λx. ¦g x¦ powr p)"
apply (cases "p < 0")
apply (subst bigtheta_inverse[symmetric], subst (1 2) powr_minus[symmetric])
unfolding bigtheta_def apply (auto simp: bigomega_iff_bigo intro!: bigo_powr)
done
lemma bigo_powr_nonneg:
fixes f :: "'a ⇒ real"
assumes "f ∈ O[F](g)" "p ≥ 0" "eventually (λx. f x ≥ 0) F" "eventually (λx. g x ≥ 0) F"
shows "(λx. f x powr p) ∈ O[F](λx. g x powr p)"
proof -
from assms(3) have "(λx. f x powr p) ∈ Θ[F](λx. ¦f x¦ powr p)"
by (intro bigthetaI_cong) (auto elim!: eventually_mono)
also have "(λx. ¦f x¦ powr p) ∈ O[F](λx. ¦g x¦ powr p)" by (intro bigo_powr) fact+
also from assms(4) have "(λx. ¦g x¦ powr p) ∈ Θ[F](λx. g x powr p)"
by (intro bigthetaI_cong) (auto elim!: eventually_mono)
finally show ?thesis .
qed
lemma zero_in_smallo [simp]: "(λ_. 0) ∈ o[F](f)"
by (intro landau_o.smallI) simp_all
lemma zero_in_bigo [simp]: "(λ_. 0) ∈ O[F](f)"
by (intro landau_o.bigI[of 1]) simp_all
lemma in_bigomega_zero [simp]: "f ∈ Ω[F](λx. 0)"
by (rule landau_omega.bigI[of 1]) simp_all
lemma in_smallomega_zero [simp]: "f ∈ ω[F](λx. 0)"
by (simp add: smallomega_iff_smallo)
lemma in_smallo_zero_iff [simp]: "f ∈ o[F](λ_. 0) ⟷ eventually (λx. f x = 0) F"
proof
assume "f ∈ o[F](λ_. 0)"
from landau_o.smallD[OF this, of 1] show "eventually (λx. f x = 0) F" by simp
next
assume "eventually (λx. f x = 0) F"
hence "∀c>0. eventually (λx. (norm (f x)) ≤ c * ¦0¦) F" by simp
thus "f ∈ o[F](λ_. 0)" unfolding smallo_def by simp
qed
lemma in_bigo_zero_iff [simp]: "f ∈ O[F](λ_. 0) ⟷ eventually (λx. f x = 0) F"
proof
assume "f ∈ O[F](λ_. 0)"
thus "eventually (λx. f x = 0) F" by (elim landau_o.bigE) simp
next
assume "eventually (λx. f x = 0) F"
hence "eventually (λx. (norm (f x)) ≤ 1 * ¦0¦) F" by simp
thus "f ∈ O[F](λ_. 0)" by (intro landau_o.bigI[of 1]) simp_all
qed
lemma zero_in_smallomega_iff [simp]: "(λ_. 0) ∈ ω[F](f) ⟷ eventually (λx. f x = 0) F"
by (simp add: smallomega_iff_smallo)
lemma zero_in_bigomega_iff [simp]: "(λ_. 0) ∈ Ω[F](f) ⟷ eventually (λx. f x = 0) F"
by (simp add: bigomega_iff_bigo)
lemma zero_in_bigtheta_iff [simp]: "(λ_. 0) ∈ Θ[F](f) ⟷ eventually (λx. f x = 0) F"
unfolding bigtheta_def by simp
lemma in_bigtheta_zero_iff [simp]: "f ∈ Θ[F](λx. 0) ⟷ eventually (λx. f x = 0) F"
unfolding bigtheta_def by simp
lemma cmult_in_bigo_iff [simp]: "(λx. c * f x) ∈ O[F](g) ⟷ c = 0 ∨ f ∈ O[F](g)"
and cmult_in_bigo_iff' [simp]: "(λx. f x * c) ∈ O[F](g) ⟷ c = 0 ∨ f ∈ O[F](g)"
and cmult_in_smallo_iff [simp]: "(λx. c * f x) ∈ o[F](g) ⟷ c = 0 ∨ f ∈ o[F](g)"
and cmult_in_smallo_iff' [simp]: "(λx. f x * c) ∈ o[F](g) ⟷ c = 0 ∨ f ∈ o[F](g)"
by (cases "c = 0", simp, simp)+
lemma bigo_const [simp]: "(λ_. c) ∈ O[F](λ_. 1)" by (rule bigoI[of _ "norm c"]) simp
lemma bigo_const_iff [simp]: "(λ_. c1) ∈ O[F](λ_. c2) ⟷ F = bot ∨ c1 = 0 ∨ c2 ≠ 0"
by (cases "c1 = 0"; cases "c2 = 0")
(auto simp: bigo_def eventually_False intro: exI[of _ 1] exI[of _ "norm c1 / norm c2"])
lemma bigomega_const_iff [simp]: "(λ_. c1) ∈ Ω[F](λ_. c2) ⟷ F = bot ∨ c1 ≠ 0 ∨ c2 = 0"
by (cases "c1 = 0"; cases "c2 = 0")
(auto simp: bigomega_def eventually_False mult_le_0_iff
intro: exI[of _ 1] exI[of _ "norm c1 / norm c2"])
lemma smallo_real_nat_transfer:
"(f :: real ⇒ real) ∈ o(g) ⟹ (λx::nat. f (real x)) ∈ o(λx. g (real x))"
by (rule landau_o.small.compose[OF _ filterlim_real_sequentially])
lemma bigo_real_nat_transfer:
"(f :: real ⇒ real) ∈ O(g) ⟹ (λx::nat. f (real x)) ∈ O(λx. g (real x))"
by (rule landau_o.big.compose[OF _ filterlim_real_sequentially])
lemma smallomega_real_nat_transfer:
"(f :: real ⇒ real) ∈ ω(g) ⟹ (λx::nat. f (real x)) ∈ ω(λx. g (real x))"
by (rule landau_omega.small.compose[OF _ filterlim_real_sequentially])
lemma bigomega_real_nat_transfer:
"(f :: real ⇒ real) ∈ Ω(g) ⟹ (λx::nat. f (real x)) ∈ Ω(λx. g (real x))"
by (rule landau_omega.big.compose[OF _ filterlim_real_sequentially])
lemma bigtheta_real_nat_transfer:
"(f :: real ⇒ real) ∈ Θ(g) ⟹ (λx::nat. f (real x)) ∈ Θ(λx. g (real x))"
unfolding bigtheta_def using bigo_real_nat_transfer bigomega_real_nat_transfer by blast
lemmas landau_real_nat_transfer [intro] =
bigo_real_nat_transfer smallo_real_nat_transfer bigomega_real_nat_transfer
smallomega_real_nat_transfer bigtheta_real_nat_transfer
lemma landau_symbol_if_at_top_eq [simp]:
assumes "landau_symbol L L' Lr"
shows "L at_top (λx::'a::linordered_semidom. if x = a then f x else g x) = L at_top (g)"
apply (rule landau_symbol.cong[OF assms])
using less_add_one[of a] apply (auto intro: eventually_mono eventually_ge_at_top[of "a + 1"])
done
lemmas landau_symbols_if_at_top_eq [simp] = landau_symbols[THEN landau_symbol_if_at_top_eq]
lemma sum_in_smallo:
assumes "f ∈ o[F](h)" "g ∈ o[F](h)"
shows "(λx. f x + g x) ∈ o[F](h)" "(λx. f x - g x) ∈ o[F](h)"
proof -
have "(λx. f x + g x) ∈ o[F](h)" if "f ∈ o[F](h)" "g ∈ o[F](h)" for f g
proof (rule landau_o.smallI)
fix c :: real assume "c > 0"
hence "c/2 > 0" by simp
from that[THEN landau_o.smallD[OF _ this]]
show "eventually (λx. norm (f x + g x) ≤ c * (norm (h x))) F"
by eventually_elim (auto intro: order.trans[OF norm_triangle_ineq])
qed
from this[of f g] this[of f "λx. -g x"] assms
show "(λx. f x + g x) ∈ o[F](h)" "(λx. f x - g x) ∈ o[F](h)" by simp_all
qed
lemma big_sum_in_smallo:
assumes "⋀x. x ∈ A ⟹ f x ∈ o[F](g)"
shows "(λx. sum (λy. f y x) A) ∈ o[F](g)"
using assms by (induction A rule: infinite_finite_induct) (auto intro: sum_in_smallo)
lemma sum_in_bigo:
assumes "f ∈ O[F](h)" "g ∈ O[F](h)"
shows "(λx. f x + g x) ∈ O[F](h)" "(λx. f x - g x) ∈ O[F](h)"
proof -
have "(λx. f x + g x) ∈ O[F](h)" if "f ∈ O[F](h)" "g ∈ O[F](h)" for f g
proof -
from that obtain c1 c2 where *: "c1 > 0" "c2 > 0"
and **: "∀⇩F x in F. norm (f x) ≤ c1 * norm (h x)"
"∀⇩F x in F. norm (g x) ≤ c2 * norm (h x)"
by (elim landau_o.bigE)
from ** have "eventually (λx. norm (f x + g x) ≤ (c1 + c2) * (norm (h x))) F"
by eventually_elim (auto simp: algebra_simps intro: order.trans[OF norm_triangle_ineq])
then show ?thesis by (rule bigoI)
qed
from assms this[of f g] this[of f "λx. - g x"]
show "(λx. f x + g x) ∈ O[F](h)" "(λx. f x - g x) ∈ O[F](h)" by simp_all
qed
lemma big_sum_in_bigo:
assumes "⋀x. x ∈ A ⟹ f x ∈ O[F](g)"
shows "(λx. sum (λy. f y x) A) ∈ O[F](g)"
using assms by (induction A rule: infinite_finite_induct) (auto intro: sum_in_bigo)
lemma le_imp_bigo_real:
assumes "c ≥ 0" "eventually (λx. f x ≤ c * (g x :: real)) F" "eventually (λx. 0 ≤ f x) F"
shows "f ∈ O[F](g)"
proof -
have "eventually (λx. norm (f x) ≤ c * norm (g x)) F"
using assms(2,3)
proof eventually_elim
case (elim x)
have "norm (f x) ≤ c * g x" using elim by simp
also have "… ≤ c * norm (g x)" by (intro mult_left_mono assms) auto
finally show ?case .
qed
thus ?thesis by (intro bigoI[of _ c]) auto
qed
context landau_symbol
begin
lemma mult_cancel_left:
assumes "f1 ∈ Θ[F](g1)" and "eventually (λx. g1 x ≠ 0) F"
notes [trans] = bigtheta_trans1 bigtheta_trans2
shows "(λx. f1 x * f2 x) ∈ L F (λx. g1 x * g2 x) ⟷ f2 ∈ L F (g2)"
proof
assume A: "(λx. f1 x * f2 x) ∈ L F (λx. g1 x * g2 x)"
from assms have nz: "eventually (λx. f1 x ≠ 0) F" by (simp add: eventually_nonzero_bigtheta)
hence "f2 ∈ Θ[F](λx. f1 x * f2 x / f1 x)"
by (intro bigthetaI_cong) (auto elim!: eventually_mono)
also from A assms nz have "(λx. f1 x * f2 x / f1 x) ∈ L F (λx. g1 x * g2 x / f1 x)"
by (intro divide_right) simp_all
also from assms nz have "(λx. g1 x * g2 x / f1 x) ∈ Θ[F](λx. g1 x * g2 x / g1 x)"
by (intro landau_theta.mult landau_theta.divide) (simp_all add: bigtheta_sym)
also from assms have "(λx. g1 x * g2 x / g1 x) ∈ Θ[F](g2)"
by (intro bigthetaI_cong) (auto elim!: eventually_mono)
finally show "f2 ∈ L F (g2)" .
next
assume "f2 ∈ L F (g2)"
hence "(λx. f1 x * f2 x) ∈ L F (λx. f1 x * g2 x)" by (rule mult_left)
also have "(λx. f1 x * g2 x) ∈ Θ[F](λx. g1 x * g2 x)"
by (intro landau_theta.mult_right assms)
finally show "(λx. f1 x * f2 x) ∈ L F (λx. g1 x * g2 x)" .
qed
lemma mult_cancel_right:
assumes "f2 ∈ Θ[F](g2)" and "eventually (λx. g2 x ≠ 0) F"
shows "(λx. f1 x * f2 x) ∈ L F (λx. g1 x * g2 x) ⟷ f1 ∈ L F (g1)"
by (subst (1 2) mult.commute) (rule mult_cancel_left[OF assms])
lemma divide_cancel_right:
assumes "f2 ∈ Θ[F](g2)" and "eventually (λx. g2 x ≠ 0) F"
shows "(λx. f1 x / f2 x) ∈ L F (λx. g1 x / g2 x) ⟷ f1 ∈ L F (g1)"
by (subst (1 2) divide_inverse, intro mult_cancel_right bigtheta_inverse) (simp_all add: assms)
lemma divide_cancel_left:
assumes "f1 ∈ Θ[F](g1)" and "eventually (λx. g1 x ≠ 0) F"
shows "(λx. f1 x / f2 x) ∈ L F (λx. g1 x / g2 x) ⟷
(λx. inverse (f2 x)) ∈ L F (λx. inverse (g2 x))"
by (simp only: divide_inverse mult_cancel_left[OF assms])
end
lemma powr_smallo_iff:
assumes "filterlim g at_top F" "F ≠ bot"
shows "(λx. g x powr p :: real) ∈ o[F](λx. g x powr q) ⟷ p < q"
proof-
from assms have "eventually (λx. g x ≥ 1) F" by (force simp: filterlim_at_top)
hence A: "eventually (λx. g x ≠ 0) F" by eventually_elim simp
have B: "(λx. g x powr q) ∈ O[F](λx. g x powr p) ⟹ (λx. g x powr p) ∉ o[F](λx. g x powr q)"
proof
assume "(λx. g x powr q) ∈ O[F](λx. g x powr p)" "(λx. g x powr p) ∈ o[F](λx. g x powr q)"
from landau_o.big_small_asymmetric[OF this] have "eventually (λx. g x = 0) F" by simp
with A have "eventually (λ_::'a. False) F" by eventually_elim simp
thus False by (simp add: eventually_False assms)
qed
show ?thesis
proof (cases p q rule: linorder_cases)
assume "p < q"
hence "(λx. g x powr p) ∈ o[F](λx. g x powr q)" using assms A
by (auto intro!: smalloI_tendsto tendsto_neg_powr simp flip: powr_diff)
with ‹p < q› show ?thesis by auto
next
assume "p = q"
hence "(λx. g x powr q) ∈ O[F](λx. g x powr p)" by (auto intro!: bigthetaD1)
with B ‹p = q› show ?thesis by auto
next
assume "p > q"
hence "(λx. g x powr q) ∈ O[F](λx. g x powr p)" using assms A
by (auto intro!: smalloI_tendsto tendsto_neg_powr landau_o.small_imp_big simp flip: powr_diff)
with B ‹p > q› show ?thesis by auto
qed
qed
lemma powr_bigo_iff:
assumes "filterlim g at_top F" "F ≠ bot"
shows "(λx. g x powr p :: real) ∈ O[F](λx. g x powr q) ⟷ p ≤ q"
proof-
from assms have "eventually (λx. g x ≥ 1) F" by (force simp: filterlim_at_top)
hence A: "eventually (λx. g x ≠ 0) F" by eventually_elim simp
have B: "(λx. g x powr q) ∈ o[F](λx. g x powr p) ⟹ (λx. g x powr p) ∉ O[F](λx. g x powr q)"
proof
assume "(λx. g x powr q) ∈ o[F](λx. g x powr p)" "(λx. g x powr p) ∈ O[F](λx. g x powr q)"
from landau_o.small_big_asymmetric[OF this] have "eventually (λx. g x = 0) F" by simp
with A have "eventually (λ_::'a. False) F" by eventually_elim simp
thus False by (simp add: eventually_False assms)
qed
show ?thesis
proof (cases p q rule: linorder_cases)
assume "p < q"
hence "(λx. g x powr p) ∈ o[F](λx. g x powr q)" using assms A
by (auto intro!: smalloI_tendsto tendsto_neg_powr simp flip: powr_diff)
with ‹p < q› show ?thesis by (auto intro: landau_o.small_imp_big)
next
assume "p = q"
hence "(λx. g x powr q) ∈ O[F](λx. g x powr p)" by (auto intro!: bigthetaD1)
with B ‹p = q› show ?thesis by auto
next
assume "p > q"
hence "(λx. g x powr q) ∈ o[F](λx. g x powr p)" using assms A
by (auto intro!: smalloI_tendsto tendsto_neg_powr simp flip: powr_diff)
with B ‹p > q› show ?thesis by (auto intro: landau_o.small_imp_big)
qed
qed
lemma powr_bigtheta_iff:
assumes "filterlim g at_top F" "F ≠ bot"
shows "(λx. g x powr p :: real) ∈ Θ[F](λx. g x powr q) ⟷ p = q"
using assms unfolding bigtheta_def by (auto simp: bigomega_iff_bigo powr_bigo_iff)
subsection ‹Flatness of real functions›
text ‹
Given two real-valued functions $f$ and $g$, we say that $f$ is flatter than $g$ if
any power of $f(x)$ is asymptotically dominated by any positive power of $g(x)$. This is
a useful notion since, given two products of powers of functions sorted by flatness, we can
compare them asymptotically by simply comparing the exponent lists lexicographically.
A simple sufficient criterion for flatness it that $\ln f(x) \in o(\ln g(x))$, which we show
now.
›
lemma ln_smallo_imp_flat:
fixes f g :: "real ⇒ real"
assumes lim_f: "filterlim f at_top at_top"
assumes lim_g: "filterlim g at_top at_top"
assumes ln_o_ln: "(λx. ln (f x)) ∈ o(λx. ln (g x))"
assumes q: "q > 0"
shows "(λx. f x powr p) ∈ o(λx. g x powr q)"
proof (rule smalloI_tendsto)
from lim_f have "eventually (λx. f x > 0) at_top"
by (simp add: filterlim_at_top_dense)
hence f_nz: "eventually (λx. f x ≠ 0) at_top" by eventually_elim simp
from lim_g have g_gt_1: "eventually (λx. g x > 1) at_top"
by (simp add: filterlim_at_top_dense)
hence g_nz: "eventually (λx. g x ≠ 0) at_top" by eventually_elim simp
thus "eventually (λx. g x powr q ≠ 0) at_top"
by eventually_elim simp
have eq: "eventually (λx. q * (p/q * (ln (f x) / ln (g x)) - 1) * ln (g x) =
p * ln (f x) - q * ln (g x)) at_top"
using g_gt_1 by eventually_elim (insert q, simp_all add: field_simps)
have "filterlim (λx. q * (p/q * (ln (f x) / ln (g x)) - 1) * ln (g x)) at_bot at_top"
by (insert q)
(rule filterlim_tendsto_neg_mult_at_bot tendsto_mult
tendsto_const tendsto_diff smalloD_tendsto[OF ln_o_ln] lim_g
filterlim_compose[OF ln_at_top] | simp)+
hence "filterlim (λx. p * ln (f x) - q * ln (g x)) at_bot at_top"
by (subst (asm) filterlim_cong[OF refl refl eq])
hence *: "((λx. exp (p * ln (f x) - q * ln (g x))) ⤏ 0) at_top"
by (rule filterlim_compose[OF exp_at_bot])
have eq: "eventually (λx. exp (p * ln (f x) - q * ln (g x)) = f x powr p / g x powr q) at_top"
using f_nz g_nz by eventually_elim (simp add: powr_def exp_diff)
show "((λx. f x powr p / g x powr q) ⤏ 0) at_top"
using * by (subst (asm) filterlim_cong[OF refl refl eq])
qed
lemma ln_smallo_imp_flat':
fixes f g :: "real ⇒ real"
assumes lim_f: "filterlim f at_top at_top"
assumes lim_g: "filterlim g at_top at_top"
assumes ln_o_ln: "(λx. ln (f x)) ∈ o(λx. ln (g x))"
assumes q: "q < 0"
shows "(λx. g x powr q) ∈ o(λx. f x powr p)"
proof -
from lim_f lim_g have "eventually (λx. f x > 0) at_top" "eventually (λx. g x > 0) at_top"
by (simp_all add: filterlim_at_top_dense)
hence "eventually (λx. f x ≠ 0) at_top" "eventually (λx. g x ≠ 0) at_top"
by (auto elim: eventually_mono)
moreover from assms have "(λx. f x powr -p) ∈ o(λx. g x powr -q)"
by (intro ln_smallo_imp_flat assms) simp_all
ultimately show ?thesis unfolding powr_minus
by (simp add: landau_o.small.inverse_cancel)
qed
subsection ‹Asymptotic Equivalence›
named_theorems asymp_equiv_intros
named_theorems asymp_equiv_simps
definition asymp_equiv :: "('a ⇒ ('b :: real_normed_field)) ⇒ 'a filter ⇒ ('a ⇒ 'b) ⇒ bool"
(‹_ ∼[_] _› [51, 10, 51] 50)
where "f ∼[F] g ⟷ ((λx. if f x = 0 ∧ g x = 0 then 1 else f x / g x) ⤏ 1) F"
abbreviation (input) asymp_equiv_at_top where
"asymp_equiv_at_top f g ≡ f ∼[at_top] g"
bundle asymp_equiv_notation
begin
notation asymp_equiv_at_top (infix ‹∼› 50)
end
lemma asymp_equivI: "((λx. if f x = 0 ∧ g x = 0 then 1 else f x / g x) ⤏ 1) F ⟹ f ∼[F] g"
by (simp add: asymp_equiv_def)
lemma asymp_equivD: "f ∼[F] g ⟹ ((λx. if f x = 0 ∧ g x = 0 then 1 else f x / g x) ⤏ 1) F"
by (simp add: asymp_equiv_def)
lemma asymp_equiv_filtermap_iff:
"f ∼[filtermap h F] g ⟷ (λx. f (h x)) ∼[F] (λx. g (h x))"
by (simp add: asymp_equiv_def filterlim_filtermap)
lemma asymp_equiv_refl [simp, asymp_equiv_intros]: "f ∼[F] f"
proof (intro asymp_equivI)
have "eventually (λx. 1 = (if f x = 0 ∧ f x = 0 then 1 else f x / f x)) F"
by (intro always_eventually) simp
moreover have "((λ_. 1) ⤏ 1) F" by simp
ultimately show "((λx. if f x = 0 ∧ f x = 0 then 1 else f x / f x) ⤏ 1) F"
by (simp add: tendsto_eventually)
qed
lemma asymp_equiv_symI:
assumes "f ∼[F] g"
shows "g ∼[F] f"
using tendsto_inverse[OF asymp_equivD[OF assms]]
by (auto intro!: asymp_equivI simp: if_distrib conj_commute cong: if_cong)
lemma asymp_equiv_sym: "f ∼[F] g ⟷ g ∼[F] f"
by (blast intro: asymp_equiv_symI)
lemma asymp_equivI':
assumes "((λx. f x / g x) ⤏ 1) F"
shows "f ∼[F] g"
proof (cases "F = bot")
case False
have "eventually (λx. f x ≠ 0) F"
proof (rule ccontr)
assume "¬eventually (λx. f x ≠ 0) F"
hence "frequently (λx. f x = 0) F" by (simp add: frequently_def)
hence "frequently (λx. f x / g x = 0) F" by (auto elim!: frequently_elim1)
from limit_frequently_eq[OF False this assms] show False by simp_all
qed
hence "eventually (λx. f x / g x = (if f x = 0 ∧ g x = 0 then 1 else f x / g x)) F"
by eventually_elim simp
with assms show "f ∼[F] g" unfolding asymp_equiv_def
by (rule Lim_transform_eventually)
qed (simp_all add: asymp_equiv_def)
lemma tendsto_imp_asymp_equiv_const:
assumes "(f ⤏ c) F" "c ≠ 0"
shows "f ∼[F] (λ_. c)"
by (rule asymp_equivI' tendsto_eq_intros assms refl)+ (use assms in auto)
lemma asymp_equiv_cong:
assumes "eventually (λx. f1 x = f2 x) F" "eventually (λx. g1 x = g2 x) F"
shows "f1 ∼[F] g1 ⟷ f2 ∼[F] g2"
unfolding asymp_equiv_def
proof (rule tendsto_cong, goal_cases)
case 1
from assms show ?case by eventually_elim simp
qed
lemma asymp_equiv_eventually_zeros:
fixes f g :: "'a ⇒ 'b :: real_normed_field"
assumes "f ∼[F] g"
shows "eventually (λx. f x = 0 ⟷ g x = 0) F"
proof -
let ?h = "λx. if f x = 0 ∧ g x = 0 then 1 else f x / g x"
have "eventually (λx. x ≠ 0) (nhds (1::'b))"
by (rule t1_space_nhds) auto
hence "eventually (λx. x ≠ 0) (filtermap ?h F)"
using assms unfolding asymp_equiv_def filterlim_def
by (rule filter_leD [rotated])
hence "eventually (λx. ?h x ≠ 0) F" by (simp add: eventually_filtermap)
thus ?thesis by eventually_elim (auto split: if_splits)
qed
lemma asymp_equiv_transfer:
assumes "f1 ∼[F] g1" "eventually (λx. f1 x = f2 x) F" "eventually (λx. g1 x = g2 x) F"
shows "f2 ∼[F] g2"
using assms(1) asymp_equiv_cong[OF assms(2,3)] by simp
lemma asymp_equiv_transfer_trans [trans]:
assumes "(λx. f x (h1 x)) ∼[F] (λx. g x (h1 x))"
assumes "eventually (λx. h1 x = h2 x) F"
shows "(λx. f x (h2 x)) ∼[F] (λx. g x (h2 x))"
by (rule asymp_equiv_transfer[OF assms(1)]) (insert assms(2), auto elim!: eventually_mono)
lemma asymp_equiv_trans [trans]:
fixes f g h
assumes "f ∼[F] g" "g ∼[F] h"
shows "f ∼[F] h"
proof -
let ?T = "λf g x. if f x = 0 ∧ g x = 0 then 1 else f x / g x"
from tendsto_mult[OF assms[THEN asymp_equivD]]
have "((λx. ?T f g x * ?T g h x) ⤏ 1) F" by simp
moreover from assms[THEN asymp_equiv_eventually_zeros]
have "eventually (λx. ?T f g x * ?T g h x = ?T f h x) F" by eventually_elim simp
ultimately show ?thesis unfolding asymp_equiv_def by (rule Lim_transform_eventually)
qed
lemma asymp_equiv_trans_lift1 [trans]:
assumes "a ∼[F] f b" "b ∼[F] c" "⋀c d. c ∼[F] d ⟹ f c ∼[F] f d"
shows "a ∼[F] f c"
using assms by (blast intro: asymp_equiv_trans)
lemma asymp_equiv_trans_lift2 [trans]:
assumes "f a ∼[F] b" "a ∼[F] c" "⋀c d. c ∼[F] d ⟹ f c ∼[F] f d"
shows "f c ∼[F] b"
using asymp_equiv_symI[OF assms(3)[OF assms(2)]] assms(1)
by (blast intro: asymp_equiv_trans)
lemma asymp_equivD_const:
assumes "f ∼[F] (λ_. c)"
shows "(f ⤏ c) F"
proof (cases "c = 0")
case False
with tendsto_mult_right[OF asymp_equivD[OF assms], of c] show ?thesis by simp
next
case True
with asymp_equiv_eventually_zeros[OF assms] show ?thesis
by (simp add: tendsto_eventually)
qed
lemma asymp_equiv_refl_ev:
assumes "eventually (λx. f x = g x) F"
shows "f ∼[F] g"
by (intro asymp_equivI tendsto_eventually)
(insert assms, auto elim!: eventually_mono)
lemma asymp_equiv_nhds_iff: "f ∼[nhds (z :: 'a :: t1_space)] g ⟷ f ∼[at z] g ∧ f z = g z"
by (auto simp: asymp_equiv_def tendsto_nhds_iff)
lemma asymp_equiv_sandwich:
fixes f g h :: "'a ⇒ 'b :: {real_normed_field, order_topology, linordered_field}"
assumes "eventually (λx. f x ≥ 0) F"
assumes "eventually (λx. f x ≤ g x) F"
assumes "eventually (λx. g x ≤ h x) F"
assumes "f ∼[F] h"
shows "g ∼[F] f" "g ∼[F] h"
proof -
show "g ∼[F] f"
proof (rule asymp_equivI, rule tendsto_sandwich)
from assms(1-3) asymp_equiv_eventually_zeros[OF assms(4)]
show "eventually (λn. (if h n = 0 ∧ f n = 0 then 1 else h n / f n) ≥
(if g n = 0 ∧ f n = 0 then 1 else g n / f n)) F"
by eventually_elim (auto intro!: divide_right_mono)
from assms(1-3) asymp_equiv_eventually_zeros[OF assms(4)]
show "eventually (λn. 1 ≤
(if g n = 0 ∧ f n = 0 then 1 else g n / f n)) F"
by eventually_elim (auto intro!: divide_right_mono)
qed (insert asymp_equiv_symI[OF assms(4)], simp_all add: asymp_equiv_def)
also note ‹f ∼[F] h›
finally show "g ∼[F] h" .
qed
lemma asymp_equiv_imp_eventually_same_sign:
fixes f g :: "real ⇒ real"
assumes "f ∼[F] g"
shows "eventually (λx. sgn (f x) = sgn (g x)) F"
proof -
from assms have "((λx. sgn (if f x = 0 ∧ g x = 0 then 1 else f x / g x)) ⤏ sgn 1) F"
unfolding asymp_equiv_def by (rule tendsto_sgn) simp_all
from order_tendstoD(1)[OF this, of "1/2"]
have "eventually (λx. sgn (if f x = 0 ∧ g x = 0 then 1 else f x / g x) > 1/2) F"
by simp
thus "eventually (λx. sgn (f x) = sgn (g x)) F"
proof eventually_elim
case (elim x)
thus ?case
by (cases "f x" "0 :: real" rule: linorder_cases;
cases "g x" "0 :: real" rule: linorder_cases) simp_all
qed
qed
lemma
fixes f g :: "_ ⇒ real"
assumes "f ∼[F] g"
shows asymp_equiv_eventually_same_sign: "eventually (λx. sgn (f x) = sgn (g x)) F" (is ?th1)
and asymp_equiv_eventually_neg_iff: "eventually (λx. f x < 0 ⟷ g x < 0) F" (is ?th2)
and asymp_equiv_eventually_pos_iff: "eventually (λx. f x > 0 ⟷ g x > 0) F" (is ?th3)
proof -
from assms have "filterlim (λx. if f x = 0 ∧ g x = 0 then 1 else f x / g x) (nhds 1) F"
by (rule asymp_equivD)
from order_tendstoD(1)[OF this zero_less_one]
show ?th1 ?th2 ?th3
by (eventually_elim; force simp: sgn_if field_split_simps split: if_splits)+
qed
lemma asymp_equiv_tendsto_transfer:
assumes "f ∼[F] g" and "(f ⤏ c) F"
shows "(g ⤏ c) F"
proof -
let ?h = "λx. (if g x = 0 ∧ f x = 0 then 1 else g x / f x) * f x"
from assms(1) have "g ∼[F] f" by (rule asymp_equiv_symI)
hence "filterlim (λx. if g x = 0 ∧ f x = 0 then 1 else g x / f x) (nhds 1) F"
by (rule asymp_equivD)
from tendsto_mult[OF this assms(2)] have "(?h ⤏ c) F" by simp
moreover
have "eventually (λx. ?h x = g x) F"
using asymp_equiv_eventually_zeros[OF assms(1)] by eventually_elim simp
ultimately show ?thesis
by (rule Lim_transform_eventually)
qed
lemma tendsto_asymp_equiv_cong:
assumes "f ∼[F] g"
shows "(f ⤏ c) F ⟷ (g ⤏ c) F"
proof -
have "(f ⤏ c * 1) F" if fg: "f ∼[F] g" and "(g ⤏ c) F" for f g :: "'a ⇒ 'b"
proof -
from that have *: "((λx. g x * (if f x = 0 ∧ g x = 0 then 1 else f x / g x)) ⤏ c * 1) F"
by (intro tendsto_intros asymp_equivD)
have "eventually (λx. g x * (if f x = 0 ∧ g x = 0 then 1 else f x / g x) = f x) F"
using asymp_equiv_eventually_zeros[OF fg] by eventually_elim simp
with * show ?thesis by (rule Lim_transform_eventually)
qed
from this[of f g] this[of g f] assms show ?thesis by (auto simp: asymp_equiv_sym)
qed
lemma smallo_imp_eventually_sgn:
fixes f g :: "real ⇒ real"
assumes "g ∈ o(f)"
shows "eventually (λx. sgn (f x + g x) = sgn (f x)) at_top"
proof -
have "0 < (1/2 :: real)" by simp
from landau_o.smallD[OF assms, OF this]
have "eventually (λx. ¦g x¦ ≤ 1/2 * ¦f x¦) at_top" by simp
thus ?thesis
proof eventually_elim
case (elim x)
thus ?case
by (cases "f x" "0::real" rule: linorder_cases;
cases "f x + g x" "0::real" rule: linorder_cases) simp_all
qed
qed
context
begin
private lemma asymp_equiv_add_rightI:
assumes "f ∼[F] g" "h ∈ o[F](g)"
shows "(λx. f x + h x) ∼[F] g"
proof -
let ?T = "λf g x. if f x = 0 ∧ g x = 0 then 1 else f x / g x"
from landau_o.smallD[OF assms(2) zero_less_one]
have ev: "eventually (λx. g x = 0 ⟶ h x = 0) F" by eventually_elim auto
have "(λx. f x + h x) ∼[F] g ⟷ ((λx. ?T f g x + h x / g x) ⤏ 1) F"
unfolding asymp_equiv_def using ev
by (intro tendsto_cong) (auto elim!: eventually_mono simp: field_split_simps)
also have "… ⟷ ((λx. ?T f g x + h x / g x) ⤏ 1 + 0) F" by simp
also have … by (intro tendsto_intros asymp_equivD assms smalloD_tendsto)
finally show "(λx. f x + h x) ∼[F] g" .
qed
lemma asymp_equiv_add_right [asymp_equiv_simps]:
assumes "h ∈ o[F](g)"
shows "(λx. f x + h x) ∼[F] g ⟷ f ∼[F] g"
proof
assume "(λx. f x + h x) ∼[F] g"
from asymp_equiv_add_rightI[OF this, of "λx. -h x"] assms show "f ∼[F] g"
by simp
qed (simp_all add: asymp_equiv_add_rightI assms)
end
lemma asymp_equiv_add_left [asymp_equiv_simps]:
assumes "h ∈ o[F](g)"
shows "(λx. h x + f x) ∼[F] g ⟷ f ∼[F] g"
using asymp_equiv_add_right[OF assms] by (simp add: add.commute)
lemma asymp_equiv_add_right' [asymp_equiv_simps]:
assumes "h ∈ o[F](g)"
shows "g ∼[F] (λx. f x + h x) ⟷ g ∼[F] f"
using asymp_equiv_add_right[OF assms] by (simp add: asymp_equiv_sym)
lemma asymp_equiv_add_left' [asymp_equiv_simps]:
assumes "h ∈ o[F](g)"
shows "g ∼[F] (λx. h x + f x) ⟷ g ∼[F] f"
using asymp_equiv_add_left[OF assms] by (simp add: asymp_equiv_sym)
lemma smallo_imp_asymp_equiv:
assumes "(λx. f x - g x) ∈ o[F](g)"
shows "f ∼[F] g"
proof -
from assms have "(λx. f x - g x + g x) ∼[F] g"
by (subst asymp_equiv_add_left) simp_all
thus ?thesis by simp
qed
lemma asymp_equiv_uminus [asymp_equiv_intros]:
"f ∼[F] g ⟹ (λx. -f x) ∼[F] (λx. -g x)"
by (simp add: asymp_equiv_def cong: if_cong)
lemma asymp_equiv_uminus_iff [asymp_equiv_simps]:
"(λx. -f x) ∼[F] g ⟷ f ∼[F] (λx. -g x)"
by (simp add: asymp_equiv_def cong: if_cong)
lemma asymp_equiv_mult [asymp_equiv_intros]:
fixes f1 f2 g1 g2 :: "'a ⇒ 'b :: real_normed_field"
assumes "f1 ∼[F] g1" "f2 ∼[F] g2"
shows "(λx. f1 x * f2 x) ∼[F] (λx. g1 x * g2 x)"
proof -
let ?T = "λf g x. if f x = 0 ∧ g x = 0 then 1 else f x / g x"
let ?S = "λx. (if f1 x = 0 ∧ g1 x = 0 then 1 - ?T f2 g2 x
else if f2 x = 0 ∧ g2 x = 0 then 1 - ?T f1 g1 x else 0)"
let ?S' = "λx. ?T (λx. f1 x * f2 x) (λx. g1 x * g2 x) x - ?T f1 g1 x * ?T f2 g2 x"
have A: "((λx. 1 - ?T f g x) ⤏ 0) F" if "f ∼[F] g" for f g :: "'a ⇒ 'b"
by (rule tendsto_eq_intros refl asymp_equivD[OF that])+ simp_all
from assms have *: "((λx. ?T f1 g1 x * ?T f2 g2 x) ⤏ 1 * 1) F"
by (intro tendsto_mult asymp_equivD)
{
have "(?S ⤏ 0) F"
by (intro filterlim_If assms[THEN A, THEN tendsto_mono[rotated]])
(auto intro: le_infI1 le_infI2)
moreover have "eventually (λx. ?S x = ?S' x) F"
using assms[THEN asymp_equiv_eventually_zeros] by eventually_elim auto
ultimately have "(?S' ⤏ 0) F" by (rule Lim_transform_eventually)
}
with * have "(?T (λx. f1 x * f2 x) (λx. g1 x * g2 x) ⤏ 1 * 1) F"
by (rule Lim_transform)
then show ?thesis by (simp add: asymp_equiv_def)
qed
lemma asymp_equiv_power [asymp_equiv_intros]:
"f ∼[F] g ⟹ (λx. f x ^ n) ∼[F] (λx. g x ^ n)"
by (induction n) (simp_all add: asymp_equiv_mult)
lemma asymp_equiv_inverse [asymp_equiv_intros]:
assumes "f ∼[F] g"
shows "(λx. inverse (f x)) ∼[F] (λx. inverse (g x))"
proof -
from tendsto_inverse[OF asymp_equivD[OF assms]]
have "((λx. if f x = 0 ∧ g x = 0 then 1 else g x / f x) ⤏ 1) F"
by (simp add: if_distrib cong: if_cong)
also have "(λx. if f x = 0 ∧ g x = 0 then 1 else g x / f x) =
(λx. if f x = 0 ∧ g x = 0 then 1 else inverse (f x) / inverse (g x))"
by (intro ext) (simp add: field_simps)
finally show ?thesis by (simp add: asymp_equiv_def)
qed
lemma asymp_equiv_inverse_iff [asymp_equiv_simps]:
"(λx. inverse (f x)) ∼[F] (λx. inverse (g x)) ⟷ f ∼[F] g"
proof
assume "(λx. inverse (f x)) ∼[F] (λx. inverse (g x))"
hence "(λx. inverse (inverse (f x))) ∼[F] (λx. inverse (inverse (g x)))" (is ?P)
by (rule asymp_equiv_inverse)
also have "?P ⟷ f ∼[F] g" by (intro asymp_equiv_cong) simp_all
finally show "f ∼[F] g" .
qed (simp_all add: asymp_equiv_inverse)
lemma asymp_equiv_divide [asymp_equiv_intros]:
assumes "f1 ∼[F] g1" "f2 ∼[F] g2"
shows "(λx. f1 x / f2 x) ∼[F] (λx. g1 x / g2 x)"
using asymp_equiv_mult[OF assms(1) asymp_equiv_inverse[OF assms(2)]] by (simp add: field_simps)
lemma asymp_equivD_strong:
assumes "f ∼[F] g" "eventually (λx. f x ≠ 0 ∨ g x ≠ 0) F"
shows "((λx. f x / g x) ⤏ 1) F"
proof -
from assms(1) have "((λx. if f x = 0 ∧ g x = 0 then 1 else f x / g x) ⤏ 1) F"
by (rule asymp_equivD)
also have "?this ⟷ ?thesis"
by (intro filterlim_cong eventually_mono[OF assms(2)]) auto
finally show ?thesis .
qed
lemma asymp_equiv_compose [asymp_equiv_intros]:
assumes "f ∼[G] g" "filterlim h G F"
shows "f ∘ h ∼[F] g ∘ h"
proof -
let ?T = "λf g x. if f x = 0 ∧ g x = 0 then 1 else f x / g x"
have "f ∘ h ∼[F] g ∘ h ⟷ ((?T f g ∘ h) ⤏ 1) F"
by (simp add: asymp_equiv_def o_def)
also have "… ⟷ (?T f g ⤏ 1) (filtermap h F)"
by (rule tendsto_compose_filtermap)
also have "…"
by (rule tendsto_mono[of _ G]) (insert assms, simp_all add: asymp_equiv_def filterlim_def)
finally show ?thesis .
qed
lemma asymp_equiv_compose':
assumes "f ∼[G] g" "filterlim h G F"
shows "(λx. f (h x)) ∼[F] (λx. g (h x))"
using asymp_equiv_compose[OF assms] by (simp add: o_def)
lemma asymp_equiv_powr_real [asymp_equiv_intros]:
fixes f g :: "'a ⇒ real"
assumes "f ∼[F] g" "eventually (λx. f x ≥ 0) F" "eventually (λx. g x ≥ 0) F"
shows "(λx. f x powr y) ∼[F] (λx. g x powr y)"
proof -
let ?T = "λf g x. if f x = 0 ∧ g x = 0 then 1 else f x / g x"
have "((λx. ?T f g x powr y) ⤏ 1 powr y) F"
by (intro tendsto_intros asymp_equivD[OF assms(1)]) simp_all
hence "((λx. ?T f g x powr y) ⤏ 1) F" by simp
moreover have "eventually (λx. ?T f g x powr y = ?T (λx. f x powr y) (λx. g x powr y) x) F"
using asymp_equiv_eventually_zeros[OF assms(1)] assms(2,3)
by eventually_elim (auto simp: powr_divide)
ultimately show ?thesis unfolding asymp_equiv_def by (rule Lim_transform_eventually)
qed
lemma asymp_equiv_norm [asymp_equiv_intros]:
fixes f g :: "'a ⇒ 'b :: real_normed_field"
assumes "f ∼[F] g"
shows "(λx. norm (f x)) ∼[F] (λx. norm (g x))"
using tendsto_norm[OF asymp_equivD[OF assms]]
by (simp add: if_distrib asymp_equiv_def norm_divide cong: if_cong)
lemma asymp_equiv_abs_real [asymp_equiv_intros]:
fixes f g :: "'a ⇒ real"
assumes "f ∼[F] g"
shows "(λx. ¦f x¦) ∼[F] (λx. ¦g x¦)"
using tendsto_rabs[OF asymp_equivD[OF assms]]
by (simp add: if_distrib asymp_equiv_def cong: if_cong)
lemma asymp_equiv_imp_eventually_le:
assumes "f ∼[F] g" "c > 1"
shows "eventually (λx. norm (f x) ≤ c * norm (g x)) F"
proof -
from order_tendstoD(2)[OF asymp_equivD[OF asymp_equiv_norm[OF assms(1)]] assms(2)]
asymp_equiv_eventually_zeros[OF assms(1)]
show ?thesis by eventually_elim (auto split: if_splits simp: field_simps)
qed
lemma asymp_equiv_imp_eventually_ge:
assumes "f ∼[F] g" "c < 1"
shows "eventually (λx. norm (f x) ≥ c * norm (g x)) F"
proof -
from order_tendstoD(1)[OF asymp_equivD[OF asymp_equiv_norm[OF assms(1)]] assms(2)]
asymp_equiv_eventually_zeros[OF assms(1)]
show ?thesis by eventually_elim (auto split: if_splits simp: field_simps)
qed
lemma asymp_equiv_imp_bigo:
assumes "f ∼[F] g"
shows "f ∈ O[F](g)"
proof (rule bigoI)
have "(3/2::real) > 1" by simp
from asymp_equiv_imp_eventually_le[OF assms this]
show "eventually (λx. norm (f x) ≤ 3/2 * norm (g x)) F"
by eventually_elim simp
qed
lemma asymp_equiv_imp_bigomega:
"f ∼[F] g ⟹ f ∈ Ω[F](g)"
using asymp_equiv_imp_bigo[of g F f] by (simp add: asymp_equiv_sym bigomega_iff_bigo)
lemma asymp_equiv_imp_bigtheta:
"f ∼[F] g ⟹ f ∈ Θ[F](g)"
by (intro bigthetaI asymp_equiv_imp_bigo asymp_equiv_imp_bigomega)
lemma asymp_equiv_at_infinity_transfer:
assumes "f ∼[F] g" "filterlim f at_infinity F"
shows "filterlim g at_infinity F"
proof -
from assms(1) have "g ∈ Θ[F](f)" by (rule asymp_equiv_imp_bigtheta[OF asymp_equiv_symI])
also from assms have "f ∈ ω[F](λ_. 1)" by (simp add: smallomega_1_conv_filterlim)
finally show ?thesis by (simp add: smallomega_1_conv_filterlim)
qed
lemma asymp_equiv_at_top_transfer:
fixes f g :: "_ ⇒ real"
assumes "f ∼[F] g" "filterlim f at_top F"
shows "filterlim g at_top F"
proof (rule filterlim_at_infinity_imp_filterlim_at_top)
show "filterlim g at_infinity F"
by (rule asymp_equiv_at_infinity_transfer[OF assms(1) filterlim_mono[OF assms(2)]])
(auto simp: at_top_le_at_infinity)
from assms(2) have "eventually (λx. f x > 0) F"
using filterlim_at_top_dense by blast
with asymp_equiv_eventually_pos_iff[OF assms(1)] show "eventually (λx. g x > 0) F"
by eventually_elim blast
qed
lemma asymp_equiv_at_bot_transfer:
fixes f g :: "_ ⇒ real"
assumes "f ∼[F] g" "filterlim f at_bot F"
shows "filterlim g at_bot F"
unfolding filterlim_uminus_at_bot
by (rule asymp_equiv_at_top_transfer[of "λx. -f x" F "λx. -g x"])
(insert assms, auto simp: filterlim_uminus_at_bot asymp_equiv_uminus)
lemma asymp_equivI'_const:
assumes "((λx. f x / g x) ⤏ c) F" "c ≠ 0"
shows "f ∼[F] (λx. c * g x)"
using tendsto_mult[OF assms(1) tendsto_const[of "inverse c"]] assms(2)
by (intro asymp_equivI') (simp add: field_simps)
lemma asymp_equivI'_inverse_const:
assumes "((λx. f x / g x) ⤏ inverse c) F" "c ≠ 0"
shows "(λx. c * f x) ∼[F] g"
using tendsto_mult[OF assms(1) tendsto_const[of "c"]] assms(2)
by (intro asymp_equivI') (simp add: field_simps)
lemma filterlim_at_bot_imp_at_infinity: "filterlim f at_bot F ⟹ filterlim f at_infinity F"
for f :: "_ ⇒ real" using at_bot_le_at_infinity filterlim_mono by blast
lemma asymp_equiv_imp_diff_smallo:
assumes "f ∼[F] g"
shows "(λx. f x - g x) ∈ o[F](g)"
proof (rule landau_o.smallI)
fix c :: real assume "c > 0"
hence c: "min c 1 > 0" by simp
let ?h = "λx. if f x = 0 ∧ g x = 0 then 1 else f x / g x"
from assms have "((λx. ?h x - 1) ⤏ 1 - 1) F"
by (intro tendsto_diff asymp_equivD tendsto_const)
from tendstoD[OF this c] show "eventually (λx. norm (f x - g x) ≤ c * norm (g x)) F"
proof eventually_elim
case (elim x)
from elim have "norm (f x - g x) ≤ norm (f x / g x - 1) * norm (g x)"
by (subst norm_mult [symmetric]) (auto split: if_splits simp add: algebra_simps)
also have "norm (f x / g x - 1) * norm (g x) ≤ c * norm (g x)" using elim
by (auto split: if_splits simp: mult_right_mono)
finally show ?case .
qed
qed
lemma asymp_equiv_altdef:
"f ∼[F] g ⟷ (λx. f x - g x) ∈ o[F](g)"
by (rule iffI[OF asymp_equiv_imp_diff_smallo smallo_imp_asymp_equiv])
lemma asymp_equiv_0_left_iff [simp]: "(λ_. 0) ∼[F] f ⟷ eventually (λx. f x = 0) F"
and asymp_equiv_0_right_iff [simp]: "f ∼[F] (λ_. 0) ⟷ eventually (λx. f x = 0) F"
by (simp_all add: asymp_equiv_altdef landau_o.small_refl_iff)
lemma asymp_equiv_sandwich_real:
fixes f g l u :: "'a ⇒ real"
assumes "l ∼[F] g" "u ∼[F] g" "eventually (λx. f x ∈ {l x..u x}) F"
shows "f ∼[F] g"
unfolding asymp_equiv_altdef
proof (rule landau_o.smallI)
fix c :: real assume c: "c > 0"
have "eventually (λx. norm (f x - g x) ≤ max (norm (l x - g x)) (norm (u x - g x))) F"
using assms(3) by eventually_elim auto
moreover have "eventually (λx. norm (l x - g x) ≤ c * norm (g x)) F"
"eventually (λx. norm (u x - g x) ≤ c * norm (g x)) F"
using assms(1,2) by (auto simp: asymp_equiv_altdef dest: landau_o.smallD[OF _ c])
hence "eventually (λx. max (norm (l x - g x)) (norm (u x - g x)) ≤ c * norm (g x)) F"
by eventually_elim simp
ultimately show "eventually (λx. norm (f x - g x) ≤ c * norm (g x)) F"
by eventually_elim (rule order.trans)
qed
lemma asymp_equiv_sandwich_real':
fixes f g l u :: "'a ⇒ real"
assumes "f ∼[F] l" "f ∼[F] u" "eventually (λx. g x ∈ {l x..u x}) F"
shows "f ∼[F] g"
using asymp_equiv_sandwich_real[of l F f u g] assms by (simp add: asymp_equiv_sym)
lemma asymp_equiv_sandwich_real'':
fixes f g l u :: "'a ⇒ real"
assumes "l1 ∼[F] u1" "u1 ∼[F] l2" "l2 ∼[F] u2"
"eventually (λx. f x ∈ {l1 x..u1 x}) F" "eventually (λx. g x ∈ {l2 x..u2 x}) F"
shows "f ∼[F] g"
by (meson assms asymp_equiv_sandwich_real asymp_equiv_sandwich_real' asymp_equiv_trans)
end