Theory Function_Growth
section ‹Comparing growth of functions on natural numbers by a preorder relation›
theory Function_Growth
imports
Main
"HOL-Library.Preorder"
"HOL-Library.Discrete"
begin
subsection ‹Motivation›
text ‹
When comparing growth of functions in computer science, it is common to adhere
on Landau Symbols (``O-Notation''). However these come at the cost of notational
oddities, particularly writing ‹f = O(g)› for ‹f ∈ O(g)› etc.
Here we suggest a different way, following Hardy (G.~H.~Hardy and J.~E.~Littlewood,
Some problems of Diophantine approximation, Acta Mathematica 37 (1914), p.~225).
We establish a quasi order relation ‹≲› on functions such that
‹f ≲ g ⟷ f ∈ O(g)›. From a didactic point of view, this does not only
avoid the notational oddities mentioned above but also emphasizes the key insight
of a growth hierarchy of functions:
‹(λn. 0) ≲ (λn. k) ≲ Discrete.log ≲ Discrete.sqrt ≲ id ≲ …›.
›
subsection ‹Model›
text ‹
Our growth functions are of type ‹ℕ ⇒ ℕ›. This is different
to the usual conventions for Landau symbols for which ‹ℝ ⇒ ℝ›
would be appropriate, but we argue that ‹ℝ ⇒ ℝ› is more
appropriate for analysis, whereas our setting is discrete.
Note that we also restrict the additional coefficients to ‹ℕ›, something
we discuss at the particular definitions.
›
subsection ‹The ‹≲› relation›
definition less_eq_fun :: "(nat ⇒ nat) ⇒ (nat ⇒ nat) ⇒ bool" (infix "≲" 50)
where
"f ≲ g ⟷ (∃c>0. ∃n. ∀m>n. f m ≤ c * g m)"
text ‹
This yields ‹f ≲ g ⟷ f ∈ O(g)›. Note that ‹c› is restricted to
‹ℕ›. This does not pose any problems since if ‹f ∈ O(g)› holds for
a ‹c ∈ ℝ›, it also holds for ‹⌈c⌉ ∈ ℕ› by transitivity.
›
lemma less_eq_funI [intro?]:
assumes "∃c>0. ∃n. ∀m>n. f m ≤ c * g m"
shows "f ≲ g"
unfolding less_eq_fun_def by (rule assms)
lemma not_less_eq_funI:
assumes "⋀c n. c > 0 ⟹ ∃m>n. c * g m < f m"
shows "¬ f ≲ g"
using assms unfolding less_eq_fun_def linorder_not_le [symmetric] by blast
lemma less_eq_funE [elim?]:
assumes "f ≲ g"
obtains n c where "c > 0" and "⋀m. m > n ⟹ f m ≤ c * g m"
using assms unfolding less_eq_fun_def by blast
lemma not_less_eq_funE:
assumes "¬ f ≲ g" and "c > 0"
obtains m where "m > n" and "c * g m < f m"
using assms unfolding less_eq_fun_def linorder_not_le [symmetric] by blast
subsection ‹The ‹≈› relation, the equivalence relation induced by ‹≲››
definition equiv_fun :: "(nat ⇒ nat) ⇒ (nat ⇒ nat) ⇒ bool" (infix "≅" 50)
where
"f ≅ g ⟷
(∃c⇩1>0. ∃c⇩2>0. ∃n. ∀m>n. f m ≤ c⇩1 * g m ∧ g m ≤ c⇩2 * f m)"
text ‹
This yields ‹f ≅ g ⟷ f ∈ Θ(g)›. Concerning ‹c⇩1› and ‹c⇩2›
restricted to \<^typ>‹nat›, see note above on ‹(≲)›.
›
lemma equiv_funI:
assumes "∃c⇩1>0. ∃c⇩2>0. ∃n. ∀m>n. f m ≤ c⇩1 * g m ∧ g m ≤ c⇩2 * f m"
shows "f ≅ g"
unfolding equiv_fun_def by (rule assms)
lemma not_equiv_funI:
assumes "⋀c⇩1 c⇩2 n. c⇩1 > 0 ⟹ c⇩2 > 0 ⟹
∃m>n. c⇩1 * f m < g m ∨ c⇩2 * g m < f m"
shows "¬ f ≅ g"
using assms unfolding equiv_fun_def linorder_not_le [symmetric] by blast
lemma equiv_funE:
assumes "f ≅ g"
obtains n c⇩1 c⇩2 where "c⇩1 > 0" and "c⇩2 > 0"
and "⋀m. m > n ⟹ f m ≤ c⇩1 * g m ∧ g m ≤ c⇩2 * f m"
using assms unfolding equiv_fun_def by blast
lemma not_equiv_funE:
fixes n c⇩1 c⇩2
assumes "¬ f ≅ g" and "c⇩1 > 0" and "c⇩2 > 0"
obtains m where "m > n"
and "c⇩1 * f m < g m ∨ c⇩2 * g m < f m"
using assms unfolding equiv_fun_def linorder_not_le [symmetric] by blast
subsection ‹The ‹≺› relation, the strict part of ‹≲››
definition less_fun :: "(nat ⇒ nat) ⇒ (nat ⇒ nat) ⇒ bool" (infix "≺" 50)
where
"f ≺ g ⟷ f ≲ g ∧ ¬ g ≲ f"
lemma less_funI:
assumes "∃c>0. ∃n. ∀m>n. f m ≤ c * g m"
and "⋀c n. c > 0 ⟹ ∃m>n. c * f m < g m"
shows "f ≺ g"
using assms unfolding less_fun_def less_eq_fun_def linorder_not_less [symmetric] by blast
lemma not_less_funI:
assumes "⋀c n. c > 0 ⟹ ∃m>n. c * g m < f m"
and "∃c>0. ∃n. ∀m>n. g m ≤ c * f m"
shows "¬ f ≺ g"
using assms unfolding less_fun_def less_eq_fun_def linorder_not_less [symmetric] by blast
lemma less_funE [elim?]:
assumes "f ≺ g"
obtains n c where "c > 0" and "⋀m. m > n ⟹ f m ≤ c * g m"
and "⋀c n. c > 0 ⟹ ∃m>n. c * f m < g m"
proof -
from assms have "f ≲ g" and "¬ g ≲ f" by (simp_all add: less_fun_def)
from ‹f ≲ g› obtain n c where *:"c > 0" "m > n ⟹ f m ≤ c * g m" for m
by (rule less_eq_funE) blast
{ fix c n :: nat
assume "c > 0"
with ‹¬ g ≲ f› obtain m where "m > n" "c * f m < g m"
by (rule not_less_eq_funE) blast
then have **: "∃m>n. c * f m < g m" by blast
} note ** = this
from * ** show thesis by (rule that)
qed
lemma not_less_funE:
assumes "¬ f ≺ g" and "c > 0"
obtains m where "m > n" and "c * g m < f m"
| d q where "⋀m. d > 0 ⟹ m > q ⟹ g q ≤ d * f q"
using assms unfolding less_fun_def linorder_not_less [symmetric] by blast
text ‹
I did not find a proof for ‹f ≺ g ⟷ f ∈ o(g)›. Maybe this only
holds if ‹f› and/or ‹g› are of a certain class of functions.
However ‹f ∈ o(g) ⟶ f ≺ g› is provable, and this yields a
handy introduction rule.
Note that D. Knuth ignores ‹o› altogether. So what \dots
Something still has to be said about the coefficient ‹c› in
the definition of ‹(≺)›. In the typical definition of ‹o›,
it occurs on the \emph{right} hand side of the ‹(>)›. The reason
is that the situation is dual to the definition of ‹O›: the definition
works since ‹c› may become arbitrary small. Since this is not possible
within \<^term>‹ℕ›, we push the coefficient to the left hand side instead such
that it may become arbitrary big instead.
›
lemma less_fun_strongI:
assumes "⋀c. c > 0 ⟹ ∃n. ∀m>n. c * f m < g m"
shows "f ≺ g"
proof (rule less_funI)
have "1 > (0::nat)" by simp
with assms [OF this] obtain n where *: "m > n ⟹ 1 * f m < g m" for m
by blast
have "∀m>n. f m ≤ 1 * g m"
proof (rule allI, rule impI)
fix m
assume "m > n"
with * have "1 * f m < g m" by simp
then show "f m ≤ 1 * g m" by simp
qed
with ‹1 > 0› show "∃c>0. ∃n. ∀m>n. f m ≤ c * g m" by blast
fix c n :: nat
assume "c > 0"
with assms obtain q where "m > q ⟹ c * f m < g m" for m by blast
then have "c * f (Suc (q + n)) < g (Suc (q + n))" by simp
moreover have "Suc (q + n) > n" by simp
ultimately show "∃m>n. c * f m < g m" by blast
qed
subsection ‹‹≲› is a preorder›
text ‹This yields all lemmas relating ‹≲›, ‹≺› and ‹≅›.›