(* Title: HOL/MetisExamples/BigO.thy ID: $Id: BigO.thy,v 1.12 2007/11/06 07:47:31 haftmann Exp $ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Testing the metis method *) header {* Big O notation *} theory BigO imports SetsAndFunctions begin subsection {* Definitions *} constdefs bigo :: "('a => 'b::ordered_idom) => ('a => 'b) set" ("(1O'(_'))") "O(f::('a => 'b)) == {h. EX c. ALL x. abs (h x) <= c * abs (f x)}" ML{*ResAtp.problem_name := "BigO__bigo_pos_const"*} lemma bigo_pos_const: "(EX (c::'a::ordered_idom). ALL x. (abs (h x)) <= (c * (abs (f x)))) = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" apply auto apply (case_tac "c = 0", simp) apply (rule_tac x = "1" in exI, simp) apply (rule_tac x = "abs c" in exI, auto) apply (metis abs_ge_minus_self abs_ge_zero abs_minus_cancel abs_of_nonneg equation_minus_iff Orderings.xt1(6) abs_mult) done (*** Now various verions with an increasing modulus ***) ML{*ResReconstruct.modulus := 1*} lemma bigo_pos_const: "(EX (c::'a::ordered_idom). ALL x. (abs (h x)) <= (c * (abs (f x)))) = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" apply auto apply (case_tac "c = 0", simp) apply (rule_tac x = "1" in exI, simp) apply (rule_tac x = "abs c" in exI, auto) proof (neg_clausify) fix c x have 0: "!!(X1::'a::ordered_idom) X2::'a::ordered_idom. ¦X1 * X2¦ = ¦X2 * X1¦" by (metis abs_mult mult_commute) have 1: "!!(X1::'a::ordered_idom) X2::'a::ordered_idom. X1 ≤ (0::'a::ordered_idom) ∨ ¦X2¦ * X1 = ¦X2 * X1¦" by (metis abs_mult_pos linorder_linear) have 2: "!!(X1::'a::ordered_idom) X2::'a::ordered_idom. ¬ (0::'a::ordered_idom) < X1 * X2 ∨ ¬ (0::'a::ordered_idom) ≤ X2 ∨ ¬ X1 ≤ (0::'a::ordered_idom)" by (metis linorder_not_less mult_nonneg_nonpos2) assume 3: "!!x::'b::type. ¦(h::'b::type => 'a::ordered_idom) x¦ ≤ (c::'a::ordered_idom) * ¦(f::'b::type => 'a::ordered_idom) x¦" assume 4: "¬ ¦(h::'b::type => 'a::ordered_idom) (x::'b::type)¦ ≤ ¦c::'a::ordered_idom¦ * ¦(f::'b::type => 'a::ordered_idom) x¦" have 5: "¬ ¦(h::'b::type => 'a::ordered_idom) (x::'b::type)¦ ≤ ¦(c::'a::ordered_idom) * (f::'b::type => 'a::ordered_idom) x¦" by (metis 4 abs_mult) have 6: "!!(X1::'a::ordered_idom) X2::'a::ordered_idom. ¬ X1 ≤ (0::'a::ordered_idom) ∨ X1 ≤ ¦X2¦" by (metis abs_ge_zero xt1(6)) have 7: "!!(X1::'a::ordered_idom) X2::'a::ordered_idom. X1 ≤ ¦X2¦ ∨ (0::'a::ordered_idom) < X1" by (metis not_leE 6) have 8: "(0::'a::ordered_idom) < ¦(h::'b::type => 'a::ordered_idom) (x::'b::type)¦" by (metis 5 7) have 9: "!!X1::'a::ordered_idom. ¬ ¦(h::'b::type => 'a::ordered_idom) (x::'b::type)¦ ≤ X1 ∨ (0::'a::ordered_idom) < X1" by (metis 8 order_less_le_trans) have 10: "(0::'a::ordered_idom) < (c::'a::ordered_idom) * ¦(f::'b::type => 'a::ordered_idom) (x::'b::type)¦" by (metis 3 9) have 11: "¬ (c::'a::ordered_idom) ≤ (0::'a::ordered_idom)" by (metis abs_ge_zero 2 10) have 12: "!!X1::'a::ordered_idom. (c::'a::ordered_idom) * ¦X1¦ = ¦X1 * c¦" by (metis mult_commute 1 11) have 13: "!!X1::'b::type. - (h::'b::type => 'a::ordered_idom) X1 ≤ (c::'a::ordered_idom) * ¦(f::'b::type => 'a::ordered_idom) X1¦" by (metis 3 abs_le_D2) have 14: "!!X1::'b::type. - (h::'b::type => 'a::ordered_idom) X1 ≤ ¦(c::'a::ordered_idom) * (f::'b::type => 'a::ordered_idom) X1¦" by (metis 0 12 13) have 15: "!!(X1::'a::ordered_idom) X2::'a::ordered_idom. ¦X1 * ¦X2¦¦ = ¦X1 * X2¦" by (metis abs_mult abs_mult_pos abs_ge_zero) have 16: "!!(X1::'a::ordered_idom) X2::'a::ordered_idom. X1 ≤ ¦X2¦ ∨ ¬ X1 ≤ X2" by (metis xt1(6) abs_ge_self) have 17: "!!(X1::'a::ordered_idom) X2::'a::ordered_idom. ¬ ¦X1¦ ≤ X2 ∨ X1 ≤ ¦X2¦" by (metis 16 abs_le_D1) have 18: "!!X1::'b::type. (h::'b::type => 'a::ordered_idom) X1 ≤ ¦(c::'a::ordered_idom) * (f::'b::type => 'a::ordered_idom) X1¦" by (metis 17 3 15) show "False" by (metis abs_le_iff 5 18 14) qed lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). ALL x. (abs (h x)) <= (c * (abs (f x)))) = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" apply auto apply (case_tac "c = 0", simp) apply (rule_tac x = "1" in exI, simp) apply (rule_tac x = "abs c" in exI, auto); ML{*ResReconstruct.modulus:=2*} proof (neg_clausify) fix c x have 0: "!!(X1::'a::ordered_idom) X2::'a::ordered_idom. ¦X1 * X2¦ = ¦X2 * X1¦" by (metis abs_mult mult_commute) assume 1: "!!x::'b::type. ¦(h::'b::type => 'a::ordered_idom) x¦ ≤ (c::'a::ordered_idom) * ¦(f::'b::type => 'a::ordered_idom) x¦" assume 2: "¬ ¦(h::'b::type => 'a::ordered_idom) (x::'b::type)¦ ≤ ¦c::'a::ordered_idom¦ * ¦(f::'b::type => 'a::ordered_idom) x¦" have 3: "¬ ¦(h::'b::type => 'a::ordered_idom) (x::'b::type)¦ ≤ ¦(c::'a::ordered_idom) * (f::'b::type => 'a::ordered_idom) x¦" by (metis 2 abs_mult) have 4: "!!(X1::'a::ordered_idom) X2::'a::ordered_idom. ¬ X1 ≤ (0::'a::ordered_idom) ∨ X1 ≤ ¦X2¦" by (metis abs_ge_zero xt1(6)) have 5: "(0::'a::ordered_idom) < ¦(h::'b::type => 'a::ordered_idom) (x::'b::type)¦" by (metis not_leE 4 3) have 6: "(0::'a::ordered_idom) < (c::'a::ordered_idom) * ¦(f::'b::type => 'a::ordered_idom) (x::'b::type)¦" by (metis 1 order_less_le_trans 5) have 7: "!!X1::'a::ordered_idom. (c::'a::ordered_idom) * ¦X1¦ = ¦X1 * c¦" by (metis abs_ge_zero linorder_not_less mult_nonneg_nonpos2 6 linorder_linear abs_mult_pos mult_commute) have 8: "!!X1::'b::type. - (h::'b::type => 'a::ordered_idom) X1 ≤ ¦(c::'a::ordered_idom) * (f::'b::type => 'a::ordered_idom) X1¦" by (metis 0 7 abs_le_D2 1) have 9: "!!(X1::'a::ordered_idom) X2::'a::ordered_idom. ¬ ¦X1¦ ≤ X2 ∨ X1 ≤ ¦X2¦" by (metis abs_ge_self xt1(6) abs_le_D1) show "False" by (metis 8 abs_ge_zero abs_mult_pos abs_mult 1 9 3 abs_le_iff) qed lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). ALL x. (abs (h x)) <= (c * (abs (f x)))) = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" apply auto apply (case_tac "c = 0", simp) apply (rule_tac x = "1" in exI, simp) apply (rule_tac x = "abs c" in exI, auto); ML{*ResReconstruct.modulus:=3*} proof (neg_clausify) fix c x assume 0: "!!x::'b::type. ¦(h::'b::type => 'a::ordered_idom) x¦ ≤ (c::'a::ordered_idom) * ¦(f::'b::type => 'a::ordered_idom) x¦" assume 1: "¬ ¦(h::'b::type => 'a::ordered_idom) (x::'b::type)¦ ≤ ¦c::'a::ordered_idom¦ * ¦(f::'b::type => 'a::ordered_idom) x¦" have 2: "!!(X1::'a::ordered_idom) X2::'a::ordered_idom. X1 ≤ ¦X2¦ ∨ (0::'a::ordered_idom) < X1" by (metis abs_ge_zero xt1(6) not_leE) have 3: "¬ (c::'a::ordered_idom) ≤ (0::'a::ordered_idom)" by (metis abs_ge_zero mult_nonneg_nonpos2 linorder_not_less order_less_le_trans 1 abs_mult 2 0) have 4: "!!(X1::'a::ordered_idom) X2::'a::ordered_idom. ¦X1 * ¦X2¦¦ = ¦X1 * X2¦" by (metis abs_ge_zero abs_mult_pos abs_mult) have 5: "!!X1::'b::type. (h::'b::type => 'a::ordered_idom) X1 ≤ ¦(c::'a::ordered_idom) * (f::'b::type => 'a::ordered_idom) X1¦" by (metis 4 0 xt1(6) abs_ge_self abs_le_D1) show "False" by (metis abs_mult mult_commute 3 abs_mult_pos linorder_linear 0 abs_le_D2 5 1 abs_le_iff) qed ML{*ResReconstruct.modulus:=1*} lemma (*bigo_pos_const:*) "(EX (c::'a::ordered_idom). ALL x. (abs (h x)) <= (c * (abs (f x)))) = (EX c. 0 < c & (ALL x. (abs(h x)) <= (c * (abs (f x)))))" apply auto apply (case_tac "c = 0", simp) apply (rule_tac x = "1" in exI, simp) apply (rule_tac x = "abs c" in exI, auto); proof (neg_clausify) fix c x (*sort/type constraint inserted by hand!*) have 0: "!!(X1::'a::ordered_idom) X2. ¦X1 * ¦X2¦¦ = ¦X1 * X2¦" by (metis abs_ge_zero abs_mult_pos abs_mult) assume 1: "!!A. ¦h A¦ ≤ c * ¦f A¦" have 2: "!!X1 X2. ¬ ¦X1¦ ≤ X2 ∨ (0::'a) ≤ X2" by (metis abs_ge_zero order_trans) have 3: "!!X1. (0::'a) ≤ c * ¦f X1¦" by (metis 1 2) have 4: "!!X1. c * ¦f X1¦ = ¦c * f X1¦" by (metis 0 abs_of_nonneg 3) have 5: "!!X1. - h X1 ≤ c * ¦f X1¦" by (metis 1 abs_le_D2) have 6: "!!X1. - h X1 ≤ ¦c * f X1¦" by (metis 4 5) have 7: "!!X1. h X1 ≤ c * ¦f X1¦" by (metis 1 abs_le_D1) have 8: "!!X1. h X1 ≤ ¦c * f X1¦" by (metis 4 7) assume 9: "¬ ¦h x¦ ≤ ¦c¦ * ¦f x¦" have 10: "¬ ¦h x¦ ≤ ¦c * f x¦" by (metis abs_mult 9) show "False" by (metis 6 8 10 abs_leI) qed ML{*ResReconstruct.recon_sorts:=true*} lemma bigo_alt_def: "O(f) = {h. EX c. (0 < c & (ALL x. abs (h x) <= c * abs (f x)))}" by (auto simp add: bigo_def bigo_pos_const) ML{*ResAtp.problem_name := "BigO__bigo_elt_subset"*} lemma bigo_elt_subset [intro]: "f : O(g) ==> 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_pos) apply (assumption)+ (*sledgehammer*); 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) apply (rule order_less_imp_le, assumption); done ML{*ResAtp.problem_name := "BigO__bigo_refl"*} lemma bigo_refl [intro]: "f : O(f)" apply(auto simp add: bigo_def) proof (neg_clausify) fix x assume 0: "!!xa. ¬ ¦f (x xa)¦ ≤ xa * ¦f (x xa)¦" have 1: "!!X2. X2 ≤ (1::'b) * X2 ∨ ¬ (1::'b) ≤ (1::'b)" by (metis mult_le_cancel_right1 order_eq_iff) have 2: "!!X2. X2 ≤ (1::'b) * X2" by (metis order_eq_iff 1) show "False" by (metis 0 2) qed ML{*ResAtp.problem_name := "BigO__bigo_zero"*} lemma bigo_zero: "0 : O(g)" apply (auto simp add: bigo_def func_zero) proof (neg_clausify) fix x assume 0: "!!xa. ¬ (0::'b) ≤ xa * ¦g (x xa)¦" have 1: "¬ (0::'b) ≤ (0::'b)" by (metis 0 mult_eq_0_iff) show "False" by (metis 1 linorder_neq_iff linorder_antisym_conv1) qed 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) + O(f) <= O(f)" apply (auto simp add: bigo_alt_def set_plus) apply (rule_tac x = "c + ca" in exI) apply auto apply (simp add: ring_distribs func_plus) apply (blast intro:order_trans abs_triangle_ineq add_mono elim:) done lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)" apply (rule equalityI) apply (rule bigo_plus_self_subset) apply (rule set_zero_plus2) apply (rule bigo_zero) done lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) + O(g)" apply (rule subsetI) apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus) apply (subst bigo_pos_const [symmetric])+ 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_distribs) 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 (rule mult_nonneg_nonneg) apply (rule add_nonneg_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 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_distribs) 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 (metis abs_not_less_zero double_less_0_iff less_not_permute linorder_not_less mult_less_0_iff) apply (rule ext) apply (auto simp add: if_splits linorder_not_le) done lemma bigo_plus_subset2 [intro]: "A <= O(f) ==> B <= O(f) ==> A + B <= O(f)" apply (subgoal_tac "A + B <= O(f) + O(f)") apply (erule order_trans) apply simp apply (auto del: subsetI simp del: bigo_plus_idemp) done ML{*ResAtp.problem_name := "BigO__bigo_plus_eq"*} lemma bigo_plus_eq: "ALL x. 0 <= f x ==> ALL 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 (*sledgehammer*); apply (rule_tac x = "max c ca" in exI) apply (rule conjI) apply (metis Orderings.less_max_iff_disj) apply clarify apply (drule_tac x = "xa" in spec)+ apply (subgoal_tac "0 <= f xa + g xa") apply (simp add: ring_distribs) 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 (blast intro: order_trans) defer 1 apply (rule abs_triangle_ineq) apply (metis add_nonneg_nonneg) apply (rule add_mono) ML{*ResAtp.problem_name := "BigO__bigo_plus_eq_simpler"*} (*Found by SPASS; SLOW*) apply (metis le_maxI2 linorder_linear linorder_not_le min_max.less_eq_less_sup.sup_absorb1 mult_le_cancel_right xt1(6)) apply (metis le_maxI2 linorder_not_le mult_le_cancel_right xt1(6)) done ML{*ResAtp.problem_name := "BigO__bigo_bounded_alt"*} lemma bigo_bounded_alt: "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" apply (auto simp add: bigo_def) (*Version 1: one-shot proof*) apply (metis OrderedGroup.abs_le_D1 Orderings.linorder_class.not_less order_less_le Orderings.xt1(12) Ring_and_Field.abs_mult) done lemma bigo_bounded_alt: "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" apply (auto simp add: bigo_def) (*Version 2: single-step proof*) proof (neg_clausify) fix x assume 0: "!!x. f x ≤ c * g x" assume 1: "!!xa. ¬ f (x xa) ≤ xa * ¦g (x xa)¦" have 2: "!!X3. c * g X3 = f X3 ∨ ¬ c * g X3 ≤ f X3" by (metis 0 order_antisym_conv) have 3: "!!X3. ¬ f (x ¦X3¦) ≤ ¦X3 * g (x ¦X3¦)¦" by (metis 1 abs_mult) have 4: "!!X1 X3::'b::ordered_idom. X3 ≤ X1 ∨ X1 ≤ ¦X3¦" by (metis linorder_linear abs_le_D1) have 5: "!!X3::'b. ¦X3¦ * ¦X3¦ = X3 * X3" by (metis abs_mult_self AC_mult.f.commute) have 6: "!!X3. ¬ X3 * X3 < (0::'b::ordered_idom)" by (metis not_square_less_zero AC_mult.f.commute) have 7: "!!X1 X3::'b. ¦X1¦ * ¦X3¦ = ¦X3 * X1¦" by (metis abs_mult AC_mult.f.commute) have 8: "!!X3::'b. X3 * X3 = ¦X3 * X3¦" by (metis abs_mult 5) have 9: "!!X3. X3 * g (x ¦X3¦) ≤ f (x ¦X3¦)" by (metis 3 4) have 10: "c * g (x ¦c¦) = f (x ¦c¦)" by (metis 2 9) have 11: "!!X3::'b. ¦X3¦ * ¦¦X3¦¦ = ¦X3¦ * ¦X3¦" by (metis abs_idempotent abs_mult 8) have 12: "!!X3::'b. ¦X3 * ¦X3¦¦ = ¦X3¦ * ¦X3¦" by (metis AC_mult.f.commute 7 11) have 13: "!!X3::'b. ¦X3 * ¦X3¦¦ = X3 * X3" by (metis 8 7 12) have 14: "!!X3. X3 ≤ ¦X3¦ ∨ X3 < (0::'b)" by (metis abs_ge_self abs_le_D1 abs_if) have 15: "!!X3. X3 ≤ ¦X3¦ ∨ ¦X3¦ < (0::'b)" by (metis abs_ge_self abs_le_D1 abs_if) have 16: "!!X3. X3 * X3 < (0::'b) ∨ X3 * ¦X3¦ ≤ X3 * X3" by (metis 15 13) have 17: "!!X3::'b. X3 * ¦X3¦ ≤ X3 * X3" by (metis 16 6) have 18: "!!X3. X3 ≤ ¦X3¦ ∨ ¬ X3 < (0::'b)" by (metis mult_le_cancel_left 17) have 19: "!!X3::'b. X3 ≤ ¦X3¦" by (metis 18 14) have 20: "¬ f (x ¦c¦) ≤ ¦f (x ¦c¦)¦" by (metis 3 10) show "False" by (metis 20 19) qed text{*So here is the easier (and more natural) problem using transitivity*} ML{*ResAtp.problem_name := "BigO__bigo_bounded_alt_trans"*} lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" apply (auto simp add: bigo_def) (*Version 1: one-shot proof*) apply (metis Orderings.leD Orderings.leI abs_ge_self abs_le_D1 abs_mult abs_of_nonneg order_le_less xt1(12)); done text{*So here is the easier (and more natural) problem using transitivity*} ML{*ResAtp.problem_name := "BigO__bigo_bounded_alt_trans"*} lemma "ALL x. 0 <= f x ==> ALL x. f x <= c * g x ==> f : O(g)" apply (auto simp add: bigo_def) (*Version 2: single-step proof*) proof (neg_clausify) fix x assume 0: "!!A::'a::type. (f::'a::type => 'b::ordered_idom) A ≤ (c::'b::ordered_idom) * (g::'a::type => 'b::ordered_idom) A" assume 1: "!!A::'b::ordered_idom. ¬ (f::'a::type => 'b::ordered_idom) ((x::'b::ordered_idom => 'a::type) A) ≤ A * ¦(g::'a::type => 'b::ordered_idom) (x A)¦" have 2: "!!X2::'a::type. ¬ (c::'b::ordered_idom) * (g::'a::type => 'b::ordered_idom) X2 < (f::'a::type => 'b::ordered_idom) X2" by (metis 0 linorder_not_le) have 3: "!!X2::'b::ordered_idom. ¬ (f::'a::type => 'b::ordered_idom) ((x::'b::ordered_idom => 'a::type) ¦X2¦) ≤ ¦X2 * (g::'a::type => 'b::ordered_idom) (x ¦X2¦)¦" by (metis abs_mult 1) have 4: "!!X2::'b::ordered_idom. ¦X2 * (g::'a::type => 'b::ordered_idom) ((x::'b::ordered_idom => 'a::type) ¦X2¦)¦ < (f::'a::type => 'b::ordered_idom) (x ¦X2¦)" by (metis 3 linorder_not_less) have 5: "!!X2::'b::ordered_idom. X2 * (g::'a::type => 'b::ordered_idom) ((x::'b::ordered_idom => 'a::type) ¦X2¦) < (f::'a::type => 'b::ordered_idom) (x ¦X2¦)" by (metis abs_less_iff 4) show "False" by (metis 2 5) qed lemma bigo_bounded: "ALL x. 0 <= f x ==> ALL x. f x <= g x ==> f : O(g)" apply (erule bigo_bounded_alt [of f 1 g]) apply simp done ML{*ResAtp.problem_name := "BigO__bigo_bounded2"*} lemma bigo_bounded2: "ALL x. lb x <= f x ==> ALL 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) prefer 2 apply (drule_tac x = x in spec)+ apply arith (*not clear that it's provable otherwise*) proof (neg_clausify) fix x assume 0: "!!y. lb y ≤ f y" assume 1: "¬ (0::'b) ≤ f x + - lb x" have 2: "!!X3. (0::'b) + X3 = X3" by (metis diff_eq_eq right_minus_eq) have 3: "¬ (0::'b) ≤ f x - lb x" by (metis 1 compare_rls(1)) have 4: "¬ (0::'b) + lb x ≤ f x" by (metis 3 le_diff_eq) show "False" by (metis 4 2 0) qed ML{*ResAtp.problem_name := "BigO__bigo_abs"*} lemma bigo_abs: "(%x. abs(f x)) =o O(f)" apply (unfold bigo_def) apply auto proof (neg_clausify) fix x assume 0: "!!xa. ¬ ¦f (x xa)¦ ≤ xa * ¦f (x xa)¦" have 1: "!!X2. X2 ≤ (1::'b) * X2 ∨ ¬ (1::'b) ≤ (1::'b)" by (metis mult_le_cancel_right1 order_eq_iff) have 2: "!!X2. X2 ≤ (1::'b) * X2" by (metis order_eq_iff 1) show "False" by (metis 0 2) qed ML{*ResAtp.problem_name := "BigO__bigo_abs2"*} lemma bigo_abs2: "f =o O(%x. abs(f x))" apply (unfold bigo_def) apply auto proof (neg_clausify) fix x assume 0: "!!xa. ¬ ¦f (x xa)¦ ≤ xa * ¦f (x xa)¦" have 1: "!!X2. X2 ≤ (1::'b) * X2 ∨ ¬ (1::'b) ≤ (1::'b)" by (metis mult_le_cancel_right1 order_eq_iff) have 2: "!!X2. X2 ≤ (1::'b) * X2" by (metis order_eq_iff 1) show "False" by (metis 0 2) qed 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) ==> (%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)" using a by (rule bigo_elt_subset) finally show "(%x. abs (f x) - abs (g x)) : O(h)". qed lemma bigo_abs5: "f =o O(g) ==> (%x. abs(f x)) =o O(g)" by (unfold bigo_def, auto) lemma bigo_elt_subset2 [intro]: "f : g +o O(h) ==> O(f) <= O(g) + O(h)" proof - assume "f : g +o O(h)" also have "... <= O(g) + O(h)" by (auto del: subsetI) also have "... = O(%x. abs(g x)) + O(%x. abs(h x))" apply (subst bigo_abs3 [symmetric])+ apply (rule refl) done also have "... = O((%x. abs(g x)) + (%x. abs(h x)))" by (rule bigo_plus_eq [symmetric], auto) finally have "f : ...". then have "O(f) <= ..." by (elim bigo_elt_subset) also have "... = O(%x. abs(g x)) + O(%x. abs(h x))" by (rule bigo_plus_eq, auto) finally show ?thesis by (simp add: bigo_abs3 [symmetric]) qed ML{*ResAtp.problem_name := "BigO__bigo_mult"*} lemma bigo_mult [intro]: "O(f)*O(g) <= O(f * g)" apply (rule subsetI) apply (subst bigo_def) apply (auto simp del: abs_mult mult_ac simp add: bigo_alt_def set_times func_times) (*sledgehammer*); apply (rule_tac x = "c * ca" in exI) apply(rule allI) apply(erule_tac x = x in allE)+ apply(subgoal_tac "c * ca * abs(f x * g x) = (c * abs(f x)) * (ca * abs(g x))") ML{*ResAtp.problem_name := "BigO__bigo_mult_simpler"*} prefer 2 apply (metis Finite_Set.AC_mult.f.assoc Finite_Set.AC_mult.f.left_commute OrderedGroup.abs_of_pos OrderedGroup.mult_left_commute Ring_and_Field.abs_mult Ring_and_Field.mult_pos_pos) apply(erule ssubst) apply (subst abs_mult) (*not qute BigO__bigo_mult_simpler_1 (a hard problem!) as abs_mult has just been done*) proof (neg_clausify) fix a c b ca x assume 0: "(0::'b::ordered_idom) < (c::'b::ordered_idom)" assume 1: "¦(a::'a => 'b::ordered_idom) (x::'a)¦ ≤ (c::'b::ordered_idom) * ¦(f::'a => 'b::ordered_idom) x¦" assume 2: "¦(b::'a => 'b::ordered_idom) (x::'a)¦ ≤ (ca::'b::ordered_idom) * ¦(g::'a => 'b::ordered_idom) x¦" assume 3: "¬ ¦(a::'a => 'b::ordered_idom) (x::'a)¦ * ¦(b::'a => 'b::ordered_idom) x¦ ≤ (c::'b::ordered_idom) * ¦(f::'a => 'b::ordered_idom) x¦ * ((ca::'b::ordered_idom) * ¦(g::'a => 'b::ordered_idom) x¦)" have 4: "¦c::'b::ordered_idom¦ = c" by (metis OrderedGroup.abs_of_pos 0) have 5: "!!X1::'b::ordered_idom. (c::'b::ordered_idom) * ¦X1¦ = ¦c * X1¦" by (metis Ring_and_Field.abs_mult 4) have 6: "(0::'b::ordered_idom) = (1::'b::ordered_idom) ∨ (0::'b::ordered_idom) < (1::'b::ordered_idom)" by (metis OrderedGroup.abs_not_less_zero Ring_and_Field.abs_one Ring_and_Field.linorder_neqE_ordered_idom) have 7: "(0::'b::ordered_idom) < (1::'b::ordered_idom)" by (metis 6 Ring_and_Field.one_neq_zero) have 8: "¦1::'b::ordered_idom¦ = (1::'b::ordered_idom)" by (metis OrderedGroup.abs_of_pos 7) have 9: "!!X1::'b::ordered_idom. (0::'b::ordered_idom) ≤ (c::'b::ordered_idom) * ¦X1¦" by (metis OrderedGroup.abs_ge_zero 5) have 10: "!!X1::'b::ordered_idom. X1 * (1::'b::ordered_idom) = X1" by (metis Ring_and_Field.mult_cancel_right2 Finite_Set.AC_mult.f.commute) have 11: "!!X1::'b::ordered_idom. ¦¦X1¦¦ = ¦X1¦ * ¦1::'b::ordered_idom¦" by (metis Ring_and_Field.abs_mult OrderedGroup.abs_idempotent 10) have 12: "!!X1::'b::ordered_idom. ¦¦X1¦¦ = ¦X1¦" by (metis 11 8 10) have 13: "!!X1::'b::ordered_idom. (0::'b::ordered_idom) ≤ ¦X1¦" by (metis OrderedGroup.abs_ge_zero 12) have 14: "¬ (0::'b::ordered_idom) ≤ (c::'b::ordered_idom) * ¦(f::'a => 'b::ordered_idom) (x::'a)¦ ∨ ¬ (0::'b::ordered_idom) ≤ ¦(b::'a => 'b::ordered_idom) x¦ ∨ ¬ ¦b x¦ ≤ (ca::'b::ordered_idom) * ¦(g::'a => 'b::ordered_idom) x¦ ∨ ¬ ¦(a::'a => 'b::ordered_idom) x¦ ≤ c * ¦f x¦" by (metis 3 Ring_and_Field.mult_mono) have 15: "¬ (0::'b::ordered_idom) ≤ ¦(b::'a => 'b::ordered_idom) (x::'a)¦ ∨ ¬ ¦b x¦ ≤ (ca::'b::ordered_idom) * ¦(g::'a => 'b::ordered_idom) x¦ ∨ ¬ ¦(a::'a => 'b::ordered_idom) x¦ ≤ (c::'b::ordered_idom) * ¦(f::'a => 'b::ordered_idom) x¦" by (metis 14 9) have 16: "¬ ¦(b::'a => 'b::ordered_idom) (x::'a)¦ ≤ (ca::'b::ordered_idom) * ¦(g::'a => 'b::ordered_idom) x¦ ∨ ¬ ¦(a::'a => 'b::ordered_idom) x¦ ≤ (c::'b::ordered_idom) * ¦(f::'a => 'b::ordered_idom) x¦" by (metis 15 13) have 17: "¬ ¦(a::'a => 'b::ordered_idom) (x::'a)¦ ≤ (c::'b::ordered_idom) * ¦(f::'a => 'b::ordered_idom) x¦" by (metis 16 2) show 18: "False" by (metis 17 1) qed ML{*ResAtp.problem_name := "BigO__bigo_mult2"*} lemma bigo_mult2 [intro]: "f *o O(g) <= O(f * g)" apply (auto simp add: bigo_def elt_set_times_def func_times abs_mult) (*sledgehammer*); apply (rule_tac x = c in exI) apply clarify apply (drule_tac x = x in spec) ML{*ResAtp.problem_name := "BigO__bigo_mult2_simpler"*} (*sledgehammer [no luck]*); apply (subgoal_tac "abs(f x) * abs(b x) <= abs(f x) * (c * abs(g x))") apply (simp add: mult_ac) apply (rule mult_left_mono, assumption) apply (rule abs_ge_zero) done ML{*ResAtp.problem_name:="BigO__bigo_mult3"*} lemma bigo_mult3: "f : O(h) ==> g : O(j) ==> f * g : O(h * j)" by (metis bigo_mult set_times_intro subset_iff) ML{*ResAtp.problem_name:="BigO__bigo_mult4"*} lemma bigo_mult4 [intro]:"f : k +o O(h) ==> g * f : (g * k) +o O(g * h)" by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib) 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 ML{*ResAtp.problem_name := "BigO__bigo_mult6"*} lemma bigo_mult6: "ALL x. f x ~= 0 ==> O(f * g) = (f::'a => ('b::ordered_field)) *o O(g)" by (metis bigo_mult2 bigo_mult5 order_antisym) (*proof requires relaxing relevance: 2007-01-25*) ML{*ResAtp.problem_name := "BigO__bigo_mult7"*} declare bigo_mult6 [simp] lemma bigo_mult7: "ALL x. f x ~= 0 ==> O(f * g) <= O(f::'a => ('b::ordered_field)) * O(g)" (*sledgehammer*) apply (subst bigo_mult6) apply assumption apply (rule set_times_mono3) apply (rule bigo_refl) done declare bigo_mult6 [simp del] ML{*ResAtp.problem_name := "BigO__bigo_mult8"*} declare bigo_mult7[intro!] lemma bigo_mult8: "ALL x. f x ~= 0 ==> O(f * g) = O(f::'a => ('b::ordered_field)) * O(g)" by (metis bigo_mult bigo_mult7 order_antisym_conv) lemma bigo_minus [intro]: "f : O(g) ==> - f : O(g)" by (auto simp add: bigo_def func_minus) lemma bigo_minus2: "f : g +o O(h) ==> -f : -g +o O(h)" apply (rule set_minus_imp_plus) apply (drule set_plus_imp_minus) apply (drule bigo_minus) apply (simp add: diff_minus) done lemma bigo_minus3: "O(-f) = O(f)" by (auto simp add: bigo_def func_minus abs_minus_cancel) lemma bigo_plus_absorb_lemma1: "f : O(g) ==> f +o O(g) <= O(g)" proof - assume a: "f : O(g)" show "f +o O(g) <= O(g)" proof - have "f : O(f)" by auto then have "f +o O(g) <= O(f) + O(g)" by (auto del: subsetI) also have "... <= O(g) + O(g)" proof - from a have "O(f) <= O(g)" by (auto del: subsetI) thus ?thesis by (auto del: subsetI) qed also have "... <= O(g)" by (simp add: bigo_plus_idemp) finally show ?thesis . qed qed lemma bigo_plus_absorb_lemma2: "f : O(g) ==> O(g) <= f +o O(g)" proof - assume a: "f : O(g)" show "O(g) <= f +o O(g)" proof - from a have "-f : O(g)" by auto then have "-f +o O(g) <= O(g)" by (elim bigo_plus_absorb_lemma1) then have "f +o (-f +o O(g)) <= f +o O(g)" by auto also have "f +o (-f +o O(g)) = O(g)" by (simp add: set_plus_rearranges) finally show ?thesis . qed qed ML{*ResAtp.problem_name:="BigO__bigo_plus_absorb"*} lemma bigo_plus_absorb [simp]: "f : O(g) ==> f +o O(g) = O(g)" by (metis bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 order_eq_iff); lemma bigo_plus_absorb2 [intro]: "f : O(g) ==> A <= O(g) ==> f +o A <= O(g)" apply (subgoal_tac "f +o A <= f +o O(g)") apply force+ done lemma bigo_add_commute_imp: "f : g +o O(h) ==> g : f +o O(h)" apply (subst set_minus_plus [symmetric]) apply (subgoal_tac "g - f = - (f - g)") apply (erule ssubst) apply (rule bigo_minus) apply (subst set_minus_plus) apply assumption apply (simp add: diff_minus add_ac) done lemma bigo_add_commute: "(f : g +o O(h)) = (g : f +o O(h))" apply (rule iffI) apply (erule bigo_add_commute_imp)+ done lemma bigo_const1: "(%x. c) : O(%x. 1)" by (auto simp add: bigo_def mult_ac) ML{*ResAtp.problem_name:="BigO__bigo_const2"*} lemma (*bigo_const2 [intro]:*) "O(%x. c) <= O(%x. 1)" by (metis bigo_const1 bigo_elt_subset); lemma bigo_const2 [intro]: "O(%x. c::'b::ordered_idom) <= O(%x. 1)"; (*??FAILS because the two occurrences of COMBK have different polymorphic types proof (neg_clausify) assume 0: "¬ O(COMBK (c::'b::ordered_idom)) ⊆ O(COMBK (1::'b::ordered_idom))" have 1: "COMBK (c::'b::ordered_idom) ∉ O(COMBK (1::'b::ordered_idom))" apply (rule notI) apply (rule 0 [THEN notE]) apply (rule bigo_elt_subset) apply assumption; sorry by (metis 0 bigo_elt_subset) loops?? show "False" by (metis 1 bigo_const1) qed *) apply (rule bigo_elt_subset) apply (rule bigo_const1) done ML{*ResAtp.problem_name := "BigO__bigo_const3"*} lemma bigo_const3: "(c::'a::ordered_field) ~= 0 ==> (%x. 1) : O(%x. c)" apply (simp add: bigo_def) proof (neg_clausify) assume 0: "(c::'a::ordered_field) ≠ (0::'a::ordered_field)" assume 1: "!!A::'a::ordered_field. ¬ (1::'a::ordered_field) ≤ A * ¦c::'a::ordered_field¦" have 2: "(0::'a::ordered_field) = ¦c::'a::ordered_field¦ ∨ ¬ (1::'a::ordered_field) ≤ (1::'a::ordered_field)" by (metis 1 field_inverse) have 3: "¦c::'a::ordered_field¦ = (0::'a::ordered_field)" by (metis linorder_neq_iff linorder_antisym_conv1 2) have 4: "(0::'a::ordered_field) = (c::'a::ordered_field)" by (metis 3 abs_eq_0) show "False" by (metis 0 4) qed lemma bigo_const4: "(c::'a::ordered_field) ~= 0 ==> O(%x. 1) <= O(%x. c)" by (rule bigo_elt_subset, rule bigo_const3, assumption) lemma bigo_const [simp]: "(c::'a::ordered_field) ~= 0 ==> O(%x. c) = O(%x. 1)" by (rule equalityI, rule bigo_const2, rule bigo_const4, assumption) ML{*ResAtp.problem_name := "BigO__bigo_const_mult1"*} lemma bigo_const_mult1: "(%x. c * f x) : O(f)" apply (simp add: bigo_def abs_mult) proof (neg_clausify) fix x assume 0: "!!xa::'b::ordered_idom. ¬ ¦c::'b::ordered_idom¦ * ¦(f::'a::type => 'b::ordered_idom) ((x::'b::ordered_idom => 'a::type) xa)¦ ≤ xa * ¦f (x xa)¦" show "False" by (metis linorder_neq_iff linorder_antisym_conv1 0) qed lemma bigo_const_mult2: "O(%x. c * f x) <= O(f)" by (rule bigo_elt_subset, rule bigo_const_mult1) ML{*ResAtp.problem_name := "BigO__bigo_const_mult3"*} lemma bigo_const_mult3: "(c::'a::ordered_field) ~= 0 ==> f : O(%x. c * f x)" apply (simp add: bigo_def) (*sledgehammer [no luck]*); apply (rule_tac x = "abs(inverse c)" in exI) apply (simp only: abs_mult [symmetric] mult_assoc [symmetric]) apply (subst left_inverse) apply (auto ); done lemma bigo_const_mult4: "(c::'a::ordered_field) ~= 0 ==> O(f) <= O(%x. c * f x)" by (rule bigo_elt_subset, rule bigo_const_mult3, assumption) lemma bigo_const_mult [simp]: "(c::'a::ordered_field) ~= 0 ==> O(%x. c * f x) = O(f)" by (rule equalityI, rule bigo_const_mult2, erule bigo_const_mult4) ML{*ResAtp.problem_name := "BigO__bigo_const_mult5"*} lemma bigo_const_mult5 [simp]: "(c::'a::ordered_field) ~= 0 ==> (%x. c) *o O(f) = O(f)" apply (auto del: subsetI) apply (rule order_trans) apply (rule bigo_mult2) apply (simp add: func_times) apply (auto intro!: subsetI simp add: bigo_def elt_set_times_def func_times) apply (rule_tac x = "%y. inverse c * x y" in exI) apply (rename_tac g d) apply safe apply (rule_tac [2] ext) prefer 2 apply (metis AC_mult.f_e.left_ident mult_assoc right_inverse) apply (simp add: mult_assoc [symmetric] abs_mult) (*couldn't get this proof without the step above; SLOW*) apply (metis AC_mult.f.assoc abs_ge_zero mult_left_mono) done ML{*ResAtp.problem_name := "BigO__bigo_const_mult6"*} lemma bigo_const_mult6 [intro]: "(%x. c) *o O(f) <= O(f)" apply (auto intro!: subsetI simp add: bigo_def elt_set_times_def func_times simp del: abs_mult mult_ac) (*sledgehammer*); 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 (subst abs_mult) 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) ==> (%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) <= O(g)" by (auto del: subsetI) finally show ?thesis . qed 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 [symmetric] diff_minus func_minus func_plus) apply (erule bigo_compose1) done subsection {* Setsum *} lemma bigo_setsum_main: "ALL x. ALL y : A x. 0 <= h x y ==> EX c. ALL x. ALL y : A x. abs(f x y) <= c * (h x y) ==> (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)" apply (auto simp add: bigo_def) apply (rule_tac x = "abs c" in exI) apply (subst abs_of_nonneg) back back apply (rule setsum_nonneg) apply force apply (subst setsum_right_distrib) apply (rule allI) apply (rule order_trans) apply (rule setsum_abs) apply (rule setsum_mono) apply (blast intro: order_trans mult_right_mono abs_ge_self) done ML{*ResAtp.problem_name := "BigO__bigo_setsum1"*} lemma bigo_setsum1: "ALL x y. 0 <= h x y ==> EX c. ALL x y. abs(f x y) <= c * (h x y) ==> (%x. SUM y : A x. f x y) =o O(%x. SUM y : A x. h x y)" apply (rule bigo_setsum_main) (*sledgehammer*); apply force apply clarsimp apply (rule_tac x = c in exI) apply force done lemma bigo_setsum2: "ALL y. 0 <= h y ==> EX c. ALL y. abs(f y) <= c * (h y) ==> (%x. SUM y : A x. f y) =o O(%x. SUM y : A x. h y)" by (rule bigo_setsum1, auto) ML{*ResAtp.problem_name := "BigO__bigo_setsum3"*} lemma bigo_setsum3: "f =o O(h) ==> (%x. SUM y : A x. (l x y) * f(k x y)) =o O(%x. SUM y : A x. abs(l x y * h(k x y)))" apply (rule bigo_setsum1) apply (rule allI)+ apply (rule abs_ge_zero) apply (unfold bigo_def) apply (auto simp add: abs_mult); (*sledgehammer*); apply (rule_tac x = c in exI) apply (rule allI)+ apply (subst mult_left_commute) apply (rule mult_left_mono) apply (erule spec) apply (rule abs_ge_zero) done lemma bigo_setsum4: "f =o g +o O(h) ==> (%x. SUM y : A x. l x y * f(k x y)) =o (%x. SUM y : A x. l x y * g(k x y)) +o O(%x. SUM y : A x. abs(l x y * h(k x y)))" apply (rule set_minus_imp_plus) apply (subst func_diff) apply (subst setsum_subtractf [symmetric]) apply (subst right_diff_distrib [symmetric]) apply (rule bigo_setsum3) apply (subst func_diff [symmetric]) apply (erule set_plus_imp_minus) done ML{*ResAtp.problem_name := "BigO__bigo_setsum5"*} lemma bigo_setsum5: "f =o O(h) ==> ALL x y. 0 <= l x y ==> ALL x. 0 <= h x ==> (%x. SUM y : A x. (l x y) * f(k x y)) =o O(%x. SUM y : A x. (l x y) * h(k x y))" apply (subgoal_tac "(%x. SUM y : A x. (l x y) * h(k x y)) = (%x. SUM y : A x. abs((l x y) * h(k x y)))") apply (erule ssubst) apply (erule bigo_setsum3) apply (rule ext) apply (rule setsum_cong2) apply (thin_tac "f ∈ O(h)") apply (metis abs_of_nonneg zero_le_mult_iff) done lemma bigo_setsum6: "f =o g +o O(h) ==> ALL x y. 0 <= l x y ==> ALL x. 0 <= h x ==> (%x. SUM y : A x. (l x y) * f(k x y)) =o (%x. SUM y : A x. (l x y) * g(k x y)) +o O(%x. SUM y : A x. (l x y) * h(k x y))" apply (rule set_minus_imp_plus) apply (subst func_diff) apply (subst setsum_subtractf [symmetric]) apply (subst right_diff_distrib [symmetric]) apply (rule bigo_setsum5) apply (subst func_diff [symmetric]) apply (drule set_plus_imp_minus) apply auto done subsection {* Misc useful stuff *} lemma bigo_useful_intro: "A <= O(f) ==> B <= O(f) ==> A + B <= O(f)" apply (subst bigo_plus_idemp [symmetric]) apply (rule set_plus_mono2) apply assumption+ done lemma bigo_useful_add: "f =o O(h) ==> g =o O(h) ==> f + g =o O(h)" apply (subst bigo_plus_idemp [symmetric]) 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)" 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) done ML{*ResAtp.problem_name := "BigO__bigo_fix"*} lemma bigo_fix: "(%x. f ((x::nat) + 1)) =o O(%x. h(x + 1)) ==> f 0 = 0 ==> f =o O(h)" apply (simp add: bigo_alt_def) (*sledgehammer*); apply clarify apply (rule_tac x = c in exI) apply safe apply (case_tac "x = 0") apply (metis OrderedGroup.abs_ge_zero OrderedGroup.abs_zero order_less_le Ring_and_Field.split_mult_pos_le) apply (subgoal_tac "x = Suc (x - 1)") apply metis apply simp done lemma bigo_fix2: "(%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_fix) apply (subst func_diff) apply (subst func_diff [symmetric]) apply (rule set_plus_imp_minus) apply simp apply (simp add: func_diff) done subsection {* Less than or equal to *} constdefs lesso :: "('a => 'b::ordered_idom) => ('a => 'b) => ('a => 'b)" (infixl "<o" 70) "f <o g == (%x. max (f x - g x) 0)" lemma bigo_lesseq1: "f =o O(h) ==> ALL x. abs (g x) <= abs (f x) ==> g =o O(h)" apply (unfold bigo_def) apply clarsimp apply (blast intro: order_trans) done lemma bigo_lesseq2: "f =o O(h) ==> ALL x. abs (g x) <= f x ==> g =o O(h)" apply (erule bigo_lesseq1) apply (blast intro: abs_ge_self order_trans) done lemma bigo_lesseq3: "f =o O(h) ==> 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_of_nonneg) apply (erule spec)+ done lemma bigo_lesseq4: "f =o O(h) ==> 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_of_nonneg) apply (erule spec)+ done ML{*ResAtp.problem_name:="BigO__bigo_lesso1"*} lemma bigo_lesso1: "ALL x. f x <= g x ==> f <o g =o O(h)" apply (unfold lesso_def) apply (subgoal_tac "(%x. max (f x - g x) 0) = 0") (*??Translation of TSTP raised an exception: Type unification failed: Variable ?'X2.0::type not of sort ord*) apply (metis bigo_zero) apply (unfold func_zero) apply (rule ext) apply (simp split: split_max) done ML{*ResAtp.problem_name := "BigO__bigo_lesso2"*} lemma bigo_lesso2: "f =o g +o O(h) ==> 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 (erule thin_rl) (*sledgehammer*); apply (case_tac "0 <= k x - g x") prefer 2 (*re-order subgoals because I don't know what to put after a structured proof*) apply (metis abs_ge_zero abs_minus_commute linorder_linear min_max.less_eq_less_sup.sup_absorb1 min_max.less_eq_less_sup.sup_commute) proof (neg_clausify) fix x assume 0: "!!A. k A ≤ f A" have 1: "!!(X1::'b::ordered_idom) X2. ¬ max X1 X2 < X1" by (metis linorder_not_less le_maxI1) (*sort inserted by hand*) assume 2: "(0::'b) ≤ k x - g x" have 3: "¬ k x - g x < (0::'b)" by (metis 2 linorder_not_less) have 4: "!!X1 X2. min X1 (k X2) ≤ f X2" by (metis min_max.less_eq_less_inf.inf_le2 min_max.less_eq_less_inf.le_inf_iff min_max.less_eq_less_inf.le_iff_inf 0) have 5: "¦g x - f x¦ = f x - g x" by (metis abs_minus_commute combine_common_factor mult_zero_right minus_add_cancel minus_zero abs_if diff_less_eq min_max.less_eq_less_inf.inf_commute 4 linorder_not_le min_max.less_eq_less_inf.le_iff_inf 3 diff_less_0_iff_less linorder_not_less) have 6: "max (0::'b) (k x - g x) = k x - g x" by (metis min_max.less_eq_less_sup.le_iff_sup 2) assume 7: "¬ max (k x - g x) (0::'b) ≤ ¦f x - g x¦" have 8: "¬ k x - g x ≤ f x - g x" by (metis 5 abs_minus_commute 7 min_max.less_eq_less_sup.sup_commute 6) show "False" by (metis min_max.less_eq_less_sup.sup_commute min_max.less_eq_less_inf.inf_commute min_max.less_eq_less_inf_sup.sup_inf_absorb min_max.less_eq_less_inf.le_iff_inf 0 max_diff_distrib_left 1 linorder_not_le 8) qed ML{*ResAtp.problem_name := "BigO__bigo_lesso3"*} lemma bigo_lesso3: "f =o g +o O(h) ==> 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 (erule thin_rl) (*sledgehammer*); apply (case_tac "0 <= f x - k x") apply (simp del: compare_rls diff_minus); apply (subst abs_of_nonneg) apply (drule_tac x = x in spec) back ML{*ResAtp.problem_name := "BigO__bigo_lesso3_simpler"*} apply (metis diff_less_0_iff_less linorder_not_le not_leE uminus_add_conv_diff xt1(12) xt1(6)) apply (metis add_minus_cancel diff_le_eq le_diff_eq uminus_add_conv_diff) apply (metis abs_ge_zero linorder_linear min_max.less_eq_less_sup.sup_absorb1 min_max.less_eq_less_sup.sup_commute) done lemma bigo_lesso4: "f <o g =o O(k::'a=>'b::ordered_field) ==> g =o h +o O(k) ==> f <o h =o O(k)" apply (unfold lesso_def) apply (drule set_plus_imp_minus) apply (drule bigo_abs5) back apply (simp add: func_diff) apply (drule bigo_useful_add) apply assumption apply (erule bigo_lesseq2) back apply (rule allI) apply (auto simp add: func_plus func_diff compare_rls split: split_max abs_split) done ML{*ResAtp.problem_name := "BigO__bigo_lesso5"*} lemma bigo_lesso5: "f <o g =o O(h) ==> EX C. ALL x. f x <= g x + C * abs(h x)" apply (simp only: lesso_def bigo_alt_def) apply clarsimp apply (metis abs_if abs_mult add_commute diff_le_eq less_not_permute) done end
lemma bigo_pos_const:
(∃c. ∀x. ¦h x¦ ≤ c * ¦f x¦) = (∃c>0::'a. ∀x. ¦h x¦ ≤ c * ¦f x¦)
lemma bigo_pos_const:
(∃c. ∀x. ¦h x¦ ≤ c * ¦f x¦) = (∃c>0::'a. ∀x. ¦h x¦ ≤ c * ¦f x¦)
lemma
(∃c. ∀x. ¦h x¦ ≤ c * ¦f x¦) = (∃c>0::'a. ∀x. ¦h x¦ ≤ c * ¦f x¦)
lemma
(∃c. ∀x. ¦h x¦ ≤ c * ¦f x¦) = (∃c>0::'a. ∀x. ¦h x¦ ≤ c * ¦f x¦)
lemma
(∃c. ∀x. ¦h x¦ ≤ c * ¦f x¦) = (∃c>0::'a. ∀x. ¦h x¦ ≤ c * ¦f x¦)
lemma bigo_alt_def:
O(f) = {h. ∃c>0::'b. ∀x. ¦h x¦ ≤ c * ¦f x¦}
lemma bigo_elt_subset:
f ∈ O(g) ==> O(f) ⊆ O(g)
lemma bigo_refl:
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:
O(f + g) ⊆ O(f) + O(g)
lemma bigo_plus_subset2:
[| A ⊆ O(f); B ⊆ O(f) |] ==> A + B ⊆ O(f)
lemma bigo_plus_eq:
[| ∀x. (0::'b) ≤ f x; ∀x. (0::'b) ≤ g x |] ==> O(f + g) = O(f) + O(g)
lemma bigo_bounded_alt:
[| ∀x. (0::'b) ≤ f x; ∀x. f x ≤ c * g x |] ==> f ∈ O(g)
lemma bigo_bounded_alt:
[| ∀x. (0::'b) ≤ f x; ∀x. f x ≤ c * g x |] ==> f ∈ O(g)
lemma
[| ∀x. (0::'b) ≤ f x; ∀x. f x ≤ c * g x |] ==> f ∈ O(g)
lemma
[| ∀x. (0::'b) ≤ f x; ∀x. f x ≤ c * g x |] ==> f ∈ O(g)
lemma bigo_bounded:
[| ∀x. (0::'b) ≤ f x; ∀x. f x ≤ g x |] ==> f ∈ O(g)
lemma bigo_bounded2:
[| ∀x. lb x ≤ f x; ∀x. f x ≤ lb x + g x |] ==> f ∈ lb +o 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:
f ∈ g +o O(h) ==> (λx. ¦f x¦) ∈ (λx. ¦g x¦) +o O(h)
lemma bigo_abs5:
f ∈ O(g) ==> (λx. ¦f x¦) ∈ O(g)
lemma bigo_elt_subset2:
f ∈ g +o O(h) ==> O(f) ⊆ O(g) + O(h)
lemma bigo_mult:
O(f) * O(g) ⊆ O(f * g)
lemma bigo_mult2:
f *o O(g) ⊆ O(f * g)
lemma bigo_mult3:
[| f ∈ O(h); g ∈ O(j) |] ==> f * g ∈ O(h * j)
lemma bigo_mult4:
f ∈ k +o O(h) ==> g * f ∈ g * k +o O(g * h)
lemma bigo_mult5:
∀x. f x ≠ (0::'b) ==> O(f * g) ⊆ f *o O(g)
lemma bigo_mult6:
∀x. f x ≠ (0::'b) ==> O(f * g) = f *o 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_minus:
f ∈ O(g) ==> - f ∈ O(g)
lemma bigo_minus2:
f ∈ g +o O(h) ==> - f ∈ - g +o O(h)
lemma bigo_minus3:
O(- f) = O(f)
lemma bigo_plus_absorb_lemma1:
f ∈ O(g) ==> f +o O(g) ⊆ O(g)
lemma bigo_plus_absorb_lemma2:
f ∈ O(g) ==> O(g) ⊆ f +o O(g)
lemma bigo_plus_absorb:
f ∈ O(g) ==> f +o O(g) = O(g)
lemma bigo_plus_absorb2:
[| f ∈ O(g); A ⊆ O(g) |] ==> f +o A ⊆ O(g)
lemma bigo_add_commute_imp:
f ∈ g +o O(h) ==> g ∈ f +o O(h)
lemma bigo_add_commute:
(f ∈ g +o O(h)) = (g ∈ f +o O(h))
lemma bigo_const1:
(λx. c) ∈ O(λx. 1::'b)
lemma
O(λx. c) ⊆ O(λx. 1::'b)
lemma bigo_const2:
O(λx. c) ⊆ O(λx. 1::'b)
lemma bigo_const3:
c ≠ (0::'a) ==> (λx. 1::'a) ∈ O(λx. c)
lemma bigo_const4:
c ≠ (0::'a) ==> O(λx. 1::'a) ⊆ O(λx. c)
lemma bigo_const:
c ≠ (0::'a) ==> O(λx. c) = O(λx. 1::'a)
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::'a) ==> f ∈ O(λx. c * f x)
lemma bigo_const_mult4:
c ≠ (0::'a) ==> O(f) ⊆ O(λx. c * f x)
lemma bigo_const_mult:
c ≠ (0::'a) ==> O(λx. c * f x) = O(f)
lemma bigo_const_mult5:
c ≠ (0::'a) ==> (λx. c) *o O(f) = O(f)
lemma bigo_const_mult6:
(λx. c) *o O(f) ⊆ O(f)
lemma bigo_const_mult7:
f ∈ O(g) ==> (λx. c * f x) ∈ O(g)
lemma bigo_compose1:
f ∈ O(g) ==> (λx. f (k x)) ∈ O(λx. g (k x))
lemma bigo_compose2:
f ∈ g +o O(h) ==> (λx. f (k x)) ∈ (λx. g (k x)) +o O(λx. h (k x))
lemma bigo_setsum_main:
[| ∀x. ∀y∈A x. (0::'c) ≤ h x y; ∃c. ∀x. ∀y∈A x. ¦f x y¦ ≤ c * h x y |]
==> (λx. setsum (f x) (A x)) ∈ O(λx. setsum (h x) (A x))
lemma bigo_setsum1:
[| ∀x y. (0::'c) ≤ 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_setsum2:
[| ∀y. (0::'b) ≤ h y; ∃c. ∀y. ¦f y¦ ≤ c * h y |]
==> (λx. setsum f (A x)) ∈ O(λx. setsum h (A x))
lemma bigo_setsum3:
f ∈ O(h)
==> (λx. ∑y∈A x. l x y * f (k x y)) ∈ O(λx. ∑y∈A x. ¦l x y * h (k x y)¦)
lemma bigo_setsum4:
f ∈ g +o O(h)
==> (λx. ∑y∈A x. l x y * f (k x y))
∈ (λx. ∑y∈A x. l x y * g (k x y)) +o O(λx. ∑y∈A x. ¦l x y * h (k x y)¦)
lemma bigo_setsum5:
[| f ∈ O(h); ∀x y. (0::'b) ≤ l x y; ∀x. (0::'b) ≤ h x |]
==> (λx. ∑y∈A x. l x y * f (k x y)) ∈ O(λx. ∑y∈A x. l x y * h (k x y))
lemma bigo_setsum6:
[| f ∈ g +o O(h); ∀x y. (0::'b) ≤ l x y; ∀x. (0::'b) ≤ h x |]
==> (λx. ∑y∈A x. l x y * f (k x y))
∈ (λx. ∑y∈A x. l x y * g (k x y)) +o O(λx. ∑y∈A x. l x y * h (k x y))
lemma bigo_useful_intro:
[| A ⊆ O(f); B ⊆ O(f) |] ==> A + B ⊆ O(f)
lemma bigo_useful_add:
[| f ∈ O(h); g ∈ O(h) |] ==> f + g ∈ O(h)
lemma bigo_useful_const_mult:
[| c ≠ (0::'a); (λx. c) * f ∈ O(h) |] ==> f ∈ O(h)
lemma bigo_fix:
[| (λx. f (x + 1)) ∈ O(λx. h (x + 1)); f 0 = (0::'a) |] ==> f ∈ O(h)
lemma bigo_fix2:
[| (λx. f (x + 1)) ∈ (λx. g (x + 1)) +o O(λx. h (x + 1)); f 0 = g 0 |]
==> f ∈ g +o O(h)
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 x ≤ f 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 x ≤ g x ==> f <o g ∈ O(h)
lemma bigo_lesso2:
[| f ∈ g +o O(h); ∀x. (0::'b) ≤ k x; ∀x. k x ≤ f x |] ==> k <o g ∈ O(h)
lemma bigo_lesso3:
[| f ∈ g +o O(h); ∀x. (0::'b) ≤ k x; ∀x. g x ≤ k x |] ==> f <o k ∈ O(h)
lemma bigo_lesso4:
[| f <o g ∈ O(k); g ∈ h +o O(k) |] ==> f <o h ∈ O(k)
lemma bigo_lesso5:
f <o g ∈ O(h) ==> ∃C. ∀x. f x ≤ g x + C * ¦h x¦