Theory BigO

Up to index of Isabelle/HOL/HOL-Complex/NumberTheory

theory BigO = SetsAndFunctions + RingLib + FiniteLib + RealLib:

(*  Title:      BigO.thy
    Authors:    Jeremy Avigad and Kevin Donnelly
*)

header {* Big O notation *}

theory BigO = SetsAndFunctions + RingLib + FiniteLib + RealLib:

subsection {* Preliminaries *}

text {*
Note: since the Big O library includes rules that demonstrate set inclusion,
to use the automated reasoners with the library one should redeclare subsetI 
as an intro rule, instead of an intro! rule *}

text {* Missing transitivity rule! *}

lemma eq_in_trans [trans]: "a = b ==> b ∈ C ==> a ∈ C";
by auto;

subsection {* Definitions *}

constdefs 

  bigo :: "('a => 'b::ordered_ring) => ('a => 'b) set"           ("(1O'(_'))")
  "O(f::('a => 'b::ordered_ring)) == {h. ∃c. ∀x. abs (h x) ≤ c * abs (f x)}"

  bigoset :: "('a => 'b::ordered_ring) set => ('a => 'b) set"    ("(1O'(_'))")
  "O(S::('a => 'b::ordered_ring) set) == 
    {h. ∃f∈S. ∃c. ∀x. abs(h x) ≤ (c * abs(f x))}";


subsection {* Basic properties *}

lemma bigo_pos_const: "(∃c. ∀x. (abs (h x)) ≤ (c * (abs (f x))))
  = (∃c. 0 < (c::'a::ordered_ring) ∧ (∀x. (abs(h x)) ≤ (c * (abs (f x)))))";
apply(auto);
apply(case_tac "c = 0");
apply(simp);
apply(rule_tac x = "1" in exI);
apply(auto simp add: zero_less_one);
apply(rule_tac x = "abs c" in exI);
apply auto;
apply (subgoal_tac "c * abs(f x) ≤ abs c * abs (f x)");
apply (erule_tac x = x in allE);
apply force;
apply (rule mult_right_mono);
apply (rule abs_ge_self);
apply (rule abs_ge_zero);
done;

lemma bigo_alt_def: "O(f::('a => 'b::ordered_ring)) = 
  {h. ∃c. (0 < c ∧ (∀x. abs (h x) ≤ c * abs (f x)))}";
by (auto simp add: bigo_def bigo_pos_const);

lemma bigoset_alt_def: "O(S::('a => 'b::ordered_ring) set) = 
    {h. ∃f∈S. ∃c. (0 < c & (∀x. abs(h x) ≤ c * abs(f x)))}";
by (auto simp add: bigoset_def bigo_pos_const);

lemma bigoset_alt_def2: "O(A::('a => 'b::ordered_ring) set) = 
    { h. ∃f∈A. h ∈ O(f) }";
  by (simp add: bigoset_def bigo_def);

lemma bigo_elt_subset [intro]: "f ∈ O(g::('a => 'b::ordered_ring))
  ==> O(f) ⊆ O(g)";
apply (auto simp add: bigo_alt_def);
apply (rule_tac x = "ca * c" in exI);
apply (rule conjI);
apply (rule mult_pos);
apply (assumption)+;
apply (rule allI);
apply (drule_tac x = "xa" in spec)+;
apply (subgoal_tac "ca * abs(f xa) ≤ ca * (c * abs(g xa))");
apply (erule order_trans);
apply (simp add: mult_ac);
apply (rule mult_left_mono, assumption);
by (rule order_less_imp_le, assumption);

lemma bigoset_elt_subset: "f ∈ O(A::('a => 'b::ordered_ring) set)
  ==> O(f) ⊆ O(A)";
apply (auto simp add: bigo_alt_def bigoset_alt_def);
apply (rule_tac x = fa in bexI);
apply (rule_tac x = "ca * c" in exI);
apply (rule conjI);
apply (rule mult_pos);
apply (assumption)+;
apply (rule allI);
apply (drule_tac x = "xa" in spec)+;
apply (subgoal_tac "ca * abs(f xa) ≤ ca * (c * abs(fa xa))");
apply (erule order_trans);
apply (simp add: mult_ac);
apply (rule mult_left_mono, assumption);
by (rule order_less_imp_le, assumption, assumption);

lemma bigoset_elt_subset2: "f ∈ A ==> f ∈ O(A)";
apply (auto simp add: bigoset_def);
apply (rule_tac x = f in bexI);
apply (rule_tac x = 1 in exI);
by auto;

lemma bigoset_mono: "A ⊆ B ==> O(A) ⊆ O(B)";
by (auto simp add: bigoset_def);

lemma bigo_refl [intro]: "f ∈ O(f)";
apply(auto simp add: bigo_def);
apply(rule_tac x = 1 in exI);
by simp;

lemma bigoset_refl: "A ⊆ O(A)";
apply (auto simp add: bigoset_def);
apply (rule_tac x = x in bexI);
apply (rule_tac x = 1 in exI);
by auto;

lemma bigo_bigo_refl: "f ∈ O(O(f))";
apply (insert bigo_refl [of f]);
apply (insert bigoset_refl [of "O(f)"]);
by (erule set_rev_mp, assumption);

lemma bigo_bigo_eq: "O(O(f)) = O(f::('a => 'b::ordered_ring))";
apply(subgoal_tac "(O(O(f)) = O(f)) = (O(f) = O(O(f)))",erule ssubst);
apply(simp add: bigo_def bigoset_def bigo_refl bigo_pos_const);
apply(auto);
apply(simp_all only: bigo_pos_const[THEN sym]);
apply(rule_tac x = "f" in exI);
apply(auto);
apply(rule_tac x = "1" in exI);
apply(auto);
apply(rule_tac x = "ca * c" in exI);
apply(auto);
apply(erule_tac x = "xa" in allE);
apply(erule_tac x = "xa" in allE);
apply(subgoal_tac "ca * abs(fa xa) ≤ ca * c * abs(f xa)");
apply force;
apply(subgoal_tac "ca * abs (fa xa) ≤ ca * (c * abs (f xa))");
apply(auto simp add: mult_assoc mult_left_mono);
apply(rule mult_left_mono);
by(auto simp add: order_less_imp_le);

lemma bigo_zero: "0 ∈ O(g::('a => 'b::ordered_ring))";
apply (auto simp add: bigo_def func_zero);
apply (rule_tac x = 0 in exI);
by auto;

lemma bigo_zero2: "O(λx.0) = {λx.0}";
apply (auto simp add: bigo_def); 
apply (rule ext);
apply auto;
done;

lemma bigo_plus_self_subset [intro]: 
  "O((f::'a => 'b::ordered_ring)) + O(f) ⊆ O(f)";
apply (auto simp add: bigo_alt_def bigoset_alt_def set_plus);
apply (rule_tac x = "c + ca" in exI);
apply auto;
apply (elim pos_plus_pos);
apply assumption;
apply (simp add: func_plus);
apply (drule_tac x = x in spec)+;
apply (subgoal_tac "abs (a x + b x) ≤ abs (a x) + abs (b x)");
apply (erule order_trans);
apply (simp add: ring_distrib);
apply (erule add_mono);
apply (assumption);
by (rule abs_triangle_ineq);

lemma bigo_plus_idemp [simp]: "O(f::('a=>'b::ordered_ring)) + O(f) = O(f)";
apply (rule equalityI);
apply (rule bigo_plus_self_subset);
apply (rule set_zero_plus2); 
by (rule bigo_zero);

lemma bigo_plus_subset_lemma:   "x ∈ O(f + g) ==> x ∈ O(f) + 
  O(g::'a=>'b::ordered_ring)";
apply(auto simp add: bigo_def bigo_pos_const func_plus set_plus);
apply(simp only: bigo_pos_const[THEN sym]);
apply(rule_tac x = "%n. if abs (g n) ≤ (abs (f n)) then x n else 0" in exI);
apply(rule conjI);
apply(rule_tac x = "c + c" in exI);
apply(clarsimp);
apply(auto);
apply(subgoal_tac "c * abs (f xa + g xa) ≤ (c + c) * abs (f xa)");
apply(erule_tac x = xa in allE);
apply(erule order_trans);
apply(simp);
apply(subgoal_tac "c * abs (f xa + g xa) ≤ c * (abs (f xa) + abs (g xa))");
apply(erule order_trans);
apply(simp add: ring_distrib);
apply(rule mult_left_mono);
apply assumption;
apply(simp add: order_less_le);
apply(rule mult_left_mono);
apply(simp add: abs_triangle_ineq);
apply(simp add: order_less_le);
apply(subgoal_tac "abs (0::'b) = 0");
apply(erule ssubst);
apply (rule nonneg_times_nonneg);
apply (rule nonneg_plus_nonneg);
apply auto;
apply(rule_tac x = "%n. if (abs (f n)) <  abs (g n) then x n else 0" in exI);
apply(rule conjI);
apply(rule_tac x = "c + c" in exI);
apply(clarsimp);
apply(auto);
apply(subgoal_tac "c * abs (f xa + g xa) ≤ (c + c) * abs (g xa)");
apply(erule_tac x = xa in allE)
apply(erule order_trans);
apply(simp);
apply(subgoal_tac "c * abs (f xa + g xa) ≤ c * (abs (f xa) + abs (g xa))");
apply(erule order_trans);
apply(simp add: ring_distrib);
apply (rule mult_left_mono);
apply(simp add: order_less_le);
apply(simp add: order_less_le);
apply (rule mult_left_mono);
apply (rule abs_triangle_ineq);
apply(simp add: order_less_le);
apply (rule nonneg_times_nonneg);
apply (rule nonneg_plus_nonneg);
apply (erule order_less_imp_le)+;
apply simp;
apply (rule ext);
apply (auto simp add: if_splits linorder_not_le);
done;

lemma bigo_plus_subset [intro]: "O(a + b) ⊆ O(a) + O(b::'a=>'b::ordered_ring)";
by (rule subsetI, rule bigo_plus_subset_lemma, assumption);

lemma bigo_plus_subset2: "O(f \<pluso> A) ⊆ O(f) + O(A)";
apply (auto simp add: bigoset_alt_def2 elt_set_plus_def set_plus);
apply (frule bigo_plus_subset_lemma);
by (auto simp add: set_plus);

lemma bigo_plus_subset3: "O((A::('a => 'b::ordered_ring) set) + B) ⊆
    O(A) + O(B)";
apply (auto simp add: bigoset_alt_def2 set_plus);
apply (frule bigo_plus_subset_lemma);
by (auto simp add: set_plus);

lemma bigo_plus_subset4 [intro]: "[| A ⊆ O(f::('a => 'b::ordered_ring));
  B ⊆ O(f) |] ==> A + B ⊆ O(f)"; 
proof -;
  assume "A ⊆ O(f::('a => 'b::ordered_ring))" and "B ⊆ O(f)";
  then have "A + B ⊆ O(f) + O(f)";
    by (auto del: subsetI simp del: bigo_plus_idemp);
  also have "… ⊆ O(f)";
    by (auto del: subsetI);
  finally show ?thesis;.;
qed;

lemma bigo_plus_subset5: "O((f::('a => 'b::ordered_ring)) \<pluso> 
    O(g::('a => 'b))) ⊆ O(f) + O(g)";
proof -;
  have "O((f::('a => 'b::ordered_ring)) \<pluso> 
    O(g::('a => 'b))) ⊆ O(f) + O(O(g))";
    by (rule bigo_plus_subset2);
  thus ?thesis; by (simp add: bigo_bigo_eq);
qed;

lemma bigo_plus_subset6: "[|∀x. 0 ≤ f x; ∀x. 0 ≤ g x |] ==> 
  O(f + g) = O(f) + O(g)";
apply (rule equalityI);
apply (rule bigo_plus_subset);
apply (simp add: bigo_alt_def set_plus func_plus);
apply clarify;
apply (rule_tac x = "max c ca" in exI);
apply (rule conjI);
apply (subgoal_tac "c ≤ max c ca");
apply (erule order_less_le_trans);
apply assumption;
apply (rule le_maxI1);
apply clarify;
apply (drule_tac x = "xa" in spec)+;
apply (subgoal_tac "0 ≤ f xa + g xa");
apply (simp add: ring_distrib abs_nonneg);
apply (subgoal_tac "abs(a xa + b xa) ≤ abs(a xa) + abs(b xa)");
apply (subgoal_tac "abs(a xa) + abs(b xa) ≤ 
    max c ca * f xa + max c ca * g xa");
apply (force);
apply (rule add_mono);
apply (subgoal_tac "c * f xa ≤ max c ca * f xa");
apply (force);
apply (rule mult_right_mono);
apply (rule le_maxI1);
apply assumption;
apply (subgoal_tac "ca * g xa ≤ max c ca * g xa");
apply (force);
apply (rule mult_right_mono);
apply (rule le_maxI2);
apply assumption;
apply (rule abs_triangle_ineq);
apply (rule nonneg_plus_nonneg);
apply assumption+;
done;

lemma bigo_elt_subset2 [intro]: "f ∈ g \<pluso> O(h::('a => 'b::ordered_ring)) ==> 
  O(f) ⊆ O(g) + O(h)";
proof -;
  assume "f ∈ g \<pluso> O(h::('a => 'b::ordered_ring))";
  then have "O(f) ⊆ O(g \<pluso> O(h))";
    apply (intro bigoset_elt_subset);
    by (rule bigoset_elt_subset2);
  also have "… ⊆ O(g) + O(h)";
    by (rule bigo_plus_subset5);
  finally show ?thesis;.;
qed;

lemma bigo_mult [intro]: "O(f)*O(g) ⊆ O(f * g::('a => 'b::ordered_ring))"
apply(auto simp add: bigo_def bigo_pos_const set_times func_times);
apply(rule_tac x = "c * ca" in exI);
apply(rule allI);
apply(erule_tac x = x in allE);
apply(erule_tac x = x in allE);
apply(subgoal_tac "c * ca * (abs(f x) * abs(g x)) = (c * abs(f x)) * (ca * abs(g x))");
apply(erule ssubst);
apply (rule mult_mono);
apply assumption+;
apply (rule nonneg_times_nonneg);
apply auto;
apply (simp add: mult_ac);
done;

text {* For reals, can make this = *}

lemma bigo_mult2 [intro]: "f \<timeso> O(g::('a => 'b::ordered_ring)) ⊆
   O(f * g)";
apply (auto simp add: bigo_def elt_set_times_def func_times abs_mult);
apply (rule_tac x = c in exI);
apply auto;
apply (drule_tac x = x in spec);
apply (subgoal_tac "abs(f x) * abs(b x) ≤ abs(f x) * (c * abs(g x))");
apply (force simp add: mult_ac);
apply (rule mult_left_mono, assumption);
by (rule abs_ge_zero);

lemma bigo_mult3: "[| f : O(h::'a=>'b::ordered_ring); g : O(j) |] ==> f * g : O(h * j)";
apply(subgoal_tac "f * g : O(h) * O(j)");
apply(insert bigo_mult[of h j]);
apply(simp only: subsetD);
by(auto simp add: bigo_def func_times set_times);

lemma bigo_mult4[intro]:"[| f : k +o O(h::'a=>'b::ordered_ring) |] ==> g * f : (
g * k) +o O(g * h)";
apply(simp add: bigo_def elt_set_plus_def func_times);
apply(clarsimp);
apply(rule_tac x = "g * b" in exI);
apply(rule conjI);
apply(rule_tac x = c in exI);
apply(clarsimp);
apply(simp add: func_times abs_mult);
apply(erule_tac x = x in allE);
apply(subgoal_tac "c * (abs (g x) * abs (h x)) = abs (g x) * (c * abs (h x))");
apply(erule ssubst);
apply(rule mult_left_mono);
apply(simp_all add: abs_ge_zero);
apply(simp add: times_ac1);
by(simp add: func_times func_plus ring_distrib);

lemma bigo_minus [intro]: "f : O(g::'a=>'b::ordered_ring) ==> - f : O(g)";
apply(auto simp add: bigo_def);
apply(rule_tac x = c in exI);
apply(rule allI);
apply(erule_tac x = x in allE);
apply(subgoal_tac "abs ((- f) x) = abs (f x)");
apply(erule ssubst);
apply(simp);
apply(auto simp add: func_minus);
done;

lemma bigo_mult5: "ALL x. f x ~= 0 ==>
    O(f * g) <= (f::'a => ('b::ordered_field)) *o O(g)";
proof -;
  assume "ALL x. f x ~= 0";
  show "O(f * g) <= f *o O(g)";
  proof;
    fix h;
    assume "h : O(f * g)";
    then have "(%x. 1 / (f x)) * h : (%x. 1 / f x) *o O(f * g)";
      by auto;
    also have "... <= O((%x. 1 / f x) * (f * g))";
      by (rule bigo_mult2);
    also have "(%x. 1 / f x) * (f * g) = g";
      apply (simp add: func_times); 
      apply (rule ext);
      apply (simp add: prems nonzero_divide_eq_eq mult_ac);
      done;
    finally; have "(%x. (1::'b) / f x) * h : O(g)";.;
    then have "f * ((%x. (1::'b) / f x) * h) : f *o O(g)";
      by auto;
    also have "f * ((%x. (1::'b) / f x) * h) = h";
      apply (simp add: func_times); 
      apply (rule ext);
      apply (simp add: prems nonzero_divide_eq_eq mult_ac);
      done;
    finally show "h : f *o O(g)";.;
  qed;
qed;

lemma bigo_mult6: "ALL x. f x ~= 0 ==>
    O(f * g) = (f::'a => ('b::ordered_field)) *o O(g)";
apply (rule equalityI);
apply (erule bigo_mult5);
apply (rule bigo_mult2);
done;

lemma bigo_mult7: "ALL x. f x ~= 0 ==>
    O(f * g) <= O(f::'a => ('b::ordered_field)) * O(g)";
  apply (subst bigo_mult6);
  apply assumption;
  apply (rule set_times_mono3);
  apply (rule bigo_refl);
done;

lemma bigo_mult8: "ALL x. f x ~= 0 ==>
    O(f * g) = O(f::'a => ('b::ordered_field)) * O(g)";
apply (rule equalityI);
apply (erule bigo_mult7);
apply (rule bigo_mult);
done;

lemma bigo_minus2: "f ∈ g \<pluso> O(h::('a => ('b::ordered_ring))) ==> 
-f ∈ -g \<pluso> O(h)";
proof -;
  assume "f ∈ g \<pluso> O(h::('a => ('b::ordered_ring)))";
  then have "f - g ∈ O(h)";
    by (rule set_plus_imp_minus);
  then have "-(f - g) ∈ O(h)";
    by (rule bigo_minus);
  also have "-(f - g) = -f - -g";
    by (simp add: diff_minus plus_ac0);
  finally have "-f - -g ∈ O(h)";.;
  then show "-f ∈ -g \<pluso> O(h)";
    by (rule set_minus_imp_plus);
qed;

lemma bigo_minus3: "O(-f) = O(f::('a => 'b::ordered_ring))";
by (auto simp add: bigo_def func_minus abs_minus_cancel);

declare subsetI [rule del, intro];

lemma bigo_plus_absorb_lemma1 [intro]: "f ∈ O(g::('a => 'b::ordered_ring)) ==> 
  f \<pluso> O(g) ⊆ O(g)";
proof -;
  assume a: "f ∈ O(g)";
  show "f \<pluso> O(g) ⊆ O(g)";
  proof -;
    have "f ∈ O(f)" by auto;
    then have "f \<pluso> O(g) ⊆ O(f) + O(g)";
      by auto;
    also have "… ⊆ O(g) + O(g)";
    proof -;
      from a have "O(f) ⊆ O(g)" by auto;
      thus ?thesis; by auto;
    qed;
    also have "… ⊆ O(g)"; by (simp add: bigo_plus_idemp);
    finally show ?thesis;.;
  qed;
qed;

lemma bigo_plus_absorb_lemma2 [intro]: "f ∈ O(g::('a => 'b::ordered_ring)) ==> 
    O(g) ⊆ f \<pluso> O(g)";
proof -;
  assume a: "f ∈ O(g::('a => 'b::ordered_ring))";
  show "O(g) ⊆ f \<pluso> O(g)";
  proof -;
    from a have "-f ∈ O(g)"; by auto;
    then have "-f \<pluso> O(g) ⊆ O(g)"; by auto;
    then have "f \<pluso> (-f \<pluso> O(g)) ⊆ f \<pluso> O(g)"; by auto;
    also have "f \<pluso> (-f \<pluso> O(g)) = O(g)";
      by (simp add: set_plus_rearranges);
    finally show ?thesis;.;
  qed;
qed;

lemma bigo_plus_absorb [simp]: "f ∈ O(g::('a => 'b::ordered_ring)) ==> 
    f \<pluso> O(g) = O(g)";
apply (rule equalityI);
apply (rule bigo_plus_absorb_lemma1, assumption);
by (rule bigo_plus_absorb_lemma2);

lemma bigo_plus_absorb2 [intro]: "f ∈ O(g::('a::type => 'b::ordered_ring)) ==> 
  A ⊆ O(g) ==> f \<pluso> A ⊆ O(g)";
apply (subgoal_tac "f \<pluso> A ⊆ f \<pluso> O(g)");
by force+;

lemma bigo_add_commute_imp: "(f : g +o O(h::'a=>'b::ordered_ring)) ==> 
  (g : f +o O(h))";
apply (subst set_minus_plus [THEN sym]);
apply (subgoal_tac "g - f = - (f - g)");
apply (erule ssubst);
apply (rule bigo_minus);
apply (subst set_minus_plus);
apply assumption;
by (simp add: diff_minus add_ac);

lemma bigo_add_commute: "(f : g +o O(h::'a=>'b::ordered_ring)) = 
 (g : f +o O(h))";
apply (rule iffI);
apply (rule bigo_add_commute_imp, assumption);
by (rule bigo_add_commute_imp, assumption);

lemma bigo_bounded: "[| ∀x. 0 ≤ f x; ∀x. f x ≤ g x |] ==> f ∈ O(g)"; 
apply (auto simp add: bigo_def);
apply (rule_tac x = 1 in exI);
apply auto;
apply (drule_tac x = x in spec)+;
apply (subgoal_tac "0 ≤ g x");
apply (simp add: abs_nonneg);
apply (erule order_trans, assumption);
done;

lemma bigo_bounded_alt: "[| ∀x. 0 ≤ f x; ∀x. f x ≤ c * g x |] ==> f ∈ O(g)"; 
apply (auto simp add: bigo_def);
apply (rule_tac x = "abs c" in exI);
apply auto;
apply (drule_tac x = x in spec)+;
apply (subgoal_tac "abs(c) * abs(g x) = abs(c * g x)");
apply (erule ssubst);
apply (subgoal_tac "0 ≤ c * g x");
apply (simp only: abs_nonneg);
apply (erule order_trans, assumption);
apply simp;
done;

lemma bigo_bounded2: "[| ∀x. lb x ≤ f x; ∀x. f x ≤ lb x + g x |] ==> 
  f ∈ lb +o O(g)";
apply (rule set_minus_imp_plus);
apply (rule bigo_bounded);
apply (auto simp add: diff_minus func_minus func_plus);
apply (drule_tac x = x in spec)+;
apply (subgoal_tac "-lb x + lb x ≤ - lb x + f x");
apply (simp add: plus_ac0);
apply (erule add_left_mono);
apply (drule_tac x = x in spec)+;
apply (subgoal_tac "-lb x + f x ≤ -lb x + (lb x + g x)");
apply (simp add: plus_ac0);
by (erule add_left_mono);

lemma bigo_bounded3: "[| ∀x. lb x ≤ f x & f x ≤ lb x + g x |] ==>
  f ∈ lb +o O(g)";
  by (rule bigo_bounded2, auto);

lemma bigo_abs: "(%x. abs(f x)) =o O(f)"; 
  apply (unfold bigo_def);
  apply auto;
  apply (rule_tac x = 1 in exI);
  apply auto;
done;

lemma bigo_abs2: "f =o O(%x. abs(f x))";
  apply (unfold bigo_def);
  apply auto;
  apply (rule_tac x = 1 in exI);
  apply auto;
done;

lemma bigo_abs3: "O(f) = O(%x. abs(f x))";
  apply (rule equalityI);
  apply (rule bigo_elt_subset);
  apply (rule bigo_abs2);
  apply (rule bigo_elt_subset);
  apply (rule bigo_abs);
done;

lemma bigo_abs4: "f =o g +o O(h::'a=>'b::ordered_ring) ==>
    (%x. abs (f x)) =o (%x. abs (g x)) +o O(h)";
  apply (drule set_plus_imp_minus);
  apply (rule set_minus_imp_plus);
  apply (subst func_diff);
proof -;
  assume a: "f - g : O(h)";
  have "(%x. abs (f x) - abs (g x)) =o O(%x. abs(abs (f x) - abs (g x)))";
    by (rule bigo_abs2);
  also have "... <= O(%x. abs (f x - g x))";
    apply (rule bigo_elt_subset);
    apply (rule bigo_bounded);
    apply force;
    apply (rule allI);
    apply (rule abs_triangle_ineq3);
    done;
  also have "... <= O(f - g)";
    apply (rule bigo_elt_subset);
    apply (subst func_diff);
    apply (rule bigo_abs);
    done;
  also have "... <= O(h)";
    by (rule bigo_elt_subset);
  finally show "(%x. abs (f x) - abs (g x)) : O(h)";.;
qed;


-- {* generalize these beyond reals *}

lemma bigo_const1: "(λx. c) ∈ O(λx. 1)";
by (auto simp add: bigo_def mult_ac);

lemma bigo_const2 [intro]: "O(λx. c) ⊆ O(λx. 1)";
apply (rule bigo_elt_subset);
by (rule bigo_const1);

lemma bigo_const3: "(c::real) ~= 0 ==> (λx. 1) ∈ O(λx. c)";
apply (simp add: bigo_def);
apply (rule_tac x = "abs(inverse c)" in exI);
by (simp); 

lemma bigo_const4: "(c::real) ~= 0 ==> O(λx. 1) ⊆ O(λx. c)";
by (rule bigo_elt_subset, rule bigo_const3, assumption);

lemma bigo_const [intro]: "(c::real) ~= 0 ==> O(λx. c) = O(λx. 1)";
by (rule equalityI, rule bigo_const2, rule bigo_const4, assumption);

lemma bigo_const_mult1: "(λx. c * f x) ∈ O(f)";
apply (simp add: bigo_def);
apply (rule_tac x = "abs(c)" in exI);
by auto;

lemma bigo_const_mult2 [intro]: "O(λx. c * f x) ⊆ O(f)";
by (rule bigo_elt_subset, rule bigo_const_mult1);

lemma bigo_const_mult3: "(c::real) ~= 0 ==> f ∈ O(λx. c * f x)";
apply (simp add: bigo_def);
apply (rule_tac x = "abs(inverse c)" in exI);
by (simp add: mult_assoc [THEN sym]);

lemma bigo_const_mult4: "(c::real) ~= 0 ==> O(f) ⊆ O(λx. c * f x)";
by (rule bigo_elt_subset, rule bigo_const_mult3, assumption);

lemma bigo_const_mult [simp]: "(c::real) ~= 0 ==> O(λx. c * f x) = O(f)";
apply (rule equalityI, rule bigo_const_mult2, rule bigo_const_mult4);
by assumption;

lemma bigo_const_mult5 [simp]: "(c::real) ~= 0 ==> 
  (λx. c) \<timeso> O(f::('a => real)) = O(f)";
apply (auto intro!: subsetI simp add: bigo_def elt_set_times_def func_times);
apply (rule_tac x = "abs(c) * ca" in exI);
apply (auto); 
apply (subgoal_tac "abs(c) * ca * abs(f x) = abs(c) * (ca * abs(f x))");  
apply (erule ssubst);
apply (rule mult_mono);
apply auto;
apply (rule_tac x = "λ n. (1 / c) * x n" in exI);
apply auto;
apply (rule_tac x = "ca / abs(c)" in exI);
apply auto;
apply (subgoal_tac "abs(x xa / c) = abs(x xa) / abs(c)");
apply (erule ssubst);
apply (rule divide_right_mono);
apply (erule spec);
apply force;
apply (simp add: real_divide_def);
done;

lemma bigo_const_mult6 [intro]: "((%x. c) *o O(f::('b => 'a::ordered_ring))) ≤ O(f)";
apply(auto intro!: subsetI
  simp add: bigo_def elt_set_times_def func_times);
apply(rule_tac x = "ca * (abs c)" in exI);
apply(rule allI);
apply (subgoal_tac "ca * abs(c) * abs(f x) = abs(c) * (ca * abs(f x))");
apply (erule ssubst);
apply (rule mult_left_mono);
apply (erule spec);
apply simp;
apply(simp add: mult_ac);
done;

lemma bigo_const_mult7 [intro]: "f =o O(g::('a => 'b::ordered_ring)) ==> 
  (%x. c * f x) =o O(g)";
proof -;
  assume "f =o O(g)";
  then have "(%x. c) * f =o (%x. c) *o O(g)";
    by auto;
  also have "(%x. c) * f = (%x. c * f x)";
    by (simp add: func_times);
  also have "(%x. c) *o O(g) =s O(g)";
    by auto;
  finally show ?thesis;.;
qed;

lemma bigo_sumr_pos: "[| ∀x. 0 ≤ h x; f ∈ O(h) |] ==>
  (%x. sumr 0 x f) : O(%x. sumr 0 x h)";
apply(auto simp add: bigo_def bigo_pos_const);
apply(simp add: bigo_pos_const[THEN sym]);
apply(rule_tac x = c in exI);
apply(rule allI);
apply(induct_tac x);
apply(simp);
apply(simp);
apply(subgoal_tac "abs (sumr 0 n h + h n) = abs (sumr 0 n h) + abs(h n)");
apply(erule ssubst);
apply(simp add: ring_distrib);
apply(subgoal_tac "abs (sumr 0 n f + f n) <= abs (sumr 0 n f) + abs(f n)");
apply(erule order_trans);
apply(rule add_mono);
apply(simp);
apply(erule_tac x = n in allE);
apply(erule_tac x = n in allE);
apply(simp);
apply(simp add: abs_triangle_ineq);
apply(simp add: sumr_ge_zero real_abs_def);
apply(subgoal_tac "- sumr 0 n h <= - 0");
apply(erule_tac x = n in allE);
apply(simp);
apply(frule_tac x = n in allE);
apply(simp);
by(simp add: sumr_ge_zero);

declare abs_nonneg [simp del];
declare abs_nonpos [simp del];

lemma bigo_sumr_pos2: "[| ∀x. 0 ≤ h x; f ∈ g \<pluso> O(h) |] 
  ==> (λx. sumr 0 x f) ∈ (λx. sumr 0 x g) \<pluso> O(λx. sumr 0 x h)";
apply(rule set_minus_imp_plus);
apply(insert set_plus_imp_minus[of f g "O(h)"]);
apply(simp add: bigo_def bigo_pos_const elt_set_plus_def func_plus func_diff_minus func_minus);
apply(erule exE);
apply(erule exE);
apply(simp only: bigo_pos_const[THEN sym]);
apply(rule_tac x = c in exI);
apply(rule allI);
apply(rotate_tac 1);
apply(simp add: sumr_diff);
apply(subgoal_tac "c * abs (sumr 0 x h) = sumr 0 x (%x. c * abs(h x))");
apply(erule ssubst);
apply(subgoal_tac "abs (sumr 0 x b) <= sumr 0 x (%x. abs(b x))");
apply(erule order_trans);
apply(rule sumr_le2);
apply(rule allI);
apply(clarify);
apply(rotate_tac 2);
apply(erule_tac x = r in allE);
apply(simp);
apply(simp add: sumr_rabs);
apply(induct_tac x);
apply(simp);
apply(simp);
apply(subgoal_tac "abs (sumr 0 n h + h n) = abs (sumr 0 n h) + abs (h n)");
apply(rotate_tac -1);
apply(erule ssubst);
apply(simp add: ring_distrib);
apply(erule conjE);
apply(subgoal_tac "0 <= sumr 0 n h");
apply(erule_tac x = n in allE);
apply(simp add: real_abs_def);
apply auto;
apply (simp add: sumr_ge_zero);
done;

declare abs_nonneg [simp];
declare abs_nonpos [simp];

lemma bigo_sumr3: "f : O(h) ==>
    (%x. sumr 0 x f) : O(%x. sumr 0 x (%y. abs(h y)))";
  apply (rule bigo_sumr_pos);
  apply force;
  apply (subst bigo_abs3 [THEN sym]);
  apply assumption;
done;

lemma bigo_sumr4: "f =o g +o O(h) ==>
    (%x. sumr 0 x f) =o (%x. sumr 0 x g) +o O(%x. sumr 0 x (%y. abs(h y)))";
  apply (rule bigo_sumr_pos2);
  apply force;
  apply (subst bigo_abs3 [THEN sym]);
  apply assumption;
done;

lemma bigo_sumr5: "[| ALL x y. 0 <= h x y; 
    EX c. ALL x y. abs(f x y) <= c * (h x y) |] ==> 
      (%x. sumr 0 (j x) (f x)) : O(%x. sumr 0 (j x) (h x))";
  apply(auto simp add: bigo_def);
  apply(rule_tac x = c in exI);
  apply(auto);
  apply(subgoal_tac "ALL y. abs (sumr 0 y (f x)) <= 
    c * abs (sumr 0 y (h x))");
  apply(erule spec);
  apply(rule allI);
  apply(subgoal_tac "0 <= sumr 0 y (h x)");
  apply(auto simp add: sumr_mult sumr_le2 sumr_ge_zero);
  apply(subgoal_tac "abs (sumr 0 y (f x)) <= sumr 0 y (%n. abs (f x n))");
  apply(subgoal_tac "sumr 0 y (%n. abs (f x n)) <= sumr 0 y (%n. c * h x n)");
  apply(simp only: le_trans);
by(auto simp add: sumr_le2 sumr_rabs);

lemma bigo_sumr6: "[| ALL x y. 0 <= h x y; 
    EX c. ALL x y. abs((f x y) - (g x y)) <= (c * (h x y)) |] ==> 
      (%x. sumr 0 (j x) (f x)) : (%x. sumr 0 (j x) (g x)) 
        +o O(%x. sumr 0 (j x) (h x))";
  apply(simp only: set_minus_plus[THEN sym] func_diff_minus2 func_plus 
    sumr_minus[THEN sym] sumr_add func_minus);
  apply(simp only: diff_minus[THEN sym]);
  apply(rule bigo_sumr5);
  by(auto);

lemma bigo_sumr7: "f =o O(h) ==>
  (%x. sumr 0 (j x) (%y. (l x y) * (f (k x y)))) =o 
    O(%x. sumr 0 (j x) (%y. abs((l x y) * (h (k x y)))))";
  apply (rule bigo_sumr5);
  apply (simp del: abs_mult);
  apply (unfold bigo_def);
  apply auto;
  apply (rule_tac x = c in exI);
  apply (rule allI)+;
  apply (subst mult_left_commute);
  apply (rule mult_left_mono);
  apply (erule spec);
  apply force;
done;

lemma bigo_sumr8: "f =o g +o O(h) ==>
  (%x. sumr 0 (j x) (%y. (l x y) * (f (k x y)))) =o 
  (%x. sumr 0 (j x) (%y. (l x y) * (g (k x y)))) +o
    O(%x. sumr 0 (j x) (%y. abs((l x y) * (h (k x y)))))";
  apply (simp only: set_minus_plus [THEN sym] func_diff sumr_diff
    right_diff_distrib [THEN sym]);
  apply (erule bigo_sumr7);
done;

lemma bigo_sumr9: "f =o g +o O(%x. 1) ==>
    ALL x y. 0 <= l x y ==>
    (%x. sumr 0 (j x) (%y. (l x y) * f(k x y))): 
      (%x. sumr 0 (j x) (%y. (l x y) * g(k x y))) +o 
        O(%x. sumr 0 (j x) (l x))";
  by (drule bigo_sumr8 [of f g "%x. 1" j l k], simp);

lemma bigo_compose1: "f =o O(g) ==> (%x. f(k x)) =o O(%x. g(k x))";
  by (unfold bigo_def, auto);

lemma bigo_compose2: "f =o g +o O(h) ==> (%x. f(k x)) =o (%x. g(k x)) +o 
    O(%x. h(k x))";
  apply (simp only: set_minus_plus [THEN sym] diff_minus func_minus
    func_plus);
  apply (erule bigo_compose1);
  done;

subsection {* Setsum *}

lemma bigo_setsum1: "[| ALL x y. (0::'a::ordered_ring) <= h x y; 
    EX c. ALL x y. abs(f x y) <= c * (h x y) |] ==> 
      (%x. setsum (f x) (A x)) : O(%x. setsum (h x) (A x))";
  apply (auto simp add: bigo_def);
  apply (rule_tac x = "abs c" in exI);
  apply (auto);
  apply (case_tac "finite (A x)");
  apply (subst abs_nonneg);
  apply (rule setsum_nonneg);
  apply assumption;
  apply force;
  apply (subst setsum_const_times [THEN sym]);
  apply (rule order_trans);
  apply (rule abs_setsum);
  apply (rule setsum_le_cong);
  apply (rule order_trans);
  apply (subgoal_tac "abs (f x xa) <= c * h x xa");
  apply (assumption);
  apply force;
  apply (rule mult_right_mono);
  apply (rule abs_ge_self);
  apply force;
  apply (simp add: setsum_def);
done;

lemma bigo_setsum3: "f =o O(h) ==>
    (%x. setsum (%y. ((l x y)::'a::ordered_ring) * f(k x y)) (A x)) =o
    O(%x. setsum (%y. abs((l x y) * h(k x y))) (A x))"; 
  apply (rule bigo_setsum1);
  apply (rule allI)+;
  apply (rule abs_ge_zero);
  apply (unfold bigo_def);
  apply auto;
  apply (rule_tac x = c in exI);
  apply (rule allI)+;
  apply (subst mult_left_commute);
  apply (rule mult_left_mono);
  apply (erule spec);
  apply force;
done;

lemma setsum_subtractf': "(∑x:A. ((f x)::'a::ordered_ring) - g x) = 
    setsum f A - setsum g A";
  apply (case_tac "finite A");
  apply (erule setsum_subtractf);
  apply (simp add: setsum_def);
done;

lemma bigo_setsum4: "f =o g +o O(h) ==>
    (%x. setsum (%y. ((l x y)::'a::ordered_ring) * f(k x y)) (A x)) =o
    (%x. setsum (%y. ((l x y)::'a::ordered_ring) * g(k x y)) (A x)) +o
    O(%x. setsum (%y. abs((l x y) * h(k x y))) (A x))";
  apply (rule set_minus_imp_plus);
  apply (subst func_diff);
  apply (subgoal_tac "(%u. (∑y:A u. l u y * f (k u y)) -
      (∑y:A u. l u y * g (k u y))) = 
    (%u. ∑y:A u. l u y * (f - g)(k u y))");
  apply (erule ssubst);
  apply (rule bigo_setsum3);
  apply (erule set_plus_imp_minus);
  apply (rule ext);
  apply (subst setsum_subtractf' [THEN sym]);
  apply (case_tac "finite (A u)"); 
  apply (rule setsum_cong2);
  apply (subst func_diff);
  apply (simp add: ring_eq_simps);
  apply (simp add: setsum_def);
done;

lemma bigo_setsum5: "f =o O(h) ==> ALL x y. 0 <= l x y ==> 
    ALL x. 0 <= h x ==> 
      (%x. setsum (%y. ((l x y)::'a::ordered_ring) * f(k x y)) (A x)) =o
      O(%x. setsum (%y. (l x y) * h(k x y)) (A x))";
  apply (subgoal_tac "(%x. setsum (%y. (l x y) * h(k x y)) (A x))
    = (%x. setsum (%y. abs((l x y) * h(k x y))) (A x))");
  apply (erule ssubst);
  apply (erule bigo_setsum3);
  apply (rule ext);
  apply (rule setsum_cong2);
  apply (subst abs_nonneg);
  apply (rule nonneg_times_nonneg);
  apply auto;
done;

lemma bigo_setsum6: "f =o g +o O(h) ==> ALL x y. 0 <= l x y ==>
    ALL x. 0 <= h x ==>
      (%x. setsum (%y. ((l x y)::'a::ordered_ring) * f(k x y)) (A x)) =o
      (%x. setsum (%y. ((l x y)::'a::ordered_ring) * g(k x y)) (A x)) +o
      O(%x. setsum (%y. (l x y) * h(k x y)) (A x))";
  apply (rule set_minus_imp_plus);
  apply (subst func_diff);
  apply (subgoal_tac "(%u. (∑y:A u. l u y * f (k u y)) -
      (∑y:A u. l u y * g (k u y))) = 
    (%u. ∑y:A u. l u y * (f - g)(k u y))");
  apply (erule ssubst);
  apply (rule bigo_setsum5);
  apply (erule set_plus_imp_minus);
  apply assumption+;
  apply (rule ext);
  apply (subst setsum_subtractf' [THEN sym]);
  apply (case_tac "finite (A u)"); 
  apply (rule setsum_cong2);
  apply (subst func_diff);
  apply (simp add: ring_eq_simps);
  apply (simp add: setsum_def);
done;

subsection {* Misc useful stuff *}

lemma bigo_useful_intro: "A <= O(f) ==> B <= O(f) ==>
  A + B <= O(f::'a=>('b::ordered_ring))";
  apply (subst bigo_plus_idemp [THEN sym]);
  apply (rule set_plus_mono2);
  apply assumption+;
done;

lemma bigo_add_useful: "f =o O(k::'a=>'b::ordered_ring) ==>
    g =o O(k) ==> f + g =o O(k)";
  apply (subst bigo_plus_idemp [THEN sym]);
  apply (rule set_plus_intro);
  apply assumption+;
done;
  
lemma bigo_useful_const_mult: "(c::'a::ordered_field) ~= 0 ==> 
    (%x. c) * f =o O(h) ==> f =o O(h::'b=>'a)";
  apply (rule subsetD);
  apply (subgoal_tac "(%x. 1 / c) *o O(h) <= O(h)");
  apply assumption;
  apply (rule bigo_const_mult6);
  apply (subgoal_tac "f = (%x. 1 / c) * ((%x. c) * f)");
  apply (erule ssubst);
  apply (erule set_times_intro2);
  apply (simp add: func_times); 
  apply (rule ext);
  apply (subst times_divide_eq_left [THEN sym]);
  apply simp;
done;

lemma bigo_fix2: "(%x. f ((x::nat) + 1)) =o O(%x. h(x + 1)) ==> f 0 = 0 ==>
    f =o O(h)";
  apply (simp add: bigo_alt_def);
  apply auto;
  apply (rule_tac x = c in exI);
  apply auto;
  apply (case_tac "x = 0");
  apply simp;
  apply (rule nonneg_times_nonneg);
  apply force;
  apply force;
  apply (subgoal_tac "x = Suc (x - 1)");
  apply (erule ssubst);back;
  apply (erule spec);
  apply simp;
done;

lemma bigo_fix3: 
    "(%x. f ((x::nat) + 1)) =o (%x. g(x + 1)) +o O(%x. h(x + 1)) ==> 
       f 0 = g 0 ==> f =o g +o O(h)";
  apply (rule set_minus_imp_plus);
  apply (rule bigo_fix2);
  apply (subst func_diff);
  apply (subst func_diff [THEN sym]);back;
  apply (erule set_plus_imp_minus);
  apply (subst func_diff);
  apply simp;
done;

lemma bigo_LIMSEQ1: "f =o O(g) ==> g ----> 0 ==> f ----> 0";
  apply (simp add: LIMSEQ_def bigo_alt_def);
  apply clarify;
  apply (drule_tac x = "r / c" in spec);
  apply (drule mp);
  apply (erule pos_div_pos);
  apply assumption;
  apply clarify;
  apply (rule_tac x = no in exI);
  apply (rule allI);
  apply (drule_tac x = n in spec)+;
  apply (rule impI);
  apply (drule mp);
  apply assumption;
  apply (rule order_le_less_trans);
  apply assumption;
  apply (rule order_less_le_trans);
  apply (subgoal_tac "c * abs(g n) < c * (r / c)");
  apply assumption;
  apply (erule mult_strict_left_mono);
  apply assumption;
  apply simp;
done;

lemma bigo_LIMSEQ2: "f =o g +o O(h) ==> h ----> 0 ==> f ----> a 
    ==> g ----> a";
  apply (drule set_plus_imp_minus);
  apply (drule bigo_LIMSEQ1);
  apply assumption;
  apply (simp only: func_diff);
  apply (erule LIMSEQ_diff_approach_zero2);
  apply assumption;
done;

subsection {* Older stuff *}
-- {* Are these superceded by the above? *}

lemma bigo_plus_cong: "[| (f::'a=>'b::ordered_ring) = g + h; h : O(j::'a=>'b::ordered_ring) |] ==>
  f : g +o O(j)";
by(auto simp add: bigo_def bigo_pos_const elt_set_plus_def func_plus);

lemma bigo_plus_cong2: "[| (f::'a=>'b::ordered_ring) = g + h; h : i +o O(j::'a=>'b::ordered_ring) |] ==>
  f : (g + i) +o O(j)";
apply(auto simp add: bigo_def bigo_pos_const elt_set_plus_def func_plus);
apply(rule_tac x = b in exI);
apply(simp only: bigo_pos_const[THEN sym]);
by(auto simp add: add_assoc);

lemma bigo_minus_cong: "[| (f::'a=>'b::ordered_ring) = g - h; h : O(j::'a=>'b::ordered_ring) |] ==>
  f : g +o O(j)";
apply(auto simp add: bigo_def bigo_pos_const elt_set_plus_def func_plus);
apply(rule_tac x = "- h" in exI);
apply(simp only: bigo_pos_const[THEN sym]);
apply(auto simp add: func_minus func_plus func_diff_minus diff_minus);
done;

lemma bigo_minus_cong2: "[| (f::'a=>'b::ordered_ring) = g - h; h : i +o O(j::'a=>'b::ordered_ring) |] ==>
  f : (g - i) +o O(j)";
apply(simp);
apply(auto simp add: bigo_def bigo_pos_const elt_set_plus_def func_plus);
apply(rule_tac x = "- b" in exI);
apply(auto);
apply(simp only: bigo_pos_const[THEN sym]);
apply(rule_tac x = c in exI);
apply(rule allI);
apply(simp add: func_minus);
apply(simp add: func_minus diff_minus func_plus ext plus_ac0);
done;

subsection {* Less than or equal to *}

constdefs 
  lesso :: "('a => 'b::ordered_ring) => ('a => 'b) => ('a => 'b)"
      (infixl "<o" 70)
  "f <o g == (%x. max (f x - g x) 0)";

lemma bigo_lesseq1: "f =o O((h::'a=>'b::ordered_ring)) ==>
    ALL x. abs (g x) <= abs (f x) ==>
      g =o O(h)";
  apply (unfold bigo_def);
  apply clarsimp;
  apply (rule_tac x = c in exI);
  apply (rule allI);
  apply (rule order_trans);
  apply (erule spec)+;
done;

lemma bigo_lesseq2: "f =o O((h::'a=>'b::ordered_ring)) ==>
    ALL x. abs (g x) <= f x ==>
      g =o O(h)";
  apply (erule bigo_lesseq1);
  apply (rule allI);
  apply (drule_tac x = x in spec);
  apply (rule order_trans);
  apply assumption;
  apply (rule abs_ge_self);
done;

lemma bigo_lesseq3: "f =o O((h::'a=>'b::ordered_ring)) ==>
    ALL x. 0 <= g x ==> ALL x. g x <= f x ==>
      g =o O(h)";
  apply (erule bigo_lesseq2);
  apply (rule allI);
  apply (subst abs_nonneg);
  apply (erule spec)+;
done;

lemma bigo_lesseq4: "f =o O((h::'a=>'b::ordered_ring)) ==>
    ALL x. 0 <= g x ==> ALL x. g x <= abs (f x) ==>
      g =o O(h)";
  apply (erule bigo_lesseq1);
  apply (rule allI);
  apply (subst abs_nonneg);back;
  apply (erule spec)+;
done;

lemma bigo_lesso1: 
    "ALL x. f x <= g x ==> f <o g =o O(h::'a=>'b::ordered_ring)";
  apply (unfold lesso_def);
  apply (subgoal_tac "(%x. max (f x - g x) 0) = 0");
  apply (erule ssubst);
  apply (rule bigo_zero);
  apply (unfold func_zero);
  apply (rule ext);
  apply (simp split: split_max);
done;

lemma bigo_lesso2: "f =o g +o O((h::'a=>'b::ordered_ring)) ==>
    ALL x. 0 <= k x ==> ALL x. k x <= f x ==>
      k <o g =o O(h)";
  apply (unfold lesso_def);
  apply (rule bigo_lesseq4);
  apply (erule set_plus_imp_minus);
  apply (rule allI);
  apply (rule le_maxI2);
  apply (rule allI);
  apply (subst func_diff);
  apply (case_tac "0 <= k x - g x");
  apply simp;
  apply (subst abs_nonneg);
  apply (drule_tac x = x in spec);back;
  apply (simp add: compare_rls);
  apply (subst diff_minus)+;
  apply (rule add_right_mono);
  apply (erule spec);
  apply (rule order_trans); 
  prefer 2;
  apply (rule abs_ge_zero);
  apply (simp add: compare_rls);
done;

lemma bigo_lesso3: "f =o g +o O((h::'a=>'b::ordered_ring)) ==>
    ALL x. 0 <= k x ==> ALL x. g x <= k x ==>
      f <o k =o O(h)";
  apply (unfold lesso_def);
  apply (rule bigo_lesseq4);
  apply (erule set_plus_imp_minus);
  apply (rule allI);
  apply (rule le_maxI2);
  apply (rule allI);
  apply (subst func_diff);
  apply (case_tac "0 <= f x - k x");
  apply simp;
  apply (subst abs_nonneg);
  apply (drule_tac x = x in spec);back;
  apply (simp add: compare_rls);
  apply (subst diff_minus)+;
  apply (rule add_left_mono);
  apply (rule le_imp_neg_le);
  apply (erule spec);
  apply (rule order_trans); 
  prefer 2;
  apply (rule abs_ge_zero);
  apply (simp add: compare_rls);
done;

(* Generalize! *)

lemma bigo_lesso4: "f <o g =o O(k::'a=>real) ==>
    g =o h +o O(k) ==> f <o h =o O(k)";
  apply (unfold lesso_def);
  apply (drule set_plus_imp_minus);
  apply (subgoal_tac "(%x. abs ((g - h) x)) : O(k)");
  apply (frule bigo_add_useful);
  apply assumption;back;back;
  apply (erule bigo_lesseq2);back;back;back;
  apply (rule allI);
  apply (simp add: func_plus func_diff);
  apply (subst abs_nonneg);back;
  apply (rule le_maxI2);
  apply (auto simp add: func_plus func_diff compare_rls 
    split: split_max abs_split);
  apply (rule subsetD);
  prefer 2;
  apply (rule bigo_abs);
  apply (erule bigo_elt_subset);
done;

lemma bigo_lesso5: "f <o g =o O(h::'a=>'b::ordered_ring) ==>
    EX C. ALL x. f x <= g x + C * abs(h x)";
  apply (simp only: lesso_def bigo_alt_def);
  apply clarsimp;
  apply (rule_tac x = c in exI);
  apply (rule allI);
  apply (drule_tac x = x in spec);
  apply (subgoal_tac "abs(max (f x - g x) 0) = max (f x - g x) 0");
  apply (clarsimp simp add: compare_rls add_ac); 
  apply (rule abs_nonneg);
  apply (rule le_maxI2);
done;

lemma lesso_add: "f <o g =o O(h::'a=>real) ==>
      k <o l =o O(h) ==> (f + k) <o (g + l) =o O(h)";
  apply (unfold lesso_def);
  apply (rule bigo_lesseq3);
  apply (erule bigo_add_useful);
  apply assumption;
  apply (rule allI);
  apply (simp split: split_max);
  apply (rule allI);
  apply (simp split: split_max add: func_plus);
done;

end;

Preliminaries

lemma eq_in_trans:

  [| a = b; bC |] ==> aC

Definitions

Basic properties

lemma bigo_pos_const:

  (∃c. ∀x. ¦h x¦ ≤ c * ¦f x¦) = (∃c. (0::'a) < c ∧ (∀x. ¦h x¦ ≤ c * ¦f x¦))

lemma bigo_alt_def:

  O(f) = {h. ∃c. (0::'b) < c ∧ (∀x. ¦h x¦ ≤ c * ¦f x¦)}

lemma bigoset_alt_def:

  O(S) = {h. ∃fS. ∃c. (0::'b) < c ∧ (∀x. ¦h x¦ ≤ c * ¦f x¦)}

lemma bigoset_alt_def2:

  O(A) = {h. ∃fA. h ∈ O(f)}

lemma bigo_elt_subset:

  f ∈ O(g) ==> O(f) ⊆ O(g)

lemma bigoset_elt_subset:

  f ∈ O(A) ==> O(f) ⊆ O(A)

lemma bigoset_elt_subset2:

  fA ==> f ∈ O(A)

lemma bigoset_mono:

  AB ==> O(A) ⊆ O(B)

lemma bigo_refl:

  f ∈ O(f)

lemma bigoset_refl:

  A ⊆ O(A)

lemma bigo_bigo_refl:

  f ∈ O(O(f))

lemma bigo_bigo_eq:

  O(O(f)) = O(f)

lemma bigo_zero:

  0 ∈ O(g)

lemma bigo_zero2:

  O(%x. 0::'b) = {%x. 0::'b}

lemma bigo_plus_self_subset:

  O(f) + O(f) ⊆ O(f)

lemma bigo_plus_idemp:

  O(f) + O(f) = O(f)

lemma bigo_plus_subset_lemma:

  x ∈ O(f + g) ==> x ∈ O(f) + O(g)

lemma bigo_plus_subset:

  O(a + b) ⊆ O(a) + O(b)

lemma bigo_plus_subset2:

  O(f + A) ⊆ O(f) + O(A)

lemma bigo_plus_subset3:

  O(A + B) ⊆ O(A) + O(B)

lemma bigo_plus_subset4:

  [| A ⊆ O(f); B ⊆ O(f) |] ==> A + B ⊆ O(f)

lemma bigo_plus_subset5:

  O(f + O(g)) ⊆ O(f) + O(g)

lemma bigo_plus_subset6:

  [| ∀x. (0::'b) ≤ f x; ∀x. (0::'b) ≤ g x |] ==> O(f + g) = O(f) + O(g)

lemma bigo_elt_subset2:

  fg + O(h) ==> O(f) ⊆ O(g) + O(h)

lemma bigo_mult:

  O(f) * O(g) ⊆ O(f * g)

lemma bigo_mult2:

  f * O(g) ⊆ O(f * g)

lemma bigo_mult3:

  [| f ∈ O(h); g ∈ O(j) |] ==> f * g ∈ O(h * j)

lemma bigo_mult4:

  fk + O(h) ==> g * f ∈ (g * k) + O(g * h)

lemma bigo_minus:

  f ∈ O(g) ==> - f ∈ O(g)

lemma bigo_mult5:

x. f x ≠ (0::'b) ==> O(f * g) ⊆ f * O(g)

lemma bigo_mult6:

x. f x ≠ (0::'b) ==> O(f * g) = f * O(g)

lemma bigo_mult7:

x. f x ≠ (0::'b) ==> O(f * g) ⊆ O(f) * O(g)

lemma bigo_mult8:

x. f x ≠ (0::'b) ==> O(f * g) = O(f) * O(g)

lemma bigo_minus2:

  fg + O(h) ==> - f ∈ - g + O(h)

lemma bigo_minus3:

  O(- f) = O(f)

lemma bigo_plus_absorb_lemma1:

  f ∈ O(g) ==> f + O(g) ⊆ O(g)

lemma bigo_plus_absorb_lemma2:

  f ∈ O(g) ==> O(g) ⊆ f + O(g)

lemma bigo_plus_absorb:

  f ∈ O(g) ==> f + O(g) = O(g)

lemma bigo_plus_absorb2:

  [| f ∈ O(g); A ⊆ O(g) |] ==> f + A ⊆ O(g)

lemma bigo_add_commute_imp:

  fg + O(h) ==> gf + O(h)

lemma bigo_add_commute:

  (fg + O(h)) = (gf + O(h))

lemma bigo_bounded:

  [| ∀x. (0::'b) ≤ f x; ∀x. f xg x |] ==> f ∈ O(g)

lemma bigo_bounded_alt:

  [| ∀x. (0::'b) ≤ f x; ∀x. f xc * g x |] ==> f ∈ O(g)

lemma bigo_bounded2:

  [| ∀x. lb xf x; ∀x. f xlb x + g x |] ==> flb + O(g)

lemma bigo_bounded3:

x. lb xf xf xlb x + g x ==> flb + O(g)

lemma bigo_abs:

  (%x. ¦f x¦) ∈ O(f)

lemma bigo_abs2:

  f ∈ O(%x. ¦f x¦)

lemma bigo_abs3:

  O(f) = O(%x. ¦f x¦)

lemma bigo_abs4:

  fg + O(h) ==> (%x. ¦f x¦) ∈ (%x. ¦g x¦) + O(h)

lemma bigo_const1:

  (%x. c) ∈ O(%x. 1::'b)

lemma bigo_const2:

  O(%x. c) ⊆ O(%x. 1::'b)

lemma bigo_const3:

  c ≠ 0 ==> (%x. 1) ∈ O(%x. c)

lemma bigo_const4:

  c ≠ 0 ==> O(%x. 1) ⊆ O(%x. c)

lemma bigo_const:

  c ≠ 0 ==> O(%x. c) = O(%x. 1)

lemma bigo_const_mult1:

  (%x. c * f x) ∈ O(f)

lemma bigo_const_mult2:

  O(%x. c * f x) ⊆ O(f)

lemma bigo_const_mult3:

  c ≠ 0 ==> f ∈ O(%x. c * f x)

lemma bigo_const_mult4:

  c ≠ 0 ==> O(f) ⊆ O(%x. c * f x)

lemma bigo_const_mult:

  c ≠ 0 ==> O(%x. c * f x) = O(f)

lemma bigo_const_mult5:

  c ≠ 0 ==> (%x. c) * O(f) = O(f)

lemma bigo_const_mult6:

  (%x. c) * O(f) ⊆ O(f)

lemma bigo_const_mult7:

  f ∈ O(g) ==> (%x. c * f x) ∈ O(g)

lemma bigo_sumr_pos:

  [| ∀x. 0 ≤ h x; f ∈ O(h) |] ==> (%x. sumr 0 x f) ∈ O(%x. sumr 0 x h)

lemma bigo_sumr_pos2:

  [| ∀x. 0 ≤ h x; fg + O(h) |]
  ==> (%x. sumr 0 x f) ∈ (%x. sumr 0 x g) + O(%x. sumr 0 x h)

lemma bigo_sumr3:

  f ∈ O(h) ==> (%x. sumr 0 x f) ∈ O(%x. sumr 0 x (%y. ¦h y¦))

lemma bigo_sumr4:

  fg + O(h)
  ==> (%x. sumr 0 x f) ∈ (%x. sumr 0 x g) + O(%x. sumr 0 x (%y. ¦h y¦))

lemma bigo_sumr5:

  [| ∀x y. 0 ≤ h x y; ∃c. ∀x y. ¦f x y¦ ≤ c * h x y |]
  ==> (%x. sumr 0 (j x) (f x)) ∈ O(%x. sumr 0 (j x) (h x))

lemma bigo_sumr6:

  [| ∀x y. 0 ≤ h x y; ∃c. ∀x y. ¦f x y - g x y¦ ≤ c * h x y |]
  ==> (%x. sumr 0 (j x) (f x))
      ∈ (%x. sumr 0 (j x) (g x)) + O(%x. sumr 0 (j x) (h x))

lemma bigo_sumr7:

  f ∈ O(h)
  ==> (%x. sumr 0 (j x) (%y. l x y * f (k x y)))
      ∈ O(%x. sumr 0 (j x) (%y. ¦l x y * h (k x y)¦))

lemma bigo_sumr8:

  fg + O(h)
  ==> (%x. sumr 0 (j x) (%y. l x y * f (k x y)))
      ∈ (%x. sumr 0 (j x) (%y. l x y * g (k x y))) +
        O(%x. sumr 0 (j x) (%y. ¦l x y * h (k x y)¦))

lemma bigo_sumr9:

  [| fg + O(%x. 1); ∀x y. 0 ≤ l x y |]
  ==> (%x. sumr 0 (j x) (%y. l x y * f (k x y)))
      ∈ (%x. sumr 0 (j x) (%y. l x y * g (k x y))) + O(%x. sumr 0 (j x) (l x))

lemma bigo_compose1:

  f ∈ O(g) ==> (%x. f (k x)) ∈ O(%x. g (k x))

lemma bigo_compose2:

  fg + O(h) ==> (%x. f (k x)) ∈ (%x. g (k x)) + O(%x. h (k x))

Setsum

lemma bigo_setsum1:

  [| ∀x y. (0::'a) ≤ h x y; ∃c. ∀x y. ¦f x y¦ ≤ c * h x y |]
  ==> (%x. setsum (f x) (A x)) ∈ O(%x. setsum (h x) (A x))

lemma bigo_setsum3:

  f ∈ O(h)
  ==> (%x. ∑yA x. l x y * f (k x y)) ∈ O(%x. ∑yA x. ¦l x y * h (k x y)¦)

lemma setsum_subtractf':

  (∑xA. f x - g x) = setsum f A - setsum g A

lemma bigo_setsum4:

  fg + O(h)
  ==> (%x. ∑yA x. l x y * f (k x y))
      ∈ (%x. ∑yA x. l x y * g (k x y)) + O(%x. ∑yA x. ¦l x y * h (k x y)¦)

lemma bigo_setsum5:

  [| f ∈ O(h); ∀x y. (0::'a) ≤ l x y; ∀x. (0::'a) ≤ h x |]
  ==> (%x. ∑yA x. l x y * f (k x y)) ∈ O(%x. ∑yA x. l x y * h (k x y))

lemma bigo_setsum6:

  [| fg + O(h); ∀x y. (0::'a) ≤ l x y; ∀x. (0::'a) ≤ h x |]
  ==> (%x. ∑yA x. l x y * f (k x y))
      ∈ (%x. ∑yA x. l x y * g (k x y)) + O(%x. ∑yA x. l x y * h (k x y))

Misc useful stuff

lemma bigo_useful_intro:

  [| A ⊆ O(f); B ⊆ O(f) |] ==> A + B ⊆ O(f)

lemma bigo_add_useful:

  [| f ∈ O(k); g ∈ O(k) |] ==> f + g ∈ O(k)

lemma bigo_useful_const_mult:

  [| c ≠ (0::'a); (%x. c) * f ∈ O(h) |] ==> f ∈ O(h)

lemma bigo_fix2:

  [| (%x. f (x + 1)) ∈ O(%x. h (x + 1)); f 0 = (0::'a) |] ==> f ∈ O(h)

lemma bigo_fix3:

  [| (%x. f (x + 1)) ∈ (%x. g (x + 1)) + O(%x. h (x + 1)); f 0 = g 0 |]
  ==> fg + O(h)

lemma bigo_LIMSEQ1:

  [| f ∈ O(g); g ----> 0 |] ==> f ----> 0

lemma bigo_LIMSEQ2:

  [| fg + O(h); h ----> 0; f ----> a |] ==> g ----> a

Older stuff

lemma bigo_plus_cong:

  [| f = g + h; h ∈ O(j) |] ==> fg + O(j)

lemma bigo_plus_cong2:

  [| f = g + h; hi + O(j) |] ==> f ∈ (g + i) + O(j)

lemma bigo_minus_cong:

  [| f = g - h; h ∈ O(j) |] ==> fg + O(j)

lemma bigo_minus_cong2:

  [| f = g - h; hi + O(j) |] ==> f ∈ (g - i) + O(j)

Less than or equal to

lemma bigo_lesseq1:

  [| f ∈ O(h); ∀x. ¦g x¦ ≤ ¦f x¦ |] ==> g ∈ O(h)

lemma bigo_lesseq2:

  [| f ∈ O(h); ∀x. ¦g x¦ ≤ f x |] ==> g ∈ O(h)

lemma bigo_lesseq3:

  [| f ∈ O(h); ∀x. (0::'b) ≤ g x; ∀x. g xf x |] ==> g ∈ O(h)

lemma bigo_lesseq4:

  [| f ∈ O(h); ∀x. (0::'b) ≤ g x; ∀x. g x ≤ ¦f x¦ |] ==> g ∈ O(h)

lemma bigo_lesso1:

x. f xg x ==> f <o g ∈ O(h)

lemma bigo_lesso2:

  [| fg + O(h); ∀x. (0::'b) ≤ k x; ∀x. k xf x |] ==> k <o g ∈ O(h)

lemma bigo_lesso3:

  [| fg + O(h); ∀x. (0::'b) ≤ k x; ∀x. g xk x |] ==> f <o k ∈ O(h)

lemma bigo_lesso4:

  [| f <o g ∈ O(k); gh + O(k) |] ==> f <o h ∈ O(k)

lemma bigo_lesso5:

  f <o g ∈ O(h) ==> ∃C. ∀x. f xg x + C * ¦h x¦

lemma lesso_add:

  [| f <o g ∈ O(h); k <o l ∈ O(h) |] ==> (f + k) <o (g + l) ∈ O(h)