Added Real_Asymp package
authorManuel Eberl <eberlm@in.tum.de>
Sun, 15 Jul 2018 14:46:57 +0200
changeset 68630 c55f6f0b3854
parent 68629 f36858fdf768
child 68637 ec8c7c9459e0
Added Real_Asymp package
CONTRIBUTORS
NEWS
src/Doc/ROOT
src/Doc/Real_Asymp/Real_Asymp_Doc.thy
src/Doc/Real_Asymp/document/root.tex
src/Doc/Real_Asymp/document/style.sty
src/HOL/ROOT
src/HOL/Real_Asymp/Eventuallize.thy
src/HOL/Real_Asymp/Inst_Existentials.thy
src/HOL/Real_Asymp/Lazy_Eval.thy
src/HOL/Real_Asymp/Multiseries_Expansion.thy
src/HOL/Real_Asymp/Multiseries_Expansion_Bounds.thy
src/HOL/Real_Asymp/Real_Asymp.thy
src/HOL/Real_Asymp/Real_Asymp_Approx.thy
src/HOL/Real_Asymp/Real_Asymp_Examples.thy
src/HOL/Real_Asymp/asymptotic_basis.ML
src/HOL/Real_Asymp/exp_log_expression.ML
src/HOL/Real_Asymp/expansion_interface.ML
src/HOL/Real_Asymp/inst_existentials.ML
src/HOL/Real_Asymp/lazy_eval.ML
src/HOL/Real_Asymp/multiseries_expansion.ML
src/HOL/Real_Asymp/multiseries_expansion_bounds.ML
src/HOL/Real_Asymp/real_asymp.ML
src/HOL/Real_Asymp/real_asymp_diag.ML
--- a/CONTRIBUTORS	Sun Jul 15 01:14:04 2018 +0100
+++ b/CONTRIBUTORS	Sun Jul 15 14:46:57 2018 +0200
@@ -6,6 +6,10 @@
 Contributions to Isabelle2018
 -----------------------------
 
+* July 2018: Manuel Eberl
+  "real_asymp" proof method for automatic proofs of real limits, "Big-O"
+  statements, etc.
+
 * June 2018: Fabian Immler
   More tool support for HOL-Types_To_Sets.
 
--- a/NEWS	Sun Jul 15 01:14:04 2018 +0100
+++ b/NEWS	Sun Jul 15 14:46:57 2018 +0200
@@ -231,6 +231,9 @@
 
 *** HOL ***
 
+* New proof method "real_asymp" to prove asymptotics or real-valued 
+  functions (limits, "Big-O", etc.) automatically.
+
 * Sledgehammer: bundled version of "vampire" (for non-commercial users)
 helps to avoid fragility of "remote_vampire" service.
 
--- a/src/Doc/ROOT	Sun Jul 15 01:14:04 2018 +0100
+++ b/src/Doc/ROOT	Sun Jul 15 14:46:57 2018 +0200
@@ -359,6 +359,20 @@
     "root.tex"
     "svmono.cls"
 
+session Real_Asymp (doc) in "Real_Asymp" = "HOL-Real_Asymp" +
+  theories
+    Real_Asymp_Doc
+  document_files (in "..")
+    "prepare_document"
+    "pdfsetup.sty"
+    "iman.sty"
+    "extra.sty"
+    "isar.sty"
+    "manual.bib"
+  document_files
+    "root.tex"
+    "style.sty"
+
 session Sledgehammer (doc) in "Sledgehammer" = Pure +
   options [document_variants = "sledgehammer"]
   document_files (in "..")
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Real_Asymp/Real_Asymp_Doc.thy	Sun Jul 15 14:46:57 2018 +0200
@@ -0,0 +1,286 @@
+(*<*)
+theory Real_Asymp_Doc
+  imports "HOL-Real_Asymp.Real_Asymp"
+begin
+
+ML_file \<open>../antiquote_setup.ML\<close>
+(*>*)
+
+section \<open>Introduction\<close>
+
+text \<open>
+  This document describes the \<^verbatim>\<open>real_asymp\<close> package that provides a number of tools
+  related to the asymptotics of real-valued functions. These tools are:
+
+    \<^item> The @{method real_asymp} method to prove limits, statements involving Landau symbols
+      (`Big-O' etc.), and asymptotic equivalence (\<open>\<sim>\<close>)
+
+    \<^item> The @{command real_limit} command to compute the limit of a real-valued function at a
+      certain point
+
+    \<^item> The @{command real_expansion} command to compute the asymptotic expansion of a
+      real-valued function at a certain point
+
+  These three tools will be described in the following sections.
+\<close>
+
+section \<open>Supported Operations\<close>
+
+text \<open>
+  The \<^verbatim>\<open>real_asymp\<close> package fully supports the following operations:
+
+    \<^item> Basic arithmetic (addition, subtraction, multiplication, division)
+
+    \<^item> Powers with constant natural exponent
+
+    \<^item> @{term exp}, @{term ln}, @{term log}, @{term sqrt}, @{term "root k"} (for a constant k)
+  
+    \<^item> @{term sin}, @{term cos}, @{term tan} at finite points
+
+    \<^item> @{term sinh}, @{term cosh}, @{term tanh}
+
+    \<^item> @{term min}, @{term max}, @{term abs}
+
+  Additionally, the following operations are supported in a `best effort' fashion using asymptotic
+  upper/lower bounds:
+
+    \<^item> Powers with variable natural exponent
+
+    \<^item> @{term sin} and @{term cos} at \<open>\<plusminus>\<infinity>\<close>
+
+    \<^item> @{term floor}, @{term ceiling}, @{term frac}, and \<open>mod\<close>
+
+  Additionally, the @{term arctan} function is partially supported. The method may fail when
+  the argument to @{term arctan} contains functions of different order of growth.
+\<close>
+
+
+section \<open>Proving Limits and Asymptotic Properties\<close>
+
+text \<open>
+  \[
+    @{method_def (HOL) real_asymp} : \<open>method\<close>
+  \]
+
+  @{rail \<open>
+    @@{method (HOL) real_asymp} opt? (@{syntax simpmod} * )
+    ;
+    opt: '(' ('verbose' | 'fallback' ) ')'
+    ;
+    @{syntax_def simpmod}: ('add' | 'del' | 'only' | 'split' (() | '!' | 'del') |
+      'cong' (() | 'add' | 'del')) ':' @{syntax thms}
+  \<close>}
+\<close>
+
+text \<open>
+  The @{method real_asymp} method is a semi-automatic proof method for proving certain statements
+  related to the asymptotic behaviour of real-valued functions. In the following, let \<open>f\<close> and \<open>g\<close>
+  be functions of type @{typ "real \<Rightarrow> real"} and \<open>F\<close> and \<open>G\<close> real filters.
+
+  The functions \<open>f\<close> and \<open>g\<close> can be built from the operations mentioned before and may contain free 
+  variables. The filters \<open>F\<close> and \<open>G\<close> can be either \<open>\<plusminus>\<infinity>\<close> or a finite point in \<open>\<real>\<close>, possibly
+  with approach from the left or from the right.
+
+  The class of statements that is supported by @{method real_asymp} then consists of:
+
+    \<^item> Limits, i.\,e.\ @{prop "filterlim f F G"}
+
+    \<^item> Landau symbols, i.\,e.\ @{prop "f \<in> O[F](g)"} and analogously for \<^emph>\<open>o\<close>, \<open>\<Omega>\<close>, \<open>\<omega>\<close>, \<open>\<Theta>\<close>
+
+    \<^item> Asymptotic equivalence, i.\,e.\ @{prop "f \<sim>[F] g"}
+
+    \<^item> Asymptotic inequalities, i.\,e.\ @{prop "eventually (\<lambda>x. f x \<le> g x) F"}
+
+  For typical problems arising in practice that do not contain free variables, @{method real_asymp}
+  tends to succeed fully automatically within fractions of seconds, e.\,g.:
+\<close>
+
+lemma \<open>filterlim (\<lambda>x::real. (1 + 1 / x) powr x) (nhds (exp 1)) at_top\<close>
+  by real_asymp
+
+text \<open>
+  What makes the method \<^emph>\<open>semi-automatic\<close> is the fact that it has to internally determine the 
+  sign of real-valued constants and identical zeroness of real-valued functions, and while the
+  internal heuristics for this work very well most of the time, there are situations where the 
+  method fails to determine the sign of a constant, in which case the user needs to go back and
+  supply more information about the sign of that constant in order to get a result.
+
+  A simple example is the following:
+\<close>
+(*<*)
+experiment
+  fixes a :: real
+begin
+(*>*)
+lemma \<open>filterlim (\<lambda>x::real. exp (a * x)) at_top at_top\<close>
+oops
+
+text \<open>
+  Here, @{method real_asymp} outputs an error message stating that it could not determine the
+  sign of the free variable @{term "a :: real"}. In this case, the user must add the assumption
+  \<open>a > 0\<close> and give it to @{method real_asymp}.
+\<close>
+lemma
+  assumes \<open>a > 0\<close>
+  shows   \<open>filterlim (\<lambda>x::real. exp (a * x)) at_top at_top\<close>
+  using assms by real_asymp
+(*<*)
+end
+(*>*)
+text \<open>
+  Additional modifications to the simpset that is used for determining signs can also be made
+  with \<open>simp add:\<close> modifiers etc.\ in the same way as when using the @{method simp} method directly.
+
+  The same situation can also occur without free variables if the constant in question is a
+  complicated expression that the simplifier does not know enough ebout,
+  e.\,g.\ @{term "pi - exp 1"}.
+
+  In order to trace problems with sign determination, the \<open>(verbose)\<close> option can be passed to
+  @{method real_asymp}. It will then print a detailed error message whenever it encounters
+  a sign determination problem that it cannot solve.
+
+  The \<open>(fallback)\<close> flag causes the method not to use asymptotic interval arithmetic, but only
+  the much simpler core mechanism of computing a single Multiseries expansion for the input
+  function. There should normally be no need to use this flag; however, the core part of the 
+  code has been tested much more rigorously than the asymptotic interval part. In case there 
+  is some implementation problem with it for a problem that can be solved without it, the 
+  \<open>(fallback)\<close> option can be used until that problem is resolved.
+\<close>
+
+
+
+section \<open>Diagnostic Commands\<close>
+
+text \<open>
+  \[\begin{array}{rcl}
+    @{command_def (HOL) real_limit} & : & \<open>context \<rightarrow>\<close> \\
+    @{command_def (HOL) real_expansion} & : & \<open>context \<rightarrow>\<close>
+  \end{array}\]
+
+  @{rail \<open>
+    @@{command (HOL) real_limit} (limitopt*)
+    ;
+    @@{command (HOL) real_expansion} (expansionopt*)
+    ;
+    limitopt : ('limit' ':' @{syntax term}) | ('facts' ':' @{syntax thms})
+    ;
+    expansionopt : 
+        ('limit' ':' @{syntax term}) |
+        ('terms' ':' @{syntax nat} ('(' 'strict' ')') ?) |
+        ('facts' ':' @{syntax thms})
+  \<close>}
+
+    \<^descr>@{command real_limit} computes the limit of the given function \<open>f(x)\<close> for as \<open>x\<close> tends
+    to the specified limit point. Additional facts can be provided with the \<open>facts\<close> option, 
+    similarly to the @{command using} command with @{method real_asymp}. The limit point given
+    by the \<open>limit\<close> option must be one of the filters @{term "at_top"}, @{term "at_bot"}, 
+    @{term "at_left"}, or @{term "at_right"}. If no limit point is given, @{term "at_top"} is used
+    by default.
+  
+    \<^medskip>
+    The output of @{command real_limit} can be \<open>\<infinity>\<close>, \<open>-\<infinity>\<close>, \<open>\<plusminus>\<infinity>\<close>, \<open>c\<close> (for some real constant \<open>c\<close>),
+    \<open>0\<^sup>+\<close>, or \<open>0\<^sup>-\<close>. The \<open>+\<close> and \<open>-\<close> in the last case indicate whether the approach is from above
+    or from below (corresponding to @{term "at_right (0::real)"} and @{term "at_left (0::real)"}); 
+    for technical reasons, this information is currently not displayed if the limit is not 0.
+  
+    \<^medskip>
+    If the given function does not tend to a definite limit (e.\,g.\ @{term "sin x"} for \<open>x \<rightarrow> \<infinity>\<close>),
+    the command might nevertheless succeed to compute an asymptotic upper and/or lower bound and
+    display the limits of these bounds instead.
+
+    \<^descr>@{command real_expansion} works similarly to @{command real_limit}, but displays the first few
+    terms of the asymptotic multiseries expansion of the given function at the given limit point 
+    and the order of growth of the remainder term.
+  
+    In addition to the options already explained for the @{command real_limit} command, it takes
+    an additional option \<open>terms\<close> that controls how many of the leading terms of the expansion are
+    printed. If the \<open>(strict)\<close> modifier is passed to the \<open>terms option\<close>, terms whose coefficient is
+    0 are dropped from the output and do not count to the number of terms to be output.
+  
+    \<^medskip>
+    By default, the first three terms are output and the \<open>strict\<close> option is disabled.
+
+  Note that these two commands are intended for diagnostic use only. While the central part
+  of their implementation – finding a multiseries expansion and reading off the limit – are the
+  same as in the @{method real_asymp} method and therefore trustworthy, there is a small amount
+  of unverified code involved in pre-processing and printing (e.\,g.\ for reducing all the
+  different options for the \<open>limit\<close> option to the @{term at_top} case).
+\<close>
+
+
+section \<open>Extensibility\<close>
+
+subsection \<open>Basic fact collections\<close>
+
+text \<open>
+  The easiest way to add support for additional operations is to add corresponding facts
+  to one of the following fact collections. However, this only works for particularly simple cases.
+
+    \<^descr>@{thm [source] real_asymp_reify_simps} specifies a list of (unconditional) equations that are 
+     unfolded as a first step of @{method real_asymp} and the related commands. This can be used to 
+     add support for operations that can be represented easily by other operations that are already
+     supported, such as @{term sinh}, which is equal to @{term "\<lambda>x. (exp x - exp (-x)) / 2"}.
+
+    \<^descr>@{thm [source] real_asymp_nat_reify} and @{thm [source] real_asymp_int_reify} is used to
+     convert operations on natural numbers or integers to operations on real numbers. This enables
+     @{method real_asymp} to also work on functions that return a natural number or an integer.
+\<close>
+
+subsection \<open>Expanding Custom Operations\<close>
+
+text \<open>
+  Support for some non-trivial new operation \<open>f(x\<^sub>1, \<dots>, x\<^sub>n)\<close> can be added dynamically at any
+  time, but it involves writing ML code and involves a significant amount of effort, especially
+  when the function has essential singularities.
+
+  The first step is to write an ML function that takes as arguments
+    \<^item> the expansion context
+    \<^item> the term \<open>t\<close> to expand (which will be of the form \<open>f(g\<^sub>1(x), \<dots>, g\<^sub>n(x))\<close>)
+    \<^item> a list of \<open>n\<close> theorems of the form @{prop "(g\<^sub>i expands_to G\<^sub>i) bs"}
+    \<^item> the current basis \<open>bs\<close>
+  and returns a theorem of the form @{prop "(t expands_to F) bs'"} and a new basis \<open>bs'\<close> which 
+  must be a superset of the original basis.
+
+  This function must then be registered as a handler for the operation by proving a vacuous lemma
+  of the form @{prop "REAL_ASYMP_CUSTOM f"} (which is only used for tagging) and passing that
+  lemma and the expansion function to @{ML [source] Exp_Log_Expression.register_custom_from_thm}
+  in a @{command local_setup} invocation.
+
+  If the expansion produced by the handler function contains new definitions, corresponding 
+  evaluation equations must be added to the fact collection @{thm [source] real_asymp_eval_eqs}.
+  These equations must have the form \<open>h p\<^sub>1 \<dots> p\<^sub>m = rhs\<close> where \<open>h\<close> must be a constant of arity \<open>m\<close>
+  (i.\,e.\ the function on the left-hand side must be fully applied) and the \<open>p\<^sub>i\<close> can be patterns
+  involving \<open>constructors\<close>.
+
+  New constructors for this pattern matching can be registered by adding a theorem of the form
+  @{prop "REAL_ASYMP_EVAL_CONSTRUCTOR c"} to the fact collection
+  @{thm [source] exp_log_eval_constructor}, but this should be quite rare in practice.
+
+  \<^medskip>
+  Note that there is currently no way to add support for custom operations in connection with
+  `oscillating' terms. The above mechanism only works if all arguments of the new operation have
+  an exact multiseries expansion.
+\<close>
+
+
+subsection \<open>Extending the Sign Oracle\<close>
+
+text \<open>
+  By default, the \<^verbatim>\<open>real_asymp\<close> package uses the simplifier with the given simpset and facts
+  in order to determine the sign of real constants. This is done by invoking the simplifier
+  on goals like \<open>c = 0\<close>, \<open>c \<noteq> 0\<close>, \<open>c > 0\<close>, or \<open>c < 0\<close> or some subset thereof, depending on the
+  context.
+
+  If the simplifier cannot prove any of them, the entire method (or command) invocation will fail.
+  It is, however, possible to dynamically add additional sign oracles that will be tried; the 
+  most obvious candidate for an oracle that one may want to add or remove dynamically are 
+  approximation-based tactics.
+
+  Adding such a tactic can be done by calling
+  @{ML [source] Multiseries_Expansion.register_sign_oracle}. Note that if the tactic cannot prove
+  a goal, it should fail as fast as possible.
+\<close>
+
+(*<*)
+end
+(*>*)
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Real_Asymp/document/root.tex	Sun Jul 15 14:46:57 2018 +0200
@@ -0,0 +1,39 @@
+\documentclass[11pt,a4paper]{article}
+\usepackage{amsfonts, amsmath, amssymb}
+\usepackage{railsetup}
+\usepackage{iman}
+\usepackage{extra}
+\usepackage{isar}
+\usepackage{isabelle}
+\usepackage{isabellesym}
+\usepackage{style}
+
+% this should be the last package used
+\usepackage{pdfsetup}
+
+% urls in roman style, theory text in math-similar italics
+\urlstyle{rm}
+\isabellestyle{it}
+
+
+\begin{document}
+
+\title{\texttt{real\_asymp}: Semi-Automatic Real Asymptotics\\ in Isabelle\slash HOL}
+\author{Manuel Eberl}
+\maketitle
+
+\tableofcontents
+\newpage
+\parindent 0pt\parskip 0.5ex
+
+\input{session}
+
+\bibliographystyle{abbrv}
+\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Real_Asymp/document/style.sty	Sun Jul 15 14:46:57 2018 +0200
@@ -0,0 +1,46 @@
+
+%% toc
+\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1}
+\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}}
+
+%% paragraphs
+\setlength{\parindent}{1em}
+
+%% references
+\newcommand{\secref}[1]{\S\ref{#1}}
+\newcommand{\figref}[1]{figure~\ref{#1}}
+
+%% logical markup
+\newcommand{\strong}[1]{{\bfseries {#1}}}
+\newcommand{\qn}[1]{\emph{#1}}
+
+%% typographic conventions
+\newcommand{\qt}[1]{``{#1}''}
+\newcommand{\ditem}[1]{\item[\isastyletext #1]}
+
+%% quote environment
+\isakeeptag{quote}
+\renewenvironment{quote}
+  {\list{}{\leftmargin2em\rightmargin0pt}\parindent0pt\parskip0pt\item\relax}
+  {\endlist}
+\renewcommand{\isatagquote}{\begin{quote}}
+\renewcommand{\endisatagquote}{\end{quote}}
+\newcommand{\quotebreak}{\\[1.2ex]}
+
+%% presentation
+\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2}
+
+%% character detail
+\renewcommand{\isadigit}[1]{\isamath{#1}}
+\binperiod
+\underscoreoff
+
+%% format
+\pagestyle{headings}
+\isabellestyle{it}
+
+
+%%% Local Variables: 
+%%% mode: latex
+%%% TeX-master: "implementation"
+%%% End: 
--- a/src/HOL/ROOT	Sun Jul 15 01:14:04 2018 +0100
+++ b/src/HOL/ROOT	Sun Jul 15 14:46:57 2018 +0200
@@ -78,6 +78,14 @@
     (*conflicting type class instantiations and dependent applications*)
     Field_as_Ring
 
+session "HOL-Real_Asymp" in Real_Asymp = HOL +
+  sessions
+    "HOL-Decision_Procs"
+  theories
+    Real_Asymp
+    Real_Asymp_Approx
+    Real_Asymp_Examples    
+
 session "HOL-Hahn_Banach" in Hahn_Banach = "HOL-Library" +
   description {*
     Author:     Gertrud Bauer, TU Munich
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real_Asymp/Eventuallize.thy	Sun Jul 15 14:46:57 2018 +0200
@@ -0,0 +1,57 @@
+section \<open>Lifting theorems onto filters\<close>
+theory Eventuallize
+  imports Complex_Main
+begin
+
+text \<open>
+  The following attribute and ML function lift a given theorem of the form
+    @{prop "\<forall>x. A x \<longrightarrow> B x \<longrightarrow> C x"}
+  to
+    @{prop "eventually A F \<Longrightarrow> eventually B F \<Longrightarrow> eventually C F"} .
+\<close>
+
+ML \<open>
+signature EVENTUALLIZE =
+sig
+val eventuallize : Proof.context -> thm -> int option -> thm
+end
+
+structure Eventuallize : EVENTUALLIZE =
+struct
+
+fun dest_All (Const (@{const_name "HOL.All"}, _) $ Abs (x, T, t)) = (x, T, t)
+  | dest_All (Const (@{const_name "HOL.All"}, T) $ f) =
+      ("x", T |> dest_funT |> fst |> dest_funT |> fst, f $ Bound 0)
+  | dest_All t = raise TERM ("dest_All", [t])
+
+fun strip_imp (@{term "(\<longrightarrow>)"} $ a $ b) = apfst (cons a) (strip_imp b)
+  | strip_imp t = ([], t)
+
+fun eventuallize ctxt thm n =
+  let
+    fun err n = raise THM ("Invalid index in eventuallize: " ^ Int.toString n, 0, [thm])
+    val n_max =
+      thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> dest_All |> #3 |> strip_imp |> fst |> length
+    val _ = case n of NONE => () | SOME n =>
+      if n >= 0 andalso n <= n_max then () else err n
+    val n = case n of SOME n => n | NONE => n_max
+    fun conv n =
+      if n < 2 then Conv.all_conv else Conv.arg_conv (conv (n - 1)) then_conv
+        Conv.rewr_conv @{thm eq_reflection[OF imp_conjL [symmetric]]}
+    val conv = Conv.arg_conv (Conv.binder_conv (K (conv n)) ctxt)
+    val thm' = Conv.fconv_rule conv thm
+    fun go n = if n < 2 then @{thm _} else @{thm eventually_conj} OF [@{thm _}, go (n - 1)]
+  in
+    @{thm eventually_rev_mp[OF _ always_eventually]} OF [go n, thm']
+  end
+
+end
+\<close>
+
+attribute_setup eventuallized = \<open>
+  Scan.lift (Scan.option Parse.nat) >>
+    (fn n => fn (ctxt, thm) =>
+      (NONE, SOME (Eventuallize.eventuallize (Context.proof_of ctxt) thm n)))
+\<close>
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real_Asymp/Inst_Existentials.thy	Sun Jul 15 14:46:57 2018 +0200
@@ -0,0 +1,21 @@
+section \<open>Tactic for instantiating existentials\<close>
+theory Inst_Existentials
+  imports Main
+begin
+
+text \<open>
+  Coinduction proofs in Isabelle often lead to proof obligations with nested conjunctions and
+  existential quantifiers, e.g. @{prop "\<exists>x y. P x y \<and> (\<exists>z. Q x y z)"} .
+
+  The following tactic allows instantiating these existentials with a given list of witnesses.
+\<close>
+
+ML_file \<open>inst_existentials.ML\<close>
+
+method_setup inst_existentials = \<open>
+  Scan.lift (Scan.repeat Parse.term) >> 
+    (fn ts => fn ctxt => SIMPLE_METHOD' (Inst_Existentials.tac ctxt 
+       (map (Syntax.read_term ctxt)  ts)))
+\<close>
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real_Asymp/Lazy_Eval.thy	Sun Jul 15 14:46:57 2018 +0200
@@ -0,0 +1,39 @@
+section \<open>Lazy evaluation within the logic\<close>
+theory Lazy_Eval
+imports
+  Complex_Main
+begin
+
+text \<open>  
+  This is infrastructure to lazily evaluate an expression (typically something corecursive)
+  within the logic by simple rewriting. A signature of supported (co-)datatype constructures 
+  upon which pattern matching is allowed and a list of function equations that are used in
+  rewriting must be provided.
+
+  One can then e.\,g.\ determine whether a given pattern matches a given expression. To do this,
+  the expression will be rewritten using the given function equations until enough constructors
+  have been exposed to decide whether the pattern matches.
+
+  This infrastructure was developed specifically for evaluating Multiseries expressions, but
+  can, in principle, be used for other purposes as well.
+\<close>
+
+lemma meta_eq_TrueE: "PROP P \<equiv> Trueprop True \<Longrightarrow> PROP P" by simp
+
+datatype cmp_result = LT | EQ | GT
+
+definition COMPARE :: "real \<Rightarrow> real \<Rightarrow> cmp_result" where
+  "COMPARE x y = (if x < y then LT else if x > y then GT else EQ)"
+
+lemma COMPARE_intros: 
+  "x < y \<Longrightarrow> COMPARE x y \<equiv> LT" "x > y \<Longrightarrow> COMPARE x y \<equiv> GT" "x = y \<Longrightarrow> COMPARE x y \<equiv> EQ"
+  by (simp_all add: COMPARE_def)
+  
+primrec CMP_BRANCH :: "cmp_result \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where
+  "CMP_BRANCH LT x y z = x"
+| "CMP_BRANCH EQ x y z = y"
+| "CMP_BRANCH GT x y z = z"  
+
+ML_file \<open>lazy_eval.ML\<close>
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real_Asymp/Multiseries_Expansion.thy	Sun Jul 15 14:46:57 2018 +0200
@@ -0,0 +1,5390 @@
+(*
+  File:    Multiseries_Expansion.thy
+  Author:  Manuel Eberl, TU München
+
+  Asymptotic bases and Multiseries expansions of real-valued functions.
+  Contains automation to prove limits and asymptotic relationships between such functions.
+*)
+section \<open>Multiseries expansions\<close>
+theory Multiseries_Expansion
+imports
+  "HOL-Library.BNF_Corec"
+  "HOL-Library.Landau_Symbols"
+  Lazy_Eval
+  Inst_Existentials
+  Eventuallize
+begin
+
+(* TODO Move *)
+lemma real_powr_at_top: 
+  assumes "(p::real) > 0"
+  shows   "filterlim (\<lambda>x. x powr p) at_top at_top"
+proof (subst filterlim_cong[OF refl refl])
+  show "LIM x at_top. exp (p * ln x) :> at_top"
+    by (rule filterlim_compose[OF exp_at_top filterlim_tendsto_pos_mult_at_top[OF tendsto_const]])
+       (simp_all add: ln_at_top assms)
+  show "eventually (\<lambda>x. x powr p = exp (p * ln x)) at_top"
+    using eventually_gt_at_top[of 0] by eventually_elim (simp add: powr_def)
+qed
+
+
+subsection \<open>Streams and lazy lists\<close>
+
+codatatype 'a msstream = MSSCons 'a "'a msstream"
+
+primrec mssnth :: "'a msstream \<Rightarrow> nat \<Rightarrow> 'a" where
+  "mssnth xs 0 = (case xs of MSSCons x xs \<Rightarrow> x)"
+| "mssnth xs (Suc n) = (case xs of MSSCons x xs \<Rightarrow> mssnth xs n)"
+
+codatatype 'a msllist = MSLNil | MSLCons 'a "'a msllist"
+  for map: mslmap
+
+fun lcoeff where
+  "lcoeff MSLNil n = 0"
+| "lcoeff (MSLCons x xs) 0 = x"
+| "lcoeff (MSLCons x xs) (Suc n) = lcoeff xs n"
+
+primcorec msllist_of_msstream :: "'a msstream \<Rightarrow> 'a msllist" where
+  "msllist_of_msstream xs = (case xs of MSSCons x xs \<Rightarrow> MSLCons x (msllist_of_msstream xs))"
+  
+lemma msllist_of_msstream_MSSCons [simp]: 
+  "msllist_of_msstream (MSSCons x xs) = MSLCons x (msllist_of_msstream xs)"
+  by (subst msllist_of_msstream.code) simp
+
+lemma lcoeff_llist_of_stream [simp]: "lcoeff (msllist_of_msstream xs) n = mssnth xs n"
+  by (induction "msllist_of_msstream xs" n arbitrary: xs rule: lcoeff.induct;
+      subst msllist_of_msstream.code) (auto split: msstream.splits)
+
+
+subsection \<open>Convergent power series\<close>
+
+subsubsection \<open>Definition\<close>
+
+definition convergent_powser :: "real msllist \<Rightarrow> bool" where
+  "convergent_powser cs \<longleftrightarrow> (\<exists>R>0. \<forall>x. abs x < R \<longrightarrow> summable (\<lambda>n. lcoeff cs n * x ^ n))"
+ 
+lemma convergent_powser_stl:
+  assumes "convergent_powser (MSLCons c cs)"
+  shows   "convergent_powser cs"
+proof -
+  from assms obtain R where "R > 0" "\<And>x. abs x < R \<Longrightarrow> summable (\<lambda>n. lcoeff (MSLCons c cs) n * x ^ n)"
+    unfolding convergent_powser_def by blast
+  hence "summable (\<lambda>n. lcoeff cs n * x ^ n)" if "abs x < R" for x
+    using that by (subst (asm) summable_powser_split_head [symmetric]) simp
+  thus ?thesis using \<open>R > 0\<close> by (auto simp: convergent_powser_def)
+qed
+  
+  
+definition powser :: "real msllist \<Rightarrow> real \<Rightarrow> real" where
+  "powser cs x = suminf (\<lambda>n. lcoeff cs n * x ^ n)"
+
+lemma isCont_powser:
+  assumes "convergent_powser cs"
+  shows   "isCont (powser cs) 0"
+proof -
+  from assms obtain R where R: "R > 0" "\<And>x. abs x < R \<Longrightarrow> summable (\<lambda>n. lcoeff cs n * x ^ n)"
+    unfolding convergent_powser_def by blast
+  hence *: "summable (\<lambda>n. lcoeff cs n * (R/2) ^ n)" by (intro R) simp_all
+  from \<open>R > 0\<close> show ?thesis unfolding powser_def
+    by (intro isCont_powser[OF *]) simp_all
+qed
+
+lemma powser_MSLNil [simp]: "powser MSLNil = (\<lambda>_. 0)"
+  by (simp add: fun_eq_iff powser_def)
+
+lemma powser_MSLCons:
+  assumes "convergent_powser (MSLCons c cs)"
+  shows   "eventually (\<lambda>x. powser (MSLCons c cs) x = x * powser cs x + c) (nhds 0)"
+proof -
+  from assms obtain R where R: "R > 0" "\<And>x. abs x < R \<Longrightarrow> summable (\<lambda>n. lcoeff (MSLCons c cs) n * x ^ n)"
+    unfolding convergent_powser_def by blast
+  from R have "powser (MSLCons c cs) x = x * powser cs x + c" if "abs x < R" for x
+    using that unfolding powser_def by (subst powser_split_head) simp_all
+  moreover have "eventually (\<lambda>x. abs x < R) (nhds 0)"
+    using \<open>R > 0\<close> filterlim_ident[of "nhds (0::real)"]
+    by (subst (asm) tendsto_iff) (simp add: dist_real_def)
+  ultimately show ?thesis by (auto elim: eventually_mono)
+qed
+
+definition convergent_powser' :: "real msllist \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> bool" where
+  "convergent_powser' cs f \<longleftrightarrow> (\<exists>R>0. \<forall>x. abs x < R \<longrightarrow> (\<lambda>n. lcoeff cs n * x ^ n) sums f x)"
+
+lemma convergent_powser'_imp_convergent_powser:
+  "convergent_powser' cs f \<Longrightarrow> convergent_powser cs" 
+  unfolding convergent_powser_def convergent_powser'_def by (auto simp: sums_iff)
+
+lemma convergent_powser'_eventually:
+  assumes "convergent_powser' cs f"
+  shows   "eventually (\<lambda>x. powser cs x = f x) (nhds 0)"
+proof -
+  from assms obtain R where "R > 0" "\<And>x. abs x < R \<Longrightarrow> (\<lambda>n. lcoeff cs n * x ^ n) sums f x"
+    unfolding convergent_powser'_def by blast
+  hence "powser cs x = f x" if "abs x < R" for x
+    using that by (simp add: powser_def sums_iff)
+  moreover have "eventually (\<lambda>x. abs x < R) (nhds 0)"
+    using \<open>R > 0\<close> filterlim_ident[of "nhds (0::real)"]
+    by (subst (asm) tendsto_iff) (simp add: dist_real_def)
+  ultimately show ?thesis by (auto elim: eventually_mono)
+qed  
+
+
+subsubsection \<open>Geometric series\<close>
+
+primcorec mssalternate :: "'a \<Rightarrow> 'a \<Rightarrow> 'a msstream" where
+  "mssalternate a b = MSSCons a (mssalternate b a)"
+
+lemma case_mssalternate [simp]: 
+  "(case mssalternate a b of MSSCons c d \<Rightarrow> f c d) = f a (mssalternate b a)"
+  by (subst mssalternate.code) simp
+
+lemma mssnth_mssalternate: "mssnth (mssalternate a b) n = (if even n then a else b)"
+  by (induction n arbitrary: a b; subst mssalternate.code) simp_all
+  
+lemma convergent_powser'_geometric:
+  "convergent_powser' (msllist_of_msstream (mssalternate 1 (-1))) (\<lambda>x. inverse (1 + x))"
+proof -
+  have "(\<lambda>n. lcoeff (msllist_of_msstream (mssalternate 1 (-1))) n * x ^ n) sums (inverse (1 + x))"
+    if "abs x < 1" for x :: real
+  proof -
+    have "(\<lambda>n. (-1) ^ n * x ^ n) sums (inverse (1 + x))"
+      using that geometric_sums[of "-x"] by (simp add: field_simps power_mult_distrib [symmetric])
+    also have "(\<lambda>n. (-1) ^ n * x ^ n) =
+                 (\<lambda>n. lcoeff (msllist_of_msstream (mssalternate 1 (-1))) n * x ^ n)"
+      by (auto simp add: mssnth_mssalternate fun_eq_iff)
+    finally show ?thesis .
+  qed
+  thus ?thesis unfolding convergent_powser'_def by (force intro!: exI[of _ 1])
+qed
+
+
+subsubsection \<open>Exponential series\<close>
+
+primcorec exp_series_stream_aux :: "real \<Rightarrow> real \<Rightarrow> real msstream" where
+  "exp_series_stream_aux acc n = MSSCons acc (exp_series_stream_aux (acc / n) (n + 1))"
+  
+lemma mssnth_exp_series_stream_aux:
+  "mssnth (exp_series_stream_aux (1 / fact n) (n + 1)) m = 1 / fact (n + m)"
+proof (induction m arbitrary: n)
+  case (0 n)
+  thus ?case by (subst exp_series_stream_aux.code) simp
+next
+  case (Suc m n)
+  from Suc.IH[of "n + 1"] show ?case
+    by (subst exp_series_stream_aux.code) (simp add: algebra_simps)
+qed
+
+definition exp_series_stream :: "real msstream" where
+  "exp_series_stream = exp_series_stream_aux 1 1"
+  
+lemma mssnth_exp_series_stream:
+  "mssnth exp_series_stream n = 1 / fact n"
+  unfolding exp_series_stream_def using mssnth_exp_series_stream_aux[of 0 n] by simp
+
+lemma convergent_powser'_exp:
+  "convergent_powser' (msllist_of_msstream exp_series_stream) exp"
+proof -
+  have "(\<lambda>n. lcoeff (msllist_of_msstream exp_series_stream) n * x ^ n) sums exp x" for x :: real
+    using exp_converges[of x] by (simp add: mssnth_exp_series_stream divide_simps)
+  thus ?thesis by (auto intro: exI[of _ "1::real"] simp: convergent_powser'_def)
+qed
+  
+
+subsubsection \<open>Logarithm series\<close>
+
+primcorec ln_series_stream_aux :: "bool \<Rightarrow> real \<Rightarrow> real msstream" where
+  "ln_series_stream_aux b n = 
+     MSSCons (if b then -inverse n else inverse n) (ln_series_stream_aux (\<not>b) (n+1))"
+
+lemma mssnth_ln_series_stream_aux:
+  "mssnth (ln_series_stream_aux b x) n = 
+     (if b then - ((-1)^n) * inverse (x + real n) else ((-1)^n) * inverse (x + real n))"
+  by (induction n arbitrary: b x; subst ln_series_stream_aux.code) simp_all
+
+definition ln_series_stream :: "real msstream" where
+  "ln_series_stream = MSSCons 0 (ln_series_stream_aux False 1)"
+  
+lemma mssnth_ln_series_stream:
+  "mssnth ln_series_stream n = (-1) ^ Suc n / real n"
+  unfolding ln_series_stream_def 
+  by (cases n) (simp_all add: mssnth_ln_series_stream_aux field_simps)
+
+lemma ln_series': 
+  assumes "abs (x::real) < 1"
+  shows   "(\<lambda>n. - ((-x)^n) / of_nat n) sums ln (1 + x)"
+proof -
+  have "\<forall>n\<ge>1. norm (-((-x)^n)) / of_nat n \<le> norm x ^ n / 1"
+    by (intro exI[of _ "1::nat"] allI impI frac_le) (simp_all add: power_abs)
+  hence "\<exists>N. \<forall>n\<ge>N. norm (-((-x)^n) / of_nat n) \<le> norm x ^ n" 
+    by (intro exI[of _ 1]) simp_all
+  moreover from assms have "summable (\<lambda>n. norm x ^ n)" 
+    by (intro summable_geometric) simp_all
+  ultimately have *: "summable (\<lambda>n. - ((-x)^n) / of_nat n)"
+    by (rule summable_comparison_test)
+  moreover from assms have "0 < 1 + x" "1 + x < 2" by simp_all
+  from ln_series[OF this] 
+    have "ln (1 + x) = (\<Sum>n. - ((-x) ^ Suc n) / real (Suc n))"
+    by (simp_all add: power_minus' mult_ac)
+  hence "ln (1 + x) = (\<Sum>n. - ((-x) ^ n / real n))"
+    by (subst (asm) suminf_split_head[OF *]) simp_all
+  ultimately show ?thesis by (simp add: sums_iff)
+qed
+
+lemma convergent_powser'_ln:
+  "convergent_powser' (msllist_of_msstream ln_series_stream) (\<lambda>x. ln (1 + x))"
+proof -
+  have "(\<lambda>n. lcoeff (msllist_of_msstream ln_series_stream)n * x ^ n) sums ln (1 + x)"
+    if "abs x < 1" for x
+  proof -
+    from that have "(\<lambda>n. - ((- x) ^ n) / real n) sums ln (1 + x)" by (rule ln_series')
+    also have "(\<lambda>n. - ((- x) ^ n) / real n) =
+                 (\<lambda>n. lcoeff (msllist_of_msstream ln_series_stream) n * x ^ n)"
+      by (auto simp: fun_eq_iff mssnth_ln_series_stream power_mult_distrib [symmetric])
+    finally show ?thesis .
+  qed
+  thus ?thesis unfolding convergent_powser'_def by (force intro!: exI[of _ 1])
+qed
+
+
+subsubsection \<open>Generalized binomial series\<close>
+
+primcorec gbinomial_series_aux :: "bool \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real msllist" where
+  "gbinomial_series_aux abort x n acc =
+     (if abort \<and> acc = 0 then MSLNil else 
+        MSLCons acc (gbinomial_series_aux abort x (n + 1) ((x - n) / (n + 1) * acc)))"
+
+lemma gbinomial_series_abort [simp]: "gbinomial_series_aux True x n 0 = MSLNil"
+  by (subst gbinomial_series_aux.code) simp
+
+definition gbinomial_series :: "bool \<Rightarrow> real \<Rightarrow> real msllist" where
+  "gbinomial_series abort x = gbinomial_series_aux abort x 0 1"
+
+
+(* TODO Move *)
+lemma gbinomial_Suc_rec:
+  "(x gchoose (Suc k)) = (x gchoose k) * ((x - of_nat k) / (of_nat k + 1))"
+proof -
+  have "((x - 1) + 1) gchoose (Suc k) = x * (x - 1 gchoose k) / (1 + of_nat k)"
+    by (subst gbinomial_factors) simp
+  also have "x * (x - 1 gchoose k) = (x - of_nat k) * (x gchoose k)"
+    by (rule gbinomial_absorb_comp [symmetric])
+  finally show ?thesis by (simp add: algebra_simps)
+qed
+
+lemma lcoeff_gbinomial_series_aux:
+  "lcoeff (gbinomial_series_aux abort x m (x gchoose m)) n = (x gchoose (n + m))"
+proof (induction n arbitrary: m)
+  case 0
+  show ?case by (subst gbinomial_series_aux.code) simp
+next
+  case (Suc n m)
+  show ?case
+  proof (cases "abort \<and> (x gchoose m) = 0")
+    case True
+    with gbinomial_mult_fact[of m x] obtain k where "x = real k" "k < m"
+      by auto
+    hence "(x gchoose Suc (n + m)) = 0"
+      using gbinomial_mult_fact[of "Suc (n + m)" x]
+      by (simp add: gbinomial_altdef_of_nat)
+    with True show ?thesis by (subst gbinomial_series_aux.code) simp
+  next
+    case False
+    hence "lcoeff (gbinomial_series_aux abort x m (x gchoose m)) (Suc n) = 
+             lcoeff (gbinomial_series_aux abort x (Suc m) ((x gchoose m) *
+             ((x - real m) / (real m + 1)))) n"
+      by (subst gbinomial_series_aux.code) (auto simp: algebra_simps)
+    also have "((x gchoose m) * ((x - real m) / (real m + 1))) = x gchoose (Suc m)" 
+      by (rule gbinomial_Suc_rec [symmetric])
+    also have "lcoeff (gbinomial_series_aux abort x (Suc m) \<dots>) n = x gchoose (n + Suc m)"
+      by (rule Suc.IH)
+    finally show ?thesis by simp
+  qed
+qed
+
+lemma lcoeff_gbinomial_series [simp]:
+  "lcoeff (gbinomial_series abort x) n = (x gchoose n)"
+  using lcoeff_gbinomial_series_aux[of abort x 0 n] by (simp add: gbinomial_series_def)
+
+lemma gbinomial_ratio_limit:
+  fixes a :: "'a :: real_normed_field"
+  assumes "a \<notin> \<nat>"
+  shows "(\<lambda>n. (a gchoose n) / (a gchoose Suc n)) \<longlonglongrightarrow> -1"
+proof (rule Lim_transform_eventually)
+  let ?f = "\<lambda>n. inverse (a / of_nat (Suc n) - of_nat n / of_nat (Suc n))"
+  from eventually_gt_at_top[of "0::nat"]
+    show "eventually (\<lambda>n. ?f n = (a gchoose n) /(a gchoose Suc n)) sequentially"
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    then obtain q where q: "n = Suc q" by (cases n) blast
+    let ?P = "\<Prod>i=0..<n. a - of_nat i"
+    from n have "(a gchoose n) / (a gchoose Suc n) = (of_nat (Suc n) :: 'a) *
+                   (?P / (\<Prod>i=0..n. a - of_nat i))"
+      by (simp add: gbinomial_prod_rev atLeastLessThanSuc_atLeastAtMost)
+    also from q have "(\<Prod>i=0..n. a - of_nat i) = ?P * (a - of_nat n)"
+      by (simp add: prod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost)
+    also have "?P / \<dots> = (?P / ?P) / (a - of_nat n)" by (rule divide_divide_eq_left[symmetric])
+    also from assms have "?P / ?P = 1" by auto
+    also have "of_nat (Suc n) * (1 / (a - of_nat n)) =
+                   inverse (inverse (of_nat (Suc n)) * (a - of_nat n))" by (simp add: field_simps)
+    also have "inverse (of_nat (Suc n)) * (a - of_nat n) =
+                 a / of_nat (Suc n) - of_nat n / of_nat (Suc n)"
+      by (simp add: field_simps del: of_nat_Suc)
+    finally show "?f n = (a gchoose n) / (a gchoose Suc n)" by simp
+  qed
+
+  have "(\<lambda>n. norm a / (of_nat (Suc n))) \<longlonglongrightarrow> 0"
+    unfolding divide_inverse
+    by (intro tendsto_mult_right_zero LIMSEQ_inverse_real_of_nat)
+  hence "(\<lambda>n. a / of_nat (Suc n)) \<longlonglongrightarrow> 0"
+    by (subst tendsto_norm_zero_iff[symmetric]) (simp add: norm_divide del: of_nat_Suc)
+  hence "?f \<longlonglongrightarrow> inverse (0 - 1)"
+    by (intro tendsto_inverse tendsto_diff LIMSEQ_n_over_Suc_n) simp_all
+  thus "?f \<longlonglongrightarrow> -1" by simp
+qed
+  
+lemma summable_gbinomial:
+  fixes a z :: real
+  assumes "norm z < 1"
+  shows   "summable (\<lambda>n. (a gchoose n) * z ^ n)"
+proof (cases "z = 0 \<or> a \<in> \<nat>")
+  case False
+  define r where "r = (norm z + 1) / 2"
+  from assms have r: "r > norm z" "r < 1" by (simp_all add: r_def field_simps)
+  from False have "abs z > 0" by auto
+  from False have "a \<notin> \<nat>" by (auto elim!: Nats_cases)
+  hence *: "(\<lambda>n. norm (z / ((a gchoose n) / (a gchoose (Suc n))))) \<longlonglongrightarrow> norm (z / (-1))"
+    by (intro tendsto_norm tendsto_divide tendsto_const gbinomial_ratio_limit) simp_all
+  hence "\<forall>\<^sub>F x in at_top. norm (z / ((a gchoose x) / (a gchoose Suc x))) > 0"
+    using assms False by (intro order_tendstoD) auto
+  hence nz: "\<forall>\<^sub>F x in at_top. (a gchoose x) \<noteq> 0" by eventually_elim auto
+      
+  have "\<forall>\<^sub>F x in at_top. norm (z / ((a gchoose x) / (a gchoose Suc x))) < r"
+    using assms r by (intro order_tendstoD(2)[OF *]) simp_all
+  with nz have "\<forall>\<^sub>F n in at_top. norm ((a gchoose (Suc n)) * z) \<le> r * norm (a gchoose n)"
+    by eventually_elim (simp add: field_simps abs_mult split: if_splits)
+  hence "\<forall>\<^sub>F n in at_top. norm (z ^ n) * norm ((a gchoose (Suc n)) * z) \<le>
+                           norm (z ^ n) * (r * norm (a gchoose n))"
+    by eventually_elim (insert False, intro mult_left_mono, auto)
+  hence "\<forall>\<^sub>F n in at_top. norm ((a gchoose (Suc n)) * z ^ (Suc n)) \<le> 
+                           r * norm ((a gchoose n) * z ^ n)"
+    by (simp add: abs_mult mult_ac)
+  then obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> norm ((a gchoose (Suc n)) * z ^ (Suc n)) \<le>
+                                          r * norm ((a gchoose n) * z ^ n)"
+    unfolding eventually_at_top_linorder by blast
+  show "summable (\<lambda>n. (a gchoose n) * z ^ n)"
+    by (intro summable_ratio_test[OF r(2), of N] N)
+next
+  case True
+  thus ?thesis
+  proof
+    assume "a \<in> \<nat>"
+    then obtain n where [simp]: "a = of_nat n" by (auto elim: Nats_cases)
+    from eventually_gt_at_top[of n] 
+      have "eventually (\<lambda>n. (a gchoose n) * z ^ n = 0) at_top"
+      by eventually_elim (simp add: binomial_gbinomial [symmetric])
+    from summable_cong[OF this] show ?thesis by simp
+  qed auto
+qed
+
+lemma gen_binomial_real:
+  fixes z :: real
+  assumes "norm z < 1"
+  shows   "(\<lambda>n. (a gchoose n) * z^n) sums (1 + z) powr a"
+proof -
+  define K where "K = 1 - (1 - norm z) / 2"
+  from assms have K: "K > 0" "K < 1" "norm z < K"
+     unfolding K_def by (auto simp: field_simps intro!: add_pos_nonneg)
+  let ?f = "\<lambda>n. a gchoose n" and ?f' = "diffs (\<lambda>n. a gchoose n)"
+  have summable_strong: "summable (\<lambda>n. ?f n * z ^ n)" if "norm z < 1" for z using that
+    by (intro summable_gbinomial)
+  with K have summable: "summable (\<lambda>n. ?f n * z ^ n)" if "norm z < K" for z using that by auto
+  hence summable': "summable (\<lambda>n. ?f' n * z ^ n)" if "norm z < K" for z using that
+    by (intro termdiff_converges[of _ K]) simp_all
+
+  define f f' where [abs_def]: "f z = (\<Sum>n. ?f n * z ^ n)" "f' z = (\<Sum>n. ?f' n * z ^ n)" for z
+  {
+    fix z :: real assume z: "norm z < K"
+    from summable_mult2[OF summable'[OF z], of z]
+      have summable1: "summable (\<lambda>n. ?f' n * z ^ Suc n)" by (simp add: mult_ac)
+    hence summable2: "summable (\<lambda>n. of_nat n * ?f n * z^n)"
+      unfolding diffs_def by (subst (asm) summable_Suc_iff)
+
+    have "(1 + z) * f' z = (\<Sum>n. ?f' n * z^n) + (\<Sum>n. ?f' n * z^Suc n)"
+      unfolding f_f'_def using summable' z by (simp add: algebra_simps suminf_mult)
+    also have "(\<Sum>n. ?f' n * z^n) = (\<Sum>n. of_nat (Suc n) * ?f (Suc n) * z^n)"
+      by (intro suminf_cong) (simp add: diffs_def)
+    also have "(\<Sum>n. ?f' n * z^Suc n) = (\<Sum>n. of_nat n * ?f n * z ^ n)"
+      using summable1 suminf_split_initial_segment[OF summable1] unfolding diffs_def
+      by (subst suminf_split_head, subst (asm) summable_Suc_iff) simp_all
+    also have "(\<Sum>n. of_nat (Suc n) * ?f (Suc n) * z^n) + (\<Sum>n. of_nat n * ?f n * z^n) =
+                 (\<Sum>n. a * ?f n * z^n)"
+      by (subst gbinomial_mult_1, subst suminf_add)
+         (insert summable'[OF z] summable2,
+          simp_all add: summable_powser_split_head algebra_simps diffs_def)
+    also have "\<dots> = a * f z" unfolding f_f'_def
+      by (subst suminf_mult[symmetric]) (simp_all add: summable[OF z] mult_ac)
+    finally have "a * f z = (1 + z) * f' z" by simp
+  } note deriv = this
+
+  have [derivative_intros]: "(f has_field_derivative f' z) (at z)" if "norm z < of_real K" for z
+    unfolding f_f'_def using K that
+    by (intro termdiffs_strong[of "?f" K z] summable_strong) simp_all
+  have "f 0 = (\<Sum>n. if n = 0 then 1 else 0)" unfolding f_f'_def by (intro suminf_cong) simp
+  also have "\<dots> = 1" using sums_single[of 0 "\<lambda>_. 1::real"] unfolding sums_iff by simp
+  finally have [simp]: "f 0 = 1" .
+
+  have "f z * (1 + z) powr (-a) = f 0 * (1 + 0) powr (-a)"
+  proof (rule DERIV_isconst3[where ?f = "\<lambda>x. f x * (1 + x) powr (-a)"])
+    fix z :: real assume z': "z \<in> {-K<..<K}"
+    hence z: "norm z < K" using K by auto
+    with K have nz: "1 + z \<noteq> 0" by (auto dest!: minus_unique)
+    from z K have "norm z < 1" by simp
+    hence "((\<lambda>z. f z * (1 + z) powr (-a)) has_field_derivative
+              f' z * (1 + z) powr (-a) - a * f z * (1 + z) powr (-a-1)) (at z)" using z
+      by (auto intro!: derivative_eq_intros)
+    also from z have "a * f z = (1 + z) * f' z" by (rule deriv)
+    also have "f' z * (1 + z) powr - a - (1 + z) * f' z * (1 + z) powr (- a - 1) = 0"
+      using \<open>norm z < 1\<close> by (auto simp add: field_simps powr_diff)
+    finally show "((\<lambda>z. f z * (1 + z) powr (-a)) has_field_derivative 0) (at z)" .
+  qed (insert K, auto)
+  also have "f 0 * (1 + 0) powr -a = 1" by simp
+  finally have "f z = (1 + z) powr a" using K
+    by (simp add: field_simps dist_real_def powr_minus)
+  with summable K show ?thesis unfolding f_f'_def by (simp add: sums_iff)
+qed
+
+lemma convergent_powser'_gbinomial:
+  "convergent_powser' (gbinomial_series abort p) (\<lambda>x. (1 + x) powr p)"
+proof -
+  have "(\<lambda>n. lcoeff (gbinomial_series abort p) n * x ^ n) sums (1 + x) powr p" if "abs x < 1" for x
+    using that gen_binomial_real[of x p] by simp
+  thus ?thesis unfolding convergent_powser'_def by (force intro!: exI[of _ 1])
+qed
+
+lemma convergent_powser'_gbinomial':
+  "convergent_powser' (gbinomial_series abort (real n)) (\<lambda>x. (1 + x) ^ n)"
+proof -
+  {
+    fix x :: real
+    have "(\<lambda>k. if k \<in> {0..n} then real (n choose k) * x ^ k else 0) sums (x + 1) ^ n"
+      using sums_If_finite_set[of "{..n}" "\<lambda>k. real (n choose k) * x ^ k"]
+      by (subst binomial_ring) simp
+    also have "(\<lambda>k. if k \<in> {0..n} then real (n choose k) * x ^ k else 0) = 
+                 (\<lambda>k. lcoeff (gbinomial_series abort (real n)) k * x ^ k)"
+      by (simp add: fun_eq_iff binomial_gbinomial [symmetric])
+    finally have "\<dots> sums (1 + x) ^ n" by (simp add: add_ac)
+  }
+  thus ?thesis unfolding convergent_powser'_def
+    by (auto intro!: exI[of _ 1])
+qed
+
+
+subsubsection \<open>Sine and cosine\<close>
+
+primcorec sin_series_stream_aux :: "bool \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real msstream" where
+  "sin_series_stream_aux b acc m = 
+    MSSCons (if b then -inverse acc else inverse acc)
+          (sin_series_stream_aux (\<not>b) (acc * (m + 1) * (m + 2)) (m + 2))"
+
+lemma mssnth_sin_series_stream_aux:
+  "mssnth (sin_series_stream_aux b (fact m) m) n = 
+     (if b then -1 else 1) * (-1) ^ n / (fact (2 * n + m))"
+proof (induction n arbitrary: b m) 
+  case (0 b m)
+  thus ?case by (subst sin_series_stream_aux.code) (simp add: field_simps)
+next
+  case (Suc n b m)
+  show ?case using Suc.IH[of "\<not>b" "m + 2"]
+    by (subst sin_series_stream_aux.code) (auto simp: field_simps)
+qed
+
+definition sin_series_stream where
+  "sin_series_stream = sin_series_stream_aux False 1 1"
+
+lemma mssnth_sin_series_stream:
+  "mssnth sin_series_stream n = (-1) ^ n / fact (2 * n + 1)"
+  using mssnth_sin_series_stream_aux[of False 1 n] unfolding sin_series_stream_def by simp
+
+definition cos_series_stream where
+  "cos_series_stream = sin_series_stream_aux False 1 0"
+
+lemma mssnth_cos_series_stream:
+  "mssnth cos_series_stream n = (-1) ^ n / fact (2 * n)"
+  using mssnth_sin_series_stream_aux[of False 0 n] unfolding cos_series_stream_def by simp
+
+lemma summable_pre_sin: "summable (\<lambda>n. mssnth sin_series_stream n * (x::real) ^ n)"
+proof -
+  have *: "summable (\<lambda>n. 1 / fact n * abs x ^ n)"
+    using exp_converges[of "abs x"] by (simp add: sums_iff field_simps)
+  {
+    fix n :: nat
+    have "\<bar>x\<bar> ^ n / fact (2 * n + 1) \<le> \<bar>x\<bar> ^ n / fact n"
+      by (intro divide_left_mono fact_mono) auto
+    hence "norm (mssnth sin_series_stream n * x ^ n) \<le> 1 / fact n * abs x ^ n"
+      by (simp add: mssnth_sin_series_stream abs_mult power_abs field_simps)
+  }
+  thus ?thesis
+    by (intro summable_comparison_test[OF _ *]) (auto intro!: exI[of _ 0])
+qed
+
+lemma summable_pre_cos: "summable (\<lambda>n. mssnth cos_series_stream n * (x::real) ^ n)"
+proof -
+  have *: "summable (\<lambda>n. 1 / fact n * abs x ^ n)"
+    using exp_converges[of "abs x"] by (simp add: sums_iff field_simps)
+  {
+    fix n :: nat
+    have "\<bar>x\<bar> ^ n / fact (2 * n) \<le> \<bar>x\<bar> ^ n / fact n"
+      by (intro divide_left_mono fact_mono) auto
+    hence "norm (mssnth cos_series_stream n * x ^ n) \<le> 1 / fact n * abs x ^ n"
+      by (simp add: mssnth_cos_series_stream abs_mult power_abs field_simps)
+  }
+  thus ?thesis
+    by (intro summable_comparison_test[OF _ *]) (auto intro!: exI[of _ 0])
+qed
+
+lemma cos_conv_pre_cos:
+  "cos x = powser (msllist_of_msstream cos_series_stream) (x ^ 2)"
+proof -
+  have "(\<lambda>n. cos_coeff (2 * n) * x ^ (2 * n)) sums cos x"
+    using cos_converges[of x]
+    by (subst sums_mono_reindex[of "\<lambda>n. 2 * n"])
+       (auto simp: strict_mono_def cos_coeff_def elim!: evenE)
+  also have "(\<lambda>n. cos_coeff (2 * n) * x ^ (2 * n)) =
+               (\<lambda>n. mssnth cos_series_stream n * (x ^ 2) ^ n)"
+    by (simp add: fun_eq_iff mssnth_cos_series_stream cos_coeff_def power_mult)
+  finally have sums: "(\<lambda>n. mssnth cos_series_stream n * x\<^sup>2 ^ n) sums cos x" .
+  thus ?thesis by (simp add: powser_def sums_iff)
+qed
+
+lemma sin_conv_pre_sin:
+  "sin x = x * powser (msllist_of_msstream sin_series_stream) (x ^ 2)"
+proof -
+  have "(\<lambda>n. sin_coeff (2 * n + 1) * x ^ (2 * n + 1)) sums sin x"
+    using sin_converges[of x]
+    by (subst sums_mono_reindex[of "\<lambda>n. 2 * n + 1"])
+       (auto simp: strict_mono_def sin_coeff_def elim!: oddE)
+  also have "(\<lambda>n. sin_coeff (2 * n + 1) * x ^ (2 * n + 1)) =
+               (\<lambda>n. x * mssnth sin_series_stream n * (x ^ 2) ^ n)"
+    by (simp add: fun_eq_iff mssnth_sin_series_stream sin_coeff_def power_mult [symmetric] algebra_simps)
+  finally have sums: "(\<lambda>n. x * mssnth sin_series_stream n * x\<^sup>2 ^ n) sums sin x" .
+  thus ?thesis using summable_pre_sin[of "x^2"] 
+    by (simp add: powser_def sums_iff suminf_mult [symmetric] mult.assoc)
+qed
+
+lemma convergent_powser_sin_series_stream:
+  "convergent_powser (msllist_of_msstream sin_series_stream)"
+  (is "convergent_powser ?cs")
+proof -
+  show ?thesis using summable_pre_sin unfolding convergent_powser_def
+    by (intro exI[of _ 1]) auto
+qed
+
+lemma convergent_powser_cos_series_stream:
+  "convergent_powser (msllist_of_msstream cos_series_stream)"
+  (is "convergent_powser ?cs")
+proof -
+  show ?thesis using summable_pre_cos unfolding convergent_powser_def
+    by (intro exI[of _ 1]) auto
+qed
+  
+  
+subsubsection \<open>Arc tangent\<close>
+
+definition arctan_coeffs :: "nat \<Rightarrow> 'a :: {real_normed_div_algebra, banach}" where
+  "arctan_coeffs n = (if odd n then (-1) ^ (n div 2) / of_nat n else 0)"
+
+lemma arctan_converges:
+  assumes "norm x < 1"
+  shows   "(\<lambda>n. arctan_coeffs n * x ^ n) sums arctan x"
+proof -
+  have A: "(\<lambda>n. diffs arctan_coeffs n * x ^ n) sums (1 / (1 + x^2))" if "norm x < 1" for x :: real
+  proof -
+    let ?f = "\<lambda>n. (if odd n then 0 else (-1) ^ (n div 2)) * x ^ n"
+    from that have "norm x ^ 2 < 1 ^ 2" by (intro power_strict_mono) simp_all
+    hence "(\<lambda>n. (-(x^2))^n) sums (1 / (1 - (-(x^2))))" by (intro geometric_sums) simp_all
+    also have "1 - (-(x^2)) = 1 + x^2" by simp
+    also have "(\<lambda>n. (-(x^2))^n) = (\<lambda>n. ?f (2*n))" by (auto simp: fun_eq_iff power_minus' power_mult)
+    also have "range (( *) (2::nat)) = {n. even n}"
+      by (auto elim!: evenE)
+    hence "(\<lambda>n. ?f (2*n)) sums (1 / (1 + x^2)) \<longleftrightarrow> ?f sums (1 / (1 + x^2))"
+      by (intro sums_mono_reindex) (simp_all add: strict_mono_Suc_iff)
+    also have "?f = (\<lambda>n. diffs arctan_coeffs n * x ^ n)"
+      by (simp add: fun_eq_iff arctan_coeffs_def diffs_def)
+    finally show "(\<lambda>n. diffs arctan_coeffs n * x ^ n) sums (1 / (1 + x^2))" .
+  qed
+  
+  have B: "summable (\<lambda>n. arctan_coeffs n * x ^ n)" if "norm x < 1" for x :: real
+  proof (rule summable_comparison_test)
+    show "\<exists>N. \<forall>n\<ge>N. norm (arctan_coeffs n * x ^ n) \<le> 1 * norm x ^ n"
+      unfolding norm_mult norm_power
+      by (intro exI[of _ "0::nat"] allI impI mult_right_mono)
+         (simp_all add: arctan_coeffs_def divide_simps)
+    from that show "summable (\<lambda>n. 1 * norm x ^ n)"
+      by (intro summable_mult summable_geometric) simp_all
+  qed
+  
+  define F :: "real \<Rightarrow> real" where "F = (\<lambda>x. \<Sum>n. arctan_coeffs n * x ^ n)"
+  have [derivative_intros]:
+    "(F has_field_derivative (1 / (1 + x^2))) (at x)" if "norm x < 1" for x :: real
+  proof -
+    define K :: real where "K = (1 + norm x) / 2"
+    from that have K: "norm x < K" "K < 1" by (simp_all add: K_def field_simps)
+    have "(F has_field_derivative (\<Sum>n. diffs arctan_coeffs n * x ^ n)) (at x)"
+      unfolding F_def using K by (intro termdiffs_strong[where K = K] B) simp_all
+    also from A[OF that] have "(\<Sum>n. diffs arctan_coeffs n * x ^ n) = 1 / (1 + x^2)"
+      by (simp add: sums_iff)
+    finally show ?thesis .
+  qed
+  from assms have "arctan x - F x = arctan 0 - F 0"
+    by (intro DERIV_isconst3[of "-1" 1 _ _ "\<lambda>x. arctan x - F x"])
+       (auto intro!: derivative_eq_intros simp: field_simps)
+  hence "F x = arctan x" by (simp add: F_def arctan_coeffs_def)
+  with B[OF assms] show ?thesis by (simp add: sums_iff F_def)
+qed
+
+primcorec arctan_series_stream_aux :: "bool \<Rightarrow> real \<Rightarrow> real msstream" where
+  "arctan_series_stream_aux b n = 
+     MSSCons (if b then -inverse n else inverse n) (arctan_series_stream_aux (\<not>b) (n + 2))"
+
+lemma mssnth_arctan_series_stream_aux: 
+  "mssnth (arctan_series_stream_aux b n) m = (-1) ^ (if b then Suc m else m) / (2*m + n)"
+  by (induction m arbitrary: b n; subst arctan_series_stream_aux.code)
+     (auto simp: field_simps split: if_splits)
+
+definition arctan_series_stream where
+  "arctan_series_stream = arctan_series_stream_aux False 1"
+
+lemma mssnth_arctan_series_stream:
+  "mssnth arctan_series_stream n = (-1) ^ n / (2 * n + 1)"
+  by (simp add: arctan_series_stream_def mssnth_arctan_series_stream_aux)
+
+lemma summable_pre_arctan:
+  assumes "norm x < 1"
+  shows   "summable (\<lambda>n. mssnth arctan_series_stream n * x ^ n)" (is "summable ?f")
+proof (rule summable_comparison_test)
+  show "summable (\<lambda>n. abs x ^ n)" using assms by (intro summable_geometric) auto
+  show "\<exists>N. \<forall>n\<ge>N. norm (?f n) \<le> abs x ^ n"
+  proof (intro exI[of _ 0] allI impI)
+    fix n :: nat
+    have "norm (?f n) = \<bar>x\<bar> ^ n / (2 * real n + 1)"
+      by (simp add: mssnth_arctan_series_stream abs_mult power_abs)
+    also have "\<dots> \<le> \<bar>x\<bar> ^ n / 1" by (intro divide_left_mono) auto
+    finally show "norm (?f n) \<le> abs x ^ n" by simp
+  qed
+qed
+
+lemma arctan_conv_pre_arctan:
+  assumes "norm x < 1"
+  shows   "arctan x = x * powser (msllist_of_msstream arctan_series_stream) (x ^ 2)"
+proof -
+  from assms have "norm x * norm x < 1 * 1" by (intro mult_strict_mono) auto
+  hence norm_less: "norm (x ^ 2) < 1" by (simp add: power2_eq_square)
+  have "(\<lambda>n. arctan_coeffs n * x ^ n) sums arctan x"
+    by (intro arctan_converges assms)
+  also have "?this \<longleftrightarrow> (\<lambda>n. arctan_coeffs (2 * n + 1) * x ^ (2 * n + 1)) sums arctan x"
+    by (intro sums_mono_reindex [symmetric])
+       (auto simp: arctan_coeffs_def strict_mono_def elim!: oddE)
+  also have "(\<lambda>n. arctan_coeffs (2 * n + 1) * x ^ (2 * n + 1)) =
+               (\<lambda>n. x * mssnth arctan_series_stream n * (x ^ 2) ^ n)"
+    by (simp add: fun_eq_iff mssnth_arctan_series_stream 
+                  power_mult [symmetric] arctan_coeffs_def mult_ac)
+  finally have "(\<lambda>n. x * mssnth arctan_series_stream n * x\<^sup>2 ^ n) sums arctan x" .
+  thus ?thesis using summable_pre_arctan[OF norm_less] assms
+    by (simp add: powser_def sums_iff suminf_mult [symmetric] mult.assoc)
+qed
+
+lemma convergent_powser_arctan: 
+  "convergent_powser (msllist_of_msstream arctan_series_stream)"
+  unfolding convergent_powser_def
+  using summable_pre_arctan by (intro exI[of _ 1]) auto
+
+lemma arctan_inverse_pos: "x > 0 \<Longrightarrow> arctan x = pi / 2 - arctan (inverse x)"
+  using arctan_inverse[of x] by (simp add: field_simps)
+    
+lemma arctan_inverse_neg: "x < 0 \<Longrightarrow> arctan x = -pi / 2 - arctan (inverse x)"
+  using arctan_inverse[of x] by (simp add: field_simps)
+
+
+
+subsection \<open>Multiseries\<close>
+
+subsubsection \<open>Asymptotic bases\<close>
+
+type_synonym basis = "(real \<Rightarrow> real) list"
+
+lemma transp_ln_smallo_ln: "transp (\<lambda>f g. (\<lambda>x::real. ln (g x)) \<in> o(\<lambda>x. ln (f x)))"
+  by (rule transpI, erule landau_o.small.trans)
+
+definition basis_wf :: "basis \<Rightarrow> bool" where
+  "basis_wf basis \<longleftrightarrow> (\<forall>f\<in>set basis. filterlim f at_top at_top) \<and> 
+                      sorted_wrt (\<lambda>f g. (\<lambda>x. ln (g x)) \<in> o(\<lambda>x. ln (f x))) basis"
+
+lemma basis_wf_Nil [simp]: "basis_wf []"
+  by (simp add: basis_wf_def)
+
+lemma basis_wf_Cons: 
+  "basis_wf (f # basis) \<longleftrightarrow> 
+     filterlim f at_top at_top \<and> (\<forall>g\<in>set basis. (\<lambda>x. ln (g x)) \<in> o(\<lambda>x. ln (f x))) \<and> basis_wf basis"
+  unfolding basis_wf_def by auto
+
+(* TODO: Move *)
+lemma sorted_wrt_snoc:
+  assumes trans: "transp P" and "sorted_wrt P xs" "P (last xs) y"
+  shows   "sorted_wrt P (xs @ [y])"
+proof (subst sorted_wrt_append, intro conjI ballI)
+  fix x y' assume x: "x \<in> set xs" and y': "y' \<in> set [y]"
+  from y' have [simp]: "y' = y" by simp
+  show "P x y'"
+  proof (cases "x = last xs")
+    case False
+    from x have eq: "xs = butlast xs @ [last xs]"
+      by (subst append_butlast_last_id) auto
+    from x and False have x': "x \<in> set (butlast xs)" by (subst (asm) eq) auto
+    have "sorted_wrt P xs" by fact
+    also note eq
+    finally have "P x (last xs)" using x'
+      by (subst (asm) sorted_wrt_append) auto
+    with \<open>P (last xs) y\<close> have "P x y" using transpD[OF trans] by blast
+    thus ?thesis by simp
+  qed (insert assms, auto)
+qed (insert assms, auto)  
+
+lemma basis_wf_snoc:
+  assumes "bs \<noteq> []"
+  assumes "basis_wf bs" "filterlim b at_top at_top"
+  assumes "(\<lambda>x. ln (b x)) \<in> o(\<lambda>x. ln (last bs x))"
+  shows   "basis_wf (bs @ [b])"
+proof -
+  have "transp (\<lambda>f g. (\<lambda>x. ln (g x)) \<in> o(\<lambda>x. ln (f x)))"
+    by (auto simp: transp_def intro: landau_o.small.trans)
+  thus ?thesis using assms
+    by (auto simp add: basis_wf_def sorted_wrt_snoc[OF transp_ln_smallo_ln])
+qed
+
+lemma basis_wf_singleton: "basis_wf [b] \<longleftrightarrow> filterlim b at_top at_top"
+  by (simp add: basis_wf_Cons)
+
+lemma basis_wf_many: "basis_wf (b # b' # bs) \<longleftrightarrow>
+  filterlim b at_top at_top \<and> (\<lambda>x. ln (b' x)) \<in> o(\<lambda>x. ln (b x)) \<and> basis_wf (b' # bs)"
+  unfolding basis_wf_def by (subst sorted_wrt2[OF transp_ln_smallo_ln]) auto
+
+
+subsubsection \<open>Monomials\<close>
+
+type_synonym monom = "real \<times> real list"
+
+definition eval_monom :: "monom \<Rightarrow> basis \<Rightarrow> real \<Rightarrow> real" where
+  "eval_monom = (\<lambda>(c,es) basis x. c * prod_list (map (\<lambda>(b,e). b x powr e) (zip basis es)))"
+  
+lemma eval_monom_Nil1 [simp]: "eval_monom (c, []) basis = (\<lambda>_. c)"
+  by (simp add: eval_monom_def split: prod.split)
+
+lemma eval_monom_Nil2 [simp]: "eval_monom monom [] = (\<lambda>_. fst monom)"
+  by (simp add: eval_monom_def split: prod.split)
+
+lemma eval_monom_Cons: 
+  "eval_monom (c, e # es) (b # basis) = (\<lambda>x. eval_monom (c, es) basis x * b x powr e)"
+  by (simp add: eval_monom_def fun_eq_iff split: prod.split)
+
+definition inverse_monom :: "monom \<Rightarrow> monom" where
+  "inverse_monom monom = (case monom of (c, es) \<Rightarrow> (inverse c, map uminus es))"
+
+lemma length_snd_inverse_monom [simp]: 
+  "length (snd (inverse_monom monom)) = length (snd monom)"
+  by (simp add: inverse_monom_def split: prod.split)
+
+lemma eval_monom_pos:
+  assumes "basis_wf basis" "fst monom > 0"
+  shows   "eventually (\<lambda>x. eval_monom monom basis x > 0) at_top"
+proof (cases monom)
+  case (Pair c es)
+  with assms have "c > 0" by simp
+  with assms(1) show ?thesis unfolding Pair
+  proof (induction es arbitrary: basis)
+    case (Cons e es)
+    note A = this
+    thus ?case
+    proof (cases basis)
+      case (Cons b basis')
+      with A(1)[of basis'] A(2,3) 
+        have A: "\<forall>\<^sub>F x in at_top. eval_monom (c, es) basis' x > 0" 
+         and B: "eventually (\<lambda>x. b x > 0) at_top"
+        by (simp_all add: basis_wf_Cons filterlim_at_top_dense)
+      thus ?thesis by eventually_elim (simp add: Cons eval_monom_Cons)
+    qed simp
+  qed simp
+qed
+
+lemma eval_monom_uminus: "eval_monom (-c, es) basis x = -eval_monom (c, es) basis x"
+  by (simp add: eval_monom_def)
+
+lemma eval_monom_neg:
+  assumes "basis_wf basis" "fst monom < 0"
+  shows   "eventually (\<lambda>x. eval_monom monom basis x < 0) at_top"    
+proof -
+  from assms have "eventually (\<lambda>x. eval_monom (-fst monom, snd monom) basis x > 0) at_top"
+    by (intro eval_monom_pos) simp_all
+  thus ?thesis by (simp add: eval_monom_uminus)
+qed
+
+lemma eval_monom_nonzero:
+  assumes "basis_wf basis" "fst monom \<noteq> 0"
+  shows   "eventually (\<lambda>x. eval_monom monom basis x \<noteq> 0) at_top"
+proof (cases "fst monom" "0 :: real" rule: linorder_cases)
+  case greater
+  with assms have "eventually (\<lambda>x. eval_monom monom basis x > 0) at_top"
+    by (simp add: eval_monom_pos)
+  thus ?thesis by eventually_elim simp
+next
+  case less
+  with assms have "eventually (\<lambda>x. eval_monom (-fst monom, snd monom) basis x > 0) at_top"
+    by (simp add: eval_monom_pos)
+  thus ?thesis by eventually_elim (simp add: eval_monom_uminus)
+qed (insert assms, simp_all)
+
+
+subsubsection \<open>Typeclass for multiseries\<close>
+
+class multiseries = plus + minus + times + uminus + inverse + 
+  fixes is_expansion :: "'a \<Rightarrow> basis \<Rightarrow> bool"
+    and expansion_level :: "'a itself \<Rightarrow> nat"
+    and eval :: "'a \<Rightarrow> real \<Rightarrow> real"
+    and zero_expansion :: 'a
+    and const_expansion :: "real \<Rightarrow> 'a"
+    and powr_expansion :: "bool \<Rightarrow> 'a \<Rightarrow> real \<Rightarrow> 'a"
+    and power_expansion :: "bool \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a"
+    and trimmed :: "'a \<Rightarrow> bool"
+    and dominant_term :: "'a \<Rightarrow> monom"
+
+  assumes is_expansion_length:
+    "is_expansion F basis \<Longrightarrow> length basis = expansion_level (TYPE('a))"
+  assumes is_expansion_zero:
+    "basis_wf basis \<Longrightarrow> length basis = expansion_level (TYPE('a)) \<Longrightarrow> 
+       is_expansion zero_expansion basis"
+  assumes is_expansion_const:
+    "basis_wf basis \<Longrightarrow> length basis = expansion_level (TYPE('a)) \<Longrightarrow> 
+       is_expansion (const_expansion c) basis"
+  assumes is_expansion_uminus:
+    "basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> is_expansion (-F) basis"
+  assumes is_expansion_add: 
+    "basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> is_expansion G basis \<Longrightarrow> 
+       is_expansion (F + G) basis"
+  assumes is_expansion_minus: 
+    "basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> is_expansion G basis \<Longrightarrow> 
+       is_expansion (F - G) basis"
+  assumes is_expansion_mult: 
+    "basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> is_expansion G basis \<Longrightarrow> 
+       is_expansion (F * G) basis"
+  assumes is_expansion_inverse:
+    "basis_wf basis \<Longrightarrow> trimmed F \<Longrightarrow> is_expansion F basis \<Longrightarrow> 
+       is_expansion (inverse F) basis"
+  assumes is_expansion_divide:
+    "basis_wf basis \<Longrightarrow> trimmed G \<Longrightarrow> is_expansion F basis \<Longrightarrow> is_expansion G basis \<Longrightarrow> 
+       is_expansion (F / G) basis"
+  assumes is_expansion_powr:
+    "basis_wf basis \<Longrightarrow> trimmed F \<Longrightarrow> fst (dominant_term F) > 0 \<Longrightarrow> is_expansion F basis \<Longrightarrow>
+       is_expansion (powr_expansion abort F p) basis"
+  assumes is_expansion_power:
+    "basis_wf basis \<Longrightarrow> trimmed F \<Longrightarrow> is_expansion F basis \<Longrightarrow>
+       is_expansion (power_expansion abort F n) basis"
+  
+  assumes is_expansion_imp_smallo:
+    "basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> filterlim b at_top at_top \<Longrightarrow>
+       (\<forall>g\<in>set basis. (\<lambda>x. ln (g x)) \<in> o(\<lambda>x. ln (b x))) \<Longrightarrow> e > 0 \<Longrightarrow> eval F \<in> o(\<lambda>x. b x powr e)"
+  assumes is_expansion_imp_smallomega:
+    "basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> filterlim b at_top at_top \<Longrightarrow> trimmed F \<Longrightarrow> 
+       (\<forall>g\<in>set basis. (\<lambda>x. ln (g x)) \<in> o(\<lambda>x. ln (b x))) \<Longrightarrow> e < 0 \<Longrightarrow> eval F \<in> \<omega>(\<lambda>x. b x powr e)"
+  assumes trimmed_imp_eventually_sgn:
+    "basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> trimmed F \<Longrightarrow>
+       eventually (\<lambda>x. sgn (eval F x) = sgn (fst (dominant_term F))) at_top"
+  assumes trimmed_imp_eventually_nz: 
+    "basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> trimmed F \<Longrightarrow> 
+       eventually (\<lambda>x. eval F x \<noteq> 0) at_top"
+  assumes trimmed_imp_dominant_term_nz: "trimmed F \<Longrightarrow> fst (dominant_term F) \<noteq> 0"
+  
+  assumes dominant_term: 
+    "basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> trimmed F \<Longrightarrow>
+       eval F \<sim>[at_top] eval_monom (dominant_term F) basis"
+  assumes dominant_term_bigo:
+    "basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow>
+       eval F \<in> O(eval_monom (1, snd (dominant_term F)) basis)"
+  assumes length_dominant_term:
+    "length (snd (dominant_term F)) = expansion_level TYPE('a)"
+  assumes fst_dominant_term_uminus [simp]: "fst (dominant_term (- F)) = -fst (dominant_term F)"
+  assumes trimmed_uminus_iff [simp]: "trimmed (-F) \<longleftrightarrow> trimmed F"
+  
+  assumes add_zero_expansion_left [simp]: "zero_expansion + F = F"
+  assumes add_zero_expansion_right [simp]: "F + zero_expansion = F"
+  
+  assumes eval_zero [simp]: "eval zero_expansion x = 0"
+  assumes eval_const [simp]: "eval (const_expansion c) x = c"
+  assumes eval_uminus [simp]: "eval (-F) = (\<lambda>x. -eval F x)"
+  assumes eval_plus [simp]: "eval (F + G) = (\<lambda>x. eval F x + eval G x)"
+  assumes eval_minus [simp]: "eval (F - G) = (\<lambda>x. eval F x - eval G x)"
+  assumes eval_times [simp]: "eval (F * G) = (\<lambda>x. eval F x * eval G x)"
+  assumes eval_inverse [simp]: "eval (inverse F) = (\<lambda>x. inverse (eval F x))"
+  assumes eval_divide [simp]: "eval (F / G) = (\<lambda>x. eval F x / eval G x)"
+  assumes eval_powr [simp]: "eval (powr_expansion abort F p) = (\<lambda>x. eval F x powr p)"
+  assumes eval_power [simp]: "eval (power_expansion abort F n) = (\<lambda>x. eval F x ^ n)"
+  assumes minus_eq_plus_uminus: "F - G = F + (-G)"
+  assumes times_const_expansion_1: "const_expansion 1 * F = F"
+  assumes trimmed_const_expansion: "trimmed (const_expansion c) \<longleftrightarrow> c \<noteq> 0"
+begin
+
+definition trimmed_pos where
+  "trimmed_pos F \<longleftrightarrow> trimmed F \<and> fst (dominant_term F) > 0"
+
+definition trimmed_neg where
+  "trimmed_neg F \<longleftrightarrow> trimmed F \<and> fst (dominant_term F) < 0"
+
+lemma trimmed_pos_uminus: "trimmed_neg F \<Longrightarrow> trimmed_pos (-F)"
+  by (simp add: trimmed_neg_def trimmed_pos_def)
+
+lemma trimmed_pos_imp_trimmed: "trimmed_pos x \<Longrightarrow> trimmed x"
+  by (simp add: trimmed_pos_def)
+
+lemma trimmed_neg_imp_trimmed: "trimmed_neg x \<Longrightarrow> trimmed x"
+  by (simp add: trimmed_neg_def)
+
+end
+
+
+subsubsection \<open>Zero-rank expansions\<close>
+
+instantiation real :: multiseries
+begin
+
+inductive is_expansion_real :: "real \<Rightarrow> basis \<Rightarrow> bool" where
+  "is_expansion_real c []"
+  
+definition expansion_level_real :: "real itself \<Rightarrow> nat" where
+  expansion_level_real_def [simp]: "expansion_level_real _ = 0"
+
+definition eval_real :: "real \<Rightarrow> real \<Rightarrow> real" where
+  eval_real_def [simp]: "eval_real c = (\<lambda>_. c)"
+
+definition zero_expansion_real :: "real" where
+  zero_expansion_real_def [simp]: "zero_expansion_real = 0"
+  
+definition const_expansion_real :: "real \<Rightarrow> real" where
+  const_expansion_real_def [simp]: "const_expansion_real x = x"
+
+definition powr_expansion_real :: "bool \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
+  powr_expansion_real_def [simp]: "powr_expansion_real _ x p = x powr p"
+
+definition power_expansion_real :: "bool \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real" where
+  power_expansion_real_def [simp]: "power_expansion_real _ x n = x ^ n"
+
+definition trimmed_real :: "real \<Rightarrow> bool" where
+  trimmed_real_def [simp]: "trimmed_real x \<longleftrightarrow> x \<noteq> 0"
+
+definition dominant_term_real :: "real \<Rightarrow> monom" where
+  dominant_term_real_def [simp]: "dominant_term_real c = (c, [])"
+
+instance proof
+  fix basis :: basis and b :: "real \<Rightarrow> real" and e F :: real
+  assume *: "basis_wf basis" "filterlim b at_top at_top" "is_expansion F basis" "0 < e"
+  have "(\<lambda>x. b x powr e) \<in> \<omega>(\<lambda>_. 1)"
+    by (intro smallomegaI_filterlim_at_infinity filterlim_at_top_imp_at_infinity) 
+       (auto intro!: filterlim_compose[OF real_powr_at_top] * )
+  with * show "eval F \<in> o(\<lambda>x. b x powr e)"
+    by (cases "F = 0") (auto elim!: is_expansion_real.cases simp: smallomega_iff_smallo)
+next
+  fix basis :: basis and b :: "real \<Rightarrow> real" and e F :: real
+  assume *: "basis_wf basis" "filterlim b at_top at_top" "is_expansion F basis" 
+            "e < 0" "trimmed F"
+  from * have pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+  have "(\<lambda>x. b x powr -e) \<in> \<omega>(\<lambda>_. 1)"
+    by (intro smallomegaI_filterlim_at_infinity filterlim_at_top_imp_at_infinity) 
+       (auto intro!: filterlim_compose[OF real_powr_at_top] * )
+  with pos have "(\<lambda>x. b x powr e) \<in> o(\<lambda>_. 1)" unfolding powr_minus
+    by (subst (asm) landau_omega.small.inverse_eq2)
+       (auto elim: eventually_mono simp: smallomega_iff_smallo)
+  with * show "eval F \<in> \<omega>(\<lambda>x. b x powr e)"
+    by (auto elim!: is_expansion_real.cases simp: smallomega_iff_smallo)
+qed (auto intro!: is_expansion_real.intros elim!: is_expansion_real.cases)
+
+end
+
+lemma eval_real: "eval (c :: real) x = c" by simp
+
+
+subsubsection \<open>Operations on multiseries\<close>
+
+lemma ms_aux_cases [case_names MSLNil MSLCons]:
+  fixes xs :: "('a \<times> real) msllist"
+  obtains "xs = MSLNil" | c e xs' where "xs = MSLCons (c, e) xs'"
+proof (cases xs)
+  case (MSLCons x xs')
+  with that(2)[of "fst x" "snd x" xs'] show ?thesis by auto
+qed auto
+
+definition ms_aux_hd_exp :: "('a \<times> real) msllist \<Rightarrow> real option" where
+  "ms_aux_hd_exp xs = (case xs of MSLNil \<Rightarrow> None | MSLCons (_, e) _ \<Rightarrow> Some e)"
+
+primrec ms_exp_gt :: "real \<Rightarrow> real option \<Rightarrow> bool" where
+  "ms_exp_gt x None = True"
+| "ms_exp_gt x (Some y) \<longleftrightarrow> x > y"
+
+lemma ms_aux_hd_exp_MSLNil [simp]: "ms_aux_hd_exp MSLNil = None"
+  by (simp add: ms_aux_hd_exp_def split: prod.split)
+  
+lemma ms_aux_hd_exp_MSLCons [simp]: "ms_aux_hd_exp (MSLCons x xs) = Some (snd x)"
+  by (simp add: ms_aux_hd_exp_def split: prod.split)
+
+coinductive is_expansion_aux :: 
+    "('a :: multiseries \<times> real) msllist \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> basis \<Rightarrow> bool" where
+  is_expansion_MSLNil: 
+    "eventually (\<lambda>x. f x = 0) at_top \<Longrightarrow> length basis = Suc (expansion_level TYPE('a)) \<Longrightarrow>
+       is_expansion_aux MSLNil f basis"
+| is_expansion_MSLCons: 
+    "is_expansion_aux xs (\<lambda>x. f x - eval C x * b x powr e) (b # basis) \<Longrightarrow>
+       is_expansion C basis \<Longrightarrow>
+       (\<And>e'. e' > e \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')) \<Longrightarrow> ms_exp_gt e (ms_aux_hd_exp xs) \<Longrightarrow>
+       is_expansion_aux (MSLCons (C, e) xs) f (b # basis)"
+
+inductive_cases is_expansion_aux_MSLCons: "is_expansion_aux (MSLCons (c, e) xs) f basis"
+ 
+lemma is_expansion_aux_basis_nonempty: "is_expansion_aux F f basis \<Longrightarrow> basis \<noteq> []"
+  by (erule is_expansion_aux.cases) auto
+
+lemma is_expansion_aux_expansion_level:
+  assumes "is_expansion_aux (G :: ('a::multiseries \<times> real) msllist) g basis"
+  shows   "length basis = Suc (expansion_level TYPE('a))"
+  using assms by (cases rule: is_expansion_aux.cases) (auto dest: is_expansion_length)  
+
+lemma is_expansion_aux_imp_smallo:
+  assumes "is_expansion_aux xs f basis" "ms_exp_gt p (ms_aux_hd_exp xs)" 
+  shows   "f \<in> o(\<lambda>x. hd basis x powr p)"
+  using assms
+proof (cases rule: is_expansion_aux.cases)
+  case is_expansion_MSLNil
+  show ?thesis by (simp add: landau_o.small.in_cong[OF is_expansion_MSLNil(2)])
+next
+  case (is_expansion_MSLCons xs C b e basis)
+  with assms have "f \<in> o(\<lambda>x. b x powr p)"
+    by (intro is_expansion_MSLCons) (simp_all add: ms_aux_hd_exp_def)
+  thus ?thesis by (simp add: is_expansion_MSLCons)
+qed
+
+lemma is_expansion_aux_imp_smallo':
+  assumes "is_expansion_aux xs f basis"
+  obtains e where "f \<in> o(\<lambda>x. hd basis x powr e)"
+proof -
+  define e where "e = (case ms_aux_hd_exp xs of None \<Rightarrow> 0 | Some e \<Rightarrow> e + 1)"
+  have "ms_exp_gt e (ms_aux_hd_exp xs)"
+    by (auto simp add: e_def ms_aux_hd_exp_def split: msllist.splits)
+  from assms this have "f \<in> o(\<lambda>x. hd basis x powr e)" by (rule is_expansion_aux_imp_smallo)
+  from that[OF this] show ?thesis .
+qed
+
+lemma is_expansion_aux_imp_smallo'':
+  assumes "is_expansion_aux xs f basis" "ms_exp_gt e' (ms_aux_hd_exp xs)"
+  obtains e where "e < e'" "f \<in> o(\<lambda>x. hd basis x powr e)"
+proof -
+  define e where "e = (case ms_aux_hd_exp xs of None \<Rightarrow> e' - 1 | Some e \<Rightarrow> (e' + e) / 2)"
+  from assms have "ms_exp_gt e (ms_aux_hd_exp xs)" "e < e'"
+    by (cases xs; simp add: e_def field_simps)+
+  from assms(1) this(1) have "f \<in> o(\<lambda>x. hd basis x powr e)" by (rule is_expansion_aux_imp_smallo)
+  from that[OF \<open>e < e'\<close> this] show ?thesis .
+qed
+
+definition trimmed_ms_aux :: "('a :: multiseries \<times> real) msllist \<Rightarrow> bool" where
+  "trimmed_ms_aux xs = (case xs of MSLNil \<Rightarrow> False | MSLCons (C, _) _ \<Rightarrow> trimmed C)"
+ 
+lemma not_trimmed_ms_aux_MSLNil [simp]: "\<not>trimmed_ms_aux MSLNil"
+  by (simp add: trimmed_ms_aux_def)
+
+lemma trimmed_ms_aux_MSLCons: "trimmed_ms_aux (MSLCons x xs) = trimmed (fst x)"
+  by (simp add: trimmed_ms_aux_def split: prod.split)
+
+lemma trimmed_ms_aux_imp_nz:
+  assumes "basis_wf basis" "is_expansion_aux xs f basis" "trimmed_ms_aux xs"
+  shows   "eventually (\<lambda>x. f x \<noteq> 0) at_top"
+proof (cases xs rule: ms_aux_cases)
+  case (MSLCons C e xs')
+  from assms this obtain b basis' where *: "basis = b # basis'"
+    "is_expansion_aux xs' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')" 
+    "ms_exp_gt e (ms_aux_hd_exp xs')" "is_expansion C basis'" "\<And>e'. e' > e \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')"
+    by (auto elim!: is_expansion_aux_MSLCons)
+  from assms(1,3) * have nz: "eventually (\<lambda>x. eval C x \<noteq> 0) at_top"
+    by (auto simp: trimmed_ms_aux_def MSLCons basis_wf_Cons
+             intro: trimmed_imp_eventually_nz[of basis'])
+  from assms(1) * have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_nz: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+  
+  from is_expansion_aux_imp_smallo''[OF *(2,3)]
+  obtain e' where e': "e' < e" "(\<lambda>x. f x - eval C x * b x powr e) \<in> o(\<lambda>x. b x powr e')"
+    unfolding list.sel by blast
+  note this(2)
+  also have "\<dots> = o(\<lambda>x. b x powr (e' - e) * b x powr e)" by (simp add: powr_add [symmetric])
+  also from assms * e' have "eval C \<in> \<omega>(\<lambda>x. b x powr (e' - e))"
+    by (intro is_expansion_imp_smallomega[OF _ *(4)])
+       (simp_all add: MSLCons basis_wf_Cons trimmed_ms_aux_def)
+  hence "(\<lambda>x. b x powr (e' - e) * b x powr e) \<in> o(\<lambda>x. eval C x * b x powr e)"
+    by (intro landau_o.small_big_mult is_expansion_imp_smallomega) 
+       (simp_all add: smallomega_iff_smallo)
+  finally have "(\<lambda>x. f x - eval C x * b x powr e + eval C x * b x powr e) \<in> 
+                  \<Theta>(\<lambda>x. eval C x * b x powr e)"
+    by (subst bigtheta_sym, subst landau_theta.plus_absorb1) simp_all
+  hence theta: "f \<in> \<Theta>(\<lambda>x. eval C x * b x powr e)" by simp
+  have "eventually (\<lambda>x. eval C x * b x powr e \<noteq> 0) at_top"
+    using b_nz nz by eventually_elim simp
+  thus ?thesis by (subst eventually_nonzero_bigtheta [OF theta])
+qed (insert assms, simp_all add: trimmed_ms_aux_def)
+
+lemma is_expansion_aux_imp_smallomega:
+  assumes "basis_wf basis" "is_expansion_aux xs f basis" "filterlim b' at_top at_top"
+          "trimmed_ms_aux xs" "\<forall>g\<in>set basis. (\<lambda>x. ln (g x)) \<in> o(\<lambda>x. ln (b' x))" "r < 0"
+  shows   "f \<in> \<omega>(\<lambda>x. b' x powr r)"
+proof (cases xs rule: ms_aux_cases)
+  case (MSLCons C e xs')
+  from assms this obtain b basis' where *: "basis = b # basis'" "trimmed C"
+    "is_expansion_aux xs' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
+    "ms_exp_gt e (ms_aux_hd_exp xs')"
+    "is_expansion C basis'" "\<And>e'. e' > e \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')"
+    by (auto elim!: is_expansion_aux_MSLCons simp: trimmed_ms_aux_def)
+  from assms * have nz: "eventually (\<lambda>x. eval C x \<noteq> 0) at_top"
+    by (intro trimmed_imp_eventually_nz[of basis']) (simp_all add: basis_wf_Cons)
+  from assms(1) * have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_nz: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+  
+  from is_expansion_aux_imp_smallo''[OF *(3,4)]
+  obtain e' where e': "e' < e" "(\<lambda>x. f x - eval C x * b x powr e) \<in> o(\<lambda>x. b x powr e')"
+    unfolding list.sel by blast
+  note this(2)
+  also have "\<dots> = o(\<lambda>x. b x powr (e' - e) * b x powr e)" by (simp add: powr_add [symmetric])
+  also from assms * e' have "eval C \<in> \<omega>(\<lambda>x. b x powr (e' - e))"
+    by (intro is_expansion_imp_smallomega[OF _ *(5)])
+       (simp_all add: MSLCons basis_wf_Cons trimmed_ms_aux_def)
+  hence "(\<lambda>x. b x powr (e' - e) * b x powr e) \<in> o(\<lambda>x. eval C x * b x powr e)"
+    by (intro landau_o.small_big_mult is_expansion_imp_smallomega) 
+       (simp_all add: smallomega_iff_smallo)
+  finally have "(\<lambda>x. f x - eval C x * b x powr e + eval C x * b x powr e) \<in> 
+                  \<Theta>(\<lambda>x. eval C x * b x powr e)"
+    by (subst bigtheta_sym, subst landau_theta.plus_absorb1) simp_all
+  hence theta: "f \<in> \<Theta>(\<lambda>x. eval C x * b x powr e)" by simp
+  also from * assms e' have "(\<lambda>x. eval C x * b x powr e) \<in> \<omega>(\<lambda>x. b x powr (e' - e) * b x powr e)"
+    by (intro landau_omega.small_big_mult is_expansion_imp_smallomega[of basis'])
+       (simp_all add: basis_wf_Cons)
+  also have "\<dots> = \<omega>(\<lambda>x. b x powr e')" by (simp add: powr_add [symmetric])
+  also from *(1) assms have "(\<lambda>x. b x powr e') \<in> \<omega>(\<lambda>x. b' x powr r)"
+    unfolding smallomega_iff_smallo by (intro ln_smallo_imp_flat') (simp_all add: basis_wf_Cons)
+  finally show ?thesis .
+qed (insert assms, simp_all add: trimmed_ms_aux_def)  
+
+definition dominant_term_ms_aux :: "('a :: multiseries \<times> real) msllist \<Rightarrow> monom" where
+  "dominant_term_ms_aux xs =
+     (case xs of MSLNil \<Rightarrow> (0, replicate (Suc (expansion_level TYPE('a))) 0)
+               | MSLCons (C, e) _ \<Rightarrow> case dominant_term C of (c, es) \<Rightarrow> (c, e # es))"
+
+lemma dominant_term_ms_aux_MSLCons: 
+  "dominant_term_ms_aux (MSLCons (C, e) xs) = map_prod id (\<lambda>es. e # es) (dominant_term C)"
+  by (simp add: dominant_term_ms_aux_def case_prod_unfold map_prod_def)
+
+lemma length_dominant_term_ms_aux [simp]:
+  "length (snd (dominant_term_ms_aux (xs :: ('a :: multiseries \<times> real) msllist))) = 
+     Suc (expansion_level TYPE('a))"
+proof (cases xs rule: ms_aux_cases)
+  case (MSLCons C e xs')
+  hence "length (snd (dominant_term_ms_aux xs)) = Suc (length (snd (dominant_term C)))"
+    by (simp add: dominant_term_ms_aux_def split: prod.splits)
+  also note length_dominant_term
+  finally show ?thesis .
+qed (simp_all add: dominant_term_ms_aux_def split: msllist.splits prod.splits)
+
+definition const_ms_aux :: "real \<Rightarrow> ('a :: multiseries \<times> real) msllist" where
+  "const_ms_aux c = MSLCons (const_expansion c, 0) MSLNil"
+
+definition uminus_ms_aux :: "('a :: uminus \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
+  "uminus_ms_aux xs = mslmap (map_prod uminus id) xs"
+  
+corec (friend) plus_ms_aux :: "('a :: plus \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
+  "plus_ms_aux xs ys =
+     (case (xs, ys) of
+        (MSLNil, MSLNil) \<Rightarrow> MSLNil
+      | (MSLNil, MSLCons y ys) \<Rightarrow> MSLCons y ys
+      | (MSLCons x xs, MSLNil) \<Rightarrow> MSLCons x xs
+      | (MSLCons x xs, MSLCons y ys) \<Rightarrow>
+          if snd x > snd y then MSLCons x (plus_ms_aux xs (MSLCons y ys))
+          else if snd x < snd y then MSLCons y (plus_ms_aux (MSLCons x xs) ys)
+          else MSLCons (fst x + fst y, snd x) (plus_ms_aux xs ys))"
+
+context
+begin
+
+friend_of_corec mslmap where
+  "mslmap (f :: 'a \<Rightarrow> 'a) xs = 
+     (case xs of MSLNil \<Rightarrow> MSLNil | MSLCons x xs \<Rightarrow> MSLCons (f x) (mslmap f xs))"
+   by (simp split: msllist.splits, transfer_prover)
+
+corec (friend) times_ms_aux
+     :: "('a :: {plus,times} \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
+  "times_ms_aux xs ys =
+     (case (xs, ys) of
+        (MSLNil, ys) \<Rightarrow> MSLNil
+      | (xs, MSLNil) \<Rightarrow> MSLNil
+      | (MSLCons x xs, MSLCons y ys) \<Rightarrow>
+           MSLCons (fst x * fst y, snd x + snd y) 
+             (plus_ms_aux (mslmap (map_prod (( *) (fst x)) ((+) (snd x))) ys)
+               (times_ms_aux xs (MSLCons y ys))))"
+
+definition scale_shift_ms_aux :: "('a :: times \<times> real) \<Rightarrow> ('a \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
+  "scale_shift_ms_aux = (\<lambda>(a,b) xs. mslmap (map_prod (( *) a) ((+) b)) xs)"
+
+lemma times_ms_aux_altdef:
+  "times_ms_aux xs ys = 
+     (case (xs, ys) of
+        (MSLNil, ys) \<Rightarrow> MSLNil
+      | (xs, MSLNil) \<Rightarrow> MSLNil
+      | (MSLCons x xs, MSLCons y ys) \<Rightarrow>
+          MSLCons (fst x * fst y, snd x + snd y)
+            (plus_ms_aux (scale_shift_ms_aux x ys) (times_ms_aux xs (MSLCons y ys))))"
+  by (subst times_ms_aux.code) (simp_all add: scale_shift_ms_aux_def split: msllist.splits)
+end
+
+corec powser_ms_aux :: "real msllist \<Rightarrow> ('a :: multiseries \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
+  "powser_ms_aux cs xs = (case cs of MSLNil \<Rightarrow> MSLNil | MSLCons c cs \<Rightarrow>
+     MSLCons (const_expansion c, 0) (times_ms_aux xs (powser_ms_aux cs xs)))"
+  
+definition minus_ms_aux :: "('a :: multiseries \<times> real) msllist \<Rightarrow> _" where
+  "minus_ms_aux xs ys = plus_ms_aux xs (uminus_ms_aux ys)"
+
+lemma is_expansion_aux_coinduct [consumes 1, case_names is_expansion_aux]:
+  fixes xs :: "('a :: multiseries \<times> real) msllist"
+  assumes "X xs f basis" "(\<And>xs f basis. X xs f basis \<Longrightarrow> 
+     (xs = MSLNil \<and> (\<forall>\<^sub>F x in at_top. f x = 0) \<and> length basis = Suc (expansion_level TYPE('a))) \<or>
+     (\<exists>xs' b e basis' C. xs = MSLCons (C, e) xs' \<and> basis = b # basis' \<and>
+        (X xs' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')) \<and> is_expansion C basis' \<and>
+        (\<forall>x>e. f \<in> o(\<lambda>xa. b xa powr x)) \<and> ms_exp_gt e (ms_aux_hd_exp xs')))"
+  shows "is_expansion_aux xs f basis"
+using assms(1)
+proof (coinduction arbitrary: xs f)
+  case (is_expansion_aux xs f)
+  from assms(2)[OF is_expansion_aux] show ?case by blast
+qed 
+
+lemma is_expansion_aux_cong:
+  assumes "is_expansion_aux F f basis"
+  assumes "eventually (\<lambda>x. f x = g x) at_top"
+  shows   "is_expansion_aux F g basis"
+  using assms
+proof (coinduction arbitrary: F f g rule: is_expansion_aux_coinduct)
+  case (is_expansion_aux F f g)
+  from this have ev: "eventually (\<lambda>x. g x = f x) at_top" by (simp add: eq_commute)
+  from ev have [simp]: "eventually (\<lambda>x. g x = 0) at_top \<longleftrightarrow> eventually (\<lambda>x. f x = 0) at_top"
+    by (rule eventually_subst')
+  from ev have [simp]: "(\<forall>\<^sub>F x in at_top. h x = g x - h' x) \<longleftrightarrow>
+                          (\<forall>\<^sub>F x in at_top. h x = f x - h' x)" for h h'
+      by (rule eventually_subst')
+  note [simp] = landau_o.small.in_cong[OF ev]
+  from is_expansion_aux(1) show ?case
+    by (cases rule: is_expansion_aux.cases) force+  
+qed
+
+lemma scale_shift_ms_aux_conv_mslmap: 
+  "scale_shift_ms_aux x = mslmap (map_prod (( *) (fst x)) ((+) (snd x)))"
+  by (rule ext) (simp add: scale_shift_ms_aux_def map_prod_def case_prod_unfold)
+
+fun inverse_ms_aux :: "('a :: multiseries \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
+  "inverse_ms_aux (MSLCons x xs) = 
+     (let c' = inverse (fst x)
+      in  scale_shift_ms_aux (c', -snd x) 
+            (powser_ms_aux (msllist_of_msstream (mssalternate 1 (-1)))
+              (scale_shift_ms_aux (c', -snd x) xs)))"
+
+(* TODO: Move? *)
+lemma inverse_prod_list: "inverse (prod_list xs :: 'a :: field) = prod_list (map inverse xs)"
+  by (induction xs) simp_all
+
+lemma eval_inverse_monom: 
+  "eval_monom (inverse_monom monom) basis = (\<lambda>x. inverse (eval_monom monom basis x))"
+  by (rule ext) (simp add: eval_monom_def inverse_monom_def zip_map2 o_def case_prod_unfold
+                   inverse_prod_list powr_minus)
+
+fun powr_ms_aux :: "bool \<Rightarrow> ('a :: multiseries \<times> real) msllist \<Rightarrow> real \<Rightarrow> ('a \<times> real) msllist" where
+  "powr_ms_aux abort (MSLCons x xs) p = 
+      scale_shift_ms_aux (powr_expansion abort (fst x) p, snd x * p)
+        (powser_ms_aux (gbinomial_series abort p)
+          (scale_shift_ms_aux (inverse (fst x), -snd x) xs))"
+
+fun power_ms_aux :: "bool \<Rightarrow> ('a :: multiseries \<times> real) msllist \<Rightarrow> nat \<Rightarrow> ('a \<times> real) msllist" where
+  "power_ms_aux abort (MSLCons x xs) n = 
+      scale_shift_ms_aux (power_expansion abort (fst x) n, snd x * real n)
+        (powser_ms_aux (gbinomial_series abort (real n))
+          (scale_shift_ms_aux (inverse (fst x), -snd x) xs))"
+
+definition sin_ms_aux' :: "('a :: multiseries \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
+  "sin_ms_aux' xs = times_ms_aux xs (powser_ms_aux (msllist_of_msstream sin_series_stream)
+                      (times_ms_aux xs xs))"
+  
+definition cos_ms_aux' :: "('a :: multiseries \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
+  "cos_ms_aux' xs = powser_ms_aux (msllist_of_msstream cos_series_stream) (times_ms_aux xs xs)"
+
+subsubsection \<open>Basic properties of multiseries operations\<close>  
+
+lemma uminus_ms_aux_MSLNil [simp]: "uminus_ms_aux MSLNil = MSLNil"
+  by (simp add: uminus_ms_aux_def)
+  
+lemma uminus_ms_aux_MSLCons: "uminus_ms_aux (MSLCons x xs) = MSLCons (-fst x, snd x) (uminus_ms_aux xs)"
+  by (simp add: uminus_ms_aux_def map_prod_def split: prod.splits)
+
+lemma uminus_ms_aux_MSLNil_iff [simp]: "uminus_ms_aux xs = MSLNil \<longleftrightarrow> xs = MSLNil"
+  by (simp add: uminus_ms_aux_def)
+  
+lemma hd_exp_uminus [simp]: "ms_aux_hd_exp (uminus_ms_aux xs) = ms_aux_hd_exp xs"
+  by (simp add: uminus_ms_aux_def ms_aux_hd_exp_def split: msllist.splits prod.split)
+  
+lemma scale_shift_ms_aux_MSLNil_iff [simp]: "scale_shift_ms_aux x F = MSLNil \<longleftrightarrow> F = MSLNil"
+  by (auto simp: scale_shift_ms_aux_def split: prod.split)
+
+lemma scale_shift_ms_aux_MSLNil [simp]: "scale_shift_ms_aux x MSLNil = MSLNil"
+  by (simp add: scale_shift_ms_aux_def split: prod.split)
+
+lemma scale_shift_ms_aux_1 [simp]:
+  "scale_shift_ms_aux (1 :: real, 0) xs = xs"
+  by (simp add: scale_shift_ms_aux_def map_prod_def msllist.map_id)
+
+lemma scale_shift_ms_aux_MSLCons: 
+  "scale_shift_ms_aux x (MSLCons y F) = MSLCons (fst x * fst y, snd x + snd y) (scale_shift_ms_aux x F)"
+  by (auto simp: scale_shift_ms_aux_def map_prod_def split: prod.split)
+
+lemma hd_exp_basis:
+  "ms_aux_hd_exp (scale_shift_ms_aux x xs) = map_option ((+) (snd x)) (ms_aux_hd_exp xs)"
+  by (auto simp: ms_aux_hd_exp_def scale_shift_ms_aux_MSLCons split: msllist.split)
+
+lemma plus_ms_aux_MSLNil_iff: "plus_ms_aux F G = MSLNil \<longleftrightarrow> F = MSLNil \<and> G = MSLNil"
+  by (subst plus_ms_aux.code) (simp split: msllist.splits)
+
+lemma plus_ms_aux_MSLNil [simp]: "plus_ms_aux F MSLNil = F" "plus_ms_aux MSLNil G = G"
+  by (subst plus_ms_aux.code, simp split: msllist.splits)+
+
+lemma plus_ms_aux_MSLCons: 
+  "plus_ms_aux (MSLCons x xs) (MSLCons y ys) = 
+     (if snd x > snd y then MSLCons x (plus_ms_aux xs (MSLCons y ys))
+      else if snd x < snd y then MSLCons y (plus_ms_aux (MSLCons x xs) ys)
+      else MSLCons (fst x + fst y, snd x) (plus_ms_aux xs ys))"
+  by (subst plus_ms_aux.code) simp
+
+lemma hd_exp_plus:
+  "ms_aux_hd_exp (plus_ms_aux xs ys) = combine_options max (ms_aux_hd_exp xs) (ms_aux_hd_exp ys)"
+  by (cases xs; cases ys)
+     (simp_all add: plus_ms_aux_MSLCons ms_aux_hd_exp_def max_def split: prod.split)
+
+lemma minus_ms_aux_MSLNil_left [simp]: "minus_ms_aux MSLNil ys = uminus_ms_aux ys"
+  by (simp add: minus_ms_aux_def)
+
+lemma minus_ms_aux_MSLNil_right [simp]: "minus_ms_aux xs MSLNil = xs"
+  by (simp add: minus_ms_aux_def)
+
+lemma times_ms_aux_MSLNil_iff: "times_ms_aux F G = MSLNil \<longleftrightarrow> F = MSLNil \<or> G = MSLNil"
+  by (subst times_ms_aux.code) (simp split: msllist.splits)
+
+lemma times_ms_aux_MSLNil [simp]: "times_ms_aux MSLNil G = MSLNil" "times_ms_aux F MSLNil = MSLNil"
+  by (subst times_ms_aux.code, simp split: msllist.splits)+
+
+lemma times_ms_aux_MSLCons: "times_ms_aux (MSLCons x xs) (MSLCons y ys) = 
+  MSLCons (fst x * fst y, snd x + snd y) (plus_ms_aux (scale_shift_ms_aux x ys)
+       (times_ms_aux xs (MSLCons y ys)))"
+  by (subst times_ms_aux_altdef) simp
+
+lemma hd_exp_times:
+  "ms_aux_hd_exp (times_ms_aux xs ys) = 
+     (case (ms_aux_hd_exp xs, ms_aux_hd_exp ys) of (Some x, Some y) \<Rightarrow> Some (x + y) | _ \<Rightarrow> None)"
+  by (cases xs; cases ys) (simp_all add: times_ms_aux_MSLCons ms_aux_hd_exp_def split: prod.split)
+
+
+subsubsection \<open>Induction upto friends for multiseries\<close>
+
+inductive ms_closure :: 
+  "(('a :: multiseries \<times> real) msllist \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> basis \<Rightarrow> bool) \<Rightarrow>
+   ('a :: multiseries \<times> real) msllist \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> basis \<Rightarrow> bool" for P where
+  ms_closure_embed: 
+    "P xs f basis \<Longrightarrow> ms_closure P xs f basis"
+| ms_closure_cong: 
+    "ms_closure P xs f basis \<Longrightarrow> eventually (\<lambda>x. f x = g x) at_top \<Longrightarrow> ms_closure P xs g basis"
+| ms_closure_MSLCons:
+    "ms_closure P xs (\<lambda>x. f x - eval C x * b x powr e) (b # basis) \<Longrightarrow>
+       is_expansion C basis \<Longrightarrow> (\<And>e'. e' > e \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')) \<Longrightarrow>
+       ms_exp_gt e (ms_aux_hd_exp xs) \<Longrightarrow>
+       ms_closure P (MSLCons (C, e) xs) f (b # basis)"
+| ms_closure_add: 
+    "ms_closure P xs f basis \<Longrightarrow> ms_closure P ys g basis \<Longrightarrow> 
+       ms_closure P (plus_ms_aux xs ys) (\<lambda>x. f x + g x) basis"
+| ms_closure_mult: 
+    "ms_closure P xs f basis \<Longrightarrow> ms_closure P ys g basis \<Longrightarrow> 
+       ms_closure P (times_ms_aux xs ys) (\<lambda>x. f x * g x) basis"
+| ms_closure_basis: 
+    "ms_closure P xs f (b # basis) \<Longrightarrow> is_expansion C basis \<Longrightarrow>
+       ms_closure P (scale_shift_ms_aux (C,e) xs) (\<lambda>x. eval C x * b x powr e * f x) (b # basis)"
+ | ms_closure_embed':
+    "is_expansion_aux xs f basis \<Longrightarrow> ms_closure P xs f basis"
+
+lemma is_expansion_aux_coinduct_upto [consumes 2, case_names A B]:
+  fixes xs :: "('a :: multiseries \<times> real) msllist"
+  assumes *: "X xs f basis" and basis: "basis_wf basis"
+  and step: "\<And>xs f basis. X xs f basis \<Longrightarrow> 
+    (xs = MSLNil \<and> eventually (\<lambda>x. f x = 0) at_top \<and> length basis = Suc (expansion_level TYPE('a))) \<or>
+    (\<exists>xs' b e basis' C. xs = MSLCons (C, e) xs' \<and> basis = b # basis' \<and>
+       ms_closure X xs' (\<lambda>x. f x - eval C x * b x powr e) (b # basis') \<and>
+       is_expansion C basis' \<and> (\<forall>e'>e. f \<in> o(\<lambda>x. b x powr e')) \<and> ms_exp_gt e (ms_aux_hd_exp xs'))"
+  shows "is_expansion_aux xs f basis"
+proof -
+  from * have "ms_closure X xs f basis" by (rule ms_closure_embed[of X xs f basis])
+  thus ?thesis
+  proof (coinduction arbitrary: xs f rule: is_expansion_aux_coinduct)
+    case (is_expansion_aux xs f)
+    from this and basis show ?case
+    proof (induction rule: ms_closure.induct)
+      case (ms_closure_embed xs f basis)
+      from step[OF ms_closure_embed(1)] show ?case by auto
+    next
+      case (ms_closure_MSLCons xs f C b e basis)
+      thus ?case by auto
+    next
+      case (ms_closure_cong xs f basis g)
+      note ev = \<open>\<forall>\<^sub>F x in at_top. f x = g x\<close>
+      hence ev_zero_iff: "eventually (\<lambda>x. f x = 0) at_top \<longleftrightarrow> eventually (\<lambda>x. g x = 0) at_top"
+        by (rule eventually_subst')
+      have *: "ms_closure X xs' (\<lambda>x. f x - c x * b x powr e) basis \<longleftrightarrow>
+                 ms_closure X xs' (\<lambda>x. g x - c x * b x powr e) basis" for xs' c b e
+        by (rule iffI; erule ms_closure.intros(2)) (insert ev, auto elim: eventually_mono)
+      from ms_closure_cong show ?case
+        by (simp add: ev_zero_iff * landau_o.small.in_cong[OF ev])     
+    next
+      case (ms_closure_embed' xs f basis)
+      from this show ?case
+        by (cases rule: is_expansion_aux.cases) (force intro: ms_closure.intros(7))+
+    next
+      
+      case (ms_closure_basis xs f b basis C e)
+      show ?case
+      proof (cases xs rule: ms_aux_cases)
+        case MSLNil
+        with ms_closure_basis show ?thesis
+          by (auto simp: scale_shift_ms_aux_def split: prod.split elim: eventually_mono)
+      next
+        case (MSLCons C' e' xs')
+        with ms_closure_basis have IH:
+          "ms_closure X (MSLCons (C', e') xs') f (b # basis)"
+          "is_expansion C basis" "xs = MSLCons (C', e') xs'"
+          "ms_closure X xs' (\<lambda>x. f x - eval C' x * b x powr e') (b # basis)"
+          "ms_exp_gt e' (ms_aux_hd_exp xs')"
+          "is_expansion C' basis" "\<And>x. x > e' \<Longrightarrow> f \<in> o(\<lambda>xa. b xa powr x)"
+          by auto
+        
+        {
+          fix e'' :: real assume *: "e + e' < e''"
+          define d where "d = (e'' - e - e') / 2"
+          from * have "d > 0" by (simp add: d_def)
+          hence "(\<lambda>x. eval C x * b x powr e * f x) \<in> o(\<lambda>x. b x powr d * b x powr e * b x powr (e'+d))"
+            using ms_closure_basis(4) IH
+            by (intro landau_o.small.mult[OF landau_o.small_big_mult] IH 
+              is_expansion_imp_smallo[OF _ IH(2)]) (simp_all add: basis_wf_Cons)
+          also have "\<dots> = o(\<lambda>x. b x powr e'')" by (simp add: d_def powr_add [symmetric])
+          finally have "(\<lambda>x. eval C x * b x powr e * f x) \<in> \<dots>" .
+        }
+        moreover have "ms_closure X xs' (\<lambda>x. f x - eval C' x * b x powr e') (b # basis)"
+          using ms_closure_basis IH by auto
+        note ms_closure.intros(6)[OF this IH(2), of e]
+        moreover have "ms_exp_gt (e + e') (ms_aux_hd_exp (scale_shift_ms_aux (C, e) xs'))"
+          using \<open>ms_exp_gt e' (ms_aux_hd_exp xs')\<close>
+          by (cases xs') (simp_all add: hd_exp_basis scale_shift_ms_aux_def )
+        ultimately show ?thesis using IH(2,6) MSLCons ms_closure_basis(4)
+          by (auto simp: scale_shift_ms_aux_MSLCons algebra_simps powr_add basis_wf_Cons
+                   intro: is_expansion_mult)
+      qed
+      
+    next
+      case (ms_closure_add xs f basis ys g)
+      show ?case
+      proof (cases xs ys rule: ms_aux_cases[case_product ms_aux_cases])
+        case MSLNil_MSLNil
+        with ms_closure_add 
+          have "eventually (\<lambda>x. f x = 0) at_top" "eventually (\<lambda>x. g x = 0) at_top"
+            by simp_all
+        hence "eventually (\<lambda>x. f x + g x = 0) at_top" by eventually_elim simp
+        with MSLNil_MSLNil ms_closure_add show ?thesis by simp
+      next
+        case (MSLNil_MSLCons c e xs')
+        with ms_closure_add have "eventually (\<lambda>x. f x = 0) at_top" by simp
+        hence ev: "eventually (\<lambda>x. f x + g x = g x) at_top" by eventually_elim simp
+        have *: "ms_closure X xs (\<lambda>x. f x + g x - y x) basis \<longleftrightarrow> 
+                   ms_closure X xs (\<lambda>x. g x - y x) basis" for y basis xs
+          by (rule iffI; erule ms_closure_cong) (insert ev, simp_all)
+        from MSLNil_MSLCons ms_closure_add
+        show ?thesis by (simp add: * landau_o.small.in_cong[OF ev])
+      next
+        case (MSLCons_MSLNil c e xs')
+        with ms_closure_add have "eventually (\<lambda>x. g x = 0) at_top" by simp
+        hence ev: "eventually (\<lambda>x. f x + g x = f x) at_top" by eventually_elim simp
+        have *: "ms_closure X xs (\<lambda>x. f x + g x - y x) basis \<longleftrightarrow> 
+                   ms_closure X xs (\<lambda>x. f x - y x) basis" for y basis xs
+          by (rule iffI; erule ms_closure_cong) (insert ev, simp_all)
+        from MSLCons_MSLNil ms_closure_add
+        show ?thesis by (simp add: * landau_o.small.in_cong[OF ev])          
+      next
+        case (MSLCons_MSLCons C1 e1 xs' C2 e2 ys')
+        with ms_closure_add obtain b basis' where IH:
+          "basis_wf (b # basis')" "basis = b # basis'"
+          "ms_closure X (MSLCons (C1, e1) xs') f (b # basis')"
+          "ms_closure X (MSLCons (C2, e2) ys') g (b # basis')"
+          "xs = MSLCons (C1, e1) xs'" "ys = MSLCons (C2, e2) ys'"
+          "ms_closure X xs' (\<lambda>x. f x - eval C1 x * b x powr e1) (b # basis')"
+          "ms_closure X ys' (\<lambda>x. g x - eval C2 x * b x powr e2) (b # basis')"
+          "is_expansion C1 basis'" "\<And>e. e > e1 \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e)"
+          "is_expansion C2 basis'" "\<And>e. e > e2 \<Longrightarrow> g \<in> o(\<lambda>x. b x powr e)"
+          "ms_exp_gt e1 (ms_aux_hd_exp xs')" "ms_exp_gt e2 (ms_aux_hd_exp ys')"
+          by blast
+        have o: "(\<lambda>x. f x + g x) \<in> o(\<lambda>x. b x powr e)" if "e > max e1 e2" for e
+          using that by (intro sum_in_smallo IH) simp_all
+      
+        show ?thesis
+        proof (cases e1 e2 rule: linorder_cases)
+          case less
+          have gt: "ms_exp_gt e2 (combine_options max (Some e1) (ms_aux_hd_exp ys'))"
+            using \<open>ms_exp_gt e2 _\<close> less by (cases "ms_aux_hd_exp ys'") auto
+          have "ms_closure X (plus_ms_aux xs ys')
+                  (\<lambda>x. f x + (g x - eval C2 x * b x powr e2)) (b # basis')"
+            by (rule ms_closure.intros(4)) (insert IH, simp_all)
+          with less IH(2,11,14) o gt show ?thesis
+            by (intro disjI2) (simp add: MSLCons_MSLCons plus_ms_aux_MSLCons algebra_simps hd_exp_plus)
+        next
+          case greater
+          have gt: "ms_exp_gt e1 (combine_options max (ms_aux_hd_exp xs') (Some e2))"
+            using \<open>ms_exp_gt e1 _\<close> greater by (cases "ms_aux_hd_exp xs'") auto
+          have "ms_closure X (plus_ms_aux xs' ys)
+                  (\<lambda>x. (f x - eval C1 x * b x powr e1) + g x) (b # basis')"
+            by (rule ms_closure.intros(4)) (insert IH, simp_all)
+          with greater IH(2,9,13) o gt show ?thesis
+            by (intro disjI2) (simp add:  MSLCons_MSLCons plus_ms_aux_MSLCons algebra_simps hd_exp_plus)
+        next
+          case equal
+          have gt: "ms_exp_gt e2 (combine_options max (ms_aux_hd_exp xs')
+                      (ms_aux_hd_exp ys'))"
+            using \<open>ms_exp_gt e1 _\<close> \<open>ms_exp_gt e2 _\<close> equal
+            by (cases "ms_aux_hd_exp xs'"; cases "ms_aux_hd_exp ys'") auto
+          have "ms_closure X (plus_ms_aux xs' ys')
+                  (\<lambda>x. (f x - eval C1 x * b x powr e1) + (g x - eval C2 x * b x powr e2)) (b # basis')"
+            by (rule ms_closure.intros(4)) (insert IH, simp_all)
+          with equal IH(1,2,9,11,13,14) o gt show ?thesis
+            by (intro disjI2) 
+               (auto intro: is_expansion_add 
+                     simp: plus_ms_aux_MSLCons basis_wf_Cons algebra_simps hd_exp_plus  MSLCons_MSLCons)
+        qed
+      qed
+      
+    next
+      
+      case (ms_closure_mult xs f basis ys g)
+      show ?case
+      proof (cases "xs = MSLNil \<or> ys = MSLNil")
+        case True
+        with ms_closure_mult 
+          have "eventually (\<lambda>x. f x = 0) at_top \<or> eventually (\<lambda>x. g x = 0) at_top" by blast
+        hence "eventually (\<lambda>x. f x * g x = 0) at_top" by (auto elim: eventually_mono)
+        with ms_closure_mult True show ?thesis by auto
+      next
+        case False
+        with ms_closure_mult obtain C1 e1 xs' C2 e2 ys' b basis' where IH:
+          "xs = MSLCons (C1, e1) xs'" "ys = MSLCons (C2, e2) ys'"
+          "basis_wf (b # basis')" "basis = b # basis'"
+          "ms_closure X (MSLCons (C1, e1) xs') f (b # basis')"
+          "ms_closure X (MSLCons (C2, e2) ys') g (b # basis')"
+          "ms_closure X xs' (\<lambda>x. f x - eval C1 x * b x powr e1) (b # basis')"
+          "ms_closure X ys' (\<lambda>x. g x - eval C2 x * b x powr e2) (b # basis')"
+          "is_expansion C1 basis'" "\<And>e. e > e1 \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e)"
+          "is_expansion C2 basis'" "\<And>e. e > e2 \<Longrightarrow> g \<in> o(\<lambda>x. b x powr e)"
+          "ms_exp_gt e1 (ms_aux_hd_exp xs')" "ms_exp_gt e2 (ms_aux_hd_exp ys')"
+          by blast
+        have o: "(\<lambda>x. f x * g x) \<in> o(\<lambda>x. b x powr e)" if "e > e1 + e2" for e
+        proof -
+          define d where "d = (e - e1 - e2) / 2"
+          from that have d: "d > 0" by (simp add: d_def)
+          hence "(\<lambda>x. f x * g x) \<in> o(\<lambda>x. b x powr (e1 + d) * b x powr (e2 + d))"
+            by (intro landau_o.small.mult IH) simp_all
+          also have "\<dots> = o(\<lambda>x. b x powr e)" by (simp add: d_def powr_add [symmetric])
+          finally show ?thesis .
+        qed
+
+        have "ms_closure X (plus_ms_aux (scale_shift_ms_aux (C1, e1) ys') (times_ms_aux xs' ys))
+                (\<lambda>x. eval C1 x * b x powr e1 * (g x - eval C2 x * b x powr e2) + 
+                     ((f x - eval C1 x * b x powr e1) * g x)) (b # basis')"
+          (is "ms_closure _ ?zs ?f _") using ms_closure_mult IH(4)
+          by (intro ms_closure.intros(4-6) IH) simp_all
+        also have "?f = (\<lambda>x. f x * g x - eval C1 x * eval C2 x * b x powr (e1 + e2))"
+          by (intro ext) (simp add: algebra_simps powr_add)
+        finally have "ms_closure X ?zs \<dots> (b # basis')" .
+        moreover from \<open>ms_exp_gt e1 _\<close> \<open>ms_exp_gt e2 _\<close>
+        have "ms_exp_gt (e1 + e2) (combine_options max (map_option ((+) e1) 
+                (ms_aux_hd_exp ys')) (case ms_aux_hd_exp xs' of None \<Rightarrow> None | Some x \<Rightarrow>
+                  (case Some e2 of None \<Rightarrow> None | Some y \<Rightarrow> Some (x + y))))"
+            by (cases "ms_aux_hd_exp xs'"; cases "ms_aux_hd_exp ys'")
+               (simp_all add: hd_exp_times hd_exp_plus hd_exp_basis IH)
+        ultimately show ?thesis using IH(1,2,3,4,9,11) o
+          by (auto simp: times_ms_aux_MSLCons basis_wf_Cons hd_exp_times hd_exp_plus hd_exp_basis
+                   intro: is_expansion_mult)
+      qed
+    qed
+  qed
+qed
+
+
+
+subsubsection \<open>Dominant terms\<close>
+
+lemma one_smallo_powr: "e > (0::real) \<Longrightarrow> (\<lambda>_. 1) \<in> o(\<lambda>x. x powr e)"
+  by (auto simp: smallomega_iff_smallo [symmetric] real_powr_at_top 
+                 smallomegaI_filterlim_at_top_norm)
+
+lemma powr_smallo_one: "e < (0::real) \<Longrightarrow> (\<lambda>x. x powr e) \<in> o(\<lambda>_. 1)"
+  by (intro smalloI_tendsto) (auto intro!: tendsto_neg_powr filterlim_ident)
+
+lemma eval_monom_smallo':
+  assumes "basis_wf (b # basis)" "e > 0"
+  shows   "eval_monom x basis \<in> o(\<lambda>x. b x powr e)"
+proof (cases x)
+  case (Pair c es)
+  from assms show ?thesis unfolding Pair
+  proof (induction es arbitrary: b basis e)
+    case (Nil b basis e)
+    hence b: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+    have "eval_monom (c, []) basis \<in> O(\<lambda>_. 1)" by simp
+    also have "(\<lambda>_. 1) \<in> o(\<lambda>x. b x powr e)"
+      by (rule landau_o.small.compose[OF _ b]) (insert Nil, simp add: one_smallo_powr)
+    finally show ?case .
+  next
+    case (Cons e' es b basis e)
+    show ?case
+    proof (cases basis)
+      case Nil
+      from Cons have b: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+      from Nil have "eval_monom (c, e' # es) basis \<in> O(\<lambda>_. 1)" by simp
+      also have "(\<lambda>_. 1) \<in> o(\<lambda>x. b x powr e)"
+        by (rule landau_o.small.compose[OF _ b]) (insert Cons.prems, simp add: one_smallo_powr)
+      finally show ?thesis .
+    next
+      case (Cons b' basis')
+      from Cons.prems have "eval_monom (c, es) basis' \<in> o(\<lambda>x. b' x powr 1)"
+        by (intro Cons.IH) (simp_all add: basis_wf_Cons Cons)
+      hence "(\<lambda>x. eval_monom (c, es) basis' x * b' x powr e') \<in> o(\<lambda>x. b' x powr 1 * b' x powr e')"
+        by (rule landau_o.small_big_mult) simp_all
+      also have "\<dots> = o(\<lambda>x. b' x powr (1 + e'))" by (simp add: powr_add)
+      also from Cons.prems have "(\<lambda>x. b' x powr (1 + e')) \<in> o(\<lambda>x. b x powr e)"
+        by (intro ln_smallo_imp_flat) (simp_all add: basis_wf_Cons Cons)
+      finally show ?thesis by (simp add: powr_add [symmetric] Cons eval_monom_Cons)
+    qed
+  qed
+qed
+
+lemma eval_monom_smallomega':
+  assumes "basis_wf (b # basis)" "e < 0"
+  shows   "eval_monom (1, es) basis \<in> \<omega>(\<lambda>x. b x powr e)"
+  using assms
+proof (induction es arbitrary: b basis e)
+  case (Nil b basis e)
+  hence b: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  have "eval_monom (1, []) basis \<in> \<Omega>(\<lambda>_. 1)" by simp
+  also have "(\<lambda>_. 1) \<in> \<omega>(\<lambda>x. b x powr e)" unfolding smallomega_iff_smallo
+    by (rule landau_o.small.compose[OF _ b]) (insert Nil, simp add: powr_smallo_one)
+  finally show ?case .
+next
+  case (Cons e' es b basis e)
+  show ?case
+  proof (cases basis)
+    case Nil
+    from Cons have b: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+    from Nil have "eval_monom (1, e' # es) basis \<in> \<Omega>(\<lambda>_. 1)" by simp
+    also have "(\<lambda>_. 1) \<in> \<omega>(\<lambda>x. b x powr e)" unfolding smallomega_iff_smallo
+      by (rule landau_o.small.compose[OF _ b]) (insert Cons.prems, simp add: powr_smallo_one)
+    finally show ?thesis .
+  next
+    case (Cons b' basis')
+    from Cons.prems have "eval_monom (1, es) basis' \<in> \<omega>(\<lambda>x. b' x powr -1)"
+      by (intro Cons.IH) (simp_all add: basis_wf_Cons Cons)
+    hence "(\<lambda>x. eval_monom (1, es) basis' x * b' x powr e') \<in> \<omega>(\<lambda>x. b' x powr -1 * b' x powr e')"
+      by (rule landau_omega.small_big_mult) simp_all
+    also have "\<dots> = \<omega>(\<lambda>x. b' x powr (e' - 1))" by (simp add: powr_diff powr_minus field_simps)
+    also from Cons.prems have "(\<lambda>x. b' x powr (e' - 1)) \<in> \<omega>(\<lambda>x. b x powr e)"
+      unfolding smallomega_iff_smallo
+      by (intro ln_smallo_imp_flat') (simp_all add: basis_wf_Cons Cons)
+    finally show ?thesis by (simp add: powr_add [symmetric] Cons eval_monom_Cons)
+  qed
+qed
+
+lemma dominant_term_ms_aux:
+  assumes basis: "basis_wf basis" and xs: "trimmed_ms_aux xs" "is_expansion_aux xs f basis"
+  shows   "f \<sim>[at_top] eval_monom (dominant_term_ms_aux xs) basis" (is ?thesis1)
+    and   "eventually (\<lambda>x. sgn (f x) = sgn (fst (dominant_term_ms_aux xs))) at_top" (is ?thesis2)
+proof -
+  include asymp_equiv_notation
+  from xs(2,1) obtain xs' C b e basis' where xs':
+    "trimmed C" "xs = MSLCons (C, e) xs'" "basis = b # basis'"
+    "is_expansion_aux xs' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
+    "is_expansion C basis'" "(\<And>e'. e < e' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e'))"
+    "ms_exp_gt e (ms_aux_hd_exp xs')"
+    by (cases rule: is_expansion_aux.cases) (auto simp: trimmed_ms_aux_MSLCons)
+  from is_expansion_aux_imp_smallo''[OF xs'(4,7)]
+    obtain e' where e': "e' < e" "(\<lambda>x. f x - eval C x * b x powr e) \<in> o(\<lambda>x. b x powr e')"
+      unfolding list.sel by blast   
+  from basis xs' have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+    
+  note e'(2)
+  also have "o(\<lambda>x. b x powr e') = o(\<lambda>x. b x powr (e' - e) * b x powr e)" 
+    by (simp add: powr_add [symmetric])
+  also from xs' basis e'(1) have "eval C \<in> \<omega>(\<lambda>x. b x powr (e' - e))"
+    by (intro is_expansion_imp_smallomega[of basis']) (simp_all add: basis_wf_Cons)
+  hence "(\<lambda>x. b x powr (e' - e) * b x powr e) \<in> o(\<lambda>x. eval C x * b x powr e)"
+    by (intro landau_o.small_big_mult) (simp_all add: smallomega_iff_smallo)
+  finally have *: "(\<lambda>x. f x - eval C x * b x powr e) \<in> o(\<lambda>x. eval C x * b x powr e)" .
+  hence "f \<sim> (\<lambda>x. eval C x * b x powr e)" by (intro smallo_imp_asymp_equiv) simp
+  also from basis xs' have "eval C \<sim> (\<lambda>x. eval_monom (dominant_term C) basis' x)"
+    by (intro dominant_term) (simp_all add: basis_wf_Cons)
+  also have "(\<lambda>a. eval_monom (dominant_term C) basis' a * b a powr e) = 
+               eval_monom (dominant_term_ms_aux xs) basis"
+    by (auto simp add: dominant_term_ms_aux_def xs' eval_monom_Cons fun_eq_iff split: prod.split)
+  finally show ?thesis1 by - (simp_all add: asymp_equiv_intros)
+  
+  have "eventually (\<lambda>x. sgn (eval C x * b x powr e + (f x - eval C x * b x powr e)) = 
+                          sgn (eval C x * b x powr e)) at_top"
+    by (intro smallo_imp_eventually_sgn *)
+  moreover have "eventually (\<lambda>x. sgn (eval C x) = sgn (fst (dominant_term C))) at_top"
+    using xs' basis by (intro trimmed_imp_eventually_sgn[of basis']) (simp_all add: basis_wf_Cons)
+  ultimately have "eventually (\<lambda>x. sgn (f x) = sgn (fst (dominant_term C))) at_top"
+    using b_pos by eventually_elim (simp_all add: sgn_mult)
+  thus ?thesis2 using xs'(2) by (auto simp: dominant_term_ms_aux_def split: prod.split)
+qed
+
+lemma eval_monom_smallo:
+  assumes "basis \<noteq> []" "basis_wf basis" "length es = length basis" "e > hd es"
+  shows   "eval_monom (c, es) basis \<in> o(\<lambda>x. hd basis x powr e)"
+proof (cases es; cases basis)
+  fix b basis' e' es' assume [simp]: "basis = b # basis'" "es = e' # es'"
+  from assms have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" 
+    using filterlim_at_top_dense by blast
+  have "(\<lambda>x. eval_monom (1, es') basis' x * b x powr e') \<in> o(\<lambda>x. b x powr (e - e') * b x powr e')"
+    using assms by (intro landau_o.small_big_mult eval_monom_smallo') auto
+  also have "(\<lambda>x. b x powr (e - e') * b x powr e') \<in> \<Theta>(\<lambda>x. b x powr e)"
+    by (intro bigthetaI_cong eventually_mono[OF b_pos]) (auto simp: powr_diff)
+  finally show ?thesis by (simp add: eval_monom_def algebra_simps)
+qed (insert assms, auto)
+
+lemma eval_monom_smallomega:
+  assumes "basis \<noteq> []" "basis_wf basis" "length es = length basis" "e < hd es"
+  shows   "eval_monom (1, es) basis \<in> \<omega>(\<lambda>x. hd basis x powr e)"
+proof (cases es; cases basis)
+  fix b basis' e' es' assume [simp]: "basis = b # basis'" "es = e' # es'"
+  from assms have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" 
+    using filterlim_at_top_dense by blast
+  have "(\<lambda>x. eval_monom (1, es') basis' x * b x powr e') \<in> \<omega>(\<lambda>x. b x powr (e - e') * b x powr e')"
+    using assms by (intro landau_omega.small_big_mult eval_monom_smallomega') auto
+  also have "(\<lambda>x. b x powr (e - e') * b x powr e') \<in> \<Theta>(\<lambda>x. b x powr e)"
+    by (intro bigthetaI_cong eventually_mono[OF b_pos]) (auto simp: powr_diff)
+  finally show ?thesis by (simp add: eval_monom_Cons)
+qed (insert assms, auto)
+
+
+subsubsection \<open>Correctness of multiseries operations\<close>
+
+lemma drop_zero_ms_aux:
+  assumes "eventually (\<lambda>x. eval C x = 0) at_top"
+  assumes "is_expansion_aux (MSLCons (C, e) xs) f basis"
+  shows   "is_expansion_aux xs f basis"
+proof (rule is_expansion_aux_cong)
+  from assms(2) show "is_expansion_aux xs (\<lambda>x. f x - eval C x * hd basis x powr e) basis"
+    by (auto elim: is_expansion_aux_MSLCons)
+  from assms(1) show "eventually (\<lambda>x. f x - eval C x * hd basis x powr e = f x) at_top"
+    by eventually_elim simp
+qed
+
+lemma dominant_term_ms_aux_bigo:
+  assumes basis: "basis_wf basis" and xs: "is_expansion_aux xs f basis"
+  shows   "f \<in> O(eval_monom (1, snd (dominant_term_ms_aux xs)) basis)" (is ?thesis1)
+proof (cases xs rule: ms_aux_cases)
+  case MSLNil
+  with assms have "eventually (\<lambda>x. f x = 0) at_top"
+    by (auto elim: is_expansion_aux.cases)
+  hence "f \<in> \<Theta>(\<lambda>_. 0)" by (intro bigthetaI_cong) auto
+  also have "(\<lambda>_. 0) \<in> O(eval_monom (1, snd (dominant_term_ms_aux xs)) basis)" by simp
+  finally show ?thesis .
+next
+  case [simp]: (MSLCons C e xs')
+  from xs obtain b basis' where xs':
+    "basis = b # basis'"
+    "is_expansion_aux xs' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
+    "is_expansion C basis'" "(\<And>e'. e < e' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e'))"
+    "ms_exp_gt e (ms_aux_hd_exp xs')"
+    by (cases rule: is_expansion_aux.cases) auto
+  from is_expansion_aux_imp_smallo''[OF xs'(2,5)]
+    obtain e' where e': "e' < e" "(\<lambda>x. f x - eval C x * b x powr e) \<in> o(\<lambda>x. b x powr e')"
+      unfolding list.sel by blast   
+  from basis xs' have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+
+  let ?h = "(\<lambda>x. eval_monom (1, snd (dominant_term C)) basis' x * b x powr e)"
+  note e'(2)
+  also have "(\<lambda>x. b x powr e') \<in> \<Theta>(\<lambda>x. b x powr (e' - e) * b x powr e)"
+    by (intro bigthetaI_cong eventually_mono[OF b_pos]) (auto simp: powr_diff)
+  also have "(\<lambda>x. b x powr (e' - e) * b x powr e) \<in> o(?h)"
+    unfolding smallomega_iff_smallo [symmetric] using e'(1) basis xs'
+    by (intro landau_omega.small_big_mult landau_omega.big_refl eval_monom_smallomega')
+       (simp_all add: basis_wf_Cons)
+  finally have "(\<lambda>x. f x - eval C x * b x powr e) \<in> O(?h)" by (rule landau_o.small_imp_big)
+  moreover have "(\<lambda>x. eval C x * b x powr e) \<in> O(?h)"
+    using basis xs' by (intro landau_o.big.mult dominant_term_bigo) (auto simp: basis_wf_Cons)
+  ultimately have "(\<lambda>x. f x - eval C x * b x powr e + eval C x * b x powr e) \<in> O(?h)"
+    by (rule sum_in_bigo)
+  hence "f \<in> O(?h)" by simp
+  also have "?h = eval_monom (1, snd (dominant_term_ms_aux xs)) basis"
+    using xs' by (auto simp: dominant_term_ms_aux_def case_prod_unfold eval_monom_Cons)
+  finally show ?thesis .
+qed
+
+lemma const_ms_aux:
+  assumes basis: "basis_wf basis" 
+      and "length basis = Suc (expansion_level TYPE('a::multiseries))"
+  shows   "is_expansion_aux (const_ms_aux c :: ('a \<times> real) msllist) (\<lambda>_. c) basis"
+proof -
+  from assms(2) obtain b basis' where [simp]: "basis = b # basis'" by (cases basis) simp_all
+  from basis have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+
+  have "(\<lambda>_. c) \<in> o(\<lambda>x. b x powr e)" if "e > 0" for e
+  proof -
+    have "(\<lambda>_. c) \<in> O(\<lambda>_. 1)" by (cases "c = 0") simp_all  
+    also from b_pos have "(\<lambda>_. 1) \<in> \<Theta>(\<lambda>x. b x powr 0)" 
+      by (intro bigthetaI_cong) (auto elim: eventually_mono)
+    also from that b_limit have "(\<lambda>x. b x powr 0) \<in> o(\<lambda>x. b x powr e)" 
+      by (subst powr_smallo_iff) auto
+    finally show ?thesis .
+  qed
+  with b_pos assms show ?thesis
+    by (auto intro!: is_expansion_aux.intros simp: const_ms_aux_def is_expansion_const basis_wf_Cons 
+             elim: eventually_mono)
+qed
+
+lemma uminus_ms_aux:
+  assumes basis: "basis_wf basis"
+  assumes F: "is_expansion_aux F f basis"
+  shows   "is_expansion_aux (uminus_ms_aux F) (\<lambda>x. -f x) basis"
+  using F
+proof (coinduction arbitrary: F f rule: is_expansion_aux_coinduct)
+  case (is_expansion_aux F f)
+  note IH = is_expansion_aux
+  show ?case
+  proof (cases F rule: ms_aux_cases)
+    case MSLNil
+    with IH show ?thesis by (auto simp: uminus_ms_aux_def elim: is_expansion_aux.cases)
+  next
+    case (MSLCons C e xs)
+    with IH obtain b basis'
+      where IH: "basis = b # basis'" 
+                "is_expansion_aux xs (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
+                "is_expansion C basis'"
+                "\<And>e'. e < e' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')" "ms_exp_gt e (ms_aux_hd_exp xs)"
+      by (auto elim: is_expansion_aux_MSLCons)
+    from basis IH have basis': "basis_wf basis'" by (simp add: basis_wf_Cons)
+    with IH basis show ?thesis
+      by (force simp: uminus_ms_aux_MSLCons MSLCons is_expansion_uminus)
+  qed
+qed
+
+lemma plus_ms_aux:
+  assumes "basis_wf basis" "is_expansion_aux F f basis" "is_expansion_aux G g basis"
+  shows   "is_expansion_aux (plus_ms_aux F G) (\<lambda>x. f x + g x) basis"
+  using assms
+proof (coinduction arbitrary: F f G g rule: is_expansion_aux_coinduct)
+  case (is_expansion_aux F f G g)
+  note IH = this
+  show ?case
+  proof (cases F G rule: ms_aux_cases[case_product ms_aux_cases])
+    assume [simp]: "F = MSLNil" "G = MSLNil"
+    with IH have *: "eventually (\<lambda>x. f x = 0) at_top" "eventually (\<lambda>x. g x = 0) at_top" 
+                 and length: "length basis = Suc (expansion_level TYPE('a))"
+      by (auto elim: is_expansion_aux.cases)
+    from * have "eventually (\<lambda>x. f x + g x = 0) at_top" by eventually_elim simp
+    with length show ?case by simp 
+  next
+    assume [simp]: "F = MSLNil"
+    fix C e G' assume G: "G = MSLCons (C, e) G'"
+    from IH have f: "eventually (\<lambda>x. f x = 0) at_top" by (auto elim: is_expansion_aux.cases)
+    from f have "eventually (\<lambda>x. f x + g x = g x) at_top" by eventually_elim simp
+    note eq = landau_o.small.in_cong[OF this]
+    from IH(3) G obtain b basis' where G':
+      "basis = b # basis'"
+      "is_expansion_aux G' (\<lambda>x. g x - eval C x * b x powr e) (b # basis')"
+      "is_expansion C basis'" "\<forall>e'>e. g \<in> o(\<lambda>x. b x powr e')" "ms_exp_gt e (ms_aux_hd_exp G')"
+      by (auto elim!: is_expansion_aux_MSLCons)
+    show ?thesis
+      by (rule disjI2, inst_existentials G' b e basis' C F f G' "\<lambda>x. g x - eval C x * b x powr e")
+         (insert G' IH(1,2), simp_all add: G eq algebra_simps)
+  next
+    assume [simp]: "G = MSLNil"
+    fix C e F' assume F: "F = MSLCons (C, e) F'"
+    from IH have g: "eventually (\<lambda>x. g x = 0) at_top" by (auto elim: is_expansion_aux.cases)
+    hence "eventually (\<lambda>x. f x + g x = f x) at_top" by eventually_elim simp
+    note eq = landau_o.small.in_cong[OF this]
+    from IH F obtain b basis' where F':
+      "basis = b # basis'"
+      "is_expansion_aux F' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
+      "is_expansion C basis'" "\<forall>e'>e. f \<in> o(\<lambda>x. b x powr e')" "ms_exp_gt e (ms_aux_hd_exp F')"
+      by (auto elim!: is_expansion_aux_MSLCons)
+    show ?thesis
+      by (rule disjI2, inst_existentials F' b e basis' C F' "\<lambda>x. f x - eval C x * b x powr e" G g)
+         (insert F g F' IH, simp_all add: eq algebra_simps)
+  next
+    fix C1 e1 F' C2 e2 G'
+    assume F: "F = MSLCons (C1, e1) F'" and G: "G = MSLCons (C2, e2) G'"
+    from IH F obtain b basis' where
+      basis': "basis = b # basis'" and F':
+      "is_expansion_aux F' (\<lambda>x. f x - eval C1 x * b x powr e1) (b # basis')"
+      "is_expansion C1 basis'" "\<And>e'. e' > e1 \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')" 
+      "ms_exp_gt e1 (ms_aux_hd_exp F')"
+      by (auto elim!: is_expansion_aux_MSLCons)
+    from IH G basis' have G':
+      "is_expansion_aux G' (\<lambda>x. g x - eval C2 x * b x powr e2) (b # basis')"
+      "is_expansion C2 basis'" "\<And>e'. e' > e2 \<Longrightarrow> g \<in> o(\<lambda>x. b x powr e')"
+      "ms_exp_gt e2 (ms_aux_hd_exp G')"
+      by (auto elim!: is_expansion_aux_MSLCons)  
+    hence *: "\<forall>x>max e1 e2. (\<lambda>x. f x + g x) \<in> o(\<lambda>xa. b xa powr x)"
+      by (intro allI impI sum_in_smallo F' G') simp_all
+      
+    consider "e1 > e2" | "e1 < e2" | "e1 = e2" by force
+    thus ?thesis
+    proof cases
+      assume e: "e1 > e2"
+      have gt: "ms_exp_gt e1 (combine_options max (ms_aux_hd_exp F') (Some e2))"
+        using \<open>ms_exp_gt e1 _\<close> e by (cases "ms_aux_hd_exp F'") simp_all
+      show ?thesis
+        by (rule disjI2, inst_existentials "plus_ms_aux F' G" b e1 basis' C1 F' 
+                           "\<lambda>x. f x - eval C1 x * b x powr e1" G g)
+           (insert e F'(1,2,4) IH(1,3) basis'(1) * gt,
+            simp_all add: F G plus_ms_aux_MSLCons algebra_simps hd_exp_plus)
+    next
+      assume e: "e1 < e2"
+      have gt: "ms_exp_gt e2 (combine_options max (Some e1) (ms_aux_hd_exp G'))"
+        using \<open>ms_exp_gt e2 _\<close> e by (cases "ms_aux_hd_exp G'") simp_all
+      show ?thesis
+        by (rule disjI2, inst_existentials "plus_ms_aux F G'" b e2 basis' C2 F f G'
+                           "\<lambda>x. g x - eval C2 x * b x powr e2")
+           (insert e G'(1,2,4) IH(1,2) basis'(1) * gt, 
+            simp_all add: F G plus_ms_aux_MSLCons algebra_simps hd_exp_plus)
+    next
+      assume e: "e1 = e2"
+      have gt: "ms_exp_gt e2 (combine_options max (ms_aux_hd_exp F') (ms_aux_hd_exp G'))"
+        using \<open>ms_exp_gt e1 _\<close> \<open>ms_exp_gt e2 _\<close> e
+        by (cases "ms_aux_hd_exp F'"; cases "ms_aux_hd_exp G'") simp_all
+      show ?thesis
+        by (rule disjI2, inst_existentials "plus_ms_aux F' G'" b e1 basis' "C1 + C2" F' 
+                         "\<lambda>x. f x - eval C1 x * b x powr e1" G' "\<lambda>x. g x - eval C2 x * b x powr e2")
+           (insert e F'(1,2,4) G'(1,2,4) IH(1) basis'(1) * gt, 
+            simp_all add: F G plus_ms_aux_MSLCons algebra_simps hd_exp_plus is_expansion_add basis_wf_Cons)
+    qed
+  qed
+qed
+
+lemma minus_ms_aux:
+  assumes "basis_wf basis" "is_expansion_aux F f basis" "is_expansion_aux G g basis"
+  shows   "is_expansion_aux (minus_ms_aux F G) (\<lambda>x. f x - g x) basis"
+proof -
+  have "is_expansion_aux (minus_ms_aux F G) (\<lambda>x. f x + (- g x)) basis"
+    unfolding minus_ms_aux_def by (intro plus_ms_aux uminus_ms_aux assms)
+  thus ?thesis by simp
+qed
+
+lemma scale_shift_ms_aux:
+  assumes basis: "basis_wf (b # basis)"
+  assumes F: "is_expansion_aux F f (b # basis)"
+  assumes C: "is_expansion C basis"
+  shows   "is_expansion_aux (scale_shift_ms_aux (C, e) F) (\<lambda>x. eval C x * b x powr e * f x) (b # basis)"
+  using F
+proof (coinduction arbitrary: F f rule: is_expansion_aux_coinduct)
+  case (is_expansion_aux F f)
+  note IH = is_expansion_aux
+  show ?case
+  proof (cases F rule: ms_aux_cases)
+    case MSLNil
+    with IH show ?thesis 
+      by (auto simp: scale_shift_ms_aux_def elim!: is_expansion_aux.cases eventually_mono)
+  next
+    case (MSLCons C' e' xs)
+    with IH have IH: "is_expansion_aux xs (\<lambda>x. f x - eval C' x * b x powr e') (b # basis)"
+                     "is_expansion C' basis" "ms_exp_gt e' (ms_aux_hd_exp xs)"
+                     "\<And>e''. e' < e'' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e'')"
+      by (auto elim: is_expansion_aux_MSLCons)
+    
+    have "(\<lambda>x. eval C x * b x powr e * f x) \<in> o(\<lambda>x. b x powr e'')" if "e'' > e + e'" for e''
+    proof -
+      define d where "d = (e'' - e - e') / 2"
+      from that have d: "d > 0" by (simp add: d_def)
+      have "(\<lambda>x. eval C x * b x powr e * f x) \<in> o(\<lambda>x. b x powr d * b x powr e * b x powr (e' + d))"
+        using d basis
+        by (intro landau_o.small.mult[OF landau_o.small_big_mult] 
+              is_expansion_imp_smallo[OF _ C] IH) (simp_all add: basis_wf_Cons)
+      also have "\<dots> = o(\<lambda>x. b x powr e'')" by (simp add: d_def powr_add [symmetric])
+      finally show ?thesis .
+    qed
+    moreover have "ms_exp_gt (e + e') (ms_aux_hd_exp (scale_shift_ms_aux (C, e) xs))"
+      using IH(3) by (cases "ms_aux_hd_exp xs") (simp_all add: hd_exp_basis)
+    ultimately show ?thesis using basis IH(1,2)
+      by (intro disjI2, inst_existentials "scale_shift_ms_aux (C, e) xs" b "e + e'" 
+                          basis "C * C'" xs "\<lambda>x. f x - eval C' x * b x powr e'")
+         (simp_all add: scale_shift_ms_aux_MSLCons MSLCons is_expansion_mult basis_wf_Cons
+                        algebra_simps powr_add C)
+  qed
+qed
+
+lemma times_ms_aux:
+  assumes basis: "basis_wf basis"
+  assumes F: "is_expansion_aux F f basis" and G: "is_expansion_aux G g basis"
+  shows   "is_expansion_aux (times_ms_aux F G) (\<lambda>x. f x * g x) basis"
+  using F G
+proof (coinduction arbitrary: F f G g rule: is_expansion_aux_coinduct_upto)
+  case (B F f G g)
+  note IH = this
+  show ?case
+  proof (cases "F = MSLNil \<or> G = MSLNil")
+    case True
+    with IH have "eventually (\<lambda>x. f x = 0) at_top \<or> eventually (\<lambda>x. g x = 0) at_top"
+      and length: "length basis = Suc (expansion_level TYPE('a))"
+      by (auto elim: is_expansion_aux.cases)
+    from this(1) have "eventually (\<lambda>x. f x * g x = 0) at_top" by (auto elim!: eventually_mono)
+    with length True show ?thesis by auto
+  next
+    case False
+    from False obtain C1 e1 F' where F: "F = MSLCons (C1, e1) F'" 
+      by (cases F rule: ms_aux_cases) simp_all
+    from False obtain C2 e2 G' where G: "G = MSLCons (C2, e2) G'" 
+      by (cases G rule: ms_aux_cases) simp_all
+    
+    from IH(1) F obtain b basis' where
+      basis': "basis = b # basis'" and F':
+      "is_expansion_aux F' (\<lambda>x. f x - eval C1 x * b x powr e1) (b # basis')"
+      "is_expansion C1 basis'" "\<And>e'. e' > e1 \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')"
+      "ms_exp_gt e1 (ms_aux_hd_exp F')"
+      by (auto elim!: is_expansion_aux_MSLCons)
+    from IH(2) G basis' have G':
+      "is_expansion_aux G' (\<lambda>x. g x - eval C2 x * b x powr e2) (b # basis')"
+      "is_expansion C2 basis'" "\<And>e'. e' > e2 \<Longrightarrow> g \<in> o(\<lambda>x. b x powr e')"
+      "ms_exp_gt e2 (ms_aux_hd_exp G')"
+      by (auto elim!: is_expansion_aux_MSLCons)  
+    let ?P = "(\<lambda>xs'' fa'' basisa''. \<exists>F f G. xs'' = times_ms_aux F G \<and>
+            (\<exists>g. fa'' = (\<lambda>x. f x * g x) \<and> basisa'' = b # basis' \<and> is_expansion_aux F f
+                  (b # basis') \<and> is_expansion_aux G g (b # basis')))"
+    have "ms_closure ?P (plus_ms_aux (scale_shift_ms_aux (C1, e1) G') (times_ms_aux F' G))
+            (\<lambda>x. eval C1 x * b x powr e1 * (g x - eval C2 x * b x powr e2) +
+                 (f x - eval C1 x * b x powr e1) * g x) (b # basis')"
+      (is "ms_closure _ ?zs _ _") using IH(2) basis' basis
+      by (intro ms_closure_add ms_closure_embed'[OF scale_shift_ms_aux] 
+            ms_closure_mult[OF ms_closure_embed' ms_closure_embed'] F' G') simp_all
+    hence "ms_closure ?P (plus_ms_aux (scale_shift_ms_aux (C1, e1) G') (times_ms_aux F' G))
+            (\<lambda>x. f x * g x - eval C1 x * eval C2 x * b x powr (e1 + e2)) (b # basis')"
+      by (simp add: algebra_simps powr_add)
+    moreover have "(\<lambda>x. f x * g x) \<in> o(\<lambda>x. b x powr e)" if "e > e1 + e2" for e
+    proof -
+      define d where "d = (e - e1 - e2) / 2"
+      from that have "d > 0" by (simp add: d_def)
+      hence "(\<lambda>x. f x * g x) \<in> o(\<lambda>x. b x powr (e1 + d) * b x powr (e2 + d))"
+        by (intro landau_o.small.mult F' G') simp_all
+      also have "\<dots> = o(\<lambda>x. b x powr e)"
+        by (simp add: d_def powr_add [symmetric])
+      finally show ?thesis .
+    qed
+    moreover from \<open>ms_exp_gt e1 _\<close> \<open>ms_exp_gt e2 _\<close>
+      have "ms_exp_gt (e1 + e2) (ms_aux_hd_exp ?zs)"
+        by (cases "ms_aux_hd_exp F'"; cases "ms_aux_hd_exp G'") 
+           (simp_all add: hd_exp_times hd_exp_plus hd_exp_basis G)    
+    ultimately show ?thesis using F'(2) G'(2) basis' basis
+      by (simp add: times_ms_aux_MSLCons basis_wf_Cons is_expansion_mult F G)
+  qed
+qed (insert basis, simp_all)
+
+
+
+
+lemma powser_ms_aux_MSLNil_iff [simp]: "powser_ms_aux cs f = MSLNil \<longleftrightarrow> cs = MSLNil"
+  by (subst powser_ms_aux.code) (simp split: msllist.splits)
+
+lemma powser_ms_aux_MSLNil [simp]: "powser_ms_aux MSLNil f = MSLNil"
+  by (subst powser_ms_aux.code) simp
+
+lemma powser_ms_aux_MSLNil' [simp]: 
+  "powser_ms_aux (MSLCons c cs) MSLNil = MSLCons (const_expansion c, 0) MSLNil"
+  by (subst powser_ms_aux.code) simp
+
+lemma powser_ms_aux_MSLCons: 
+  "powser_ms_aux (MSLCons c cs) f = MSLCons (const_expansion c, 0) (times_ms_aux f (powser_ms_aux cs f))"
+  by (subst powser_ms_aux.code) simp
+
+lemma hd_exp_powser: "ms_aux_hd_exp (powser_ms_aux cs f) = (if cs = MSLNil then None else Some 0)"
+  by (subst powser_ms_aux.code) (simp split: msllist.splits)
+  
+lemma powser_ms_aux:
+  assumes "convergent_powser cs" and basis: "basis_wf basis"
+  assumes G: "is_expansion_aux G g basis" "ms_exp_gt 0 (ms_aux_hd_exp G)"
+  shows   "is_expansion_aux (powser_ms_aux cs G) (powser cs \<circ> g) basis"
+using assms(1)
+proof (coinduction arbitrary: cs rule: is_expansion_aux_coinduct_upto)
+  case (B cs)
+  show ?case
+  proof (cases cs)
+    case MSLNil
+    with G show ?thesis by (auto simp: is_expansion_aux_expansion_level)
+  next
+    case (MSLCons c cs')
+    from is_expansion_aux_basis_nonempty[OF G(1)]
+      obtain b basis' where basis': "basis = b # basis'" by (cases basis) simp_all
+    from B have conv: "convergent_powser cs'" by (auto simp: MSLCons dest: convergent_powser_stl)  
+    from basis basis' have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+    hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+    
+    from G(2) have "g \<in> o(\<lambda>x. hd basis x powr 0)"
+      by (intro is_expansion_aux_imp_smallo[OF G(1)]) simp
+    with basis' have "g \<in> o(\<lambda>x. b x powr 0)" by simp
+    also have "(\<lambda>x. b x powr 0) \<in> \<Theta>(\<lambda>x. 1)" 
+       using b_pos by (intro bigthetaI_cong) (auto elim!: eventually_mono)
+    finally have g_limit: "(g \<longlongrightarrow> 0) at_top" by (auto dest: smalloD_tendsto)
+  
+    let ?P = "(\<lambda>xs'' fa'' basisa''. \<exists>cs. xs'' = powser_ms_aux cs G \<and>
+                     fa'' = powser cs \<circ> g \<and> basisa'' = b # basis' \<and> convergent_powser cs)"
+    have "ms_closure ?P (times_ms_aux G (powser_ms_aux cs' G)) 
+            (\<lambda>x. g x * (powser cs' \<circ> g) x) basis" using conv
+      by (intro ms_closure_mult [OF ms_closure_embed' ms_closure_embed] G)
+         (auto simp add: basis' o_def)
+    also have "(\<lambda>x. g x * (powser cs' \<circ> g) x) = (\<lambda>x. x * powser cs' x) \<circ> g"
+      by (simp add: o_def)
+    finally have "ms_closure ?P (times_ms_aux G (powser_ms_aux cs' G)) 
+                    (\<lambda>x. (powser cs \<circ> g) x - c * b x powr 0) basis" unfolding o_def
+    proof (rule ms_closure_cong)
+      note b_pos
+      moreover have "eventually (\<lambda>x. powser cs (g x) = g x * powser cs' (g x) + c) at_top"
+      proof -
+        from B have "eventually (\<lambda>x. powser cs x = x * powser cs' x + c) (nhds 0)"
+          by (simp add: MSLCons powser_MSLCons)
+        from this and g_limit show ?thesis by (rule eventually_compose_filterlim)
+      qed
+      ultimately show "eventually (\<lambda>x. g x * powser cs' (g x) = 
+                         powser cs (g x) - c * b x powr 0) at_top"
+        by eventually_elim simp
+    qed 
+    moreover from basis G have "is_expansion (const_expansion c :: 'a) basis'"
+      by (intro is_expansion_const)
+         (auto dest: is_expansion_aux_expansion_level simp: basis' basis_wf_Cons)
+    moreover have "powser cs \<circ> g \<in> o(\<lambda>x. b x powr e)" if "e > 0" for e
+    proof -
+      have "((powser cs \<circ> g) \<longlongrightarrow> powser cs 0) at_top" unfolding o_def
+        by (intro isCont_tendsto_compose[OF _ g_limit] isCont_powser B)
+      hence "powser cs \<circ> g \<in> O(\<lambda>_. 1)" 
+        by (intro bigoI_tendsto[where c = "powser cs 0"]) (simp_all add: o_def)
+      also from b_pos have  "O(\<lambda>_. 1) = O(\<lambda>x. b x powr 0)" 
+        by (intro landau_o.big.cong) (auto elim: eventually_mono)
+      also from that b_limit have "(\<lambda>x. b x powr 0) \<in> o(\<lambda>x. b x powr e)"
+        by (subst powr_smallo_iff) simp_all
+      finally show ?thesis .
+    qed
+    moreover from G have "ms_exp_gt 0 (ms_aux_hd_exp (times_ms_aux G (powser_ms_aux cs' G)))"
+      by (cases "ms_aux_hd_exp G") (simp_all add: hd_exp_times MSLCons hd_exp_powser)
+    ultimately show ?thesis
+      by (simp add: MSLCons powser_ms_aux_MSLCons basis')
+  qed
+qed (insert assms, simp_all)
+  
+lemma powser_ms_aux':
+  assumes powser: "convergent_powser' cs f" and basis: "basis_wf basis"
+  assumes G: "is_expansion_aux G g basis" "ms_exp_gt 0 (ms_aux_hd_exp G)"
+  shows   "is_expansion_aux (powser_ms_aux cs G) (f \<circ> g) basis"
+proof (rule is_expansion_aux_cong)
+  from is_expansion_aux_basis_nonempty[OF G(1)] basis
+    have "filterlim (hd basis) at_top at_top" by (cases basis) (simp_all add: basis_wf_Cons)
+  hence pos: "eventually (\<lambda>x. hd basis x > 0) at_top" by (simp add: filterlim_at_top_dense)
+  from G(2) have "g \<in> o(\<lambda>x. hd basis x powr 0)"
+    by (intro is_expansion_aux_imp_smallo[OF G(1)]) simp
+  also have "(\<lambda>x. hd basis x powr 0) \<in> \<Theta>(\<lambda>x. 1)" 
+     using pos by (intro bigthetaI_cong) (auto elim!: eventually_mono)
+  finally have g_limit: "(g \<longlongrightarrow> 0) at_top" by (auto dest: smalloD_tendsto)
+  
+  from powser have "eventually (\<lambda>x. powser cs x = f x) (nhds 0)"
+    by (rule convergent_powser'_eventually)
+  from this and g_limit show "eventually (\<lambda>x. (powser cs \<circ> g) x = (f \<circ> g) x) at_top" 
+    unfolding o_def by (rule eventually_compose_filterlim)
+qed (rule assms powser_ms_aux convergent_powser'_imp_convergent_powser)+
+
+lemma inverse_ms_aux:
+  assumes basis: "basis_wf basis"
+  assumes F: "is_expansion_aux F f basis" "trimmed_ms_aux F"
+  shows   "is_expansion_aux (inverse_ms_aux F) (\<lambda>x. inverse (f x)) basis"
+proof (cases F rule: ms_aux_cases)
+  case (MSLCons C e F')
+  from F MSLCons obtain b basis' where F': "basis = b # basis'" "trimmed C"
+    "is_expansion_aux F' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
+    "is_expansion C basis'" "(\<And>e'. e < e' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e'))" 
+    "ms_exp_gt e (ms_aux_hd_exp F')"
+    by (auto elim!: is_expansion_aux_MSLCons simp: trimmed_ms_aux_def)
+  define f' where "f' = (\<lambda>x. f x - eval C x * b x powr e)"
+  from basis F' have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+  moreover from F'(6)
+    have "ms_exp_gt 0 (ms_aux_hd_exp (scale_shift_ms_aux (inverse C, - e) F'))"
+    by (cases "ms_aux_hd_exp F'") (simp_all add: hd_exp_basis)
+  ultimately have
+       "is_expansion_aux (inverse_ms_aux F) 
+          (\<lambda>x. eval (inverse C) x * b x powr (-e) * 
+                 ((\<lambda>x. inverse (1 + x)) \<circ> (\<lambda>x. eval (inverse C) x * b x powr (-e) * f' x)) x)
+          (b # basis')" (is "is_expansion_aux _ ?h _")
+    unfolding MSLCons inverse_ms_aux.simps Let_def fst_conv snd_conv f'_def eval_inverse [symmetric]
+    using basis F(2) F'(1,3,4)
+    by (intro scale_shift_ms_aux powser_ms_aux' is_expansion_inverse)
+       (simp_all add: convergent_powser'_geometric basis_wf_Cons trimmed_ms_aux_def MSLCons)
+  also have "b # basis' = basis" by (simp add: F')
+  finally show ?thesis
+  proof (rule is_expansion_aux_cong, goal_cases)
+    case 1
+    from basis F' have "eventually (\<lambda>x. eval C x \<noteq> 0) at_top" 
+      by (intro trimmed_imp_eventually_nz[of basis']) (simp_all add: basis_wf_Cons)
+    with b_pos show ?case by eventually_elim (simp add: o_def field_simps powr_minus f'_def)
+  qed
+qed (insert assms, auto simp: trimmed_ms_aux_def)
+
+lemma hd_exp_inverse: 
+  "xs \<noteq> MSLNil \<Longrightarrow> ms_aux_hd_exp (inverse_ms_aux xs) = map_option uminus (ms_aux_hd_exp xs)"
+  by (cases xs) (auto simp: Let_def hd_exp_basis hd_exp_powser inverse_ms_aux.simps)
+
+lemma eval_pos_if_dominant_term_pos:
+  assumes "basis_wf basis" "is_expansion F basis" "trimmed F"
+          "fst (dominant_term F) > 0"
+  shows   "eventually (\<lambda>x. eval F x > 0) at_top"
+proof -
+  have "eval F \<sim>[at_top] eval_monom (dominant_term F) basis"
+    by (intro dominant_term assms)
+  from trimmed_imp_eventually_sgn[OF assms(1-3)]
+    have "\<forall>\<^sub>F x in at_top. sgn (eval F x) = sgn (fst (dominant_term F))" .
+  moreover from assms
+    have "eventually (\<lambda>x. eval_monom (dominant_term F) basis x > 0) at_top"
+    by (intro eval_monom_pos)
+  ultimately show ?thesis by eventually_elim (insert assms, simp add: sgn_if split: if_splits)
+qed
+
+lemma eval_pos_if_dominant_term_pos':
+  assumes "basis_wf basis" "trimmed_ms_aux F" "is_expansion_aux F f basis" 
+          "fst (dominant_term_ms_aux F) > 0"
+  shows   "eventually (\<lambda>x. f x > 0) at_top"
+proof -
+  have "f \<sim>[at_top] eval_monom (dominant_term_ms_aux F) basis"
+    by (intro dominant_term_ms_aux assms)
+  from dominant_term_ms_aux(2)[OF assms(1-3)]
+    have "\<forall>\<^sub>F x in at_top. sgn (f x) = sgn (fst (dominant_term_ms_aux F))" .
+  moreover from assms
+    have "eventually (\<lambda>x. eval_monom (dominant_term_ms_aux F) basis x > 0) at_top"
+    by (intro eval_monom_pos)
+  ultimately show ?thesis by eventually_elim (insert assms, simp add: sgn_if split: if_splits)
+qed
+  
+lemma eval_neg_if_dominant_term_neg':
+  assumes "basis_wf basis" "trimmed_ms_aux F" "is_expansion_aux F f basis" 
+          "fst (dominant_term_ms_aux F) < 0"
+  shows   "eventually (\<lambda>x. f x < 0) at_top"
+proof -
+  have "f \<sim>[at_top] eval_monom (dominant_term_ms_aux F) basis"
+    by (intro dominant_term_ms_aux assms)
+  from dominant_term_ms_aux(2)[OF assms(1-3)]
+    have "\<forall>\<^sub>F x in at_top. sgn (f x) = sgn (fst (dominant_term_ms_aux F))" .
+  moreover from assms
+  have "eventually (\<lambda>x. eval_monom (dominant_term_ms_aux F) basis x < 0) at_top"
+    by (intro eval_monom_neg)
+  ultimately show ?thesis by eventually_elim (insert assms, simp add: sgn_if split: if_splits)
+qed
+
+lemma fst_dominant_term_ms_aux_MSLCons: 
+  "fst (dominant_term_ms_aux (MSLCons x xs)) = fst (dominant_term (fst x))"
+  by (auto simp: dominant_term_ms_aux_def split: prod.splits)
+
+lemma powr_ms_aux:
+  assumes basis: "basis_wf basis"
+  assumes F: "is_expansion_aux F f basis" "trimmed_ms_aux F" "fst (dominant_term_ms_aux F) > 0"
+  shows   "is_expansion_aux (powr_ms_aux abort F p) (\<lambda>x.  f x powr p) basis"
+proof (cases F rule: ms_aux_cases)
+  case (MSLCons C e F')
+  from F MSLCons obtain b basis' where F': "basis = b # basis'" "trimmed C" "fst (dominant_term C) > 0"
+    "is_expansion_aux F' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
+    "is_expansion C basis'" "(\<And>e'. e < e' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e'))"
+    "ms_exp_gt e (ms_aux_hd_exp F')"
+    by (auto elim!: is_expansion_aux_MSLCons simp: trimmed_ms_aux_def dominant_term_ms_aux_def 
+             split: prod.splits)
+  define f' where "f' = (\<lambda>x. f x - eval C x * b x powr e)"
+  from basis F' have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+  moreover from F' have "ms_exp_gt 0 (ms_aux_hd_exp (scale_shift_ms_aux (inverse C, - e) F'))"
+    by (cases "ms_aux_hd_exp F'") (simp_all add: hd_exp_basis)
+  ultimately have
+       "is_expansion_aux (powr_ms_aux abort F p) 
+          (\<lambda>x. eval (powr_expansion abort C p) x * b x powr (e * p) * 
+                 ((\<lambda>x. (1 + x) powr p) \<circ> (\<lambda>x. eval (inverse C) x * b x powr (-e) * f' x)) x)
+          (b # basis')" (is "is_expansion_aux _ ?h _")
+    unfolding MSLCons powr_ms_aux.simps Let_def fst_conv snd_conv f'_def eval_inverse [symmetric]
+        using basis F'(1-5)
+    by (intro scale_shift_ms_aux powser_ms_aux' is_expansion_inverse is_expansion_powr)
+       (simp_all add: MSLCons basis_wf_Cons convergent_powser'_gbinomial)
+  also have "b # basis' = basis" by (simp add: F')
+  finally show ?thesis
+  proof (rule is_expansion_aux_cong, goal_cases)
+    case 1
+    from basis F' have "eventually (\<lambda>x. eval C x > 0) at_top" 
+      by (intro eval_pos_if_dominant_term_pos[of basis']) (simp_all add: basis_wf_Cons)
+    moreover from basis F have "eventually (\<lambda>x. f x > 0) at_top"
+      by (intro eval_pos_if_dominant_term_pos'[of basis F])
+    ultimately show ?case using b_pos
+      by eventually_elim 
+         (simp add: powr_powr [symmetric] powr_minus powr_mult powr_divide f'_def field_simps)
+  qed
+qed (insert assms, auto simp: trimmed_ms_aux_def)
+
+lemma power_ms_aux:
+  assumes basis: "basis_wf basis"
+  assumes F: "is_expansion_aux F f basis" "trimmed_ms_aux F"
+  shows   "is_expansion_aux (power_ms_aux abort F n) (\<lambda>x. f x ^ n) basis"
+proof (cases F rule: ms_aux_cases)
+  case (MSLCons C e F')
+  from F MSLCons obtain b basis' where F': "basis = b # basis'" "trimmed C"
+    "is_expansion_aux F' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
+    "is_expansion C basis'" "(\<And>e'. e < e' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e'))" 
+    "ms_exp_gt e (ms_aux_hd_exp F')"
+    by (auto elim!: is_expansion_aux_MSLCons simp: trimmed_ms_aux_def dominant_term_ms_aux_def 
+             split: prod.splits)
+  define f' where "f' = (\<lambda>x. f x - eval C x * b x powr e)"
+  from basis F' have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+  moreover have "ms_exp_gt 0 (ms_aux_hd_exp (scale_shift_ms_aux (inverse C, - e) F'))"
+    using F'(6) by (cases "ms_aux_hd_exp F'") (simp_all add: hd_exp_basis)
+  ultimately have
+       "is_expansion_aux (power_ms_aux abort F n) 
+          (\<lambda>x. eval (power_expansion abort C n) x * b x powr (e * real n) * 
+                 ((\<lambda>x. (1 + x) ^ n) \<circ> (\<lambda>x. eval (inverse C) x * b x powr (-e) * f' x)) x)
+          (b # basis')" (is "is_expansion_aux _ ?h _")
+    unfolding MSLCons power_ms_aux.simps Let_def fst_conv snd_conv f'_def eval_inverse [symmetric]
+        using basis F'(1-4)
+    by (intro scale_shift_ms_aux powser_ms_aux' is_expansion_inverse is_expansion_power)
+       (simp_all add: MSLCons basis_wf_Cons convergent_powser'_gbinomial')
+  also have "b # basis' = basis" by (simp add: F')
+  finally show ?thesis
+  proof (rule is_expansion_aux_cong, goal_cases)
+    case 1
+    from F'(1,2,4) basis have "eventually (\<lambda>x. eval C x \<noteq> 0) at_top"
+      using trimmed_imp_eventually_nz[of basis' C] by (simp add: basis_wf_Cons)
+    thus ?case using b_pos
+      by eventually_elim
+         (simp add: field_simps f'_def powr_minus powr_powr [symmetric] powr_realpow 
+                    power_mult_distrib [symmetric] power_divide [symmetric]
+               del: power_mult_distrib power_divide)
+  qed
+qed (insert assms, auto simp: trimmed_ms_aux_def)
+
+lemma snd_dominant_term_ms_aux_MSLCons:
+  "snd (dominant_term_ms_aux (MSLCons (C, e) xs)) = e # snd (dominant_term C)"
+  by (simp add: dominant_term_ms_aux_def case_prod_unfold)
+
+
+subsubsection \<open>Type-class instantiations\<close>
+
+datatype 'a ms = MS "('a \<times> real) msllist" "real \<Rightarrow> real"
+
+instantiation ms :: (uminus) uminus
+begin
+
+primrec uminus_ms where
+  "-(MS xs f) = MS (uminus_ms_aux xs) (\<lambda>x. -f x)"
+  
+instance ..
+end
+
+instantiation ms :: (plus) plus
+begin
+
+fun plus_ms :: "'a ms \<Rightarrow> 'a ms \<Rightarrow> 'a ms" where
+  "MS xs f + MS ys g = MS (plus_ms_aux xs ys) (\<lambda>x. f x + g x)"
+
+instance ..
+end
+
+instantiation ms :: ("{plus,uminus}") minus
+begin
+
+definition minus_ms :: "'a ms \<Rightarrow> 'a ms \<Rightarrow> 'a ms" where
+  "F - G = F + -(G :: 'a ms)"
+
+instance ..
+end
+
+instantiation ms :: ("{plus,times}") times
+begin
+
+fun times_ms :: "'a ms \<Rightarrow> 'a ms \<Rightarrow> 'a ms" where
+  "MS xs f * MS ys g = MS (times_ms_aux xs ys) (\<lambda>x. f x * g x)"
+
+instance ..
+end
+
+instantiation ms :: (multiseries) inverse
+begin
+
+fun inverse_ms :: "'a ms \<Rightarrow> 'a ms" where
+  "inverse_ms (MS xs f) = MS (inverse_ms_aux xs) (\<lambda>x. inverse (f x))"
+
+fun divide_ms :: "'a ms \<Rightarrow> 'a ms \<Rightarrow> 'a ms" where
+  "divide_ms (MS xs f) (MS ys g) = MS (times_ms_aux xs (inverse_ms_aux ys)) (\<lambda>x. f x / g x)"
+
+instance ..
+end
+
+
+instantiation ms :: (multiseries) multiseries
+begin
+
+definition expansion_level_ms :: "'a ms itself \<Rightarrow> nat" where
+  expansion_level_ms_def [simp]: "expansion_level_ms _ = Suc (expansion_level (TYPE('a)))"
+
+definition zero_expansion_ms :: "'a ms" where
+  "zero_expansion = MS MSLNil (\<lambda>_. 0)"
+
+definition const_expansion_ms :: "real \<Rightarrow> 'a ms" where
+  "const_expansion c = MS (const_ms_aux c) (\<lambda>_. c)"
+
+primrec is_expansion_ms :: "'a ms \<Rightarrow> basis \<Rightarrow> bool" where
+  "is_expansion (MS xs f) = is_expansion_aux xs f"
+
+lemma is_expansion_ms_def': "is_expansion F = (case F of MS xs f \<Rightarrow> is_expansion_aux xs f)"
+  by (simp add: is_expansion_ms_def split: ms.splits)
+
+primrec eval_ms :: "'a ms \<Rightarrow> real \<Rightarrow> real" where
+  "eval_ms (MS _ f) = f"
+  
+lemma eval_ms_def': "eval F = (case F of MS _ f \<Rightarrow> f)"
+  by (cases F) simp_all
+
+primrec powr_expansion_ms :: "bool \<Rightarrow> 'a ms \<Rightarrow> real \<Rightarrow> 'a ms" where
+  "powr_expansion_ms abort (MS xs f) p = MS (powr_ms_aux abort xs p) (\<lambda>x. f x powr p)"
+  
+primrec power_expansion_ms :: "bool \<Rightarrow> 'a ms \<Rightarrow> nat \<Rightarrow> 'a ms" where
+  "power_expansion_ms abort (MS xs f) n = MS (power_ms_aux abort xs n) (\<lambda>x. f x ^ n)"
+
+primrec trimmed_ms :: "'a ms \<Rightarrow> bool" where
+  "trimmed_ms (MS xs _) \<longleftrightarrow> trimmed_ms_aux xs"
+
+primrec dominant_term_ms :: "'a ms \<Rightarrow> monom" where
+  "dominant_term_ms (MS xs _) = dominant_term_ms_aux xs"
+
+lemma length_dominant_term_ms: 
+  "length (snd (dominant_term (F :: 'a ms))) = Suc (expansion_level TYPE('a))"
+  by (cases F) (simp_all add: length_dominant_term)  
+
+instance 
+proof (standard, goal_cases length_basis zero const uminus plus minus times 
+         inverse divide powr power smallo smallomega trimmed dominant dominant_bigo)
+  case (smallo basis F b e)
+  from \<open>is_expansion F basis\<close> obtain xs f where F: "F = MS xs f" "is_expansion_aux xs f basis"
+    by (simp add: is_expansion_ms_def' split: ms.splits)
+  from F(2) have nonempty: "basis \<noteq> []" by (rule is_expansion_aux_basis_nonempty)
+  with smallo have filterlim_hd_basis: "filterlim (hd basis) at_top at_top"
+    by (cases basis) (simp_all add: basis_wf_Cons)
+  from F(2) obtain e' where "f \<in> o(\<lambda>x. hd basis x powr e')"
+    by (erule is_expansion_aux_imp_smallo')
+  also from smallo nonempty filterlim_hd_basis have "(\<lambda>x. hd basis x powr e') \<in> o(\<lambda>x. b x powr e)"
+    by (intro ln_smallo_imp_flat) auto
+  finally show ?case by (simp add: F)
+next
+  case (smallomega basis F b e)
+  obtain xs f where F: "F = MS xs f" by (cases F) simp_all
+  from this smallomega have *: "is_expansion_aux xs f basis" by simp
+  with smallomega F have "f \<in> \<omega>(\<lambda>x. b x powr e)"
+    by (intro is_expansion_aux_imp_smallomega [OF _ *])
+       (simp_all add: is_expansion_ms_def' split: ms.splits)
+  with F show ?case by simp
+next
+  case (minus basis F G)
+  thus ?case
+    by (simp add: minus_ms_def is_expansion_ms_def' add_uminus_conv_diff [symmetric] 
+                  plus_ms_aux uminus_ms_aux del: add_uminus_conv_diff split: ms.splits)
+next
+  case (divide basis F G)
+  have "G / F = G * inverse F" by (cases F; cases G) (simp add: divide_inverse)
+  with divide show ?case
+    by (simp add: is_expansion_ms_def' divide_inverse times_ms_aux inverse_ms_aux split: ms.splits)
+next
+  fix c :: real
+  show "trimmed (const_expansion c :: 'a ms) \<longleftrightarrow> c \<noteq> 0"
+    by (simp add: const_expansion_ms_def trimmed_ms_aux_def const_ms_aux_def 
+                  trimmed_const_expansion split: msllist.splits)
+next
+  fix F :: "'a ms" assume "trimmed F"
+  thus "fst (dominant_term F) \<noteq> 0"
+    by (cases F) (auto simp: dominant_term_ms_aux_def trimmed_ms_aux_MSLCons case_prod_unfold 
+                    trimmed_imp_dominant_term_nz split: msllist.splits)
+next
+  fix F :: "'a ms"
+  have "times_ms_aux (MSLCons (const_expansion 1, 0) MSLNil) xs = xs" for xs :: "('a \<times> real) msllist"
+  proof (coinduction arbitrary: xs rule: msllist.coinduct_upto)
+    case Eq_real_prod_msllist
+    have "map_prod (( *) (const_expansion 1 :: 'a)) ((+) (0::real)) = (\<lambda>x. x)"
+      by (rule ext) (simp add: map_prod_def times_const_expansion_1 case_prod_unfold)
+    moreover have "mslmap \<dots> = (\<lambda>x. x)" by (rule ext) (simp add: msllist.map_ident)
+    ultimately have "scale_shift_ms_aux (const_expansion 1 :: 'a, 0) = (\<lambda>x. x)"
+      by (simp add: scale_shift_ms_aux_conv_mslmap)
+    thus ?case
+      by (cases xs rule: ms_aux_cases)
+         (auto simp: times_ms_aux_MSLCons times_const_expansion_1
+            times_ms_aux.corec.cong_refl)
+  qed
+  thus "const_expansion 1 * F = F"
+    by (cases F) (simp add: const_expansion_ms_def const_ms_aux_def)
+next
+  fix F :: "'a ms"
+  show "fst (dominant_term (- F)) = - fst (dominant_term F)"
+       "trimmed (-F) \<longleftrightarrow> trimmed F"
+     by (cases F; simp add: dominant_term_ms_aux_def case_prod_unfold uminus_ms_aux_MSLCons
+           trimmed_ms_aux_def split: msllist.splits)+
+next
+  fix F :: "'a ms"
+  show "zero_expansion + F = F" "F + zero_expansion = F"
+    by (cases F; simp add: zero_expansion_ms_def)+
+qed (auto simp: eval_ms_def' const_expansion_ms_def trimmed_ms_aux_imp_nz is_expansion_ms_def' 
+                 const_ms_aux uminus_ms_aux plus_ms_aux times_ms_aux inverse_ms_aux 
+                 is_expansion_aux_expansion_level dominant_term_ms_aux length_dominant_term_ms
+                 minus_ms_def powr_ms_aux power_ms_aux zero_expansion_ms_def 
+                 is_expansion_aux.intros(1) dominant_term_ms_aux_bigo
+          split: ms.splits)
+
+end
+
+
+subsubsection \<open>Changing the evaluation function of a multiseries\<close>
+
+definition modify_eval :: "(real \<Rightarrow> real) \<Rightarrow> 'a ms \<Rightarrow> 'a ms" where
+  "modify_eval f F = (case F of MS xs _ \<Rightarrow> MS xs f)"
+
+lemma eval_modify_eval [simp]: "eval (modify_eval f F) = f"
+  by (cases F) (simp add: modify_eval_def)
+
+
+subsubsection \<open>Scaling expansions\<close>
+
+definition scale_ms :: "real \<Rightarrow> 'a :: multiseries \<Rightarrow> 'a" where
+  "scale_ms c F = const_expansion c * F"
+
+lemma scale_ms_real [simp]: "scale_ms c (c' :: real) = c * c'"
+  by (simp add: scale_ms_def)
+
+definition scale_ms_aux :: "real \<Rightarrow> ('a :: multiseries \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
+  "scale_ms_aux c xs = scale_shift_ms_aux (const_expansion c, 0) xs"
+
+
+lemma scale_ms_aux_eq_MSLNil_iff [simp]: "scale_ms_aux x xs = MSLNil \<longleftrightarrow> xs = MSLNil"
+  unfolding scale_ms_aux_def by simp
+
+lemma times_ms_aux_singleton:
+  "times_ms_aux (MSLCons (c, e) MSLNil) xs = scale_shift_ms_aux (c, e) xs"
+proof (coinduction arbitrary: xs rule: msllist.coinduct_strong)
+  case (Eq_msllist xs)
+  thus ?case
+    by (cases xs rule: ms_aux_cases)
+       (simp_all add: scale_shift_ms_aux_def times_ms_aux_MSLCons)
+qed
+
+lemma scale_ms [simp]: "scale_ms c (MS xs f) = MS (scale_ms_aux c xs) (\<lambda>x. c * f x)"
+  by (simp add: scale_ms_def scale_ms_aux_def const_expansion_ms_def const_ms_aux_def 
+        times_ms_aux_singleton)
+
+lemma scale_ms_aux_MSLNil [simp]: "scale_ms_aux c MSLNil = MSLNil" 
+  by (simp add: scale_ms_aux_def)
+
+lemma scale_ms_aux_MSLCons: 
+  "scale_ms_aux c (MSLCons (c', e) xs) = MSLCons (scale_ms c c', e) (scale_ms_aux c xs)"
+  by (simp add: scale_ms_aux_def scale_shift_ms_aux_MSLCons scale_ms_def)
+
+
+subsubsection \<open>Additional convenience rules\<close>
+
+lemma trimmed_pos_real_iff [simp]: "trimmed_pos (x :: real) \<longleftrightarrow> x > 0"
+  by (auto simp: trimmed_pos_def)
+
+lemma trimmed_pos_ms_iff: 
+  "trimmed_pos (MS xs f) \<longleftrightarrow> (case xs of MSLNil \<Rightarrow> False | MSLCons x xs \<Rightarrow> trimmed_pos (fst x))"
+  by (auto simp add: trimmed_pos_def dominant_term_ms_aux_def trimmed_ms_aux_def
+           split: msllist.splits prod.splits)
+
+lemma not_trimmed_pos_MSLNil [simp]: "\<not>trimmed_pos (MS MSLNil f)"
+  by (simp add: trimmed_pos_ms_iff)
+
+lemma trimmed_pos_MSLCons [simp]: "trimmed_pos (MS (MSLCons x xs) f) = trimmed_pos (fst x)"
+  by (simp add: trimmed_pos_ms_iff)
+
+lemma drop_zero_ms:
+  assumes "eventually (\<lambda>x. eval C x = 0) at_top"
+  assumes "is_expansion (MS (MSLCons (C, e) xs) f) basis"
+  shows   "is_expansion (MS xs f) basis"
+  using assms by (auto simp: is_expansion_ms_def intro: drop_zero_ms_aux)
+
+lemma expansion_level_eq_Nil: "length [] = expansion_level TYPE(real)" by simp
+lemma expansion_level_eq_Cons: 
+  "length xs = expansion_level TYPE('a::multiseries) \<Longrightarrow>
+     length (x # xs) = expansion_level TYPE('a ms)" by simp
+
+lemma basis_wf_insert_ln: 
+  assumes "basis_wf [b]"
+  shows   "basis_wf [\<lambda>x. ln (b x)]" (is ?thesis1)
+          "basis_wf [b, \<lambda>x. ln (b x)]" (is ?thesis2)
+          "is_expansion (MS (MSLCons (1::real,1) MSLNil) (\<lambda>x. ln (b x))) [\<lambda>x. ln (b x)]" (is ?thesis3)
+proof -
+  have "ln \<in> o(\<lambda>x. x :: real)"
+    using ln_x_over_x_tendsto_0 by (intro smalloI_tendsto) auto
+  with assms show ?thesis1 ?thesis2
+    by (auto simp: basis_wf_Cons
+          intro!: landau_o.small.compose[of _ _ _ "\<lambda>x. ln (b x)"] filterlim_compose[OF ln_at_top])
+  from assms have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence ev: "eventually (\<lambda>x. b x > 1) at_top" by (simp add: filterlim_at_top_dense)
+  have "(\<lambda>x::real. x) \<in> \<Theta>(\<lambda>x. x powr 1)"
+    by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto
+  also have "(\<lambda>x::real. x powr 1) \<in> o(\<lambda>a. a powr e)" if "e > 1" for e
+    by (subst powr_smallo_iff) (auto simp: that filterlim_ident)
+  finally show ?thesis3 using assms ev
+    by (auto intro!: is_expansion_aux.intros landau_o.small.compose[of _ at_top _ "\<lambda>x. ln (b x)"]
+                     filterlim_compose[OF ln_at_top b_limit] is_expansion_real.intros
+             elim!: eventually_mono)
+qed
+
+lemma basis_wf_lift_modification:
+  assumes "basis_wf (b' # b # bs)" "basis_wf (b # bs')"
+  shows   "basis_wf (b' # b # bs')"
+  using assms by (simp add: basis_wf_many)
+
+lemma default_basis_wf: "basis_wf [\<lambda>x. x]" 
+  by (simp add: basis_wf_Cons filterlim_ident)
+
+  
+subsubsection \<open>Lifting expansions\<close>
+
+definition is_lifting where
+  "is_lifting f basis basis' \<longleftrightarrow> 
+     (\<forall>C. eval (f C) = eval C \<and> (is_expansion C basis \<longrightarrow> is_expansion (f C) basis') \<and>
+          trimmed (f C) = trimmed C \<and> fst (dominant_term (f C)) = fst (dominant_term C))"
+  
+lemma is_liftingD:
+  assumes "is_lifting f basis basis'"
+  shows   "eval (f C) = eval C" "is_expansion C basis \<Longrightarrow> is_expansion (f C) basis'"
+          "trimmed (f C) \<longleftrightarrow> trimmed C" "fst (dominant_term (f C)) = fst (dominant_term C)"
+  using assms [unfolded is_lifting_def] unfolding is_lifting_def by blast+
+  
+
+definition lift_ms :: "'a :: multiseries \<Rightarrow> 'a ms" where
+  "lift_ms C = MS (MSLCons (C, 0) MSLNil) (eval C)"
+
+lemma is_expansion_lift:
+  assumes "basis_wf (b # basis)" "is_expansion C basis"
+  shows   "is_expansion (lift_ms C) (b # basis)"
+proof -
+  from assms have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+  moreover from assms have "eval C \<in> o(\<lambda>x. b x powr e)" if "e > 0" for e
+    using that by (intro is_expansion_imp_smallo[OF _ assms(2)]) (simp_all add: basis_wf_Cons)
+  ultimately show ?thesis using assms
+    by (auto intro!: is_expansion_aux.intros simp: lift_ms_def is_expansion_length 
+             elim: eventually_mono)
+qed
+
+lemma eval_lift_ms [simp]: "eval (lift_ms C) = eval C"
+  by (simp add: lift_ms_def)
+
+lemma is_lifting_lift:
+  assumes "basis_wf (b # basis)"
+  shows   "is_lifting lift_ms basis (b # basis)"
+  using is_expansion_lift[OF assms] unfolding is_lifting_def
+  by (auto simp: lift_ms_def trimmed_ms_aux_MSLCons dominant_term_ms_aux_def case_prod_unfold)
+
+
+definition map_ms_aux :: 
+    "('a :: multiseries \<Rightarrow> 'b :: multiseries) \<Rightarrow> 
+       ('a \<times> real) msllist \<Rightarrow> ('b \<times> real) msllist" where
+  "map_ms_aux f xs = mslmap (\<lambda>(c,e). (f c, e)) xs"
+
+lemma map_ms_aux_MSLNil [simp]: "map_ms_aux f MSLNil = MSLNil"
+  by (simp add: map_ms_aux_def)
+
+lemma map_ms_aux_MSLCons: "map_ms_aux f (MSLCons (C, e) xs) = MSLCons (f C, e) (map_ms_aux f xs)"
+  by (simp add: map_ms_aux_def)
+
+lemma hd_exp_map [simp]: "ms_aux_hd_exp (map_ms_aux f xs) = ms_aux_hd_exp xs"
+  by (simp add: ms_aux_hd_exp_def map_ms_aux_def split: msllist.splits)
+
+lemma map_ms_altdef: "map_ms f (MS xs g) = MS (map_ms_aux f xs) g"
+  by (simp add: map_ms_aux_def map_prod_def)
+
+lemma eval_map_ms [simp]: "eval (map_ms f F) = eval F"
+  by (cases F) simp_all
+
+lemma is_expansion_map_aux:
+  fixes f :: "'a :: multiseries \<Rightarrow> 'b :: multiseries"
+  assumes "is_expansion_aux xs g (b # basis)"
+  assumes "\<And>C. is_expansion C basis \<Longrightarrow> is_expansion (f C) basis'"
+  assumes "length basis' = expansion_level TYPE('b)"
+  assumes "\<And>C. eval (f C) = eval C"
+  shows   "is_expansion_aux (map_ms_aux f xs) g (b # basis')"
+  using assms(1)
+proof (coinduction arbitrary: xs g rule: is_expansion_aux_coinduct)
+  case (is_expansion_aux xs g)
+  show ?case
+  proof (cases xs rule: ms_aux_cases)
+    case MSLNil
+    with is_expansion_aux show ?thesis
+      by (auto simp: assms elim: is_expansion_aux.cases)
+  next
+    case (MSLCons c e xs')
+    with is_expansion_aux show ?thesis
+      by (auto elim!: is_expansion_aux_MSLCons simp: map_ms_aux_MSLCons assms)
+  qed    
+qed
+
+lemma is_expansion_map:
+  fixes f :: "'a :: multiseries \<Rightarrow> 'b :: multiseries"
+    and F :: "'a ms"
+  assumes "is_expansion G (b # basis)"
+  assumes "\<And>C. is_expansion C basis \<Longrightarrow> is_expansion (f C) basis'"
+  assumes "\<And>C. eval (f C) = eval C"
+  assumes "length basis' = expansion_level TYPE('b)"
+  shows   "is_expansion (map_ms f G) (b # basis')"
+  using assms by (cases G, simp only: map_ms_altdef) (auto intro!: is_expansion_map_aux)
+
+lemma is_expansion_map_final: 
+  fixes f :: "'a :: multiseries \<Rightarrow> 'b :: multiseries"
+    and F :: "'a ms"
+  assumes "is_lifting f basis basis'"
+  assumes "is_expansion G (b # basis)"
+  assumes "length basis' = expansion_level TYPE('b)"
+  shows   "is_expansion (map_ms f G) (b # basis')"
+  by (rule is_expansion_map[OF assms(2)]) (insert assms, simp_all add: is_lifting_def)
+
+lemma fst_dominant_term_map_ms: 
+  "is_lifting f basis basis' \<Longrightarrow> fst (dominant_term (map_ms f C)) = fst (dominant_term C)"
+  by (cases C)
+     (simp add: dominant_term_ms_aux_def case_prod_unfold is_lifting_def split: msllist.splits)
+
+lemma trimmed_map_ms:
+  "is_lifting f basis basis' \<Longrightarrow> trimmed (map_ms f C) \<longleftrightarrow> trimmed C"
+  by (cases C) (simp add: trimmed_ms_aux_def is_lifting_def split: msllist.splits)
+
+lemma is_lifting_map: 
+  fixes f :: "'a :: multiseries \<Rightarrow> 'b :: multiseries"
+    and F :: "'a ms"
+  assumes "is_lifting f basis basis'"
+  assumes "length basis' = expansion_level TYPE('b)"
+  shows   "is_lifting (map_ms f) (b # basis) (b # basis')"
+  unfolding is_lifting_def
+  by (intro allI conjI impI is_expansion_map_final[OF assms(1)]) 
+     (insert assms, simp_all add: is_lifting_def fst_dominant_term_map_ms[OF assms(1)] 
+        trimmed_map_ms[OF assms(1)])
+
+lemma is_lifting_id: "is_lifting (\<lambda>x. x) basis basis"
+  by (simp add: is_lifting_def)
+
+lemma is_lifting_trans: 
+  "is_lifting f basis basis' \<Longrightarrow> is_lifting g basis' basis'' \<Longrightarrow> is_lifting (\<lambda>x. g (f x)) basis basis''"
+  by (auto simp: is_lifting_def)
+
+lemma is_expansion_X: "is_expansion (MS (MSLCons (1 :: real, 1) MSLNil) (\<lambda>x. x)) [\<lambda>x. x]"
+proof -
+  have "(\<lambda>x::real. x) \<in> \<Theta>(\<lambda>x. x powr 1)"
+    by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto
+  also have "(\<lambda>x::real. x powr 1) \<in> o(\<lambda>a. a powr e)" if "e > 1" for e
+    by (subst powr_smallo_iff) (auto simp: that filterlim_ident)
+  finally show ?thesis  
+    by (auto intro!: is_expansion_aux.intros is_expansion_real.intros 
+                     eventually_mono[OF eventually_gt_at_top[of "0::real"]])
+qed
+
+
+inductive expands_to :: "(real \<Rightarrow> real) \<Rightarrow> 'a :: multiseries \<Rightarrow> basis \<Rightarrow> bool" 
+    (infix "(expands'_to)" 50) where
+  "is_expansion F basis \<Longrightarrow> eventually (\<lambda>x. eval F x = f x) at_top \<Longrightarrow> (f expands_to F) basis"
+
+lemma dominant_term_expands_to:
+  assumes "basis_wf basis" "(f expands_to F) basis" "trimmed F"
+  shows   "f \<sim>[at_top] eval_monom (dominant_term F) basis"
+proof -
+  from assms have "eval F \<sim>[at_top] f" by (intro asymp_equiv_refl_ev) (simp add: expands_to.simps)
+  also from dominant_term[OF assms(1) _ assms(3)] assms(2)
+    have "eval F \<sim>[at_top] eval_monom (dominant_term F) basis" by (simp add: expands_to.simps)
+  finally show ?thesis by (subst asymp_equiv_sym) simp_all
+qed
+
+lemma expands_to_cong:
+  "(f expands_to F) basis \<Longrightarrow> eventually (\<lambda>x. f x = g x) at_top \<Longrightarrow> (g expands_to F) basis"
+  by (auto simp: expands_to.simps elim: eventually_elim2)
+
+lemma expands_to_cong':
+  assumes "(f expands_to MS xs g) basis" "eventually (\<lambda>x. g x = g' x) at_top"
+  shows   "(f expands_to MS xs g') basis"
+proof -
+  have "is_expansion_aux xs g' basis"
+    by (rule is_expansion_aux_cong [OF _ assms(2)]) (insert assms(1), simp add: expands_to.simps)
+  with assms show ?thesis
+    by (auto simp: expands_to.simps elim: eventually_elim2)
+qed
+
+lemma eval_expands_to: "(f expands_to F) basis \<Longrightarrow> eventually (\<lambda>x. eval F x = f x) at_top"
+  by (simp add: expands_to.simps)
+
+lemma expands_to_real_compose:
+  assumes "(f expands_to (c::real)) bs"
+  shows   "((\<lambda>x. g (f x)) expands_to g c) bs"
+  using assms by (auto simp: expands_to.simps is_expansion_real.simps elim: eventually_mono)
+
+lemma expands_to_lift_compose:
+  assumes "(f expands_to MS (MSLCons (C, e) xs) f') bs'"
+  assumes "eventually (\<lambda>x. f' x - h x = 0) at_top" "e = 0"
+  assumes "((\<lambda>x. g (h x)) expands_to G) bs" "basis_wf (b # bs)"
+  shows   "((\<lambda>x. g (f x)) expands_to lift_ms G) (b # bs)"
+proof -
+  from assms have "\<forall>\<^sub>F x in at_top. f' x = f x" "\<forall>\<^sub>F x in at_top. eval G x = g (h x)"
+    by (auto simp: expands_to.simps)
+  with assms(2) have "\<forall>\<^sub>F x in at_top. eval G x = g (f x)"
+    by eventually_elim simp
+  with assms show ?thesis
+    by (intro expands_to.intros is_expansion_lift) (auto simp: expands_to.simps)
+qed
+
+lemma expands_to_zero: 
+    "basis_wf basis \<Longrightarrow> length basis = expansion_level TYPE('a) \<Longrightarrow>
+       ((\<lambda>_. 0) expands_to (zero_expansion :: 'a :: multiseries)) basis"
+  by (auto simp add: expands_to.simps is_expansion_zero)
+  
+lemma expands_to_const: 
+    "basis_wf basis \<Longrightarrow> length basis = expansion_level TYPE('a) \<Longrightarrow>
+       ((\<lambda>_. c) expands_to (const_expansion c :: 'a :: multiseries)) basis"
+  by (auto simp add: expands_to.simps is_expansion_const)
+
+lemma expands_to_X: "((\<lambda>x. x) expands_to MS (MSLCons (1 :: real, 1) MSLNil) (\<lambda>x. x)) [\<lambda>x. x]"
+  using is_expansion_X by (simp add: expands_to.simps)
+    
+lemma expands_to_uminus:
+  "basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow> ((\<lambda>x. -f x) expands_to -F) basis"
+  by (auto simp: expands_to.simps is_expansion_uminus)
+
+lemma expands_to_add:
+  "basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow> (g expands_to G) basis \<Longrightarrow>
+     ((\<lambda>x. f x + g x) expands_to F + G) basis"
+  by (auto simp: expands_to.simps is_expansion_add elim: eventually_elim2)
+
+lemma expands_to_minus:
+  "basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow> (g expands_to G) basis \<Longrightarrow>
+     ((\<lambda>x. f x - g x) expands_to F - G) basis"
+  by (auto simp: expands_to.simps is_expansion_minus elim: eventually_elim2)
+
+lemma expands_to_mult:
+  "basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow> (g expands_to G) basis \<Longrightarrow>
+     ((\<lambda>x. f x * g x) expands_to F * G) basis"
+  by (auto simp: expands_to.simps is_expansion_mult elim: eventually_elim2)
+
+lemma expands_to_inverse:
+  "trimmed F \<Longrightarrow> basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow>
+     ((\<lambda>x. inverse (f x)) expands_to inverse F) basis"
+  by (auto simp: expands_to.simps is_expansion_inverse)
+
+lemma expands_to_divide:
+  "trimmed G \<Longrightarrow> basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow> (g expands_to G) basis \<Longrightarrow>
+     ((\<lambda>x. f x / g x) expands_to F / G) basis"
+  by (auto simp: expands_to.simps is_expansion_divide elim: eventually_elim2)
+
+lemma expands_to_powr_0:
+  "eventually (\<lambda>x. f x = 0) at_top \<Longrightarrow> g \<equiv> g \<Longrightarrow> basis_wf bs \<Longrightarrow>
+     length bs = expansion_level TYPE('a) \<Longrightarrow>
+     ((\<lambda>x. f x powr g x) expands_to (zero_expansion :: 'a :: multiseries)) bs"
+  by (erule (1) expands_to_cong[OF expands_to_zero]) simp_all
+
+lemma expands_to_powr_const:
+  "trimmed_pos F \<Longrightarrow> basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow> p \<equiv> p \<Longrightarrow>
+     ((\<lambda>x. f x powr p) expands_to powr_expansion abort F p) basis"
+  by (auto simp: expands_to.simps is_expansion_powr trimmed_pos_def elim: eventually_mono)
+
+lemma expands_to_powr_const_0:
+  "eventually (\<lambda>x. f x = 0) at_top \<Longrightarrow> basis_wf bs \<Longrightarrow> 
+     length bs = expansion_level TYPE('a :: multiseries) \<Longrightarrow>
+     p \<equiv> p \<Longrightarrow> ((\<lambda>x. f x powr p) expands_to (zero_expansion :: 'a)) bs"
+  by (rule expands_to_cong[OF expands_to_zero]) auto
+  
+
+lemma expands_to_powr:
+  assumes "trimmed_pos F" "basis_wf basis'" "(f expands_to F) basis'"
+  assumes "((\<lambda>x. exp (ln (f x) * g x)) expands_to E) basis"
+  shows   "((\<lambda>x. f x powr g x) expands_to E) basis"
+using assms(4)
+proof (rule expands_to_cong)
+  from eval_pos_if_dominant_term_pos[of basis' F] assms(1-4)
+    show "eventually (\<lambda>x. exp (ln (f x) * g x) = f x powr g x) at_top"
+    by (auto simp: expands_to.simps trimmed_pos_def powr_def elim: eventually_elim2)
+qed
+
+lemma expands_to_ln_powr:
+  assumes "trimmed_pos F" "basis_wf basis'" "(f expands_to F) basis'"
+  assumes "((\<lambda>x. ln (f x) * g x) expands_to E) basis"
+  shows   "((\<lambda>x. ln (f x powr g x)) expands_to E) basis"
+using assms(4)
+proof (rule expands_to_cong)
+  from eval_pos_if_dominant_term_pos[of basis' F] assms(1-4)
+    show "eventually (\<lambda>x. ln (f x) * g x = ln (f x powr g x)) at_top"
+    by (auto simp: expands_to.simps trimmed_pos_def powr_def elim: eventually_elim2)
+qed
+  
+lemma expands_to_exp_ln:
+  assumes "trimmed_pos F" "basis_wf basis" "(f expands_to F) basis"
+  shows   "((\<lambda>x. exp (ln (f x))) expands_to F) basis"
+using assms(3)
+proof (rule expands_to_cong)
+  from eval_pos_if_dominant_term_pos[of basis F] assms
+    show "eventually (\<lambda>x. f x = exp (ln (f x))) at_top"
+    by (auto simp: expands_to.simps trimmed_pos_def powr_def elim: eventually_elim2)
+qed
+
+lemma expands_to_power:
+  assumes "trimmed F" "basis_wf basis" "(f expands_to F) basis"
+  shows   "((\<lambda>x. f x ^ n) expands_to power_expansion abort F n) basis"
+  using assms by (auto simp: expands_to.simps is_expansion_power elim: eventually_mono)
+
+lemma expands_to_power_0:
+  assumes "eventually (\<lambda>x. f x = 0) at_top" "basis_wf basis"
+          "length basis = expansion_level TYPE('a :: multiseries)" "n \<equiv> n"
+  shows   "((\<lambda>x. f x ^ n) expands_to (const_expansion (0 ^ n) :: 'a)) basis"
+  by (rule expands_to_cong[OF expands_to_const]) (insert assms, auto elim: eventually_mono)
+
+lemma expands_to_0th_root:
+  assumes  "n = 0" "basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)" "f \<equiv> f"
+  shows   "((\<lambda>x. root n (f x)) expands_to (zero_expansion :: 'a)) basis"
+  by (rule expands_to_cong[OF expands_to_zero]) (insert assms, auto)
+
+lemma expands_to_root_0:
+  assumes  "n > 0" "eventually (\<lambda>x. f x = 0) at_top"
+           "basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)"
+  shows   "((\<lambda>x. root n (f x)) expands_to (zero_expansion :: 'a)) basis"
+  by (rule expands_to_cong[OF expands_to_zero]) (insert assms, auto elim: eventually_mono)
+
+lemma expands_to_root:
+  assumes  "n > 0" "trimmed_pos F" "basis_wf basis" "(f expands_to F) basis"
+  shows   "((\<lambda>x. root n (f x)) expands_to powr_expansion False F (inverse (real n))) basis"
+proof -
+  have "((\<lambda>x. f x powr (inverse (real n))) expands_to 
+          powr_expansion False F (inverse (real n))) basis"
+    using assms(2-) by (rule expands_to_powr_const)
+  moreover have "eventually (\<lambda>x. f x powr (inverse (real n)) = root n (f x)) at_top"
+    using eval_pos_if_dominant_term_pos[of basis F] assms
+    by (auto simp: trimmed_pos_def expands_to.simps root_powr_inverse field_simps
+             elim: eventually_elim2)
+  ultimately show ?thesis by (rule expands_to_cong)
+qed
+
+lemma expands_to_root_neg:
+  assumes  "n > 0" "trimmed_neg F" "basis_wf basis" "(f expands_to F) basis"
+  shows   "((\<lambda>x. root n (f x)) expands_to
+             -powr_expansion False (-F) (inverse (real n))) basis"
+proof (rule expands_to_cong)
+  show "((\<lambda>x. -root n (-f x)) expands_to
+          -powr_expansion False (-F) (inverse (real n))) basis"
+    using assms by (intro expands_to_uminus expands_to_root trimmed_pos_uminus) auto
+qed (simp_all add: real_root_minus)
+
+lemma expands_to_sqrt:
+  assumes "trimmed_pos F" "basis_wf basis" "(f expands_to F) basis"
+  shows   "((\<lambda>x. sqrt (f x)) expands_to powr_expansion False F (1/2)) basis"
+  using expands_to_root[of 2 F basis f] assms by (simp add: sqrt_def)
+
+lemma expands_to_exp_real:
+  "(f expands_to (F :: real)) basis \<Longrightarrow> ((\<lambda>x. exp (f x)) expands_to exp F) basis"
+  by (auto simp: expands_to.simps is_expansion_real.simps elim: eventually_mono)
+
+lemma expands_to_exp_MSLNil:
+  assumes "(f expands_to (MS MSLNil f' :: 'a :: multiseries ms)) bs" "basis_wf bs"
+  shows   "((\<lambda>x. exp (f x)) expands_to (const_expansion 1 :: 'a ms)) bs"
+proof -
+  from assms have
+     "eventually (\<lambda>x. f' x = 0) at_top" "eventually (\<lambda>x. f' x = f x) at_top"
+     and [simp]: "length bs = Suc (expansion_level(TYPE('a)))"
+    by (auto simp: expands_to.simps elim: is_expansion_aux.cases)
+  from this(1,2) have "eventually (\<lambda>x. 1 = exp (f x)) at_top"
+    by eventually_elim simp
+  thus ?thesis by (auto simp: expands_to.simps intro!: is_expansion_const assms)
+qed
+
+lemma expands_to_exp_zero:
+  "(g expands_to MS xs f) basis \<Longrightarrow> eventually (\<lambda>x. f x = 0) at_top \<Longrightarrow> basis_wf basis \<Longrightarrow>
+     length basis = expansion_level TYPE('a :: multiseries) \<Longrightarrow>
+     ((\<lambda>x. exp (g x)) expands_to (const_expansion 1 :: 'a)) basis"
+  by (auto simp: expands_to.simps intro!: is_expansion_const elim: eventually_elim2)
+
+lemma expands_to_sin_real: 
+  "(f expands_to (F :: real)) basis \<Longrightarrow> ((\<lambda>x. sin (f x)) expands_to sin F) basis"
+  by (auto simp: expands_to.simps is_expansion_real.simps elim: eventually_mono)
+
+lemma expands_to_cos_real: 
+  "(f expands_to (F :: real)) basis \<Longrightarrow> ((\<lambda>x. cos (f x)) expands_to cos F) basis"
+  by (auto simp: expands_to.simps is_expansion_real.simps elim: eventually_mono)
+   
+lemma expands_to_sin_MSLNil:
+  "(f expands_to MS (MSLNil:: ('a \<times> real) msllist) g) basis \<Longrightarrow> basis_wf basis \<Longrightarrow> 
+     ((\<lambda>x. sin (f x)) expands_to MS (MSLNil :: ('a :: multiseries \<times> real) msllist) (\<lambda>x. 0)) basis"
+  by (auto simp: expands_to.simps dominant_term_ms_aux_def intro!: is_expansion_aux.intros
+           elim: eventually_elim2 is_expansion_aux.cases)
+
+lemma expands_to_cos_MSLNil:
+  "(f expands_to MS (MSLNil:: ('a :: multiseries \<times> real) msllist) g) basis \<Longrightarrow> basis_wf basis \<Longrightarrow> 
+     ((\<lambda>x. cos (f x)) expands_to (const_expansion 1 :: 'a ms)) basis"
+  by (auto simp: expands_to.simps dominant_term_ms_aux_def const_expansion_ms_def 
+           intro!: const_ms_aux elim: is_expansion_aux.cases eventually_elim2)
+
+lemma sin_ms_aux':
+  assumes "ms_exp_gt 0 (ms_aux_hd_exp xs)" "basis_wf basis" "is_expansion_aux xs f basis"
+  shows   "is_expansion_aux (sin_ms_aux' xs) (\<lambda>x. sin (f x)) basis"
+  unfolding sin_ms_aux'_def sin_conv_pre_sin power2_eq_square using assms
+  by (intro times_ms_aux powser_ms_aux[unfolded o_def] convergent_powser_sin_series_stream)
+     (auto simp: hd_exp_times add_neg_neg split: option.splits)
+
+lemma cos_ms_aux':
+  assumes "ms_exp_gt 0 (ms_aux_hd_exp xs)" "basis_wf basis" "is_expansion_aux xs f basis"
+  shows   "is_expansion_aux (cos_ms_aux' xs) (\<lambda>x. cos (f x)) basis"
+  unfolding cos_ms_aux'_def cos_conv_pre_cos power2_eq_square using assms
+  by (intro times_ms_aux powser_ms_aux[unfolded o_def] convergent_powser_cos_series_stream)
+     (auto simp: hd_exp_times add_neg_neg split: option.splits)
+
+lemma expands_to_sin_ms_neg_exp: 
+  assumes "e < 0" "basis_wf basis" "(f expands_to MS (MSLCons (C, e) xs) g) basis"
+  shows   "((\<lambda>x. sin (f x)) expands_to MS (sin_ms_aux' (MSLCons (C, e) xs)) (\<lambda>x. sin (g x))) basis"
+  using assms by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def
+                       intro!: sin_ms_aux' elim: eventually_mono)
+
+lemma expands_to_cos_ms_neg_exp: 
+  assumes "e < 0" "basis_wf basis" "(f expands_to MS (MSLCons (C, e) xs) g) basis"
+  shows   "((\<lambda>x. cos (f x)) expands_to MS (cos_ms_aux' (MSLCons (C, e) xs)) (\<lambda>x. cos (g x))) basis"
+  using assms by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def
+                       intro!: cos_ms_aux' elim: eventually_mono)
+
+lemma is_expansion_sin_ms_zero_exp: 
+  fixes F :: "('a :: multiseries \<times> real) msllist"
+  assumes basis: "basis_wf (b # basis)" 
+  assumes F: "is_expansion (MS (MSLCons (C, 0) xs) f) (b # basis)"
+  assumes Sin: "((\<lambda>x. sin (eval C x)) expands_to (Sin :: 'a)) basis"
+  assumes Cos: "((\<lambda>x. cos (eval C x)) expands_to (Cos :: 'a)) basis"
+  shows   "is_expansion 
+             (MS (plus_ms_aux (scale_shift_ms_aux (Sin, 0) (cos_ms_aux' xs))
+                              (scale_shift_ms_aux (Cos, 0) (sin_ms_aux' xs))) 
+             (\<lambda>x. sin (f x))) (b # basis)" (is "is_expansion (MS ?A _) _")
+proof -
+  let ?g = "\<lambda>x. f x - eval C x * b x powr 0"
+  let ?h = "\<lambda>x. eval Sin x * b x powr 0 * cos (?g x) + eval Cos x * b x powr 0 * sin (?g x)"
+  from basis have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+  from F have F': "is_expansion_aux xs ?g (b # basis)" 
+                  "ms_exp_gt 0 (ms_aux_hd_exp xs)" "is_expansion C basis"
+    by (auto elim!: is_expansion_aux_MSLCons) 
+  have "is_expansion_aux ?A ?h (b # basis)"
+    using basis Sin Cos unfolding F'(1)
+    by (intro plus_ms_aux scale_shift_ms_aux cos_ms_aux' sin_ms_aux' F')
+       (simp_all add: expands_to.simps)
+  moreover from b_pos eval_expands_to[OF Sin] eval_expands_to[OF Cos] 
+    have "eventually (\<lambda>x. ?h x = sin (f x)) at_top"
+    by eventually_elim (simp add: sin_add [symmetric])
+  ultimately have "is_expansion_aux ?A (\<lambda>x. sin (f x)) (b # basis)"
+    by (rule is_expansion_aux_cong)
+  thus ?thesis by simp
+qed
+
+lemma is_expansion_cos_ms_zero_exp: 
+  fixes F :: "('a :: multiseries \<times> real) msllist"
+  assumes basis: "basis_wf (b # basis)" 
+  assumes F: "is_expansion (MS (MSLCons (C, 0) xs) f) (b # basis)"
+  assumes Sin: "((\<lambda>x. sin (eval C x)) expands_to (Sin :: 'a)) basis"
+  assumes Cos: "((\<lambda>x. cos (eval C x)) expands_to (Cos :: 'a)) basis"
+  shows   "is_expansion 
+             (MS (minus_ms_aux (scale_shift_ms_aux (Cos, 0) (cos_ms_aux' xs))
+                              (scale_shift_ms_aux (Sin, 0) (sin_ms_aux' xs))) 
+             (\<lambda>x. cos (f x))) (b # basis)" (is "is_expansion (MS ?A _) _")
+proof -
+  let ?g = "\<lambda>x. f x - eval C x * b x powr 0"
+  let ?h = "\<lambda>x. eval Cos x * b x powr 0 * cos (?g x) - eval Sin x * b x powr 0 * sin (?g x)"
+  from basis have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+  from F have F': "is_expansion_aux xs ?g (b # basis)" 
+                  "ms_exp_gt 0 (ms_aux_hd_exp xs)" "is_expansion C basis"
+    by (auto elim!: is_expansion_aux_MSLCons) 
+  have "is_expansion_aux ?A ?h (b # basis)"
+    using basis Sin Cos unfolding F'(1)
+    by (intro minus_ms_aux scale_shift_ms_aux cos_ms_aux' sin_ms_aux' F')
+       (simp_all add: expands_to.simps)
+  moreover from b_pos eval_expands_to[OF Sin] eval_expands_to[OF Cos] 
+    have "eventually (\<lambda>x. ?h x = cos (f x)) at_top"
+    by eventually_elim (simp add: cos_add [symmetric])
+  ultimately have "is_expansion_aux ?A (\<lambda>x. cos (f x)) (b # basis)"
+    by (rule is_expansion_aux_cong)
+  thus ?thesis by simp
+qed
+
+lemma expands_to_sin_ms_zero_exp: 
+  assumes e: "e = 0" and basis: "basis_wf (b # basis)" 
+  assumes F: "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)"
+  assumes Sin: "((\<lambda>x. sin (c x)) expands_to Sin) basis"
+  assumes Cos: "((\<lambda>x. cos (c x)) expands_to Cos) basis"
+  assumes C: "eval C = c"
+  shows   "((\<lambda>x. sin (f x)) expands_to 
+             MS (plus_ms_aux (scale_shift_ms_aux (Sin, 0) (cos_ms_aux' xs))
+                              (scale_shift_ms_aux (Cos, 0) (sin_ms_aux' xs)))
+                (\<lambda>x. sin (g x))) (b # basis)"
+  using is_expansion_sin_ms_zero_exp[of b basis C xs g Sin Cos] assms
+  by (auto simp: expands_to.simps elim: eventually_mono)
+    
+lemma expands_to_cos_ms_zero_exp: 
+  assumes e: "e = 0" and basis: "basis_wf (b # basis)" 
+  assumes F: "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)"
+  assumes Sin: "((\<lambda>x. sin (c x)) expands_to Sin) basis"
+  assumes Cos: "((\<lambda>x. cos (c x)) expands_to Cos) basis"
+  assumes C: "eval C = c"
+  shows   "((\<lambda>x. cos (f x)) expands_to 
+             MS (minus_ms_aux (scale_shift_ms_aux (Cos, 0) (cos_ms_aux' xs))
+                              (scale_shift_ms_aux (Sin, 0) (sin_ms_aux' xs)))
+                (\<lambda>x. cos (g x))) (b # basis)"
+  using is_expansion_cos_ms_zero_exp[of b basis C xs g Sin Cos] assms
+  by (auto simp: expands_to.simps elim: eventually_mono)
+  
+
+
+lemma expands_to_sgn_zero:
+  assumes "eventually (\<lambda>x. f x = 0) at_top" "basis_wf basis"
+          "length basis = expansion_level TYPE('a :: multiseries)"
+  shows   "((\<lambda>x. sgn (f x)) expands_to (zero_expansion :: 'a)) basis"
+  by (rule expands_to_cong[OF expands_to_zero]) (insert assms, auto simp: sgn_eq_0_iff)
+
+lemma expands_to_sgn_pos:
+  assumes "trimmed_pos F" "(f expands_to F) basis" "basis_wf basis"
+          "length basis = expansion_level TYPE('a :: multiseries)"
+  shows   "((\<lambda>x. sgn (f x)) expands_to (const_expansion 1 :: 'a)) basis"
+proof (rule expands_to_cong[OF expands_to_const])
+  from trimmed_imp_eventually_sgn[of basis F] assms
+    have "eventually (\<lambda>x. sgn (eval F x) = 1) at_top" 
+      by (simp add: expands_to.simps trimmed_pos_def)
+  moreover from assms have "eventually (\<lambda>x. eval F x = f x) at_top"
+    by (simp add: expands_to.simps)
+  ultimately show "eventually (\<lambda>x. 1 = sgn (f x)) at_top" by eventually_elim simp
+qed (insert assms, auto)
+
+lemma expands_to_sgn_neg:
+  assumes "trimmed_neg F" "(f expands_to F) basis" "basis_wf basis"
+          "length basis = expansion_level TYPE('a :: multiseries)"
+  shows   "((\<lambda>x. sgn (f x)) expands_to (const_expansion (-1) :: 'a)) basis"
+proof (rule expands_to_cong[OF expands_to_const])
+  from trimmed_imp_eventually_sgn[of basis F] assms
+    have "eventually (\<lambda>x. sgn (eval F x) = -1) at_top" 
+      by (simp add: expands_to.simps trimmed_neg_def)
+  moreover from assms have "eventually (\<lambda>x. eval F x = f x) at_top"
+    by (simp add: expands_to.simps)
+  ultimately show "eventually (\<lambda>x. -1 = sgn (f x)) at_top" by eventually_elim simp
+qed (insert assms, auto)
+
+
+
+lemma expands_to_abs_pos:
+  assumes "trimmed_pos F" "basis_wf basis" "(f expands_to F) basis"
+  shows   "((\<lambda>x. abs (f x)) expands_to F) basis"
+using assms(3)
+proof (rule expands_to_cong)
+  from trimmed_imp_eventually_sgn[of basis F] assms
+    have "eventually (\<lambda>x. sgn (eval F x) = 1) at_top" 
+    by (simp add: expands_to.simps trimmed_pos_def)
+  with assms(3) show "eventually (\<lambda>x. f x = abs (f x)) at_top" 
+    by (auto simp: expands_to.simps sgn_if split: if_splits elim: eventually_elim2)
+qed
+
+lemma expands_to_abs_neg:
+  assumes "trimmed_neg F" "basis_wf basis" "(f expands_to F) basis"
+  shows   "((\<lambda>x. abs (f x)) expands_to -F) basis"
+using expands_to_uminus[OF assms(2,3)]
+proof (rule expands_to_cong)
+  from trimmed_imp_eventually_sgn[of basis F] assms
+    have "eventually (\<lambda>x. sgn (eval F x) = -1) at_top" 
+    by (simp add: expands_to.simps trimmed_neg_def)
+  with assms(3) show "eventually (\<lambda>x. -f x = abs (f x)) at_top" 
+    by (auto simp: expands_to.simps sgn_if split: if_splits elim: eventually_elim2)
+qed
+
+lemma expands_to_imp_eventually_nz:
+  assumes "basis_wf basis" "(f expands_to F) basis" "trimmed F"
+  shows   "eventually (\<lambda>x. f x \<noteq> 0) at_top"
+  using trimmed_imp_eventually_nz[OF assms(1), of F] assms(2,3)
+  by (auto simp: expands_to.simps elim: eventually_elim2)
+
+lemma expands_to_imp_eventually_pos:
+  assumes "basis_wf basis" "(f expands_to F) basis" "trimmed_pos F"
+  shows   "eventually (\<lambda>x. f x > 0) at_top"
+  using eval_pos_if_dominant_term_pos[of basis F] assms 
+  by (auto simp: expands_to.simps trimmed_pos_def elim: eventually_elim2)
+
+lemma expands_to_imp_eventually_neg:
+  assumes "basis_wf basis" "(f expands_to F) basis" "trimmed_neg F"
+  shows   "eventually (\<lambda>x. f x < 0) at_top"
+  using eval_pos_if_dominant_term_pos[of basis "-F"] assms
+  by (auto simp: expands_to.simps trimmed_neg_def is_expansion_uminus elim: eventually_elim2)
+
+lemma expands_to_imp_eventually_gt:
+  assumes "basis_wf basis" "((\<lambda>x. f x - g x) expands_to F) basis" "trimmed_pos F"
+  shows   "eventually (\<lambda>x. f x > g x) at_top"
+  using expands_to_imp_eventually_pos[OF assms] by simp
+
+lemma expands_to_imp_eventually_lt:
+  assumes "basis_wf basis" "((\<lambda>x. f x - g x) expands_to F) basis" "trimmed_neg F"
+  shows   "eventually (\<lambda>x. f x < g x) at_top"
+  using expands_to_imp_eventually_neg[OF assms] by simp
+
+lemma eventually_diff_zero_imp_eq:
+  fixes f g :: "real \<Rightarrow> real"
+  assumes "eventually (\<lambda>x. f x - g x = 0) at_top"
+  shows   "eventually (\<lambda>x. f x = g x) at_top"
+  using assms by eventually_elim simp
+
+lemma smallo_trimmed_imp_eventually_less:
+  fixes f g :: "real \<Rightarrow> real"
+  assumes "f \<in> o(g)" "(g expands_to G) bs" "basis_wf bs" "trimmed_pos G"
+  shows   "eventually (\<lambda>x. f x < g x) at_top"
+proof -
+  from assms(2-4) have pos: "eventually (\<lambda>x. g x > 0) at_top"
+    using expands_to_imp_eventually_pos by blast
+  have "(1 / 2 :: real) > 0" by simp
+  from landau_o.smallD[OF assms(1) this] and pos
+    show ?thesis by eventually_elim (simp add: abs_if split: if_splits)
+qed
+
+lemma smallo_trimmed_imp_eventually_greater:
+  fixes f g :: "real \<Rightarrow> real"
+  assumes "f \<in> o(g)" "(g expands_to G) bs" "basis_wf bs" "trimmed_neg G"
+  shows   "eventually (\<lambda>x. f x > g x) at_top"
+proof -
+  from assms(2-4) have pos: "eventually (\<lambda>x. g x < 0) at_top"
+    using expands_to_imp_eventually_neg by blast
+  have "(1 / 2 :: real) > 0" by simp
+  from landau_o.smallD[OF assms(1) this] and pos
+    show ?thesis by eventually_elim (simp add: abs_if split: if_splits)
+qed
+
+lemma expands_to_min_lt:
+  assumes "(f expands_to F) basis" "eventually (\<lambda>x. f x < g x) at_top"
+  shows   "((\<lambda>x. min (f x) (g x)) expands_to F) basis"
+  using assms(1)
+  by (rule expands_to_cong) (insert assms(2), auto simp: min_def elim: eventually_mono)
+
+lemma expands_to_min_eq:
+  assumes "(f expands_to F) basis" "eventually (\<lambda>x. f x = g x) at_top"
+  shows   "((\<lambda>x. min (f x) (g x)) expands_to F) basis"
+  using assms(1)
+  by (rule expands_to_cong) (insert assms(2), auto simp: min_def elim: eventually_mono)
+
+lemma expands_to_min_gt:
+  assumes "(g expands_to G) basis" "eventually (\<lambda>x. f x > g x) at_top"
+  shows   "((\<lambda>x. min (f x) (g x)) expands_to G) basis"
+  using assms(1)
+  by (rule expands_to_cong) (insert assms(2), auto simp: min_def elim: eventually_mono)
+    
+lemma expands_to_max_gt:
+  assumes "(f expands_to F) basis" "eventually (\<lambda>x. f x > g x) at_top"
+  shows   "((\<lambda>x. max (f x) (g x)) expands_to F) basis"
+  using assms(1)
+  by (rule expands_to_cong) (insert assms(2), auto simp: max_def elim: eventually_mono)
+    
+lemma expands_to_max_eq:
+  assumes "(f expands_to F) basis" "eventually (\<lambda>x. f x = g x) at_top"
+  shows   "((\<lambda>x. max (f x) (g x)) expands_to F) basis"
+  using assms(1)
+  by (rule expands_to_cong) (insert assms(2), auto simp: max_def elim: eventually_mono)
+
+lemma expands_to_max_lt:
+  assumes "(g expands_to G) basis" "eventually (\<lambda>x. f x < g x) at_top"
+  shows   "((\<lambda>x. max (f x) (g x)) expands_to G) basis"
+  using assms(1)
+  by (rule expands_to_cong) (insert assms(2), auto simp: max_def elim: eventually_mono)
+
+
+lemma expands_to_lift:
+  "is_lifting f basis basis' \<Longrightarrow> (c expands_to C) basis \<Longrightarrow> (c expands_to (f C)) basis'"
+  by (simp add: is_lifting_def expands_to.simps)
+    
+lemma expands_to_basic_powr: 
+  assumes "basis_wf (b # basis)" "length basis = expansion_level TYPE('a::multiseries)"
+  shows   "((\<lambda>x. b x powr e) expands_to 
+             MS (MSLCons (const_expansion 1 :: 'a, e) MSLNil) (\<lambda>x. b x powr e)) (b # basis)"
+  using assms by (auto simp: expands_to.simps basis_wf_Cons powr_smallo_iff
+                       intro!: is_expansion_aux.intros is_expansion_const powr_smallo_iff)
+
+lemma expands_to_basic_inverse: 
+  assumes "basis_wf (b # basis)" "length basis = expansion_level TYPE('a::multiseries)"
+  shows   "((\<lambda>x. inverse (b x)) expands_to 
+             MS (MSLCons (const_expansion 1 :: 'a, -1) MSLNil) (\<lambda>x. b x powr -1)) (b # basis)"
+proof -
+  from assms have "eventually (\<lambda>x. b x > 0) at_top" 
+    by (simp add: basis_wf_Cons filterlim_at_top_dense)
+  moreover have "(\<lambda>a. inverse (a powr 1)) \<in> o(\<lambda>a. a powr e')" if "e' > -1" for e' :: real
+    using that by (simp add: landau_o.small.inverse_eq2 powr_add [symmetric] one_smallo_powr)
+  ultimately show ?thesis using assms
+    by (auto simp: expands_to.simps basis_wf_Cons powr_minus
+             elim: eventually_mono
+             intro!: is_expansion_aux.intros is_expansion_const
+                     landau_o.small.compose[of _ at_top _ b])
+qed
+
+lemma expands_to_div':
+  assumes "basis_wf basis" "(f expands_to F) basis" "((\<lambda>x. inverse (g x)) expands_to G) basis"
+  shows   "((\<lambda>x. f x / g x) expands_to F * G) basis"
+  using expands_to_mult[OF assms] by (simp add: divide_simps)
+
+lemma expands_to_basic: 
+  assumes "basis_wf (b # basis)" "length basis = expansion_level TYPE('a::multiseries)"
+  shows   "(b expands_to MS (MSLCons (const_expansion 1 :: 'a, 1) MSLNil) b) (b # basis)"
+proof -
+  from assms have "eventually (\<lambda>x. b x > 0) at_top" 
+    by (simp add: basis_wf_Cons filterlim_at_top_dense)
+  moreover {
+    fix e' :: real assume e': "e' > 1"
+    have "(\<lambda>x::real. x) \<in> \<Theta>(\<lambda>x. x powr 1)"
+      by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto
+    also have "(\<lambda>x. x powr 1) \<in> o(\<lambda>x. x powr e')"
+      using e' by (subst powr_smallo_iff) (auto simp: filterlim_ident)
+    finally have "(\<lambda>x. x) \<in> o(\<lambda>x. x powr e')" .
+  }
+  ultimately show ?thesis using assms
+    by (auto simp: expands_to.simps basis_wf_Cons elim: eventually_mono
+             intro!: is_expansion_aux.intros is_expansion_const
+                     landau_o.small.compose[of _ at_top _ b])
+qed
+
+lemma expands_to_lift':
+  assumes "basis_wf (b # basis)" "(f expands_to MS xs g) basis"
+  shows   "(f expands_to (MS (MSLCons (MS xs g, 0) MSLNil) g)) (b # basis)"
+proof -
+  have "is_lifting lift_ms basis (b # basis)" by (rule is_lifting_lift) fact+
+  from expands_to_lift[OF this assms(2)] show ?thesis by (simp add: lift_ms_def)
+qed
+
+lemma expands_to_lift'':
+  assumes "basis_wf (b # basis)" "(f expands_to F) basis"
+  shows   "(f expands_to (MS (MSLCons (F, 0) MSLNil) (eval F))) (b # basis)"
+proof -
+  have "is_lifting lift_ms basis (b # basis)" by (rule is_lifting_lift) fact+
+  from expands_to_lift[OF this assms(2)] show ?thesis by (simp add: lift_ms_def)
+qed
+
+lemma expands_to_drop_zero:
+  assumes "eventually (\<lambda>x. eval C x = 0) at_top"
+  assumes "(f expands_to (MS (MSLCons (C, e) xs) g)) basis"
+  shows   "(f expands_to (MS xs g)) basis"
+  using assms drop_zero_ms[of C e xs g basis] by (simp add: expands_to.simps)
+
+lemma expands_to_hd'':
+  assumes "(f expands_to (MS (MSLCons (C, e) xs) g)) (b # basis)"
+  shows   "(eval C expands_to C) basis"
+  using assms by (auto simp: expands_to.simps elim: is_expansion_aux_MSLCons)
+
+lemma expands_to_hd:
+  assumes "(f expands_to (MS (MSLCons (MS ys h, e) xs) g)) (b # basis)"
+  shows   "(h expands_to MS ys h) basis"
+  using assms by (auto simp: expands_to.simps elim: is_expansion_aux_MSLCons)
+
+lemma expands_to_hd':
+  assumes "(f expands_to (MS (MSLCons (c :: real, e) xs) g)) (b # basis)"
+  shows   "((\<lambda>_. c) expands_to c) basis"
+  using assms by (auto simp: expands_to.simps elim: is_expansion_aux_MSLCons)
+
+lemma expands_to_trim_cong:
+  assumes "(f expands_to (MS (MSLCons (C, e) xs) g)) (b # basis)"
+  assumes "(eval C expands_to C') basis"
+  shows   "(f expands_to (MS (MSLCons (C', e) xs) g)) (b # basis)"
+proof -
+  from assms(1) have "is_expansion_aux xs (\<lambda>x. g x - eval C x * b x powr e) (b # basis)"
+    by (auto simp: expands_to.simps elim!: is_expansion_aux_MSLCons)
+  hence "is_expansion_aux xs (\<lambda>x. g x - eval C' x * b x powr e) (b # basis)"
+    by (rule is_expansion_aux_cong)
+       (insert assms(2), auto simp: expands_to.simps elim: eventually_mono)
+  with assms show ?thesis
+    by (auto simp: expands_to.simps elim!: is_expansion_aux_MSLCons intro!: is_expansion_aux.intros)
+qed
+
+lemma is_expansion_aux_expands_to:
+  assumes "(f expands_to MS xs g) basis"
+  shows   "is_expansion_aux xs f basis"
+proof -
+  from assms have "is_expansion_aux xs g basis" "eventually (\<lambda>x. g x = f x) at_top" 
+    by (simp_all add: expands_to.simps)
+  thus ?thesis by (rule is_expansion_aux_cong)
+qed
+
+lemma ln_series_stream_aux_code:
+  "ln_series_stream_aux True c = MSSCons (- inverse c) (ln_series_stream_aux False (c + 1))"
+  "ln_series_stream_aux False c = MSSCons (inverse c) (ln_series_stream_aux True (c + 1))"
+  by (subst ln_series_stream_aux.code, simp)+
+
+definition powser_ms_aux' where
+  "powser_ms_aux' cs xs = powser_ms_aux (msllist_of_msstream cs) xs"
+
+lemma ln_ms_aux:
+  fixes C L :: "'a :: multiseries"
+  assumes trimmed: "trimmed_pos C" and basis: "basis_wf (b # basis)"
+  assumes F: "is_expansion_aux (MSLCons (C, e) xs) f (b # basis)"
+  assumes L: "((\<lambda>x. ln (eval C x) + e * ln (b x)) expands_to L) basis"
+  shows   "is_expansion_aux 
+             (MSLCons (L, 0) (times_ms_aux (scale_shift_ms_aux (inverse C, - e) xs)
+               (powser_ms_aux' (ln_series_stream_aux False 1) 
+                 (scale_shift_ms_aux (inverse C, - e) xs))))
+             (\<lambda>x. ln (f x)) (b # basis)"
+    (is "is_expansion_aux ?G _ _")
+proof -
+  from F have F': 
+    "is_expansion_aux xs (\<lambda>x. f x - eval C x * b x powr e) (b # basis)"
+    "is_expansion C basis" "\<forall>e'>e. f \<in> o(\<lambda>x. b x powr e')" "ms_exp_gt e (ms_aux_hd_exp xs)"
+    by (auto elim!: is_expansion_aux_MSLCons)
+  from basis have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 1) at_top" by (simp add: filterlim_at_top_dense)
+  from eval_pos_if_dominant_term_pos[of basis C] trimmed F' basis
+    have C_pos: "eventually (\<lambda>x. eval C x > 0) at_top"
+      by (auto simp: basis_wf_Cons trimmed_pos_def)
+  from eval_pos_if_dominant_term_pos'[OF basis _ F] trimmed
+  have f_pos: "eventually (\<lambda>x. f x > 0) at_top"
+    by (simp add: trimmed_pos_def trimmed_ms_aux_def dominant_term_ms_aux_def case_prod_unfold)
+  
+  from F' have "ms_exp_gt 0 (map_option ((+) (-e)) (ms_aux_hd_exp xs))" 
+    by (cases "ms_aux_hd_exp xs") simp_all
+  hence "is_expansion_aux (powser_ms_aux' ln_series_stream (scale_shift_ms_aux (inverse C, -e) xs))
+          ((\<lambda>x. ln (1 + x)) \<circ> (\<lambda>x. eval (inverse C) x * b x powr -e * 
+              (f x - eval C x * b x powr e))) (b # basis)" (is "is_expansion_aux _ ?g _")
+    unfolding powser_ms_aux'_def using powser_ms_aux' basis F' trimmed
+    by (intro powser_ms_aux' convergent_powser'_ln scale_shift_ms_aux is_expansion_inverse F')
+       (simp_all add: trimmed_pos_def hd_exp_basis basis_wf_Cons)
+  moreover have "eventually (\<lambda>x. ?g x = ln (f x) - eval L x * b x powr 0) at_top"
+    using b_pos C_pos f_pos eval_expands_to[OF L]
+    by eventually_elim 
+       (simp add: powr_minus algebra_simps ln_mult ln_inverse expands_to.simps ln_powr)
+  ultimately 
+    have "is_expansion_aux (powser_ms_aux' ln_series_stream (scale_shift_ms_aux (inverse C, -e) xs))
+            (\<lambda>x. ln (f x) - eval L x * b x powr 0) (b # basis)" 
+      by (rule is_expansion_aux_cong)
+  hence *: "is_expansion_aux (times_ms_aux (scale_shift_ms_aux (inverse C, - e) xs)
+              (powser_ms_aux' (ln_series_stream_aux False 1) (scale_shift_ms_aux (inverse C, - e) xs)))
+              (\<lambda>x. ln (f x) - eval L x * b x powr 0) (b # basis)"
+    unfolding ln_series_stream_def powser_ms_aux'_def powser_ms_aux_MSLCons msllist_of_msstream_MSSCons
+    by (rule drop_zero_ms_aux [rotated]) simp_all
+  moreover from F' have exp: "ms_exp_gt 0 (map_option ((+) (-e)) (ms_aux_hd_exp xs))"
+    by (cases "ms_aux_hd_exp xs") simp_all
+  moreover have "(\<lambda>x. ln (f x)) \<in> o(\<lambda>x. b x powr e')" if "e' > 0" for e'
+  proof -
+    from is_expansion_aux_imp_smallo[OF *, of e'] that exp
+    have "(\<lambda>x. ln (f x) - eval L x * b x powr 0) \<in> o(\<lambda>x. b x powr e')" 
+      by (cases "ms_aux_hd_exp xs") 
+         (simp_all add: hd_exp_times hd_exp_powser hd_exp_basis powser_ms_aux'_def)
+    moreover {
+      from L F' basis that have "eval L \<in> o(\<lambda>x. b x powr e')"
+        by (intro is_expansion_imp_smallo[of basis]) (simp_all add: basis_wf_Cons expands_to.simps)
+      also have "eval L \<in> o(\<lambda>x. b x powr e') \<longleftrightarrow> (\<lambda>x. eval L x * b x powr 0) \<in> o(\<lambda>x. b x powr e')"
+        using b_pos by (intro landau_o.small.in_cong) (auto elim: eventually_mono)
+      finally have \<dots> .
+    } ultimately have "(\<lambda>x. ln (f x) - eval L x * b x powr 0 + eval L x * b x powr 0) \<in> 
+                          o(\<lambda>x. b x powr e')" by (rule sum_in_smallo)
+    thus ?thesis by simp
+  qed
+  ultimately show "is_expansion_aux ?G (\<lambda>x. ln (f x)) (b # basis)" using L
+    by (intro is_expansion_aux.intros)
+       (auto simp: expands_to.simps hd_exp_times hd_exp_powser hd_exp_basis
+                   powser_ms_aux'_def split: option.splits)
+qed
+
+lemma expands_to_ln:
+  fixes C L :: "'a :: multiseries"
+  assumes trimmed: "trimmed_pos (MS (MSLCons (C, e) xs) g)" and basis: "basis_wf (b # basis)"
+  assumes F: "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)"
+  assumes L: "((\<lambda>x. ln (h x) + e * ln (b x)) expands_to L) basis"
+  and h: "eval C = h"
+  shows   "((\<lambda>x. ln (f x)) expands_to MS
+             (MSLCons (L, 0) (times_ms_aux (scale_shift_ms_aux (inverse C, - e) xs)
+               (powser_ms_aux' (ln_series_stream_aux False 1) 
+                 (scale_shift_ms_aux (inverse C, - e) xs)))) (\<lambda>x. ln (f x))) (b # basis)"
+  using is_expansion_aux_expands_to[OF F] assms by (auto simp: expands_to.simps intro!: ln_ms_aux)
+
+lemma trimmed_lifting:
+  assumes "is_lifting f basis basis'"
+  assumes "trimmed F"
+  shows   "trimmed (f F)"
+  using assms by (simp add: is_lifting_def)
+
+lemma trimmed_pos_lifting:
+  assumes "is_lifting f basis basis'"
+  assumes "trimmed_pos F"
+  shows   "trimmed_pos (f F)"
+  using assms by (simp add: is_lifting_def trimmed_pos_def)
+
+lemma expands_to_ln_aux_0:
+  assumes e: "e = 0"
+  assumes L1: "((\<lambda>x. ln (g x)) expands_to L) basis"
+  shows "((\<lambda>x. ln (g x) + e * ln (b x)) expands_to L) basis"
+  using assms by simp
+
+lemma expands_to_ln_aux_1:
+  assumes e: "e = 1"
+  assumes basis: "basis_wf (b # basis)"
+  assumes L1: "((\<lambda>x. ln (g x)) expands_to L1) basis"
+  assumes L2: "((\<lambda>x. ln (b x)) expands_to L2) basis"
+  shows "((\<lambda>x. ln (g x) + e * ln (b x)) expands_to L1 + L2) basis"
+  using assms by (intro expands_to_add) (simp_all add: basis_wf_Cons)
+
+lemma expands_to_ln_eventually_1:
+  assumes "basis_wf basis" "length basis = expansion_level TYPE('a)"
+  assumes "eventually (\<lambda>x. f x - 1 = 0) at_top"
+  shows   "((\<lambda>x. ln (f x)) expands_to (zero_expansion :: 'a :: multiseries)) basis"
+  by (rule expands_to_cong [OF expands_to_zero]) (insert assms, auto elim: eventually_mono)  
+
+lemma expands_to_ln_aux:
+  assumes basis: "basis_wf (b # basis)"
+  assumes L1: "((\<lambda>x. ln (g x)) expands_to L1) basis"
+  assumes L2: "((\<lambda>x. ln (b x)) expands_to L2) basis"
+  shows "((\<lambda>x. ln (g x) + e * ln (b x)) expands_to L1 + scale_ms e L2) basis"
+proof -
+  from L1 have "length basis = expansion_level TYPE('a)"
+    by (auto simp: expands_to.simps is_expansion_length)
+  with assms show ?thesis unfolding scale_ms_def
+    by (intro expands_to_add assms expands_to_mult expands_to_const) 
+       (simp_all add: basis_wf_Cons)
+qed
+
+lemma expands_to_ln_to_expands_to_ln_eval:
+  assumes "((\<lambda>x. ln (f x) + F x) expands_to L) basis"
+  shows   "((\<lambda>x. ln (eval (MS xs f) x) + F x) expands_to L) basis"
+  using assms by simp
+
+lemma expands_to_ln_const:
+  "((\<lambda>x. ln (eval (C :: real) x)) expands_to ln C) []"
+  by (simp add: expands_to.simps is_expansion_real.intros)
+
+lemma expands_to_meta_eq_cong:
+  assumes "(f expands_to F) basis" "F \<equiv> G"
+  shows   "(f expands_to G) basis"
+  using assms by simp
+    
+lemma expands_to_meta_eq_cong':
+  assumes "(g expands_to F) basis" "f \<equiv> g"
+  shows   "(f expands_to F) basis"
+  using assms by simp
+
+
+lemma uminus_ms_aux_MSLCons_eval:
+  "uminus_ms_aux (MSLCons (c, e) xs) = MSLCons (-c, e) (uminus_ms_aux xs)"
+  by (simp add: uminus_ms_aux_MSLCons)
+
+lemma scale_shift_ms_aux_MSLCons_eval:
+  "scale_shift_ms_aux (c, e) (MSLCons (c', e') xs) =
+     MSLCons (c * c', e + e') (scale_shift_ms_aux (c, e) xs)"
+  by (simp add: scale_shift_ms_aux_MSLCons)
+
+lemma times_ms_aux_MSLCons_eval: "times_ms_aux (MSLCons (c1, e1) xs) (MSLCons (c2, e2) ys) = 
+  MSLCons (c1 * c2, e1 + e2) (plus_ms_aux (scale_shift_ms_aux (c1, e1) ys) 
+    (times_ms_aux xs (MSLCons (c2, e2) ys)))"
+  by (simp add: times_ms_aux_MSLCons)
+
+lemma plus_ms_aux_MSLCons_eval:
+  "plus_ms_aux (MSLCons (c1, e1) xs) (MSLCons (c2, e2) ys) =
+     CMP_BRANCH (COMPARE e1 e2) 
+       (MSLCons (c2, e2) (plus_ms_aux (MSLCons (c1, e1) xs) ys))
+       (MSLCons (c1 + c2, e1) (plus_ms_aux xs ys))
+       (MSLCons (c1, e1) (plus_ms_aux xs (MSLCons (c2, e2) ys)))"
+  by (simp add: CMP_BRANCH_def COMPARE_def plus_ms_aux_MSLCons)
+
+lemma inverse_ms_aux_MSLCons: "inverse_ms_aux (MSLCons (C, e) xs) = 
+   scale_shift_ms_aux (inverse C, - e)
+     (powser_ms_aux' (mssalternate 1 (- 1))
+       (scale_shift_ms_aux (inverse C, - e)
+         xs))"
+  by (simp add: Let_def inverse_ms_aux.simps powser_ms_aux'_def)
+
+lemma powr_ms_aux_MSLCons: "powr_ms_aux abort (MSLCons (C, e) xs) p = 
+  scale_shift_ms_aux (powr_expansion abort C p, e * p)
+     (powser_ms_aux (gbinomial_series abort p)
+       (scale_shift_ms_aux (inverse C, - e) xs))"
+  by simp
+
+lemma power_ms_aux_MSLCons: "power_ms_aux abort (MSLCons (C, e) xs) n = 
+  scale_shift_ms_aux (power_expansion abort C n, e * real n)
+     (powser_ms_aux (gbinomial_series abort (real n))
+       (scale_shift_ms_aux (inverse C, - e) xs))"
+  by simp
+
+lemma minus_ms_aux_MSLCons_eval:
+  "minus_ms_aux (MSLCons (c1, e1) xs) (MSLCons (c2, e2) ys) = 
+     CMP_BRANCH (COMPARE e1 e2) 
+       (MSLCons (-c2, e2) (minus_ms_aux (MSLCons (c1, e1) xs) ys))
+       (MSLCons (c1 - c2, e1) (minus_ms_aux xs ys))
+       (MSLCons (c1, e1) (minus_ms_aux xs (MSLCons (c2, e2) ys)))"
+  by (simp add: minus_ms_aux_def plus_ms_aux_MSLCons_eval uminus_ms_aux_MSLCons minus_eq_plus_uminus)
+  
+lemma minus_ms_altdef: "MS xs f - MS ys g = MS (minus_ms_aux xs ys) (\<lambda>x. f x - g x)"
+  by (simp add: minus_ms_def minus_ms_aux_def)
+
+lemma const_expansion_ms_eval: "const_expansion c = MS (MSLCons (const_expansion c, 0) MSLNil) (\<lambda>_. c)"
+  by (simp add: const_expansion_ms_def const_ms_aux_def)
+ 
+lemma powser_ms_aux'_MSSCons:
+  "powser_ms_aux' (MSSCons c cs) xs = 
+     MSLCons (const_expansion c, 0) (times_ms_aux xs (powser_ms_aux' cs xs))"
+  by (simp add: powser_ms_aux'_def powser_ms_aux_MSLCons)
+
+lemma sin_ms_aux'_altdef:
+  "sin_ms_aux' xs = times_ms_aux xs (powser_ms_aux' sin_series_stream (times_ms_aux xs xs))"
+  by (simp add: sin_ms_aux'_def powser_ms_aux'_def)
+
+lemma cos_ms_aux'_altdef:
+  "cos_ms_aux' xs = powser_ms_aux' cos_series_stream (times_ms_aux xs xs)"
+  by (simp add: cos_ms_aux'_def powser_ms_aux'_def)
+
+lemma sin_series_stream_aux_code:
+  "sin_series_stream_aux True acc m = 
+     MSSCons (-inverse acc) (sin_series_stream_aux False (acc * (m + 1) * (m + 2)) (m + 2))"
+  "sin_series_stream_aux False acc m = 
+     MSSCons (inverse acc) (sin_series_stream_aux True (acc * (m + 1) * (m + 2)) (m + 2))"
+  by (subst sin_series_stream_aux.code; simp)+
+    
+lemma arctan_series_stream_aux_code:
+  "arctan_series_stream_aux True n = MSSCons (-inverse n) (arctan_series_stream_aux False (n + 2))"
+  "arctan_series_stream_aux False n = MSSCons (inverse n) (arctan_series_stream_aux True (n + 2))"
+  by (subst arctan_series_stream_aux.code; simp)+
+
+
+subsubsection \<open>Composition with an asymptotic power series\<close>
+
+context
+  fixes h :: "real \<Rightarrow> real" and F :: "real filter"
+begin
+  
+coinductive asymp_powser :: 
+    "(real \<Rightarrow> real) \<Rightarrow> real msstream \<Rightarrow> bool" where
+  "asymp_powser f cs \<Longrightarrow> f' \<in> O[F](\<lambda>_. 1) \<Longrightarrow>
+     eventually (\<lambda>x. f' x = c + f x * h x) F \<Longrightarrow> asymp_powser f' (MSSCons c cs)"
+
+lemma asymp_powser_coinduct [case_names asymp_powser]:
+  "P x1 x2 \<Longrightarrow> (\<And>x1 x2.  P x1 x2 \<Longrightarrow> \<exists>f cs c.
+       x2 = MSSCons c cs \<and> asymp_powser f cs \<and>
+       x1 \<in> O[F](\<lambda>_. 1) \<and> (\<forall>\<^sub>F x in F. x1 x = c + f x * h x)) \<Longrightarrow> asymp_powser x1 x2"
+  using asymp_powser.coinduct[of P x1 x2] by blast
+    
+lemma asymp_powser_coinduct' [case_names asymp_powser]:
+  "P x1 x2 \<Longrightarrow> (\<And>x1 x2.  P x1 x2 \<Longrightarrow> \<exists>f cs c.
+       x2 = MSSCons c cs \<and> P f cs \<and>
+       x1 \<in> O[F](\<lambda>_. 1) \<and> (\<forall>\<^sub>F x in F. x1 x = c + f x * h x)) \<Longrightarrow> asymp_powser x1 x2"
+  using asymp_powser.coinduct[of P x1 x2] by blast
+  
+lemma asymp_powser_transfer:
+  assumes "asymp_powser f cs" "eventually (\<lambda>x. f x = f' x) F"
+  shows   "asymp_powser f' cs"
+  using assms(1)
+proof (cases rule: asymp_powser.cases)
+  case (1 f cs' c)
+  have "asymp_powser f' (MSSCons c cs')"
+    by (rule asymp_powser.intros) 
+       (insert 1 assms(2), auto elim: eventually_elim2 dest: landau_o.big.in_cong)
+  thus ?thesis by (simp add: 1)
+qed
+
+lemma sum_bigo_imp_asymp_powser:
+  assumes filterlim_h: "filterlim h (at 0) F"
+  assumes "(\<And>n. (\<lambda>x. f x - (\<Sum>k<n. mssnth cs k * h x ^ k)) \<in> O[F](\<lambda>x. h x ^ n))"
+  shows   "asymp_powser f cs"
+  using assms(2)
+proof (coinduction arbitrary: f cs rule: asymp_powser_coinduct')
+  case (asymp_powser f cs)
+  from filterlim_h have h_nz:"eventually (\<lambda>x. h x \<noteq> 0) F"
+    by (auto simp: filterlim_at)
+  show ?case
+  proof (cases cs)
+    case (MSSCons c cs')
+    define f' where "f' = (\<lambda>x. (f x - c) / h x)"  
+    have "(\<lambda>x. f' x - (\<Sum>k<n. mssnth cs' k * h x ^ k)) \<in> O[F](\<lambda>x. h x ^ n)" for n
+    proof -
+      have "(\<lambda>x. h x * (f' x - (\<Sum>i<n. mssnth cs' i * h x ^ i))) \<in>
+                \<Theta>[F](\<lambda>x. f x - c - h x * (\<Sum>i<n. mssnth cs' i * h x ^ i))"
+        using h_nz by (intro bigthetaI_cong) (auto elim!: eventually_mono simp: f'_def field_simps)
+      also from spec[OF asymp_powser, of "Suc n"]
+        have "(\<lambda>x. f x - c - h x * (\<Sum>i<n. mssnth cs' i * h x ^ i)) \<in> O[F](\<lambda>x. h x * h x ^ n)"
+        unfolding sum_lessThan_Suc_shift
+        by (simp add: MSSCons algebra_simps sum_distrib_left sum_distrib_right)
+      finally show ?thesis
+        by (subst (asm) landau_o.big.mult_cancel_left) (insert h_nz, auto)
+    qed
+    moreover from h_nz have "\<forall>\<^sub>F x in F. f x = c + f' x * h x"
+      by eventually_elim (simp add: f'_def)
+    ultimately show ?thesis using spec[OF asymp_powser, of 0]
+      by (auto simp: MSSCons intro!: exI[of _ f'])
+  qed
+qed
+
+end
+
+
+lemma asymp_powser_o:
+  assumes "asymp_powser h F f cs" assumes "filterlim g F G"
+  shows   "asymp_powser (h \<circ> g) G (f \<circ> g) cs"
+using assms(1)
+proof (coinduction arbitrary: f cs rule: asymp_powser_coinduct')
+  case (asymp_powser f cs)
+  thus ?case
+  proof (cases rule: asymp_powser.cases)
+    case (1 f' cs' c)
+    from assms(2) have "filtermap g G \<le> F" by (simp add: filterlim_def)
+    moreover have "eventually (\<lambda>x. f x = c + f' x * h x) F" by fact
+    ultimately have "eventually (\<lambda>x. f x = c + f' x * h x) (filtermap g G)" by (rule filter_leD)
+    hence "eventually (\<lambda>x. f (g x) = c + f' (g x) * h (g x)) G" by (simp add: eventually_filtermap)
+    moreover from 1 have "f \<circ> g \<in> O[G](\<lambda>_. 1)" unfolding o_def
+      by (intro landau_o.big.compose[OF _ assms(2)]) auto
+    ultimately show ?thesis using 1 by force
+  qed
+qed
+
+lemma asymp_powser_compose:
+  assumes "asymp_powser h F f cs" assumes "filterlim g F G"
+  shows   "asymp_powser (\<lambda>x. h (g x)) G (\<lambda>x. f (g x)) cs"
+  using asymp_powser_o[OF assms] by (simp add: o_def)
+
+lemma hd_exp_powser': "ms_aux_hd_exp (powser_ms_aux' cs f) = Some 0"
+  by (simp add: powser_ms_aux'_def hd_exp_powser)
+
+lemma powser_ms_aux_asymp_powser:
+  assumes "asymp_powser h at_top f cs" and basis: "basis_wf bs"
+  assumes "is_expansion_aux xs h bs" "ms_exp_gt 0 (ms_aux_hd_exp xs)"
+  shows   "is_expansion_aux (powser_ms_aux' cs xs) f bs"
+using assms(1)
+proof (coinduction arbitrary: cs f rule: is_expansion_aux_coinduct_upto)
+  show "basis_wf bs" by fact
+next
+  case (B cs f)
+  thus ?case
+  proof (cases rule: asymp_powser.cases)
+    case (1 f' cs' c)
+    from assms(3) obtain b bs' where [simp]: "bs = b # bs'"
+      by (cases bs) (auto dest: is_expansion_aux_basis_nonempty)
+    from basis have b_at_top: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+    hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (auto simp: filterlim_at_top_dense)
+    
+    have A: "f \<in> o(\<lambda>x. b x powr e)" if "e > 0" for e :: real
+    proof -
+      have "f \<in> O(\<lambda>_. 1)" by fact
+      also have "(\<lambda>_::real. 1) \<in> o(\<lambda>x. b x powr e)"
+        by (rule landau_o.small.compose[OF _ b_at_top]) (insert that, simp_all add: one_smallo_powr)
+      finally show ?thesis .
+    qed
+    have "ms_closure (\<lambda>xsa'' fa'' basis''.
+            \<exists>cs. xsa'' = powser_ms_aux' cs xs \<and> basis'' = b # bs' \<and> asymp_powser h at_top fa'' cs)
+           (times_ms_aux xs (powser_ms_aux' cs' xs)) (\<lambda>x. h x * f' x) bs" 
+      (is "ms_closure ?Q ?ys _ _")
+      by (rule ms_closure_mult[OF ms_closure_embed' ms_closure_embed])
+         (insert assms 1, auto intro!: exI[of _ cs'])
+    moreover have "eventually (\<lambda>x. h x * f' x = f x - c * b x powr 0) at_top"
+      using b_pos and \<open>\<forall>\<^sub>F x in at_top. f x = c + f' x * h x\<close>
+      by eventually_elim simp
+    ultimately have "ms_closure ?Q ?ys (\<lambda>x. f x - c * b x powr 0) bs"
+      by (rule ms_closure_cong)
+    with 1 assms A show ?thesis
+      using is_expansion_aux_expansion_level[OF assms(3)]
+      by (auto simp: powser_ms_aux'_MSSCons basis_wf_Cons hd_exp_times hd_exp_powser'
+               intro!: is_expansion_const ms_closure_add ms_closure_mult split: option.splits)
+  qed
+qed
+
+
+subsubsection \<open>Arc tangent\<close>
+
+definition arctan_ms_aux where
+  "arctan_ms_aux xs = times_ms_aux xs (powser_ms_aux' arctan_series_stream (times_ms_aux xs xs))"
+
+lemma arctan_ms_aux_0:
+  assumes "ms_exp_gt 0 (ms_aux_hd_exp xs)" "basis_wf basis" "is_expansion_aux xs f basis"
+  shows   "is_expansion_aux (arctan_ms_aux xs) (\<lambda>x. arctan (f x)) basis" 
+proof (rule is_expansion_aux_cong)
+  let ?g = "powser (msllist_of_msstream arctan_series_stream)"
+  show "is_expansion_aux (arctan_ms_aux xs)
+          (\<lambda>x. f x * (?g \<circ> (\<lambda>x. f x * f x)) x) basis"
+    unfolding arctan_ms_aux_def powser_ms_aux'_def using assms
+    by (intro times_ms_aux powser_ms_aux powser_ms_aux convergent_powser_arctan)
+       (auto simp: hd_exp_times add_neg_neg split: option.splits)
+  
+  have "f \<in> o(\<lambda>x. hd basis x powr 0)" using assms
+    by (intro is_expansion_aux_imp_smallo[OF assms(3)]) auto
+  also have "(\<lambda>x. hd basis x powr 0) \<in> O(\<lambda>_. 1)"
+    by (intro bigoI[of _ 1]) auto
+  finally have "filterlim f (nhds 0) at_top"
+    by (auto dest!: smalloD_tendsto)
+  from order_tendstoD(2)[OF tendsto_norm [OF this], of 1]
+    have "eventually (\<lambda>x. norm (f x) < 1) at_top" by simp
+  thus "eventually (\<lambda>x. f x * (?g \<circ> (\<lambda>x. f x * f x)) x = arctan (f x)) at_top"
+    by eventually_elim (simp add: arctan_conv_pre_arctan power2_eq_square)
+qed
+
+lemma arctan_ms_aux_inf:
+  assumes "\<exists>e>0. ms_aux_hd_exp xs = Some e" "fst (dominant_term_ms_aux xs) > 0" "trimmed_ms_aux xs" 
+          "basis_wf basis" "is_expansion_aux xs f basis"
+  shows   "is_expansion_aux (minus_ms_aux (const_ms_aux (pi / 2))
+             (arctan_ms_aux (inverse_ms_aux xs))) (\<lambda>x. arctan (f x)) basis"
+    (is "is_expansion_aux ?xs' _ _")
+proof (rule is_expansion_aux_cong)
+  from \<open>trimmed_ms_aux xs\<close> have [simp]: "xs \<noteq> MSLNil" by auto
+  thus "is_expansion_aux ?xs' (\<lambda>x. pi / 2 - arctan (inverse (f x))) basis" using assms
+    by (intro minus_ms_aux arctan_ms_aux_0 inverse_ms_aux const_ms_aux)
+       (auto simp: hd_exp_inverse is_expansion_aux_expansion_level)
+next
+  from assms(2-5) have "eventually (\<lambda>x. f x > 0) at_top"
+    by (intro eval_pos_if_dominant_term_pos') simp_all
+  thus "eventually (\<lambda>x. ((\<lambda>x. pi/2 - arctan (inverse (f x)))) x = arctan (f x)) at_top"
+    unfolding o_def by eventually_elim (subst arctan_inverse_pos, simp_all)
+qed
+
+lemma arctan_ms_aux_neg_inf:
+  assumes "\<exists>e>0. ms_aux_hd_exp xs = Some e" "fst (dominant_term_ms_aux xs) < 0" "trimmed_ms_aux xs" 
+          "basis_wf basis" "is_expansion_aux xs f basis"
+  shows   "is_expansion_aux (minus_ms_aux (const_ms_aux (-pi / 2))
+             (arctan_ms_aux (inverse_ms_aux xs))) (\<lambda>x. arctan (f x)) basis"
+    (is "is_expansion_aux ?xs' _ _")
+proof (rule is_expansion_aux_cong)
+  from \<open>trimmed_ms_aux xs\<close> have [simp]: "xs \<noteq> MSLNil" by auto
+  thus "is_expansion_aux ?xs' (\<lambda>x. -pi / 2 - arctan (inverse (f x))) basis" using assms
+    by (intro minus_ms_aux arctan_ms_aux_0 inverse_ms_aux const_ms_aux)
+       (auto simp: hd_exp_inverse is_expansion_aux_expansion_level)
+next
+  from assms(2-5) have "eventually (\<lambda>x. f x < 0) at_top"
+    by (intro eval_neg_if_dominant_term_neg') simp_all
+  thus "eventually (\<lambda>x. ((\<lambda>x. -pi/2 - arctan (inverse (f x)))) x = arctan (f x)) at_top"
+    unfolding o_def by eventually_elim (subst arctan_inverse_neg, simp_all)
+qed
+
+lemma expands_to_arctan_zero:
+  assumes "((\<lambda>_::real. 0::real) expands_to C) bs" "eventually (\<lambda>x. f x = 0) at_top"
+  shows   "((\<lambda>x::real. arctan (f x)) expands_to C) bs"
+  using assms by (auto simp: expands_to.simps elim: eventually_elim2)
+
+lemma expands_to_arctan_ms_neg_exp: 
+  assumes "e < 0" "basis_wf basis" "(f expands_to MS (MSLCons (C, e) xs) g) basis"
+  shows   "((\<lambda>x. arctan (f x)) expands_to
+               MS (arctan_ms_aux (MSLCons (C, e) xs)) (\<lambda>x. arctan (g x))) basis"
+  using assms by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def
+                       intro!: arctan_ms_aux_0 elim: eventually_mono)
+
+lemma expands_to_arctan_ms_pos_exp_pos: 
+  assumes "e > 0" "trimmed_pos (MS (MSLCons (C, e) xs) g)" "basis_wf basis" 
+          "(f expands_to MS (MSLCons (C, e) xs) g) basis"
+  shows   "((\<lambda>x. arctan (f x)) expands_to MS (minus_ms_aux (const_ms_aux (pi / 2))
+             (arctan_ms_aux (inverse_ms_aux (MSLCons (C, e) xs))))
+               (\<lambda>x. arctan (g x))) basis"
+  using arctan_ms_aux_inf[of "MSLCons (C, e) xs" basis g] assms
+  by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def
+                 dominant_term_ms_aux_MSLCons trimmed_pos_def elim: eventually_mono)
+
+lemma expands_to_arctan_ms_pos_exp_neg: 
+  assumes "e > 0" "trimmed_neg (MS (MSLCons (C, e) xs) g)" "basis_wf basis" 
+          "(f expands_to MS (MSLCons (C, e) xs) g) basis"
+  shows   "((\<lambda>x. arctan (f x)) expands_to MS (minus_ms_aux (const_ms_aux (-pi/2))
+             (arctan_ms_aux (inverse_ms_aux (MSLCons (C, e) xs))))
+               (\<lambda>x. arctan (g x))) basis"
+  using arctan_ms_aux_neg_inf[of "MSLCons (C, e) xs" basis g] assms
+  by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def
+                 dominant_term_ms_aux_MSLCons trimmed_neg_def elim: eventually_mono)                     
+
+
+subsubsection \<open>Exponential function\<close>
+
+(* TODO: Move *)
+lemma ln_smallo_powr:
+  assumes "(e::real) > 0"
+  shows   "(\<lambda>x. ln x) \<in> o(\<lambda>x. x powr e)"
+proof -
+  have *: "ln \<in> o(\<lambda>x::real. x)"
+    using ln_x_over_x_tendsto_0 by (intro smalloI_tendsto) auto
+  have "(\<lambda>x. ln x) \<in> \<Theta>(\<lambda>x::real. ln x powr 1)"
+    by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 1]]) auto
+  also have "(\<lambda>x::real. ln x powr 1) \<in> o(\<lambda>x. x powr e)"
+    by (intro ln_smallo_imp_flat filterlim_ident ln_at_top assms
+              landau_o.small.compose[of _ at_top _ ln] *)
+  finally show ?thesis .
+qed
+
+lemma basis_wf_insert_exp_pos:
+  assumes "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)" 
+          "basis_wf (b # basis)" "trimmed (MS (MSLCons (C, e) xs) g)" "e - 0 > 0"
+  shows   "(\<lambda>x. ln (b x)) \<in> o(f)"
+proof -
+  from assms have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
+
+  from assms have "is_expansion_aux xs (\<lambda>x. g x - eval C x * b x powr e) (b # basis)" 
+                  "ms_exp_gt e (ms_aux_hd_exp xs)"
+      by (auto elim!: is_expansion_aux_MSLCons simp: expands_to.simps)
+  from is_expansion_aux_imp_smallo''[OF this] obtain e' where e':
+    "e' < e" "(\<lambda>x. g x - eval C x * b x powr e) \<in> o(\<lambda>x. b x powr e')" unfolding list.sel .
+  
+  define eps where "eps = e/2"
+  have "(\<lambda>x. ln (b x)) \<in> o(\<lambda>x. b x powr (e - eps))" unfolding eps_def
+    by (rule landau_o.small.compose[OF _ b_limit])
+       (insert e'(1) \<open>e - 0 > 0\<close>, simp add: ln_smallo_powr)
+  also from assms e'(1) have "eval C \<in> \<omega>(\<lambda>x. b x powr -eps)" unfolding eps_def
+    by (intro is_expansion_imp_smallomega[of basis])
+       (auto simp: basis_wf_Cons expands_to.simps trimmed_pos_def trimmed_ms_aux_MSLCons
+             elim!: is_expansion_aux_MSLCons)
+  hence "(\<lambda>x. b x powr -eps * b x powr e) \<in> o(\<lambda>x. eval C x * b x powr e)"
+    by (intro landau_o.small_big_mult) (simp_all add: smallomega_iff_smallo)
+  hence "(\<lambda>x. b x powr (e - eps)) \<in> o(\<lambda>x. eval C x * b x powr e)"
+    by (simp add: powr_add [symmetric] field_simps)
+  finally have "(\<lambda>x. ln (b x)) \<in> o(\<lambda>x. eval C x * b x powr e)" .
+  also {
+    from e' have "(\<lambda>x. g x - eval C x * b x powr e) \<in> o(\<lambda>x. b x powr (e' - e) * b x powr e)"
+      by (simp add: powr_add [symmetric])
+    also from assms e'(1) have "eval C \<in> \<omega>(\<lambda>x. b x powr (e' - e))" unfolding eps_def
+      by (intro is_expansion_imp_smallomega[of basis])
+         (auto simp: basis_wf_Cons expands_to.simps trimmed_pos_def trimmed_ms_aux_MSLCons
+               elim!: is_expansion_aux_MSLCons)
+    hence "(\<lambda>x. b x powr (e' - e) * b x powr e) \<in> o(\<lambda>x. eval C x * b x powr e)"
+      by (intro landau_o.small_big_mult is_expansion_imp_smallo) 
+         (simp_all add: smallomega_iff_smallo)
+    finally have "o(\<lambda>x. eval C x * b x powr e) = 
+                    o(\<lambda>x. g x - eval C x * b x powr e + eval C x * b x powr e)"
+      by (intro landau_o.small.plus_absorb1 [symmetric]) simp_all
+  }
+  also have "o(\<lambda>x. g x - eval C x * b x powr e + eval C x * b x powr e) = o(g)" by simp
+  also from assms have "g \<in> \<Theta>(f)" by (intro bigthetaI_cong) (simp add: expands_to.simps)
+  finally show "(\<lambda>x. ln (b x)) \<in> o(f)" .
+qed
+
+lemma basis_wf_insert_exp_uminus:
+  "(\<lambda>x. ln (b x)) \<in> o(f) \<Longrightarrow> (\<lambda>x. ln (b x)) \<in> o(\<lambda>x. -f x :: real)"
+  by simp
+
+lemma basis_wf_insert_exp_uminus':
+  "f \<in> o(\<lambda>x. ln (b x)) \<Longrightarrow> (\<lambda>x. -f x) \<in> o(\<lambda>x. ln (b x))"
+  by simp
+    
+lemma expands_to_exp_insert_pos:
+  fixes C :: "'a :: multiseries"
+  assumes "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)" 
+          "basis_wf (b # basis)" "trimmed_pos (MS (MSLCons (C, e) xs) g)" 
+          "(\<lambda>x. ln (b x)) \<in> o(f)"
+  shows   "filterlim (\<lambda>x. exp (f x)) at_top at_top"
+          "((\<lambda>x. ln (exp (f x))) expands_to MS (MSLCons (C, e) xs) g) (b # basis)"
+          "((\<lambda>x. exp (f x)) expands_to 
+             MS (MSLCons (const_expansion 1 :: 'a ms, 1) MSLNil) (\<lambda>x. exp (g x)))
+             ((\<lambda>x. exp (f x)) # b # basis)" (is ?thesis3)
+proof -
+  have "ln \<in> \<omega>(\<lambda>x::real. 1)"
+    by (intro smallomegaI_filterlim_at_top_norm)
+       (auto intro!: filterlim_compose[OF filterlim_abs_real] ln_at_top)
+  with assms have "(\<lambda>x. 1) \<in> o(\<lambda>x. ln (b x))"
+    by (intro landau_o.small.compose[of _ at_top _ b])
+       (simp_all add: basis_wf_Cons smallomega_iff_smallo)
+  also note \<open>(\<lambda>x. ln (b x)) \<in> o(f)\<close>
+  finally have f: "f \<in> \<omega>(\<lambda>_. 1)" by (simp add: smallomega_iff_smallo)
+  hence f: "filterlim (\<lambda>x. abs (f x)) at_top at_top"
+    by (auto dest: smallomegaD_filterlim_at_top_norm)
+  
+  define c where "c = fst (dominant_term C)"
+  define es where "es = snd (dominant_term C)"
+  from trimmed_imp_eventually_sgn[OF assms(2), of "MS (MSLCons (C, e) xs) g"] assms 
+    have "\<forall>\<^sub>F x in at_top. abs (f x) = f x"
+      by (auto simp: dominant_term_ms_aux_MSLCons trimmed_pos_def expands_to.simps sgn_if 
+               split: if_splits elim: eventually_elim2)
+  from filterlim_cong[OF refl refl this] f 
+    show *: "filterlim (\<lambda>x. exp (f x)) at_top at_top" 
+    by (auto intro: filterlim_compose[OF exp_at_top])
+  
+  {
+    fix e' :: real assume e': "e' > 1"
+    from assms have "(\<lambda>x. exp (g x)) \<in> \<Theta>(\<lambda>x. exp (f x) powr 1)"
+      by (intro bigthetaI_cong) (auto elim: eventually_mono simp: expands_to.simps)
+    also from e' * have "(\<lambda>x. exp (f x) powr 1) \<in> o(\<lambda>x. exp (f x) powr e')"
+      by (subst powr_smallo_iff) (insert *, simp_all)
+    finally have "(\<lambda>x. exp (g x)) \<in> o(\<lambda>x. exp (f x) powr e')" .
+  }
+  with assms show ?thesis3
+    by (auto intro!: is_expansion_aux.intros is_expansion_const 
+             simp: expands_to.simps is_expansion_aux_expansion_level
+             dest: is_expansion_aux_expansion_level)
+qed (insert assms, simp_all)
+    
+lemma expands_to_exp_insert_neg:
+  fixes C :: "'a :: multiseries"
+  assumes "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)" 
+          "basis_wf (b # basis)" "trimmed_neg (MS (MSLCons (C, e) xs) g)" 
+          "(\<lambda>x. ln (b x)) \<in> o(f)"
+  shows   "filterlim (\<lambda>x. exp (-f x)) at_top at_top" (is ?thesis1)
+          "((\<lambda>x. ln (exp (-f x))) expands_to MS (MSLCons (-C, e) (uminus_ms_aux xs)) (\<lambda>x. - g x)) 
+             (b # basis)"
+          "trimmed_pos (MS (MSLCons (-C, e) (uminus_ms_aux xs)) (\<lambda>x. - g x))"
+          "((\<lambda>x. exp (f x)) expands_to 
+             MS (MSLCons (const_expansion 1 :: 'a ms, -1) MSLNil) (\<lambda>x. exp (g x)))
+             ((\<lambda>x. exp (-f x)) # b # basis)" (is ?thesis3)
+proof -
+  have "ln \<in> \<omega>(\<lambda>x::real. 1)"
+    by (intro smallomegaI_filterlim_at_top_norm)
+       (auto intro!: filterlim_compose[OF filterlim_abs_real] ln_at_top)
+  with assms have "(\<lambda>x. 1) \<in> o(\<lambda>x. ln (b x))"
+    by (intro landau_o.small.compose[of _ at_top _ b])
+       (simp_all add: basis_wf_Cons smallomega_iff_smallo)
+  also note \<open>(\<lambda>x. ln (b x)) \<in> o(f)\<close>
+  finally have f: "f \<in> \<omega>(\<lambda>_. 1)" by (simp add: smallomega_iff_smallo)
+  hence f: "filterlim (\<lambda>x. abs (f x)) at_top at_top" 
+    by (auto dest: smallomegaD_filterlim_at_top_norm)
+  from trimmed_imp_eventually_sgn[OF assms(2), of "MS (MSLCons (C, e) xs) g"] assms 
+    have "\<forall>\<^sub>F x in at_top. abs (f x) = -f x"
+      by (auto simp: dominant_term_ms_aux_MSLCons trimmed_neg_def expands_to.simps sgn_if 
+               split: if_splits elim: eventually_elim2)
+  from filterlim_cong[OF refl refl this] f 
+    show *: "filterlim (\<lambda>x. exp (-f x)) at_top at_top" 
+    by (auto intro: filterlim_compose[OF exp_at_top])
+  
+  have "((\<lambda>x. -f x) expands_to -MS (MSLCons (C, e) xs) g) (b # basis)"
+    by (intro expands_to_uminus assms)
+  thus "((\<lambda>x. ln (exp (-f x))) expands_to MS (MSLCons (-C, e) (uminus_ms_aux xs)) (\<lambda>x. - g x)) 
+          (b # basis)"
+    by (simp add: uminus_ms_aux_MSLCons)
+      
+  {
+    fix e' :: real assume e': "e' > -1"
+    from assms have "(\<lambda>x. exp (g x)) \<in> \<Theta>(\<lambda>x. exp (-f x) powr -1)"
+      by (intro bigthetaI_cong) 
+         (auto elim: eventually_mono simp: expands_to.simps powr_minus exp_minus)
+    also from e' * have "(\<lambda>x. exp (-f x) powr -1) \<in> o(\<lambda>x. exp (-f x) powr e')"
+      by (subst powr_smallo_iff) (insert *, simp_all)
+    finally have "(\<lambda>x. exp (g x)) \<in> o(\<lambda>x. exp (-f x) powr e')" .
+  }
+  with assms show ?thesis3
+    by (auto intro!: is_expansion_aux.intros is_expansion_const 
+             simp: expands_to.simps is_expansion_aux_expansion_level powr_minus exp_minus
+             dest: is_expansion_aux_expansion_level)
+qed (insert assms, simp_all add: trimmed_pos_def trimmed_neg_def 
+       trimmed_ms_aux_MSLCons dominant_term_ms_aux_MSLCons)
+
+lemma is_expansion_aux_exp_neg:
+  assumes "is_expansion_aux F f basis" "basis_wf basis" "ms_exp_gt 0 (ms_aux_hd_exp F)"
+  shows   "is_expansion_aux (powser_ms_aux' exp_series_stream F) (\<lambda>x. exp (f x)) basis"
+  using powser_ms_aux'[OF convergent_powser'_exp assms(2,1,3)]
+  by (simp add: o_def powser_ms_aux'_def)
+
+lemma is_expansion_exp_neg:
+  assumes "is_expansion (MS (MSLCons (C, e) xs) f) basis" "basis_wf basis" "e < 0"
+  shows   "is_expansion (MS (powser_ms_aux' exp_series_stream (MSLCons (C, e) xs)) 
+             (\<lambda>x. exp (f x))) basis"
+  using is_expansion_aux_exp_neg[of "MSLCons (C, e) xs" f basis] assms by simp
+
+lemma expands_to_exp_neg:
+  assumes "(f expands_to MS (MSLCons (C, e) xs) g) basis" "basis_wf basis" "e - 0 < 0"
+  shows   "((\<lambda>x. exp (f x)) expands_to MS (powser_ms_aux' exp_series_stream (MSLCons (C, e) xs)) 
+             (\<lambda>x. exp (g x))) basis"
+  using assms is_expansion_exp_neg[of C e xs g basis] by (auto simp: expands_to.simps)
+
+lemma basis_wf_insert_exp_near:
+  assumes "basis_wf (b # basis)" "basis_wf ((\<lambda>x. exp (f x)) # basis)" "f \<in> o(\<lambda>x. ln (b x))"
+  shows   "basis_wf (b # (\<lambda>x. exp (f x)) # basis)"
+  using assms by (simp_all add: basis_wf_Cons)
+
+
+definition scale_ms_aux' where "scale_ms_aux' c F = scale_shift_ms_aux (c, 0) F"
+definition shift_ms_aux where "shift_ms_aux e F = scale_shift_ms_aux (const_expansion 1, e) F"
+
+lemma scale_ms_1 [simp]: "scale_ms 1 F = F"
+  by (simp add: scale_ms_def times_const_expansion_1)
+
+lemma scale_ms_aux_1 [simp]: "scale_ms_aux 1 xs = xs"
+  by (simp add: scale_ms_aux_def scale_shift_ms_aux_def map_prod_def msllist.map_ident
+        case_prod_unfold times_const_expansion_1)
+    
+lemma scale_ms_aux'_1 [simp]: "scale_ms_aux' (const_expansion 1) xs = xs"
+  by (simp add: scale_ms_aux'_def scale_shift_ms_aux_def map_prod_def msllist.map_ident
+        case_prod_unfold times_const_expansion_1)
+    
+lemma shift_ms_aux_0 [simp]: "shift_ms_aux 0 xs = xs"
+  by (simp add: shift_ms_aux_def scale_shift_ms_aux_def map_prod_def case_prod_unfold 
+        times_const_expansion_1 msllist.map_ident)
+  
+lemma scale_ms_aux'_MSLNil [simp]: "scale_ms_aux' C MSLNil = MSLNil"
+  by (simp add: scale_ms_aux'_def)
+
+lemma scale_ms_aux'_MSLCons [simp]: 
+  "scale_ms_aux' C (MSLCons (C', e) xs) = MSLCons (C * C', e) (scale_ms_aux' C xs)"
+  by (simp add: scale_ms_aux'_def scale_shift_ms_aux_MSLCons)
+
+lemma shift_ms_aux_MSLNil [simp]: "shift_ms_aux e MSLNil = MSLNil"
+  by (simp add: shift_ms_aux_def)
+
+lemma shift_ms_aux_MSLCons [simp]: 
+  "shift_ms_aux e (MSLCons (C, e') xs) = MSLCons (C, e' + e) (shift_ms_aux e xs)"
+  by (simp add: shift_ms_aux_def scale_shift_ms_aux_MSLCons times_const_expansion_1)
+
+lemma scale_ms_aux:
+  fixes xs :: "('a :: multiseries \<times> real) msllist"
+  assumes "is_expansion_aux xs f basis" "basis_wf basis"
+  shows   "is_expansion_aux (scale_ms_aux c xs) (\<lambda>x. c * f x) basis"
+proof (cases basis)
+  case (Cons b basis')
+  show ?thesis
+  proof (rule is_expansion_aux_cong)
+    show "is_expansion_aux (scale_ms_aux c xs) 
+            (\<lambda>x. eval (const_expansion c :: 'a) x * b x powr 0 * f x) basis"
+      using assms unfolding scale_ms_aux_def Cons
+      by (intro scale_shift_ms_aux is_expansion_const) 
+         (auto simp: basis_wf_Cons dest: is_expansion_aux_expansion_level)
+  next
+    from assms have "eventually (\<lambda>x. b x > 0) at_top" 
+      by (simp add: basis_wf_Cons Cons filterlim_at_top_dense)
+    thus "eventually (\<lambda>x. eval (const_expansion c :: 'a) x * b x powr 0 * f x = c * f x) at_top"
+      by eventually_elim simp
+  qed
+qed (insert assms, auto dest: is_expansion_aux_basis_nonempty)
+
+lemma scale_ms_aux':
+  assumes "is_expansion_aux xs f (b # basis)" "is_expansion C basis" "basis_wf (b # basis)"
+  shows   "is_expansion_aux (scale_ms_aux' C xs) (\<lambda>x. eval C x * f x) (b # basis)"
+proof (rule is_expansion_aux_cong)
+  show "is_expansion_aux (scale_ms_aux' C xs) (\<lambda>x. eval C x * b x powr 0 * f x) (b # basis)"
+    using assms unfolding scale_ms_aux'_def Cons by (intro scale_shift_ms_aux) simp_all
+next
+  from assms have "eventually (\<lambda>x. b x > 0) at_top" 
+    by (simp add: basis_wf_Cons filterlim_at_top_dense)
+  thus "eventually (\<lambda>x. eval C x * b x powr 0 * f x = eval C x * f x) at_top"
+    by eventually_elim simp
+qed
+
+lemma shift_ms_aux:
+  fixes xs :: "('a :: multiseries \<times> real) msllist"
+  assumes "is_expansion_aux xs f (b # basis)" "basis_wf (b # basis)"
+  shows   "is_expansion_aux (shift_ms_aux e xs) (\<lambda>x. b x powr e * f x) (b # basis)"
+  unfolding shift_ms_aux_def 
+  using scale_shift_ms_aux[OF assms(2,1) is_expansion_const[of _ 1], of e] assms
+  by (auto dest!: is_expansion_aux_expansion_level simp: basis_wf_Cons)
+  
+lemma expands_to_exp_0:
+  assumes "(f expands_to MS (MSLCons (MS ys c, e) xs) g) (b # basis)"
+          "basis_wf (b # basis)" "e - 0 = 0" "((\<lambda>x. exp (c x)) expands_to E) basis"
+  shows   "((\<lambda>x. exp (f x)) expands_to
+             MS (scale_ms_aux' E (powser_ms_aux' exp_series_stream xs)) (\<lambda>x. exp (g x))) (b # basis)"
+          (is "(_ expands_to MS ?H _) _")
+proof -
+  from assms have "is_expansion_aux ?H (\<lambda>x. eval E x * exp (g x - c x * b x powr 0)) (b # basis)"
+    by (intro scale_ms_aux' is_expansion_aux_exp_neg)
+       (auto elim!: is_expansion_aux_MSLCons simp: expands_to.simps)
+  moreover {
+    from \<open>basis_wf (b # basis)\<close> have "eventually (\<lambda>x. b x > 0) at_top"
+      by (simp add: filterlim_at_top_dense basis_wf_Cons)
+    moreover from assms(4) have "eventually (\<lambda>x. eval E x = exp (c x)) at_top"
+      by (simp add: expands_to.simps)
+    ultimately have "eventually (\<lambda>x. eval E x * exp (g x - c x * b x powr 0) = exp (g x)) at_top"
+      by eventually_elim (simp add: exp_diff)
+  }
+  ultimately have "is_expansion_aux ?H (\<lambda>x. exp (g x)) (b # basis)"
+    by (rule is_expansion_aux_cong)
+  with assms(1) show ?thesis by (simp add: expands_to.simps)
+qed
+
+lemma expands_to_exp_0_real:
+  assumes "(f expands_to MS (MSLCons (c::real, e) xs) g) [b]" "basis_wf [b]" "e - 0 = 0"
+  shows   "((\<lambda>x. exp (f x)) expands_to 
+             MS (scale_ms_aux (exp c) (powser_ms_aux' exp_series_stream xs)) (\<lambda>x. exp (g x))) [b]"
+          (is "(_ expands_to MS ?H _) _")
+proof -
+  from assms have "is_expansion_aux ?H (\<lambda>x. exp c * exp (g x - c * b x powr 0)) [b]"
+    by (intro scale_ms_aux is_expansion_aux_exp_neg)
+       (auto elim!: is_expansion_aux_MSLCons simp: expands_to.simps)
+  moreover from \<open>basis_wf [b]\<close> have "eventually (\<lambda>x. b x > 0) at_top"
+    by (simp add: filterlim_at_top_dense basis_wf_Cons)
+  hence "eventually (\<lambda>x. exp c * exp (g x - c * b x powr 0) = exp (g x)) at_top"
+    by eventually_elim (simp add: exp_diff)
+  ultimately have "is_expansion_aux ?H (\<lambda>x. exp (g x)) [b]"
+    by (rule is_expansion_aux_cong)
+  with assms(1) show ?thesis by (simp add: expands_to.simps)
+qed
+
+lemma expands_to_exp_0_pull_out1:
+  assumes "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)"
+  assumes "((\<lambda>x. ln (b x)) expands_to L) basis" "basis_wf (b # basis)" "e - 0 = 0" "c \<equiv> c"
+  shows   "((\<lambda>x. f x - c * ln (b x)) expands_to 
+             MS (MSLCons (C - scale_ms c L, e) xs) (\<lambda>x. g x - c * ln (b x))) (b # basis)"
+proof -
+  from \<open>basis_wf (b # basis)\<close> have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  have "(\<lambda>x. c * ln (b x)) \<in> o(\<lambda>x. b x powr e')" if "e' > 0" for e'
+    using that by (intro landau_o.small.compose[OF _ b_limit]) (simp add: ln_smallo_powr)
+  hence *: "(\<lambda>x. g x - c * ln (b x)) \<in> o(\<lambda>x. b x powr e')" if "e' > 0" for e' using assms(1,4) that
+    by (intro sum_in_smallo) (auto simp: expands_to.simps elim!: is_expansion_aux_MSLCons)
+  
+  have "is_expansion_aux xs (\<lambda>x. (g x - c * b x powr 0 * eval L x) -
+           eval (C - scale_ms c L) x * b x powr e) (b # basis)" 
+    (is "is_expansion_aux _ ?h _") using assms
+    by (auto simp: expands_to.simps algebra_simps scale_ms_def elim!: is_expansion_aux_MSLCons)
+  moreover from b_limit have "eventually (\<lambda>x. b x > 0) at_top"
+    by (simp add: filterlim_at_top_dense)
+  hence "eventually (\<lambda>x. ?h x = g x - c * ln (b x) - 
+           eval (C - const_expansion c * L) x * b x powr e) at_top" 
+    (is "eventually (\<lambda>x. _ = ?h x) _") using assms(2)
+    by (auto simp: expands_to.simps scale_ms_def elim: eventually_elim2)
+  ultimately have **: "is_expansion_aux xs ?h (b # basis)" by (rule is_expansion_aux_cong)
+  from assms(1,4) have "ms_exp_gt 0 (ms_aux_hd_exp xs)" "is_expansion C basis"
+    by (auto simp: expands_to.simps elim!: is_expansion_aux_MSLCons)
+  moreover from assms(1) have "length basis = expansion_level TYPE('a)"
+    by (auto simp: expands_to.simps dest: is_expansion_aux_expansion_level)
+  ultimately have "is_expansion_aux (MSLCons (C - scale_ms c L, e) xs) 
+                     (\<lambda>x. g x - c * ln (b x)) (b # basis)" using assms unfolding scale_ms_def
+    by (intro is_expansion_aux.intros is_expansion_minus is_expansion_mult is_expansion_const * **)
+       (simp_all add: basis_wf_Cons expands_to.simps)
+  with assms(1) show ?thesis by (auto simp: expands_to.simps)
+qed
+    
+lemma expands_to_exp_0_pull_out2:
+  assumes "((\<lambda>x. exp (f x - c * ln (b x))) expands_to MS xs g) (b # basis)"
+  assumes "basis_wf (b # basis)"
+  shows   "((\<lambda>x. exp (f x)) expands_to MS (shift_ms_aux c xs) 
+             (\<lambda>x. b x powr c * g x)) (b # basis)"
+proof (rule expands_to.intros)
+  show "is_expansion (MS (shift_ms_aux c xs) (\<lambda>x. b x powr c * g x)) (b # basis)"
+    using assms unfolding expands_to.simps by (auto intro!: shift_ms_aux)
+  from assms have "eventually (\<lambda>x. b x > 0) at_top" 
+    by (simp add: basis_wf_Cons filterlim_at_top_dense)
+  with assms(1)
+  show "\<forall>\<^sub>F x in at_top. eval (MS (shift_ms_aux c xs) (\<lambda>x. b x powr c * g x)) x = exp (f x)"
+    by (auto simp: expands_to.simps exp_diff powr_def elim: eventually_elim2)
+qed
+
+lemma expands_to_exp_0_pull_out2_nat:
+  assumes "((\<lambda>x. exp (f x - c * ln (b x))) expands_to MS xs g) (b # basis)"
+  assumes "basis_wf (b # basis)" "c = real n"
+  shows   "((\<lambda>x. exp (f x)) expands_to MS (shift_ms_aux c xs) 
+             (\<lambda>x. b x ^ n * g x)) (b # basis)"
+using expands_to_exp_0_pull_out2[OF assms(1-2)]
+proof (rule expands_to_cong')
+  from assms(2) have "eventually (\<lambda>x. b x > 0) at_top" 
+    by (simp add: filterlim_at_top_dense basis_wf_Cons)
+  with assms(3) show "eventually (\<lambda>x. b x powr c * g x = b x ^ n * g x) at_top"
+    by (auto elim!: eventually_mono simp: powr_realpow)
+qed
+
+lemma expands_to_exp_0_pull_out2_neg_nat:
+  assumes "((\<lambda>x. exp (f x - c * ln (b x))) expands_to MS xs g) (b # basis)"
+  assumes "basis_wf (b # basis)" "c = -real n"
+  shows   "((\<lambda>x. exp (f x)) expands_to MS (shift_ms_aux c xs) 
+             (\<lambda>x. g x / b x ^ n)) (b # basis)"
+using expands_to_exp_0_pull_out2[OF assms(1-2)]
+proof (rule expands_to_cong')
+  from assms(2) have "eventually (\<lambda>x. b x > 0) at_top" 
+    by (simp add: filterlim_at_top_dense basis_wf_Cons)
+  with assms(3) show "eventually (\<lambda>x. b x powr c * g x = g x / b x ^ n) at_top"
+    by (auto elim!: eventually_mono simp: powr_realpow powr_minus divide_simps)
+qed
+
+lemma eval_monom_collapse: "c * eval_monom (c', es) basis x = eval_monom (c * c', es) basis x"
+  by (simp add: eval_monom_def)
+
+lemma eval_monom_1_conv: "eval_monom a basis = (\<lambda>x. fst a * eval_monom (1, snd a) basis x)"
+  by (simp add: eval_monom_def case_prod_unfold)
+
+
+subsubsection \<open>Comparing exponent lists\<close>
+
+primrec flip_cmp_result where
+  "flip_cmp_result LT = GT"
+| "flip_cmp_result GT = LT"
+| "flip_cmp_result EQ = EQ"
+
+fun compare_lists :: "real list \<Rightarrow> real list \<Rightarrow> cmp_result" where
+  "compare_lists [] [] = EQ"
+| "compare_lists (a # as) (b # bs) = 
+     (if a < b then LT else if b < a then GT else compare_lists as bs)"
+| "compare_lists _ _ = undefined"
+
+declare compare_lists.simps [simp del]
+  
+lemma compare_lists_Nil [simp]: "compare_lists [] [] = EQ" by (fact compare_lists.simps)
+
+lemma compare_lists_Cons [simp]:
+  "a < b \<Longrightarrow> compare_lists (a # as) (b # bs) = LT"
+  "a > b \<Longrightarrow> compare_lists (a # as) (b # bs) = GT"
+  "a = b \<Longrightarrow> compare_lists (a # as) (b # bs) = compare_lists as bs"
+  by (simp_all add: compare_lists.simps(2))
+
+lemma flip_compare_lists:
+  "length as = length bs \<Longrightarrow> flip_cmp_result (compare_lists as bs) = compare_lists bs as"
+  by (induction as bs rule: compare_lists.induct) (auto simp: compare_lists.simps(2))
+
+lemma compare_lists_induct [consumes 1, case_names Nil Eq Neq]:
+  fixes as bs :: "real list"
+  assumes "length as = length bs"
+  assumes "P [] []"
+  assumes "\<And>a as bs. P as bs \<Longrightarrow> P (a # as) (a # bs)"
+  assumes "\<And>a as b bs. a \<noteq> b \<Longrightarrow> P (a # as) (b # bs)"
+  shows   "P as bs"
+  using assms(1)
+proof (induction as bs rule: list_induct2)
+  case (Cons a as b bs)
+  thus ?case by (cases "a < b") (auto intro: assms)
+qed (simp_all add: assms)
+  
+
+lemma eval_monom_smallo_eval_monom:
+  assumes "length es1 = length es2" "length es2 = length basis" "basis_wf basis"
+  assumes "compare_lists es1 es2 = LT"
+  shows   "eval_monom (1, es1) basis \<in> o(eval_monom (1, es2) basis)"
+using assms
+proof (induction es1 es2 basis rule: list_induct3)
+  case (Cons e1 es1 e2 es2 b basis)
+  show ?case
+  proof (cases "e1 = e2")
+    case True
+    from Cons.prems have "eventually (\<lambda>x. b x > 0) at_top" 
+      by (simp add: basis_wf_Cons filterlim_at_top_dense)
+    with Cons True show ?thesis
+      by (auto intro: landau_o.small_big_mult simp: eval_monom_Cons basis_wf_Cons)
+  next
+    case False
+    with Cons.prems have "e1 < e2" by (cases e1 e2 rule: linorder_cases) simp_all
+    with Cons.prems have
+       "(\<lambda>x. eval_monom (1, es1) basis x * eval_monom (inverse_monom (1, es2)) basis x *
+           b x powr e1) \<in> o(\<lambda>x. b x powr ((e2 - e1)/2) * b x powr ((e2 - e1)/2) * b x powr e1)"
+      (is "?f \<in> _") by (intro landau_o.small_big_mult[OF _ landau_o.big_refl] landau_o.small.mult 
+                          eval_monom_smallo') simp_all
+    also have "\<dots> = o(\<lambda>x. b x powr e2)" by (simp add: powr_add [symmetric])
+    also have "?f = (\<lambda>x. eval_monom (1, e1 # es1) (b # basis) x / eval_monom (1, es2) basis x)"
+      by (simp add: eval_inverse_monom field_simps eval_monom_Cons)
+    also have "\<dots> \<in> o(\<lambda>x. b x powr e2) \<longleftrightarrow> 
+                 eval_monom (1, e1 # es1) (b # basis) \<in> o(eval_monom (1, e2 # es2) (b # basis))"
+      using Cons.prems 
+      by (subst landau_o.small.divide_eq2)
+         (simp_all add: eval_monom_Cons mult_ac eval_monom_nonzero basis_wf_Cons)
+    finally show ?thesis .
+  qed
+qed simp_all
+
+lemma eval_monom_eq:
+  assumes "length es1 = length es2" "length es2 = length basis" "basis_wf basis"
+  assumes "compare_lists es1 es2 = EQ"
+  shows   "eval_monom (1, es1) basis = eval_monom (1, es2) basis"
+  using assms
+  by (induction es1 es2 basis rule: list_induct3)
+     (auto simp: basis_wf_Cons eval_monom_Cons compare_lists.simps(2) split: if_splits)
+
+definition compare_expansions :: "'a :: multiseries \<Rightarrow> 'a \<Rightarrow> cmp_result \<times> real \<times> real" where
+  "compare_expansions F G =
+     (case compare_lists (snd (dominant_term F)) (snd (dominant_term G)) of
+        LT \<Rightarrow> (LT, 0, 0)
+      | GT \<Rightarrow> (GT, 0, 0)
+      | EQ \<Rightarrow> (EQ, fst (dominant_term F),  fst (dominant_term G)))"
+
+lemma compare_expansions_real:
+  "compare_expansions (c1 :: real) c2 = (EQ, c1, c2)"
+  by (simp add: compare_expansions_def)
+
+lemma compare_expansions_MSLCons_eval:
+  "compare_expansions (MS (MSLCons (C1, e1) xs) f) (MS (MSLCons (C2, e2) ys) g) =
+     CMP_BRANCH (COMPARE e1 e2) (LT, 0, 0) (compare_expansions C1 C2) (GT, 0, 0)"
+  by (simp add: compare_expansions_def dominant_term_ms_aux_def case_prod_unfold 
+        COMPARE_def CMP_BRANCH_def split: cmp_result.splits)
+
+lemma compare_expansions_LT_I:
+  assumes "e1 - e2 < 0"
+  shows   "compare_expansions (MS (MSLCons (C1, e1) xs) f) (MS (MSLCons (C2, e2) ys) g) = (LT, 0, 0)"
+  using assms by (simp add: compare_expansions_MSLCons_eval CMP_BRANCH_def COMPARE_def)
+
+lemma compare_expansions_GT_I:
+  assumes "e1 - e2 > 0"
+  shows   "compare_expansions (MS (MSLCons (C1, e1) xs) f) (MS (MSLCons (C2, e2) ys) g) = (GT, 0, 0)"
+  using assms by (simp add: compare_expansions_MSLCons_eval CMP_BRANCH_def COMPARE_def)
+
+lemma compare_expansions_same_exp:
+  assumes "e1 - e2 = 0" "compare_expansions C1 C2 = res"
+  shows   "compare_expansions (MS (MSLCons (C1, e1) xs) f) (MS (MSLCons (C2, e2) ys) g) = res"
+  using assms by (simp add: compare_expansions_MSLCons_eval CMP_BRANCH_def COMPARE_def)
+  
+
+lemma compare_expansions_GT_flip:
+  "length (snd (dominant_term F)) = length (snd (dominant_term G)) \<Longrightarrow>
+     compare_expansions F G = (GT, c) \<longleftrightarrow> compare_expansions G F = (LT, c)"
+  using flip_compare_lists[of "snd (dominant_term F)" "snd (dominant_term G)"]
+  by (auto simp: compare_expansions_def split: cmp_result.splits)
+
+lemma compare_expansions_LT:
+  assumes "compare_expansions F G = (LT, c, c')" "trimmed G"
+          "(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis"
+  shows   "f \<in> o(g)"
+proof -
+  from assms have "f \<in> \<Theta>(eval F)"
+    by (auto simp: expands_to.simps eq_commute intro!: bigthetaI_cong)
+  also have "eval F \<in> O(eval_monom (1, snd (dominant_term F)) basis)"
+    using assms by (intro dominant_term_bigo) (auto simp: expands_to.simps)
+  also have "eval_monom (1, snd (dominant_term F)) basis \<in> 
+               o(eval_monom (1, snd (dominant_term G)) basis)" using assms 
+    by (intro eval_monom_smallo_eval_monom)
+       (auto simp: length_dominant_term expands_to.simps is_expansion_length compare_expansions_def
+             split: cmp_result.splits)
+  also have "eval_monom (1, snd (dominant_term G)) basis \<in> \<Theta>(eval_monom (dominant_term G) basis)"
+    by (subst (2) eval_monom_1_conv, subst landau_theta.cmult)
+       (insert assms, simp_all add: trimmed_imp_dominant_term_nz)    
+  also have "eval_monom (dominant_term G) basis \<in> \<Theta>(g)"
+    by (intro asymp_equiv_imp_bigtheta[OF asymp_equiv_symI] dominant_term_expands_to
+              asymp_equiv_imp_bigtheta assms)
+  finally show ?thesis .
+qed
+
+lemma compare_expansions_GT:
+  assumes "compare_expansions F G = (GT, c, c')" "trimmed F"
+          "(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis"
+  shows   "g \<in> o(f)"
+  by (rule compare_expansions_LT[OF _ assms(2,4,3)])
+     (insert assms, simp_all add: compare_expansions_GT_flip length_dominant_term)
+
+lemma compare_expansions_EQ:
+  assumes "compare_expansions F G = (EQ, c, c')" "trimmed F" "trimmed G"
+          "(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis"
+  shows   "(\<lambda>x. c' * f x) \<sim>[at_top] (\<lambda>x. c * g x)"
+proof -
+  from assms(1) have c: "c = fst (dominant_term F)" "c' = fst (dominant_term G)"
+    by (auto simp: compare_expansions_def split: cmp_result.splits)
+  have "(\<lambda>x. c' * f x) \<sim>[at_top] (\<lambda>x. c' * (c * eval_monom (1, snd (dominant_term F)) basis x))" 
+    unfolding c by (rule asymp_equiv_mult, rule asymp_equiv_refl, subst eval_monom_collapse)
+                   (auto intro: dominant_term_expands_to assms)
+  also have "eval_monom (1, snd (dominant_term F)) basis =
+               eval_monom (1, snd (dominant_term G)) basis" using assms
+    by (intro eval_monom_eq)
+       (simp_all add: compare_expansions_def length_dominant_term is_expansion_length 
+          expands_to.simps split: cmp_result.splits)
+  also have "(\<lambda>x. c' * (c * \<dots> x)) = (\<lambda>x. c * (c' * \<dots> x))" by (simp add: mult_ac)
+  also have "\<dots> \<sim>[at_top] (\<lambda>x. c * g x)"
+    unfolding c by (rule asymp_equiv_mult, rule asymp_equiv_refl, subst eval_monom_collapse, 
+                       rule asymp_equiv_symI)  (auto intro: dominant_term_expands_to assms)
+  finally show ?thesis by (simp add: asymp_equiv_sym)
+qed
+
+lemma compare_expansions_EQ_imp_bigo:
+  assumes "compare_expansions F G = (EQ, c, c')" "trimmed G"
+          "(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis"
+  shows   "f \<in> O(g)"
+proof -
+  from assms(1) have c: "c = fst (dominant_term F)" "c' = fst (dominant_term G)"
+    by (auto simp: compare_expansions_def split: cmp_result.splits)
+  from assms(3) have [simp]: "expansion_level TYPE('a) = length basis"
+    by (auto simp: expands_to.simps is_expansion_length)
+
+  have "f \<in> \<Theta>(eval F)"
+    using assms by (intro bigthetaI_cong) (auto simp: expands_to.simps eq_commute)
+  also have "eval F \<in> O(eval_monom (1, snd (dominant_term F)) basis)"
+    using assms by (intro dominant_term_bigo assms) (auto simp: expands_to.simps)
+  also have "eval_monom (1, snd (dominant_term F)) basis =
+               eval_monom (1, snd (dominant_term G)) basis" using assms
+    by (intro eval_monom_eq) (auto simp: compare_expansions_def case_prod_unfold
+                                length_dominant_term split: cmp_result.splits)
+  also have "\<dots> \<in> O(eval_monom (dominant_term G) basis)" using assms(2)
+    by (auto simp add: eval_monom_def case_prod_unfold dest: trimmed_imp_dominant_term_nz)
+  also have "eval_monom (dominant_term G) basis \<in> \<Theta>(eval G)"
+    by (rule asymp_equiv_imp_bigtheta, rule asymp_equiv_symI, rule dominant_term)
+       (insert assms, auto simp: expands_to.simps)
+  also have "eval G \<in> \<Theta>(g)"
+    using assms by (intro bigthetaI_cong) (auto simp: expands_to.simps)
+  finally show ?thesis .
+qed
+
+lemma compare_expansions_EQ_same:
+  assumes "compare_expansions F G = (EQ, c, c')" "trimmed F" "trimmed G"
+          "(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis"
+          "c = c'"
+  shows   "f \<sim>[at_top] g"
+proof -
+  from assms have [simp]: "c' \<noteq> 0" 
+    by (auto simp: compare_expansions_def trimmed_imp_dominant_term_nz split: cmp_result.splits)
+  have "(\<lambda>x. inverse c * (c' * f x)) \<sim>[at_top] (\<lambda>x. inverse c * (c * g x))"
+    by (rule asymp_equiv_mult[OF asymp_equiv_refl]) (rule compare_expansions_EQ[OF assms(1-6)])
+  with assms(7) show ?thesis by (simp add: divide_simps)
+qed
+  
+lemma compare_expansions_EQ_imp_bigtheta:
+  assumes "compare_expansions F G = (EQ, c, c')" "trimmed F" "trimmed G"
+          "(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis"
+  shows   "f \<in> \<Theta>(g)"
+proof -
+  from assms have "(\<lambda>x. c' * f x) \<in> \<Theta>(\<lambda>x. c * g x)"
+    by (intro asymp_equiv_imp_bigtheta compare_expansions_EQ)
+  moreover from assms have "c \<noteq> 0" "c' \<noteq> 0"
+    by (auto simp: compare_expansions_def trimmed_imp_dominant_term_nz split: cmp_result.splits)
+  ultimately show ?thesis by simp
+qed
+
+lemma expands_to_MSLCons_0_asymp_equiv_hd:
+  assumes "(f expands_to (MS (MSLCons (C, 0) xs) g)) basis" "trimmed (MS (MSLCons (C, 0) xs) f)"
+          "basis_wf basis"
+  shows   "f \<sim>[at_top] eval C"
+proof -
+  from assms(1) is_expansion_aux_basis_nonempty obtain b basis' where [simp]: "basis = b # basis'"
+    by (cases basis) (auto simp: expands_to.simps)
+  from assms have b_pos: "eventually (\<lambda>x. b x > 0) at_top" 
+    by (simp add: filterlim_at_top_dense basis_wf_Cons)
+  from assms have "f \<sim>[at_top] eval_monom (dominant_term (MS (MSLCons (C, 0) xs) g)) basis" 
+    by (intro dominant_term_expands_to) simp_all
+  also have "\<dots> = (\<lambda>x. eval_monom (dominant_term C) basis' x * b x powr 0)"
+    by (simp_all add: dominant_term_ms_aux_def case_prod_unfold eval_monom_Cons)
+  also have "eventually (\<lambda>x. \<dots> x = eval_monom (dominant_term C) basis' x) at_top"
+    using b_pos by eventually_elim simp
+  also from assms have "eval_monom (dominant_term C) basis' \<sim>[at_top] eval C"
+    by (intro asymp_equiv_symI [OF dominant_term_expands_to] 
+          expands_to_hd''[of f C 0 xs g b basis']) (auto simp: trimmed_ms_aux_MSLCons basis_wf_Cons)
+  finally show ?thesis .
+qed
+    
+  
+lemma compare_expansions_LT':
+  assumes "compare_expansions (MS ys h) G \<equiv> (LT, c, c')" "trimmed G"
+          "(f expands_to (MS (MSLCons (MS ys h, e) xs) f')) (b # basis)" "(g expands_to G) basis"
+          "e = 0" "basis_wf (b # basis)"
+  shows   "h \<in> o(g)"   
+proof -
+  from assms show ?thesis
+    by (intro compare_expansions_LT[OF _ assms(2) _ assms(4)])
+       (auto intro: expands_to_hd'' simp: trimmed_ms_aux_MSLCons basis_wf_Cons expands_to.simps 
+             elim!: is_expansion_aux_MSLCons)
+qed
+
+lemma compare_expansions_GT':
+  assumes "compare_expansions C G \<equiv> (GT, c, c')"
+          "trimmed (MS (MSLCons (C, e) xs) f')"
+          "(f expands_to (MS (MSLCons (C, e) xs) f')) (b # basis)" "(g expands_to G) basis"
+          "e = 0" "basis_wf (b # basis)"
+  shows   "g \<in> o(f)"
+proof -
+  from assms have "g \<in> o(eval C)"
+    by (intro compare_expansions_GT[of C G c c' _ basis])
+       (auto intro: expands_to_hd'' simp: trimmed_ms_aux_MSLCons basis_wf_Cons)
+  also from assms have "f \<in> \<Theta>(eval C)"
+    by (intro asymp_equiv_imp_bigtheta expands_to_MSLCons_0_asymp_equiv_hd[of f C xs f' "b#basis"])
+       auto
+  finally show ?thesis .
+qed
+
+lemma compare_expansions_EQ':
+  assumes "compare_expansions C G = (EQ, c, c')" 
+          "trimmed (MS (MSLCons (C, e) xs) f')" "trimmed G"
+          "(f expands_to (MS (MSLCons (C, e) xs) f')) (b # basis)" "(g expands_to G) basis"
+          "e = 0" "basis_wf (b # basis)"
+  shows   "f \<sim>[at_top] (\<lambda>x. c / c' * g x)" 
+proof -
+  from assms have [simp]: "c' \<noteq> 0" 
+    by (auto simp: compare_expansions_def trimmed_imp_dominant_term_nz split: cmp_result.splits)
+  from assms have "f \<sim>[at_top] eval C"
+    by (intro expands_to_MSLCons_0_asymp_equiv_hd[of f C xs f' "b#basis"]) auto
+  also from assms(2,4,7) have *: "(\<lambda>x. c' * eval C x) \<sim>[at_top] (\<lambda>x. c * g x)"
+    by (intro compare_expansions_EQ[OF assms(1) _ assms(3) _ assms(5)])
+       (auto intro: expands_to_hd'' simp: trimmed_ms_aux_MSLCons basis_wf_Cons)
+  have "(\<lambda>x. inverse c' * (c' * eval C x)) \<sim>[at_top] (\<lambda>x. inverse c' * (c * g x))"
+    by (rule asymp_equiv_mult) (insert *, simp_all)
+  hence "eval C \<sim>[at_top] (\<lambda>x. c / c' * g x)" by (simp add: divide_simps)
+  finally show ?thesis .
+qed
+
+lemma expands_to_insert_ln: 
+  assumes "basis_wf [b]"
+  shows   "((\<lambda>x. ln (b x)) expands_to MS (MSLCons (1::real,1) MSLNil) (\<lambda>x. ln (b x))) [\<lambda>x. ln (b x)]"
+proof -
+  from assms have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b: "eventually (\<lambda>x. b x > 1) at_top" by (simp add: filterlim_at_top_dense)
+  have "(\<lambda>x::real. x) \<in> \<Theta>(\<lambda>x. x powr 1)"
+    by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto
+  also have "(\<lambda>x::real. x powr 1) \<in> o(\<lambda>a. a powr e)" if "e > 1" for e
+    by (subst powr_smallo_iff) (auto simp: that filterlim_ident)
+  finally show ?thesis using b assms
+    by (auto intro!: is_expansion_aux.intros landau_o.small.compose[of _ at_top _ "\<lambda>x. ln (b x)"]
+                     filterlim_compose[OF ln_at_top b_limit] is_expansion_real.intros
+             elim!: eventually_mono simp: expands_to.simps)
+qed
+
+lemma trimmed_pos_insert_ln:
+  assumes "basis_wf [b]"
+  shows   "trimmed_pos (MS (MSLCons (1::real, 1) MSLNil) (\<lambda>x. ln (b x)))"
+  by (simp_all add: trimmed_ms_aux_def)
+
+
+primrec compare_list_0 where
+  "compare_list_0 [] = EQ"
+| "compare_list_0 (c # cs) = CMP_BRANCH (COMPARE c 0) LT (compare_list_0 cs) GT"
+
+
+lemma compare_reals_diff_sgnD:
+  "a - (b :: real) < 0 \<Longrightarrow> a < b" "a - b = 0 \<Longrightarrow> a = b" "a - b > 0 \<Longrightarrow> a > b"
+  by simp_all
+
+lemma expands_to_real_imp_filterlim:
+  assumes "(f expands_to (c :: real)) basis"
+  shows   "(f \<longlongrightarrow> c) at_top"
+  using assms by (auto elim!: expands_to.cases simp: eq_commute[of c] Lim_eventually)
+
+lemma expands_to_MSLNil_imp_filterlim:
+  assumes "(f expands_to MS MSLNil f') basis"
+  shows   "(f \<longlongrightarrow> 0) at_top"
+proof -
+  from assms have "eventually (\<lambda>x. f' x = 0) at_top" "eventually (\<lambda>x. f' x = f x) at_top"
+    by (auto elim!: expands_to.cases is_expansion_aux.cases[of MSLNil])
+  hence "eventually (\<lambda>x. f x = 0) at_top" by eventually_elim auto
+  thus ?thesis by (simp add: Lim_eventually)
+qed
+
+lemma expands_to_neg_exponent_imp_filterlim:
+  assumes "(f expands_to MS (MSLCons (C, e) xs) f') basis" "basis_wf basis" "e < 0"
+  shows   "(f \<longlongrightarrow> 0) at_top"
+proof -
+  let ?F = "MS (MSLCons (C, e) xs) f'"
+  let ?es = "snd (dominant_term ?F)"
+  from assms have exp: "is_expansion ?F basis"
+    by (simp add: expands_to.simps)
+  from assms have "f \<in> \<Theta>(f')" by (intro bigthetaI_cong) (simp add: expands_to.simps eq_commute)
+  also from dominant_term_bigo[OF assms(2) exp]
+    have "f' \<in> O(eval_monom (1, ?es) basis)" by simp
+  also have "(eval_monom (1, ?es) basis) \<in> o(\<lambda>x. hd basis x powr 0)" using exp assms
+    by (intro eval_monom_smallo)
+       (auto simp: is_expansion_aux_basis_nonempty dominant_term_ms_aux_def
+                   case_prod_unfold length_dominant_term is_expansion_aux_expansion_level)
+  also have "(\<lambda>x. hd basis x powr 0) \<in> O(\<lambda>_. 1)"
+    by (intro landau_o.bigI[of 1]) auto
+  finally show ?thesis by (auto dest!: smalloD_tendsto)
+qed
+
+lemma expands_to_pos_exponent_imp_filterlim:
+  assumes "(f expands_to MS (MSLCons (C, e) xs) f') basis" "trimmed (MS (MSLCons (C, e) xs) f')"
+          "basis_wf basis" "e > 0"
+  shows   "filterlim f at_infinity at_top"
+proof -
+  let ?F = "MS (MSLCons (C, e) xs) f'"
+  let ?es = "snd (dominant_term ?F)"
+  from assms have exp: "is_expansion ?F basis"
+    by (simp add: expands_to.simps)
+  with assms have "filterlim (hd basis) at_top at_top"
+    using is_expansion_aux_basis_nonempty[of "MSLCons (C, e) xs" f' basis]
+    by (cases basis) (simp_all add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. hd basis x > 0) at_top" using filterlim_at_top_dense by blast
+
+  from assms have "f \<in> \<Theta>(f')" by (intro bigthetaI_cong) (simp add: expands_to.simps eq_commute)
+  also from dominant_term[OF assms(3) exp assms(2)]
+    have "f' \<in> \<Theta>(eval_monom (dominant_term ?F) basis)" by (simp add: asymp_equiv_imp_bigtheta)
+  also have "(eval_monom (dominant_term ?F) basis) \<in> \<Theta>(eval_monom (1, ?es) basis)"
+    using assms by (simp add: eval_monom_def case_prod_unfold dominant_term_ms_aux_def 
+                              trimmed_imp_dominant_term_nz trimmed_ms_aux_def)
+  also have "eval_monom (1, ?es) basis \<in> \<omega>(\<lambda>x. hd basis x powr 0)" using exp assms
+    by (intro eval_monom_smallomega)
+       (auto simp: is_expansion_aux_basis_nonempty dominant_term_ms_aux_def
+                   case_prod_unfold length_dominant_term is_expansion_aux_expansion_level)
+  also have "(\<lambda>x. hd basis x powr 0) \<in> \<Theta>(\<lambda>_. 1)"
+    by (intro bigthetaI_cong eventually_mono[OF b_pos]) auto
+  finally show ?thesis by (auto dest!: smallomegaD_filterlim_at_infinity)
+qed
+
+lemma expands_to_zero_exponent_imp_filterlim:
+  assumes "(f expands_to MS (MSLCons (C, e) xs) f') basis"
+          "basis_wf basis" "e = 0"
+  shows   "filterlim (eval C) at_infinity at_top \<Longrightarrow> filterlim f at_infinity at_top"
+    and   "filterlim (eval C) (nhds L) at_top \<Longrightarrow> filterlim f (nhds L) at_top"
+proof -
+  from assms obtain b basis' where *:
+    "basis = b # basis'" "is_expansion C basis'" "ms_exp_gt 0 (ms_aux_hd_exp xs)"
+    "is_expansion_aux xs (\<lambda>x. f' x - eval C x * b x powr 0) basis"
+    by (auto elim!: expands_to.cases is_expansion_aux_MSLCons)
+
+  from *(1) assms have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
+  hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" using filterlim_at_top_dense by blast
+
+  from assms(1) have "eventually (\<lambda>x. f' x = f x) at_top" by (simp add: expands_to.simps)
+  hence "eventually (\<lambda>x. f' x - eval C x * b x powr 0 = f x - eval C x) at_top"
+    using b_pos by eventually_elim auto
+  from *(4) and this have "is_expansion_aux xs (\<lambda>x. f x - eval C x) basis"
+    by (rule is_expansion_aux_cong)
+  hence "(\<lambda>x. f x - eval C x) \<in> o(\<lambda>x. hd basis x powr 0)"
+    by (rule is_expansion_aux_imp_smallo) fact
+  also have "(\<lambda>x. hd basis x powr 0) \<in> \<Theta>(\<lambda>_. 1)"
+    by (intro bigthetaI_cong eventually_mono[OF b_pos]) (auto simp: *(1))
+  finally have lim: "filterlim (\<lambda>x. f x - eval C x) (nhds 0) at_top"
+    by (auto dest: smalloD_tendsto)
+
+  show "filterlim f at_infinity at_top" if "filterlim (eval C) at_infinity at_top"
+    using tendsto_add_filterlim_at_infinity[OF lim that] by simp
+  show "filterlim f (nhds L) at_top" if "filterlim (eval C) (nhds L) at_top"
+    using tendsto_add[OF lim that] by simp
+qed
+
+lemma expands_to_lift_function:
+  assumes "eventually (\<lambda>x. f x - eval c x = 0) at_top"
+          "((\<lambda>x. g (eval (c :: 'a :: multiseries) x)) expands_to G) bs'"
+  shows   "((\<lambda>x. g (f x)) expands_to G) bs'"
+  by (rule expands_to_cong[OF assms(2)]) (insert assms, auto elim: eventually_mono) 
+
+lemma drop_zero_ms':
+  fixes c f
+  assumes "c = (0::real)" "(f expands_to MS (MSLCons (c, e) xs) g) basis"
+  shows   "(f expands_to MS xs g) basis"
+  using assms drop_zero_ms[of c e xs g basis] by (simp add: expands_to.simps)
+
+lemma trimmed_realI: "c \<noteq> (0::real) \<Longrightarrow> trimmed c"
+  by simp
+
+lemma trimmed_pos_realI: "c > (0::real) \<Longrightarrow> trimmed_pos c"
+  by simp
+
+lemma trimmed_neg_realI: "c < (0::real) \<Longrightarrow> trimmed_neg c"
+  by (simp add: trimmed_neg_def)
+
+lemma trimmed_pos_hd_coeff: "trimmed_pos (MS (MSLCons (C, e) xs) f) \<Longrightarrow> trimmed_pos C"
+  by simp
+
+lemma lift_trimmed: "trimmed C \<Longrightarrow> trimmed (MS (MSLCons (C, e) xs) f)"
+  by (auto simp: trimmed_ms_aux_def)
+
+lemma lift_trimmed_pos: "trimmed_pos C \<Longrightarrow> trimmed_pos (MS (MSLCons (C, e) xs) f)"
+  by simp
+
+lemma lift_trimmed_neg: "trimmed_neg C \<Longrightarrow> trimmed_neg (MS (MSLCons (C, e) xs) f)"
+  by (simp add: trimmed_neg_def fst_dominant_term_ms_aux_MSLCons trimmed_ms_aux_MSLCons)
+
+lemma lift_trimmed_pos':
+  "trimmed_pos C \<Longrightarrow> MS (MSLCons (C, e) xs) f \<equiv> MS (MSLCons (C, e) xs) f \<Longrightarrow>
+     trimmed_pos (MS (MSLCons (C, e) xs) f)"
+  by simp
+
+lemma lift_trimmed_neg':
+  "trimmed_neg C \<Longrightarrow> MS (MSLCons (C, e) xs) f \<equiv> MS (MSLCons (C, e) xs) f \<Longrightarrow>
+     trimmed_neg (MS (MSLCons (C, e) xs) f)"
+  by (simp add: lift_trimmed_neg)
+
+lemma trimmed_eq_cong: "trimmed C \<Longrightarrow> C \<equiv> C' \<Longrightarrow> trimmed C'"
+  and trimmed_pos_eq_cong: "trimmed_pos C \<Longrightarrow> C \<equiv> C' \<Longrightarrow> trimmed_pos C'"
+  and trimmed_neg_eq_cong: "trimmed_neg C \<Longrightarrow> C \<equiv> C' \<Longrightarrow> trimmed_neg C'"
+  by simp_all
+
+lemma trimmed_hd: "trimmed (MS (MSLCons (C, e) xs) f) \<Longrightarrow> trimmed C"
+  by (simp add: trimmed_ms_aux_MSLCons)
+
+lemma trim_lift_eq:
+  assumes "A \<equiv> MS (MSLCons (C, e) xs) f" "C \<equiv> D"
+  shows   "A \<equiv> MS (MSLCons (D, e) xs) f" 
+  using assms by simp
+
+lemma basis_wf_manyI: 
+  "filterlim b' at_top at_top \<Longrightarrow> (\<lambda>x. ln (b x)) \<in> o(\<lambda>x. ln (b' x)) \<Longrightarrow>
+     basis_wf (b # basis) \<Longrightarrow> basis_wf (b' # b # basis)"
+  by (simp_all add: basis_wf_many)
+
+lemma ln_smallo_ln_exp: "(\<lambda>x. ln (b x)) \<in> o(f) \<Longrightarrow> (\<lambda>x. ln (b x)) \<in> o(\<lambda>x. ln (exp (f x :: real)))"
+  by simp
+
+
+subsection \<open>Reification and other technical details\<close>
+
+text \<open>
+  The following is used by the automation in order to avoid writing terms like $x^2$ or $x^{-2}$
+  as @{term "\<lambda>x::real. x powr 2"} etc.\ but as the more agreeable @{term "\<lambda>x::real. x ^ 2"} or
+  @{term "\<lambda>x::real. inverse (x ^ 2)"}.
+\<close>
+
+lemma intyness_0: "0 \<equiv> real 0"
+  and intyness_1: "1 \<equiv> real 1"
+  and intyness_numeral: "num \<equiv> num \<Longrightarrow> numeral num \<equiv> real (numeral num)"
+  and intyness_uminus:  "x \<equiv> real n \<Longrightarrow> -x \<equiv> -real n"
+  and intyness_of_nat:  "n \<equiv> n \<Longrightarrow> real n \<equiv> real n"
+  by simp_all
+    
+lemma intyness_simps:
+  "real a + real b = real (a + b)"
+  "real a * real b = real (a * b)"
+  "real a ^ n = real (a ^ n)"
+  "1 = real 1" "0 = real 0" "numeral num = real (numeral num)" 
+  by simp_all
+
+lemma odd_Numeral1: "odd (Numeral1)"
+  by simp
+
+lemma even_addI:
+  "even a \<Longrightarrow> even b \<Longrightarrow> even (a + b)"
+  "odd a \<Longrightarrow> odd b \<Longrightarrow> even (a + b)"
+  by auto
+
+lemma odd_addI:
+  "even a \<Longrightarrow> odd b \<Longrightarrow> odd (a + b)"
+  "odd a \<Longrightarrow> even b \<Longrightarrow> odd (a + b)"
+  by auto
+
+lemma even_diffI:
+  "even a \<Longrightarrow> even b \<Longrightarrow> even (a - b :: 'a :: ring_parity)"
+  "odd a \<Longrightarrow> odd b \<Longrightarrow> even (a - b)"
+  by auto
+
+lemma odd_diffI:
+  "even a \<Longrightarrow> odd b \<Longrightarrow> odd (a - b :: 'a :: ring_parity)"
+  "odd a \<Longrightarrow> even b \<Longrightarrow> odd (a - b)"
+  by auto
+
+lemma even_multI: "even a \<Longrightarrow> even (a * b)" "even b \<Longrightarrow> even (a * b)"
+  by auto
+
+lemma odd_multI: "odd a \<Longrightarrow> odd b \<Longrightarrow> odd (a * b)"
+  by auto
+
+lemma even_uminusI: "even a \<Longrightarrow> even (-a :: 'a :: ring_parity)"
+  and odd_uminusI:  "odd a \<Longrightarrow> odd (-a :: 'a :: ring_parity)"
+  by auto
+
+lemma odd_powerI: "odd a \<Longrightarrow> odd (a ^ n)"
+  by auto
+
+
+text \<open>
+  The following theorem collection is used to pre-process functions for multiseries expansion.
+\<close>
+named_theorems real_asymp_reify_simps
+
+lemmas [real_asymp_reify_simps] =
+  sinh_field_def cosh_field_def tanh_real_altdef arsinh_def arcosh_def artanh_def
+
+
+text \<open>
+  This is needed in order to handle things like @{term "\<lambda>n. f n ^ n"}.
+\<close>
+definition powr_nat :: "real \<Rightarrow> real \<Rightarrow> real" where 
+  "powr_nat x y = 
+     (if y = 0 then 1
+      else if x < 0 then cos (pi * y) * (-x) powr y else x powr y)"
+  
+lemma powr_nat_of_nat: "powr_nat x (of_nat n) = x ^ n"
+  by (cases "x > 0") (auto simp: powr_nat_def powr_realpow not_less power_mult_distrib [symmetric])
+
+lemma powr_nat_conv_powr: "x > 0 \<Longrightarrow> powr_nat x y = x powr y"
+  by (simp_all add: powr_nat_def)
+
+lemma reify_power: "x ^ n \<equiv> powr_nat x (of_nat n)"
+  by (simp add: powr_nat_of_nat)
+
+
+lemma sqrt_conv_root [real_asymp_reify_simps]: "sqrt x = root 2 x"
+  by (simp add: sqrt_def)
+
+lemma tan_conv_sin_cos [real_asymp_reify_simps]: "tan x = sin x / cos x"
+  by (simp add: tan_def)
+
+definition rfloor :: "real \<Rightarrow> real" where "rfloor x = real_of_int (floor x)"
+definition rceil :: "real \<Rightarrow> real" where "rceil x = real_of_int (ceiling x)"
+definition rnatmod :: "real \<Rightarrow> real \<Rightarrow> real"
+  where "rnatmod x y = real (nat \<lfloor>x\<rfloor> mod nat \<lfloor>y\<rfloor>)"
+definition rintmod :: "real \<Rightarrow> real \<Rightarrow> real"
+  where "rintmod x y = real_of_int (\<lfloor>x\<rfloor> mod \<lfloor>y\<rfloor>)"
+
+lemmas [real_asymp_reify_simps] =
+  ln_exp log_def rfloor_def [symmetric] rceil_def [symmetric]
+
+lemma expands_to_powr_nat_0_0:
+  assumes "eventually (\<lambda>x. f x = 0) at_top" "eventually (\<lambda>x. g x = 0) at_top"
+          "basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)"
+  shows   "((\<lambda>x. powr_nat (f x) (g x)) expands_to (const_expansion 1 :: 'a)) basis"
+proof (rule expands_to_cong [OF expands_to_const])
+  from assms(1,2) show "eventually (\<lambda>x. 1 = powr_nat (f x) (g x)) at_top"
+    by eventually_elim (simp add: powr_nat_def)
+qed fact+
+
+lemma expands_to_powr_nat_0:
+  assumes "eventually (\<lambda>x. f x = 0) at_top" "(g expands_to G) basis" "trimmed G"
+          "basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)"
+  shows   "((\<lambda>x. powr_nat (f x) (g x)) expands_to (zero_expansion :: 'a)) basis"
+proof (rule expands_to_cong [OF expands_to_zero])
+  from assms have "eventually (\<lambda>x. g x \<noteq> 0) at_top"
+    using expands_to_imp_eventually_nz[of basis g G] by auto
+  with assms(1) show "eventually (\<lambda>x. 0 = powr_nat (f x) (g x)) at_top"
+    by eventually_elim (simp add: powr_nat_def)
+qed (insert assms, auto elim!: eventually_mono simp: powr_nat_def)
+
+lemma expands_to_powr_nat:
+  assumes "trimmed_pos F" "basis_wf basis'" "(f expands_to F) basis'"
+  assumes "((\<lambda>x. exp (ln (f x) * g x)) expands_to E) basis"
+  shows   "((\<lambda>x. powr_nat (f x) (g x)) expands_to E) basis"
+using assms(4)
+proof (rule expands_to_cong)
+  from eval_pos_if_dominant_term_pos[of basis' F] assms(1-4)
+    show "eventually (\<lambda>x. exp (ln (f x) * g x) = powr_nat (f x) (g x)) at_top"
+    by (auto simp: expands_to.simps trimmed_pos_def powr_def powr_nat_def elim: eventually_elim2)
+qed
+
+
+subsection \<open>Some technical lemmas\<close>
+
+lemma landau_meta_eq_cong: "f \<in> L(g) \<Longrightarrow> f' \<equiv> f \<Longrightarrow> g' \<equiv> g \<Longrightarrow> f' \<in> L(g')"
+  and asymp_equiv_meta_eq_cong: "f \<sim>[at_top] g \<Longrightarrow> f' \<equiv> f \<Longrightarrow> g' \<equiv> g \<Longrightarrow> f' \<sim>[at_top] g'"
+  by simp_all
+    
+lemma filterlim_mono': "filterlim f F G \<Longrightarrow> F \<le> F' \<Longrightarrow> filterlim f F' G"
+  by (erule (1) filterlim_mono) simp_all
+
+lemma at_within_le_nhds: "at x within A \<le> nhds x"
+  by (simp add: at_within_def)
+
+lemma at_within_le_at: "at x within A \<le> at x"
+  by (rule at_le) simp_all
+
+lemma at_right_to_top': "at_right c = filtermap (\<lambda>x::real. c + inverse x) at_top"
+  by (subst at_right_to_0, subst at_right_to_top) (simp add: filtermap_filtermap add_ac)
+
+lemma at_left_to_top': "at_left c = filtermap (\<lambda>x::real. c - inverse x) at_top"
+  by (subst at_left_minus, subst at_right_to_0, subst at_right_to_top) 
+     (simp add: filtermap_filtermap add_ac)
+
+lemma at_left_to_top: "at_left 0 = filtermap (\<lambda>x::real. - inverse x) at_top"
+  by (simp add: at_left_to_top')
+
+lemma filterlim_conv_filtermap:
+  "G = filtermap g G' \<Longrightarrow>
+     PROP (Trueprop (filterlim f F G)) \<equiv> PROP (Trueprop (filterlim (\<lambda>x. f (g x)) F G'))"
+  by (simp add: filterlim_filtermap)
+
+lemma eventually_conv_filtermap:
+  "G = filtermap g G' \<Longrightarrow> 
+     PROP (Trueprop (eventually P G)) \<equiv> PROP (Trueprop (eventually (\<lambda>x. P (g x)) G'))"
+  by (simp add: eventually_filtermap)
+
+lemma eventually_lt_imp_eventually_le:
+  "eventually (\<lambda>x. f x < (g x :: real)) F \<Longrightarrow> eventually (\<lambda>x. f x \<le> g x) F"
+  by (erule eventually_mono) simp
+    
+lemma smallo_imp_smallomega: "f \<in> o[F](g) \<Longrightarrow> g \<in> \<omega>[F](f)"
+  by (simp add: smallomega_iff_smallo)
+
+lemma bigo_imp_bigomega: "f \<in> O[F](g) \<Longrightarrow> g \<in> \<Omega>[F](f)"
+  by (simp add: bigomega_iff_bigo)
+
+context
+  fixes L L' :: "real filter \<Rightarrow> (real \<Rightarrow> _) \<Rightarrow> _" and Lr :: "real filter \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> _"
+  assumes LS: "landau_symbol L L' Lr"
+begin
+  
+interpretation landau_symbol L L' Lr by (fact LS)
+
+lemma landau_symbol_at_top_imp_at_bot:
+  "(\<lambda>x. f (-x)) \<in> L' at_top (\<lambda>x. g (-x)) \<Longrightarrow> f \<in> L at_bot g"
+  by (simp add: in_filtermap_iff at_bot_mirror)
+
+lemma landau_symbol_at_top_imp_at_right_0:
+  "(\<lambda>x. f (inverse x)) \<in> L' at_top (\<lambda>x. g (inverse x)) \<Longrightarrow> f \<in> L (at_right 0) g"
+  by (simp add: in_filtermap_iff at_right_to_top')
+    
+lemma landau_symbol_at_top_imp_at_left_0:
+  "(\<lambda>x. f ( -inverse x)) \<in> L' at_top (\<lambda>x. g (-inverse x)) \<Longrightarrow> f \<in> L (at_left 0) g"
+  by (simp add: in_filtermap_iff at_left_to_top')    
+
+lemma landau_symbol_at_top_imp_at_right:
+  "(\<lambda>x. f (a + inverse x)) \<in> L' at_top (\<lambda>x. g (a + inverse x)) \<Longrightarrow> f \<in> L (at_right a) g"
+  by (simp add: in_filtermap_iff at_right_to_top')
+    
+lemma landau_symbol_at_top_imp_at_left:
+  "(\<lambda>x. f (a - inverse x)) \<in> L' at_top (\<lambda>x. g (a - inverse x)) \<Longrightarrow> f \<in> L (at_left a) g"
+  by (simp add: in_filtermap_iff at_left_to_top')
+    
+lemma landau_symbol_at_top_imp_at:
+  "(\<lambda>x. f (a - inverse x)) \<in> L' at_top (\<lambda>x. g (a - inverse x)) \<Longrightarrow>
+   (\<lambda>x. f (a + inverse x)) \<in> L' at_top (\<lambda>x. g (a + inverse x)) \<Longrightarrow>
+   f \<in> L (at a) g"
+  by (simp add: sup at_eq_sup_left_right 
+        landau_symbol_at_top_imp_at_left landau_symbol_at_top_imp_at_right)
+      
+lemma landau_symbol_at_top_imp_at_0:
+  "(\<lambda>x. f (-inverse x)) \<in> L' at_top (\<lambda>x. g (-inverse x)) \<Longrightarrow>
+   (\<lambda>x. f (inverse x)) \<in> L' at_top (\<lambda>x. g (inverse x)) \<Longrightarrow>
+   f \<in> L (at 0) g"
+  by (rule landau_symbol_at_top_imp_at) simp_all
+
+end
+
+
+context
+  fixes f g :: "real \<Rightarrow> real"
+begin
+
+lemma asymp_equiv_at_top_imp_at_bot:
+  "(\<lambda>x. f (- x)) \<sim>[at_top] (\<lambda>x. g (-x)) \<Longrightarrow> f \<sim>[at_bot] g"
+  by (simp add: asymp_equiv_def filterlim_at_bot_mirror)
+
+lemma asymp_equiv_at_top_imp_at_right_0:
+  "(\<lambda>x. f (inverse x)) \<sim>[at_top] (\<lambda>x. g (inverse x)) \<Longrightarrow> f \<sim>[at_right 0] g"
+  by (simp add: at_right_to_top asymp_equiv_filtermap_iff)
+
+lemma asymp_equiv_at_top_imp_at_left_0:
+  "(\<lambda>x. f (-inverse x)) \<sim>[at_top] (\<lambda>x. g (-inverse x)) \<Longrightarrow> f \<sim>[at_left 0] g"
+  by (simp add: at_left_to_top asymp_equiv_filtermap_iff)
+
+lemma asymp_equiv_at_top_imp_at_right:
+  "(\<lambda>x. f (a + inverse x)) \<sim>[at_top] (\<lambda>x. g (a + inverse x)) \<Longrightarrow> f \<sim>[at_right a] g"
+  by (simp add: at_right_to_top' asymp_equiv_filtermap_iff)
+
+lemma asymp_equiv_at_top_imp_at_left:
+  "(\<lambda>x. f (a - inverse x)) \<sim>[at_top] (\<lambda>x. g (a - inverse x)) \<Longrightarrow> f \<sim>[at_left a] g"
+  by (simp add: at_left_to_top' asymp_equiv_filtermap_iff)
+
+lemma asymp_equiv_at_top_imp_at:
+  "(\<lambda>x. f (a - inverse x)) \<sim>[at_top] (\<lambda>x. g (a - inverse x)) \<Longrightarrow>
+   (\<lambda>x. f (a + inverse x)) \<sim>[at_top] (\<lambda>x. g (a + inverse x)) \<Longrightarrow> f \<sim>[at a] g"
+  unfolding asymp_equiv_def
+  using asymp_equiv_at_top_imp_at_left[of a] asymp_equiv_at_top_imp_at_right[of a]
+  by (intro filterlim_split_at) (auto simp: asymp_equiv_def)
+
+lemma asymp_equiv_at_top_imp_at_0:
+  "(\<lambda>x. f (-inverse x)) \<sim>[at_top] (\<lambda>x. g (-inverse x)) \<Longrightarrow>
+   (\<lambda>x. f (inverse x)) \<sim>[at_top] (\<lambda>x. g (inverse x)) \<Longrightarrow> f \<sim>[at 0] g"
+  by (rule asymp_equiv_at_top_imp_at) auto
+
+end
+
+
+lemmas eval_simps =
+  eval_const eval_plus eval_minus eval_uminus eval_times eval_inverse eval_divide
+  eval_map_ms eval_lift_ms eval_real_def eval_ms.simps
+
+lemma real_eqI: "a - b = 0 \<Longrightarrow> a = (b::real)"
+  by simp
+
+lemma lift_ms_simps:
+  "lift_ms (MS xs f) = MS (MSLCons (MS xs f, 0) MSLNil) f"
+  "lift_ms (c :: real) = MS (MSLCons (c, 0) MSLNil) (\<lambda>_. c)"
+  by (simp_all add: lift_ms_def)
+
+lemmas landau_reduce_to_top = 
+  landau_symbols [THEN landau_symbol_at_top_imp_at_bot]
+  landau_symbols [THEN landau_symbol_at_top_imp_at_left_0]
+  landau_symbols [THEN landau_symbol_at_top_imp_at_right_0]
+  landau_symbols [THEN landau_symbol_at_top_imp_at_left]
+  landau_symbols [THEN landau_symbol_at_top_imp_at_right]
+  asymp_equiv_at_top_imp_at_bot
+  asymp_equiv_at_top_imp_at_left_0
+  asymp_equiv_at_top_imp_at_right_0
+  asymp_equiv_at_top_imp_at_left
+  asymp_equiv_at_top_imp_at_right
+
+lemmas landau_at_top_imp_at = 
+  landau_symbols [THEN landau_symbol_at_top_imp_at_0]
+  landau_symbols [THEN landau_symbol_at_top_imp_at]
+  asymp_equiv_at_top_imp_at_0
+  asymp_equiv_at_top_imp_at
+
+
+lemma nhds_leI: "x = y \<Longrightarrow> nhds x \<le> nhds y"
+  by simp
+    
+lemma of_nat_diff_real: "of_nat (a - b) = max (0::real) (of_nat a - of_nat b)"
+  and of_nat_max_real: "of_nat (max a b) = max (of_nat a) (of_nat b :: real)"
+  and of_nat_min_real: "of_nat (min a b) = min (of_nat a) (of_nat b :: real)"
+  and of_int_max_real: "of_int (max c d) = max (of_int c) (of_int d :: real)"
+  and of_int_min_real: "of_int (min c d) = min (of_int c) (of_int d :: real)"
+  by simp_all
+
+named_theorems real_asymp_nat_reify
+
+lemmas [real_asymp_nat_reify] = 
+  of_nat_add of_nat_mult of_nat_diff_real of_nat_max_real of_nat_min_real of_nat_power 
+  of_nat_Suc of_nat_numeral
+
+lemma of_nat_div_real [real_asymp_nat_reify]:
+  "of_nat (a div b) = max 0 (rfloor (of_nat a / of_nat b))"
+  by (simp add: rfloor_def floor_divide_of_nat_eq)
+
+lemma of_nat_mod_real [real_asymp_nat_reify]: "of_nat (a mod b) = rnatmod (of_nat a) (of_nat b)"
+  by (simp add: rnatmod_def)
+
+lemma real_nat [real_asymp_nat_reify]: "real (nat a) = max 0 (of_int a)"
+  by simp
+
+
+named_theorems real_asymp_int_reify
+
+lemmas [real_asymp_int_reify] = 
+  of_int_add of_int_mult of_int_diff of_int_minus of_int_max_real of_int_min_real
+  of_int_power of_int_of_nat_eq of_int_numeral
+
+lemma of_int_div_real [real_asymp_int_reify]:
+  "of_int (a div b) = rfloor (of_int a / of_int b)"
+  by (simp add: rfloor_def floor_divide_of_int_eq)
+
+named_theorems real_asymp_preproc
+
+lemmas [real_asymp_preproc] =
+  of_nat_add of_nat_mult of_nat_diff_real of_nat_max_real of_nat_min_real of_nat_power 
+  of_nat_Suc of_nat_numeral
+  of_int_add of_int_mult of_int_diff of_int_minus of_int_max_real of_int_min_real
+  of_int_power of_int_of_nat_eq of_int_numeral real_nat
+
+lemma of_int_mod_real [real_asymp_int_reify]: "of_int (a mod b) = rintmod (of_int a) (of_int b)"
+  by (simp add: rintmod_def)
+
+
+lemma filterlim_of_nat_at_top:
+  "filterlim f F at_top \<Longrightarrow> filterlim (\<lambda>x. f (of_nat x :: real)) F at_top"
+  by (erule filterlim_compose[OF _filterlim_real_sequentially])
+
+lemma asymp_equiv_real_nat_transfer:
+  "f \<sim>[at_top] g \<Longrightarrow> (\<lambda>x. f (of_nat x :: real)) \<sim>[at_top] (\<lambda>x. g (of_nat x))"
+  unfolding asymp_equiv_def by (rule filterlim_of_nat_at_top)
+
+lemma eventually_nat_real:
+  assumes "eventually P (at_top :: real filter)"
+  shows   "eventually (\<lambda>x. P (real x)) (at_top :: nat filter)"
+  using assms filterlim_real_sequentially
+  unfolding filterlim_def le_filter_def eventually_filtermap by auto
+
+lemmas real_asymp_nat_intros =
+  filterlim_of_nat_at_top eventually_nat_real smallo_real_nat_transfer bigo_real_nat_transfer
+  smallomega_real_nat_transfer bigomega_real_nat_transfer bigtheta_real_nat_transfer
+  asymp_equiv_real_nat_transfer
+
+
+lemma filterlim_of_int_at_top:
+  "filterlim f F at_top \<Longrightarrow> filterlim (\<lambda>x. f (of_int x :: real)) F at_top"
+  by (erule filterlim_compose[OF _ filterlim_real_of_int_at_top])
+
+(* TODO Move *)
+lemma filterlim_real_of_int_at_bot [tendsto_intros]:
+  "filterlim real_of_int at_bot at_bot"
+  unfolding filterlim_at_bot
+proof
+  fix C :: real
+  show "eventually (\<lambda>n. real_of_int n \<le> C) at_bot"
+    using eventually_le_at_bot[of "\<lfloor>C\<rfloor>"] by eventually_elim linarith
+qed
+
+lemma filterlim_of_int_at_bot:
+  "filterlim f F at_bot \<Longrightarrow> filterlim (\<lambda>x. f (of_int x :: real)) F at_bot"
+  by (erule filterlim_compose[OF _ filterlim_real_of_int_at_bot])
+find_theorems of_int at_bot
+
+lemma asymp_equiv_real_int_transfer_at_top:
+  "f \<sim>[at_top] g \<Longrightarrow> (\<lambda>x. f (of_int x :: real)) \<sim>[at_top] (\<lambda>x. g (of_int x))"
+  unfolding asymp_equiv_def by (rule filterlim_of_int_at_top)
+
+lemma asymp_equiv_real_int_transfer_at_bot:
+  "f \<sim>[at_bot] g \<Longrightarrow> (\<lambda>x. f (of_int x :: real)) \<sim>[at_bot] (\<lambda>x. g (of_int x))"
+  unfolding asymp_equiv_def by (rule filterlim_of_int_at_bot)
+
+lemma eventually_int_real_at_top:
+  assumes "eventually P (at_top :: real filter)"
+  shows   "eventually (\<lambda>x. P (of_int x)) (at_top :: int filter)"
+  using assms filterlim_of_int_at_top filterlim_iff filterlim_real_of_int_at_top by blast
+
+lemma eventually_int_real_at_bot:
+  assumes "eventually P (at_bot :: real filter)"
+  shows   "eventually (\<lambda>x. P (of_int x)) (at_bot :: int filter)"
+  using assms filterlim_of_int_at_bot filterlim_iff filterlim_real_of_int_at_bot by blast
+
+
+lemmas real_asymp_int_intros =
+  filterlim_of_int_at_bot filterlim_of_int_at_top
+  landau_symbols[THEN landau_symbol.compose[OF _ _ filterlim_real_of_int_at_top]]
+  landau_symbols[THEN landau_symbol.compose[OF _ _ filterlim_real_of_int_at_bot]]
+  asymp_equiv_real_int_transfer_at_top asymp_equiv_real_int_transfer_at_bot
+
+(* TODO: Move? *)
+lemma tendsto_discrete_iff:
+  "filterlim f (nhds (c :: 'a :: discrete_topology)) F \<longleftrightarrow> eventually (\<lambda>x. f x = c) F"
+  by (auto simp: nhds_discrete filterlim_principal)
+
+lemma tendsto_of_nat_iff:
+  "filterlim (\<lambda>n. of_nat (f n)) (nhds (of_nat c :: 'a :: real_normed_div_algebra)) F \<longleftrightarrow>
+     eventually (\<lambda>n. f n = c) F"
+proof
+  assume "((\<lambda>n. of_nat (f n)) \<longlongrightarrow> (of_nat c :: 'a)) F"
+  hence "eventually (\<lambda>n. \<bar>real (f n) - real c\<bar> < 1/2) F"
+    by (force simp: tendsto_iff dist_of_nat dest: spec[of _ "1/2"])
+  thus "eventually (\<lambda>n. f n = c) F"
+    by eventually_elim (simp add: abs_if split: if_splits)
+next
+  assume "eventually (\<lambda>n. f n = c) F"
+  hence "eventually (\<lambda>n. of_nat (f n) = (of_nat c :: 'a)) F"
+    by eventually_elim simp
+  thus "filterlim (\<lambda>n. of_nat (f n)) (nhds (of_nat c :: 'a)) F"
+    by (rule Lim_eventually)
+qed
+
+lemma tendsto_of_int_iff:
+  "filterlim (\<lambda>n. of_int (f n)) (nhds (of_int c :: 'a :: real_normed_div_algebra)) F \<longleftrightarrow>
+     eventually (\<lambda>n. f n = c) F"
+proof
+  assume "((\<lambda>n. of_int (f n)) \<longlongrightarrow> (of_int c :: 'a)) F"
+  hence "eventually (\<lambda>n. \<bar>real_of_int (f n) - of_int c\<bar> < 1/2) F"
+    by (force simp: tendsto_iff dist_of_int dest: spec[of _ "1/2"])
+  thus "eventually (\<lambda>n. f n = c) F"
+    by eventually_elim (simp add: abs_if split: if_splits)
+next
+  assume "eventually (\<lambda>n. f n = c) F"
+  hence "eventually (\<lambda>n. of_int (f n) = (of_int c :: 'a)) F"
+    by eventually_elim simp
+  thus "filterlim (\<lambda>n. of_int (f n)) (nhds (of_int c :: 'a)) F"
+    by (rule Lim_eventually)
+qed
+
+lemma filterlim_at_top_int_iff_filterlim_real:
+  "filterlim f at_top F \<longleftrightarrow> filterlim (\<lambda>x. real_of_int (f x)) at_top F" (is "?lhs = ?rhs")
+proof
+  assume ?rhs
+  hence "filterlim (\<lambda>x. floor (real_of_int (f x))) at_top F"
+    by (intro filterlim_compose[OF filterlim_floor_sequentially])
+  also have "?this \<longleftrightarrow> ?lhs" by (intro filterlim_cong refl) auto
+  finally show ?lhs .
+qed (auto intro: filterlim_compose[OF filterlim_real_of_int_at_top])
+
+lemma filterlim_floor_at_bot: "filterlim (floor :: real \<Rightarrow> int) at_bot at_bot"
+proof (subst filterlim_at_bot, rule allI)
+  fix C :: int show "eventually (\<lambda>x::real. \<lfloor>x\<rfloor> \<le> C) at_bot"
+    using eventually_le_at_bot[of "real_of_int C"] by eventually_elim linarith
+qed
+
+lemma filterlim_at_bot_int_iff_filterlim_real:
+  "filterlim f at_bot F \<longleftrightarrow> filterlim (\<lambda>x. real_of_int (f x)) at_bot F" (is "?lhs = ?rhs")
+proof
+  assume ?rhs
+  hence "filterlim (\<lambda>x. floor (real_of_int (f x))) at_bot F"
+    by (intro filterlim_compose[OF filterlim_floor_at_bot])
+  also have "?this \<longleftrightarrow> ?lhs" by (intro filterlim_cong refl) auto
+  finally show ?lhs .
+qed (auto intro: filterlim_compose[OF filterlim_real_of_int_at_bot])
+(* END TODO *)
+
+
+lemma real_asymp_real_nat_transfer:
+  "filterlim (\<lambda>n. real (f n)) at_top F \<Longrightarrow> filterlim f at_top F"
+  "filterlim (\<lambda>n. real (f n)) (nhds (real c)) F \<Longrightarrow> eventually (\<lambda>n. f n = c) F"
+  "eventually (\<lambda>n. real (f n) \<le> real (g n)) F \<Longrightarrow> eventually (\<lambda>n. f n \<le> g n) F"
+  "eventually (\<lambda>n. real (f n) < real (g n)) F \<Longrightarrow> eventually (\<lambda>n. f n < g n) F"
+  by (simp_all add: filterlim_sequentially_iff_filterlim_real tendsto_of_nat_iff)
+
+lemma real_asymp_real_int_transfer:
+  "filterlim (\<lambda>n. real_of_int (f n)) at_top F \<Longrightarrow> filterlim f at_top F"
+  "filterlim (\<lambda>n. real_of_int (f n)) at_bot F \<Longrightarrow> filterlim f at_bot F"
+  "filterlim (\<lambda>n. real_of_int (f n)) (nhds (of_int c)) F \<Longrightarrow> eventually (\<lambda>n. f n = c) F"
+  "eventually (\<lambda>n. real_of_int (f n) \<le> real_of_int (g n)) F \<Longrightarrow> eventually (\<lambda>n. f n \<le> g n) F"
+  "eventually (\<lambda>n. real_of_int (f n) < real_of_int (g n)) F \<Longrightarrow> eventually (\<lambda>n. f n < g n) F"
+  by (simp_all add: tendsto_of_int_iff filterlim_at_top_int_iff_filterlim_real
+                    filterlim_at_bot_int_iff_filterlim_real)
+  
+
+lemma eventually_at_left_at_right_imp_at:
+  "eventually P (at_left a) \<Longrightarrow> eventually P (at_right a) \<Longrightarrow> eventually P (at (a::real))"
+  by (simp add: eventually_at_split)
+
+lemma filtermap_at_left_shift: "filtermap (\<lambda>x. x - d) (at_left a) = at_left (a - d)"
+  for a d :: "real"
+  by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric])
+    
+lemma at_left_to_0: "at_left a = filtermap (\<lambda>x. x + a) (at_left 0)"
+  for a :: real using filtermap_at_left_shift[of "-a" 0] by simp
+    
+lemma filterlim_at_leftI: 
+  assumes "filterlim (\<lambda>x. f x - c) (at_left 0) F"
+  shows   "filterlim f (at_left (c::real)) F"
+proof -
+  from assms have "filtermap (\<lambda>x. f x - c) F \<le> at_left 0" by (simp add: filterlim_def)
+  hence "filtermap (\<lambda>x. x + c) (filtermap (\<lambda>x. f x - c) F) \<le> filtermap (\<lambda>x. x + c) (at_left 0)"
+    by (rule filtermap_mono)
+  thus ?thesis by (subst at_left_to_0) (simp_all add: filterlim_def filtermap_filtermap)
+qed
+
+lemma filterlim_at_rightI: 
+  assumes "filterlim (\<lambda>x. f x - c) (at_right 0) F"
+  shows   "filterlim f (at_right (c::real)) F"
+proof -
+  from assms have "filtermap (\<lambda>x. f x - c) F \<le> at_right 0" by (simp add: filterlim_def)
+  hence "filtermap (\<lambda>x. x + c) (filtermap (\<lambda>x. f x - c) F) \<le> filtermap (\<lambda>x. x + c) (at_right 0)"
+    by (rule filtermap_mono)
+  thus ?thesis by (subst at_right_to_0) (simp_all add: filterlim_def filtermap_filtermap)
+qed
+
+lemma filterlim_atI':
+  assumes "filterlim (\<lambda>x. f x - c) (at 0) F"
+  shows   "filterlim f (at (c::real)) F"
+proof -
+  from assms have "filtermap (\<lambda>x. f x - c) F \<le> at 0" by (simp add: filterlim_def)
+  hence "filtermap (\<lambda>x. x + c) (filtermap (\<lambda>x. f x - c) F) \<le> filtermap (\<lambda>x. x + c) (at 0)"
+    by (rule filtermap_mono)
+  thus ?thesis by (subst at_to_0) (simp_all add: filterlim_def filtermap_filtermap)
+qed
+
+lemma gbinomial_series_aux_altdef:
+  "gbinomial_series_aux False x n acc =
+     MSLCons acc (gbinomial_series_aux False x (n + 1) ((x - n) / (n + 1) * acc))"
+  "gbinomial_series_aux True x n acc =
+     (if acc = 0 then MSLNil else 
+        MSLCons acc (gbinomial_series_aux True x (n + 1) ((x - n) / (n + 1) * acc)))"
+  by (subst gbinomial_series_aux.code, simp)+
+
+
+
+subsection \<open>Proof method setup\<close>
+
+text \<open>
+  The following theorem collection is the list of rewrite equations to be used by the
+  lazy evaluation package. If something is missing here, normalisation of terms into
+  head normal form will fail.
+\<close>
+
+named_theorems real_asymp_eval_eqs
+
+lemmas [real_asymp_eval_eqs] =
+  msllist.sel fst_conv snd_conv CMP_BRANCH.simps plus_ms.simps minus_ms_altdef 
+  minus_ms_aux_MSLNil_left minus_ms_aux_MSLNil_right minus_ms_aux_MSLCons_eval times_ms.simps
+  uminus_ms.simps divide_ms.simps inverse_ms.simps inverse_ms_aux_MSLCons const_expansion_ms_eval
+  const_expansion_real_def times_ms_aux_MSLNil times_ms_aux_MSLCons_eval scale_shift_ms_aux_MSLCons_eval 
+  scale_shift_ms_aux_MSLNil uminus_ms_aux_MSLCons_eval uminus_ms_aux_MSLNil 
+  scale_ms_aux_MSLNil scale_ms_aux_MSLCons scale_ms_def shift_ms_aux_MSLNil shift_ms_aux_MSLCons
+  scale_ms_aux'_MSLNil scale_ms_aux'_MSLCons exp_series_stream_def exp_series_stream_aux.code
+  plus_ms_aux_MSLNil plus_ms_aux_MSLCons_eval map_ms_altdef map_ms_aux_MSLCons
+  map_ms_aux_MSLNil lift_ms_simps powser_ms_aux_MSLNil powser_ms_aux_MSLCons
+  ln_series_stream_aux_code gbinomial_series_def gbinomial_series_aux_altdef
+  mssalternate.code powr_expansion_ms.simps powr_expansion_real_def powr_ms_aux_MSLCons
+  power_expansion_ms.simps power_expansion_real_def power_ms_aux_MSLCons
+  powser_ms_aux'_MSSCons sin_ms_aux'_altdef cos_ms_aux'_altdef const_ms_aux_def
+  compare_expansions_MSLCons_eval compare_expansions_real zero_expansion_ms_def
+  zero_expansion_real_def sin_series_stream_def cos_series_stream_def arctan_series_stream_def
+  arctan_ms_aux_def sin_series_stream_aux_code arctan_series_stream_aux_code if_True if_False
+                                               
+text \<open>
+  The following constant and theorem collection exist in order to register constructors for
+  lazy evaluation. All constants marked in such a way will be passed to the lazy evaluation
+  package as free constructors upon which pattern-matching can be done.
+\<close>
+definition REAL_ASYMP_EVAL_CONSTRUCTOR :: "'a \<Rightarrow> bool"
+  where [simp]: "REAL_ASYMP_EVAL_CONSTRUCTOR x = True"
+
+named_theorems exp_log_eval_constructor
+
+lemma exp_log_eval_constructors [exp_log_eval_constructor]:
+  "REAL_ASYMP_EVAL_CONSTRUCTOR MSLNil"
+  "REAL_ASYMP_EVAL_CONSTRUCTOR MSLCons"
+  "REAL_ASYMP_EVAL_CONSTRUCTOR MSSCons"
+  "REAL_ASYMP_EVAL_CONSTRUCTOR LT"
+  "REAL_ASYMP_EVAL_CONSTRUCTOR EQ"
+  "REAL_ASYMP_EVAL_CONSTRUCTOR GT"
+  "REAL_ASYMP_EVAL_CONSTRUCTOR Pair"
+  "REAL_ASYMP_EVAL_CONSTRUCTOR True"
+  "REAL_ASYMP_EVAL_CONSTRUCTOR False"
+  "REAL_ASYMP_EVAL_CONSTRUCTOR MS"
+  by simp_all
+
+text \<open>
+  The following constant exists in order to mark functions as having plug-in support
+  for multiseries expansions (i.\,e.\ this can be used to add support for arbitrary functions
+  later on)
+\<close>
+definition REAL_ASYMP_CUSTOM :: "'a \<Rightarrow> bool"
+  where [simp]: "REAL_ASYMP_CUSTOM x = True"
+
+lemmas [simp del] = ms.map inverse_ms_aux.simps divide_ms.simps
+
+definition expansion_with_remainder_term :: "(real \<Rightarrow> real) \<Rightarrow> (real \<Rightarrow> real) set \<Rightarrow> bool" where
+  "expansion_with_remainder_term _ _ = True"
+
+notation (output) expansion_with_remainder_term (infixl "+" 10)
+
+ML_file \<open>asymptotic_basis.ML\<close>
+ML_file \<open>exp_log_expression.ML\<close>
+ML_file \<open>expansion_interface.ML\<close>
+ML_file \<open>multiseries_expansion.ML\<close>
+ML_file \<open>real_asymp.ML\<close>
+
+method_setup real_asymp = 
+  \<open>(Scan.lift (Scan.optional (Args.parens (Args.$$$ "verbose") >> K true) false) --|
+    Method.sections Simplifier.simp_modifiers) >> 
+     (fn verbose => fn ctxt => SIMPLE_METHOD' (Real_Asymp_Basic.tac verbose ctxt))\<close>
+  "Semi-automatic decision procedure for asymptotics of exp-log functions"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real_Asymp/Multiseries_Expansion_Bounds.thy	Sun Jul 15 14:46:57 2018 +0200
@@ -0,0 +1,699 @@
+section \<open>Asymptotic real interval arithmetic\<close>
+(*
+  File:     Multiseries_Expansion_Bounds.thy
+  Author:   Manuel Eberl, TU München
+
+  Automatic computation of upper and lower expansions for real-valued functions.
+  Allows limited handling of functions involving oscillating functions like sin, cos, floor, etc.
+*)
+theory Multiseries_Expansion_Bounds
+imports
+  Multiseries_Expansion
+begin
+
+lemma expands_to_cong_reverse:
+  "eventually (\<lambda>x. f x = g x) at_top \<Longrightarrow> (g expands_to F) bs \<Longrightarrow> (f expands_to F) bs"
+  using expands_to_cong[of g F bs f] by (simp add: eq_commute)
+
+lemma expands_to_trivial_bounds: "(f expands_to F) bs \<Longrightarrow> eventually (\<lambda>x. f x \<in> {f x..f x}) at_top"
+  by simp
+
+lemma eventually_in_atLeastAtMostI:
+  assumes "eventually (\<lambda>x. f x \<ge> l x) at_top" "eventually (\<lambda>x. f x \<le> u x) at_top"
+  shows   "eventually (\<lambda>x. f x \<in> {l x..u x}) at_top"
+  using assms by eventually_elim simp_all
+
+lemma tendsto_sandwich':
+  fixes l f u :: "'a \<Rightarrow> 'b :: order_topology"
+  shows "eventually (\<lambda>x. l x \<le> f x) F \<Longrightarrow> eventually (\<lambda>x. f x \<le> u x) F \<Longrightarrow>
+           (l \<longlongrightarrow> L1) F \<Longrightarrow> (u \<longlongrightarrow> L2) F \<Longrightarrow> L1 = L2 \<Longrightarrow> (f \<longlongrightarrow> L1) F"
+  using tendsto_sandwich[of l f F u L1] by simp
+
+(* TODO: Move? *)
+lemma filterlim_at_bot_mono:
+  fixes l f u :: "'a \<Rightarrow> 'b :: linorder_topology"
+  assumes "filterlim u at_bot F" and "eventually (\<lambda>x. f x \<le> u x) F"
+  shows   "filterlim f at_bot F"
+  unfolding filterlim_at_bot
+proof
+  fix Z :: 'b
+  from assms(1) have "eventually (\<lambda>x. u x \<le> Z) F" by (auto simp: filterlim_at_bot)
+  with assms(2) show "eventually (\<lambda>x. f x \<le> Z) F" by eventually_elim simp
+qed
+
+context
+begin
+
+qualified lemma eq_zero_imp_nonneg: "x = (0::real) \<Longrightarrow> x \<ge> 0"
+  by simp
+
+qualified lemma exact_to_bound: "(f expands_to F) bs \<Longrightarrow> eventually (\<lambda>x. f x \<le> f x) at_top"
+  by simp
+
+qualified lemma expands_to_abs_nonneg: "(f expands_to F) bs \<Longrightarrow> eventually (\<lambda>x. abs (f x) \<ge> 0) at_top"
+  by simp
+
+qualified lemma eventually_nonpos_flip: "eventually (\<lambda>x. f x \<le> (0::real)) F \<Longrightarrow> eventually (\<lambda>x. -f x \<ge> 0) F"
+  by simp
+
+qualified lemma bounds_uminus:
+  fixes a b :: "real \<Rightarrow> real"
+  assumes "eventually (\<lambda>x. a x \<le> b x) at_top"
+  shows   "eventually (\<lambda>x. -b x \<le> -a x) at_top"
+  using assms by eventually_elim simp
+
+qualified lemma
+  fixes a b c d :: "real \<Rightarrow> real"
+  assumes "eventually (\<lambda>x. a x \<le> b x) at_top" "eventually (\<lambda>x. c x \<le> d x) at_top"
+  shows   combine_bounds_add:  "eventually (\<lambda>x. a x + c x \<le> b x + d x) at_top"
+    and   combine_bounds_diff: "eventually (\<lambda>x. a x - d x \<le> b x - c x) at_top"
+  by (use assms in eventually_elim; simp add: add_mono diff_mono)+
+
+qualified lemma
+  fixes a b c d :: "real \<Rightarrow> real"
+  assumes "eventually (\<lambda>x. a x \<le> b x) at_top" "eventually (\<lambda>x. c x \<le> d x) at_top"
+  shows   combine_bounds_min: "eventually (\<lambda>x. min (a x) (c x) \<le> min (b x) (d x)) at_top"
+    and   combine_bounds_max: "eventually (\<lambda>x. max (a x) (c x) \<le> max (b x) (d x)) at_top"
+  by (blast intro: eventually_elim2[OF assms] min.mono max.mono)+
+
+
+qualified lemma trivial_bounds_sin:  "\<forall>x::real. sin x \<in> {-1..1}"
+  and trivial_bounds_cos:  "\<forall>x::real. cos x \<in> {-1..1}"
+  and trivial_bounds_frac: "\<forall>x::real. frac x \<in> {0..1}"
+  by (auto simp: less_imp_le[OF frac_lt_1])
+
+qualified lemma trivial_boundsI:
+  fixes f g:: "real \<Rightarrow> real"
+  assumes "\<forall>x. f x \<in> {l..u}" and "g \<equiv> g"
+  shows   "eventually (\<lambda>x. f (g x) \<ge> l) at_top" "eventually (\<lambda>x. f (g x) \<le> u) at_top"
+  using assms by auto
+
+qualified lemma
+  fixes f f' :: "real \<Rightarrow> real"
+  shows transfer_lower_bound:
+          "eventually (\<lambda>x. g x \<ge> l x) at_top \<Longrightarrow> f \<equiv> g \<Longrightarrow> eventually (\<lambda>x. f x \<ge> l x) at_top"
+  and   transfer_upper_bound:
+          "eventually (\<lambda>x. g x \<le> u x) at_top \<Longrightarrow> f \<equiv> g \<Longrightarrow> eventually (\<lambda>x. f x \<le> u x) at_top"
+  by simp_all  
+
+qualified lemma mono_bound:
+  fixes f g h :: "real \<Rightarrow> real"
+  assumes "mono h" "eventually (\<lambda>x. f x \<le> g x) at_top"
+  shows   "eventually (\<lambda>x. h (f x) \<le> h (g x)) at_top"
+  by (intro eventually_mono[OF assms(2)] monoD[OF assms(1)])
+
+qualified lemma
+  fixes f l :: "real \<Rightarrow> real"
+  assumes "(l expands_to L) bs" "trimmed_pos L" "basis_wf bs" "eventually (\<lambda>x. l x \<le> f x) at_top"
+  shows   expands_to_lb_ln: "eventually (\<lambda>x. ln (l x) \<le> ln (f x)) at_top"
+    and   expands_to_ub_ln: 
+            "eventually (\<lambda>x. f x \<le> u x) at_top \<Longrightarrow> eventually (\<lambda>x. ln (f x) \<le> ln (u x)) at_top"
+proof -
+  from assms(3,1,2) have pos: "eventually (\<lambda>x. l x > 0) at_top"
+    by (rule expands_to_imp_eventually_pos)  
+  from pos and assms(4)
+    show "eventually (\<lambda>x. ln (l x) \<le> ln (f x)) at_top" by eventually_elim simp
+  assume "eventually (\<lambda>x. f x \<le> u x) at_top"
+  with pos and assms(4) show "eventually (\<lambda>x. ln (f x) \<le> ln (u x)) at_top"
+    by eventually_elim simp
+qed
+
+qualified lemma eventually_sgn_ge_1D:
+  assumes "eventually (\<lambda>x::real. sgn (f x) \<ge> l x) at_top"
+          "(l expands_to (const_expansion 1 :: 'a :: multiseries)) bs"
+  shows   "((\<lambda>x. sgn (f x)) expands_to (const_expansion 1 :: 'a)) bs"
+proof (rule expands_to_cong)
+  from assms(2) have "eventually (\<lambda>x. l x = 1) at_top"
+    by (simp add: expands_to.simps)
+  with assms(1) show "eventually (\<lambda>x. 1 = sgn (f x)) at_top"
+    by eventually_elim (auto simp: sgn_if split: if_splits)
+qed (insert assms, auto simp: expands_to.simps)
+
+qualified lemma eventually_sgn_le_neg1D:
+  assumes "eventually (\<lambda>x::real. sgn (f x) \<le> u x) at_top"
+          "(u expands_to (const_expansion (-1) :: 'a :: multiseries)) bs"
+  shows   "((\<lambda>x. sgn (f x)) expands_to (const_expansion (-1) :: 'a)) bs"
+proof (rule expands_to_cong)
+  from assms(2) have "eventually (\<lambda>x. u x = -1) at_top"
+    by (simp add: expands_to.simps eq_commute)
+  with assms(1) show "eventually (\<lambda>x. -1 = sgn (f x)) at_top"
+    by eventually_elim (auto simp: sgn_if split: if_splits)
+qed (insert assms, auto simp: expands_to.simps)
+
+
+qualified lemma expands_to_squeeze:
+  assumes "eventually (\<lambda>x. l x \<le> f x) at_top" "eventually (\<lambda>x. f x \<le> g x) at_top"
+          "(l expands_to L) bs" "(g expands_to L) bs"
+  shows   "(f expands_to L) bs"
+proof (rule expands_to_cong[OF assms(3)])
+  from assms have "eventually (\<lambda>x. eval L x = l x) at_top" "eventually (\<lambda>x. eval L x = g x) at_top"
+    by (simp_all add: expands_to.simps)
+  with assms(1,2) show "eventually (\<lambda>x. l x = f x) at_top"
+    by eventually_elim simp
+qed 
+
+
+qualified lemma mono_exp_real: "mono (exp :: real \<Rightarrow> real)"
+  by (auto intro!: monoI)
+
+qualified lemma mono_sgn_real: "mono (sgn :: real \<Rightarrow> real)"
+  by (auto intro!: monoI simp: sgn_if)
+
+qualified lemma mono_arctan_real: "mono (arctan :: real \<Rightarrow> real)"
+  by (auto intro!: monoI arctan_monotone')
+
+qualified lemma mono_root_real: "n \<equiv> n \<Longrightarrow> mono (root n :: real \<Rightarrow> real)"
+  by (cases n) (auto simp: mono_def)
+
+qualified lemma mono_rfloor: "mono rfloor" and mono_rceil: "mono rceil"
+  by (auto intro!: monoI simp: rfloor_def floor_mono rceil_def ceiling_mono)
+
+qualified lemma lower_bound_cong:
+  "eventually (\<lambda>x. f x = g x) at_top \<Longrightarrow> eventually (\<lambda>x. l x \<le> g x) at_top \<Longrightarrow>
+     eventually (\<lambda>x. l x \<le> f x) at_top"
+  by (erule (1) eventually_elim2) simp
+
+qualified lemma upper_bound_cong:
+  "eventually (\<lambda>x. f x = g x) at_top \<Longrightarrow> eventually (\<lambda>x. g x \<le> u x) at_top \<Longrightarrow>
+     eventually (\<lambda>x. f x \<le> u x) at_top"
+  by (erule (1) eventually_elim2) simp
+
+qualified lemma
+  assumes "eventually (\<lambda>x. f x = (g x :: real)) at_top"
+  shows   eventually_eq_min: "eventually (\<lambda>x. min (f x) (g x) = f x) at_top"
+    and   eventually_eq_max: "eventually (\<lambda>x. max (f x) (g x) = f x) at_top"
+  by (rule eventually_mono[OF assms]; simp)+
+
+qualified lemma rfloor_bound:
+    "eventually (\<lambda>x. l x \<le> f x) at_top \<Longrightarrow> eventually (\<lambda>x. l x - 1 \<le> rfloor (f x)) at_top"
+    "eventually (\<lambda>x. f x \<le> u x) at_top \<Longrightarrow> eventually (\<lambda>x. rfloor (f x) \<le> u x) at_top"
+  and rceil_bound:
+    "eventually (\<lambda>x. l x \<le> f x) at_top \<Longrightarrow> eventually (\<lambda>x. l x \<le> rceil (f x)) at_top"
+    "eventually (\<lambda>x. f x \<le> u x) at_top \<Longrightarrow> eventually (\<lambda>x. rceil (f x) \<le> u x + 1) at_top"
+  unfolding rfloor_def rceil_def by (erule eventually_mono, linarith)+
+
+qualified lemma natmod_trivial_lower_bound:
+  fixes f g :: "real \<Rightarrow> real"
+  assumes "f \<equiv> f" "g \<equiv> g"
+  shows "eventually (\<lambda>x. rnatmod (f x) (g x) \<ge> 0) at_top"
+  by (simp add: rnatmod_def)
+
+qualified lemma natmod_upper_bound:
+  fixes f g :: "real \<Rightarrow> real"
+  assumes "f \<equiv> f" and "eventually (\<lambda>x. l2 x \<le> g x) at_top" and "eventually (\<lambda>x. g x \<le> u2 x) at_top"
+  assumes "eventually (\<lambda>x. l2 x - 1 \<ge> 0) at_top"
+  shows "eventually (\<lambda>x. rnatmod (f x) (g x) \<le> u2 x - 1) at_top"
+  using assms(2-)
+proof eventually_elim
+  case (elim x)
+  have "rnatmod (f x) (g x) = real (nat \<lfloor>f x\<rfloor> mod nat \<lfloor>g x\<rfloor>)"
+    by (simp add: rnatmod_def)
+  also have "nat \<lfloor>f x\<rfloor> mod nat \<lfloor>g x\<rfloor> < nat \<lfloor>g x\<rfloor>"
+    using elim by (intro mod_less_divisor) auto
+  hence "real (nat \<lfloor>f x\<rfloor> mod nat \<lfloor>g x\<rfloor>) \<le> u2 x - 1"
+    using elim by linarith
+  finally show ?case .
+qed
+
+qualified lemma natmod_upper_bound':
+  fixes f g :: "real \<Rightarrow> real"
+  assumes "g \<equiv> g" "eventually (\<lambda>x. u1 x \<ge> 0) at_top" and "eventually (\<lambda>x. f x \<le> u1 x) at_top"
+  shows "eventually (\<lambda>x. rnatmod (f x) (g x) \<le> u1 x) at_top"
+  using assms(2-)
+proof eventually_elim
+  case (elim x)
+  have "rnatmod (f x) (g x) = real (nat \<lfloor>f x\<rfloor> mod nat \<lfloor>g x\<rfloor>)"
+    by (simp add: rnatmod_def)
+  also have "nat \<lfloor>f x\<rfloor> mod nat \<lfloor>g x\<rfloor> \<le> nat \<lfloor>f x\<rfloor>"
+    by auto
+  hence "real (nat \<lfloor>f x\<rfloor> mod nat \<lfloor>g x\<rfloor>) \<le> real (nat \<lfloor>f x\<rfloor>)"
+    by simp
+  also have "\<dots> \<le> u1 x" using elim by linarith
+  finally show "rnatmod (f x) (g x) \<le> \<dots>" .
+qed
+
+qualified lemma expands_to_natmod_nonpos:
+  fixes f g :: "real \<Rightarrow> real"
+  assumes "g \<equiv> g" "eventually (\<lambda>x. u1 x \<le> 0) at_top" "eventually (\<lambda>x. f x \<le> u1 x) at_top"
+          "((\<lambda>_. 0) expands_to C) bs"
+  shows "((\<lambda>x. rnatmod (f x) (g x)) expands_to C) bs"
+  by (rule expands_to_cong[OF assms(4)])
+     (insert assms, auto elim: eventually_elim2 simp: rnatmod_def)
+  
+
+qualified lemma eventually_atLeastAtMostI:
+  fixes f l r :: "real \<Rightarrow> real"
+  assumes "eventually (\<lambda>x. l x \<le> f x) at_top" "eventually (\<lambda>x. f x \<le> u x) at_top"
+  shows   "eventually (\<lambda>x. f x \<in> {l x..u x}) at_top"
+  using assms by eventually_elim simp
+
+qualified lemma eventually_atLeastAtMostD:
+  fixes f l r :: "real \<Rightarrow> real"
+  assumes "eventually (\<lambda>x. f x \<in> {l x..u x}) at_top"
+  shows   "eventually (\<lambda>x. l x \<le> f x) at_top" "eventually (\<lambda>x. f x \<le> u x) at_top" 
+  using assms by (simp_all add: eventually_conj_iff)
+
+qualified lemma eventually_0_imp_prod_zero:
+  fixes f g :: "real \<Rightarrow> real"
+  assumes "eventually (\<lambda>x. f x = 0) at_top"
+  shows   "eventually (\<lambda>x. f x * g x = 0) at_top" "eventually (\<lambda>x. g x * f x = 0) at_top"
+  by (use assms in eventually_elim; simp)+
+
+qualified lemma mult_right_bounds:
+  fixes f g :: "real \<Rightarrow> real"
+  shows "\<forall>x. f x \<in> {l x..u x} \<longrightarrow> g x \<ge> 0 \<longrightarrow> f x * g x \<in> {l x * g x..u x * g x}"
+    and "\<forall>x. f x \<in> {l x..u x} \<longrightarrow> g x \<le> 0 \<longrightarrow> f x * g x \<in> {u x * g x..l x * g x}"
+  by (auto intro: mult_right_mono mult_right_mono_neg)
+
+qualified lemma mult_left_bounds:
+  fixes f g :: "real \<Rightarrow> real"
+  shows "\<forall>x. g x \<in> {l x..u x} \<longrightarrow> f x \<ge> 0 \<longrightarrow> f x * g x \<in> {f x * l x..f x * u x}"
+    and "\<forall>x. g x \<in> {l x..u x} \<longrightarrow> f x \<le> 0 \<longrightarrow> f x * g x \<in> {f x * u x..f x * l x}"
+  by (auto intro: mult_left_mono mult_left_mono_neg)
+
+qualified lemma mult_mono_nonpos_nonneg:
+  "a \<le> c \<Longrightarrow> d \<le> b \<Longrightarrow> a \<le> 0 \<Longrightarrow> d \<ge> 0 \<Longrightarrow> a * b \<le> c * (d :: 'a :: linordered_ring)"
+  using mult_mono[of "-c" "-a" d b] by simp
+
+qualified lemma mult_mono_nonneg_nonpos:
+  "c \<le> a \<Longrightarrow> b \<le> d \<Longrightarrow> a \<ge> 0 \<Longrightarrow> d \<le> 0 \<Longrightarrow> a * b \<le> c * (d :: 'a :: {comm_ring, linordered_ring})"
+  using mult_mono[of c a "-d" "-b"] by (simp add: mult.commute)
+
+qualified lemma mult_mono_nonpos_nonpos:
+  "c \<le> a \<Longrightarrow> d \<le> b \<Longrightarrow> c \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> a * b \<le> c * (d :: 'a :: linordered_ring)"
+  using mult_mono[of "-a" "-c" "-b" "-d"] by simp
+
+qualified lemmas mult_monos =
+  mult_mono mult_mono_nonpos_nonneg mult_mono_nonneg_nonpos mult_mono_nonpos_nonpos
+
+
+qualified lemma mult_bounds_real:
+  fixes f g l1 u1 l2 u2 l u :: "real \<Rightarrow> real"
+  defines "P \<equiv> \<lambda>l u x. f x \<in> {l1 x..u1 x} \<longrightarrow> g x \<in> {l2 x..u2 x} \<longrightarrow> f x * g x \<in> {l..u}"
+  shows   "\<forall>x. l1 x \<ge> 0 \<longrightarrow> l2 x \<ge> 0 \<longrightarrow> P (l1 x * l2 x) (u1 x * u2 x) x"
+    and   "\<forall>x. u1 x \<le> 0 \<longrightarrow> l2 x \<ge> 0 \<longrightarrow> P (l1 x * u2 x) (u1 x * l2 x) x"
+    and   "\<forall>x. l1 x \<ge> 0 \<longrightarrow> u2 x \<le> 0 \<longrightarrow> P (u1 x * l2 x) (l1 x * u2 x) x"
+    and   "\<forall>x. u1 x \<le> 0 \<longrightarrow> u2 x \<le> 0 \<longrightarrow> P (u1 x * u2 x) (l1 x * l2 x) x"
+
+    and   "\<forall>x. l1 x \<le> 0 \<longrightarrow> u1 x \<ge> 0 \<longrightarrow> l2 x \<ge> 0 \<longrightarrow> P (l1 x * u2 x) (u1 x * u2 x) x"
+    and   "\<forall>x. l1 x \<le> 0 \<longrightarrow> u1 x \<ge> 0 \<longrightarrow> u2 x \<le> 0 \<longrightarrow> P (u1 x * l2 x) (l1 x * l2 x) x"
+    and   "\<forall>x. l1 x \<ge> 0 \<longrightarrow> l2 x \<le> 0 \<longrightarrow> u2 x \<ge> 0 \<longrightarrow> P (u1 x * l2 x) (u1 x * u2 x) x"
+    and   "\<forall>x. u1 x \<le> 0 \<longrightarrow> l2 x \<le> 0 \<longrightarrow> u2 x \<ge> 0 \<longrightarrow> P (l1 x * u2 x) (l1 x * l2 x) x"
+
+    and   "\<forall>x. l1 x \<le> 0 \<longrightarrow> u1 x \<ge> 0 \<longrightarrow> l2 x \<le> 0 \<longrightarrow> u2 x \<ge> 0 \<longrightarrow> l x \<le> l1 x * u2 x \<longrightarrow>
+             l x \<le> u1 x * l2 x \<longrightarrow> u x \<ge> l1 x* l2 x \<longrightarrow> u x \<ge> u1 x * u2 x \<longrightarrow> P (l x) (u x) x"
+proof goal_cases
+  case 1
+  show ?case by (auto intro: mult_mono simp: P_def)
+next
+  case 2
+  show ?case by (auto intro: mult_monos(2) simp: P_def)
+next
+  case 3
+  show ?case unfolding P_def
+    by (subst (1 2 3) mult.commute) (auto intro: mult_monos(2) simp: P_def)
+next
+  case 4
+  show ?case by (auto intro: mult_monos(4) simp: P_def)
+next
+  case 5
+  show ?case by (auto intro: mult_monos(1,2) simp: P_def)
+next
+  case 6
+  show ?case by (auto intro: mult_monos(3,4) simp: P_def)
+next
+  case 7
+  show ?case unfolding P_def
+    by (subst (1 2 3) mult.commute) (auto intro: mult_monos(1,2))
+next
+  case 8
+  show ?case unfolding P_def
+    by (subst (1 2 3) mult.commute) (auto intro: mult_monos(3,4))
+next
+  case 9
+  show ?case
+  proof (safe, goal_cases)
+    case (1 x)
+    from 1(1-4) show ?case unfolding P_def
+      by (cases "f x \<ge> 0"; cases "g x \<ge> 0";
+          fastforce intro: mult_monos 1(5,6)[THEN order.trans] 1(7,8)[THEN order.trans[rotated]])
+  qed
+qed
+
+
+qualified lemma powr_bounds_real:
+  fixes f g l1 u1 l2 u2 l u :: "real \<Rightarrow> real"
+  defines "P \<equiv> \<lambda>l u x. f x \<in> {l1 x..u1 x} \<longrightarrow> g x \<in> {l2 x..u2 x} \<longrightarrow> f x powr g x \<in> {l..u}"
+  shows   "\<forall>x. l1 x \<ge> 0 \<longrightarrow> l1 x \<ge> 1 \<longrightarrow> l2 x \<ge> 0 \<longrightarrow> P (l1 x powr l2 x) (u1 x powr u2 x) x"
+    and   "\<forall>x. l1 x \<ge> 0 \<longrightarrow> u1 x \<le> 1 \<longrightarrow> l2 x \<ge> 0 \<longrightarrow> P (l1 x powr u2 x) (u1 x powr l2 x) x"
+    and   "\<forall>x. l1 x \<ge> 0 \<longrightarrow> l1 x \<ge> 1 \<longrightarrow> u2 x \<le> 0 \<longrightarrow> P (u1 x powr l2 x) (l1 x powr u2 x) x"
+    and   "\<forall>x. l1 x > 0 \<longrightarrow> u1 x \<le> 1 \<longrightarrow> u2 x \<le> 0 \<longrightarrow> P (u1 x powr u2 x) (l1 x powr l2 x) x"
+
+    and   "\<forall>x. l1 x \<ge> 0 \<longrightarrow> l1 x \<le> 1 \<longrightarrow> u1 x \<ge> 1 \<longrightarrow> l2 x \<ge> 0 \<longrightarrow> P (l1 x powr u2 x) (u1 x powr u2 x) x"
+    and   "\<forall>x. l1 x > 0 \<longrightarrow> l1 x \<le> 1 \<longrightarrow> u1 x \<ge> 1 \<longrightarrow> u2 x \<le> 0 \<longrightarrow> P (u1 x powr l2 x) (l1 x powr l2 x) x"
+    and   "\<forall>x. l1 x \<ge> 0 \<longrightarrow> l1 x \<ge> 1 \<longrightarrow> l2 x \<le> 0 \<longrightarrow> u2 x \<ge> 0 \<longrightarrow> P (u1 x powr l2 x) (u1 x powr u2 x) x"
+    and   "\<forall>x. l1 x > 0 \<longrightarrow> u1 x \<le> 1 \<longrightarrow> l2 x \<le> 0 \<longrightarrow> u2 x \<ge> 0 \<longrightarrow> P (l1 x powr u2 x) (l1 x powr l2 x) x"
+
+    and   "\<forall>x. l1 x > 0 \<longrightarrow> l1 x \<le> 1 \<longrightarrow> u1 x \<ge> 1 \<longrightarrow> l2 x \<le> 0 \<longrightarrow> u2 x \<ge> 0 \<longrightarrow> l x \<le> l1 x powr u2 x \<longrightarrow>
+             l x \<le> u1 x powr l2 x \<longrightarrow> u x \<ge> l1 x powr l2 x \<longrightarrow> u x \<ge> u1 x powr u2 x \<longrightarrow> P (l x) (u x) x"
+proof goal_cases
+  case 1
+  show ?case by (auto simp: P_def powr_def intro: mult_monos)
+next
+  case 2
+  show ?case by (auto simp: P_def powr_def intro: mult_monos)
+next
+  case 3
+  show ?case by (auto simp: P_def powr_def intro: mult_monos)
+next
+  case 4
+  show ?case by (auto simp: P_def powr_def intro: mult_monos)
+next
+  case 5
+  note comm = mult.commute[of _ "ln x" for x :: real]
+  show ?case by (auto simp: P_def powr_def comm intro: mult_monos)
+next
+  case 6
+  note comm = mult.commute[of _ "ln x" for x :: real]
+  show ?case by (auto simp: P_def powr_def comm intro: mult_monos)
+next
+  case 7
+  show ?case by (auto simp: P_def powr_def intro: mult_monos)
+next
+  case 8  
+  show ?case 
+    by (auto simp: P_def powr_def intro: mult_monos)
+next
+  case 9
+  show ?case unfolding P_def
+  proof (safe, goal_cases)
+    case (1 x)
+    define l' where "l' = (\<lambda>x. min (ln (l1 x) * u2 x) (ln (u1 x) * l2 x))"
+    define u' where "u' = (\<lambda>x. max (ln (l1 x) * l2 x) (ln (u1 x) * u2 x))"
+    from 1 spec[OF mult_bounds_real(9)[of "\<lambda>x. ln (l1 x)" "\<lambda>x. ln (u1 x)" l2 u2 l' u' 
+                                          "\<lambda>x. ln (f x)" g], of x]
+      have "ln (f x) * g x \<in> {l' x..u' x}" by (auto simp: powr_def mult.commute l'_def u'_def)
+    with 1 have "f x powr g x \<in> {exp (l' x)..exp (u' x)}"
+      by (auto simp: powr_def mult.commute)
+    also from 1 have "exp (l' x) = min (l1 x powr u2 x) (u1 x powr l2 x)"
+      by (auto simp add: l'_def powr_def min_def mult.commute)
+    also from 1 have "exp (u' x) = max (l1 x powr l2 x) (u1 x powr u2 x)"
+      by (auto simp add: u'_def powr_def max_def mult.commute)
+    finally show ?case using 1(5-9) by auto
+  qed
+qed
+
+qualified lemma powr_right_bounds:
+  fixes f g :: "real \<Rightarrow> real"
+  shows "\<forall>x. l x \<ge> 0 \<longrightarrow> f x \<in> {l x..u x} \<longrightarrow> g x \<ge> 0 \<longrightarrow> f x powr g x \<in> {l x powr g x..u x powr g x}"
+    and "\<forall>x. l x > 0 \<longrightarrow> f x \<in> {l x..u x} \<longrightarrow> g x \<le> 0 \<longrightarrow> f x powr g x \<in> {u x powr g x..l x powr g x}"
+  by (auto intro: powr_mono2 powr_mono2')
+
+(* TODO Move *)
+lemma powr_mono': "a \<le> (b::real) \<Longrightarrow> x \<ge> 0 \<Longrightarrow> x \<le> 1 \<Longrightarrow> x powr b \<le> x powr a"
+  using powr_mono[of "-b" "-a" "inverse x"] by (auto simp: powr_def ln_inverse ln_div divide_simps)
+
+qualified lemma powr_left_bounds:
+  fixes f g :: "real \<Rightarrow> real"
+  shows "\<forall>x. f x > 0 \<longrightarrow> g x \<in> {l x..u x} \<longrightarrow> f x \<ge> 1 \<longrightarrow> f x powr g x \<in> {f x powr l x..f x powr u x}"
+    and "\<forall>x. f x > 0 \<longrightarrow> g x \<in> {l x..u x} \<longrightarrow> f x \<le> 1 \<longrightarrow> f x powr g x \<in> {f x powr u x..f x powr l x}"
+  by (auto intro: powr_mono powr_mono')
+
+qualified lemma powr_nat_bounds_transfer_ge:
+  "\<forall>x. l1 x \<ge> 0 \<longrightarrow> f x \<ge> l1 x \<longrightarrow> f x powr g x \<ge> l x \<longrightarrow> powr_nat (f x) (g x) \<ge> l x"
+  by (auto simp: powr_nat_def)
+
+qualified lemma powr_nat_bounds_transfer_le:
+  "\<forall>x. l1 x > 0 \<longrightarrow> f x \<ge> l1 x \<longrightarrow> f x powr g x \<le> u x \<longrightarrow> powr_nat (f x) (g x) \<le> u x"
+  "\<forall>x. l1 x \<ge> 0 \<longrightarrow> l2 x > 0 \<longrightarrow> f x \<ge> l1 x \<longrightarrow> g x \<ge> l2 x \<longrightarrow> 
+      f x powr g x \<le> u x \<longrightarrow> powr_nat (f x) (g x) \<le> u x"
+  "\<forall>x. l1 x \<ge> 0 \<longrightarrow> u2 x < 0 \<longrightarrow> f x \<ge> l1 x \<longrightarrow> g x \<le> u2 x \<longrightarrow>
+      f x powr g x \<le> u x \<longrightarrow> powr_nat (f x) (g x) \<le> u x"
+  "\<forall>x. l1 x \<ge> 0 \<longrightarrow> f x \<ge> l1 x \<longrightarrow>  f x powr g x \<le> u x \<longrightarrow> u x \<le> u' x \<longrightarrow> 1 \<le> u' x \<longrightarrow> 
+     powr_nat (f x) (g x) \<le> u' x"
+  by (auto simp: powr_nat_def)
+
+lemma abs_powr_nat_le: "abs (powr_nat x y) \<le> powr_nat (abs x) y"
+  by (simp add: powr_nat_def abs_mult)
+
+qualified lemma powr_nat_bounds_ge_neg:
+  assumes "powr_nat (abs x) y \<le> u"
+  shows   "powr_nat x y \<ge> -u" "powr_nat x y \<le> u"
+proof -
+  have "abs (powr_nat x y) \<le> powr_nat (abs x) y"
+    by (rule abs_powr_nat_le)
+  also have "\<dots> \<le> u" using assms by auto
+  finally show "powr_nat x y \<ge> -u" "powr_nat x y \<le> u" by auto
+qed
+
+qualified lemma powr_nat_bounds_transfer_abs [eventuallized]:
+  "\<forall>x. powr_nat (abs (f x)) (g x) \<le> u x \<longrightarrow> powr_nat (f x) (g x) \<ge> -u x"
+  "\<forall>x. powr_nat (abs (f x)) (g x) \<le> u x \<longrightarrow> powr_nat (f x) (g x) \<le> u x"
+  using powr_nat_bounds_ge_neg by blast+
+
+qualified lemma eventually_powr_const_nonneg:
+  "f \<equiv> f \<Longrightarrow> p \<equiv> p \<Longrightarrow> eventually (\<lambda>x. f x powr p \<ge> (0::real)) at_top"
+  by simp
+
+qualified lemma eventually_powr_const_mono_nonneg:
+  assumes "p \<ge> (0 :: real)" "eventually (\<lambda>x. l x \<ge> 0) at_top" "eventually (\<lambda>x. l x \<le> f x) at_top"
+          "eventually (\<lambda>x. f x \<le> g x) at_top"
+  shows   "eventually (\<lambda>x. f x powr p \<le> g x powr p) at_top"
+  using assms(2-4) by eventually_elim (auto simp: assms(1) intro!: powr_mono2)
+
+qualified lemma eventually_powr_const_mono_nonpos:
+  assumes "p \<le> (0 :: real)" "eventually (\<lambda>x. l x > 0) at_top" "eventually (\<lambda>x. l x \<le> f x) at_top"
+          "eventually (\<lambda>x. f x \<le> g x) at_top"
+  shows   "eventually (\<lambda>x. f x powr p \<ge> g x powr p) at_top"
+  using assms(2-4) by eventually_elim (auto simp: assms(1) intro!: powr_mono2')
+
+
+qualified lemma eventually_power_mono:
+  assumes "eventually (\<lambda>x. 0 \<le> l x) at_top" "eventually (\<lambda>x. l x \<le> f x) at_top"
+          "eventually (\<lambda>x. f x \<le> g x) at_top" "n \<equiv> n"
+  shows   "eventually (\<lambda>x. f x ^ n \<le> (g x :: real) ^ n) at_top"
+  using assms(1-3) by eventually_elim (auto intro: power_mono)
+
+qualified lemma eventually_mono_power_odd:
+  assumes "odd n" "eventually (\<lambda>x. f x \<le> (g x :: real)) at_top"
+  shows   "eventually (\<lambda>x. f x ^ n \<le> g x ^ n) at_top"
+  using assms(2) by eventually_elim (insert assms(1), auto intro: power_mono_odd)
+
+
+qualified lemma eventually_lower_bound_power_even_nonpos:
+  assumes "even n" "eventually (\<lambda>x. u x \<le> (0::real)) at_top"
+          "eventually (\<lambda>x. f x \<le> u x) at_top"
+  shows   "eventually (\<lambda>x. u x ^ n \<le> f x ^ n) at_top"
+  using assms(2-) by eventually_elim (rule power_mono_even, auto simp: assms(1))
+
+qualified lemma eventually_upper_bound_power_even_nonpos:
+  assumes "even n" "eventually (\<lambda>x. u x \<le> (0::real)) at_top"
+          "eventually (\<lambda>x. l x \<le> f x) at_top" "eventually (\<lambda>x. f x \<le> u x) at_top"
+  shows   "eventually (\<lambda>x. f x ^ n \<le> l x ^ n) at_top"
+  using assms(2-) by eventually_elim (rule power_mono_even, auto simp: assms(1))
+
+qualified lemma eventually_lower_bound_power_even_indet:
+  assumes "even n" "f \<equiv> f"
+  shows   "eventually (\<lambda>x. (0::real) \<le> f x ^ n) at_top"
+  using assms by (simp add: zero_le_even_power)
+
+
+qualified lemma eventually_lower_bound_power_indet:
+  assumes "eventually (\<lambda>x. l x \<le> f x) at_top"
+  assumes "eventually (\<lambda>x. l' x \<le> l x ^ n) at_top" "eventually (\<lambda>x. l' x \<le> 0) at_top"
+  shows   "eventually (\<lambda>x. l' x \<le> (f x ^ n :: real)) at_top"
+  using assms
+proof eventually_elim
+  case (elim x)
+  thus ?case
+    using power_mono_odd[of n "l x" "f x"] zero_le_even_power[of n "f x"]
+    by (cases "even n") auto
+qed
+
+qualified lemma eventually_upper_bound_power_indet:
+  assumes "eventually (\<lambda>x. l x \<le> f x) at_top" "eventually (\<lambda>x. f x \<le> u x) at_top"
+          "eventually (\<lambda>x. u' x \<ge> -l x) at_top" "eventually (\<lambda>x. u' x \<ge> u x) at_top" "n \<equiv> n"
+  shows   "eventually (\<lambda>x. f x ^ n \<le> (u' x ^ n :: real)) at_top"
+  using assms(1-4)
+proof eventually_elim
+  case (elim x)
+  have "f x ^ n \<le> \<bar>f x ^ n\<bar>" by simp
+  also have "\<dots> = \<bar>f x\<bar> ^ n" by (simp add: power_abs)
+  also from elim have "\<dots> \<le> u' x ^ n" by (intro power_mono) auto
+  finally show ?case .
+qed
+
+qualified lemma expands_to_imp_exp_ln_eq:
+  assumes "(l expands_to L) bs" "eventually (\<lambda>x. l x \<le> f x) at_top"
+          "trimmed_pos L" "basis_wf bs"
+  shows   "eventually (\<lambda>x. exp (ln (f x)) = f x) at_top"
+proof -
+  from expands_to_imp_eventually_pos[of bs l L] assms
+    have "eventually (\<lambda>x. l x > 0) at_top" by simp
+  with assms(2) show ?thesis by eventually_elim simp
+qed
+
+qualified lemma expands_to_imp_ln_powr_eq:
+  assumes "(l expands_to L) bs" "eventually (\<lambda>x. l x \<le> f x) at_top"
+          "trimmed_pos L" "basis_wf bs"
+  shows   "eventually (\<lambda>x. ln (f x powr g x) = ln (f x) * g x) at_top"
+proof -
+  from expands_to_imp_eventually_pos[of bs l L] assms
+    have "eventually (\<lambda>x. l x > 0) at_top" by simp
+  with assms(2) show ?thesis by eventually_elim (simp add: powr_def)
+qed
+
+qualified lemma inverse_bounds_real:
+  fixes f l u :: "real \<Rightarrow> real"
+  shows "\<forall>x. l x > 0 \<longrightarrow> l x \<le> f x \<longrightarrow> f x \<le> u x \<longrightarrow> inverse (f x) \<in> {inverse (u x)..inverse (l x)}"
+    and "\<forall>x. u x < 0 \<longrightarrow> l x \<le> f x \<longrightarrow> f x \<le> u x \<longrightarrow> inverse (f x) \<in> {inverse (u x)..inverse (l x)}"
+  by (auto simp: field_simps)
+
+qualified lemma pos_imp_inverse_pos: "\<forall>x::real. f x > 0 \<longrightarrow> inverse (f x) > (0::real)"
+  and neg_imp_inverse_neg: "\<forall>x::real. f x < 0 \<longrightarrow> inverse (f x) < (0::real)"
+  by auto
+
+qualified lemma transfer_divide_bounds:
+  fixes f g :: "real \<Rightarrow> real"
+  shows "Trueprop (eventually (\<lambda>x. f x \<in> {h x * inverse (i x)..j x}) at_top) \<equiv>
+         Trueprop (eventually (\<lambda>x. f x \<in> {h x / i x..j x}) at_top)"
+    and "Trueprop (eventually (\<lambda>x. f x \<in> {j x..h x * inverse (i x)}) at_top) \<equiv>
+         Trueprop (eventually (\<lambda>x. f x \<in> {j x..h x / i x}) at_top)"
+    and "Trueprop (eventually (\<lambda>x. f x * inverse (g x) \<in> A x) at_top) \<equiv>
+         Trueprop (eventually (\<lambda>x. f x / g x \<in> A x) at_top)"
+    and "Trueprop (((\<lambda>x. f x * inverse (g x)) expands_to F) bs) \<equiv>
+         Trueprop (((\<lambda>x. f x / g x) expands_to F) bs)"
+  by (simp_all add: field_simps)
+
+qualified lemma eventually_le_self: "eventually (\<lambda>x::real. f x \<le> (f x :: real)) at_top"
+  by simp
+
+qualified lemma max_eventually_eq:
+  "eventually (\<lambda>x. f x = (g x :: real)) at_top \<Longrightarrow> eventually (\<lambda>x. max (f x) (g x) = f x) at_top"
+  by (erule eventually_mono) simp
+
+qualified lemma min_eventually_eq:
+  "eventually (\<lambda>x. f x = (g x :: real)) at_top \<Longrightarrow> eventually (\<lambda>x. min (f x) (g x) = f x) at_top"
+  by (erule eventually_mono) simp
+
+qualified lemma
+  assumes "eventually (\<lambda>x. f x = (g x :: 'a :: preorder)) F"
+  shows   eventually_eq_imp_ge: "eventually (\<lambda>x. f x \<ge> g x) F"
+    and   eventually_eq_imp_le: "eventually (\<lambda>x. f x \<le> g x) F"
+  by (use assms in eventually_elim; simp)+
+
+qualified lemma eventually_less_imp_le:
+  assumes "eventually (\<lambda>x. f x < (g x :: 'a :: order)) F"
+  shows   "eventually (\<lambda>x. f x \<le> g x) F"
+  using assms by eventually_elim auto
+
+qualified lemma
+  fixes f :: "real \<Rightarrow> real"
+  shows eventually_abs_ge_0: "eventually (\<lambda>x. abs (f x) \<ge> 0) at_top"
+    and nonneg_abs_bounds: "\<forall>x. l x \<ge> 0 \<longrightarrow> f x \<in> {l x..u x} \<longrightarrow> abs (f x) \<in> {l x..u x}" 
+    and nonpos_abs_bounds: "\<forall>x. u x \<le> 0 \<longrightarrow> f x \<in> {l x..u x} \<longrightarrow> abs (f x) \<in> {-u x..-l x}" 
+    and indet_abs_bounds:
+          "\<forall>x. l x \<le> 0 \<longrightarrow> u x \<ge> 0 \<longrightarrow> f x \<in> {l x..u x} \<longrightarrow> -l x \<le> h x \<longrightarrow> u x \<le> h x \<longrightarrow> 
+             abs (f x) \<in> {0..h x}"
+  by auto
+
+qualified lemma eventually_le_abs_nonneg:
+  "eventually (\<lambda>x. l x \<ge> 0) at_top \<Longrightarrow> eventually (\<lambda>x. f x \<ge> l x) at_top \<Longrightarrow>
+     eventually (\<lambda>x::real. l x \<le> (\<bar>f x\<bar> :: real)) at_top"
+  by (auto elim: eventually_elim2)
+
+qualified lemma eventually_le_abs_nonpos:
+  "eventually (\<lambda>x. u x \<le> 0) at_top \<Longrightarrow> eventually (\<lambda>x. f x \<le> u x) at_top \<Longrightarrow>
+     eventually (\<lambda>x::real. -u x \<le> (\<bar>f x\<bar> :: real)) at_top"
+  by (auto elim: eventually_elim2)
+
+qualified lemma
+  fixes f g h :: "'a \<Rightarrow> 'b :: order"
+  shows eventually_le_less:"eventually (\<lambda>x. f x \<le> g x) F \<Longrightarrow> eventually (\<lambda>x. g x < h x) F \<Longrightarrow> 
+           eventually (\<lambda>x. f x < h x) F"
+  and   eventually_less_le:"eventually (\<lambda>x. f x < g x) F \<Longrightarrow> eventually (\<lambda>x. g x \<le> h x) F \<Longrightarrow>
+           eventually (\<lambda>x. f x < h x) F"
+  by (erule (1) eventually_elim2; erule (1) order.trans le_less_trans less_le_trans)+
+
+qualified lemma eventually_pos_imp_nz [eventuallized]: "\<forall>x. f x > 0 \<longrightarrow> f x \<noteq> (0 :: 'a :: linordered_semiring)"
+  and eventually_neg_imp_nz [eventuallized]: "\<forall>x. f x < 0 \<longrightarrow> f x \<noteq> 0"
+  by auto
+
+qualified lemma
+  fixes f g l1 l2 u1 :: "'a \<Rightarrow> real"
+  assumes "eventually (\<lambda>x. l1 x \<le> f x) F"
+  assumes "eventually (\<lambda>x. f x \<le> u1 x) F"
+  assumes "eventually (\<lambda>x. abs (g x) \<ge> l2 x) F"
+  assumes "eventually (\<lambda>x. l2 x \<ge> 0) F"
+  shows   bounds_smallo_imp_smallo: "l1 \<in> o[F](l2) \<Longrightarrow> u1 \<in> o[F](l2) \<Longrightarrow> f \<in> o[F](g)"
+    and   bounds_bigo_imp_bigo:     "l1 \<in> O[F](l2) \<Longrightarrow> u1 \<in> O[F](l2) \<Longrightarrow> f \<in> O[F](g)"
+proof -
+  assume *: "l1 \<in> o[F](l2)" "u1 \<in> o[F](l2)"
+  show "f \<in> o[F](g)"
+  proof (rule landau_o.smallI, goal_cases)
+    case (1 c)
+    from *[THEN landau_o.smallD[OF _ 1]] and assms show ?case
+    proof eventually_elim
+      case (elim x)
+      from elim have "norm (f x) \<le> c * l2 x" by simp
+      also have "\<dots> \<le> c * norm (g x)" using 1 elim by (intro mult_left_mono) auto
+      finally show ?case .
+    qed
+  qed
+next
+  assume *: "l1 \<in> O[F](l2)" "u1 \<in> O[F](l2)"
+  then obtain C1 C2 where **: "C1 > 0" "C2 > 0" "eventually (\<lambda>x. norm (l1 x) \<le> C1 * norm (l2 x)) F"
+    "eventually (\<lambda>x. norm (u1 x) \<le> C2 * norm (l2 x)) F" by (auto elim!: landau_o.bigE)
+  from this(3,4) and assms have "eventually (\<lambda>x. norm (f x) \<le> max C1 C2 * norm (g x)) F"
+  proof eventually_elim
+    case (elim x)
+    from elim have "norm (f x) \<le> max C1 C2 * l2 x" by (subst max_mult_distrib_right) auto
+    also have "\<dots> \<le> max C1 C2 * norm (g x)" using elim ** by (intro mult_left_mono) auto
+    finally show ?case .
+  qed
+  thus "f \<in> O[F](g)" by (rule bigoI)
+qed
+
+qualified lemma real_root_mono: "mono (root n)"
+  by (cases n) (auto simp: mono_def)
+
+(* TODO: support for rintmod *)
+
+ML_file \<open>multiseries_expansion_bounds.ML\<close>
+
+method_setup real_asymp = \<open>
+let
+  type flags = {verbose : bool, fallback : bool}
+  fun join_flags
+        {verbose = verbose1, fallback = fallback1}
+        {verbose = verbose2, fallback = fallback2} =
+    {verbose = verbose1 orelse verbose2, fallback = fallback1 orelse fallback2}
+  val parse_flag =
+    (Args.$$$ "verbose" >> K {verbose = true, fallback = false}) ||
+    (Args.$$$ "fallback" >> K {verbose = false, fallback = true})
+  val default_flags = {verbose = false, fallback = false}
+  val parse_flags =
+    Scan.optional (Args.parens (Parse.list1 parse_flag)) [] >>
+      (fn xs => fold join_flags xs default_flags)
+in
+  Scan.lift parse_flags --| Method.sections Simplifier.simp_modifiers >>
+    (fn flags => SIMPLE_METHOD' o
+      (if #fallback flags then Real_Asymp_Basic.tac else Real_Asymp_Bounds.tac) (#verbose flags))
+end
+\<close> "Semi-automatic decision procedure for asymptotics of exp-log functions"
+
+end
+
+lemma "filterlim (\<lambda>x::real. sin x / x) (nhds 0) at_top"
+  by real_asymp
+
+lemma "(\<lambda>x::real. exp (sin x)) \<in> O(\<lambda>_. 1)"
+  by real_asymp
+
+lemma "(\<lambda>x::real. exp (real_of_int (floor x))) \<in> \<Theta>(\<lambda>x. exp x)"
+  by real_asymp
+
+lemma "(\<lambda>n::nat. n div 2 * ln (n div 2)) \<in> \<Theta>(\<lambda>n::nat. n * ln n)"
+  by real_asymp
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real_Asymp/Real_Asymp.thy	Sun Jul 15 14:46:57 2018 +0200
@@ -0,0 +1,49 @@
+theory Real_Asymp
+  imports Multiseries_Expansion_Bounds
+keywords
+  "real_limit" "real_expansion" :: diag  
+begin
+
+ML_file \<open>real_asymp_diag.ML\<close>
+
+(*<*)
+(* Hide constants and lemmas to avoid polluting global namespace *)
+hide_const (open)
+gbinomial_series lcoeff mssnth MSSCons MSLNil MSLCons mslmap convergent_powser convergent_powser' 
+powser mssalternate basis_wf eval_monom inverse_monom is_expansion expansion_level eval 
+zero_expansion const_expansion powr_expansion power_expansion trimmed dominant_term trimmed_pos
+trimmed_neg ms_aux_hd_exp ms_exp_gt modify_eval is_lifting expands_to asymp_powser flip_cmp_result
+compare_lists compare_expansions compare_list_0 powr_nat rfloor rceil rnatmod
+
+hide_fact (open)
+convergent_powser_stl isCont_powser powser_MSLNil powser_MSLCons
+convergent_powser'_imp_convergent_powser convergent_powser'_eventually 
+convergent_powser'_geometric convergent_powser'_exp convergent_powser'_ln gbinomial_series_abort 
+convergent_powser'_gbinomial convergent_powser'_gbinomial' convergent_powser_sin_series_stream
+convergent_powser_cos_series_stream convergent_powser_arctan basis_wf_Nil basis_wf_Cons
+basis_wf_snoc basis_wf_singleton basis_wf_many eval_monom_Nil1 eval_monom_Nil2 eval_monom_Cons 
+length_snd_inverse_monom eval_monom_pos eval_monom_uminus eval_monom_neg eval_monom_nonzero 
+eval_real eval_inverse_monom eval_monom_smallo' eval_monom_smallomega' eval_monom_smallo
+eval_monom_smallomega hd_exp_powser hd_exp_inverse eval_pos_if_dominant_term_pos 
+eval_pos_if_dominant_term_pos' eval_neg_if_dominant_term_neg' eval_modify_eval trimmed_pos_real_iff 
+trimmed_pos_ms_iff not_trimmed_pos_MSLNil trimmed_pos_MSLCons is_liftingD is_lifting_lift hd_exp_map 
+is_expansion_map is_expansion_map_aux is_expansion_map_final trimmed_map_ms is_lifting_map 
+is_lifting_id is_lifting_trans is_expansion_X dominant_term_expands_to trimmed_lifting
+trimmed_pos_lifting hd_exp_powser' compare_lists.simps compare_lists_Nil compare_lists_Cons
+flip_compare_lists compare_lists_induct eval_monom_smallo_eval_monom eval_monom_eq
+compare_expansions_real compare_expansions_MSLCons_eval compare_expansions_LT_I 
+compare_expansions_GT_I compare_expansions_same_exp compare_expansions_GT_flip compare_expansions_LT
+compare_expansions_GT compare_expansions_EQ compare_expansions_EQ_imp_bigo 
+compare_expansions_EQ_same compare_expansions_EQ_imp_bigtheta compare_expansions_LT' 
+compare_expansions_GT' compare_expansions_EQ' compare_list_0.simps compare_reals_diff_sgnD 
+trimmed_realI trimmed_pos_realI trimmed_neg_realI trimmed_pos_hd_coeff lift_trimmed lift_trimmed_pos
+lift_trimmed_neg lift_trimmed_pos' lift_trimmed_neg' trimmed_eq_cong trimmed_pos_eq_cong
+trimmed_neg_eq_cong trimmed_hd trim_lift_eq basis_wf_manyI trimmed_pos_uminus
+trimmed_pos_imp_trimmed trimmed_neg_imp_trimmed intyness_0 intyness_1 intyness_numeral
+intyness_uminus intyness_of_nat intyness_simps powr_nat_of_nat powr_nat_conv_powr reify_power
+sqrt_conv_root tan_conv_sin_cos landau_meta_eq_cong 
+asymp_equiv_meta_eq_cong eventually_lt_imp_eventually_le eventually_conv_filtermap 
+filterlim_conv_filtermap eval_simps real_eqI lift_ms_simps nhds_leI 
+(*>*)
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real_Asymp/Real_Asymp_Approx.thy	Sun Jul 15 14:46:57 2018 +0200
@@ -0,0 +1,168 @@
+(*
+  File:     Real_Asymp_Approx.thy
+  Author:   Manuel Eberl, TU München
+  
+  Integrates the approximation method into the real_asymp framework as a sign oracle
+  to automatically prove positivity/negativity of real constants
+*)
+theory Real_Asymp_Approx
+  imports Real_Asymp "HOL-Decision_Procs.Approximation"
+begin
+
+text \<open>
+  For large enough constants (such as @{term "exp (exp 10000 :: real)"}), the 
+  @{method approximation} method can require a huge amount of time and memory, effectively not
+  terminating and causing the entire prover environment to crash.
+
+  To avoid causing such situations, we first check the plausibility of the fact to prove using
+  floating-point arithmetic and only run the approximation method if the floating point evaluation
+  supports the claim. In particular, exorbitantly large numbers like the one above will lead to
+  floating point overflow, which makes the check fail.
+\<close>
+
+ML \<open>
+signature REAL_ASYMP_APPROX = 
+sig
+  val real_asymp_approx : bool Config.T
+  val sign_oracle_tac : Proof.context -> int -> tactic
+  val eval : term -> real
+end
+
+structure Real_Asymp_Approx : REAL_ASYMP_APPROX =
+struct
+
+val real_asymp_approx =
+  Attrib.setup_config_bool @{binding real_asymp_approx} (K true)
+
+val nan = Real.fromInt 0 / Real.fromInt 0
+fun clamp n = if n < 0 then 0 else n
+
+fun eval_nat (@{term "(+) :: nat => _"} $ a $ b) = bin (op +) (a, b)
+  | eval_nat (@{term "(-) :: nat => _"} $ a $ b) = bin (clamp o op -) (a, b)
+  | eval_nat (@{term "( * ) :: nat => _"} $ a $ b) = bin (op *) (a, b)
+  | eval_nat (@{term "(div) :: nat => _"} $ a $ b) = bin Int.div (a, b)
+  | eval_nat (@{term "(^) :: nat => _"} $ a $ b) = bin (fn (a,b) => Integer.pow a b) (a, b)
+  | eval_nat (t as (@{term "numeral :: _ => nat"} $ _)) = snd (HOLogic.dest_number t)
+  | eval_nat (@{term "0 :: nat"}) = 0
+  | eval_nat (@{term "1 :: nat"}) = 1
+  | eval_nat (@{term "Nat.Suc"} $ a) = un (fn n => n + 1) a
+  | eval_nat _ = ~1
+and un f a =
+      let
+        val a = eval_nat a 
+      in
+        if a >= 0 then f a else ~1
+      end
+and bin f (a, b) =
+      let
+        val (a, b) = apply2 eval_nat (a, b) 
+      in
+        if a >= 0 andalso b >= 0 then f (a, b) else ~1
+      end
+
+fun sgn n =
+  if n < Real.fromInt 0 then Real.fromInt (~1) else Real.fromInt 1
+
+fun eval (@{term "(+) :: real => _"} $ a $ b) = eval a + eval b
+  | eval (@{term "(-) :: real => _"} $ a $ b) = eval a - eval b
+  | eval (@{term "( * ) :: real => _"} $ a $ b) = eval a * eval b
+  | eval (@{term "(/) :: real => _"} $ a $ b) =
+      let val a = eval a; val b = eval b in
+        if Real.==(b, Real.fromInt 0) then nan else a / b
+      end
+  | eval (@{term "inverse :: real => _"} $ a) = Real.fromInt 1 / eval a
+  | eval (@{term "uminus :: real => _"} $ a) = Real.~ (eval a)
+  | eval (@{term "exp :: real => _"} $ a) = Math.exp (eval a)
+  | eval (@{term "ln :: real => _"} $ a) =
+      let val a = eval a in if a > Real.fromInt 0 then Math.ln a else nan end
+  | eval (@{term "(powr) :: real => _"} $ a $ b) =
+      let
+        val a = eval a; val b = eval b
+      in
+        if a < Real.fromInt 0 orelse not (Real.isFinite a) orelse not (Real.isFinite b) then
+          nan
+        else if Real.==(a, Real.fromInt 0) then
+          Real.fromInt 0
+        else 
+          Math.pow (a, b)
+      end
+  | eval (@{term "(^) :: real => _"} $ a $ b) =
+      let
+        fun powr x y = 
+          if not (Real.isFinite x) orelse y < 0 then
+            nan
+          else if x < Real.fromInt 0 andalso y mod 2 = 1 then 
+            Real.~ (Math.pow (Real.~ x, Real.fromInt y))
+          else
+            Math.pow (Real.abs x, Real.fromInt y)
+      in
+        powr (eval a) (eval_nat b)
+      end
+  | eval (@{term "root :: nat => real => _"} $ n $ a) =
+      let val a = eval a; val n = eval_nat n in
+        if n = 0 then Real.fromInt 0
+          else sgn a * Math.pow (Real.abs a, Real.fromInt 1 / Real.fromInt n) end
+  | eval (@{term "sqrt :: real => _"} $ a) =
+      let val a = eval a in sgn a * Math.sqrt (abs a) end
+  | eval (@{term "log :: real => _"} $ a $ b) =
+      let
+        val (a, b) = apply2 eval (a, b)
+      in
+        if b > Real.fromInt 0 andalso a > Real.fromInt 0 andalso Real.!= (a, Real.fromInt 1) then
+          Math.ln b / Math.ln a
+        else
+          nan
+      end
+  | eval (t as (@{term "numeral :: _ => real"} $ _)) =
+      Real.fromInt (snd (HOLogic.dest_number t))
+  | eval (@{term "0 :: real"}) = Real.fromInt 0
+  | eval (@{term "1 :: real"}) = Real.fromInt 1
+  | eval _ = nan
+
+fun sign_oracle_tac ctxt i =
+  let
+    fun tac {context = ctxt, concl, ...} =
+      let
+        val b =
+          case HOLogic.dest_Trueprop (Thm.term_of concl) of
+            @{term "(<) :: real \<Rightarrow> _"} $ lhs $ rhs =>
+              let
+                val (x, y) = apply2 eval (lhs, rhs)
+              in
+                Real.isFinite x andalso Real.isFinite y andalso x < y
+              end
+          | @{term "HOL.Not"} $ (@{term "(=) :: real \<Rightarrow> _"} $ lhs $ rhs) =>
+              let
+                val (x, y) = apply2 eval (lhs, rhs)
+              in
+                Real.isFinite x andalso Real.isFinite y andalso Real.!= (x, y)
+              end
+          | _ => false
+     in
+       if b then HEADGOAL (Approximation.approximation_tac 10 [] NONE ctxt) else no_tac
+     end
+  in
+    if Config.get ctxt real_asymp_approx then
+      Subgoal.FOCUS tac ctxt i
+    else
+      no_tac
+  end
+
+end
+\<close>
+
+setup \<open>
+  Context.theory_map (
+    Multiseries_Expansion.register_sign_oracle
+      (@{binding approximation_tac}, Real_Asymp_Approx.sign_oracle_tac))
+\<close>
+
+lemma "filterlim (\<lambda>n. (1 + (2 / 3 :: real) ^ (n + 1)) ^ 2 ^ n / 2 powr (4 / 3) ^ (n - 1))
+         at_top at_top"
+proof -
+  have [simp]: "ln 4 = 2 * ln (2 :: real)"
+    using ln_realpow[of 2 2] by simp
+  show ?thesis by (real_asymp simp: field_simps ln_div)
+qed
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real_Asymp/Real_Asymp_Examples.thy	Sun Jul 15 14:46:57 2018 +0200
@@ -0,0 +1,720 @@
+section \<open>Examples for the \<open>real_asymp\<close> method\<close>
+theory Real_Asymp_Examples
+  imports Real_Asymp
+begin
+
+context
+  includes asymp_equiv_notation
+begin
+
+subsection \<open>Dominik Gruntz's examples\<close>
+
+lemma "((\<lambda>x::real. exp x * (exp (1/x - exp (-x)) - exp (1/x))) \<longlongrightarrow> -1) at_top"
+  by real_asymp
+  
+lemma "((\<lambda>x::real. exp x * (exp (1/x + exp (-x) + exp (-(x^2))) -
+                     exp (1/x - exp (-exp x)))) \<longlongrightarrow> 1) at_top"
+  by real_asymp
+  
+lemma "filterlim (\<lambda>x::real. exp (exp (x - exp (-x)) / (1 - 1/x)) - exp (exp x)) at_top at_top"
+  by real_asymp
+
+text \<open>This example is notable because it produces an expansion of level 9.\<close>
+lemma "filterlim (\<lambda>x::real. exp (exp (exp x / (1 - 1/x))) - 
+                    exp (exp (exp x / (1 - 1/x - ln x powr (-ln x))))) at_bot at_top"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. exp (exp (exp (x + exp (-x)))) / exp (exp (exp x))) at_top at_top"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. exp (exp (exp x)) / (exp (exp (exp (x - exp (-exp x)))))) at_top at_top"
+  by real_asymp
+
+lemma "((\<lambda>x::real. exp (exp (exp x)) / exp (exp (exp x - exp (-exp (exp x))))) \<longlongrightarrow> 1) at_top"
+  by real_asymp
+
+lemma "((\<lambda>x::real. exp (exp x) / exp (exp x - exp (-exp (exp x)))) \<longlongrightarrow> 1) at_top"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. ln x ^ 2 * exp (sqrt (ln x) * ln (ln x) ^ 2 * 
+         exp (sqrt (ln (ln x)) * ln (ln (ln x)) ^ 3)) / sqrt x) (at_right 0) at_top"
+  by real_asymp
+
+lemma "((\<lambda>x::real. (x * ln x * ln (x * exp x - x^2) ^ 2) / 
+                    ln (ln (x^2 + 2*exp (exp (3*x^3*ln x))))) \<longlongrightarrow> 1/3) at_top"
+  by real_asymp
+
+lemma "((\<lambda>x::real. (exp (x * exp (-x) / (exp (-x) + exp (-(2*x^2)/(x+1)))) - exp x) / x)
+          \<longlongrightarrow> -exp 2) at_top"
+  by real_asymp
+  
+lemma "((\<lambda>x::real. (3 powr x + 5 powr x) powr (1/x)) \<longlongrightarrow> 5) at_top"
+  by real_asymp
+ 
+lemma "filterlim (\<lambda>x::real. x / (ln (x powr (ln x powr (ln 2 / ln x))))) at_top at_top"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. exp (exp (2*ln (x^5 + x) * ln (ln x))) / 
+                     exp (exp (10*ln x * ln (ln x)))) at_top at_top"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. 4/9 * (exp (exp (5/2*x powr (-5/7) + 21/8*x powr (6/11) + 
+              2*x powr (-8) + 54/17*x powr (49/45))) ^ 8) / (ln (ln (-ln (4/3*x powr (-5/14))))))
+         at_top at_top"
+  by real_asymp
+
+lemma "((\<lambda>x::real. (exp (4*x*exp (-x) / 
+                    (1/exp x + 1/exp (2*x^2/(x+1)))) - exp x) / ((exp x)^4)) \<longlongrightarrow> 1) at_top "
+  by real_asymp
+
+lemma "((\<lambda>x::real. exp (x*exp (-x) / (exp (-x) + exp (-2*x^2/(x+1))))/exp x) \<longlongrightarrow> 1) at_top"
+  by real_asymp
+
+lemma "((\<lambda>x::real. exp (exp (-x/(1+exp (-x)))) * exp (-x/(1+exp (-x/(1+exp (-x))))) * 
+           exp (exp (-x+exp (-x/(1+exp (-x))))) / (exp (-x/(1+exp (-x))))^2 - exp x + x) 
+         \<longlongrightarrow> 2) at_top"
+  by real_asymp
+
+lemma "((\<lambda>x::real. (ln(ln x + ln (ln x)) - ln (ln x)) / 
+                      (ln (ln x + ln (ln (ln x)))) * ln x) \<longlongrightarrow> 1) at_top"
+  by real_asymp
+  
+lemma "((\<lambda>x::real. exp (ln (ln (x + exp (ln x * ln (ln x)))) / 
+                       (ln (ln (ln (exp x + x + ln x)))))) \<longlongrightarrow> exp 1) at_top"
+  by real_asymp
+
+lemma "((\<lambda>x::real. exp x * (sin (1/x + exp (-x)) - sin (1/x + exp (-(x^2))))) \<longlongrightarrow> 1) at_top"
+  by real_asymp
+
+lemma "((\<lambda>x::real. exp (exp x) * (exp (sin (1/x + exp (-exp x))) - exp (sin (1/x)))) \<longlongrightarrow> 1) at_top"
+  by real_asymp
+
+lemma "((\<lambda>x::real. max x (exp x) / ln (min (exp (-x)) (exp (-exp x)))) \<longlongrightarrow> -1) at_top"
+  by real_asymp
+
+text \<open>The following example is taken from the paper by Richardson \<^emph>\<open>et al\<close>.\<close>
+lemma 
+  defines "f \<equiv> (\<lambda>x::real. ln (ln (x * exp (x * exp x) + 1)) - exp (exp (ln (ln x) + 1 / x)))"
+  shows   "(f \<longlongrightarrow> 0) at_top" (is ?thesis1)
+          "f \<sim> (\<lambda>x. -(ln x ^ 2) / (2*x))" (is ?thesis2)
+  unfolding f_def by real_asymp+
+
+  
+subsection \<open>Asymptotic inequalities related to the Akra–Bazzi theorem\<close>
+  
+definition "akra_bazzi_asymptotic1 b hb e p x \<longleftrightarrow>
+  (1 - hb * inverse b * ln x powr -(1+e)) powr p * (1 + ln (b*x + hb*x/ln x powr (1+e)) powr (-e/2))
+  \<ge> 1 + (ln x powr (-e/2) :: real)"
+definition "akra_bazzi_asymptotic1' b hb e p x \<longleftrightarrow>
+  (1 + hb * inverse b * ln x powr -(1+e)) powr p * (1 + ln (b*x + hb*x/ln x powr (1+e)) powr (-e/2))
+  \<ge> 1 + (ln x powr (-e/2) :: real)"
+definition "akra_bazzi_asymptotic2 b hb e p x \<longleftrightarrow>
+  (1 + hb * inverse b * ln x powr -(1+e)) powr p * (1 - ln (b*x + hb*x/ln x powr (1+e)) powr (-e/2))
+      \<le> 1 - ln x powr (-e/2 :: real)"
+definition "akra_bazzi_asymptotic2' b hb e p x \<longleftrightarrow>
+  (1 - hb * inverse b * ln x powr -(1+e)) powr p * (1 - ln (b*x + hb*x/ln x powr (1+e)) powr (-e/2))
+      \<le> 1 - ln x powr (-e/2 :: real)"
+definition "akra_bazzi_asymptotic3 e x \<longleftrightarrow> (1 + (ln x powr (-e/2))) / 2 \<le> (1::real)"
+definition "akra_bazzi_asymptotic4 e x \<longleftrightarrow> (1 - (ln x powr (-e/2))) * 2 \<ge> (1::real)"
+definition "akra_bazzi_asymptotic5 b hb e x \<longleftrightarrow> 
+  ln (b*x - hb*x*ln x powr -(1+e)) powr (-e/2::real) < 1"
+
+definition "akra_bazzi_asymptotic6 b hb e x \<longleftrightarrow> hb / ln x powr (1 + e :: real) < b/2"
+definition "akra_bazzi_asymptotic7 b hb e x \<longleftrightarrow> hb / ln x powr (1 + e :: real) < (1 - b) / 2"
+definition "akra_bazzi_asymptotic8 b hb e x \<longleftrightarrow> x*(1 - b - hb / ln x powr (1 + e :: real)) > 1"
+
+definition "akra_bazzi_asymptotics b hb e p x \<longleftrightarrow>
+  akra_bazzi_asymptotic1 b hb e p x \<and> akra_bazzi_asymptotic1' b hb e p x \<and>
+  akra_bazzi_asymptotic2 b hb e p x \<and> akra_bazzi_asymptotic2' b hb e p x \<and>
+  akra_bazzi_asymptotic3 e x \<and> akra_bazzi_asymptotic4 e x \<and> akra_bazzi_asymptotic5 b hb e x \<and>
+  akra_bazzi_asymptotic6 b hb e x \<and> akra_bazzi_asymptotic7 b hb e x \<and> 
+  akra_bazzi_asymptotic8 b hb e x"
+
+lemmas akra_bazzi_asymptotic_defs = 
+  akra_bazzi_asymptotic1_def akra_bazzi_asymptotic1'_def 
+  akra_bazzi_asymptotic2_def akra_bazzi_asymptotic2'_def akra_bazzi_asymptotic3_def
+  akra_bazzi_asymptotic4_def akra_bazzi_asymptotic5_def akra_bazzi_asymptotic6_def
+  akra_bazzi_asymptotic7_def akra_bazzi_asymptotic8_def akra_bazzi_asymptotics_def
+  
+lemma akra_bazzi_asymptotics:
+  assumes "\<And>b. b \<in> set bs \<Longrightarrow> b \<in> {0<..<1}" and "e > 0"
+  shows "eventually (\<lambda>x. \<forall>b\<in>set bs. akra_bazzi_asymptotics b hb e p x) at_top"
+proof (intro eventually_ball_finite ballI)
+  fix b assume "b \<in> set bs"
+  with assms have "b \<in> {0<..<1}" by simp
+  with \<open>e > 0\<close> show "eventually (\<lambda>x. akra_bazzi_asymptotics b hb e p x) at_top"
+    unfolding akra_bazzi_asymptotic_defs 
+    by (intro eventually_conj; real_asymp simp: mult_neg_pos)
+qed simp
+  
+lemma
+  fixes b \<epsilon> :: real
+  assumes "b \<in> {0<..<1}" and "\<epsilon> > 0"
+  shows "filterlim (\<lambda>x. (1 - H / (b * ln x powr (1 + \<epsilon>))) powr p * 
+           (1 + ln (b * x + H * x / ln x powr (1+\<epsilon>)) powr (-\<epsilon>/2)) - 
+           (1 + ln x powr (-\<epsilon>/2))) (at_right 0) at_top"
+  using assms by (real_asymp simp: mult_neg_pos)
+
+context
+  fixes b e p :: real
+  assumes assms: "b > 0" "e > 0"
+begin
+
+lemmas [simp] = mult_neg_pos
+
+real_limit "(\<lambda>x. ((1 - 1 / (b * ln x powr (1 + e))) powr p * 
+               (1 + ln (b * x +  x / ln x powr (1+e)) powr (-e/2)) - 
+               (1 + ln x powr (-e/2))) * ln x powr ((e / 2) + 1))"
+  facts: assms
+
+end
+
+context
+  fixes b \<epsilon> :: real
+  assumes assms: "b > 0" "\<epsilon> > 0" "\<epsilon> < 1 / 4"
+begin
+
+real_expansion "(\<lambda>x. (1 - H / (b * ln x powr (1 + \<epsilon>))) powr p * 
+           (1 + ln (b * x + H * x / ln x powr (1+\<epsilon>)) powr (-\<epsilon>/2)) - 
+           (1 + ln x powr (-\<epsilon>/2)))"
+  terms: 10 (strict)
+  facts: assms
+
+end
+
+context
+  fixes e :: real
+  assumes e: "e > 0" "e < 1 / 4"
+begin
+
+real_expansion "(\<lambda>x. (1 - 1 / (1/2 * ln x powr (1 + e))) * 
+           (1 + ln (1/2 * x + x / ln x powr (1+e)) powr (-e/2)) - 
+           (1 + ln x powr (-e/2)))"
+  terms: 10 (strict)
+  facts: e
+
+end
+
+
+subsection \<open>Concrete Mathematics\<close>
+  
+text \<open>The following inequalities are discussed on p.\ 441 in Concrete Mathematics.\<close>
+lemma 
+  fixes c \<epsilon> :: real
+  assumes "0 < \<epsilon>" "\<epsilon> < 1" "1 < c"
+    shows "(\<lambda>_::real. 1) \<in> o(\<lambda>x. ln (ln x))"
+      and "(\<lambda>x::real. ln (ln x)) \<in> o(\<lambda>x. ln x)"
+      and "(\<lambda>x::real. ln x) \<in> o(\<lambda>x. x powr \<epsilon>)"
+      and "(\<lambda>x::real. x powr \<epsilon>) \<in> o(\<lambda>x. x powr c)"
+      and "(\<lambda>x. x powr c) \<in> o(\<lambda>x. x powr ln x)"
+      and "(\<lambda>x. x powr ln x) \<in> o(\<lambda>x. c powr x)"
+      and "(\<lambda>x. c powr x) \<in> o(\<lambda>x. x powr x)"
+      and "(\<lambda>x. x powr x) \<in> o(\<lambda>x. c powr (c powr x))"
+  using assms by (real_asymp (verbose))+
+
+    
+subsection \<open>Stack Exchange\<close>
+
+text \<open>
+  http://archives.math.utk.edu/visual.calculus/1/limits.15/
+\<close>
+lemma "filterlim (\<lambda>x::real. (x * sin x) / abs x) (at_right 0) (at 0)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. x^2 / (sqrt (x^2 + 12) - sqrt (12))) (nhds (12 / sqrt 3)) (at 0)"
+proof -
+  note [simp] = powr_half_sqrt sqrt_def (* TODO: Better simproc for sqrt/root? *)
+  have "sqrt (12 :: real) = sqrt (4 * 3)"
+    by simp
+  also have "\<dots> = 2 * sqrt 3" by (subst real_sqrt_mult) simp
+  finally show ?thesis by real_asymp
+qed
+
+
+text \<open>@{url "http://math.stackexchange.com/questions/625574"}\<close>
+lemma "(\<lambda>x::real. (1 - 1/2 * x^2 - cos (x / (1 - x^2))) / x^4) \<midarrow>0\<rightarrow> 23/24"
+  by real_asymp
+
+
+text \<open>@{url "http://math.stackexchange.com/questions/122967"}\<close>
+
+real_limit "\<lambda>x. (x + 1) powr (1 + 1 / x) - x powr (1 + 1 / (x + a))"
+
+lemma "((\<lambda>x::real. ((x + 1) powr (1 + 1/x) - x powr (1 + 1 / (x + a)))) \<longlongrightarrow> 1) at_top"
+  by real_asymp
+
+
+real_limit "\<lambda>x. x ^ 2 * (arctan x - pi / 2) + x"
+
+lemma "filterlim (\<lambda>x::real. x^2 * (arctan x - pi/2) + x) (at_right 0) at_top"
+  by real_asymp
+
+
+real_limit "\<lambda>x. (root 3 (x ^ 3 + x ^ 2 + x + 1) - sqrt (x ^ 2 + x + 1) * ln (exp x + x) / x)"
+
+lemma "((\<lambda>x::real. root 3 (x^3 + x^2 + x + 1) - sqrt (x^2 + x + 1) * ln (exp x + x) / x)
+           \<longlongrightarrow> -1/6) at_top"
+  by real_asymp
+
+
+context
+  fixes a b :: real
+  assumes ab: "a > 0" "b > 0"
+begin
+real_limit "\<lambda>x. ((a powr x - x * ln a) / (b powr x - x * ln b)) powr (1 / x ^ 2)"
+  limit: "at_right 0" facts: ab
+real_limit "\<lambda>x. ((a powr x - x * ln a) / (b powr x - x * ln b)) powr (1 / x ^ 2)"
+  limit: "at_left 0" facts: ab
+lemma "(\<lambda>x. ((a powr x - x * ln a) / (b powr x - x * ln b)) powr (1 / x ^ 2))
+         \<midarrow>0\<rightarrow> exp (ln a * ln a / 2 - ln b * ln b / 2)" using ab by real_asymp
+end
+
+
+text \<open>@{url "http://math.stackexchange.com/questions/547538"}\<close>
+lemma "(\<lambda>x::real. ((x+4) powr (3/2) + exp x - 9) / x) \<midarrow>0\<rightarrow> 4"
+  by real_asymp
+
+text \<open>@{url "https://www.freemathhelp.com/forum/threads/93513"}\<close>
+lemma "((\<lambda>x::real. ((3 powr x + 4 powr x) / 4) powr (1/x)) \<longlongrightarrow> 4) at_top"
+  by real_asymp
+    
+lemma "((\<lambda>x::real. x powr (3/2) * (sqrt (x + 1) + sqrt (x - 1) - 2 * sqrt x)) \<longlongrightarrow> -1/4) at_top"
+  by real_asymp
+
+
+text \<open>@{url "https://www.math.ucdavis.edu/~kouba/CalcOneDIRECTORY/limcondirectory/LimitConstant.html"}\<close>
+lemma  "(\<lambda>x::real. (cos (2*x) - 1) / (cos x - 1)) \<midarrow>0\<rightarrow> 4"
+  by real_asymp
+
+lemma "(\<lambda>x::real. tan (2*x) / (x - pi/2)) \<midarrow>pi/2\<rightarrow> 2"
+  by real_asymp
+
+
+text \<open>@url{"https://www.math.ucdavis.edu/~kouba/CalcOneDIRECTORY/liminfdirectory/LimitInfinity.html"}\<close>
+lemma "filterlim (\<lambda>x::real. (3 powr x + 3 powr (2*x)) powr (1/x)) (nhds 9) at_top"
+  using powr_def[of 3 "2::real"] by real_asymp
+
+text \<open>@{url "https://www.math.ucdavis.edu/~kouba/CalcOneDIRECTORY/lhopitaldirectory/LHopital.html"}\<close>
+lemma "filterlim (\<lambda>x::real. (x^2 - 1) / (x^2 + 3*x - 4)) (nhds (2/5)) (at 1)"
+  by real_asymp
+    
+lemma "filterlim (\<lambda>x::real. (x - 4) / (sqrt x - 2)) (nhds 4) (at 4)"
+  by real_asymp
+    
+lemma "filterlim (\<lambda>x::real. sin x / x) (at_left 1) (at 0)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. (3 powr x - 2 powr x) / (x^2 - x)) (nhds (ln (2/3))) (at 0)"
+  by (real_asymp simp: ln_div)
+    
+lemma "filterlim (\<lambda>x::real. (1/x - 1/3) / (x^2 - 9)) (nhds (-1/54)) (at 3)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. (x * tan x) / sin (3*x)) (nhds 0) (at 0)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. sin (x^2) / (x * tan x)) (at 1) (at 0)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. (x^2 * exp x) / tan x ^ 2) (at 1) (at 0)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. exp (-1/x^2) / x^2) (at 0) (at 0)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. exp x / (5 * x + 200)) at_top at_top"
+  by real_asymp
+    
+lemma "filterlim (\<lambda>x::real. (3 + ln x) / (x^2 + 7)) (at 0) at_top"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. (x^2 + 3*x - 10) / (7*x^2 - 5*x + 4)) (at_right (1/7)) at_top"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. (ln x ^ 2) / exp (2*x)) (at_right 0) at_top"
+  by real_asymp
+    
+lemma "filterlim (\<lambda>x::real. (3 * x + 2 powr x) / (2 * x + 3 powr x)) (at 0) at_top"
+  by real_asymp
+    
+lemma "filterlim (\<lambda>x::real. (exp x + 2 / x) / (exp x + 5 / x)) (at_left 1) at_top"
+  by real_asymp
+    
+lemma "filterlim (\<lambda>x::real. sqrt (x^2 + 1) - sqrt (x + 1)) at_top at_top"
+  by real_asymp
+
+  
+lemma "filterlim (\<lambda>x::real. x * ln x) (at_left 0) (at_right 0)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. x * ln x ^ 2) (at_right 0) (at_right 0)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. ln x * tan x) (at_left 0) (at_right 0)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. x powr sin x) (at_left 1) (at_right 0)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. (1 + 3 / x) powr x) (at_left (exp 3)) at_top"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. (1 - x) powr (1 / x)) (at_left (exp (-1))) (at_right 0)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. (tan x) powr x^2) (at_left 1) (at_right 0)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. x powr (1 / sqrt x)) (at_right 1) at_top"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. ln x powr (1/x)) (at_right 1) at_top"
+  by (real_asymp (verbose))
+
+lemma "filterlim (\<lambda>x::real. x powr (x powr x)) (at_right 0) (at_right 0)"
+  by (real_asymp (verbose))
+
+
+text \<open>@{url "http://calculus.nipissingu.ca/problems/limit_problems.html"}\<close>
+lemma "((\<lambda>x::real. (x^2 - 1) / \<bar>x - 1\<bar>) \<longlongrightarrow> -2) (at_left 1)"
+      "((\<lambda>x::real. (x^2 - 1) / \<bar>x - 1\<bar>) \<longlongrightarrow> 2) (at_right 1)"
+  by real_asymp+
+
+lemma "((\<lambda>x::real. (root 3 x - 1) / \<bar>sqrt x - 1\<bar>) \<longlongrightarrow> -2 / 3) (at_left 1)"
+      "((\<lambda>x::real. (root 3 x - 1) / \<bar>sqrt x - 1\<bar>) \<longlongrightarrow> 2 / 3) (at_right 1)"
+  by real_asymp+
+
+
+text \<open>@{url "https://math.stackexchange.com/questions/547538"}\<close>
+
+lemma "(\<lambda>x::real. ((x + 4) powr (3/2) + exp x - 9) / x) \<midarrow>0\<rightarrow> 4"
+  by real_asymp
+
+text \<open>@{url "https://math.stackexchange.com/questions/625574"}\<close>
+lemma "(\<lambda>x::real. (1 - x^2 / 2 - cos (x / (1 - x^2))) / x^4) \<midarrow>0\<rightarrow> 23/24"
+  by real_asymp
+
+text \<open>@{url "https://www.mapleprimes.com/questions/151308-A-Hard-Limit-Question"}\<close>
+lemma "(\<lambda>x::real. x / (x - 1) - 1 / ln x) \<midarrow>1\<rightarrow> 1 / 2"
+  by real_asymp
+
+text \<open>@{url "https://www.freemathhelp.com/forum/threads/93513-two-extremely-difficult-limit-problems"}\<close>
+lemma "((\<lambda>x::real. ((3 powr x + 4 powr x) / 4) powr (1/x)) \<longlongrightarrow> 4) at_top"
+  by real_asymp
+
+lemma "((\<lambda>x::real. x powr 1.5 * (sqrt (x + 1) + sqrt (x - 1) - 2 * sqrt x)) \<longlongrightarrow> -1/4) at_top"
+  by real_asymp
+
+text \<open>@{url "https://math.stackexchange.com/questions/1390833"}\<close>
+context
+  fixes a b :: real
+  assumes ab: "a + b > 0" "a + b = 1"
+begin
+real_limit "\<lambda>x. (a * x powr (1 / x) + b) powr (x / ln x)"
+  facts: ab
+end
+
+
+subsection \<open>Unsorted examples\<close>
+
+context
+  fixes a :: real
+  assumes a: "a > 1"
+begin
+
+text \<open>
+  It seems that Mathematica fails to expand the following example when \<open>a\<close> is a variable.
+\<close>
+lemma "(\<lambda>x. 1 - (1 - 1 / x powr a) powr x) \<sim> (\<lambda>x. x powr (1 - a))"
+  using a by real_asymp
+
+real_limit "\<lambda>x. (1 - (1 - 1 / x powr a) powr x) * x powr (a - 1)"
+  facts: a
+
+lemma "(\<lambda>n. log 2 (real ((n + 3) choose 3) / 4) + 1) \<sim> (\<lambda>n. 3 / ln 2 * ln n)"
+proof -
+  have "(\<lambda>n. log 2 (real ((n + 3) choose 3) / 4) + 1) =
+          (\<lambda>n. log 2 ((real n + 1) * (real n + 2) * (real n + 3) / 24) + 1)"
+    by (subst binomial_gbinomial) 
+       (simp add: gbinomial_pochhammer' numeral_3_eq_3 pochhammer_Suc add_ac)
+  moreover have "\<dots> \<sim> (\<lambda>n. 3 / ln 2 * ln n)"
+    by (real_asymp simp: divide_simps)
+  ultimately show ?thesis by simp
+qed
+
+end
+
+lemma "(\<lambda>x::real. exp (sin x) / ln (cos x)) \<sim>[at 0] (\<lambda>x. -2 / x ^ 2)"
+  by real_asymp
+
+real_limit "\<lambda>x. ln (1 + 1 / x) * x"
+real_limit "\<lambda>x. ln x powr ln x / x"
+real_limit "\<lambda>x. (arctan x - pi/2) * x"
+real_limit "\<lambda>x. arctan (1/x) * x"
+
+context
+  fixes c :: real assumes c: "c \<ge> 2"
+begin
+lemma c': "c^2 - 3 > 0"
+proof -
+  from c have "c^2 \<ge> 2^2" by (rule power_mono) auto
+  thus ?thesis by simp
+qed
+
+real_limit "\<lambda>x. (c^2 - 3) * exp (-x)"
+real_limit "\<lambda>x. (c^2 - 3) * exp (-x)" facts: c'
+end
+
+lemma "c < 0 \<Longrightarrow> ((\<lambda>x::real. exp (c*x)) \<longlongrightarrow> 0) at_top"
+  by real_asymp
+
+lemma "filterlim (\<lambda>x::real. -exp (1/x)) (at_left 0) (at_left 0)"
+  by real_asymp
+
+lemma "((\<lambda>t::real. t^2 / (exp t - 1)) \<longlongrightarrow> 0) at_top"
+  by real_asymp
+
+lemma "(\<lambda>n. (1 + 1 / real n) ^ n) \<longlonglongrightarrow> exp 1"
+  by real_asymp
+  
+lemma "((\<lambda>x::real. (1 + y / x) powr x) \<longlongrightarrow> exp y) at_top"
+  by real_asymp
+
+lemma "eventually (\<lambda>x::real. x < x^2) at_top"
+  by real_asymp
+  
+lemma "eventually (\<lambda>x::real. 1 / x^2 \<ge> 1 / x) (at_left 0)"
+  by real_asymp
+
+lemma "A > 1 \<Longrightarrow> (\<lambda>n. ((1 + 1 / sqrt A) / 2) powr real_of_int (2 ^ Suc n)) \<longlonglongrightarrow> 0"
+  by (real_asymp simp: divide_simps add_pos_pos)  
+
+lemma 
+  assumes "c > (1 :: real)" "k \<noteq> 0"
+  shows   "(\<lambda>n. real n ^ k) \<in> o(\<lambda>n. c ^ n)"
+  using assms by (real_asymp (verbose))
+
+lemma "(\<lambda>x::real. exp (-(x^4))) \<in> o(\<lambda>x. exp (-(x^4)) + 1 / x ^ n)"
+  by real_asymp
+    
+lemma "(\<lambda>x::real. x^2) \<in> o[at_right 0](\<lambda>x. x)"
+  by real_asymp
+    
+lemma "(\<lambda>x::real. x^2) \<in> o[at_left 0](\<lambda>x. x)"
+  by real_asymp
+
+lemma "(\<lambda>x::real. 2 * x + x ^ 2) \<in> \<Theta>[at_left 0](\<lambda>x. x)"
+  by real_asymp
+  
+lemma "(\<lambda>x::real. inverse (x - 1)^2) \<in> \<omega>[at 1](\<lambda>x. x)"
+  by real_asymp
+
+lemma "(\<lambda>x::real. sin (1 / x)) \<sim> (\<lambda>x::real. 1 / x)"
+  by real_asymp
+
+lemma
+  assumes "q \<in> {0<..<1}"
+  shows   "LIM x at_left 1. log q (1 - x) :> at_top"
+  using assms by real_asymp
+
+context
+  fixes x k :: real
+  assumes xk: "x > 1" "k > 0"
+begin
+
+lemmas [simp] = add_pos_pos
+
+real_expansion "\<lambda>l. sqrt (1 + l * (l + 1) * 4 * pi^2 * k^2 * (log x (l + 1) - log x l)^2)"
+  terms: 2 facts: xk
+
+lemma
+  "(\<lambda>l. sqrt (1 + l * (l + 1) * 4 * pi^2 * k^2 * (log x (l + 1) - log x l)^2) -
+          sqrt (1 + 4 * pi\<^sup>2 * k\<^sup>2 / (ln x ^ 2))) \<in> O(\<lambda>l. 1 / l ^ 2)"
+  using xk by (real_asymp simp: inverse_eq_divide power_divide root_powr_inverse)
+
+end
+
+lemma "(\<lambda>x. (2 * x^3 - 128) / (sqrt x - 2)) \<midarrow>4\<rightarrow> 384"
+  by real_asymp
+
+lemma 
+  "((\<lambda>x::real. (x^2 - 1) / abs (x - 1)) \<longlongrightarrow> 2) (at_right 1)"
+  "((\<lambda>x::real. (x^2 - 1) / abs (x - 1)) \<longlongrightarrow> -2) (at_left 1)"
+  by real_asymp+
+
+lemma "(\<lambda>x::real. (root 3 x - 1) / (sqrt x - 1)) \<midarrow>1\<rightarrow> 2/3"
+  by real_asymp
+
+lemma 
+  fixes a b :: real
+  assumes "a > 1" "b > 1" "a \<noteq> b"
+  shows   "((\<lambda>x. ((a powr x - x * ln a) / (b powr x - x * ln b)) powr (1/x^3)) \<longlongrightarrow> 1) at_top"
+  using assms by real_asymp
+
+
+context
+  fixes a b :: real
+  assumes ab: "a > 0" "b > 0"
+begin
+
+lemma 
+  "((\<lambda>x. ((a powr x - x * ln a) / (b powr x - x * ln b)) powr (1 / x ^ 2)) \<longlongrightarrow>
+     exp ((ln a ^ 2 - ln b ^ 2) / 2)) (at 0)"
+  using ab by (real_asymp simp: power2_eq_square)
+
+end
+
+real_limit "\<lambda>x. 1 / sin (1/x) ^ 2 + 1 / tan (1/x) ^ 2 - 2 * x ^ 2"
+
+real_limit "\<lambda>x. ((1 / x + 4) powr 1.5 + exp (1 / x) - 9) * x"
+
+lemma "x > (1 :: real) \<Longrightarrow>
+         ((\<lambda>n. abs (x powr n / (n * (1 + x powr (2 * n)))) powr (1 / n)) \<longlongrightarrow> 1 / x) at_top"
+  by (real_asymp simp add: exp_minus field_simps)
+
+lemma "x = (1 :: real) \<Longrightarrow>
+         ((\<lambda>n. abs (x powr n / (n * (1 + x powr (2 * n)))) powr (1 / n)) \<longlongrightarrow> 1 / x) at_top"
+  by (real_asymp simp add: exp_minus field_simps)
+
+lemma "x > (0 :: real) \<Longrightarrow> x < 1 \<Longrightarrow>
+         ((\<lambda>n. abs (x powr n / (n * (1 + x powr (2 * n)))) powr (1 / n)) \<longlongrightarrow> x) at_top"
+  by real_asymp
+
+lemma "(\<lambda>x. (exp (sin b) - exp (sin (b * x))) * tan (pi * x / 2)) \<midarrow>1\<rightarrow>
+         (2*b/pi * exp (sin b) * cos b)"
+  by real_asymp
+
+(* SLOW *)
+lemma "filterlim (\<lambda>x::real. (sin (tan x) - tan (sin x))) (at 0) (at 0)"
+  by real_asymp
+
+(* SLOW *)
+lemma "(\<lambda>x::real. 1 / sin x powr (tan x ^ 2)) \<midarrow>pi/2\<rightarrow> exp (1 / 2)"
+  by (real_asymp simp: exp_minus)
+
+lemma "a > 0 \<Longrightarrow> b > 0 \<Longrightarrow> c > 0 \<Longrightarrow>
+  filterlim (\<lambda>x::real. ((a powr x + b powr x + c powr x) / 3) powr (3 / x))
+    (nhds (a * b * c)) (at 0)"
+  by (real_asymp simp: exp_add add_divide_distrib exp_minus algebra_simps)
+
+real_expansion "\<lambda>x. arctan (sin (1 / x))"
+
+real_expansion "\<lambda>x. ln (1 + 1 / x)"
+  terms: 5 (strict)
+
+real_expansion "\<lambda>x. sqrt (x + 1) - sqrt (x - 1)"
+  terms: 3 (strict)
+
+
+text \<open>
+  In this example, the first 7 terms of \<open>tan (sin x)\<close> and \<open>sin (tan x)\<close> coincide, which makes
+  the calculation a bit slow.
+\<close>
+real_limit "\<lambda>x. (tan (sin x) - sin (tan x)) / x ^ 7" limit: "at_right 0"
+
+(* SLOW *)
+real_expansion "\<lambda>x. sin (tan (1/x)) - tan (sin (1/x))"
+  terms: 1 (strict)
+
+(* SLOW *)
+real_expansion "\<lambda>x. tan (1 / x)"
+  terms: 6
+
+real_expansion "\<lambda>x::real. arctan x"
+
+real_expansion "\<lambda>x. sin (tan (1/x))"
+
+real_expansion "\<lambda>x. (sin (-1 / x) ^ 2) powr sin (-1/x)"
+
+real_limit "\<lambda>x. exp (max (sin x) 1)"
+
+lemma "filterlim (\<lambda>x::real. 1 - 1 / x + ln x) at_top at_top"
+  by (force intro: tendsto_diff filterlim_tendsto_add_at_top
+        real_tendsto_divide_at_top filterlim_ident ln_at_top)
+
+lemma "filterlim (\<lambda>i::real. i powr (-i/(i-1)) - i powr (-1/(i-1))) (at_left 1) (at_right 0)"
+  by real_asymp
+
+lemma "filterlim (\<lambda>i::real. i powr (-i/(i-1)) - i powr (-1/(i-1))) (at_right (-1)) at_top"
+  by real_asymp
+
+lemma "eventually (\<lambda>i::real. i powr (-i/(i-1)) - i powr (-1/(i-1)) < 1) (at_right 0)"
+  and "eventually (\<lambda>i::real. i powr (-i/(i-1)) - i powr (-1/(i-1)) > -1) at_top"
+  by real_asymp+
+
+end
+
+
+subsection \<open>Interval arithmetic tests\<close>
+
+lemma "filterlim (\<lambda>x::real. x + sin x) at_top at_top"
+      "filterlim (\<lambda>x::real. sin x + x) at_top at_top"
+      "filterlim (\<lambda>x::real. x + (sin x + sin x)) at_top at_top"
+  by real_asymp+
+
+lemma "filterlim (\<lambda>x::real. 2 * x + sin x * x) at_infinity at_top"
+      "filterlim (\<lambda>x::real. 2 * x + (sin x + 1) * x) at_infinity at_top"
+      "filterlim (\<lambda>x::real. 3 * x + (sin x - 1) * x) at_infinity at_top"
+      "filterlim (\<lambda>x::real. 2 * x + x * sin x) at_infinity at_top"
+      "filterlim (\<lambda>x::real. 2 * x + x * (sin x + 1)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. 3 * x + x * (sin x - 1)) at_infinity at_top"
+
+      "filterlim (\<lambda>x::real. x + sin x * sin x) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + sin x * (sin x + 1)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + sin x * (sin x - 1)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + sin x * (sin x + 2)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + sin x * (sin x - 2)) at_infinity at_top"
+
+      "filterlim (\<lambda>x::real. x + (sin x - 1) * sin x) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x - 1) * (sin x + 1)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x - 1) * (sin x - 1)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x - 1) * (sin x + 2)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x - 1) * (sin x - 2)) at_infinity at_top"
+
+      "filterlim (\<lambda>x::real. x + (sin x - 2) * sin x) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x - 2) * (sin x + 1)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x - 2) * (sin x - 1)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x - 2) * (sin x + 2)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x - 2) * (sin x - 2)) at_infinity at_top"
+
+      "filterlim (\<lambda>x::real. x + (sin x + 1) * sin x) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x + 1) * (sin x + 1)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x + 1) * (sin x - 1)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x + 1) * (sin x + 2)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x + 1) * (sin x - 2)) at_infinity at_top"
+
+      "filterlim (\<lambda>x::real. x + (sin x + 2) * sin x) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x + 2) * (sin x + 1)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x + 2) * (sin x - 1)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x + 2) * (sin x + 2)) at_infinity at_top"
+      "filterlim (\<lambda>x::real. x + (sin x + 2) * (sin x - 2)) at_infinity at_top"
+  by real_asymp+
+
+lemma "filterlim (\<lambda>x::real. x * inverse (sin x + 2)) at_top at_top"
+      "filterlim (\<lambda>x::real. x * inverse (sin x - 2)) at_bot at_top"
+      "filterlim (\<lambda>x::real. x / inverse (sin x + 2)) at_top at_top"
+      "filterlim (\<lambda>x::real. x / inverse (sin x - 2)) at_bot at_top"
+  by real_asymp+
+
+lemma "filterlim (\<lambda>x::real. \<lfloor>x\<rfloor>) at_top at_top"
+      "filterlim (\<lambda>x::real. \<lceil>x\<rceil>) at_top at_top"
+      "filterlim (\<lambda>x::real. nat \<lfloor>x\<rfloor>) at_top at_top"
+      "filterlim (\<lambda>x::real. nat \<lceil>x\<rceil>) at_top at_top"
+      "filterlim (\<lambda>x::int. nat x) at_top at_top"
+      "filterlim (\<lambda>x::int. nat x) at_top at_top"
+      "filterlim (\<lambda>n::nat. real (n mod 42) + real n) at_top at_top"
+  by real_asymp+
+
+lemma "(\<lambda>n. real_of_int \<lfloor>n\<rfloor>) \<sim>[at_top] (\<lambda>n. real_of_int n)"
+      "(\<lambda>n. sqrt \<lfloor>n\<rfloor>) \<sim>[at_top] (\<lambda>n. sqrt n)"
+  by real_asymp+
+
+lemma
+  "c > 1 \<Longrightarrow> (\<lambda>n::nat. 2 ^ (n div c) :: int) \<in> o(\<lambda>n. 2 ^ n)"
+  by (real_asymp simp: field_simps)
+
+lemma
+  "c > 1 \<Longrightarrow> (\<lambda>n::nat. 2 ^ (n div c) :: nat) \<in> o(\<lambda>n. 2 ^ n)"
+  by (real_asymp simp: field_simps)
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real_Asymp/asymptotic_basis.ML	Sun Jul 15 14:46:57 2018 +0200
@@ -0,0 +1,242 @@
+signature ASYMPTOTIC_BASIS = sig
+
+type basis_info = {wf_thm : thm, head : term}
+type basis_ln_info = {ln_thm : thm, trimmed_thm : thm}
+datatype basis' = SSng of basis_info | SCons of (basis_info * basis_ln_info * basis')
+datatype basis = SEmpty | SNE of basis'
+type lifting = thm
+
+exception BASIS of string * basis
+
+val basis_size' : basis' -> int
+val basis_size : basis -> int
+val tl_basis' : basis' -> basis
+val tl_basis : basis -> basis
+val get_basis_wf_thm' : basis' -> thm
+val get_basis_wf_thm : basis -> thm
+val get_ln_info : basis -> basis_ln_info option
+val get_basis_head' : basis' -> term
+val get_basis_head : basis -> term
+val split_basis' : basis' -> basis_info * basis_ln_info option * basis
+val split_basis : basis -> (basis_info * basis_ln_info option * basis) option
+val get_basis_list' : basis' -> term list
+val get_basis_list : basis -> term list
+val get_basis_term : basis -> term
+val extract_basis_list : thm -> term list
+
+val basis_eq' : basis' -> basis' -> bool
+val basis_eq : basis -> basis -> bool
+
+val mk_expansion_level_eq_thm' : basis' -> thm
+val mk_expansion_level_eq_thm : basis -> thm
+
+val check_basis' : basis' -> basis'
+val check_basis : basis -> basis
+
+val combine_lifts : lifting -> lifting -> lifting
+val mk_lifting : term list -> basis -> lifting
+val lift_expands_to_thm : lifting -> thm -> thm
+val lift_trimmed_thm : lifting -> thm -> thm
+val lift_trimmed_pos_thm : lifting -> thm -> thm
+val lift : basis -> thm -> thm
+
+val lift_modification' : basis' -> basis' -> basis'
+val lift_modification : basis -> basis -> basis
+
+val insert_ln' : basis' -> basis'
+val insert_ln : basis -> basis
+
+val default_basis : basis
+
+end
+
+structure Asymptotic_Basis : ASYMPTOTIC_BASIS = struct
+
+type lifting = thm
+
+val concl_of' = Thm.concl_of #> HOLogic.dest_Trueprop
+val dest_fun = dest_comb #> fst
+val dest_arg = dest_comb #> snd
+
+type basis_info = {wf_thm : thm, head : term}
+type basis_ln_info = {ln_thm : thm, trimmed_thm : thm}
+
+datatype basis' = SSng of basis_info | SCons of (basis_info * basis_ln_info * basis')
+datatype basis = SEmpty | SNE of basis'
+
+val basis_size' =
+  let
+    fun go acc (SSng _) = acc
+      | go acc (SCons (_, _, tl)) = go (acc + 1) tl
+  in
+    go 1
+  end
+
+fun basis_size SEmpty = 0
+  | basis_size (SNE b) = basis_size' b
+
+fun tl_basis' (SSng _) = SEmpty
+  | tl_basis' (SCons (_, _, tl)) = SNE tl
+
+fun tl_basis SEmpty = error "tl_basis"
+  | tl_basis (SNE basis) = tl_basis' basis
+
+fun get_basis_wf_thm' (SSng i) = #wf_thm i
+  | get_basis_wf_thm' (SCons (i, _, _)) = #wf_thm i
+
+fun get_basis_wf_thm SEmpty = @{thm basis_wf_Nil}
+  | get_basis_wf_thm (SNE basis) = get_basis_wf_thm' basis
+
+fun get_ln_info (SNE (SCons (_, info, _))) = SOME info
+  | get_ln_info _ = NONE
+
+fun get_basis_head' (SSng i) = #head i
+  | get_basis_head' (SCons (i, _, _)) = #head i
+
+fun get_basis_head SEmpty = error "get_basis_head"
+  | get_basis_head (SNE basis') = get_basis_head' basis'
+
+fun split_basis' (SSng i) = (i, NONE, SEmpty)
+  | split_basis' (SCons (i, ln_thm, tl)) = (i, SOME ln_thm, SNE tl)
+
+fun split_basis SEmpty = NONE
+  | split_basis (SNE basis) = SOME (split_basis' basis)
+
+fun get_basis_list' (SSng i) = [#head i]
+  | get_basis_list' (SCons (i, _, tl)) = #head i :: get_basis_list' tl
+
+fun get_basis_list SEmpty = []
+  | get_basis_list (SNE basis) = get_basis_list' basis
+
+val get_basis_term = HOLogic.mk_list @{typ "real => real"} o get_basis_list
+
+fun extract_basis_list thm =
+  let
+    val basis =
+      case HOLogic.dest_Trueprop (Thm.concl_of thm) of
+        Const (@{const_name "is_expansion"}, _) $ _ $ basis => basis
+      | Const (@{const_name "expands_to"}, _) $ _ $ _ $ basis => basis
+      | Const (@{const_name "basis_wf"}, _) $ basis => basis
+      | _ => raise THM ("get_basis", 1, [thm])
+  in
+    HOLogic.dest_list basis |> map Envir.eta_contract
+  end
+
+fun basis_eq' (SSng i) (SSng i') = #head i = #head i'
+  | basis_eq' (SCons (i,_,tl)) (SCons (i',_,tl')) = #head i aconv #head i' andalso basis_eq' tl tl'
+  | basis_eq' _ _ = false
+
+fun basis_eq SEmpty SEmpty = true
+  | basis_eq (SNE x) (SNE y) = basis_eq' x y
+  | basis_eq _ _ = false
+
+fun mk_expansion_level_eq_thm' (SSng _) = @{thm expansion_level_eq_Cons[OF expansion_level_eq_Nil]}
+  | mk_expansion_level_eq_thm' (SCons (_, _, tl)) = 
+      mk_expansion_level_eq_thm' tl RS @{thm expansion_level_eq_Cons}
+
+fun mk_expansion_level_eq_thm SEmpty = @{thm expansion_level_eq_Nil}
+  | mk_expansion_level_eq_thm (SNE basis) = mk_expansion_level_eq_thm' basis
+
+fun dest_wf_thm_head thm = thm |> concl_of' |> dest_arg |> dest_fun |> dest_arg
+
+fun abconv (t, t') = Envir.beta_eta_contract t aconv Envir.beta_eta_contract t'
+
+exception BASIS of (string * basis)
+
+fun check_basis' (basis as (SSng {head, wf_thm})) =
+      if abconv (dest_wf_thm_head wf_thm, head) then basis 
+        else raise BASIS ("Head mismatch", SNE basis)
+  | check_basis' (basis' as (SCons ({head, wf_thm}, {ln_thm, trimmed_thm}, basis))) =
+  case strip_comb (concl_of' ln_thm) of
+    (_, [ln_fun, ln_exp, ln_basis]) =>
+      let
+        val _ = if abconv (dest_wf_thm_head wf_thm, head) then () else 
+          raise BASIS ("Head mismatch", SNE basis')
+        val _ = if eq_list abconv (HOLogic.dest_list ln_basis, get_basis_list' basis) 
+          then () else raise BASIS ("Incorrect basis in ln_thm", SNE basis')
+        val _ = if abconv (ln_fun, @{term "\<lambda>(f::real\<Rightarrow>real) x. ln (f x)"} $ head) then () else
+          raise BASIS ("Wrong function in ln_expansion", SNE basis')
+        val _ = if abconv (ln_exp, trimmed_thm |> concl_of' |> dest_arg) then () else
+          raise BASIS ("Wrong expansion in trimmed_thm", SNE basis')
+        val _ = check_basis' basis
+      in
+        basis'
+      end
+  | _ => raise BASIS ("Malformed ln_thm", SNE basis')
+
+fun check_basis SEmpty = SEmpty
+  | check_basis (SNE basis) = SNE (check_basis' basis)
+
+fun combine_lifts thm1 thm2 = @{thm is_lifting_trans} OF [thm1, thm2]
+
+fun mk_lifting bs basis =
+  let
+    fun mk_lifting_aux [] basis =
+      (case split_basis basis of
+         NONE => @{thm is_lifting_id}
+       | SOME (_, _, basis') =>
+           combine_lifts (mk_lifting_aux [] basis') 
+             (get_basis_wf_thm basis RS @{thm is_lifting_lift}))
+    | mk_lifting_aux (b :: bs) basis =
+        (case split_basis basis of
+           NONE => raise Match
+         | SOME ({head = b', ...}, _, basis') =>
+             if b aconv b' then
+               if eq_list (op aconv) (get_basis_list basis', bs) then
+                 @{thm is_lifting_id}
+               else
+                 @{thm is_lifting_map} OF
+                   [mk_lifting_aux bs basis', mk_expansion_level_eq_thm basis']
+             else
+               combine_lifts (mk_lifting_aux (b :: bs) basis')
+                 (get_basis_wf_thm basis RS @{thm is_lifting_lift}))
+    val bs' = get_basis_list basis
+    fun err () = raise TERM ("mk_lifting", map (HOLogic.mk_list @{typ "real => real"}) [bs, bs'])
+  in
+    if subset (op aconv) (bs, bs') then
+      mk_lifting_aux bs basis handle Match => err ()
+    else
+      err ()
+  end
+
+fun lift_expands_to_thm lifting thm = @{thm expands_to_lift} OF [lifting, thm]
+fun lift_trimmed_thm lifting thm = @{thm trimmed_lifting} OF [lifting, thm]
+fun lift_trimmed_pos_thm lifting thm = @{thm trimmed_pos_lifting} OF [lifting, thm]
+fun apply_lifting lifting thm = @{thm expands_to_lift} OF [lifting, thm]
+fun lift basis thm = apply_lifting (mk_lifting (extract_basis_list thm) basis) thm
+
+fun lift_modification' (SSng s) _ = raise BASIS ("lift_modification", SNE (SSng s))
+  | lift_modification' (SCons ({wf_thm, head}, {ln_thm, trimmed_thm}, _)) new_tail =
+      let
+        val wf_thm' = @{thm basis_wf_lift_modification} OF [wf_thm, get_basis_wf_thm' new_tail]
+        val lifting = mk_lifting (extract_basis_list ln_thm) (SNE new_tail)
+        val ln_thm' = apply_lifting lifting ln_thm
+        val trimmed_thm' = lift_trimmed_pos_thm lifting trimmed_thm
+      in
+        SCons ({wf_thm = wf_thm', head = head},
+          {ln_thm = ln_thm', trimmed_thm = trimmed_thm'}, new_tail)
+      end
+
+fun lift_modification (SNE basis) (SNE new_tail) = SNE (lift_modification' basis new_tail)
+  | lift_modification _ _ = raise BASIS ("lift_modification", SEmpty)
+
+fun insert_ln' (SSng {wf_thm, head}) = 
+      let
+        val head' = Envir.eta_contract
+          (Abs ("x", @{typ real}, @{term "ln :: real \<Rightarrow> real"} $ (betapply (head, Bound 0))))
+        val info1 = {wf_thm = wf_thm RS @{thm basis_wf_insert_ln(2)}, head = head}
+        val info2 = {wf_thm = wf_thm RS @{thm basis_wf_insert_ln(1)}, head = head'}
+        val ln_thm = wf_thm RS @{thm expands_to_insert_ln}
+        val trimmed_thm = wf_thm RS @{thm trimmed_pos_insert_ln}
+      in 
+       SCons (info1, {ln_thm = ln_thm, trimmed_thm = trimmed_thm}, SSng info2)
+      end
+  | insert_ln' (basis as (SCons (_, _, tail))) = lift_modification' basis (insert_ln' tail)
+
+fun insert_ln SEmpty = raise BASIS ("Empty basis", SEmpty)
+  | insert_ln (SNE basis) = check_basis (SNE (insert_ln' basis))
+
+val default_basis = 
+  SNE (SSng {head = @{term "\<lambda>x::real. x"}, wf_thm = @{thm default_basis_wf}})
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Real_Asymp/exp_log_expression.ML	Sun Jul 15 14:46:57 2018 +0200
@@ -0,0 +1,691 @@
+signature EXP_LOG_EXPRESSION = sig
+
+exception DUP
+
+datatype expr = 
+    ConstExpr of term
+  | X
+  | Uminus of expr
+  | Add of expr * expr
+  | Minus of expr * expr
+  | Inverse of expr
+  | Mult of expr * expr
+  | Div of expr * expr
+  | Ln of expr
+  | Exp of expr
+  | Power of expr * term
+  | LnPowr of expr * expr
+  | ExpLn of expr
+  | Powr of expr * expr
+  | Powr_Nat of expr * expr
+  | Powr' of expr * term
+  | Root of expr * term
+  | Absolute of expr
+  | Sgn of expr
+  | Min of expr * expr
+  | Max of expr * expr
+  | Floor of expr
+  | Ceiling of expr
+  | Frac of expr
+  | NatMod of expr * expr
+  | Sin of expr
+  | Cos of expr
+  | ArcTan of expr
+  | Custom of string * term * expr list
+
+val preproc_term_conv : Proof.context -> conv
+val expr_to_term : expr -> term
+val reify : Proof.context -> term -> expr * thm
+val reify_simple : Proof.context -> term -> expr * thm
+
+type custom_handler = 
+  Lazy_Eval.eval_ctxt -> term -> thm list * Asymptotic_Basis.basis -> thm * Asymptotic_Basis.basis
+
+val register_custom : 
+  binding -> term -> custom_handler -> local_theory -> local_theory
+val register_custom_from_thm : 
+  binding -> thm -> custom_handler -> local_theory -> local_theory
+val expand_custom : Proof.context -> string -> custom_handler option
+
+val to_mathematica : expr -> string
+val to_maple : expr -> string
+val to_maxima : expr -> string
+val to_sympy : expr -> string
+val to_sage : expr -> string
+
+val reify_mathematica : Proof.context -> term -> string
+val reify_maple : Proof.context -> term -> string
+val reify_maxima : Proof.context -> term -> string
+val reify_sympy : Proof.context -> term -> string
+val reify_sage : Proof.context -> term -> string
+
+val limit_mathematica : string -> string
+val limit_maple : string -> string
+val limit_maxima : string -> string
+val limit_sympy : string -> string
+val limit_sage : string -> string
+
+end
+
+structure Exp_Log_Expression : EXP_LOG_EXPRESSION = struct
+
+
+datatype expr = 
+    ConstExpr of term
+  | X 
+  | Uminus of expr
+  | Add of expr * expr
+  | Minus of expr * expr
+  | Inverse of expr
+  | Mult of expr * expr
+  | Div of expr * expr
+  | Ln of expr
+  | Exp of expr
+  | Power of expr * term
+  | LnPowr of expr * expr
+  | ExpLn of expr
+  | Powr of expr * expr
+  | Powr_Nat of expr * expr
+  | Powr' of expr * term
+  | Root of expr * term
+  | Absolute of expr
+  | Sgn of expr
+  | Min of expr * expr
+  | Max of expr * expr
+  | Floor of expr
+  | Ceiling of expr
+  | Frac of expr
+  | NatMod of expr * expr
+  | Sin of expr
+  | Cos of expr
+  | ArcTan of expr
+  | Custom of string * term * expr list
+
+type custom_handler = 
+  Lazy_Eval.eval_ctxt -> term -> thm list * Asymptotic_Basis.basis -> thm * Asymptotic_Basis.basis
+
+type entry = {
+  name : string, 
+  pat : term,
+  net_pat : term,
+  expand : custom_handler
+}
+
+type entry' = {
+  pat : term,
+  net_pat : term,
+  expand : custom_handler
+}
+
+exception DUP
+
+structure Custom_Funs = Generic_Data
+(
+  type T = {
+    name_table : entry' Name_Space.table,
+    net : entry Item_Net.T
+  }
+  fun eq_entry ({name = name1, ...}, {name = name2, ...}) = (name1 = name2)
+  val empty = 
+    {
+      name_table = Name_Space.empty_table "Exp-Log Custom Function",
+      net = Item_Net.init eq_entry (fn {net_pat, ...} => [net_pat])
+    }
+  
+  fun merge ({name_table = tbl1, net = net1}, {name_table = tbl2, net = net2}) = 
+    {name_table = Name_Space.join_tables (fn _ => raise DUP) (tbl1, tbl2),
+     net = Item_Net.merge (net1, net2)}
+  val extend = I
+)
+
+fun rewrite' ctxt old_prems bounds thms ct =
+  let
+    val thy = Proof_Context.theory_of ctxt
+    fun apply_rule t thm =
+      let
+        val lhs = thm |> Thm.concl_of |> Logic.dest_equals |> fst
+        val _ = Pattern.first_order_match thy (lhs, t) (Vartab.empty, Vartab.empty)
+        val insts = (lhs, t) |> apply2 (Thm.cterm_of ctxt) |> Thm.first_order_match
+        val thm = Thm.instantiate insts thm
+        val prems = Thm.prems_of thm
+        val frees = fold Term.add_frees prems []
+      in
+        if exists (member op = bounds o fst) frees then
+          NONE
+        else
+          let
+            val thm' = thm OF (map (Thm.assume o Thm.cterm_of ctxt) prems)
+            val prems' = fold (insert op aconv) prems old_prems
+            val crhs = thm |> Thm.concl_of |> Logic.dest_equals |> snd |> Thm.cterm_of ctxt
+          in                          
+            SOME (thm', crhs, prems')
+          end
+      end
+        handle Pattern.MATCH => NONE
+    fun rewrite_subterm prems ct (Abs (x, _, _)) =
+      let
+        val (u, ctxt') = yield_singleton Variable.variant_fixes x ctxt;
+        val (v, ct') = Thm.dest_abs (SOME u) ct;
+        val (thm, prems) = rewrite' ctxt' prems (x :: bounds) thms ct'
+      in
+        if Thm.is_reflexive thm then
+          (Thm.reflexive ct, prems)
+        else
+          (Thm.abstract_rule x v thm, prems)
+      end
+    | rewrite_subterm prems ct (_ $ _) =
+      let
+        val (cs, ct) = Thm.dest_comb ct
+        val (thm, prems') = rewrite' ctxt prems bounds thms cs
+        val (thm', prems'') = rewrite' ctxt prems' bounds thms ct
+      in
+        (Thm.combination thm thm', prems'')
+      end
+    | rewrite_subterm prems ct _ = (Thm.reflexive ct, prems)
+    val t = Thm.term_of ct
+  in
+    case get_first (apply_rule t) thms of
+      NONE => rewrite_subterm old_prems ct t
+    | SOME (thm, rhs, prems) =>
+        case rewrite' ctxt prems bounds thms rhs of
+          (thm', prems) => (Thm.transitive thm thm', prems)
+  end
+
+fun rewrite ctxt thms ct =
+  let
+    val thm1 = Thm.eta_long_conversion ct
+    val rhs = thm1 |> Thm.cprop_of |> Thm.dest_comb |> snd
+    val (thm2, prems) = rewrite' ctxt [] [] thms rhs
+    val rhs = thm2 |> Thm.cprop_of |> Thm.dest_comb |> snd
+    val thm3 = Thm.eta_conversion rhs
+    val thm = Thm.transitive thm1 (Thm.transitive thm2 thm3)
+  in
+    fold (fn prem => fn thm => Thm.implies_intr (Thm.cterm_of ctxt prem) thm) prems thm
+  end
+
+fun preproc_term_conv ctxt = 
+  let
+    val thms = Named_Theorems.get ctxt @{named_theorems "real_asymp_reify_simps"}
+    val thms = map (fn thm => thm RS @{thm HOL.eq_reflection}) thms
+  in
+    rewrite ctxt thms
+  end
+
+fun register_custom' binding pat expand context =
+  let
+    val n = pat |> fastype_of |> strip_type |> fst |> length
+    val maxidx = Term.maxidx_of_term pat
+    val vars = map (fn i => Var ((Name.uu_, maxidx + i), @{typ real})) (1 upto n)
+    val net_pat = Library.foldl betapply (pat, vars)
+    val {name_table = tbl, net = net} = Custom_Funs.get context
+    val entry' = {pat = pat, net_pat = net_pat, expand = expand}
+    val (name, tbl) = Name_Space.define context true (binding, entry') tbl
+    val entry = {name = name, pat = pat, net_pat = net_pat, expand = expand}
+    val net = Item_Net.update entry net
+  in
+    Custom_Funs.put {name_table = tbl, net = net} context
+  end
+
+fun register_custom binding pat expand = 
+  let
+    fun decl phi =
+      register_custom' binding (Morphism.term phi pat) expand
+  in
+    Local_Theory.declaration {syntax = false, pervasive = false} decl
+  end
+
+fun register_custom_from_thm binding thm expand = 
+  let
+    val pat = thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> dest_comb |> snd
+  in
+    register_custom binding pat expand
+  end
+
+fun expand_custom ctxt name =
+  let
+    val {name_table, ...} = Custom_Funs.get (Context.Proof ctxt)
+  in
+    case Name_Space.lookup name_table name of
+      NONE => NONE
+    | SOME {expand, ...} => SOME expand
+  end
+
+fun expr_to_term e = 
+  let
+    fun expr_to_term' (ConstExpr c) = c
+      | expr_to_term' X = Bound 0
+      | expr_to_term' (Add (a, b)) = 
+          @{term "(+) :: real => _"} $ expr_to_term' a $ expr_to_term' b
+      | expr_to_term' (Mult (a, b)) = 
+          @{term "( *) :: real => _"} $ expr_to_term' a $ expr_to_term' b
+      | expr_to_term' (Minus (a, b)) = 
+          @{term "(-) :: real => _"} $ expr_to_term' a $ expr_to_term' b
+      | expr_to_term' (Div (a, b)) = 
+          @{term "(/) :: real => _"} $ expr_to_term' a $ expr_to_term' b
+      | expr_to_term' (Uminus a) = 
+          @{term "uminus :: real => _"} $ expr_to_term' a
+      | expr_to_term' (Inverse a) = 
+          @{term "inverse :: real => _"} $ expr_to_term' a
+      | expr_to_term' (Ln a) = 
+          @{term "ln :: real => _"} $ expr_to_term' a
+      | expr_to_term' (Exp a) = 
+          @{term "exp :: real => _"} $ expr_to_term' a
+      | expr_to_term' (Powr (a,b)) = 
+          @{term "(powr) :: real => _"} $ expr_to_term' a $ expr_to_term' b
+      | expr_to_term' (Powr_Nat (a,b)) = 
+          @{term "powr_nat :: real => _"} $ expr_to_term' a $ expr_to_term' b
+      | expr_to_term' (LnPowr (a,b)) = 
+          @{term "ln :: real => _"} $ 
+            (@{term "(powr) :: real => _"} $ expr_to_term' a $ expr_to_term' b)
+      | expr_to_term' (ExpLn a) = 
+          @{term "exp :: real => _"} $ (@{term "ln :: real => _"} $ expr_to_term' a)
+      | expr_to_term' (Powr' (a,b)) = 
+          @{term "(powr) :: real => _"} $ expr_to_term' a $ b
+      | expr_to_term' (Power (a,b)) = 
+          @{term "(^) :: real => _"} $ expr_to_term' a $ b
+      | expr_to_term' (Floor a) =
+          @{term Multiseries_Expansion.rfloor} $ expr_to_term' a
+      | expr_to_term' (Ceiling a) =
+          @{term Multiseries_Expansion.rceil} $ expr_to_term' a
+      | expr_to_term' (Frac a) =
+          @{term "Archimedean_Field.frac :: real \<Rightarrow> real"} $ expr_to_term' a
+      | expr_to_term' (NatMod (a,b)) = 
+          @{term "Multiseries_Expansion.rnatmod"} $ expr_to_term' a $ expr_to_term' b
+      | expr_to_term' (Root (a,b)) = 
+          @{term "root :: nat \<Rightarrow> real \<Rightarrow> _"} $ b $ expr_to_term' a
+      | expr_to_term' (Sin a) = 
+          @{term "sin :: real => _"} $ expr_to_term' a
+      | expr_to_term' (ArcTan a) = 
+          @{term "arctan :: real => _"} $ expr_to_term' a
+      | expr_to_term' (Cos a) = 
+          @{term "cos :: real => _"} $ expr_to_term' a
+      | expr_to_term' (Absolute a) = 
+          @{term "abs :: real => _"} $ expr_to_term' a
+      | expr_to_term' (Sgn a) =
+          @{term "sgn :: real => _"} $ expr_to_term' a
+      | expr_to_term' (Min (a,b)) = 
+          @{term "min :: real => _"} $ expr_to_term' a $ expr_to_term' b
+      | expr_to_term' (Max (a,b)) = 
+          @{term "max :: real => _"} $ expr_to_term' a $ expr_to_term' b
+      | expr_to_term' (Custom (_, t, args)) = Envir.beta_eta_contract (
+          fold (fn e => fn t => betapply (t, expr_to_term' e )) args t)
+  in
+    Abs ("x", @{typ "real"}, expr_to_term' e)
+  end
+
+fun reify_custom ctxt t =
+  let
+    val thy = Proof_Context.theory_of ctxt
+    val t = Envir.beta_eta_contract t
+    val t' = Envir.beta_eta_contract (Term.abs ("x", @{typ real}) t)
+    val {net, ...} = Custom_Funs.get (Context.Proof ctxt)
+    val entries = Item_Net.re