| author | wenzelm | 
| Tue, 05 Mar 2024 20:58:41 +0100 | |
| changeset 79793 | 6f08aef43dc5 | 
| parent 77879 | dd222e2af01a | 
| child 79930 | 7bac6bd83cc3 | 
| permissions | -rw-r--r-- | 
| 68630 | 1  | 
signature MULTISERIES_EXPANSION = sig  | 
2  | 
||
3  | 
type expansion_thm = thm  | 
|
4  | 
type trimmed_thm = thm  | 
|
5  | 
type expr = Exp_Log_Expression.expr  | 
|
6  | 
type basis = Asymptotic_Basis.basis  | 
|
7  | 
||
8  | 
datatype trim_mode = Simple_Trim | Pos_Trim | Neg_Trim | Sgn_Trim  | 
|
9  | 
||
10  | 
datatype zeroness = IsZero | IsNonZero | IsPos | IsNeg  | 
|
11  | 
||
12  | 
datatype intyness = Nat of thm | Neg_Nat of thm | No_Nat  | 
|
13  | 
datatype parity = Even of thm | Odd of thm | Unknown_Parity  | 
|
14  | 
||
15  | 
datatype limit =  | 
|
16  | 
Zero_Limit of bool option  | 
|
17  | 
| Finite_Limit of term  | 
|
18  | 
| Infinite_Limit of bool option  | 
|
19  | 
||
20  | 
datatype trim_result =  | 
|
21  | 
Trimmed of zeroness * trimmed_thm option  | 
|
22  | 
| Aborted of order  | 
|
23  | 
||
24  | 
val get_intyness : Proof.context -> cterm -> intyness  | 
|
25  | 
val get_parity : cterm -> parity  | 
|
26  | 
||
27  | 
val get_expansion : thm -> term  | 
|
28  | 
val get_coeff : term -> term  | 
|
29  | 
val get_exponent : term -> term  | 
|
30  | 
val get_expanded_fun : thm -> term  | 
|
31  | 
val get_eval : term -> term  | 
|
32  | 
val expands_to_hd : thm -> thm  | 
|
33  | 
||
34  | 
val mk_eval_ctxt : Proof.context -> Lazy_Eval.eval_ctxt  | 
|
35  | 
val expand : Lazy_Eval.eval_ctxt -> expr -> basis -> expansion_thm * basis  | 
|
36  | 
val expand_term : Lazy_Eval.eval_ctxt -> term -> basis -> expansion_thm * basis  | 
|
37  | 
val expand_terms : Lazy_Eval.eval_ctxt -> term list -> basis -> expansion_thm list * basis  | 
|
38  | 
||
39  | 
val limit_of_expansion : bool * bool -> Lazy_Eval.eval_ctxt -> thm * basis -> limit * thm  | 
|
40  | 
val compute_limit : Lazy_Eval.eval_ctxt -> term -> limit * thm  | 
|
41  | 
val compare_expansions :  | 
|
42  | 
Lazy_Eval.eval_ctxt -> expansion_thm * expansion_thm * basis ->  | 
|
43  | 
order * thm * expansion_thm * expansion_thm  | 
|
44  | 
||
45  | 
(* TODO DEBUG *)  | 
|
46  | 
datatype comparison_result =  | 
|
47  | 
Cmp_Dominated of order * thm list * zeroness * trimmed_thm * expansion_thm * expansion_thm  | 
|
48  | 
| Cmp_Asymp_Equiv of thm * thm  | 
|
49  | 
val compare_expansions' :  | 
|
50  | 
Lazy_Eval.eval_ctxt ->  | 
|
51  | 
thm * thm * Asymptotic_Basis.basis ->  | 
|
52  | 
comparison_result  | 
|
53  | 
||
54  | 
val prove_at_infinity : Lazy_Eval.eval_ctxt -> thm * basis -> thm  | 
|
55  | 
val prove_at_top : Lazy_Eval.eval_ctxt -> thm * basis -> thm  | 
|
56  | 
val prove_at_bot : Lazy_Eval.eval_ctxt -> thm * basis -> thm  | 
|
57  | 
val prove_nhds : Lazy_Eval.eval_ctxt -> thm * basis -> thm  | 
|
58  | 
val prove_at_0 : Lazy_Eval.eval_ctxt -> thm * basis -> thm  | 
|
59  | 
val prove_at_left_0 : Lazy_Eval.eval_ctxt -> thm * basis -> thm  | 
|
60  | 
val prove_at_right_0 : Lazy_Eval.eval_ctxt -> thm * basis -> thm  | 
|
61  | 
||
62  | 
val prove_smallo : Lazy_Eval.eval_ctxt -> thm * thm * basis -> thm  | 
|
63  | 
val prove_bigo : Lazy_Eval.eval_ctxt -> thm * thm * basis -> thm  | 
|
64  | 
val prove_bigtheta : Lazy_Eval.eval_ctxt -> thm * thm * basis -> thm  | 
|
65  | 
val prove_asymp_equiv : Lazy_Eval.eval_ctxt -> thm * thm * basis -> thm  | 
|
66  | 
||
67  | 
val prove_asymptotic_relation : Lazy_Eval.eval_ctxt -> thm * thm * basis -> order * thm  | 
|
68  | 
val prove_eventually_less : Lazy_Eval.eval_ctxt -> thm * thm * basis -> thm  | 
|
69  | 
val prove_eventually_greater : Lazy_Eval.eval_ctxt -> thm * thm * basis -> thm  | 
|
70  | 
val prove_eventually_nonzero : Lazy_Eval.eval_ctxt -> thm * basis -> thm  | 
|
71  | 
||
72  | 
val extract_terms : int * bool -> Lazy_Eval.eval_ctxt -> basis -> term -> term * term option  | 
|
73  | 
||
74  | 
(* Internal functions *)  | 
|
75  | 
val check_expansion : Exp_Log_Expression.expr -> expansion_thm -> expansion_thm  | 
|
76  | 
||
77  | 
val zero_expansion : basis -> expansion_thm  | 
|
78  | 
val const_expansion : Lazy_Eval.eval_ctxt -> basis -> term -> expansion_thm  | 
|
79  | 
val ln_expansion :  | 
|
80  | 
Lazy_Eval.eval_ctxt -> trimmed_thm -> expansion_thm -> basis -> expansion_thm * basis  | 
|
81  | 
val exp_expansion : Lazy_Eval.eval_ctxt -> expansion_thm -> basis -> expansion_thm * basis  | 
|
82  | 
val powr_expansion :  | 
|
83  | 
Lazy_Eval.eval_ctxt -> expansion_thm * expansion_thm * basis -> expansion_thm * basis  | 
|
84  | 
val powr_const_expansion :  | 
|
85  | 
Lazy_Eval.eval_ctxt -> expansion_thm * term * basis -> expansion_thm  | 
|
86  | 
val powr_nat_expansion :  | 
|
87  | 
Lazy_Eval.eval_ctxt -> expansion_thm * expansion_thm * basis -> expansion_thm * basis  | 
|
88  | 
val power_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * term * basis -> expansion_thm  | 
|
89  | 
val root_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * term * basis -> expansion_thm  | 
|
90  | 
||
91  | 
val sgn_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * basis -> expansion_thm  | 
|
92  | 
val min_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * expansion_thm * basis -> expansion_thm  | 
|
93  | 
val max_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * expansion_thm * basis -> expansion_thm  | 
|
94  | 
val arctan_expansion : Lazy_Eval.eval_ctxt -> basis -> expansion_thm -> expansion_thm  | 
|
95  | 
||
96  | 
val ev_zeroness_oracle : Lazy_Eval.eval_ctxt -> term -> thm option  | 
|
97  | 
val zeroness_oracle : bool -> trim_mode option -> Lazy_Eval.eval_ctxt -> term -> zeroness * thm option  | 
|
98  | 
||
99  | 
val whnf_expansion : Lazy_Eval.eval_ctxt -> expansion_thm -> term option * expansion_thm * thm  | 
|
100  | 
val simplify_expansion : Lazy_Eval.eval_ctxt -> expansion_thm -> expansion_thm  | 
|
101  | 
val simplify_term : Lazy_Eval.eval_ctxt -> term -> term  | 
|
102  | 
||
103  | 
val trim_expansion_while_greater :  | 
|
104  | 
bool -> term list option -> bool -> trim_mode option -> Lazy_Eval.eval_ctxt ->  | 
|
105  | 
thm * Asymptotic_Basis.basis -> thm * trim_result * (zeroness * thm) list  | 
|
106  | 
val trim_expansion : bool -> trim_mode option -> Lazy_Eval.eval_ctxt -> expansion_thm * basis ->  | 
|
107  | 
expansion_thm * zeroness * trimmed_thm option  | 
|
108  | 
val try_drop_leading_term_ex : bool -> Lazy_Eval.eval_ctxt -> expansion_thm -> expansion_thm option  | 
|
109  | 
||
110  | 
val try_prove_real_eq : bool -> Lazy_Eval.eval_ctxt -> term * term -> thm option  | 
|
111  | 
val try_prove_ev_eq : Lazy_Eval.eval_ctxt -> term * term -> thm option  | 
|
112  | 
val prove_compare_expansions : order -> thm list -> thm  | 
|
113  | 
||
114  | 
val simplify_trimmed_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * trimmed_thm ->  | 
|
115  | 
expansion_thm * trimmed_thm  | 
|
116  | 
val retrim_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * basis -> expansion_thm * thm  | 
|
117  | 
val retrim_pos_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * basis * trimmed_thm ->  | 
|
118  | 
expansion_thm * thm * trimmed_thm  | 
|
119  | 
||
120  | 
val register_sign_oracle :  | 
|
121  | 
binding * (Proof.context -> int -> tactic) -> Context.generic -> Context.generic  | 
|
122  | 
val get_sign_oracles :  | 
|
123  | 
Context.generic -> (string * (Proof.context -> int -> tactic)) list  | 
|
124  | 
||
125  | 
val solve_eval_eq : thm -> thm  | 
|
126  | 
||
127  | 
end  | 
|
128  | 
||
129  | 
structure Multiseries_Expansion : MULTISERIES_EXPANSION = struct  | 
|
130  | 
||
131  | 
open Asymptotic_Basis  | 
|
132  | 
open Exp_Log_Expression  | 
|
133  | 
open Lazy_Eval  | 
|
134  | 
||
135  | 
structure Data = Generic_Data  | 
|
136  | 
(  | 
|
137  | 
type T = (Proof.context -> int -> tactic) Name_Space.table;  | 
|
| 
74113
 
228adc502803
proper name space "kind": this is a formal name, not comment;
 
wenzelm 
parents: 
69597 
diff
changeset
 | 
138  | 
val empty : T = Name_Space.empty_table "sign_oracle_tactic";  | 
| 68630 | 139  | 
fun merge (tactics1, tactics2) : T = Name_Space.merge_tables (tactics1, tactics2);  | 
140  | 
);  | 
|
141  | 
||
142  | 
fun register_sign_oracle (s, tac) ctxt =  | 
|
143  | 
Data.map (Name_Space.define ctxt false (s, tac) #> snd) ctxt  | 
|
144  | 
||
145  | 
fun get_sign_oracles ctxt = Name_Space.fold_table cons (Data.get ctxt) []  | 
|
146  | 
||
147  | 
fun apply_sign_oracles ctxt tac =  | 
|
148  | 
let  | 
|
149  | 
val oracles = get_sign_oracles (Context.Proof ctxt)  | 
|
150  | 
    fun tac' {context = ctxt, concl, ...} =
 | 
|
| 69597 | 151  | 
if Thm.term_of concl = \<^term>\<open>HOL.Trueprop HOL.False\<close> then  | 
| 68630 | 152  | 
no_tac  | 
153  | 
else  | 
|
154  | 
FIRST (map (fn tac => HEADGOAL (snd tac ctxt)) oracles)  | 
|
155  | 
in  | 
|
156  | 
tac THEN_ALL_NEW (Subgoal.FOCUS_PREMS tac' ctxt)  | 
|
157  | 
end  | 
|
158  | 
||
159  | 
||
160  | 
type expansion_thm = thm  | 
|
161  | 
type trimmed_thm = thm  | 
|
162  | 
||
163  | 
val dest_fun = dest_comb #> fst  | 
|
164  | 
val dest_arg = dest_comb #> snd  | 
|
165  | 
val concl_of' = Thm.concl_of #> HOLogic.dest_Trueprop  | 
|
166  | 
||
167  | 
fun get_expansion thm =  | 
|
168  | 
thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> Term.dest_comb |> fst |> Term.dest_comb |> snd  | 
|
169  | 
||
170  | 
fun get_expanded_fun thm = thm |> concl_of' |> dest_fun |> dest_fun |> dest_arg  | 
|
171  | 
||
172  | 
(*  | 
|
173  | 
The following function is useful in order to detect whether a given real constant is  | 
|
174  | 
an integer, which allows us to use the "f(x) ^ n" operation instead of "f(x) powr n".  | 
|
175  | 
This usually leads to nicer results.  | 
|
176  | 
*)  | 
|
177  | 
datatype intyness = Nat of thm | Neg_Nat of thm | No_Nat  | 
|
178  | 
||
179  | 
fun get_intyness ctxt ct =  | 
|
| 69597 | 180  | 
if Thm.typ_of_cterm ct = \<^typ>\<open>Real.real\<close> then  | 
| 68630 | 181  | 
let  | 
182  | 
      val ctxt' = put_simpset HOL_basic_ss ctxt addsimps @{thms intyness_simps}
 | 
|
183  | 
val conv =  | 
|
184  | 
Simplifier.rewrite ctxt then_conv Simplifier.rewrite ctxt'  | 
|
185  | 
      fun flip (Nat thm) = Neg_Nat (thm RS @{thm intyness_uminus})
 | 
|
186  | 
| flip _ = No_Nat  | 
|
187  | 
fun get_intyness' ct =  | 
|
188  | 
case Thm.term_of ct of  | 
|
| 69597 | 189  | 
          \<^term>\<open>0::real\<close> => Nat @{thm intyness_0}
 | 
190  | 
        | \<^term>\<open>1::real\<close> => Nat @{thm intyness_1}
 | 
|
191  | 
| Const (\<^const_name>\<open>numeral\<close>, _) $ _ =>  | 
|
| 68630 | 192  | 
            Nat (Thm.reflexive (Thm.dest_arg ct) RS @{thm intyness_numeral})
 | 
| 69597 | 193  | 
| Const (\<^const_name>\<open>uminus\<close>, _) $ _ => flip (get_intyness' (Thm.dest_arg ct))  | 
194  | 
| Const (\<^const_name>\<open>of_nat\<close>, _) $ _ =>  | 
|
| 68630 | 195  | 
            Nat (Thm.reflexive (Thm.dest_arg ct) RS @{thm intyness_of_nat})
 | 
196  | 
| _ => No_Nat  | 
|
197  | 
val thm = conv ct  | 
|
198  | 
val ct' = thm |> Thm.cprop_of |> Thm.dest_equals_rhs  | 
|
199  | 
in  | 
|
200  | 
case get_intyness' ct' of  | 
|
201  | 
        Nat thm' => Nat (Thm.transitive thm thm' RS @{thm HOL.meta_eq_to_obj_eq})
 | 
|
202  | 
      | Neg_Nat thm' => Neg_Nat (Thm.transitive thm thm' RS @{thm HOL.meta_eq_to_obj_eq})
 | 
|
203  | 
| No_Nat => No_Nat  | 
|
204  | 
end  | 
|
205  | 
handle CTERM _ => No_Nat  | 
|
206  | 
else  | 
|
207  | 
No_Nat  | 
|
208  | 
||
209  | 
datatype parity = Even of thm | Odd of thm | Unknown_Parity  | 
|
210  | 
||
211  | 
(* TODO: powers *)  | 
|
212  | 
fun get_parity ct =  | 
|
213  | 
let  | 
|
214  | 
fun inst thm cts =  | 
|
215  | 
let  | 
|
216  | 
val tvs = Term.add_tvars (Thm.concl_of thm) []  | 
|
217  | 
in  | 
|
218  | 
case tvs of  | 
|
219  | 
[v] =>  | 
|
220  | 
let  | 
|
| 77879 | 221  | 
val thm' = Thm.instantiate (TVars.make1 (v, Thm.ctyp_of_cterm ct), Vars.empty) thm  | 
| 68630 | 222  | 
val vs = take (length cts) (rev (Term.add_vars (Thm.concl_of thm') []))  | 
223  | 
in  | 
|
| 74282 | 224  | 
Thm.instantiate (TVars.empty, Vars.make (vs ~~ cts)) thm'  | 
| 68630 | 225  | 
end  | 
226  | 
        | _ => raise THM ("get_parity", 0, [thm])
 | 
|
227  | 
end  | 
|
228  | 
val get_num = Thm.dest_arg o Thm.dest_arg  | 
|
229  | 
in  | 
|
230  | 
case Thm.term_of ct of  | 
|
| 69597 | 231  | 
      Const (\<^const_name>\<open>Groups.zero\<close>, _) => Even (inst @{thm even_zero} [])
 | 
232  | 
    | Const (\<^const_name>\<open>Groups.one\<close>, _) => Odd (inst @{thm odd_one} [])
 | 
|
233  | 
| Const (\<^const_name>\<open>Num.numeral_class.numeral\<close>, _) $ \<^term>\<open>Num.One\<close> =>  | 
|
| 68630 | 234  | 
        Odd (inst @{thm odd_Numeral1} [])
 | 
| 69597 | 235  | 
| Const (\<^const_name>\<open>Num.numeral_class.numeral\<close>, _) $ (\<^term>\<open>Num.Bit0\<close> $ _) =>  | 
| 68630 | 236  | 
        Even (inst @{thm even_numeral} [get_num ct])
 | 
| 69597 | 237  | 
| Const (\<^const_name>\<open>Num.numeral_class.numeral\<close>, _) $ (\<^term>\<open>Num.Bit1\<close> $ _) =>  | 
| 68630 | 238  | 
        Odd (inst @{thm odd_numeral} [get_num ct])
 | 
| 69597 | 239  | 
| Const (\<^const_name>\<open>Groups.uminus\<close>, _) $ _ => (  | 
| 68630 | 240  | 
case get_parity (Thm.dest_arg ct) of  | 
241  | 
          Even thm => Even (@{thm even_uminusI} OF [thm])
 | 
|
242  | 
        | Odd thm => Odd (@{thm odd_uminusI} OF [thm])
 | 
|
243  | 
| _ => Unknown_Parity)  | 
|
| 69597 | 244  | 
| Const (\<^const_name>\<open>Groups.plus\<close>, _) $ _ $ _ => (  | 
| 68630 | 245  | 
case apply2 get_parity (Thm.dest_binop ct) of  | 
246  | 
          (Even thm1, Even thm2) => Even (@{thm even_addI(1)} OF [thm1, thm2])
 | 
|
247  | 
        | (Odd thm1, Odd thm2) => Even (@{thm even_addI(2)} OF [thm1, thm2])
 | 
|
248  | 
        | (Even thm1, Odd thm2) => Odd (@{thm odd_addI(1)} OF [thm1, thm2])
 | 
|
249  | 
        | (Odd thm1, Even thm2) => Odd (@{thm odd_addI(2)} OF [thm1, thm2])
 | 
|
250  | 
| _ => Unknown_Parity)  | 
|
| 69597 | 251  | 
| Const (\<^const_name>\<open>Groups.minus\<close>, _) $ _ $ _ => (  | 
| 68630 | 252  | 
case apply2 get_parity (Thm.dest_binop ct) of  | 
253  | 
          (Even thm1, Even thm2) => Even (@{thm even_diffI(1)} OF [thm1, thm2])
 | 
|
254  | 
        | (Odd thm1, Odd thm2) => Even (@{thm even_diffI(2)} OF [thm1, thm2])
 | 
|
255  | 
        | (Even thm1, Odd thm2) => Odd (@{thm odd_diffI(1)} OF [thm1, thm2])
 | 
|
256  | 
        | (Odd thm1, Even thm2) => Odd (@{thm odd_diffI(2)} OF [thm1, thm2])
 | 
|
257  | 
| _ => Unknown_Parity)  | 
|
| 69597 | 258  | 
| Const (\<^const_name>\<open>Groups.times\<close>, _) $ _ $ _ => (  | 
| 68630 | 259  | 
case apply2 get_parity (Thm.dest_binop ct) of  | 
260  | 
          (Even thm1, _) => Even (@{thm even_multI(1)} OF [thm1])
 | 
|
261  | 
        | (_, Even thm2) => Even (@{thm even_multI(2)} OF [thm2])
 | 
|
262  | 
        | (Odd thm1, Odd thm2) => Odd (@{thm odd_multI} OF [thm1, thm2])
 | 
|
263  | 
| _ => Unknown_Parity)  | 
|
| 69597 | 264  | 
| Const (\<^const_name>\<open>Power.power\<close>, _) $ _ $ _ =>  | 
| 68630 | 265  | 
let  | 
266  | 
val (a, n) = Thm.dest_binop ct  | 
|
267  | 
in  | 
|
268  | 
case get_parity a of  | 
|
269  | 
            Odd thm => Odd (inst @{thm odd_powerI} [a, n] OF [thm])
 | 
|
270  | 
| _ => Unknown_Parity  | 
|
271  | 
end  | 
|
272  | 
| _ => Unknown_Parity  | 
|
273  | 
end  | 
|
274  | 
||
275  | 
fun simplify_term' facts ctxt =  | 
|
276  | 
let  | 
|
277  | 
val ctxt = Simplifier.add_prems facts ctxt  | 
|
278  | 
in  | 
|
279  | 
Thm.cterm_of ctxt #> Simplifier.rewrite ctxt #>  | 
|
280  | 
Thm.concl_of #> Logic.dest_equals #> snd  | 
|
281  | 
end  | 
|
282  | 
||
283  | 
fun simplify_term ectxt = simplify_term' (get_facts ectxt) (get_ctxt ectxt)  | 
|
284  | 
||
285  | 
fun simplify_eval ctxt =  | 
|
286  | 
  simplify_term' [] (put_simpset HOL_basic_ss ctxt addsimps @{thms eval_simps})
 | 
|
287  | 
||
288  | 
datatype zeroness = IsZero | IsNonZero | IsPos | IsNeg  | 
|
289  | 
||
290  | 
||
291  | 
(* Caution: The following functions assume that the given expansion is in normal form already  | 
|
292  | 
as far as needed. *)  | 
|
293  | 
||
294  | 
(* Returns the leading coefficient of the given expansion. This coefficient is a multiseries. *)  | 
|
295  | 
fun try_get_coeff expr =  | 
|
296  | 
case expr of  | 
|
| 69597 | 297  | 
Const (\<^const_name>\<open>MS\<close>, _) $ (  | 
298  | 
Const (\<^const_name>\<open>MSLCons\<close>, _) $ (  | 
|
299  | 
Const (\<^const_name>\<open>Pair\<close>, _) $ c $ _) $ _) $ _ =>  | 
|
| 68630 | 300  | 
SOME c  | 
301  | 
| _ => NONE  | 
|
302  | 
||
303  | 
fun get_coeff expr =  | 
|
304  | 
expr |> dest_comb |> fst |> dest_comb |> snd |> dest_comb |> fst |> dest_comb |> snd  | 
|
305  | 
|> dest_comb |> fst |> dest_comb |> snd  | 
|
306  | 
||
307  | 
(* Returns the coefficient of the leading term in the expansion (i.e. a real number) *)  | 
|
308  | 
fun get_lead_coeff expr =  | 
|
309  | 
case try_get_coeff expr of  | 
|
310  | 
NONE => expr  | 
|
311  | 
| SOME c => get_lead_coeff c  | 
|
312  | 
||
313  | 
(* Returns the exponent (w.r.t. the fastest-growing basis element) of the leading term *)  | 
|
314  | 
fun get_exponent expr =  | 
|
315  | 
expr |> dest_comb |> fst |> dest_comb |> snd |> dest_comb |> fst |> dest_comb |> snd  | 
|
316  | 
|> dest_comb |> snd  | 
|
317  | 
||
318  | 
(* Returns the list of exponents of the leading term *)  | 
|
319  | 
fun get_exponents exp =  | 
|
| 69597 | 320  | 
if fastype_of exp = \<^typ>\<open>real\<close> then  | 
| 68630 | 321  | 
[]  | 
322  | 
else  | 
|
323  | 
get_exponent exp :: get_exponents (get_coeff exp)  | 
|
324  | 
||
325  | 
(* Returns the function that the expansion corresponds to *)  | 
|
326  | 
fun get_eval expr =  | 
|
| 69597 | 327  | 
if fastype_of expr = \<^typ>\<open>real\<close> then  | 
328  | 
    Abs ("x", \<^typ>\<open>real\<close>, expr)
 | 
|
| 68630 | 329  | 
else  | 
330  | 
expr |> dest_comb |> snd  | 
|
331  | 
||
332  | 
val eval_simps = @{thms eval_simps [THEN eq_reflection]}
 | 
|
333  | 
||
334  | 
(* Tries to prove that the given function is eventually zero *)  | 
|
335  | 
fun ev_zeroness_oracle ectxt t =  | 
|
336  | 
let  | 
|
337  | 
val ctxt = Lazy_Eval.get_ctxt ectxt  | 
|
338  | 
val goal =  | 
|
| 69597 | 339  | 
betapply (\<^term>\<open>\<lambda>f::real \<Rightarrow> real. eventually (\<lambda>x. f x = 0) at_top\<close>, t)  | 
| 68630 | 340  | 
|> HOLogic.mk_Trueprop  | 
341  | 
    fun tac {context = ctxt, ...} =
 | 
|
342  | 
HEADGOAL (Method.insert_tac ctxt (get_facts ectxt))  | 
|
343  | 
THEN Local_Defs.unfold_tac ctxt eval_simps  | 
|
344  | 
THEN HEADGOAL (Simplifier.asm_full_simp_tac ctxt)  | 
|
345  | 
in  | 
|
346  | 
try (Goal.prove ctxt [] [] goal) tac  | 
|
347  | 
end  | 
|
348  | 
||
349  | 
(*  | 
|
350  | 
Encodes the kind of trimming/zeroness checking operation to be performed.  | 
|
351  | 
Simple_Trim only checks for zeroness/non-zeroness. Pos_Trim/Neg_Trim try to prove either  | 
|
352  | 
zeroness or positivity (resp. negativity). Sgn_Trim tries all three possibilities (positive,  | 
|
353  | 
negative, zero). *)  | 
|
354  | 
datatype trim_mode = Simple_Trim | Pos_Trim | Neg_Trim | Sgn_Trim  | 
|
355  | 
||
356  | 
(*  | 
|
357  | 
Checks (and proves) whether the given term (assumed to be a real number) is zero, positive,  | 
|
358  | 
or negative, depending on given flags. The "fail" flag determines whether an exception is  | 
|
359  | 
thrown if this fails.  | 
|
360  | 
*)  | 
|
361  | 
fun zeroness_oracle fail mode ectxt exp =  | 
|
362  | 
let  | 
|
363  | 
val ctxt = Lazy_Eval.get_ctxt ectxt  | 
|
| 69597 | 364  | 
val eq = (exp, \<^term>\<open>0::real\<close>) |> HOLogic.mk_eq  | 
| 68630 | 365  | 
val goal1 = (IsZero, eq |> HOLogic.mk_Trueprop)  | 
366  | 
val goal2 =  | 
|
367  | 
case mode of  | 
|
368  | 
SOME Pos_Trim =>  | 
|
| 69597 | 369  | 
(IsPos, \<^term>\<open>(<) (0::real)\<close> $ exp |> HOLogic.mk_Trueprop)  | 
| 68630 | 370  | 
| SOME Sgn_Trim =>  | 
| 69597 | 371  | 
(IsPos, \<^term>\<open>(<) (0::real)\<close> $ exp |> HOLogic.mk_Trueprop)  | 
| 68630 | 372  | 
| SOME Neg_Trim =>  | 
| 69597 | 373  | 
(IsNeg, betapply (\<^term>\<open>\<lambda>x. x < (0::real)\<close>, exp) |> HOLogic.mk_Trueprop)  | 
| 68630 | 374  | 
| _ =>  | 
375  | 
(IsNonZero, eq |> HOLogic.mk_not |> HOLogic.mk_Trueprop)  | 
|
376  | 
val goals =  | 
|
377  | 
(if mode = SOME Sgn_Trim then  | 
|
| 69597 | 378  | 
[(IsNeg, betapply (\<^term>\<open>\<lambda>x. x < (0::real)\<close>, exp) |> HOLogic.mk_Trueprop)]  | 
| 68630 | 379  | 
else  | 
380  | 
[])  | 
|
381  | 
val goals = goal2 :: goals  | 
|
382  | 
    fun tac {context = ctxt, ...} =
 | 
|
383  | 
HEADGOAL (Method.insert_tac ctxt (get_facts ectxt))  | 
|
384  | 
THEN Local_Defs.unfold_tac ctxt eval_simps  | 
|
385  | 
THEN HEADGOAL (apply_sign_oracles ctxt (Simplifier.asm_full_simp_tac ctxt))  | 
|
386  | 
fun prove (res, goal) = try (fn goal => (res, SOME (Goal.prove ctxt [] [] goal tac))) goal  | 
|
387  | 
fun err () =  | 
|
388  | 
let  | 
|
389  | 
val mode_msg =  | 
|
390  | 
case mode of  | 
|
391  | 
SOME Simple_Trim => "whether the following constant is zero"  | 
|
392  | 
| SOME Pos_Trim => "whether the following constant is zero or positive"  | 
|
393  | 
| SOME Neg_Trim => "whether the following constant is zero or negative"  | 
|
394  | 
| SOME Sgn_Trim => "the sign of the following constant"  | 
|
395  | 
| _ => raise Match  | 
|
396  | 
val t = simplify_term' (get_facts ectxt) ctxt exp  | 
|
397  | 
val _ =  | 
|
398  | 
if #verbose (#ctxt ectxt) then  | 
|
399  | 
let  | 
|
400  | 
              val p = Pretty.str ("real_asymp failed to determine " ^ mode_msg ^ ":")
 | 
|
401  | 
val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt t)]  | 
|
402  | 
in  | 
|
403  | 
Pretty.writeln p  | 
|
404  | 
end else ()  | 
|
405  | 
in  | 
|
406  | 
        raise TERM ("zeroness_oracle", [t])
 | 
|
407  | 
end  | 
|
408  | 
in  | 
|
409  | 
case prove goal1 of  | 
|
410  | 
SOME res => res  | 
|
411  | 
| NONE =>  | 
|
412  | 
if mode = NONE then  | 
|
413  | 
(IsNonZero, NONE)  | 
|
414  | 
else  | 
|
415  | 
case get_first prove (goal2 :: goals) of  | 
|
416  | 
NONE => if fail then err () else (IsNonZero, NONE)  | 
|
417  | 
| SOME res => res  | 
|
418  | 
end  | 
|
419  | 
||
420  | 
(* Tries to prove a given equality of real numbers. *)  | 
|
421  | 
fun try_prove_real_eq fail ectxt (lhs, rhs) =  | 
|
| 69597 | 422  | 
case zeroness_oracle false NONE ectxt (\<^term>\<open>(-) :: real => _\<close> $ lhs $ rhs) of  | 
| 68630 | 423  | 
    (IsZero, SOME thm) => SOME (thm RS @{thm real_eqI})
 | 
424  | 
| _ =>  | 
|
425  | 
if not fail then NONE else  | 
|
426  | 
let  | 
|
427  | 
val ctxt = get_ctxt ectxt  | 
|
428  | 
val ts = map (simplify_term' (get_facts ectxt) ctxt) [lhs, rhs]  | 
|
429  | 
val _ =  | 
|
430  | 
if #verbose (#ctxt ectxt) then  | 
|
431  | 
let  | 
|
432  | 
val p =  | 
|
433  | 
                Pretty.str ("real_asymp failed to prove that the following two numbers are equal:")
 | 
|
434  | 
val p = Pretty.chunks (p :: map (Pretty.indent 2 o Syntax.pretty_term ctxt) ts)  | 
|
435  | 
in  | 
|
436  | 
Pretty.writeln p  | 
|
437  | 
end else ()  | 
|
438  | 
in  | 
|
439  | 
        raise TERM ("try_prove_real_eq", [lhs, rhs])
 | 
|
440  | 
end  | 
|
441  | 
||
442  | 
(* Tries to prove a given eventual equality of real functions. *)  | 
|
443  | 
fun try_prove_ev_eq ectxt (f, g) =  | 
|
444  | 
let  | 
|
| 69597 | 445  | 
val t = Envir.beta_eta_contract (\<^term>\<open>\<lambda>(f::real=>real) g x. f x - g x\<close> $ f $ g)  | 
| 68630 | 446  | 
in  | 
447  | 
    Option.map (fn thm => thm RS @{thm eventually_diff_zero_imp_eq}) (ev_zeroness_oracle ectxt t)
 | 
|
448  | 
end  | 
|
449  | 
||
| 69597 | 450  | 
fun real_less a b = \<^term>\<open>(<) :: real \<Rightarrow> real \<Rightarrow> bool\<close> $ a $ b  | 
451  | 
fun real_eq a b = \<^term>\<open>(=) :: real \<Rightarrow> real \<Rightarrow> bool\<close> $ a $ b  | 
|
452  | 
fun real_neq a b = \<^term>\<open>(\<noteq>) :: real \<Rightarrow> real \<Rightarrow> bool\<close> $ a $ b  | 
|
| 68630 | 453  | 
|
454  | 
(* The hook that is called by the Lazy_Eval module whenever two real numbers have to be compared *)  | 
|
455  | 
fun real_sgn_hook ({pctxt = ctxt, facts, verbose, ...}) t =
 | 
|
456  | 
let  | 
|
457  | 
val get_rhs = Thm.concl_of #> Logic.dest_equals #> snd  | 
|
458  | 
    fun tac {context = ctxt, ...} = 
 | 
|
459  | 
HEADGOAL (Method.insert_tac ctxt (Net.content facts)  | 
|
460  | 
THEN' (apply_sign_oracles ctxt (Simplifier.asm_full_simp_tac ctxt)))  | 
|
461  | 
fun prove_first err [] [] =  | 
|
462  | 
          if not verbose then raise TERM ("real_sgn_hook", [t])
 | 
|
463  | 
            else let val _ = err () in raise TERM ("real_sgn_hook", [t]) end
 | 
|
464  | 
| prove_first err (goal :: goals) (thm :: thms) =  | 
|
465  | 
(case try (Goal.prove ctxt [] [] goal) tac of  | 
|
466  | 
SOME thm' =>  | 
|
467  | 
let val thm'' = thm' RS thm in SOME (get_rhs thm'', Conv.rewr_conv thm'') end  | 
|
468  | 
| NONE => prove_first err goals thms)  | 
|
469  | 
| prove_first _ _ _ = raise Match  | 
|
470  | 
in  | 
|
471  | 
case t of  | 
|
| 69597 | 472  | 
\<^term>\<open>(=) :: real => _\<close> $ a $ \<^term>\<open>0 :: real\<close> =>  | 
| 68630 | 473  | 
let  | 
474  | 
val goals =  | 
|
| 69597 | 475  | 
map (fn c => HOLogic.mk_Trueprop (c a \<^term>\<open>0 :: real\<close>)) [real_neq, real_eq]  | 
| 68630 | 476  | 
fun err () =  | 
477  | 
let  | 
|
478  | 
val facts' = Net.content facts  | 
|
479  | 
val a' = simplify_term' facts' ctxt a  | 
|
480  | 
              val p = Pretty.str ("real_asymp failed to determine whether the following " ^
 | 
|
481  | 
"constant is zero: ")  | 
|
482  | 
val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt a')]  | 
|
483  | 
in  | 
|
484  | 
Pretty.writeln p  | 
|
485  | 
end  | 
|
486  | 
in  | 
|
487  | 
          prove_first err goals @{thms Eq_FalseI Eq_TrueI}
 | 
|
488  | 
end  | 
|
| 69597 | 489  | 
| Const (\<^const_name>\<open>COMPARE\<close>, _) $ a $ b =>  | 
| 68630 | 490  | 
let  | 
491  | 
val goals = map HOLogic.mk_Trueprop [real_less a b, real_less b a, real_eq a b]  | 
|
492  | 
fun err () =  | 
|
493  | 
let  | 
|
494  | 
val facts' = Net.content facts  | 
|
495  | 
val (a', b') = apply2 (simplify_term' facts' ctxt) (a, b)  | 
|
496  | 
              val p = Pretty.str ("real_asymp failed to compare" ^
 | 
|
497  | 
"the following two constants: ")  | 
|
498  | 
val p = Pretty.chunks (p :: map (Pretty.indent 2 o Syntax.pretty_term ctxt) [a', b'])  | 
|
499  | 
in  | 
|
500  | 
Pretty.writeln p  | 
|
501  | 
end  | 
|
502  | 
in  | 
|
503  | 
          prove_first err goals @{thms COMPARE_intros}
 | 
|
504  | 
end  | 
|
505  | 
| _ => NONE  | 
|
506  | 
end  | 
|
507  | 
||
508  | 
(*  | 
|
509  | 
Returns the datatype constructors registered for use with the Lazy_Eval package.  | 
|
510  | 
All constructors on which pattern matching is performed need to be registered for evaluation  | 
|
511  | 
to work. It should be rare for users to add additional ones.  | 
|
512  | 
*)  | 
|
513  | 
fun get_constructors ctxt =  | 
|
514  | 
let  | 
|
| 69597 | 515  | 
val thms = Named_Theorems.get ctxt \<^named_theorems>\<open>exp_log_eval_constructor\<close>  | 
| 68630 | 516  | 
fun go _ [] acc = rev acc  | 
517  | 
| go f (x :: xs) acc =  | 
|
518  | 
case f x of  | 
|
519  | 
NONE => go f xs acc  | 
|
520  | 
| SOME y => go f xs (y :: acc)  | 
|
521  | 
fun map_option f xs = go f xs []  | 
|
522  | 
fun dest_constructor thm =  | 
|
523  | 
case Thm.concl_of thm of  | 
|
| 69597 | 524  | 
Const (\<^const_name>\<open>HOL.Trueprop\<close>, _) $  | 
525  | 
(Const (\<^const_name>\<open>REAL_ASYMP_EVAL_CONSTRUCTOR\<close>, _) $ Const (c, T)) =>  | 
|
| 68630 | 526  | 
SOME (c, length (fst (strip_type T)))  | 
527  | 
| _ => NONE  | 
|
528  | 
in  | 
|
529  | 
thms |> map_option dest_constructor  | 
|
530  | 
end  | 
|
531  | 
||
532  | 
(*  | 
|
533  | 
Creates an evaluation context with the correct setup of constructors, equations, and hooks.  | 
|
534  | 
*)  | 
|
535  | 
fun mk_eval_ctxt ctxt =  | 
|
536  | 
let  | 
|
| 69597 | 537  | 
val eval_eqs = (Named_Theorems.get ctxt \<^named_theorems>\<open>real_asymp_eval_eqs\<close>)  | 
| 68630 | 538  | 
val constructors = get_constructors ctxt  | 
539  | 
in  | 
|
540  | 
Lazy_Eval.mk_eval_ctxt ctxt constructors eval_eqs  | 
|
541  | 
|> add_hook real_sgn_hook  | 
|
542  | 
end  | 
|
543  | 
||
544  | 
(* A pattern for determining the leading coefficient of a multiseries *)  | 
|
545  | 
val exp_pat =  | 
|
546  | 
let  | 
|
547  | 
    val anypat = AnyPat ("_", 0)
 | 
|
548  | 
in  | 
|
| 69597 | 549  | 
ConsPat (\<^const_name>\<open>MS\<close>,  | 
550  | 
[ConsPat (\<^const_name>\<open>MSLCons\<close>,  | 
|
551  | 
[ConsPat (\<^const_name>\<open>Pair\<close>, [anypat, anypat]), anypat]), anypat])  | 
|
| 68630 | 552  | 
end  | 
553  | 
||
554  | 
(*  | 
|
555  | 
Evaluates an expansion to (weak) head normal form, so that the leading coefficient and  | 
|
556  | 
exponent can be read off.  | 
|
557  | 
*)  | 
|
558  | 
fun whnf_expansion ectxt thm =  | 
|
559  | 
let  | 
|
560  | 
val ctxt = get_ctxt ectxt  | 
|
561  | 
val exp = get_expansion thm  | 
|
562  | 
val (_, _, conv) = match ectxt exp_pat exp (SOME [])  | 
|
563  | 
val eq_thm = conv (Thm.cterm_of ctxt exp)  | 
|
564  | 
val exp' = eq_thm |> Thm.concl_of |> Logic.dest_equals |> snd  | 
|
565  | 
in  | 
|
566  | 
case exp' of  | 
|
| 69597 | 567  | 
Const (\<^const_name>\<open>MS\<close>, _) $ (Const (\<^const_name>\<open>MSLCons\<close>, _) $  | 
568  | 
(Const (\<^const_name>\<open>Pair\<close>, _) $ c $ _) $ _) $ _ =>  | 
|
| 68630 | 569  | 
          (SOME c, @{thm expands_to_meta_eq_cong} OF [thm, eq_thm], eq_thm)
 | 
| 69597 | 570  | 
| Const (\<^const_name>\<open>MS\<close>, _) $ Const (\<^const_name>\<open>MSLNil\<close>, _) $ _ =>  | 
| 68630 | 571  | 
        (NONE, @{thm expands_to_meta_eq_cong} OF [thm, eq_thm], eq_thm)
 | 
572  | 
    | _ => raise TERM ("whnf_expansion", [exp'])
 | 
|
573  | 
end  | 
|
574  | 
||
575  | 
fun try_lift_function ectxt (thm, SEmpty) _ = (NONE, thm)  | 
|
576  | 
| try_lift_function ectxt (thm, basis) cont =  | 
|
577  | 
case whnf_expansion ectxt thm of  | 
|
578  | 
(SOME c, thm, _) =>  | 
|
579  | 
let  | 
|
580  | 
val f = get_expanded_fun thm  | 
|
581  | 
val T = fastype_of c  | 
|
| 69597 | 582  | 
val t = Const (\<^const_name>\<open>eval\<close>, T --> \<^typ>\<open>real \<Rightarrow> real\<close>) $ c  | 
583  | 
val t = Term.betapply (Term.betapply (\<^term>\<open>\<lambda>(f::real\<Rightarrow>real) g x. f x - g x\<close>, f), t)  | 
|
| 68630 | 584  | 
in  | 
585  | 
case ev_zeroness_oracle ectxt t of  | 
|
586  | 
NONE => (NONE, thm)  | 
|
587  | 
| SOME zero_thm =>  | 
|
588  | 
let  | 
|
589  | 
              val thm' = cont ectxt (thm RS @{thm expands_to_hd''}, tl_basis basis)
 | 
|
590  | 
              val thm'' = @{thm expands_to_lift_function} OF [zero_thm, thm']
 | 
|
591  | 
in  | 
|
592  | 
(SOME (lift basis thm''), thm)  | 
|
593  | 
end  | 
|
594  | 
end  | 
|
595  | 
| _ => (NONE, thm)  | 
|
596  | 
||
597  | 
(* Turns an expansion theorem into an expansion theorem for the leading coefficient. *)  | 
|
598  | 
fun expands_to_hd thm = thm RS  | 
|
| 69597 | 599  | 
(if fastype_of (get_expansion thm) = \<^typ>\<open>real ms\<close> then  | 
| 68630 | 600  | 
     @{thm expands_to_hd'}
 | 
601  | 
else  | 
|
602  | 
     @{thm expands_to_hd})
 | 
|
603  | 
||
604  | 
fun simplify_expansion ectxt thm =  | 
|
605  | 
let  | 
|
606  | 
val exp = get_expansion thm  | 
|
607  | 
val ctxt = get_ctxt ectxt  | 
|
608  | 
val eq_thm = Simplifier.rewrite ctxt (Thm.cterm_of ctxt exp)  | 
|
609  | 
in  | 
|
610  | 
    @{thm expands_to_meta_eq_cong} OF [thm, eq_thm]
 | 
|
611  | 
end  | 
|
612  | 
||
613  | 
(*  | 
|
614  | 
Simplifies a trimmed expansion and returns the simplified expansion theorem and  | 
|
615  | 
the trimming theorem for that simplified expansion.  | 
|
616  | 
*)  | 
|
617  | 
fun simplify_trimmed_expansion ectxt (thm, trimmed_thm) =  | 
|
618  | 
let  | 
|
619  | 
val exp = get_expansion thm  | 
|
620  | 
val ctxt = get_ctxt ectxt  | 
|
621  | 
val eq_thm = Simplifier.rewrite ctxt (Thm.cterm_of ctxt exp)  | 
|
622  | 
val trimmed_cong_thm =  | 
|
623  | 
case trimmed_thm |> concl_of' |> dest_fun of  | 
|
| 69597 | 624  | 
        Const (\<^const_name>\<open>trimmed\<close>, _) => @{thm trimmed_eq_cong}
 | 
625  | 
      | Const (\<^const_name>\<open>trimmed_pos\<close>, _) => @{thm trimmed_pos_eq_cong}
 | 
|
626  | 
      | Const (\<^const_name>\<open>trimmed_neg\<close>, _) => @{thm trimmed_neg_eq_cong}
 | 
|
| 68630 | 627  | 
      | _ => raise THM ("simplify_trimmed_expansion", 2, [thm, trimmed_thm])
 | 
628  | 
in  | 
|
629  | 
    (@{thm expands_to_meta_eq_cong} OF [thm, eq_thm], 
 | 
|
630  | 
trimmed_cong_thm OF [trimmed_thm, eq_thm])  | 
|
631  | 
end  | 
|
632  | 
||
633  | 
(*  | 
|
634  | 
Re-normalises a trimmed expansion (so that the leading term with its (real) coefficient and  | 
|
635  | 
all exponents can be read off. This may be necessary after lifting a trimmed expansion to  | 
|
636  | 
a larger basis.  | 
|
637  | 
*)  | 
|
638  | 
fun retrim_expansion ectxt (thm, basis) =  | 
|
639  | 
let  | 
|
640  | 
val (c, thm, eq_thm) = whnf_expansion ectxt thm  | 
|
641  | 
in  | 
|
642  | 
case c of  | 
|
643  | 
NONE => (thm, eq_thm)  | 
|
644  | 
| SOME c =>  | 
|
| 69597 | 645  | 
if fastype_of c = \<^typ>\<open>real\<close> then  | 
| 68630 | 646  | 
(thm, eq_thm)  | 
647  | 
else  | 
|
648  | 
let  | 
|
649  | 
          val c_thm = thm RS @{thm expands_to_hd''}
 | 
|
650  | 
val (c_thm', eq_thm') = retrim_expansion ectxt (c_thm, tl_basis basis)  | 
|
651  | 
          val thm = @{thm expands_to_trim_cong} OF [thm, c_thm']
 | 
|
652  | 
in  | 
|
653  | 
          (thm, @{thm trim_lift_eq} OF [eq_thm, eq_thm'])
 | 
|
654  | 
end  | 
|
655  | 
end  | 
|
656  | 
||
657  | 
fun retrim_pos_expansion ectxt (thm, basis, trimmed_thm) =  | 
|
658  | 
let  | 
|
659  | 
val (thm', eq_thm) = retrim_expansion ectxt (thm, basis)  | 
|
660  | 
in  | 
|
661  | 
    (thm', eq_thm, @{thm trimmed_pos_eq_cong} OF [trimmed_thm, eq_thm])
 | 
|
662  | 
end  | 
|
663  | 
||
664  | 
(*  | 
|
665  | 
Tries to determine whether the leading term is (identically) zero and drops it if it is.  | 
|
666  | 
If "fail" is set, an exception is thrown when that term is a real number and zeroness cannot  | 
|
667  | 
be determined. (Which typically indicates missing facts or case distinctions)  | 
|
668  | 
*)  | 
|
669  | 
fun try_drop_leading_term_ex fail ectxt thm =  | 
|
670  | 
let  | 
|
671  | 
val exp = get_expansion thm  | 
|
672  | 
in  | 
|
| 69597 | 673  | 
if fastype_of exp = \<^typ>\<open>real\<close> then  | 
| 68630 | 674  | 
NONE  | 
| 69597 | 675  | 
else if fastype_of (get_coeff exp) = \<^typ>\<open>real\<close> then  | 
| 68630 | 676  | 
case zeroness_oracle fail (SOME Simple_Trim) ectxt (get_coeff exp) of  | 
677  | 
        (IsZero, SOME zero_thm) => SOME (@{thm drop_zero_ms'} OF [zero_thm, thm])
 | 
|
678  | 
| _ => NONE  | 
|
679  | 
else  | 
|
680  | 
let  | 
|
681  | 
val c = get_coeff exp  | 
|
682  | 
val T = fastype_of c  | 
|
| 69597 | 683  | 
val t = Const (\<^const_name>\<open>eval\<close>, T --> \<^typ>\<open>real \<Rightarrow> real\<close>) $ c  | 
| 68630 | 684  | 
in  | 
685  | 
case ev_zeroness_oracle ectxt t of  | 
|
686  | 
          SOME zero_thm => SOME (@{thm expands_to_drop_zero} OF [zero_thm, thm])
 | 
|
687  | 
| _ => NONE  | 
|
688  | 
end  | 
|
689  | 
end  | 
|
690  | 
||
691  | 
(*  | 
|
692  | 
Tries to drop the leading term of an expansion. If this is not possible, an exception  | 
|
693  | 
is thrown and an informative error message is printed.  | 
|
694  | 
*)  | 
|
695  | 
fun try_drop_leading_term ectxt thm =  | 
|
696  | 
let  | 
|
697  | 
fun err () =  | 
|
698  | 
let  | 
|
699  | 
val ctxt = get_ctxt ectxt  | 
|
700  | 
val exp = get_expansion thm  | 
|
701  | 
val c = get_coeff exp  | 
|
702  | 
val t =  | 
|
| 69597 | 703  | 
if fastype_of c = \<^typ>\<open>real\<close> then c else c |> dest_arg  | 
| 68630 | 704  | 
val t = simplify_term' (get_facts ectxt) ctxt t  | 
705  | 
val _ =  | 
|
706  | 
if #verbose (#ctxt ectxt) then  | 
|
707  | 
let  | 
|
708  | 
              val p = Pretty.str ("real_asymp failed to prove that the following term is zero: ")
 | 
|
709  | 
val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt t)]  | 
|
710  | 
in  | 
|
711  | 
Pretty.writeln p  | 
|
712  | 
end else ()  | 
|
713  | 
in  | 
|
714  | 
        raise TERM ("try_drop_leading_term", [t])
 | 
|
715  | 
end  | 
|
716  | 
in  | 
|
717  | 
case try_drop_leading_term_ex true ectxt thm of  | 
|
718  | 
NONE => err ()  | 
|
719  | 
| SOME thm => thm  | 
|
720  | 
end  | 
|
721  | 
||
722  | 
||
723  | 
datatype trim_result =  | 
|
724  | 
Trimmed of zeroness * trimmed_thm option  | 
|
725  | 
| Aborted of order  | 
|
726  | 
||
727  | 
fun cstrip_assms ct =  | 
|
728  | 
case Thm.term_of ct of  | 
|
| 69597 | 729  | 
\<^term>\<open>(==>)\<close> $ _ $ _ => cstrip_assms (snd (Thm.dest_implies ct))  | 
| 68630 | 730  | 
| _ => ct  | 
731  | 
||
732  | 
(*  | 
|
733  | 
Trims an expansion (i.e. drops leading zero terms) and provides a trimmedness theorem.  | 
|
734  | 
Optionally, a list of exponents can be given to instruct the function to only trim until  | 
|
735  | 
the exponents of the leading term are lexicographically less than (or less than or equal) than  | 
|
736  | 
the given ones. This is useful to avoid unnecessary trimming.  | 
|
737  | 
||
738  | 
The "strict" flag indicates whether the trimming should already be aborted when the  | 
|
739  | 
exponents are lexicographically equal or not.  | 
|
740  | 
||
741  | 
The "fail" flag is passed on to the zeroness oracle and determines whether a failure to determine  | 
|
742  | 
the sign of a real number leads to an exception.  | 
|
743  | 
||
744  | 
"mode" indicates what kind of trimmedness theorem will be returned: Simple_Trim only gives the  | 
|
745  | 
default trimmedness theorem, whereas Pos_Trim/Neg_Trim/Sgn_Trim will give trimmed_pos or  | 
|
746  | 
trimmed_neg. Giving "None" as mode will produce no trimmedness theorem; it will only drop  | 
|
747  | 
leading zero terms until zeroness cannot be proven anymore, upon which it will stop.  | 
|
748  | 
||
749  | 
The main result of the function is the trimmed expansion theorem.  | 
|
750  | 
||
751  | 
The function returns whether the trimming has been aborted or not. If was aborted, either  | 
|
752  | 
LESS or EQUAL will be returned, indicating whether the exponents of the leading term are  | 
|
753  | 
now lexicographically smaller or equal to the given ones. In the other case, the zeroness  | 
|
754  | 
of the leading coefficient is returned (zero, non-zero, positive, negative) together with a  | 
|
755  | 
trimmedness theorem.  | 
|
756  | 
||
757  | 
Lastly, a list of the exponent comparison results and associated theorems is also returned, so  | 
|
758  | 
that the caller can reconstruct the result of the lexicographic ordering without doing the  | 
|
759  | 
exponent comparisons again.  | 
|
760  | 
*)  | 
|
761  | 
fun trim_expansion_while_greater strict es fail mode ectxt (thm, basis) =  | 
|
762  | 
let  | 
|
763  | 
val (_, thm, _) = whnf_expansion ectxt thm  | 
|
764  | 
val thm = simplify_expansion ectxt thm  | 
|
765  | 
val cexp = thm |> Thm.cprop_of |> cstrip_assms |> Thm.dest_arg |> Thm.dest_fun |> Thm.dest_arg  | 
|
766  | 
val c = try_get_coeff (get_expansion thm)  | 
|
767  | 
fun lift_trimmed_thm nz thm =  | 
|
768  | 
let  | 
|
769  | 
val cexp = thm |> Thm.cprop_of |> cstrip_assms |> Thm.dest_arg |> Thm.dest_fun |> Thm.dest_arg  | 
|
770  | 
val lift_thm =  | 
|
771  | 
case nz of  | 
|
772  | 
            IsNonZero => @{thm trimmed_eq_cong[rotated, OF _ lift_trimmed]}
 | 
|
773  | 
          | IsPos => @{thm trimmed_pos_eq_cong[rotated, OF _ lift_trimmed_pos]}
 | 
|
774  | 
          | IsNeg => @{thm trimmed_neg_eq_cong[rotated, OF _ lift_trimmed_neg]}
 | 
|
775  | 
          | _ => raise TERM ("Unexpected zeroness result in trim_expansion", [])
 | 
|
776  | 
in  | 
|
777  | 
Thm.reflexive cexp RS lift_thm  | 
|
778  | 
end  | 
|
779  | 
fun trimmed_real_thm nz = Thm.reflexive cexp RS (  | 
|
780  | 
case nz of  | 
|
781  | 
        IsNonZero => @{thm trimmed_eq_cong[rotated, OF _ lift_trimmed[OF trimmed_realI]]}
 | 
|
782  | 
      | IsPos => @{thm trimmed_pos_eq_cong[rotated, OF _ lift_trimmed_pos[OF trimmed_pos_realI]]}
 | 
|
783  | 
      | IsNeg => @{thm trimmed_neg_eq_cong[rotated, OF _ lift_trimmed_neg[OF trimmed_neg_realI]]}
 | 
|
784  | 
      | _ => raise TERM ("Unexpected zeroness result in trim_expansion", []))
 | 
|
785  | 
fun do_trim es =  | 
|
786  | 
let  | 
|
787  | 
val c = the c  | 
|
788  | 
val T = fastype_of c  | 
|
| 69597 | 789  | 
val t = Const (\<^const_name>\<open>eval\<close>, T --> \<^typ>\<open>real \<Rightarrow> real\<close>) $ c  | 
| 68630 | 790  | 
in  | 
| 69597 | 791  | 
if T = \<^typ>\<open>real\<close> then (  | 
| 68630 | 792  | 
case zeroness_oracle fail mode ectxt c of  | 
793  | 
(IsZero, SOME zero_thm) =>  | 
|
794  | 
trim_expansion_while_greater strict es fail mode ectxt  | 
|
795  | 
                (@{thm drop_zero_ms'} OF [zero_thm, thm], basis)
 | 
|
796  | 
| (nz, SOME nz_thm) => (thm, Trimmed (nz, SOME (nz_thm RS trimmed_real_thm nz)), [])  | 
|
797  | 
| (nz, NONE) => (thm, Trimmed (nz, NONE), []))  | 
|
798  | 
else  | 
|
799  | 
case trim_expansion_while_greater strict (Option.map tl es) fail mode ectxt  | 
|
800  | 
                 (thm RS @{thm expands_to_hd''}, tl_basis basis) of
 | 
|
801  | 
(c_thm', Aborted ord, thms) =>  | 
|
802  | 
              (@{thm expands_to_trim_cong} OF [thm, c_thm'], Aborted ord, thms)
 | 
|
803  | 
| (c_thm', Trimmed (nz, trimmed_thm), thms) =>  | 
|
804  | 
let  | 
|
805  | 
                val thm = (@{thm expands_to_trim_cong} OF [thm, c_thm'])
 | 
|
806  | 
fun err () =  | 
|
807  | 
                  raise TERM ("trim_expansion: zero coefficient should have been trimmed", [c])
 | 
|
808  | 
in  | 
|
809  | 
case (nz, trimmed_thm) of  | 
|
810  | 
(IsZero, _) =>  | 
|
811  | 
if #verbose (#ctxt ectxt) then  | 
|
812  | 
let  | 
|
813  | 
val ctxt = get_ctxt ectxt  | 
|
814  | 
val t' = t |> simplify_eval ctxt |> simplify_term' (get_facts ectxt) ctxt  | 
|
815  | 
                        val p = Pretty.str ("trim_expansion failed to recognise zeroness of " ^
 | 
|
816  | 
"the following term:")  | 
|
817  | 
val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt t')]  | 
|
818  | 
val _ = Pretty.writeln p  | 
|
819  | 
in  | 
|
820  | 
err ()  | 
|
821  | 
end  | 
|
822  | 
else err ()  | 
|
823  | 
| (_, SOME trimmed_thm) =>  | 
|
824  | 
(thm, Trimmed (nz, SOME (trimmed_thm RS lift_trimmed_thm nz thm)), thms)  | 
|
825  | 
| (_, NONE) => (thm, Trimmed (nz, NONE), thms)  | 
|
826  | 
end  | 
|
827  | 
end  | 
|
| 69597 | 828  | 
val minus = \<^term>\<open>(-) :: real => real => real\<close>  | 
| 68630 | 829  | 
in  | 
830  | 
case (c, es) of  | 
|
831  | 
(NONE, _) => (thm, Trimmed (IsZero, NONE), [])  | 
|
832  | 
| (SOME c, SOME (e' :: _)) =>  | 
|
833  | 
let  | 
|
834  | 
val e = get_exponent (get_expansion thm)  | 
|
835  | 
in  | 
|
836  | 
case zeroness_oracle true (SOME Sgn_Trim) ectxt (minus $ e $ e') of  | 
|
837  | 
(IsPos, SOME pos_thm) => (  | 
|
838  | 
case try_drop_leading_term_ex false ectxt thm of  | 
|
839  | 
SOME thm =>  | 
|
840  | 
trim_expansion_while_greater strict es fail mode ectxt (thm, basis)  | 
|
841  | 
              | NONE => do_trim NONE |> @{apply 3(3)} (fn thms => (IsPos, pos_thm) :: thms))
 | 
|
842  | 
| (IsNeg, SOME neg_thm) => (thm, Aborted LESS, [(IsNeg, neg_thm)])  | 
|
843  | 
| (IsZero, SOME zero_thm) =>  | 
|
| 69597 | 844  | 
if not strict andalso fastype_of c = \<^typ>\<open>real\<close> then  | 
| 68630 | 845  | 
(thm, Aborted EQUAL, [(IsZero, zero_thm)])  | 
846  | 
else (  | 
|
847  | 
case try_drop_leading_term_ex false ectxt thm of  | 
|
848  | 
SOME thm => trim_expansion_while_greater strict es fail mode ectxt (thm, basis)  | 
|
849  | 
                | NONE => (do_trim es |> @{apply 3(3)} (fn thms => (IsZero, zero_thm) :: thms)))
 | 
|
850  | 
| _ => do_trim NONE  | 
|
851  | 
end  | 
|
852  | 
| _ => (  | 
|
853  | 
case try_drop_leading_term_ex false ectxt thm of  | 
|
854  | 
SOME thm => trim_expansion_while_greater strict es fail mode ectxt (thm, basis)  | 
|
855  | 
| NONE => do_trim NONE)  | 
|
856  | 
end  | 
|
857  | 
||
858  | 
(*  | 
|
859  | 
Trims an expansion without any stopping criterion.  | 
|
860  | 
*)  | 
|
861  | 
fun trim_expansion fail mode ectxt (thm, basis) =  | 
|
862  | 
case trim_expansion_while_greater false NONE fail mode ectxt (thm, basis) of  | 
|
863  | 
(thm, Trimmed (zeroness, trimmed_thm), _) => (thm, zeroness, trimmed_thm)  | 
|
864  | 
| _ => raise Match  | 
|
865  | 
||
866  | 
(*  | 
|
867  | 
Determines the sign of an expansion that has already been trimmed.  | 
|
868  | 
*)  | 
|
869  | 
fun determine_trimmed_sgn ectxt exp =  | 
|
| 69597 | 870  | 
if fastype_of exp = \<^typ>\<open>real\<close> then  | 
| 68630 | 871  | 
(case zeroness_oracle true (SOME Sgn_Trim) ectxt exp of  | 
872  | 
       (IsPos, SOME thm) => (IsPos, thm RS @{thm trimmed_pos_realI})
 | 
|
873  | 
     | (IsNeg, SOME thm) => (IsNeg, thm RS @{thm trimmed_neg_realI})
 | 
|
874  | 
     | _ => raise TERM ("determine_trimmed_sgn", []))
 | 
|
875  | 
else  | 
|
876  | 
let  | 
|
877  | 
val ct = Thm.cterm_of (get_ctxt ectxt) exp  | 
|
878  | 
in  | 
|
879  | 
(case determine_trimmed_sgn ectxt (get_coeff exp) of  | 
|
880  | 
         (IsPos, thm) => (IsPos, @{thm lift_trimmed_pos'} OF [thm, Thm.reflexive ct])
 | 
|
881  | 
       | (IsNeg, thm) => (IsNeg, @{thm lift_trimmed_neg'} OF [thm, Thm.reflexive ct])
 | 
|
882  | 
       | _ => raise TERM ("determine_trimmed_sgn", []))
 | 
|
883  | 
end  | 
|
884  | 
||
885  | 
fun mk_compare_expansions_const T =  | 
|
| 69597 | 886  | 
Const (\<^const_name>\<open>compare_expansions\<close>,  | 
887  | 
T --> T --> \<^typ>\<open>cmp_result \<times> real \<times> real\<close>)  | 
|
| 68630 | 888  | 
|
889  | 
datatype comparison_result =  | 
|
890  | 
Cmp_Dominated of order * thm list * zeroness * trimmed_thm * expansion_thm * expansion_thm  | 
|
891  | 
| Cmp_Asymp_Equiv of thm * thm  | 
|
892  | 
||
893  | 
fun compare_expansions' _ (thm1, thm2, SEmpty) = Cmp_Asymp_Equiv (thm1, thm2)  | 
|
894  | 
| compare_expansions' ectxt (thm1, thm2, basis) =  | 
|
895  | 
let  | 
|
896  | 
fun lift_trimmed_thm nz =  | 
|
897  | 
case nz of  | 
|
898  | 
        IsPos => @{thm lift_trimmed_pos}
 | 
|
899  | 
      | IsNeg => @{thm lift_trimmed_neg}
 | 
|
900  | 
      | _ => raise TERM ("Unexpected zeroness result in compare_expansions'", [])
 | 
|
901  | 
val (e1, e2) = apply2 (get_expansion #> get_exponent) (thm1, thm2)  | 
|
| 69597 | 902  | 
val e = \<^term>\<open>(-) :: real => _\<close> $ e1 $ e2  | 
| 68630 | 903  | 
fun trim thm = trim_expansion true (SOME Sgn_Trim) ectxt (thm, basis)  | 
904  | 
val try_drop = Option.map (whnf_expansion ectxt #> #2) o try_drop_leading_term_ex false ectxt  | 
|
905  | 
fun handle_result ord zeroness trimmed_thm thm1 thm2 =  | 
|
906  | 
let  | 
|
907  | 
val (e1, e2) = apply2 (get_expansion #> get_exponent) (thm1, thm2)  | 
|
| 69597 | 908  | 
val e = \<^term>\<open>(-) :: real => _\<close> $ e1 $ e2  | 
| 68630 | 909  | 
val mode = if ord = LESS then Neg_Trim else Pos_Trim  | 
910  | 
in  | 
|
911  | 
case zeroness_oracle true (SOME mode) ectxt e of  | 
|
912  | 
(_, SOME e_thm) => Cmp_Dominated (ord, [e_thm], zeroness, trimmed_thm, thm1, thm2)  | 
|
913  | 
| _ => raise Match  | 
|
914  | 
end  | 
|
915  | 
fun recurse e_zero_thm =  | 
|
916  | 
case basis of  | 
|
917  | 
SNE (SSng _) => Cmp_Asymp_Equiv (thm1, thm2)  | 
|
918  | 
| _ =>  | 
|
919  | 
let  | 
|
920  | 
          val (thm1', thm2') = apply2 (fn thm => thm RS @{thm expands_to_hd''}) (thm1, thm2)
 | 
|
921  | 
val (thm1', thm2') = apply2 (whnf_expansion ectxt #> #2) (thm1', thm2')  | 
|
922  | 
in  | 
|
923  | 
case compare_expansions' ectxt (thm1', thm2', tl_basis basis) of  | 
|
924  | 
Cmp_Dominated (order, e_thms, zeroness, trimmed_thm, thm1', thm2') =>  | 
|
925  | 
Cmp_Dominated (order, e_zero_thm :: e_thms, zeroness,  | 
|
926  | 
trimmed_thm RS lift_trimmed_thm zeroness,  | 
|
927  | 
                @{thm expands_to_trim_cong} OF [thm1, thm1'],
 | 
|
928  | 
                @{thm expands_to_trim_cong} OF [thm2, thm2'])
 | 
|
929  | 
| Cmp_Asymp_Equiv (thm1', thm2') => Cmp_Asymp_Equiv  | 
|
930  | 
              (@{thm expands_to_trim_cong} OF [thm1, thm1'],
 | 
|
931  | 
                @{thm expands_to_trim_cong} OF [thm2, thm2'])
 | 
|
932  | 
end  | 
|
933  | 
in  | 
|
934  | 
case zeroness_oracle false (SOME Sgn_Trim) ectxt e of  | 
|
935  | 
(IsPos, SOME _) => (  | 
|
936  | 
case try_drop thm1 of  | 
|
937  | 
SOME thm1 => compare_expansions' ectxt (thm1, thm2, basis)  | 
|
938  | 
| NONE => (  | 
|
939  | 
case trim thm1 of  | 
|
940  | 
(thm1, zeroness, SOME trimmed_thm) =>  | 
|
941  | 
handle_result GREATER zeroness trimmed_thm thm1 thm2  | 
|
942  | 
            | _ => raise TERM ("compare_expansions", map get_expansion [thm1, thm2])))
 | 
|
943  | 
| (IsNeg, SOME _) => (  | 
|
944  | 
case try_drop thm2 of  | 
|
945  | 
SOME thm2 => compare_expansions' ectxt (thm1, thm2, basis)  | 
|
946  | 
| NONE => (  | 
|
947  | 
case trim thm2 of  | 
|
948  | 
(thm2, zeroness, SOME trimmed_thm) =>  | 
|
949  | 
handle_result LESS zeroness trimmed_thm thm1 thm2  | 
|
950  | 
            | _ => raise TERM ("compare_expansions", map get_expansion [thm1, thm2])))
 | 
|
951  | 
| (IsZero, SOME e_zero_thm) => (  | 
|
952  | 
case try_drop thm1 of  | 
|
953  | 
SOME thm1 => compare_expansions' ectxt (thm1, thm2, basis)  | 
|
954  | 
| NONE => (  | 
|
955  | 
case try_drop thm2 of  | 
|
956  | 
SOME thm2 => compare_expansions' ectxt (thm1, thm2, basis)  | 
|
957  | 
| NONE => recurse e_zero_thm))  | 
|
958  | 
| _ =>  | 
|
959  | 
case try_drop thm1 of  | 
|
960  | 
SOME thm1 => compare_expansions' ectxt (thm1, thm2, basis)  | 
|
961  | 
| NONE => (  | 
|
962  | 
case try_drop thm2 of  | 
|
963  | 
SOME thm2 => compare_expansions' ectxt (thm1, thm2, basis)  | 
|
964  | 
            | NONE => raise TERM ("compare_expansions", [e1, e2]))
 | 
|
965  | 
end  | 
|
966  | 
||
967  | 
(* Uses a list of exponent comparison results to show that compare_expansions has a given result.*)  | 
|
968  | 
fun prove_compare_expansions ord [thm] = (  | 
|
969  | 
case ord of  | 
|
970  | 
        LESS => @{thm compare_expansions_LT_I} OF [thm]
 | 
|
971  | 
      | GREATER => @{thm compare_expansions_GT_I} OF [thm]
 | 
|
972  | 
      | EQUAL => @{thm compare_expansions_same_exp[OF _ compare_expansions_real]} OF [thm])
 | 
|
973  | 
| prove_compare_expansions ord (thm :: thms) =  | 
|
974  | 
      @{thm compare_expansions_same_exp} OF [thm, prove_compare_expansions ord thms]
 | 
|
975  | 
| prove_compare_expansions _ [] = raise Match  | 
|
976  | 
||
| 69597 | 977  | 
val ev_zero_pos_thm = Eventuallize.eventuallize \<^context>  | 
| 68630 | 978  | 
  @{lemma "\<forall>x::real. f x = 0 \<longrightarrow> g x > 0 \<longrightarrow> f x < g x" by auto} NONE
 | 
979  | 
  OF @{thms _ expands_to_imp_eventually_pos}
 | 
|
980  | 
||
| 69597 | 981  | 
val ev_zero_neg_thm = Eventuallize.eventuallize \<^context>  | 
| 68630 | 982  | 
  @{lemma "\<forall>x::real. f x = 0 \<longrightarrow> g x < 0 \<longrightarrow> f x > g x" by auto} NONE
 | 
983  | 
  OF @{thms _ expands_to_imp_eventually_neg}
 | 
|
984  | 
||
| 69597 | 985  | 
val ev_zero_zero_thm = Eventuallize.eventuallize \<^context>  | 
| 68630 | 986  | 
  @{lemma "\<forall>x::real. f x = 0 \<longrightarrow> g x = 0 \<longrightarrow> f x = g x" by auto} NONE
 | 
987  | 
||
988  | 
fun compare_expansions_trivial ectxt (thm1, thm2, basis) =  | 
|
989  | 
case try_prove_ev_eq ectxt (apply2 get_expanded_fun (thm1, thm2)) of  | 
|
990  | 
SOME thm => SOME (EQUAL, thm, thm1, thm2)  | 
|
991  | 
| NONE =>  | 
|
992  | 
case apply2 (ev_zeroness_oracle ectxt o get_expanded_fun) (thm1, thm2) of  | 
|
993  | 
(NONE, NONE) => NONE  | 
|
994  | 
| (SOME zero1_thm, NONE) => (  | 
|
995  | 
case trim_expansion true (SOME Sgn_Trim) ectxt (thm2, basis) of  | 
|
996  | 
(thm2, IsPos, SOME trimmed2_thm) =>  | 
|
997  | 
SOME (LESS, ev_zero_pos_thm OF  | 
|
998  | 
[zero1_thm, get_basis_wf_thm basis, thm2, trimmed2_thm], thm1, thm2)  | 
|
999  | 
| (thm2, IsNeg, SOME trimmed2_thm) =>  | 
|
1000  | 
SOME (GREATER, ev_zero_neg_thm OF  | 
|
1001  | 
[zero1_thm, get_basis_wf_thm basis, thm2, trimmed2_thm], thm1, thm2)  | 
|
1002  | 
          | _ => raise TERM ("Unexpected zeroness result in compare_expansions", []))
 | 
|
1003  | 
| (NONE, SOME zero2_thm) => (  | 
|
1004  | 
case trim_expansion true (SOME Sgn_Trim) ectxt (thm1, basis) of  | 
|
1005  | 
(thm1, IsPos, SOME trimmed1_thm) =>  | 
|
1006  | 
SOME (GREATER, ev_zero_pos_thm OF  | 
|
1007  | 
[zero2_thm, get_basis_wf_thm basis, thm1, trimmed1_thm], thm1, thm2)  | 
|
1008  | 
| (thm1, IsNeg, SOME trimmed1_thm) =>  | 
|
1009  | 
SOME (LESS, ev_zero_neg_thm OF  | 
|
1010  | 
[zero2_thm, get_basis_wf_thm basis, thm1, trimmed1_thm], thm1, thm2)  | 
|
1011  | 
          | _ => raise TERM ("Unexpected zeroness result in compare_expansions", []))
 | 
|
1012  | 
| (SOME zero1_thm, SOME zero2_thm) =>  | 
|
1013  | 
SOME (EQUAL, ev_zero_zero_thm OF [zero1_thm, zero2_thm] , thm1, thm2)  | 
|
1014  | 
||
1015  | 
fun compare_expansions ectxt (thm1, thm2, basis) =  | 
|
1016  | 
case compare_expansions_trivial ectxt (thm1, thm2, basis) of  | 
|
1017  | 
SOME res => res  | 
|
1018  | 
| NONE =>  | 
|
1019  | 
let  | 
|
1020  | 
val (_, thm1, _) = whnf_expansion ectxt thm1  | 
|
1021  | 
val (_, thm2, _) = whnf_expansion ectxt thm2  | 
|
1022  | 
in  | 
|
1023  | 
case compare_expansions' ectxt (thm1, thm2, basis) of  | 
|
1024  | 
Cmp_Dominated (order, e_thms, zeroness, trimmed_thm, thm1, thm2) =>  | 
|
1025  | 
let  | 
|
1026  | 
val wf_thm = get_basis_wf_thm basis  | 
|
1027  | 
val cmp_thm = prove_compare_expansions order e_thms  | 
|
1028  | 
val trimmed_thm' = trimmed_thm RS  | 
|
1029  | 
              (if zeroness = IsPos then @{thm trimmed_pos_imp_trimmed}
 | 
|
1030  | 
                 else @{thm trimmed_neg_imp_trimmed})
 | 
|
1031  | 
val smallo_thm =  | 
|
1032  | 
              (if order = LESS then @{thm compare_expansions_LT} else @{thm compare_expansions_GT}) OF
 | 
|
1033  | 
[cmp_thm, trimmed_thm', thm1, thm2, wf_thm]  | 
|
1034  | 
val thm' =  | 
|
1035  | 
              if zeroness = IsPos then @{thm smallo_trimmed_imp_eventually_less} 
 | 
|
1036  | 
              else @{thm smallo_trimmed_imp_eventually_greater}
 | 
|
1037  | 
val result_thm =  | 
|
1038  | 
thm' OF [smallo_thm, if order = LESS then thm2 else thm1, wf_thm, trimmed_thm]  | 
|
1039  | 
in  | 
|
1040  | 
(order, result_thm, thm1, thm2)  | 
|
1041  | 
end  | 
|
1042  | 
| Cmp_Asymp_Equiv (thm1, thm2) =>  | 
|
1043  | 
let  | 
|
1044  | 
            val thm = @{thm expands_to_minus} OF [get_basis_wf_thm basis, thm1, thm2]
 | 
|
1045  | 
val (order, result_thm) =  | 
|
1046  | 
case trim_expansion true (SOME Sgn_Trim) ectxt (thm, basis) of  | 
|
1047  | 
(thm, IsPos, SOME pos_thm) => (GREATER,  | 
|
1048  | 
                  @{thm expands_to_imp_eventually_gt} OF [get_basis_wf_thm basis, thm, pos_thm])
 | 
|
1049  | 
| (thm, IsNeg, SOME neg_thm) => (LESS,  | 
|
1050  | 
                  @{thm expands_to_imp_eventually_lt} OF [get_basis_wf_thm basis, thm, neg_thm])
 | 
|
1051  | 
              | _ => raise TERM ("Unexpected zeroness result in prove_eventually_less", [])
 | 
|
1052  | 
in  | 
|
1053  | 
(order, result_thm, thm1, thm2)  | 
|
1054  | 
end  | 
|
1055  | 
end  | 
|
1056  | 
||
1057  | 
||
1058  | 
||
1059  | 
(*  | 
|
1060  | 
Throws an exception and prints an error message indicating that the leading term could  | 
|
1061  | 
not be determined to be either zero or non-zero.  | 
|
1062  | 
*)  | 
|
1063  | 
fun raise_trimming_error ectxt thm =  | 
|
1064  | 
let  | 
|
1065  | 
val ctxt = get_ctxt ectxt  | 
|
1066  | 
fun lead_coeff exp =  | 
|
| 69597 | 1067  | 
if fastype_of exp = \<^typ>\<open>real\<close> then exp else lead_coeff (get_coeff exp)  | 
| 68630 | 1068  | 
val c = lead_coeff (get_expansion thm)  | 
1069  | 
fun err () =  | 
|
1070  | 
let  | 
|
1071  | 
val t = simplify_term' (get_facts ectxt) ctxt c  | 
|
1072  | 
val _ =  | 
|
1073  | 
if #verbose (#ctxt ectxt) then  | 
|
1074  | 
let  | 
|
1075  | 
val p = Pretty.str  | 
|
1076  | 
                ("real_asymp failed to determine whether the following constant is zero:")
 | 
|
1077  | 
val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt t)]  | 
|
1078  | 
in  | 
|
1079  | 
Pretty.writeln p  | 
|
1080  | 
end else ()  | 
|
1081  | 
in  | 
|
1082  | 
        raise TERM ("zeroness_oracle", [t])
 | 
|
1083  | 
end  | 
|
1084  | 
in  | 
|
1085  | 
err ()  | 
|
1086  | 
end  | 
|
1087  | 
||
1088  | 
||
1089  | 
(* TODO Here be dragons *)  | 
|
1090  | 
fun solve_eval_eq thm =  | 
|
1091  | 
  case try (fn _ => @{thm refl} RS thm) () of
 | 
|
1092  | 
SOME thm' => thm'  | 
|
1093  | 
| NONE =>  | 
|
1094  | 
      case try (fn _ => @{thm eval_real_def} RS thm) () of
 | 
|
1095  | 
SOME thm' => thm'  | 
|
1096  | 
      | NONE => @{thm eval_ms.simps} RS thm
 | 
|
1097  | 
||
1098  | 
(*  | 
|
1099  | 
Returns an expansion theorem for the logarithm of the given expansion.  | 
|
1100  | 
May add one additional element to the basis at the end.  | 
|
1101  | 
*)  | 
|
1102  | 
fun ln_expansion _ _ _ SEmpty = raise TERM ("ln_expansion: empty basis", [])
 | 
|
1103  | 
| ln_expansion ectxt trimmed_thm thm (SNE basis) =  | 
|
1104  | 
let  | 
|
1105  | 
fun trailing_exponent expr (SSng _) = get_exponent expr  | 
|
1106  | 
| trailing_exponent expr (SCons (_, _, tl)) = trailing_exponent (get_coeff expr) tl  | 
|
1107  | 
val e = trailing_exponent (get_expansion thm) basis  | 
|
1108  | 
fun ln_expansion_aux trimmed_thm zero_thm thm basis =  | 
|
1109  | 
let  | 
|
| 69597 | 1110  | 
val t = betapply (\<^term>\<open>\<lambda>(f::real \<Rightarrow> real) x. f x - 1 :: real\<close>, get_expanded_fun thm)  | 
| 68630 | 1111  | 
in  | 
1112  | 
case ev_zeroness_oracle ectxt t of  | 
|
1113  | 
NONE => ln_expansion_aux' trimmed_thm zero_thm thm basis  | 
|
1114  | 
| SOME zero_thm =>  | 
|
1115  | 
            @{thm expands_to_ln_eventually_1} OF 
 | 
|
1116  | 
[get_basis_wf_thm' basis, mk_expansion_level_eq_thm' basis, zero_thm]  | 
|
1117  | 
end  | 
|
1118  | 
    and ln_expansion_aux' trimmed_thm zero_thm thm (SSng {wf_thm, ...}) =
 | 
|
1119  | 
          ( @{thm expands_to_ln} OF
 | 
|
1120  | 
[trimmed_thm, wf_thm, thm,  | 
|
1121  | 
              @{thm expands_to_ln_aux_0} OF [zero_thm, @{thm expands_to_ln_const}]])
 | 
|
1122  | 
|> solve_eval_eq  | 
|
1123  | 
      | ln_expansion_aux' trimmed_thm zero_thm thm (SCons ({wf_thm, ...}, {ln_thm, ...}, basis')) =
 | 
|
1124  | 
let  | 
|
1125  | 
val c_thm =  | 
|
1126  | 
              ln_expansion_aux (trimmed_thm RS @{thm trimmed_pos_hd_coeff}) zero_thm 
 | 
|
1127  | 
(expands_to_hd thm) basis'  | 
|
1128  | 
val e = get_exponent (get_expansion thm)  | 
|
1129  | 
val c_thm' =  | 
|
1130  | 
case zeroness_oracle true NONE ectxt e of  | 
|
1131  | 
(IsZero, SOME thm) =>  | 
|
1132  | 
                  @{thm expands_to_ln_to_expands_to_ln_eval [OF expands_to_ln_aux_0]} OF [thm,c_thm]
 | 
|
1133  | 
| _ =>  | 
|
| 69597 | 1134  | 
case try_prove_real_eq false ectxt (e, \<^term>\<open>1::real\<close>) of  | 
| 68630 | 1135  | 
SOME thm =>  | 
1136  | 
                    @{thm expands_to_ln_to_expands_to_ln_eval [OF expands_to_ln_aux_1]}
 | 
|
1137  | 
OF [thm, wf_thm, c_thm, ln_thm]  | 
|
1138  | 
| NONE =>  | 
|
1139  | 
                    @{thm expands_to_ln_to_expands_to_ln_eval [OF expands_to_ln_aux]} 
 | 
|
1140  | 
OF [wf_thm, c_thm, ln_thm]  | 
|
1141  | 
in  | 
|
1142  | 
            (@{thm expands_to_ln} OF [trimmed_thm, wf_thm, thm, c_thm'])
 | 
|
1143  | 
|> solve_eval_eq  | 
|
1144  | 
end  | 
|
1145  | 
in  | 
|
1146  | 
case zeroness_oracle true NONE ectxt e of  | 
|
1147  | 
(IsZero, SOME zero_thm) => (ln_expansion_aux trimmed_thm zero_thm thm basis, SNE basis)  | 
|
1148  | 
| _ =>  | 
|
1149  | 
let  | 
|
1150  | 
val basis' = insert_ln (SNE basis)  | 
|
1151  | 
val lifting = mk_lifting (get_basis_list' basis) basis'  | 
|
1152  | 
val thm' = lift_expands_to_thm lifting thm  | 
|
1153  | 
val trimmed_thm' = lift_trimmed_pos_thm lifting trimmed_thm  | 
|
1154  | 
val (thm'', eq_thm) = retrim_expansion ectxt (thm', basis')  | 
|
1155  | 
          val trimmed_thm'' = @{thm trimmed_pos_eq_cong} OF [trimmed_thm', eq_thm]
 | 
|
1156  | 
in  | 
|
1157  | 
ln_expansion ectxt trimmed_thm'' thm'' basis'  | 
|
1158  | 
end  | 
|
1159  | 
end  | 
|
1160  | 
||
1161  | 
(*  | 
|
1162  | 
Handles a possible basis change after expanding exp(c(x)) for an expansion of the form  | 
|
1163  | 
f(x) = c(x) + g(x). Expanding exp(c(x)) may have inserted an additional basis element. If the  | 
|
1164  | 
old basis was b :: bs (i.e. c is an expansion w.r.t. bs) and the updated one is bs' (which  | 
|
1165  | 
agrees with bs except for one additional element b'), we need to argue that b :: bs' is still  | 
|
1166  | 
well-formed. This may require us to show that ln(b') is o(ln(b)), which the function takes  | 
|
1167  | 
as an argument.  | 
|
1168  | 
*)  | 
|
1169  | 
fun adjust_exp_basis basis basis' ln_smallo_thm =  | 
|
1170  | 
if length (get_basis_list basis) = length (get_basis_list basis') + 1 then  | 
|
1171  | 
basis  | 
|
1172  | 
else  | 
|
1173  | 
let  | 
|
1174  | 
val SNE (SCons (info, ln_info, tail)) = basis  | 
|
1175  | 
val SNE tail' = basis'  | 
|
1176  | 
val wf_thms = map get_basis_wf_thm [basis, basis']  | 
|
1177  | 
val wf_thm' =  | 
|
1178  | 
case  | 
|
1179  | 
get_first (fn f => try f ())  | 
|
1180  | 
            [fn _ => @{thm basis_wf_lift_modification} OF wf_thms,
 | 
|
1181  | 
             fn _ => @{thm basis_wf_insert_exp_near} OF (wf_thms @ [ln_smallo_thm]),
 | 
|
1182  | 
             fn _ => @{thm basis_wf_insert_exp_near} OF (wf_thms @ 
 | 
|
1183  | 
               [ln_smallo_thm RS @{thm basis_wf_insert_exp_uminus'}])] of
 | 
|
1184  | 
SOME wf_thm => wf_thm  | 
|
1185  | 
        | _ => raise TERM ("Lifting basis modification in exp_expansion failed.", map Thm.concl_of (wf_thms @ [ln_smallo_thm]))
 | 
|
1186  | 
      val info' = {wf_thm = wf_thm', head = #head info}
 | 
|
1187  | 
val lifting = mk_lifting (get_basis_list' tail) basis'  | 
|
1188  | 
val ln_info' =  | 
|
1189  | 
        {trimmed_thm = lift_trimmed_pos_thm lifting (#trimmed_thm ln_info),
 | 
|
1190  | 
ln_thm = lift_expands_to_thm lifting (#ln_thm ln_info)}  | 
|
1191  | 
in  | 
|
1192  | 
SNE (SCons (info', ln_info', tail'))  | 
|
1193  | 
end  | 
|
1194  | 
||
1195  | 
(* inserts the exponential of a given function at the beginning of the given basis *)  | 
|
1196  | 
fun insert_exp _ _ _ _ _ SEmpty = raise TERM ("insert_exp", [])
 | 
|
1197  | 
| insert_exp t ln_thm ln_smallo_thm ln_trimmed_thm lim_thm (SNE basis) =  | 
|
1198  | 
let  | 
|
| 69597 | 1199  | 
val head = Envir.beta_eta_contract (\<^term>\<open>\<lambda>(f::real\<Rightarrow>real) x. exp (f x)\<close> $ t)  | 
| 68630 | 1200  | 
        val ln_smallo_thm = ln_smallo_thm RS @{thm ln_smallo_ln_exp}
 | 
1201  | 
        val wf_thm = @{thm basis_wf_manyI} OF [lim_thm, ln_smallo_thm, get_basis_wf_thm' basis]
 | 
|
1202  | 
        val basis' = SNE (SCons ({wf_thm = wf_thm, head = head}, 
 | 
|
1203  | 
          {ln_thm = ln_thm, trimmed_thm = ln_trimmed_thm} , basis))
 | 
|
1204  | 
in  | 
|
1205  | 
check_basis basis'  | 
|
1206  | 
end  | 
|
1207  | 
||
1208  | 
(*  | 
|
1209  | 
Returns an expansion of the exponential of the given expansion. This may add several  | 
|
1210  | 
new basis elements at any position of the basis (except at the very end  | 
|
1211  | 
*)  | 
|
1212  | 
fun exp_expansion _ thm SEmpty = (thm RS @{thm expands_to_exp_real}, SEmpty)
 | 
|
1213  | 
| exp_expansion ectxt thm basis =  | 
|
1214  | 
let  | 
|
1215  | 
val (_, thm, _) = whnf_expansion ectxt thm  | 
|
1216  | 
in  | 
|
1217  | 
case ev_zeroness_oracle ectxt (get_eval (get_expansion thm)) of  | 
|
1218  | 
SOME zero_thm =>  | 
|
1219  | 
          (@{thm expands_to_exp_zero} OF 
 | 
|
1220  | 
[thm, zero_thm, get_basis_wf_thm basis, mk_expansion_level_eq_thm basis], basis)  | 
|
1221  | 
| NONE =>  | 
|
1222  | 
let  | 
|
1223  | 
val ln =  | 
|
1224  | 
Option.map (fn x => (#ln_thm x, #trimmed_thm x)) (get_ln_info basis)  | 
|
1225  | 
val ln = Option.map (fn (x, y) => retrim_pos_expansion ectxt (x, basis, y)) ln  | 
|
| 69597 | 1226  | 
val es' = \<^term>\<open>0::real\<close> :: (  | 
| 68630 | 1227  | 
case ln of  | 
1228  | 
NONE => []  | 
|
1229  | 
| SOME (ln_thm, _, _) => get_exponents (get_expansion ln_thm))  | 
|
1230  | 
val trim_result =  | 
|
1231  | 
trim_expansion_while_greater true (SOME es') false (SOME Simple_Trim) ectxt (thm, basis)  | 
|
1232  | 
in  | 
|
1233  | 
exp_expansion' ectxt trim_result ln basis  | 
|
1234  | 
end  | 
|
1235  | 
end  | 
|
1236  | 
and exp_expansion' _ (thm, _, _) _ SEmpty = (thm RS @{thm expands_to_exp_real}, SEmpty)
 | 
|
1237  | 
| exp_expansion' ectxt (thm, trim_result, e_thms) ln basis =  | 
|
1238  | 
let  | 
|
1239  | 
val exp = get_expansion thm  | 
|
1240  | 
val wf_thm = get_basis_wf_thm basis  | 
|
1241  | 
val f = get_expanded_fun thm  | 
|
1242  | 
fun exp_expansion_insert ln_smallo_thm = (  | 
|
1243  | 
case determine_trimmed_sgn ectxt exp of  | 
|
1244  | 
(IsPos, trimmed_thm) =>  | 
|
1245  | 
let  | 
|
1246  | 
val [lim_thm, ln_thm', thm'] =  | 
|
1247  | 
              @{thms expands_to_exp_insert_pos}
 | 
|
1248  | 
|> map (fn thm' => thm' OF [thm, wf_thm, trimmed_thm, ln_smallo_thm])  | 
|
1249  | 
val basis' = insert_exp f ln_thm' ln_smallo_thm trimmed_thm lim_thm basis  | 
|
1250  | 
in  | 
|
1251  | 
(thm', basis')  | 
|
1252  | 
end  | 
|
1253  | 
| (IsNeg, trimmed_thm) =>  | 
|
1254  | 
let  | 
|
1255  | 
val [lim_thm, ln_thm', ln_trimmed_thm, thm'] =  | 
|
1256  | 
              @{thms expands_to_exp_insert_neg}
 | 
|
1257  | 
|> map (fn thm' => thm' OF [thm, wf_thm, trimmed_thm, ln_smallo_thm])  | 
|
1258  | 
            val ln_smallo_thm = ln_smallo_thm RS @{thm basis_wf_insert_exp_uminus}
 | 
|
| 69597 | 1259  | 
val f' = Envir.beta_eta_contract (\<^term>\<open>\<lambda>(f::real\<Rightarrow>real) x. -f x\<close> $ f)  | 
| 68630 | 1260  | 
val basis' = insert_exp f' ln_thm' ln_smallo_thm ln_trimmed_thm lim_thm basis  | 
1261  | 
in  | 
|
1262  | 
(thm', basis')  | 
|
1263  | 
end  | 
|
1264  | 
      | _ => raise TERM ("Unexpected zeroness result in exp_expansion", []))
 | 
|
1265  | 
fun lexord (IsNeg :: _) = LESS  | 
|
1266  | 
| lexord (IsPos :: _) = GREATER  | 
|
1267  | 
| lexord (IsZero :: xs) = lexord xs  | 
|
1268  | 
| lexord [] = EQUAL  | 
|
1269  | 
| lexord _ = raise Match  | 
|
1270  | 
val compare_result = lexord (map fst e_thms)  | 
|
1271  | 
in  | 
|
1272  | 
case (trim_result, e_thms, compare_result) of  | 
|
1273  | 
(Aborted _, (IsNeg, e_neg_thm) :: _, _) =>  | 
|
1274  | 
(* leading exponent is negative; we can simply Taylor-expand exp(x) around 0 *)  | 
|
1275  | 
        (@{thm expands_to_exp_neg} OF [thm, get_basis_wf_thm basis, e_neg_thm], basis)
 | 
|
1276  | 
| (Trimmed (_, SOME trimmed_thm), (IsPos, e_pos_thm) :: _, GREATER) =>  | 
|
1277  | 
(* leading exponent is positive; exp(f(x)) or exp(-f(x)) is new basis element *)  | 
|
1278  | 
let  | 
|
1279  | 
val ln_smallo_thm =  | 
|
1280  | 
            @{thm basis_wf_insert_exp_pos} OF [thm, get_basis_wf_thm basis, trimmed_thm, e_pos_thm]
 | 
|
1281  | 
in  | 
|
1282  | 
exp_expansion_insert ln_smallo_thm  | 
|
1283  | 
end  | 
|
1284  | 
| (Trimmed (_, SOME trimmed_thm), _, GREATER) =>  | 
|
1285  | 
(* leading exponent is zero, but f(x) grows faster than ln(b(x)), so  | 
|
1286  | 
exp(f(x)) or exp(-f(x)) must still be new basis elements *)  | 
|
1287  | 
let  | 
|
1288  | 
val ln_thm =  | 
|
1289  | 
case ln of  | 
|
1290  | 
SOME (ln_thm, _, _) => ln_thm  | 
|
1291  | 
            | NONE => raise TERM ("TODO blubb", [])
 | 
|
1292  | 
          val ln_thm = @{thm expands_to_lift''} OF [get_basis_wf_thm basis, ln_thm]
 | 
|
1293  | 
val ln_smallo_thm =  | 
|
1294  | 
             @{thm compare_expansions_GT} OF [prove_compare_expansions GREATER (map snd e_thms),
 | 
|
1295  | 
trimmed_thm, thm, ln_thm, get_basis_wf_thm basis]  | 
|
1296  | 
in  | 
|
1297  | 
exp_expansion_insert ln_smallo_thm  | 
|
1298  | 
end  | 
|
1299  | 
| (Aborted LESS, (IsZero, e_zero_thm) :: e_thms', _) =>  | 
|
1300  | 
(* leading exponent is zero and f(x) grows more slowly than ln(b(x)), so  | 
|
1301  | 
we can write f(x) = c(x) + g(x) and therefore exp(f(x)) = exp(c(x)) * exp(g(x)).  | 
|
1302  | 
The former is treated by a recursive call; the latter by Taylor expansion. *)  | 
|
1303  | 
let  | 
|
1304  | 
val (ln_thm, trimmed_thm) =  | 
|
1305  | 
case ln of  | 
|
1306  | 
SOME (ln_thm, _, trimmed_thm) =>  | 
|
1307  | 
                (ln_thm, trimmed_thm RS @{thm trimmed_pos_imp_trimmed})
 | 
|
1308  | 
            | NONE => raise TERM ("TODO foo", [])
 | 
|
1309  | 
val c_thm = expands_to_hd thm  | 
|
1310  | 
val ln_smallo_thm =  | 
|
1311  | 
            @{thm compare_expansions_LT} OF [prove_compare_expansions LESS (map snd e_thms'),
 | 
|
1312  | 
trimmed_thm, c_thm, ln_thm, get_basis_wf_thm (tl_basis basis)]  | 
|
1313  | 
val (c_thm, c_basis) = exp_expansion ectxt c_thm (tl_basis basis)  | 
|
1314  | 
val basis' = adjust_exp_basis basis c_basis ln_smallo_thm  | 
|
1315  | 
val wf_thm = get_basis_wf_thm basis'  | 
|
1316  | 
val thm' = lift basis' thm  | 
|
1317  | 
val (thm'', _) = retrim_expansion ectxt (thm', basis')  | 
|
1318  | 
in  | 
|
1319  | 
          (@{thm expands_to_exp_0} OF [thm'', wf_thm, e_zero_thm, c_thm], basis')
 | 
|
1320  | 
end  | 
|
1321  | 
| (Trimmed _, [(IsZero, e_zero_thm)], EQUAL) =>  | 
|
1322  | 
(* f(x) can be written as c + g(x) where c is just a real constant.  | 
|
1323  | 
We can therefore write exp(f(x)) = exp(c) * exp(g(x)), where the latter is  | 
|
1324  | 
a simple Taylor expansion. *)  | 
|
1325  | 
        (@{thm expands_to_exp_0_real} OF [thm, wf_thm, e_zero_thm], basis)
 | 
|
1326  | 
| (Trimmed _, (_, e_zero_thm) :: _, EQUAL) =>  | 
|
1327  | 
(* f(x) is asymptotically equivalent to c * ln(b(x)), so we can write f(x) as  | 
|
1328  | 
c * ln(b(x)) + g(x) and therefore exp(f(x)) = b(x)^c * exp(g(x)). The second  | 
|
1329  | 
factor is handled by a recursive call *)  | 
|
1330  | 
let  | 
|
1331  | 
val ln_thm =  | 
|
1332  | 
case ln of  | 
|
1333  | 
SOME (ln_thm, _, _) => ln_thm  | 
|
1334  | 
            | NONE => raise TERM ("TODO blargh", [])
 | 
|
1335  | 
val c =  | 
|
1336  | 
case (thm, ln_thm) |> apply2 (get_expansion #> get_lead_coeff) of  | 
|
| 69597 | 1337  | 
(c1, c2) => \<^term>\<open>(/) :: real => _\<close> $ c1 $ c2  | 
| 68630 | 1338  | 
val c = Thm.cterm_of (get_ctxt ectxt) c  | 
1339  | 
||
1340  | 
val thm' =  | 
|
1341  | 
            @{thm expands_to_exp_0_pull_out1} 
 | 
|
1342  | 
OF [thm, ln_thm, wf_thm, e_zero_thm, Thm.reflexive c]  | 
|
1343  | 
val (thm'', basis') = exp_expansion ectxt thm' basis  | 
|
1344  | 
          val pat = ConsPat ("MS", [AnyPat ("_", 0), AnyPat ("_", 0)])
 | 
|
1345  | 
val (_, _, conv) = match ectxt pat (get_expansion thm'') (SOME [])  | 
|
1346  | 
val eq_thm = conv (Thm.cterm_of (get_ctxt ectxt) (get_expansion thm''))  | 
|
1347  | 
          val thm''' = @{thm expands_to_meta_eq_cong} OF [thm'', eq_thm]
 | 
|
1348  | 
val thm'''' =  | 
|
1349  | 
case get_intyness (get_ctxt ectxt) c of  | 
|
1350  | 
No_Nat =>  | 
|
1351  | 
                @{thm expands_to_exp_0_pull_out2} OF [thm''', get_basis_wf_thm basis']
 | 
|
1352  | 
| Nat nat_thm =>  | 
|
1353  | 
                @{thm expands_to_exp_0_pull_out2_nat} OF 
 | 
|
1354  | 
[thm''', get_basis_wf_thm basis', nat_thm]  | 
|
1355  | 
| Neg_Nat nat_thm =>  | 
|
1356  | 
                @{thm expands_to_exp_0_pull_out2_neg_nat} OF 
 | 
|
1357  | 
[thm''', get_basis_wf_thm basis', nat_thm]  | 
|
1358  | 
in  | 
|
1359  | 
(thm'''', basis')  | 
|
1360  | 
end  | 
|
1361  | 
| (Trimmed (IsZero, _), [], _) =>  | 
|
1362  | 
(* Expansion is empty, i.e. f(x) is identically zero *)  | 
|
1363  | 
        (@{thm expands_to_exp_MSLNil} OF [thm, get_basis_wf_thm basis], basis)
 | 
|
1364  | 
| (Trimmed (_, NONE), _, GREATER) =>  | 
|
1365  | 
(* We could not determine whether f(x) grows faster than ln(b(x)) or not. *)  | 
|
1366  | 
raise_trimming_error ectxt thm  | 
|
1367  | 
| _ => raise Match  | 
|
1368  | 
end  | 
|
1369  | 
||
1370  | 
fun powr_expansion ectxt (thm1, thm2, basis) =  | 
|
1371  | 
case ev_zeroness_oracle ectxt (get_expanded_fun thm1) of  | 
|
1372  | 
SOME zero_thm =>  | 
|
1373  | 
          (@{thm expands_to_powr_0} OF
 | 
|
1374  | 
[zero_thm, Thm.reflexive (Thm.cterm_of (get_ctxt ectxt) (get_expanded_fun thm2)),  | 
|
1375  | 
get_basis_wf_thm basis, mk_expansion_level_eq_thm basis],  | 
|
1376  | 
basis)  | 
|
1377  | 
| NONE =>  | 
|
1378  | 
let  | 
|
1379  | 
val (thm1, _, SOME trimmed_thm) =  | 
|
1380  | 
trim_expansion true (SOME Pos_Trim) ectxt (thm1, basis)  | 
|
1381  | 
val (ln_thm, basis') = ln_expansion ectxt trimmed_thm thm1 basis  | 
|
1382  | 
val thm2' = lift basis' thm2 |> simplify_expansion ectxt  | 
|
1383  | 
            val mult_thm = @{thm expands_to_mult} OF [get_basis_wf_thm basis', ln_thm, thm2']
 | 
|
1384  | 
val (exp_thm, basis'') = exp_expansion ectxt mult_thm basis'  | 
|
1385  | 
            val thm = @{thm expands_to_powr} OF 
 | 
|
1386  | 
[trimmed_thm, get_basis_wf_thm basis, thm1, exp_thm]  | 
|
1387  | 
in  | 
|
1388  | 
(thm, basis'')  | 
|
1389  | 
end  | 
|
1390  | 
||
1391  | 
fun powr_nat_expansion ectxt (thm1, thm2, basis) =  | 
|
1392  | 
case ev_zeroness_oracle ectxt (get_expanded_fun thm1) of  | 
|
1393  | 
SOME zero_thm => (  | 
|
1394  | 
case ev_zeroness_oracle ectxt (get_expanded_fun thm2) of  | 
|
1395  | 
            SOME zero'_thm => (@{thm expands_to_powr_nat_0_0} OF
 | 
|
1396  | 
[zero_thm, zero'_thm, get_basis_wf_thm basis, mk_expansion_level_eq_thm basis], basis)  | 
|
1397  | 
| NONE => (  | 
|
1398  | 
case trim_expansion true (SOME Simple_Trim) ectxt (thm2, basis) of  | 
|
1399  | 
(thm2, _, SOME trimmed_thm) =>  | 
|
1400  | 
                  (@{thm expands_to_powr_nat_0} OF [zero_thm, thm2, trimmed_thm, 
 | 
|
1401  | 
get_basis_wf_thm basis, mk_expansion_level_eq_thm basis], basis)))  | 
|
1402  | 
| NONE =>  | 
|
1403  | 
let  | 
|
1404  | 
val (thm1, _, SOME trimmed_thm) =  | 
|
1405  | 
trim_expansion true (SOME Pos_Trim) ectxt (thm1, basis)  | 
|
1406  | 
val (ln_thm, basis') = ln_expansion ectxt trimmed_thm thm1 basis  | 
|
1407  | 
val thm2' = lift basis' thm2 |> simplify_expansion ectxt  | 
|
1408  | 
            val mult_thm = @{thm expands_to_mult} OF [get_basis_wf_thm basis', ln_thm, thm2']
 | 
|
1409  | 
val (exp_thm, basis'') = exp_expansion ectxt mult_thm basis'  | 
|
1410  | 
            val thm = @{thm expands_to_powr_nat} OF 
 | 
|
1411  | 
[trimmed_thm, get_basis_wf_thm basis, thm1, exp_thm]  | 
|
1412  | 
in  | 
|
1413  | 
(thm, basis'')  | 
|
1414  | 
end  | 
|
1415  | 
||
1416  | 
fun is_numeral t =  | 
|
1417  | 
let  | 
|
1418  | 
val _ = HOLogic.dest_number t  | 
|
1419  | 
in  | 
|
1420  | 
true  | 
|
1421  | 
end  | 
|
1422  | 
handle TERM _ => false  | 
|
1423  | 
||
1424  | 
fun power_expansion ectxt (thm, n, basis) =  | 
|
1425  | 
case ev_zeroness_oracle ectxt (get_expanded_fun thm) of  | 
|
1426  | 
        SOME zero_thm => @{thm expands_to_power_0} OF
 | 
|
1427  | 
[zero_thm, get_basis_wf_thm basis, mk_expansion_level_eq_thm basis,  | 
|
1428  | 
Thm.reflexive (Thm.cterm_of (get_ctxt ectxt) n)]  | 
|
1429  | 
| NONE => (  | 
|
1430  | 
case trim_expansion true (SOME Simple_Trim) ectxt (thm, basis) of  | 
|
1431  | 
(thm', _, SOME trimmed_thm) =>  | 
|
1432  | 
let  | 
|
1433  | 
val ctxt = get_ctxt ectxt  | 
|
1434  | 
val thm =  | 
|
1435  | 
                  if is_numeral n then @{thm expands_to_power[where abort = True]}
 | 
|
1436  | 
                    else @{thm expands_to_power[where abort = False]}
 | 
|
1437  | 
val thm =  | 
|
1438  | 
Drule.infer_instantiate' ctxt [NONE, NONE, NONE, SOME (Thm.cterm_of ctxt n)] thm  | 
|
1439  | 
in  | 
|
1440  | 
thm OF [trimmed_thm, get_basis_wf_thm basis, thm']  | 
|
1441  | 
end  | 
|
1442  | 
          | _ => raise TERM ("Unexpected zeroness result in power_expansion", []))
 | 
|
1443  | 
||
1444  | 
fun powr_const_expansion ectxt (thm, p, basis) =  | 
|
1445  | 
let  | 
|
1446  | 
val pthm = Thm.reflexive (Thm.cterm_of (get_ctxt ectxt) p)  | 
|
1447  | 
in  | 
|
1448  | 
case ev_zeroness_oracle ectxt (get_expanded_fun thm) of  | 
|
1449  | 
      SOME zero_thm => @{thm expands_to_powr_const_0} OF 
 | 
|
1450  | 
[zero_thm, get_basis_wf_thm basis, mk_expansion_level_eq_thm basis, pthm]  | 
|
1451  | 
| NONE =>  | 
|
1452  | 
case trim_expansion true (SOME Pos_Trim) ectxt (thm, basis) of  | 
|
1453  | 
          (_, _, NONE) => raise TERM ("Unexpected zeroness result for powr", [])
 | 
|
1454  | 
| (thm, _, SOME trimmed_thm) =>  | 
|
1455  | 
            (if is_numeral p then @{thm expands_to_powr_const[where abort = True]}
 | 
|
1456  | 
                 else @{thm expands_to_powr_const[where abort = False]})
 | 
|
1457  | 
OF [trimmed_thm, get_basis_wf_thm basis, thm, pthm]  | 
|
1458  | 
end  | 
|
1459  | 
||
1460  | 
fun sgn_expansion ectxt (thm, basis) =  | 
|
1461  | 
let  | 
|
1462  | 
val thms = [get_basis_wf_thm basis, mk_expansion_level_eq_thm basis]  | 
|
1463  | 
in  | 
|
1464  | 
case ev_zeroness_oracle ectxt (get_expanded_fun thm) of  | 
|
1465  | 
      SOME zero_thm => @{thm expands_to_sgn_zero} OF (zero_thm :: thms)
 | 
|
1466  | 
| NONE =>  | 
|
1467  | 
case trim_expansion true (SOME Sgn_Trim) ectxt (thm, basis) of  | 
|
1468  | 
(thm, IsPos, SOME trimmed_thm) =>  | 
|
1469  | 
            @{thm expands_to_sgn_pos} OF ([trimmed_thm, thm] @ thms)
 | 
|
1470  | 
| (thm, IsNeg, SOME trimmed_thm) =>  | 
|
1471  | 
            @{thm expands_to_sgn_neg} OF ([trimmed_thm, thm] @ thms)
 | 
|
1472  | 
        | _ => raise TERM ("Unexpected zeroness result in sgn_expansion", [])
 | 
|
1473  | 
end  | 
|
1474  | 
||
1475  | 
(*  | 
|
1476  | 
Returns an expansion of the sine and cosine of the given expansion. Fails if that function  | 
|
1477  | 
goes to infinity.  | 
|
1478  | 
*)  | 
|
1479  | 
fun sin_cos_expansion _ thm SEmpty =  | 
|
1480  | 
      (thm RS @{thm expands_to_sin_real}, thm RS @{thm expands_to_cos_real})
 | 
|
1481  | 
| sin_cos_expansion ectxt thm basis =  | 
|
1482  | 
let  | 
|
1483  | 
val exp = get_expansion thm  | 
|
1484  | 
val e = get_exponent exp  | 
|
1485  | 
in  | 
|
1486  | 
case zeroness_oracle true (SOME Sgn_Trim) ectxt e of  | 
|
1487  | 
          (IsPos, _) => raise THM ("sin_cos_expansion", 0, [thm])
 | 
|
1488  | 
| (IsNeg, SOME e_thm) =>  | 
|
1489  | 
let  | 
|
1490  | 
val [thm1, thm2] =  | 
|
1491  | 
map (fn thm' => thm' OF [e_thm, get_basis_wf_thm basis, thm])  | 
|
1492  | 
                  @{thms expands_to_sin_ms_neg_exp expands_to_cos_ms_neg_exp}
 | 
|
1493  | 
in  | 
|
1494  | 
(thm1, thm2)  | 
|
1495  | 
end  | 
|
1496  | 
| (IsZero, SOME e_thm) =>  | 
|
1497  | 
let  | 
|
1498  | 
val (sin_thm, cos_thm) = (sin_cos_expansion ectxt (expands_to_hd thm) (tl_basis basis))  | 
|
1499  | 
fun mk_thm thm' =  | 
|
1500  | 
(thm' OF [e_thm, get_basis_wf_thm basis, thm, sin_thm, cos_thm]) |> solve_eval_eq  | 
|
1501  | 
val [thm1, thm2] =  | 
|
1502  | 
                map mk_thm @{thms expands_to_sin_ms_zero_exp expands_to_cos_ms_zero_exp}
 | 
|
1503  | 
in  | 
|
1504  | 
(thm1, thm2)  | 
|
1505  | 
end  | 
|
1506  | 
        | _ => raise TERM ("Unexpected zeroness result in sin_exp_expansion", [])
 | 
|
1507  | 
end  | 
|
1508  | 
||
1509  | 
fun abconv (t, t') = Envir.beta_eta_contract t aconv Envir.beta_eta_contract t'  | 
|
1510  | 
||
1511  | 
(*  | 
|
1512  | 
Makes sure that an expansion theorem really talks about the right function.  | 
|
1513  | 
This is basically a sanity check to make things fail early and in the right place.  | 
|
1514  | 
*)  | 
|
1515  | 
fun check_expansion e thm =  | 
|
1516  | 
if abconv (expr_to_term e, get_expanded_fun thm) then  | 
|
1517  | 
thm  | 
|
1518  | 
else  | 
|
1519  | 
(* TODO Remove Debugging stuff *)  | 
|
| 69597 | 1520  | 
let val _ = \<^print> e  | 
1521  | 
val _ = \<^print> (get_expanded_fun thm)  | 
|
| 68630 | 1522  | 
in  | 
1523  | 
    raise TERM ("check_expansion", [Thm.concl_of thm, expr_to_term e])
 | 
|
1524  | 
end  | 
|
1525  | 
||
1526  | 
fun minmax_expansion max [less_thm, eq_thm, gt_thm] ectxt (thm1, thm2, basis) = (  | 
|
1527  | 
case compare_expansions ectxt (thm1, thm2, basis) of  | 
|
1528  | 
(LESS, less_thm', thm1, thm2) => less_thm OF [if max then thm2 else thm1, less_thm']  | 
|
1529  | 
| (GREATER, gt_thm', thm1, thm2) => gt_thm OF [if max then thm1 else thm2, gt_thm']  | 
|
1530  | 
| (EQUAL, eq_thm', thm1, _) => eq_thm OF [thm1, eq_thm'])  | 
|
1531  | 
| minmax_expansion _ _ _ _ = raise Match  | 
|
1532  | 
||
1533  | 
val min_expansion =  | 
|
1534  | 
  minmax_expansion false @{thms expands_to_min_lt expands_to_min_eq expands_to_min_gt}
 | 
|
1535  | 
val max_expansion =  | 
|
1536  | 
  minmax_expansion true @{thms expands_to_max_lt expands_to_max_eq expands_to_max_gt}
 | 
|
1537  | 
||
1538  | 
fun zero_expansion basis =  | 
|
1539  | 
  @{thm expands_to_zero} OF [get_basis_wf_thm basis, mk_expansion_level_eq_thm basis]
 | 
|
1540  | 
||
| 69597 | 1541  | 
fun const_expansion _ basis \<^term>\<open>0 :: real\<close> = zero_expansion basis  | 
| 68630 | 1542  | 
| const_expansion ectxt basis t =  | 
1543  | 
let  | 
|
1544  | 
val ctxt = get_ctxt ectxt  | 
|
1545  | 
val thm = Drule.infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt t)]  | 
|
1546  | 
                @{thm expands_to_const}
 | 
|
1547  | 
in  | 
|
1548  | 
thm OF [get_basis_wf_thm basis, mk_expansion_level_eq_thm basis]  | 
|
1549  | 
end  | 
|
1550  | 
||
1551  | 
fun root_expansion ectxt (thm, n, basis) =  | 
|
1552  | 
let  | 
|
1553  | 
val ctxt = get_ctxt ectxt  | 
|
1554  | 
    fun tac {context = ctxt, ...} =
 | 
|
1555  | 
HEADGOAL (Method.insert_tac ctxt (get_facts ectxt))  | 
|
1556  | 
THEN Local_Defs.unfold_tac ctxt eval_simps  | 
|
1557  | 
THEN HEADGOAL (Simplifier.asm_full_simp_tac ctxt)  | 
|
1558  | 
fun prove goal =  | 
|
1559  | 
try (Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (Term.betapply (goal, n)))) tac  | 
|
1560  | 
fun err () =  | 
|
1561  | 
let  | 
|
1562  | 
val t = simplify_term' (get_facts ectxt) ctxt n  | 
|
1563  | 
val _ =  | 
|
1564  | 
if #verbose (#ctxt ectxt) then  | 
|
1565  | 
let  | 
|
1566  | 
              val p = Pretty.str ("real_asymp failed to determine whether the following constant " ^
 | 
|
1567  | 
"is zero or not:")  | 
|
1568  | 
val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt t)]  | 
|
1569  | 
in  | 
|
1570  | 
Pretty.writeln p  | 
|
1571  | 
end else ()  | 
|
1572  | 
in  | 
|
1573  | 
        raise TERM ("zeroness_oracle", [n])
 | 
|
1574  | 
end  | 
|
1575  | 
fun aux nz_thm =  | 
|
1576  | 
case trim_expansion true (SOME Sgn_Trim) ectxt (thm, basis) of  | 
|
1577  | 
(thm, IsPos, SOME trimmed_thm) =>  | 
|
1578  | 
          @{thm expands_to_root} OF [nz_thm, trimmed_thm, get_basis_wf_thm basis, thm]
 | 
|
1579  | 
| (thm, IsNeg, SOME trimmed_thm) =>  | 
|
1580  | 
          @{thm expands_to_root_neg} OF [nz_thm, trimmed_thm, get_basis_wf_thm basis, thm]
 | 
|
1581  | 
      | _ => raise TERM ("Unexpected zeroness result in root_expansion", [])
 | 
|
1582  | 
in  | 
|
| 69597 | 1583  | 
case prove \<^term>\<open>\<lambda>n::nat. n = 0\<close> of  | 
| 68630 | 1584  | 
SOME zero_thm =>  | 
1585  | 
        @{thm expands_to_0th_root} OF
 | 
|
1586  | 
[zero_thm, get_basis_wf_thm basis, mk_expansion_level_eq_thm basis,  | 
|
1587  | 
Thm.reflexive (Thm.cterm_of ctxt (get_expanded_fun thm))]  | 
|
1588  | 
| NONE =>  | 
|
| 69597 | 1589  | 
case prove \<^term>\<open>\<lambda>n::nat. n > 0\<close> of  | 
| 68630 | 1590  | 
NONE => err ()  | 
1591  | 
| SOME nz_thm =>  | 
|
1592  | 
case ev_zeroness_oracle ectxt (get_expanded_fun thm) of  | 
|
1593  | 
              SOME zero_thm => @{thm expands_to_root_0} OF
 | 
|
1594  | 
[nz_thm, zero_thm, get_basis_wf_thm basis, mk_expansion_level_eq_thm basis]  | 
|
1595  | 
| NONE => aux nz_thm  | 
|
1596  | 
end  | 
|
1597  | 
||
1598  | 
||
1599  | 
fun arctan_expansion _ SEmpty thm =  | 
|
1600  | 
      @{thm expands_to_real_compose[where g = arctan]} OF [thm]
 | 
|
1601  | 
| arctan_expansion ectxt basis thm =  | 
|
1602  | 
case ev_zeroness_oracle ectxt (get_expanded_fun thm) of  | 
|
1603  | 
        SOME zero_thm => @{thm expands_to_arctan_zero} OF [zero_expansion basis, zero_thm]
 | 
|
1604  | 
| NONE =>  | 
|
1605  | 
let  | 
|
1606  | 
val (thm, _, _) = trim_expansion true (SOME Simple_Trim) ectxt (thm, basis)  | 
|
1607  | 
val e = get_exponent (get_expansion thm)  | 
|
1608  | 
fun cont ectxt (thm, basis) = arctan_expansion ectxt basis thm  | 
|
1609  | 
in  | 
|
1610  | 
case zeroness_oracle true (SOME Sgn_Trim) ectxt e of  | 
|
1611  | 
(IsNeg, SOME neg_thm) =>  | 
|
1612  | 
                  @{thm expands_to_arctan_ms_neg_exp} OF [neg_thm, get_basis_wf_thm basis, thm]
 | 
|
1613  | 
| (IsPos, SOME e_pos_thm) => (  | 
|
1614  | 
case determine_trimmed_sgn ectxt (get_expansion thm) of  | 
|
1615  | 
(IsPos, trimmed_thm) =>  | 
|
1616  | 
                      @{thm expands_to_arctan_ms_pos_exp_pos} OF 
 | 
|
1617  | 
[e_pos_thm, trimmed_thm, get_basis_wf_thm basis, thm]  | 
|
1618  | 
| (IsNeg, trimmed_thm) =>  | 
|
1619  | 
                      @{thm expands_to_arctan_ms_pos_exp_neg} OF 
 | 
|
1620  | 
[e_pos_thm, trimmed_thm, get_basis_wf_thm basis, thm]  | 
|
1621  | 
                  | _ => raise TERM ("Unexpected trim result during expansion of arctan", []))
 | 
|
1622  | 
| (IsZero, _) => (  | 
|
1623  | 
case try_lift_function ectxt (thm, basis) cont of  | 
|
1624  | 
(SOME thm', _) => thm'  | 
|
1625  | 
| _ =>  | 
|
1626  | 
let  | 
|
1627  | 
val _ = if get_verbose ectxt then  | 
|
1628  | 
writeln "Unsupported occurrence of arctan" else ()  | 
|
1629  | 
in  | 
|
1630  | 
                        raise TERM ("Unsupported occurence of arctan", [])
 | 
|
1631  | 
end)  | 
|
1632  | 
              | _ => raise TERM ("Unexpected trim result during expansion of arctan", [])
 | 
|
1633  | 
end  | 
|
1634  | 
||
1635  | 
(* Returns an expansion theorem for a function that is already a basis element *)  | 
|
1636  | 
fun expand_basic _ t SEmpty = raise TERM ("expand_basic", [t])
 | 
|
1637  | 
| expand_basic thm t basis =  | 
|
1638  | 
if abconv (get_basis_head basis, t) then  | 
|
1639  | 
thm (get_basis_wf_thm basis) (mk_expansion_level_eq_thm (tl_basis basis))  | 
|
1640  | 
else  | 
|
1641  | 
        @{thm expands_to_lift'} OF [get_basis_wf_thm basis, expand_basic thm t (tl_basis basis)]
 | 
|
1642  | 
||
1643  | 
fun expand_unary ectxt thm e basis =  | 
|
1644  | 
let  | 
|
1645  | 
val (thm', basis') = expand' ectxt e basis |> apfst (simplify_expansion ectxt)  | 
|
1646  | 
in  | 
|
1647  | 
(thm OF [get_basis_wf_thm basis', thm'], basis')  | 
|
1648  | 
end  | 
|
1649  | 
and expand_binary ectxt thm (e1, e2) basis =  | 
|
1650  | 
let  | 
|
1651  | 
val (thm1, basis') = expand' ectxt e1 basis |> apfst (simplify_expansion ectxt)  | 
|
1652  | 
val (thm2, basis'') = expand' ectxt e2 basis' |> apfst (simplify_expansion ectxt)  | 
|
1653  | 
val thm1 = lift basis'' thm1 |> simplify_expansion ectxt  | 
|
1654  | 
in  | 
|
1655  | 
(thm OF [get_basis_wf_thm basis'', thm1, thm2], basis'')  | 
|
1656  | 
end  | 
|
1657  | 
and trim_nz mode ectxt e basis =  | 
|
1658  | 
let  | 
|
1659  | 
val (thm, basis') = expand' ectxt e basis |> apfst (simplify_expansion ectxt)  | 
|
1660  | 
val (thm', nz, trimmed_thm) = trim_expansion true (SOME mode) ectxt (thm, basis')  | 
|
1661  | 
in  | 
|
1662  | 
case trimmed_thm of  | 
|
1663  | 
          NONE => raise TERM ("expand: zero denominator", [get_expansion thm])
 | 
|
1664  | 
| SOME trimmed_thm => (thm', basis', nz, trimmed_thm)  | 
|
1665  | 
end  | 
|
1666  | 
and expand'' ectxt (ConstExpr c) basis = (const_expansion ectxt basis c, basis)  | 
|
1667  | 
  | expand'' _ X basis = (lift basis @{thm expands_to_X}, basis)
 | 
|
1668  | 
  | expand'' ectxt (Uminus e) basis = expand_unary ectxt @{thm expands_to_uminus} e basis
 | 
|
1669  | 
  | expand'' ectxt (Add e12) basis = expand_binary ectxt @{thm expands_to_add} e12 basis
 | 
|
1670  | 
  | expand'' ectxt (Minus e12) basis = expand_binary ectxt @{thm expands_to_minus} e12 basis
 | 
|
1671  | 
  | expand'' ectxt (Mult e12) basis = expand_binary ectxt @{thm expands_to_mult} e12 basis
 | 
|
1672  | 
| expand'' ectxt (Powr' (e, p)) basis = (* TODO zero basis *)  | 
|
1673  | 
let  | 
|
1674  | 
val (thm, basis') = expand' ectxt e basis |> apfst (simplify_expansion ectxt)  | 
|
1675  | 
in  | 
|
1676  | 
(powr_const_expansion ectxt (thm, p, basis'), basis')  | 
|
1677  | 
end  | 
|
1678  | 
| expand'' ectxt (Powr (e1, e2)) basis =  | 
|
1679  | 
let  | 
|
1680  | 
val (thm2, basis1) = expand' ectxt e2 basis |> apfst (simplify_expansion ectxt)  | 
|
1681  | 
val (thm1, basis2) = expand' ectxt e1 basis1 |> apfst (simplify_expansion ectxt)  | 
|
1682  | 
in  | 
|
1683  | 
powr_expansion ectxt (thm1, thm2, basis2)  | 
|
1684  | 
end  | 
|
1685  | 
| expand'' ectxt (Powr_Nat (e1, e2)) basis =  | 
|
1686  | 
let  | 
|
1687  | 
val (thm2, basis1) = expand' ectxt e2 basis |> apfst (simplify_expansion ectxt)  | 
|
1688  | 
val (thm1, basis2) = expand' ectxt e1 basis1 |> apfst (simplify_expansion ectxt)  | 
|
1689  | 
in  | 
|
1690  | 
powr_nat_expansion ectxt (thm1, thm2, basis2)  | 
|
1691  | 
end  | 
|
1692  | 
| expand'' ectxt (LnPowr (e1, e2)) basis =  | 
|
1693  | 
let (* TODO zero base *)  | 
|
1694  | 
val (thm2, basis1) = expand' ectxt e2 basis |> apfst (simplify_expansion ectxt)  | 
|
1695  | 
val (thm1, basis2, _, trimmed_thm) = trim_nz Pos_Trim ectxt e1 basis1  | 
|
1696  | 
val (ln_thm, basis3) = ln_expansion ectxt trimmed_thm thm1 basis2  | 
|
1697  | 
val thm2' = lift basis3 thm2 |> simplify_expansion ectxt  | 
|
1698  | 
        val mult_thm = @{thm expands_to_mult} OF [get_basis_wf_thm basis3, ln_thm, thm2']
 | 
|
1699  | 
        val thm = @{thm expands_to_ln_powr} OF 
 | 
|
1700  | 
[trimmed_thm, get_basis_wf_thm basis2, thm1, mult_thm]  | 
|
1701  | 
in  | 
|
1702  | 
(thm, basis3)  | 
|
1703  | 
end  | 
|
1704  | 
| expand'' ectxt (ExpLn e) basis =  | 
|
1705  | 
let  | 
|
1706  | 
val (thm, basis', _, trimmed_thm) = trim_nz Pos_Trim ectxt e basis  | 
|
1707  | 
        val thm = @{thm expands_to_exp_ln} OF [trimmed_thm, get_basis_wf_thm basis', thm]
 | 
|
1708  | 
in  | 
|
1709  | 
(thm, basis')  | 
|
1710  | 
end  | 
|
1711  | 
| expand'' ectxt (Power (e, n)) basis =  | 
|
1712  | 
let  | 
|
1713  | 
val (thm, basis') = expand' ectxt e basis |> apfst (simplify_expansion ectxt)  | 
|
1714  | 
in  | 
|
1715  | 
(power_expansion ectxt (thm, n, basis'), basis')  | 
|
1716  | 
end  | 
|
1717  | 
| expand'' ectxt (Root (e, n)) basis =  | 
|
1718  | 
let  | 
|
1719  | 
val (thm, basis') = expand' ectxt e basis |> apfst (simplify_expansion ectxt)  | 
|
1720  | 
in  | 
|
1721  | 
(root_expansion ectxt (thm, n, basis'), basis')  | 
|
1722  | 
end  | 
|
1723  | 
| expand'' ectxt (Inverse e) basis =  | 
|
1724  | 
(case trim_nz Simple_Trim ectxt e basis of  | 
|
1725  | 
(thm, basis', _, trimmed_thm) =>  | 
|
1726  | 
           (@{thm expands_to_inverse} OF [trimmed_thm, get_basis_wf_thm basis', thm], basis'))
 | 
|
1727  | 
| expand'' ectxt (Div (e1, e2)) basis =  | 
|
1728  | 
let  | 
|
1729  | 
val (thm1, basis') = expand' ectxt e1 basis  | 
|
1730  | 
val (thm2, basis'', _, trimmed_thm) = trim_nz Simple_Trim ectxt e2 basis'  | 
|
1731  | 
val thm1 = lift basis'' thm1  | 
|
1732  | 
in  | 
|
1733  | 
        (@{thm expands_to_divide} OF [trimmed_thm, get_basis_wf_thm basis'', thm1, thm2], basis'')
 | 
|
1734  | 
end  | 
|
1735  | 
| expand'' ectxt (Ln e) basis =  | 
|
1736  | 
let  | 
|
1737  | 
val (thm, basis', _, trimmed_thm) = trim_nz Pos_Trim ectxt e basis  | 
|
1738  | 
in  | 
|
1739  | 
ln_expansion ectxt trimmed_thm thm basis'  | 
|
1740  | 
end  | 
|
1741  | 
| expand'' ectxt (Exp e) basis =  | 
|
1742  | 
let  | 
|
1743  | 
val (thm, basis') = expand' ectxt e basis  | 
|
1744  | 
in  | 
|
1745  | 
exp_expansion ectxt thm basis'  | 
|
1746  | 
end  | 
|
1747  | 
| expand'' ectxt (Absolute e) basis =  | 
|
1748  | 
let  | 
|
1749  | 
val (thm, basis', nz, trimmed_thm) = trim_nz Sgn_Trim ectxt e basis  | 
|
1750  | 
val thm' =  | 
|
1751  | 
case nz of  | 
|
1752  | 
            IsPos => @{thm expands_to_abs_pos} 
 | 
|
1753  | 
          | IsNeg => @{thm expands_to_abs_neg}
 | 
|
1754  | 
          | _ => raise TERM ("Unexpected trim result during expansion of abs", [])
 | 
|
1755  | 
in  | 
|
1756  | 
(thm' OF [trimmed_thm, get_basis_wf_thm basis', thm], basis')  | 
|
1757  | 
end  | 
|
1758  | 
| expand'' ectxt (Sgn e) basis =  | 
|
1759  | 
let  | 
|
1760  | 
val (thm, basis') = expand' ectxt e basis  | 
|
1761  | 
in  | 
|
1762  | 
(sgn_expansion ectxt (thm, basis'), basis')  | 
|
1763  | 
end  | 
|
1764  | 
| expand'' ectxt (Min (e1, e2)) basis = (  | 
|
1765  | 
case try_prove_ev_eq ectxt (apply2 expr_to_term (e1, e2)) of  | 
|
1766  | 
SOME eq_thm =>  | 
|
1767  | 
expand' ectxt e1 basis  | 
|
1768  | 
          |> apfst (fn thm => @{thm expands_to_min_eq} OF [thm, eq_thm])
 | 
|
1769  | 
| NONE =>  | 
|
1770  | 
let  | 
|
1771  | 
val (thm1, basis') = expand' ectxt e1 basis  | 
|
1772  | 
val (thm2, basis'') = expand' ectxt e2 basis'  | 
|
1773  | 
val thm1' = lift basis'' thm1  | 
|
1774  | 
in  | 
|
1775  | 
(min_expansion ectxt (thm1', thm2, basis''), basis'')  | 
|
1776  | 
end)  | 
|
1777  | 
| expand'' ectxt (Max (e1, e2)) basis = (  | 
|
1778  | 
case try_prove_ev_eq ectxt (apply2 expr_to_term (e1, e2)) of  | 
|
1779  | 
SOME eq_thm =>  | 
|
1780  | 
expand' ectxt e1 basis  | 
|
1781  | 
          |> apfst (fn thm => @{thm expands_to_max_eq} OF [thm, eq_thm])
 | 
|
1782  | 
| NONE =>  | 
|
1783  | 
let  | 
|
1784  | 
val (thm1, basis') = expand' ectxt e1 basis  | 
|
1785  | 
val (thm2, basis'') = expand' ectxt e2 basis'  | 
|
1786  | 
val thm1' = lift basis'' thm1  | 
|
1787  | 
in  | 
|
1788  | 
(max_expansion ectxt (thm1', thm2, basis''), basis'')  | 
|
1789  | 
end)  | 
|
1790  | 
| expand'' ectxt (Sin e) basis =  | 
|
1791  | 
let  | 
|
1792  | 
val (thm, basis', _, _) = trim_nz Simple_Trim ectxt e basis (* TODO could be relaxed *)  | 
|
1793  | 
in  | 
|
1794  | 
(sin_cos_expansion ectxt thm basis' |> fst, basis')  | 
|
1795  | 
end  | 
|
1796  | 
| expand'' ectxt (Cos e) basis =  | 
|
1797  | 
let  | 
|
1798  | 
val (thm, basis', _, _) = trim_nz Simple_Trim ectxt e basis (* TODO could be relaxed *)  | 
|
1799  | 
in  | 
|
1800  | 
(sin_cos_expansion ectxt thm basis' |> snd, basis')  | 
|
1801  | 
end  | 
|
1802  | 
| expand'' _ (Floor _) _ =  | 
|
1803  | 
      raise TERM ("floor not supported.", [])
 | 
|
1804  | 
| expand'' _ (Ceiling _) _ =  | 
|
1805  | 
      raise TERM ("ceiling not supported.", [])
 | 
|
1806  | 
| expand'' _ (Frac _) _ =  | 
|
1807  | 
      raise TERM ("frac not supported.", [])
 | 
|
1808  | 
| expand'' _ (NatMod _) _ =  | 
|
1809  | 
      raise TERM ("mod not supported.", [])
 | 
|
1810  | 
| expand'' ectxt (ArcTan e) basis =  | 
|
1811  | 
let  | 
|
1812  | 
(* TODO: what if it's zero *)  | 
|
1813  | 
val (thm, basis') = expand' ectxt e basis  | 
|
1814  | 
in  | 
|
1815  | 
(arctan_expansion ectxt basis' thm, basis')  | 
|
1816  | 
end  | 
|
1817  | 
| expand'' ectxt (Custom (name, t, args)) basis =  | 
|
1818  | 
let  | 
|
1819  | 
fun expand_args acc basis [] = (rev acc, basis)  | 
|
1820  | 
| expand_args acc basis (arg :: args) =  | 
|
1821  | 
case expand' ectxt arg basis of  | 
|
1822  | 
(thm, basis') => expand_args (thm :: acc) basis' args  | 
|
1823  | 
in  | 
|
1824  | 
case expand_custom (get_ctxt ectxt) name of  | 
|
1825  | 
          NONE => raise TERM ("Unsupported custom function: " ^ name, [t])
 | 
|
1826  | 
| SOME e => e ectxt t (expand_args [] basis args)  | 
|
1827  | 
end  | 
|
1828  | 
||
1829  | 
and expand' ectxt (e' as (Inverse e)) basis =  | 
|
1830  | 
let  | 
|
1831  | 
val t = expr_to_term e  | 
|
1832  | 
fun thm wf_thm len_thm =  | 
|
1833  | 
          @{thm expands_to_basic_inverse} OF [wf_thm, len_thm]
 | 
|
1834  | 
in  | 
|
1835  | 
if member abconv (get_basis_list basis) t then  | 
|
1836  | 
(expand_basic thm t basis, basis) |> apfst (check_expansion e')  | 
|
1837  | 
else  | 
|
1838  | 
expand'' ectxt e' basis |> apfst (check_expansion e')  | 
|
1839  | 
end  | 
|
1840  | 
| expand' ectxt (Div (e1, e2)) basis =  | 
|
1841  | 
let  | 
|
1842  | 
val (thm1, basis') = expand' ectxt e1 basis  | 
|
1843  | 
val t = expr_to_term e2  | 
|
1844  | 
fun thm wf_thm len_thm =  | 
|
1845  | 
          @{thm expands_to_basic_inverse} OF [wf_thm, len_thm]
 | 
|
1846  | 
in  | 
|
1847  | 
if member abconv (get_basis_list basis') t then  | 
|
1848  | 
          (@{thm expands_to_div'} OF [get_basis_wf_thm basis', thm1, expand_basic thm t basis'], 
 | 
|
1849  | 
basis')  | 
|
1850  | 
else  | 
|
1851  | 
let  | 
|
1852  | 
val (thm2, basis'', _, trimmed_thm) = trim_nz Simple_Trim ectxt e2 basis'  | 
|
1853  | 
val thm1 = lift basis'' thm1  | 
|
1854  | 
in  | 
|
1855  | 
            (@{thm expands_to_divide} OF [trimmed_thm, get_basis_wf_thm basis'', thm1, thm2], 
 | 
|
1856  | 
basis'')  | 
|
1857  | 
end  | 
|
1858  | 
end  | 
|
1859  | 
| expand' ectxt (e' as (Powr' (e, p))) basis =  | 
|
1860  | 
let  | 
|
1861  | 
val t = expr_to_term e  | 
|
1862  | 
val ctxt = get_ctxt ectxt  | 
|
1863  | 
fun thm wf_thm len_thm =  | 
|
1864  | 
(Drule.infer_instantiate' ctxt [NONE, NONE, SOME (Thm.cterm_of ctxt p)]  | 
|
1865  | 
            @{thm expands_to_basic_powr}) OF [wf_thm, len_thm]
 | 
|
1866  | 
in  | 
|
1867  | 
if member abconv (get_basis_list basis) t then  | 
|
1868  | 
(expand_basic thm t basis, basis) |> apfst (check_expansion e')  | 
|
1869  | 
else  | 
|
1870  | 
expand'' ectxt e' basis |> apfst (check_expansion e')  | 
|
1871  | 
end  | 
|
1872  | 
| expand' ectxt e basis =  | 
|
1873  | 
let  | 
|
1874  | 
val t = expr_to_term e  | 
|
1875  | 
        fun thm wf_thm len_thm = @{thm expands_to_basic} OF [wf_thm, len_thm]
 | 
|
1876  | 
in  | 
|
1877  | 
if member abconv (get_basis_list basis) t then  | 
|
1878  | 
(expand_basic thm t basis, basis) |> apfst (check_expansion e)  | 
|
1879  | 
else  | 
|
1880  | 
expand'' ectxt e basis |> apfst (check_expansion e)  | 
|
1881  | 
end  | 
|
1882  | 
||
1883  | 
fun expand ectxt e basis =  | 
|
1884  | 
expand' ectxt e basis  | 
|
1885  | 
|> apfst (simplify_expansion ectxt)  | 
|
1886  | 
|> apfst (check_expansion e)  | 
|
1887  | 
||
1888  | 
fun expand_term ectxt t basis =  | 
|
1889  | 
let  | 
|
1890  | 
val ctxt = get_ctxt ectxt  | 
|
1891  | 
val (e, eq_thm) = reify ctxt t  | 
|
1892  | 
val (thm, basis) = expand ectxt e basis  | 
|
1893  | 
in  | 
|
1894  | 
    (@{thm expands_to_meta_eq_cong'} OF [thm, eq_thm], basis)
 | 
|
1895  | 
end  | 
|
1896  | 
||
1897  | 
fun expand_terms ectxt ts basis =  | 
|
1898  | 
let  | 
|
1899  | 
val ctxt = get_ctxt ectxt  | 
|
1900  | 
val e_eq_thms = map (reify ctxt) ts  | 
|
1901  | 
fun step (e, eq_thm) (thms, basis) =  | 
|
1902  | 
let  | 
|
1903  | 
val (thm, basis) = expand' ectxt e basis  | 
|
1904  | 
        val thm = @{thm expands_to_meta_eq_cong'} OF [simplify_expansion ectxt thm, eq_thm]
 | 
|
1905  | 
in  | 
|
1906  | 
(thm :: thms, basis)  | 
|
1907  | 
end  | 
|
1908  | 
val (thms, basis) = fold step e_eq_thms ([], basis)  | 
|
1909  | 
fun lift thm = lift_expands_to_thm (mk_lifting (extract_basis_list thm) basis) thm  | 
|
1910  | 
in  | 
|
1911  | 
(map lift (rev thms), basis)  | 
|
1912  | 
end  | 
|
1913  | 
||
1914  | 
datatype limit =  | 
|
1915  | 
Zero_Limit of bool option  | 
|
1916  | 
| Finite_Limit of term  | 
|
1917  | 
| Infinite_Limit of bool option  | 
|
1918  | 
||
| 69597 | 1919  | 
fun is_empty_expansion (Const (\<^const_name>\<open>MS\<close>, _) $ Const (\<^const_name>\<open>MSLNil\<close>, _) $ _) = true  | 
| 68630 | 1920  | 
| is_empty_expansion _ = false  | 
1921  | 
||
1922  | 
fun limit_of_expansion_aux ectxt basis thm =  | 
|
1923  | 
let  | 
|
1924  | 
val n = length (get_basis_list basis)  | 
|
1925  | 
val (thm, res, e_thms) =  | 
|
| 69597 | 1926  | 
trim_expansion_while_greater false (SOME (replicate n \<^term>\<open>0::real\<close>)) true  | 
| 68630 | 1927  | 
(SOME Simple_Trim) ectxt (thm, basis)  | 
1928  | 
val trimmed_thm = case res of Trimmed (_, trimmed_thm) => trimmed_thm | _ => NONE  | 
|
1929  | 
val res = case res of Trimmed _ => GREATER | Aborted res => res  | 
|
1930  | 
val exp = get_expansion thm  | 
|
1931  | 
val _ = if res = GREATER andalso is_none trimmed_thm andalso not (is_empty_expansion exp) then  | 
|
1932  | 
              raise TERM ("limit_of_expansion", [get_expansion thm]) else ()
 | 
|
1933  | 
fun go thm _ _ [] = (  | 
|
1934  | 
case zeroness_oracle false (SOME Simple_Trim) ectxt (get_expansion thm) of  | 
|
1935  | 
            (IsZero, _) => (Zero_Limit NONE, @{thm expands_to_real_imp_filterlim} OF [thm])
 | 
|
| 69597 | 1936  | 
          | _ => (Finite_Limit \<^term>\<open>0::real\<close>, @{thm expands_to_real_imp_filterlim} OF [thm]))
 | 
| 68630 | 1937  | 
| go thm _ basis ((IsNeg, neg_thm) :: _) = (Zero_Limit NONE,  | 
1938  | 
          @{thm expands_to_neg_exponent_imp_filterlim} OF
 | 
|
1939  | 
            [thm, get_basis_wf_thm basis, neg_thm RS @{thm compare_reals_diff_sgnD(1)}])
 | 
|
1940  | 
| go thm trimmed_thm basis ((IsPos, pos_thm) :: _) = (Infinite_Limit NONE,  | 
|
1941  | 
          @{thm expands_to_pos_exponent_imp_filterlim} OF
 | 
|
1942  | 
[thm, the trimmed_thm, get_basis_wf_thm basis,  | 
|
1943  | 
             pos_thm RS @{thm compare_reals_diff_sgnD(3)}])
 | 
|
1944  | 
| go thm trimmed_thm basis ((IsZero, zero_thm) :: e_thms) =  | 
|
1945  | 
let  | 
|
1946  | 
             val thm' = thm RS @{thm expands_to_hd''}
 | 
|
1947  | 
             val trimmed_thm' = Option.map (fn thm => thm RS @{thm trimmed_hd}) trimmed_thm
 | 
|
1948  | 
val (lim, lim_thm) = go thm' trimmed_thm' (tl_basis basis) e_thms  | 
|
1949  | 
val lim_lift_thm =  | 
|
1950  | 
case lim of  | 
|
1951  | 
                  Infinite_Limit _ => @{thm expands_to_zero_exponent_imp_filterlim(1)}
 | 
|
1952  | 
                | _ => @{thm expands_to_zero_exponent_imp_filterlim(2)}
 | 
|
1953  | 
val lim_thm' =  | 
|
1954  | 
lim_lift_thm OF [thm, get_basis_wf_thm basis,  | 
|
1955  | 
                 zero_thm RS @{thm compare_reals_diff_sgnD(2)}, lim_thm]
 | 
|
1956  | 
in  | 
|
1957  | 
(lim, lim_thm')  | 
|
1958  | 
end  | 
|
1959  | 
| go _ _ _ _ = raise Match  | 
|
1960  | 
in  | 
|
1961  | 
if is_empty_expansion exp then  | 
|
1962  | 
      (Zero_Limit NONE, thm RS @{thm expands_to_MSLNil_imp_filterlim}, thm)
 | 
|
1963  | 
else  | 
|
1964  | 
case go thm trimmed_thm basis e_thms of  | 
|
1965  | 
(lim, lim_thm) => (lim, lim_thm, thm)  | 
|
1966  | 
end  | 
|
1967  | 
||
1968  | 
(*  | 
|
1969  | 
Determines the limit of a function from its expansion. The two flags control whether the  | 
|
1970  | 
the sign of the approach should be determined for the infinite case (i.e. at_top/at_bot instead  | 
|
1971  | 
of just at_infinity) and the zero case (i.e. at_right 0/at_left 0 instead of just nhds 0)  | 
|
1972  | 
*)  | 
|
1973  | 
fun limit_of_expansion (sgn_zero, sgn_inf) ectxt (thm, basis) =  | 
|
1974  | 
let  | 
|
1975  | 
val (lim, lim_thm, thm) = limit_of_expansion_aux ectxt basis thm  | 
|
1976  | 
in  | 
|
1977  | 
case lim of  | 
|
1978  | 
Zero_Limit _ => (  | 
|
1979  | 
if sgn_zero then  | 
|
1980  | 
case trim_expansion false (SOME Sgn_Trim) ectxt (thm, basis) of  | 
|
1981  | 
(thm, IsPos, SOME pos_thm) => (Zero_Limit (SOME true),  | 
|
1982  | 
              @{thm tendsto_imp_filterlim_at_right[OF _ expands_to_imp_eventually_pos]} OF
 | 
|
1983  | 
[lim_thm, get_basis_wf_thm basis, thm, pos_thm])  | 
|
1984  | 
| (thm, IsNeg, SOME neg_thm) => (Zero_Limit (SOME false),  | 
|
1985  | 
              @{thm tendsto_imp_filterlim_at_left[OF _ expands_to_imp_eventually_neg]} OF
 | 
|
1986  | 
[lim_thm, get_basis_wf_thm basis, thm, neg_thm])  | 
|
1987  | 
| _ => (Zero_Limit NONE, lim_thm)  | 
|
1988  | 
else (Zero_Limit NONE, lim_thm))  | 
|
1989  | 
| Infinite_Limit _ => (  | 
|
1990  | 
if sgn_inf then  | 
|
1991  | 
case trim_expansion false (SOME Sgn_Trim) ectxt (thm, basis) of  | 
|
1992  | 
(thm, IsPos, SOME pos_thm) => (Infinite_Limit (SOME true),  | 
|
1993  | 
              (@{thm filterlim_at_infinity_imp_filterlim_at_top[OF _ expands_to_imp_eventually_pos]} OF
 | 
|
1994  | 
[lim_thm, get_basis_wf_thm basis, thm, pos_thm]))  | 
|
1995  | 
| (thm, IsNeg, SOME neg_thm) => (Infinite_Limit (SOME false),  | 
|
1996  | 
              @{thm filterlim_at_infinity_imp_filterlim_at_bot[OF _ expands_to_imp_eventually_neg]} OF
 | 
|
1997  | 
[lim_thm, get_basis_wf_thm basis, thm, neg_thm])  | 
|
1998  | 
| _ => (Infinite_Limit NONE, lim_thm)  | 
|
1999  | 
else (Infinite_Limit NONE, lim_thm))  | 
|
2000  | 
| Finite_Limit c => (Finite_Limit c, lim_thm)  | 
|
2001  | 
end  | 
|
2002  | 
||
2003  | 
fun compute_limit ectxt t =  | 
|
2004  | 
case expand_term ectxt t default_basis of  | 
|
2005  | 
(thm, basis) => limit_of_expansion (true, true) ectxt (thm, basis)  | 
|
2006  | 
||
2007  | 
fun prove_at_infinity ectxt (thm, basis) =  | 
|
2008  | 
let  | 
|
2009  | 
    fun err () = raise TERM ("prove_at_infinity", [get_expanded_fun thm])
 | 
|
2010  | 
val (thm, _, SOME trimmed_thm) = trim_expansion true (SOME Simple_Trim) ectxt (thm, basis)  | 
|
2011  | 
fun go basis thm trimmed_thm =  | 
|
| 69597 | 2012  | 
if fastype_of (get_expansion thm) = \<^typ>\<open>real\<close> then  | 
| 68630 | 2013  | 
err ()  | 
2014  | 
else  | 
|
2015  | 
case zeroness_oracle true (SOME Pos_Trim) ectxt (get_exponent (get_expansion thm)) of  | 
|
2016  | 
(IsPos, SOME pos_thm) =>  | 
|
2017  | 
            @{thm expands_to_pos_exponent_imp_filterlim} OF
 | 
|
2018  | 
[thm, trimmed_thm, get_basis_wf_thm basis, pos_thm]  | 
|
2019  | 
| (IsZero, SOME zero_thm) =>  | 
|
2020  | 
            @{thm expands_to_zero_exponent_imp_filterlim(1)} OF
 | 
|
2021  | 
[thm, get_basis_wf_thm basis, zero_thm,  | 
|
2022  | 
                 go (tl_basis basis) (thm RS @{thm expands_to_hd''})
 | 
|
2023  | 
                   (trimmed_thm RS @{thm trimmed_hd})]
 | 
|
2024  | 
| _ => err ()  | 
|
2025  | 
in  | 
|
2026  | 
go basis thm trimmed_thm  | 
|
2027  | 
end  | 
|
2028  | 
||
2029  | 
fun prove_at_top_at_bot mode ectxt (thm, basis) =  | 
|
2030  | 
let  | 
|
2031  | 
val s = if mode = Pos_Trim then "prove_at_top" else "prove_at_bot"  | 
|
2032  | 
fun err () = raise TERM (s, [get_expanded_fun thm])  | 
|
2033  | 
val (thm, _, SOME trimmed_thm) = trim_expansion true (SOME mode) ectxt (thm, basis)  | 
|
2034  | 
val trimmed_thm' = trimmed_thm RS  | 
|
2035  | 
      (if mode = Pos_Trim then @{thm trimmed_pos_imp_trimmed} else @{thm trimmed_neg_imp_trimmed})
 | 
|
2036  | 
fun go basis thm trimmed_thm =  | 
|
| 69597 | 2037  | 
if fastype_of (get_expansion thm) = \<^typ>\<open>real\<close> then  | 
| 68630 | 2038  | 
err ()  | 
2039  | 
else  | 
|
2040  | 
case zeroness_oracle true (SOME Pos_Trim) ectxt (get_exponent (get_expansion thm)) of  | 
|
2041  | 
(IsPos, SOME pos_thm) =>  | 
|
2042  | 
            @{thm expands_to_pos_exponent_imp_filterlim} OF
 | 
|
2043  | 
[thm, trimmed_thm, get_basis_wf_thm basis, pos_thm]  | 
|
2044  | 
| (IsZero, SOME zero_thm) =>  | 
|
2045  | 
            @{thm expands_to_zero_exponent_imp_filterlim(1)} OF
 | 
|
2046  | 
[thm, get_basis_wf_thm basis, zero_thm,  | 
|
2047  | 
                 go (tl_basis basis) (thm RS @{thm expands_to_hd''})
 | 
|
2048  | 
                   (trimmed_thm RS @{thm trimmed_hd})]
 | 
|
2049  | 
| _ => err ()  | 
|
2050  | 
val lim_thm = go basis thm trimmed_thm'  | 
|
2051  | 
val add_sign_thm =  | 
|
2052  | 
if mode = Pos_Trim then  | 
|
2053  | 
        @{thm filterlim_at_infinity_imp_filterlim_at_top[OF _ expands_to_imp_eventually_pos]}
 | 
|
2054  | 
else  | 
|
2055  | 
        @{thm filterlim_at_infinity_imp_filterlim_at_bot[OF _ expands_to_imp_eventually_neg]}
 | 
|
2056  | 
in  | 
|
2057  | 
add_sign_thm OF [lim_thm, get_basis_wf_thm basis, thm, trimmed_thm]  | 
|
2058  | 
end  | 
|
2059  | 
||
2060  | 
val prove_at_top = prove_at_top_at_bot Pos_Trim  | 
|
2061  | 
val prove_at_bot = prove_at_top_at_bot Neg_Trim  | 
|
2062  | 
||
2063  | 
||
2064  | 
fun prove_at_aux mode ectxt (thm, basis) =  | 
|
2065  | 
let  | 
|
2066  | 
val (s, add_sign_thm) =  | 
|
2067  | 
case mode of  | 
|
2068  | 
Simple_Trim =>  | 
|
2069  | 
          ("prove_at_0", @{thm Topological_Spaces.filterlim_atI[OF _ expands_to_imp_eventually_nz]})
 | 
|
2070  | 
| Pos_Trim =>  | 
|
2071  | 
          ("prove_at_right_0",
 | 
|
2072  | 
             @{thm tendsto_imp_filterlim_at_right[OF _ expands_to_imp_eventually_pos]})
 | 
|
2073  | 
| Neg_Trim =>  | 
|
2074  | 
          ("prove_at_left_0",
 | 
|
2075  | 
             @{thm tendsto_imp_filterlim_at_left[OF _ expands_to_imp_eventually_neg]})
 | 
|
2076  | 
fun err () = raise TERM (s, [get_expanded_fun thm])  | 
|
2077  | 
val (thm, _, SOME trimmed_thm) = trim_expansion true (SOME mode) ectxt (thm, basis)  | 
|
2078  | 
fun go basis thm =  | 
|
| 69597 | 2079  | 
if fastype_of (get_expansion thm) = \<^typ>\<open>real\<close> then  | 
| 68630 | 2080  | 
err ()  | 
2081  | 
else  | 
|
2082  | 
case zeroness_oracle true (SOME Neg_Trim) ectxt (get_exponent (get_expansion thm)) of  | 
|
2083  | 
(IsNeg, SOME neg_thm) =>  | 
|
2084  | 
            @{thm expands_to_neg_exponent_imp_filterlim} OF
 | 
|
2085  | 
[thm, get_basis_wf_thm basis, neg_thm]  | 
|
2086  | 
| (IsZero, SOME zero_thm) =>  | 
|
2087  | 
            @{thm expands_to_zero_exponent_imp_filterlim(2)} OF
 | 
|
2088  | 
[thm, get_basis_wf_thm basis, zero_thm,  | 
|
2089  | 
                 go (tl_basis basis) (thm RS @{thm expands_to_hd''})]
 | 
|
2090  | 
| _ => err ()  | 
|
2091  | 
val lim_thm = go basis thm  | 
|
2092  | 
in  | 
|
2093  | 
add_sign_thm OF [lim_thm, get_basis_wf_thm basis, thm, trimmed_thm]  | 
|
2094  | 
end  | 
|
2095  | 
||
2096  | 
val prove_at_0 = prove_at_aux Simple_Trim  | 
|
2097  | 
val prove_at_left_0 = prove_at_aux Neg_Trim  | 
|
2098  | 
val prove_at_right_0 = prove_at_aux Pos_Trim  | 
|
2099  | 
||
2100  | 
||
2101  | 
fun prove_nhds ectxt (thm, basis) =  | 
|
2102  | 
let  | 
|
2103  | 
fun simplify (a, b, c) = (a, simplify_expansion ectxt b, c)  | 
|
2104  | 
fun go thm basis =  | 
|
| 69597 | 2105  | 
if fastype_of (get_expansion thm) = \<^typ>\<open>real\<close> then  | 
| 68630 | 2106  | 
        @{thm expands_to_real_imp_filterlim} OF [thm]
 | 
2107  | 
else  | 
|
2108  | 
case whnf_expansion ectxt thm |> simplify of  | 
|
2109  | 
          (NONE, thm, _) => @{thm expands_to_MSLNil_imp_filterlim} OF [thm]
 | 
|
2110  | 
| (SOME _, thm, _) => (  | 
|
2111  | 
case zeroness_oracle true (SOME Sgn_Trim) ectxt (get_exponent (get_expansion thm)) of  | 
|
2112  | 
(IsZero, SOME zero_thm) =>  | 
|
2113  | 
                @{thm expands_to_zero_exponent_imp_filterlim(2)} OF
 | 
|
2114  | 
[thm, get_basis_wf_thm basis, zero_thm,  | 
|
2115  | 
                    go (thm RS @{thm expands_to_hd''}) (tl_basis basis)]
 | 
|
2116  | 
| (IsNeg, SOME neg_thm) =>  | 
|
2117  | 
                @{thm expands_to_neg_exponent_imp_filterlim} OF
 | 
|
2118  | 
[thm, get_basis_wf_thm basis, neg_thm]  | 
|
2119  | 
| (IsPos, _) =>  | 
|
2120  | 
go (try_drop_leading_term ectxt thm) basis  | 
|
2121  | 
            | _ => raise TERM ("Unexpected zeroness result in prove_nhds",
 | 
|
2122  | 
[get_exponent (get_expansion thm)]))  | 
|
2123  | 
in  | 
|
2124  | 
go thm basis  | 
|
2125  | 
end  | 
|
2126  | 
||
2127  | 
fun prove_equivalent theta ectxt (thm1, thm2, basis) =  | 
|
2128  | 
let  | 
|
2129  | 
val ((thm1, _, SOME trimmed_thm1), (thm2, _, SOME trimmed_thm2)) =  | 
|
2130  | 
apply2 (trim_expansion true (SOME Simple_Trim) ectxt) ((thm1, basis), (thm2, basis))  | 
|
| 69597 | 2131  | 
val pat = ConsPat (\<^const_name>\<open>Pair\<close>, [ConsPat (\<^const_name>\<open>Lazy_Eval.cmp_result.EQ\<close>, []),  | 
2132  | 
                ConsPat (\<^const_name>\<open>Pair\<close>, [AnyPat ("_", 0), AnyPat ("_", 0)])])
 | 
|
| 68630 | 2133  | 
val (exp1, exp2) = apply2 get_expansion (thm1, thm2)  | 
2134  | 
val T = fastype_of exp1  | 
|
2135  | 
val t = mk_compare_expansions_const T $ exp1 $ exp2  | 
|
2136  | 
fun eq_thm conv = HOLogic.mk_obj_eq (conv (Thm.cterm_of (get_ctxt ectxt) t))  | 
|
2137  | 
val imp_thm =  | 
|
2138  | 
      if theta then @{thm compare_expansions_EQ_imp_bigtheta}
 | 
|
2139  | 
      else @{thm compare_expansions_EQ_same}
 | 
|
2140  | 
in  | 
|
2141  | 
case match ectxt pat t (SOME []) of  | 
|
2142  | 
(SOME _, t, conv) =>  | 
|
2143  | 
let  | 
|
2144  | 
val [_, c1, c2] = HOLogic.strip_tuple t  | 
|
2145  | 
val c12_thm = if theta then [] else [the (try_prove_real_eq true ectxt (c1, c2))]  | 
|
2146  | 
in  | 
|
2147  | 
imp_thm OF ([eq_thm conv, trimmed_thm1, trimmed_thm2, thm1, thm2, get_basis_wf_thm basis]  | 
|
2148  | 
@ c12_thm)  | 
|
2149  | 
end  | 
|
2150  | 
    | _ => raise TERM ("prove_equivalent", map get_expanded_fun [thm1, thm2])
 | 
|
2151  | 
end  | 
|
2152  | 
||
2153  | 
val prove_bigtheta = prove_equivalent true  | 
|
2154  | 
val prove_asymp_equiv = prove_equivalent false  | 
|
2155  | 
||
2156  | 
fun print_trimming_error s ectxt exp =  | 
|
2157  | 
let  | 
|
2158  | 
val c = get_coeff exp  | 
|
| 69597 | 2159  | 
val t = if fastype_of c = \<^typ>\<open>real\<close> then c else get_eval c  | 
| 68630 | 2160  | 
in  | 
2161  | 
if #verbose (#ctxt ectxt) then  | 
|
2162  | 
let  | 
|
2163  | 
val ctxt = get_ctxt ectxt  | 
|
2164  | 
val p = Pretty.str "real_asymp failed to show zeroness of the following expression:"  | 
|
2165  | 
val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt t)]  | 
|
2166  | 
val _ = Pretty.writeln p  | 
|
2167  | 
in  | 
|
2168  | 
raise TERM (s, [t])  | 
|
2169  | 
end  | 
|
2170  | 
else  | 
|
2171  | 
raise TERM (s, [t])  | 
|
2172  | 
end  | 
|
2173  | 
||
2174  | 
fun prove_smallo ectxt (thm1, thm2, basis) =  | 
|
2175  | 
let  | 
|
2176  | 
val (thm2, _, SOME trimmed_thm) = trim_expansion true (SOME Simple_Trim) ectxt (thm2, basis)  | 
|
2177  | 
val es = get_exponents (get_expansion thm2)  | 
|
2178  | 
in  | 
|
2179  | 
case trim_expansion_while_greater true (SOME es) false NONE ectxt (thm1, basis) of  | 
|
2180  | 
(thm1, Aborted LESS, thms) =>  | 
|
2181  | 
        @{thm compare_expansions_LT} OF [prove_compare_expansions LESS (map snd thms),
 | 
|
2182  | 
trimmed_thm, thm1, thm2, get_basis_wf_thm basis]  | 
|
2183  | 
| (thm1, Aborted _, _) =>  | 
|
2184  | 
print_trimming_error "prove_smallo" ectxt (get_expansion thm1)  | 
|
2185  | 
| (thm1, Trimmed _, _) =>  | 
|
2186  | 
print_trimming_error "prove_smallo" ectxt (get_expansion thm1)  | 
|
2187  | 
end  | 
|
2188  | 
||
2189  | 
fun prove_bigo ectxt (thm1, thm2, basis) =  | 
|
2190  | 
let  | 
|
2191  | 
val (thm2, _, SOME trimmed_thm) = trim_expansion true (SOME Simple_Trim) ectxt (thm2, basis)  | 
|
2192  | 
val es = get_exponents (get_expansion thm2)  | 
|
2193  | 
in  | 
|
2194  | 
case trim_expansion_while_greater false (SOME es) false NONE ectxt (thm1, basis) of  | 
|
2195  | 
(thm1, Aborted LESS, thms) =>  | 
|
2196  | 
        @{thm landau_o.small_imp_big[OF compare_expansions_LT]} OF
 | 
|
2197  | 
[prove_compare_expansions LESS (map snd thms), trimmed_thm, thm1, thm2,  | 
|
2198  | 
get_basis_wf_thm basis]  | 
|
2199  | 
| (thm1, Aborted EQ, thms) =>  | 
|
2200  | 
        @{thm compare_expansions_EQ_imp_bigo} OF [prove_compare_expansions EQ (map snd thms),
 | 
|
2201  | 
trimmed_thm, thm1, thm2, get_basis_wf_thm basis]  | 
|
2202  | 
| (thm1, Trimmed _, _) =>  | 
|
2203  | 
print_trimming_error "prove_bigo" ectxt (get_expansion thm1)  | 
|
2204  | 
end  | 
|
2205  | 
||
2206  | 
||
2207  | 
fun prove_asymptotic_relation_aux mode f ectxt (thm1, thm2, basis) = f (  | 
|
2208  | 
let  | 
|
2209  | 
    val thm = @{thm expands_to_minus} OF [get_basis_wf_thm basis, thm1, thm2]
 | 
|
2210  | 
in  | 
|
2211  | 
case ev_zeroness_oracle ectxt (get_expanded_fun thm) of  | 
|
2212  | 
      SOME zero_thm => (EQUAL, zero_thm RS @{thm eventually_diff_zero_imp_eq})
 | 
|
2213  | 
| _ => (  | 
|
2214  | 
case trim_expansion true (SOME mode) ectxt (thm, basis) of  | 
|
2215  | 
(thm, IsPos, SOME pos_thm) =>  | 
|
2216  | 
          (GREATER, @{thm expands_to_imp_eventually_gt} OF [get_basis_wf_thm basis, thm, pos_thm])
 | 
|
2217  | 
| (thm, IsNeg, SOME neg_thm) =>  | 
|
2218  | 
          (LESS, @{thm expands_to_imp_eventually_lt} OF [get_basis_wf_thm basis, thm, neg_thm])
 | 
|
2219  | 
      | _ => raise TERM ("Unexpected zeroness result in prove_asymptotic_relation", []))
 | 
|
2220  | 
end)  | 
|
2221  | 
||
2222  | 
val prove_eventually_greater = prove_asymptotic_relation_aux Pos_Trim snd  | 
|
2223  | 
val prove_eventually_less = prove_asymptotic_relation_aux Neg_Trim snd  | 
|
2224  | 
val prove_asymptotic_relation = prove_asymptotic_relation_aux Sgn_Trim I  | 
|
2225  | 
||
2226  | 
fun prove_eventually_nonzero ectxt (thm, basis) =  | 
|
2227  | 
case trim_expansion true (SOME Simple_Trim) ectxt (thm, basis) of  | 
|
2228  | 
(thm, _, SOME trimmed_thm) =>  | 
|
2229  | 
      @{thm expands_to_imp_eventually_nz} OF [get_basis_wf_thm basis, thm, trimmed_thm]
 | 
|
2230  | 
  | _ => raise TERM ("prove_eventually_nonzero", [get_expanded_fun thm])
 | 
|
2231  | 
||
2232  | 
fun extract_terms (n, strict) ectxt basis t =  | 
|
2233  | 
let  | 
|
2234  | 
val bs = get_basis_list basis  | 
|
| 69597 | 2235  | 
    fun mk_constfun c = (Abs ("x", \<^typ>\<open>real\<close>, c))
 | 
2236  | 
val const_0 = mk_constfun \<^term>\<open>0 :: real\<close>  | 
|
2237  | 
val const_1 = mk_constfun \<^term>\<open>1 :: real\<close>  | 
|
2238  | 
fun uminus t = Term.betapply (\<^term>\<open>\<lambda>(f::real\<Rightarrow>real) x. -f x\<close>, t)  | 
|
| 68630 | 2239  | 
fun betapply2 a b c = Term.betapply (Term.betapply (a, b), c)  | 
2240  | 
||
2241  | 
fun mk_sum' [] acc = acc  | 
|
2242  | 
| mk_sum' ((t, sgn) :: ts) acc = mk_sum' ts (  | 
|
2243  | 
if sgn then  | 
|
| 69597 | 2244  | 
betapply2 \<^term>\<open>%(f::real=>real) g x. f x - g x\<close> acc t  | 
| 68630 | 2245  | 
else  | 
| 69597 | 2246  | 
betapply2 \<^term>\<open>%(f::real=>real) g x. f x + g x\<close> acc t)  | 
| 68630 | 2247  | 
fun mk_sum [] = const_0  | 
2248  | 
| mk_sum ((t, sgn) :: ts) = mk_sum' ts (if sgn then uminus t else t)  | 
|
2249  | 
||
2250  | 
fun mk_mult a b =  | 
|
2251  | 
if a aconv const_0 then  | 
|
2252  | 
const_0  | 
|
2253  | 
else if b aconv const_0 then  | 
|
2254  | 
const_0  | 
|
| 69597 | 2255  | 
else if a aconv \<^term>\<open>\<lambda>_::real. 1 :: real\<close> then  | 
| 68630 | 2256  | 
b  | 
| 69597 | 2257  | 
else if b aconv \<^term>\<open>\<lambda>_::real. 1 :: real\<close> then  | 
| 68630 | 2258  | 
a  | 
| 69597 | 2259  | 
else if a aconv \<^term>\<open>\<lambda>_::real. -1 :: real\<close> then  | 
2260  | 
Term.betapply (\<^term>\<open>\<lambda>(f::real\<Rightarrow>real) x. -f x\<close>, b)  | 
|
2261  | 
else if b aconv \<^term>\<open>\<lambda>_::real. -1 :: real\<close> then  | 
|
2262  | 
Term.betapply (\<^term>\<open>\<lambda>(f::real\<Rightarrow>real) x. -f x\<close>, a)  | 
|
| 68630 | 2263  | 
else  | 
| 69597 | 2264  | 
        Abs ("x", \<^typ>\<open>real\<close>, \<^term>\<open>(*) :: real => _\<close> $
 | 
| 68630 | 2265  | 
(Term.betapply (a, Bound 0)) $ (Term.betapply (b, Bound 0)))  | 
2266  | 
||
2267  | 
fun mk_powr b e =  | 
|
| 69597 | 2268  | 
if e = \<^term>\<open>0 :: real\<close> then  | 
| 68630 | 2269  | 
const_1  | 
2270  | 
else  | 
|
2271  | 
let  | 
|
2272  | 
val n = HOLogic.dest_number e |> snd  | 
|
2273  | 
in  | 
|
2274  | 
if n >= 0 then  | 
|
| 69597 | 2275  | 
Term.betapply (Term.betapply (\<^term>\<open>%(b::real=>real) e x. b x ^ e\<close>, b),  | 
2276  | 
HOLogic.mk_number \<^typ>\<open>nat\<close> n)  | 
|
| 68630 | 2277  | 
else  | 
| 69597 | 2278  | 
Term.betapply (Term.betapply (\<^term>\<open>%(b::real=>real) e x. b x powr e\<close>, b), e)  | 
| 68630 | 2279  | 
end  | 
2280  | 
handle TERM _ =>  | 
|
| 69597 | 2281  | 
Term.betapply (Term.betapply (\<^term>\<open>%(b::real=>real) e x. b x powr e\<close>, b), e)  | 
| 68630 | 2282  | 
|
2283  | 
fun mk_scale_elem b e acc =  | 
|
2284  | 
let  | 
|
2285  | 
val e = simplify_term ectxt e  | 
|
2286  | 
in  | 
|
| 69597 | 2287  | 
if e = \<^term>\<open>0 :: real\<close> then  | 
| 68630 | 2288  | 
acc  | 
| 69597 | 2289  | 
else if e = \<^term>\<open>1 :: real\<close> then  | 
| 68630 | 2290  | 
mk_mult acc b  | 
2291  | 
else  | 
|
2292  | 
mk_mult acc (mk_powr b e)  | 
|
2293  | 
end  | 
|
2294  | 
||
2295  | 
fun mk_scale_elems [] _ acc = acc  | 
|
2296  | 
| mk_scale_elems (b :: bs) (e :: es) acc =  | 
|
2297  | 
mk_scale_elems bs es (mk_scale_elem b e acc)  | 
|
2298  | 
| mk_scale_elems _ _ _ = raise Match  | 
|
2299  | 
||
2300  | 
fun mk_summand c es =  | 
|
2301  | 
let  | 
|
| 69597 | 2302  | 
val es = mk_scale_elems bs es \<^term>\<open>\<lambda>_::real. 1 :: real\<close>  | 
| 68630 | 2303  | 
in  | 
2304  | 
case c of  | 
|
| 69597 | 2305  | 
Const (\<^const_name>\<open>uminus\<close>, _) $ c => ((c, true), es)  | 
| 68630 | 2306  | 
| _ => ((c, false), es)  | 
2307  | 
end  | 
|
2308  | 
||
2309  | 
fun go _ _ _ acc 0 = (acc, 0)  | 
|
2310  | 
| go 0 es t acc n =  | 
|
2311  | 
let  | 
|
2312  | 
val c = simplify_term ectxt t  | 
|
2313  | 
in  | 
|
| 69597 | 2314  | 
if strict andalso c = \<^term>\<open>0 :: real\<close> then  | 
| 68630 | 2315  | 
(acc, n)  | 
2316  | 
else  | 
|
2317  | 
(mk_summand c (rev es) :: acc, n - 1)  | 
|
2318  | 
end  | 
|
2319  | 
| go m es t acc n =  | 
|
2320  | 
case Lazy_Eval.whnf ectxt t |> fst of  | 
|
| 69597 | 2321  | 
Const (\<^const_name>\<open>MS\<close>, _) $ cs $ _ =>  | 
| 68630 | 2322  | 
go' m es (simplify_term ectxt cs) acc n  | 
2323  | 
          | _ => raise TERM("extract_terms", [t])
 | 
|
2324  | 
and go' _ _ _ acc 0 = (acc, 0)  | 
|
2325  | 
| go' m es cs acc n =  | 
|
2326  | 
case Lazy_Eval.whnf ectxt cs |> fst of  | 
|
| 69597 | 2327  | 
Const (\<^const_name>\<open>MSLNil\<close>, _) => (acc, n)  | 
2328  | 
| Const (\<^const_name>\<open>MSLCons\<close>, _) $ c $ cs => (  | 
|
| 68630 | 2329  | 
case Lazy_Eval.whnf ectxt c |> fst |> HOLogic.dest_prod of  | 
2330  | 
(c, e) =>  | 
|
2331  | 
case go (m - 1) (e :: es) c acc n of  | 
|
2332  | 
(acc, n) => go' m es (simplify_term ectxt cs) acc n)  | 
|
2333  | 
          | _ => raise TERM("extract_terms", [t])
 | 
|
2334  | 
val (summands, remaining) = go (basis_size basis) [] t [] (n + 1)  | 
|
2335  | 
val (summands, error) =  | 
|
2336  | 
if remaining = 0 then (rev (tl summands), SOME (snd (hd summands))) else (rev summands, NONE)  | 
|
2337  | 
val summands = map (fn ((c, sgn), es) => (mk_mult (mk_constfun c) es, sgn)) summands  | 
|
| 69597 | 2338  | 
val error = Option.map (fn err => Term.betapply (\<^term>\<open>\<lambda>f::real\<Rightarrow>real. O(f)\<close>, err)) error  | 
| 68630 | 2339  | 
val expansion = mk_sum summands  | 
2340  | 
in  | 
|
2341  | 
(expansion, error)  | 
|
2342  | 
end  | 
|
2343  | 
||
2344  | 
end  | 
|
2345  | 
||
2346  | 
||
2347  | 
structure Multiseries_Expansion_Basic : EXPANSION_INTERFACE =  | 
|
2348  | 
struct  | 
|
2349  | 
open Multiseries_Expansion;  | 
|
2350  | 
||
2351  | 
type T = expansion_thm  | 
|
2352  | 
||
2353  | 
val expand_term = expand_term  | 
|
2354  | 
val expand_terms = expand_terms  | 
|
2355  | 
||
2356  | 
val prove_nhds = prove_nhds  | 
|
2357  | 
val prove_at_infinity = prove_at_infinity  | 
|
2358  | 
val prove_at_top = prove_at_top  | 
|
2359  | 
val prove_at_bot = prove_at_bot  | 
|
2360  | 
val prove_at_0 = prove_at_0  | 
|
2361  | 
val prove_at_right_0 = prove_at_right_0  | 
|
2362  | 
val prove_at_left_0 = prove_at_left_0  | 
|
2363  | 
val prove_eventually_nonzero = prove_eventually_nonzero  | 
|
2364  | 
||
2365  | 
val prove_eventually_less = prove_eventually_less  | 
|
2366  | 
val prove_eventually_greater = prove_eventually_greater  | 
|
2367  | 
||
2368  | 
val prove_smallo = prove_smallo  | 
|
2369  | 
val prove_bigo = prove_bigo  | 
|
2370  | 
val prove_bigtheta = prove_bigtheta  | 
|
2371  | 
val prove_asymp_equiv = prove_asymp_equiv  | 
|
2372  | 
||
2373  | 
end  |