| 
68630
 | 
     1  | 
signature REAL_ASYMP = sig
  | 
| 
 | 
     2  | 
val tac : bool -> Proof.context -> int -> tactic
  | 
| 
 | 
     3  | 
end
  | 
| 
 | 
     4  | 
  | 
| 
 | 
     5  | 
functor Real_Asymp (Exp : EXPANSION_INTERFACE) : REAL_ASYMP = struct
  | 
| 
 | 
     6  | 
  | 
| 
 | 
     7  | 
open Lazy_Eval
  | 
| 
 | 
     8  | 
  | 
| 
 | 
     9  | 
val dest_arg = dest_comb #> snd
  | 
| 
 | 
    10  | 
  | 
| 
 | 
    11  | 
fun prove_limit_at_top ectxt f filter =
  | 
| 
 | 
    12  | 
  let
  | 
| 
 | 
    13  | 
    val ctxt = get_ctxt ectxt
  | 
| 
 | 
    14  | 
    val basis = Asymptotic_Basis.default_basis
  | 
| 
 | 
    15  | 
    val prover =
  | 
| 
 | 
    16  | 
      case filter of
  | 
| 
69597
 | 
    17  | 
        Const (\<^const_name>\<open>Topological_Spaces.nhds\<close>, _) $ _ => SOME Exp.prove_nhds
  | 
| 
 | 
    18  | 
      | \<^term>\<open>at (0 :: real)\<close> => SOME Exp.prove_at_0
  | 
| 
 | 
    19  | 
      | \<^term>\<open>at_left (0 :: real)\<close> => SOME Exp.prove_at_left_0
  | 
| 
 | 
    20  | 
      | \<^term>\<open>at_right (0 :: real)\<close> => SOME Exp.prove_at_right_0
  | 
| 
 | 
    21  | 
      | \<^term>\<open>at_infinity :: real filter\<close> => SOME Exp.prove_at_infinity
  | 
| 
 | 
    22  | 
      | \<^term>\<open>at_top :: real filter\<close> => SOME Exp.prove_at_top
  | 
| 
 | 
    23  | 
      | \<^term>\<open>at_bot :: real filter\<close> => SOME Exp.prove_at_bot
  | 
| 
68630
 | 
    24  | 
      | _ => NONE
  | 
| 
 | 
    25  | 
    val lim_thm = Option.map (fn prover => prover ectxt (Exp.expand_term ectxt f basis)) prover
  | 
| 
 | 
    26  | 
  in
  | 
| 
 | 
    27  | 
    case lim_thm of
  | 
| 
 | 
    28  | 
      NONE => no_tac
  | 
| 
 | 
    29  | 
    | SOME lim_thm =>
  | 
| 
 | 
    30  | 
        HEADGOAL (
  | 
| 
 | 
    31  | 
          resolve_tac ctxt [lim_thm, lim_thm RS @{thm filterlim_mono'}]
 | 
| 
 | 
    32  | 
          THEN_ALL_NEW (TRY o resolve_tac ctxt @{thms at_within_le_nhds at_within_le_at nhds_leI}))
 | 
| 
 | 
    33  | 
  end
  | 
| 
 | 
    34  | 
  | 
| 
 | 
    35  | 
fun prove_eventually_at_top ectxt p =
  | 
| 
 | 
    36  | 
  case Envir.eta_long [] p of
  | 
| 
69597
 | 
    37  | 
    Abs (x, \<^typ>\<open>Real.real\<close>, Const (rel, _) $ f $ g) => ((
  | 
| 
68630
 | 
    38  | 
      let
  | 
| 
69597
 | 
    39  | 
        val (f, g) = apply2 (fn t => Abs (x, \<^typ>\<open>Real.real\<close>, t)) (f, g)
  | 
| 
 | 
    40  | 
        val _ = if rel = \<^const_name>\<open>Orderings.less\<close> 
  | 
| 
 | 
    41  | 
                    orelse rel = \<^const_name>\<open>Orderings.less_eq\<close> then ()
  | 
| 
68630
 | 
    42  | 
                  else raise TERM ("prove_eventually_at_top", [p])
 | 
| 
 | 
    43  | 
        val ctxt = get_ctxt ectxt
  | 
| 
 | 
    44  | 
        val basis = Asymptotic_Basis.default_basis
  | 
| 
 | 
    45  | 
        val ([thm1, thm2], basis) = Exp.expand_terms ectxt [f, g] basis
  | 
| 
 | 
    46  | 
        val thm = Exp.prove_eventually_less ectxt (thm1, thm2, basis)
  | 
| 
 | 
    47  | 
      in
  | 
| 
 | 
    48  | 
        HEADGOAL (resolve_tac ctxt [thm, thm RS @{thm eventually_lt_imp_eventually_le}])
 | 
| 
 | 
    49  | 
      end)
  | 
| 
 | 
    50  | 
    handle TERM _ => no_tac | THM _ => no_tac)
  | 
| 
 | 
    51  | 
  | _ => raise TERM ("prove_eventually_at_top", [p])
 | 
| 
 | 
    52  | 
  | 
| 
 | 
    53  | 
fun prove_landau ectxt l f g =
  | 
| 
 | 
    54  | 
  let
  | 
| 
 | 
    55  | 
    val ctxt = get_ctxt ectxt
  | 
| 
 | 
    56  | 
    val l' = l |> dest_Const |> fst
  | 
| 
 | 
    57  | 
    val basis = Asymptotic_Basis.default_basis
  | 
| 
 | 
    58  | 
    val ([thm1, thm2], basis) = Exp.expand_terms ectxt [f, g] basis
  | 
| 
 | 
    59  | 
    val prover =
  | 
| 
 | 
    60  | 
      case l' of
  | 
| 
69597
 | 
    61  | 
        \<^const_name>\<open>smallo\<close> => Exp.prove_smallo
  | 
| 
 | 
    62  | 
      | \<^const_name>\<open>bigo\<close> => Exp.prove_bigo
  | 
| 
 | 
    63  | 
      | \<^const_name>\<open>bigtheta\<close> => Exp.prove_bigtheta
  | 
| 
 | 
    64  | 
      | \<^const_name>\<open>asymp_equiv\<close> => Exp.prove_asymp_equiv
  | 
| 
68630
 | 
    65  | 
      | _ => raise TERM ("prove_landau", [f, g])
 | 
| 
 | 
    66  | 
  in
  | 
| 
 | 
    67  | 
    HEADGOAL (resolve_tac ctxt [prover ectxt (thm1, thm2, basis)])
  | 
| 
 | 
    68  | 
  end
  | 
| 
 | 
    69  | 
  | 
| 
 | 
    70  | 
val filter_substs = 
  | 
| 
 | 
    71  | 
  @{thms at_left_to_top at_right_to_top at_left_to_top' at_right_to_top' at_bot_mirror}
 | 
| 
 | 
    72  | 
val filterlim_substs = map (fn thm => thm RS @{thm filterlim_conv_filtermap}) filter_substs
 | 
| 
 | 
    73  | 
val eventually_substs = map (fn thm => thm RS @{thm eventually_conv_filtermap}) filter_substs
 | 
| 
 | 
    74  | 
  | 
| 
 | 
    75  | 
fun preproc_exp_log_natintfun_conv ctxt =
  | 
| 
 | 
    76  | 
  let
  | 
| 
 | 
    77  | 
    fun reify_power_conv x _ ct =
  | 
| 
 | 
    78  | 
      let
  | 
| 
 | 
    79  | 
        val thm = Conv.rewr_conv @{thm reify_power} ct
 | 
| 
 | 
    80  | 
      in
  | 
| 
 | 
    81  | 
        if exists_subterm (fn t => t aconv x) (Thm.term_of ct |> dest_arg) then
  | 
| 
 | 
    82  | 
          thm
  | 
| 
 | 
    83  | 
        else
  | 
| 
 | 
    84  | 
          raise CTERM ("reify_power_conv", [ct])
 | 
| 
 | 
    85  | 
      end
  | 
| 
 | 
    86  | 
    fun conv (x, ctxt) =
  | 
| 
 | 
    87  | 
      let
  | 
| 
 | 
    88  | 
        val thms1 =
  | 
| 
69597
 | 
    89  | 
           Named_Theorems.get ctxt \<^named_theorems>\<open>real_asymp_nat_reify\<close>
  | 
| 
68630
 | 
    90  | 
        val thms2 =
  | 
| 
69597
 | 
    91  | 
           Named_Theorems.get ctxt \<^named_theorems>\<open>real_asymp_int_reify\<close>
  | 
| 
68630
 | 
    92  | 
        val ctxt' = put_simpset HOL_basic_ss ctxt addsimps (thms1 @ thms2)
  | 
| 
 | 
    93  | 
      in
  | 
| 
74560
 | 
    94  | 
        Conv.repeat_changed_conv
  | 
| 
 | 
    95  | 
         (Simplifier.rewrite ctxt'
  | 
| 
68630
 | 
    96  | 
          then_conv Conv.bottom_conv (Conv.try_conv o reify_power_conv (Thm.term_of x)) ctxt)
  | 
| 
 | 
    97  | 
      end
  | 
| 
 | 
    98  | 
  in
  | 
| 
 | 
    99  | 
    Thm.eta_long_conversion
  | 
| 
 | 
   100  | 
    then_conv Conv.abs_conv conv ctxt 
  | 
| 
 | 
   101  | 
    then_conv Thm.eta_conversion
  | 
| 
 | 
   102  | 
  end
  | 
| 
 | 
   103  | 
  | 
| 
 | 
   104  | 
fun preproc_tac ctxt =
  | 
| 
 | 
   105  | 
  let
  | 
| 
 | 
   106  | 
    fun natint_tac {context = ctxt, concl = goal, ...} =
 | 
| 
 | 
   107  | 
      let
  | 
| 
 | 
   108  | 
        val conv = preproc_exp_log_natintfun_conv ctxt
  | 
| 
 | 
   109  | 
        val conv =
  | 
| 
 | 
   110  | 
          case Thm.term_of goal of
  | 
| 
69597
 | 
   111  | 
            \<^term>\<open>HOL.Trueprop\<close> $ t => (case t of
  | 
| 
 | 
   112  | 
              Const (\<^const_name>\<open>Filter.filterlim\<close>, _) $ _ $ _ $ _ =>
  | 
| 
68630
 | 
   113  | 
                Conv.fun_conv (Conv.fun_conv (Conv.arg_conv conv))
  | 
| 
69597
 | 
   114  | 
            | Const (\<^const_name>\<open>Filter.eventually\<close>, _) $ _ $ _ =>
  | 
| 
68630
 | 
   115  | 
                Conv.fun_conv (Conv.arg_conv conv)
  | 
| 
69597
 | 
   116  | 
            | Const (\<^const_name>\<open>Set.member\<close>, _) $ _ $ (_ $ _ $ _) =>
  | 
| 
68630
 | 
   117  | 
                Conv.combination_conv (Conv.arg_conv conv) (Conv.arg_conv conv)
  | 
| 
69597
 | 
   118  | 
            | Const (\<^const_name>\<open>Landau_Symbols.asymp_equiv\<close>, _) $ _ $ _ $ _ =>
  | 
| 
68630
 | 
   119  | 
                Conv.combination_conv (Conv.fun_conv (Conv.arg_conv conv)) conv
  | 
| 
 | 
   120  | 
            | _ => Conv.all_conv)
  | 
| 
 | 
   121  | 
          | _ => Conv.all_conv
  | 
| 
 | 
   122  | 
      in
  | 
| 
 | 
   123  | 
        HEADGOAL (CONVERSION (Conv.try_conv (Conv.arg_conv conv)))
  | 
| 
 | 
   124  | 
      end
  | 
| 
 | 
   125  | 
  in
  | 
| 
 | 
   126  | 
    SELECT_GOAL (Local_Defs.unfold_tac ctxt @{thms real_asymp_preproc})
 | 
| 
 | 
   127  | 
    THEN' TRY o resolve_tac ctxt @{thms real_asymp_real_nat_transfer real_asymp_real_int_transfer}
 | 
| 
 | 
   128  | 
    THEN' TRY o resolve_tac ctxt 
  | 
| 
 | 
   129  | 
      @{thms filterlim_at_leftI filterlim_at_rightI filterlim_atI' landau_reduce_to_top}
 | 
| 
 | 
   130  | 
    THEN' TRY o resolve_tac ctxt @{thms smallo_imp_smallomega bigo_imp_bigomega}
 | 
| 
 | 
   131  | 
    THEN' TRY o Subgoal.FOCUS_PREMS natint_tac ctxt
  | 
| 
 | 
   132  | 
    THEN' TRY o resolve_tac ctxt @{thms real_asymp_nat_intros real_asymp_int_intros}
 | 
| 
 | 
   133  | 
  end
  | 
| 
 | 
   134  | 
  | 
| 
 | 
   135  | 
datatype ('a, 'b) sum = Inl of 'a | Inr of 'b
 | 
| 
 | 
   136  | 
  | 
| 
 | 
   137  | 
fun prove_eventually ectxt p filter =
  | 
| 
 | 
   138  | 
  case filter of
  | 
| 
69597
 | 
   139  | 
    \<^term>\<open>Filter.at_top :: real filter\<close> => (prove_eventually_at_top ectxt p
  | 
| 
68630
 | 
   140  | 
      handle TERM _ => no_tac | THM _ => no_tac)
  | 
| 
 | 
   141  | 
  | _ => HEADGOAL (CONVERSION (Conv.rewrs_conv eventually_substs) 
  | 
| 
 | 
   142  | 
         THEN' tac' (#verbose (#ctxt ectxt)) (Inr ectxt))
  | 
| 
 | 
   143  | 
and prove_limit ectxt f filter filter' =
  | 
| 
 | 
   144  | 
  case filter' of
  | 
| 
69597
 | 
   145  | 
    \<^term>\<open>Filter.at_top :: real filter\<close> => (prove_limit_at_top ectxt f filter 
  | 
| 
68630
 | 
   146  | 
      handle TERM _ => no_tac | THM _ => no_tac)
  | 
| 
 | 
   147  | 
  | _ => HEADGOAL (CONVERSION (Conv.rewrs_conv filterlim_substs) 
  | 
| 
 | 
   148  | 
         THEN' tac' (#verbose (#ctxt ectxt)) (Inr ectxt))
  | 
| 
 | 
   149  | 
and tac' verbose ctxt_or_ectxt =
  | 
| 
 | 
   150  | 
  let
  | 
| 
 | 
   151  | 
    val ctxt = case ctxt_or_ectxt of Inl ctxt => ctxt | Inr ectxt => get_ctxt ectxt
  | 
| 
 | 
   152  | 
    fun tac {context = ctxt, prems, concl = goal, ...} =
 | 
| 
 | 
   153  | 
      (if verbose then print_tac ctxt "real_asymp: Goal after preprocessing" else all_tac) THEN
  | 
| 
 | 
   154  | 
      let
  | 
| 
 | 
   155  | 
        val ectxt = 
  | 
| 
 | 
   156  | 
          case ctxt_or_ectxt of 
  | 
| 
 | 
   157  | 
            Inl _ => 
  | 
| 
 | 
   158  | 
              Multiseries_Expansion.mk_eval_ctxt ctxt |> add_facts prems |> set_verbose verbose
  | 
| 
 | 
   159  | 
          | Inr ectxt => ectxt
  | 
| 
 | 
   160  | 
      in
  | 
| 
 | 
   161  | 
        case Thm.term_of goal of
  | 
| 
69597
 | 
   162  | 
          \<^term>\<open>HOL.Trueprop\<close> $ t => ((case t of
  | 
| 
 | 
   163  | 
            \<^term>\<open>Filter.filterlim :: (real \<Rightarrow> real) \<Rightarrow> _\<close> $ f $ filter $ filter' =>
  | 
| 
68630
 | 
   164  | 
              (prove_limit ectxt f filter filter' handle TERM _ => no_tac | THM _ => no_tac)
  | 
| 
69597
 | 
   165  | 
          | \<^term>\<open>Filter.eventually :: (real \<Rightarrow> bool) \<Rightarrow> _\<close> $ p $ filter =>
  | 
| 
68630
 | 
   166  | 
              (prove_eventually ectxt p filter handle TERM _ => no_tac | THM _ => no_tac)
  | 
| 
69597
 | 
   167  | 
          | \<^term>\<open>Set.member :: (real => real) => _\<close> $ f $ 
  | 
| 
 | 
   168  | 
              (l $ \<^term>\<open>at_top :: real filter\<close> $ g) =>
  | 
| 
68630
 | 
   169  | 
                (prove_landau ectxt l f g handle TERM _ => no_tac | THM _ => no_tac)
  | 
| 
69597
 | 
   170  | 
          | (l as \<^term>\<open>Landau_Symbols.asymp_equiv :: (real\<Rightarrow>real)\<Rightarrow>_\<close>) $ f $ _ $ g =>
  | 
| 
68630
 | 
   171  | 
              (prove_landau ectxt l f g handle TERM _ => no_tac | THM _ => no_tac)
  | 
| 
 | 
   172  | 
          | _ => no_tac) THEN distinct_subgoals_tac)
  | 
| 
 | 
   173  | 
        | _ => no_tac
  | 
| 
 | 
   174  | 
      end
  | 
| 
 | 
   175  | 
    fun tac' i = Subgoal.FOCUS_PREMS tac ctxt i handle TERM _ => no_tac | THM _ => no_tac
  | 
| 
 | 
   176  | 
    val at_tac =
  | 
| 
 | 
   177  | 
      HEADGOAL (resolve_tac ctxt 
  | 
| 
 | 
   178  | 
        @{thms filterlim_split_at eventually_at_left_at_right_imp_at landau_at_top_imp_at
 | 
| 
 | 
   179  | 
                 asymp_equiv_at_top_imp_at})
  | 
| 
 | 
   180  | 
      THEN PARALLEL_ALLGOALS tac'
  | 
| 
 | 
   181  | 
  in
  | 
| 
 | 
   182  | 
    (preproc_tac ctxt
  | 
| 
 | 
   183  | 
     THEN' preproc_tac ctxt
  | 
| 
 | 
   184  | 
     THEN' (SELECT_GOAL at_tac ORELSE' tac'))
  | 
| 
 | 
   185  | 
    THEN_ALL_NEW (TRY o SELECT_GOAL (SOLVE (HEADGOAL (Simplifier.asm_full_simp_tac ctxt))))
  | 
| 
 | 
   186  | 
  end
  | 
| 
 | 
   187  | 
and tac verbose ctxt = tac' verbose (Inl ctxt)
  | 
| 
 | 
   188  | 
  | 
| 
 | 
   189  | 
end
  | 
| 
 | 
   190  | 
  | 
| 
 | 
   191  | 
structure Real_Asymp_Basic = Real_Asymp(Multiseries_Expansion_Basic)
  |