diff --git a/Formal_ineqs/arith/arith_cache.hl b/Formal_ineqs/arith/arith_cache.hl index 43f56e92..9a54c540 100644 --- a/Formal_ineqs/arith/arith_cache.hl +++ b/Formal_ineqs/arith/arith_cache.hl @@ -1,12 +1,18 @@ -(* =========================================================== *) -(* Cached natural arithmetic *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Cached natural arithmetic *) +(* -------------------------------------------------------------------------- *) (* Dependencies *) -needs "arith_options.hl";; -needs "arith/arith_num.hl";; +needs "Formal_ineqs/arith_options.hl";; +needs "Formal_ineqs/arith/arith_num.hl";; module Arith_cache = struct @@ -109,34 +115,34 @@ let tm1_tm2_hash tm1 tm2 = (* SUC *) -let raw_suc_conv_hash tm = +let raw_suc_conv_hash tm = let _ = suc_counter := !suc_counter + 1 in (* let _ = suc_list := tm :: !suc_list in *) Arith_num.raw_suc_conv_hash tm;; (* x = 0 *) -let raw_eq0_hash_conv tm = +let raw_eq0_hash_conv tm = let _ = eq0_counter := !eq0_counter + 1 in (* let _ = eq0_list := tm :: !eq0_list in *) Arith_num.raw_eq0_hash_conv tm;; (* PRE *) -let raw_pre_hash_conv tm = +let raw_pre_hash_conv tm = let _ = pre_counter := !pre_counter + 1 in Arith_num.raw_pre_hash_conv tm;; (* x > 0 *) -let raw_gt0_hash_conv tm = +let raw_gt0_hash_conv tm = let _ = gt0_counter := !gt0_counter + 1 in Arith_num.raw_gt0_hash_conv tm;; (* x < y *) -let raw_lt_hash_conv tm = +let raw_lt_hash_conv tm = let _ = lt_counter := !lt_counter + 1 in Arith_num.raw_lt_hash_conv tm;; (* x <= y *) -let raw_le_hash_conv tm = +let raw_le_hash_conv tm = let _ = le_counter := !le_counter + 1 in let hash = op_tm_hash tm in try @@ -147,7 +153,7 @@ let raw_le_hash_conv tm = result;; (* x + y *) -let raw_add_conv_hash tm = +let raw_add_conv_hash tm = let _ = add_counter := !add_counter + 1 in let hash = op_tm_hash tm in try @@ -158,7 +164,7 @@ let raw_add_conv_hash tm = result;; (* x - y *) -let raw_sub_hash_conv tm = +let raw_sub_hash_conv tm = let _ = sub_counter := !sub_counter + 1 in let hash = op_tm_hash tm in try @@ -168,7 +174,7 @@ let raw_sub_hash_conv tm = let _ = my_add sub_table hash result in result;; -let raw_sub_and_le_hash_conv tm1 tm2 = +let raw_sub_and_le_hash_conv tm1 tm2 = let _ = sub_le_counter := !sub_le_counter + 1 in let hash = tm1_tm2_hash tm1 tm2 in try @@ -179,7 +185,7 @@ let raw_sub_and_le_hash_conv tm1 tm2 = result;; (* x * y *) -let raw_mul_conv_hash tm = +let raw_mul_conv_hash tm = let _ = mul_counter := !mul_counter + 1 in let hash = op_tm_hash tm in try @@ -190,7 +196,7 @@ let raw_mul_conv_hash tm = result;; (* x / y *) -let raw_div_hash_conv tm = +let raw_div_hash_conv tm = let _ = div_counter := !div_counter + 1 in let hash = op_tm_hash tm in try @@ -201,11 +207,11 @@ let raw_div_hash_conv tm = result;; (* EVEN, ODD *) -let raw_even_hash_conv tm = +let raw_even_hash_conv tm = let _ = even_counter := !even_counter + 1 in Arith_num.raw_even_hash_conv tm;; -let raw_odd_hash_conv tm = +let raw_odd_hash_conv tm = let _ = odd_counter := !odd_counter + 1 in Arith_num.raw_odd_hash_conv tm;; diff --git a/Formal_ineqs/arith/arith_float.hl b/Formal_ineqs/arith/arith_float.hl index 3e848398..1d9668a5 100644 --- a/Formal_ineqs/arith/arith_float.hl +++ b/Formal_ineqs/arith/arith_float.hl @@ -1,4186 +1,4192 @@ -(* =========================================================== *) -(* Formal floating point arithmetic *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -(* Dependencies *) -needs "arith/arith_nat.hl";; -needs "arith/num_exp_theory.hl";; -needs "arith/float_theory.hl";; -needs "arith/interval_arith.hl";; -needs "misc/misc_vars.hl";; - -needs "tests/log.hl";; - -(* FLOOR_DIV_DIV *) -needs "Library/floor.ml";; - -module type Arith_float_sig = - sig - val mk_num_exp : term -> term -> term - val dest_num_exp : term -> (term * term) - val dest_float : term -> (string * term * term) - val make_float : string -> term -> term -> term - val mk_float : int -> int -> term - - (* num_exp normalization/denormalization *) - val normalize : term -> thm * bool - val denormalize : term -> thm - - val lo_num_conv : int -> term -> thm - val hi_num_conv : int -> term -> thm - val num_exp_lo : int -> term -> thm - val num_exp_hi : int -> term -> thm - - val float_lt0 : term -> thm - val float_gt0 : term -> thm - val float_lt : term -> term -> thm - val float_le0 : term -> thm - val float_ge0 : term -> thm - val float_le : term -> term -> thm - val float_min : term -> term -> thm - val float_max : term -> term -> thm - val float_min_max : term -> term -> (thm * thm) - val float_mul_eq : term -> term -> thm - val float_mul_lo : int -> term -> term -> thm - val float_mul_hi : int -> term -> term -> thm - val float_div_lo : int -> term -> term -> thm - val float_div_hi : int -> term -> term -> thm - val float_add_lo : int -> term -> term -> thm - val float_add_hi : int -> term -> term -> thm - val float_sub_lo : int -> term -> term -> thm - val float_sub_hi : int -> term -> term -> thm - val float_sqrt_lo : int -> term -> thm - val float_sqrt_hi : int -> term -> thm - - val float_prove_le : term -> term -> bool * thm - val float_prove_lt : term -> term -> bool * thm - val float_prove_le_interval : term -> thm -> bool * thm - val float_prove_ge_interval : term -> thm -> bool * thm - val float_prove_lt_interval : term -> thm -> bool * thm - val float_prove_gt_interval : term -> thm -> bool * thm - val float_compare_interval : term -> thm -> int * thm - - val init_logs : unit -> unit - val reset_stat : unit -> unit - val reset_cache : unit -> unit - val print_stat : unit -> unit - - val dest_float_interval : term -> term * term * term - val mk_float_interval_small_num : int -> thm - val mk_float_interval_num : num -> thm - - val float_lo : int -> term -> thm - val float_hi : int -> term -> thm - - val float_interval_round : int -> thm -> thm - - val float_interval_abs : thm -> thm - val float_interval_neg : thm -> thm - val float_interval_mul : int -> thm -> thm -> thm - val float_interval_div : int -> thm -> thm -> thm - val float_interval_add : int -> thm -> thm -> thm - val float_interval_sub : int -> thm -> thm -> thm - val float_interval_sqrt : int -> thm -> thm - - val float_abs : term -> thm - - val FLOAT_TO_NUM_CONV : term -> thm -end;; - - -module Arith_float : Arith_float_sig = struct - -open Num;; -open Big_int;; -open Misc_functions;; -open Arith_nat;; -open Num_exp_theory;; -open Float_theory;; -open Interval_arith;; -open Misc_vars;; - -prioritize_real();; - -(* interval *) - -let APPROX_INTERVAL' = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL) APPROX_INTERVAL;; - -let interval_const = `interval_arith` and - num_exp_const = `num_exp`;; - -let b0_const = (fst o dest_comb o lhand o concl) (Arith_num.def_array.(0));; -let b0_name = (fst o dest_const) b0_const;; -let base_const = mk_small_numeral Arith_num.arith_base;; - -let NUM_REMOVE = prove(mk_eq(mk_comb(Arith_num.num_const, n_var_num), n_var_num), - REWRITE_TAC[Arith_num.num_def; NUMERAL]);; - -(* B0 n = base * n *) -let b0_thm = prove(mk_eq(mk_comb(b0_const, n_var_num), - mk_binop mul_op_num base_const n_var_num), - REWRITE_TAC[Arith_num.def_array.(0)] THEN - TRY ARITH_TAC THEN - GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV) [BIT0] THEN - ARITH_TAC);; - -let dest_num_exp tm = - let ltm, e_tm = dest_comb tm in - rand ltm, e_tm;; - -let num_exp_const = `num_exp`;; -let mk_num_exp n_tm e_tm = mk_binop num_exp_const n_tm e_tm;; - -(* float_num s n e -> "s", n, e *) -let dest_float tm = - let ltm, e_tm = dest_comb tm in - let ltm, n_tm = dest_comb ltm in - let float_tm, s_tm = dest_comb ltm in - if (fst o dest_const) float_tm <> "float_num" then - failwith "dest_float: not float" - else - (fst o dest_const) s_tm, n_tm, e_tm;; - -(* "s", n, e -> float_num s n e *) -let make_float = - let float_const = `float_num` in - fun s n_tm e_tm -> - let s_tm = if s = "T" then t_const else f_const in - mk_comb (mk_comb (mk_comb (float_const, s_tm), n_tm), e_tm);; - -(* Creates a float term with the value (n * base^e) *) -let mk_float = - let float_const = `float_num` in - fun n e -> - let n, s = if n < 0 then -n, t_const else n, f_const in - let n_tm = rand (Arith_nat.mk_small_numeral_array n) in - let e_tm = rand (Arith_nat.mk_small_numeral_array (e + Float_theory.min_exp)) in - mk_comb(mk_comb(mk_comb (float_const, s), n_tm), e_tm);; - - -(************************************) - -let NUM_EXP_EXP' = SPEC_ALL NUM_EXP_EXP;; -let NUM_EXP_0' = (SPEC_ALL o REWRITE_RULE[NUMERAL]) NUM_EXP_0;; -let NUM_EXP_LE' = (UNDISCH_ALL o SPEC_ALL) NUM_EXP_LE;; -let NUM_EXP_LT' = (UNDISCH_ALL o SPEC_ALL) NUM_EXP_LT;; - -(* B0 n = num_exp n bits *) -let normal_lemma1 = prove(mk_eq(mk_comb(b0_const, n_var_num), `num_exp n 1`), - REWRITE_TAC[Arith_num.def_array.(0); num_exp] THEN - TRY ARITH_TAC THEN - GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV) [BIT0] THEN - ARITH_TAC);; - -let NORMAL_LEMMA1 = NUMERALS_TO_NUM normal_lemma1;; - -let normal_lemma2 = prove(mk_eq (mk_comb (b0_const, `num_exp n e`), `num_exp n (SUC e)`), - REWRITE_TAC[normal_lemma1; NUM_EXP_EXP] THEN ARITH_TAC);; - -let rec normalize tm = - if (is_comb tm) then - let ltm, rtm = dest_comb tm in - let lname = (fst o dest_const) ltm in - if (lname = b0_name) then - let lth = INST[rtm, n_var_num] NORMAL_LEMMA1 in - let rth, flag = normalize rtm in - if flag then - let ltm, lexp = (dest_comb o snd o dest_eq o concl) lth in - let ltm, rtm = dest_comb ltm in - let rn, rexp = (dest_comb o snd o dest_eq o concl) rth in - let rn = rand rn in - let th1 = AP_THM (AP_TERM ltm rth) lexp in - let th2 = INST[rexp, e1_var_num; lexp, e2_var_num; rn, n_var_num] NUM_EXP_EXP' in - let th3 = TRANS lth (TRANS th1 th2) in - let ltm, rtm = (dest_comb o snd o dest_eq o concl) th3 in - let add_th = raw_add_conv_hash rtm in - let th4 = AP_TERM ltm add_th in - (TRANS th3 th4, true) - else - (lth, true) - else - (REFL tm, false) - else - (REFL tm, false);; - -(* Converts a raw numeral to a num_exp expression *) -let to_num_exp tm = - let x, flag = normalize tm in - if flag then x - else - INST[tm, n_var_num] NUM_EXP_0';; - -(************************************) - -let SYM_NUM_EXP_0' = SYM NUM_EXP_0';; -let NUM_EXP_n0 = prove(`!e. num_exp 0 e = 0`, REWRITE_TAC[num_exp; MULT_CLAUSES]);; -let NUM_EXP_n0' = (REWRITE_RULE[NUMERAL] o SPEC_ALL) NUM_EXP_n0;; - -let NUM_EXP_DENORM = (UNDISCH_ALL o prove) - (mk_imp(`e = _0 <=> F`, mk_eq(`num_exp n e`, mk_comb (b0_const, `num_exp n (PRE e)`))), - REWRITE_TAC[] THEN ONCE_REWRITE_TAC[SYM (REWRITE_CONV[NUMERAL] `0`)] THEN - REWRITE_TAC[num_exp; b0_thm] THEN - REWRITE_TAC[ARITH_RULE (mk_eq(mk_binop mul_op_num base_const `n * a:num`, - mk_binop mul_op_num `n:num` (mk_binop mul_op_num base_const `a:num`)))] THEN - REWRITE_TAC[GSYM EXP] THEN - SIMP_TAC[ARITH_RULE `~(e = 0) ==> SUC (PRE e) = e`]);; - -(* Converts num_exp n e to a numeral by adding e B0's *) -let rec denormalize tm = - let ltm, etm = dest_comb tm in - let ntm = rand ltm in - if (etm = zero_const) then - INST[ntm, n_var_num] SYM_NUM_EXP_0' - else - if ntm = zero_const then - INST[etm, e_var_num] NUM_EXP_n0' - else - let e_th = raw_eq0_hash_conv etm in - let th0' = INST[etm, e_var_num; ntm, n_var_num] NUM_EXP_DENORM in - let th0 = MY_PROVE_HYP e_th th0' in - let b0_tm, rtm = dest_comb(rand(concl th0)) in - let ltm, pre_tm = dest_comb rtm in - let pre_th = raw_pre_hash_conv pre_tm in - let th1 = AP_TERM ltm pre_th in - let th2 = denormalize (rand(concl th1)) in - TRANS th0 (AP_TERM b0_tm (TRANS th1 th2));; - -(***************************************) - -let rec comb_number tm n = - if (is_comb tm) then comb_number ((snd o dest_comb) tm) (n + 1) else n;; - -let make_lo_thm i = - let th_concl = mk_binop `(<=):num->num->bool` - (mk_comb (Arith_num.const_array.(0), n_var_num)) - (mk_comb (Arith_num.const_array.(i), n_var_num)) in - prove(th_concl, - REWRITE_TAC[Arith_num.def_array.(i); Arith_num.def_array.(0)] THEN - REWRITE_TAC[ARITH_LE; LE_REFL] THEN - ARITH_TAC);; - -let lo_thm_array = Array.init Arith_num.arith_base make_lo_thm;; -let lo_thm_table = Hashtbl.create Arith_num.arith_base;; - -for i = 0 to Arith_num.arith_base - 1 do - Hashtbl.add lo_thm_table Arith_num.const_array.(i) lo_thm_array.(i); -done;; - -let make_lo_thm2 i = - let th_concl = mk_imp (`n <= m:num`, - mk_binop `(<=):num->num->bool` - (mk_comb (Arith_num.const_array.(0), n_var_num)) - (mk_comb (Arith_num.const_array.(i), m_var_num))) in - (UNDISCH_ALL o prove) (th_concl, - REWRITE_TAC[Arith_num.def_array.(i); Arith_num.def_array.(0); ARITH_LE] THEN - ARITH_TAC);; - -let lo_thm2_array = Array.init Arith_num.arith_base make_lo_thm2;; -let lo_thm2_table = Hashtbl.create Arith_num.arith_base;; - -for i = 0 to Arith_num.arith_base - 1 do - Hashtbl.add lo_thm2_table Arith_num.const_array.(i) lo_thm2_array.(i); -done;; - - -let make_hi_thm i = - let th_concl = mk_imp (`n < m:num`, - mk_binop `(<):num->num->bool` - (mk_comb (Arith_num.const_array.(i), n_var_num)) - (mk_comb (Arith_num.const_array.(0), m_var_num))) in - (UNDISCH_ALL o prove) (th_concl, - REWRITE_TAC[Arith_num.def_array.(i); Arith_num.def_array.(0); ARITH_LT] THEN - ARITH_TAC);; - -let hi_thm_array = Array.init Arith_num.arith_base make_hi_thm;; -let hi_thm_table = Hashtbl.create Arith_num.arith_base;; - -for i = 0 to Arith_num.arith_base - 1 do - Hashtbl.add hi_thm_table Arith_num.const_array.(i) hi_thm_array.(i); -done;; - -(***************************************) - -let LE_REFL' = SPEC_ALL LE_REFL;; -let LE_TRANS' = (UNDISCH_ALL o SPEC_ALL o REWRITE_RULE[GSYM IMP_IMP]) LE_TRANS;; - -let lo_num_conv p tm = - let n = comb_number tm 0 in - if (n <= p) then - INST[tm, n_var_num] LE_REFL' - else - let rec lo_bound n tm = - let btm, rtm = dest_comb tm in - let th0 = INST[rtm, n_var_num] (Hashtbl.find lo_thm_table btm) in - if n > 1 then - let rth = lo_bound (n - 1) rtm in - let xtm = rand (rator (concl rth)) in - let th1' = INST[xtm, n_var_num; rtm, m_var_num] (Hashtbl.find lo_thm2_table btm) in - let th1 = MY_PROVE_HYP rth th1' in - th1 - else - th0 in - - lo_bound (n - p) tm;; - -let N_LT_SUC = ARITH_RULE `n < SUC n`;; -let LT_IMP_LE' = (UNDISCH_ALL o SPEC_ALL) LT_IMP_LE;; -let N_LT_SUC = ARITH_RULE `n < SUC n`;; -let LT_LE_TRANS = (UNDISCH_ALL o ARITH_RULE) `n < e ==> e <= m ==> n < m:num`;; - -(* Generates a theorem |- n <= m such that m contains at most p non-zero digits *) -let hi_num_conv p tm = - let n = comb_number tm 0 in - if (n <= p) then - INST[tm, n_var_num] LE_REFL' - else - let k = n - p in - - let rec check_b0s n tm = - let btm, rtm = dest_comb tm in - if ((fst o dest_const) btm = b0_name) then - if n > 1 then check_b0s (n - 1) rtm else true - else - false in - - if (check_b0s k tm) then - INST[tm, n_var_num] LE_REFL' - else - let rec hi_bound n tm = - if n > 0 then - let btm, rtm = dest_comb tm in - let r_th = hi_bound (n - 1) rtm in - let xtm = rand (concl r_th) in - let th0 = INST[rtm, n_var_num; xtm, m_var_num] (Hashtbl.find hi_thm_table btm) in - MY_PROVE_HYP r_th th0 - else - let th0 = INST[tm, n_var_num] N_LT_SUC in - let ltm, suc_tm = dest_comb (concl th0) in - let suc_th = raw_suc_conv_hash suc_tm in - EQ_MP (AP_TERM ltm suc_th) th0 in - - let th = hi_bound k tm in - let m_tm, l_tm = dest_comb (concl th) in - MY_PROVE_HYP th (INST[rand m_tm, m_var_num; l_tm, n_var_num] LT_IMP_LE');; - -(* Generates a theorem |- n < m such that m contains at most p non-zero digits *) -let hi_lt_num_conv p tm = - let n = comb_number tm 0 in - if (n <= p) then - let th0 = INST[tm, n_var_num] N_LT_SUC in - let ltm, rtm = dest_comb(concl th0) in - let suc_th = raw_suc_conv_hash rtm in - EQ_MP (AP_TERM ltm suc_th) th0 - else - let k = n - p in - - let rec check_b0s n tm = - let btm, rtm = dest_comb tm in - if ((fst o dest_const) btm = b0_name) then - if n > 1 then check_b0s (n - 1) rtm else true - else - false in - - if (check_b0s k tm) then - let th0 = INST[tm, n_var_num] N_LT_SUC in - let ltm, rtm = dest_comb (concl th0) in - let suc_th = raw_suc_conv_hash rtm in - let suc_tm = rand(concl suc_th) in - let th1 = hi_num_conv p suc_tm in - let th2 = EQ_MP (AP_TERM ltm suc_th) th0 in - let th = INST[tm, n_var_num; suc_tm, e_var_num; rand(concl th1), m_var_num] LT_LE_TRANS in - MY_PROVE_HYP th1 (MY_PROVE_HYP th2 th) - - else - let rec hi_bound n tm = - if n > 0 then - let btm, rtm = dest_comb tm in - let r_th = hi_bound (n - 1) rtm in - let xtm = rand (concl r_th) in - let th0 = INST[rtm, n_var_num; xtm, m_var_num] (Hashtbl.find hi_thm_table btm) in - MY_PROVE_HYP r_th th0 - else - let th0 = INST[tm, n_var_num] N_LT_SUC in - let ltm, suc_tm = dest_comb (concl th0) in - let suc_th = raw_suc_conv_hash suc_tm in - EQ_MP (AP_TERM ltm suc_th) th0 in - hi_bound k tm;; - -(*****************************************) - -let num_exp_lo p tm = - let ltm, e_tm = dest_comb tm in - let n_tm = rand ltm in - let n_th = lo_num_conv p n_tm in - let m_tm = rand (rator (concl n_th)) in - let m_norm, flag = normalize m_tm in - - let th0' = INST[m_tm, m_var_num; n_tm, n_var_num; e_tm, e_var_num] NUM_EXP_LE' in - let th0 = MY_PROVE_HYP n_th th0' in - if flag then - let th1 = AP_THM (AP_TERM (rator ltm) m_norm) e_tm in - let m_tm, me_tm = (dest_comb o rand o concl) m_norm in - let th2 = INST[me_tm, e1_var_num; e_tm, e2_var_num; rand m_tm, n_var_num] NUM_EXP_EXP' in - let th3 = TRANS th1 th2 in - let ltm, rtm = (dest_comb o rand o concl) th3 in - let th_add = raw_add_conv_hash rtm in - let th4 = TRANS th3 (AP_TERM ltm th_add) in - EQ_MP (AP_THM (AP_TERM le_op_num th4) tm) th0 - else - th0;; - -let num_exp_hi p tm = - let ltm, e_tm = dest_comb tm in - let n_tm = rand ltm in - let n_th = hi_num_conv p n_tm in - let m_tm = rand (concl n_th) in - let m_norm, flag = normalize m_tm in - - let th0' = INST[m_tm, n_var_num; n_tm, m_var_num; e_tm, e_var_num] NUM_EXP_LE' in - let th0 = MY_PROVE_HYP n_th th0' in - if flag then - let th1 = AP_THM (AP_TERM (rator ltm) m_norm) e_tm in - let m_tm, me_tm = (dest_comb o rand o concl) m_norm in - let th2 = INST[me_tm, e1_var_num; e_tm, e2_var_num; rand m_tm, n_var_num] NUM_EXP_EXP' in - let th3 = TRANS th1 th2 in - let ltm, rtm = (dest_comb o rand o concl) th3 in - let th_add = raw_add_conv_hash rtm in - let th4 = TRANS th3 (AP_TERM ltm th_add) in - EQ_MP (AP_TERM (rator (concl th0)) th4) th0 - else - th0;; - -let num_exp_hi_lt p tm = - let ltm, e_tm = dest_comb tm in - let n_tm = rand ltm in - let n_th = hi_lt_num_conv p n_tm in - let m_tm = rand (concl n_th) in - let m_norm, flag = normalize m_tm in - - let th0' = INST[m_tm, n_var_num; n_tm, m_var_num; e_tm, e_var_num] NUM_EXP_LT' in - let th0 = MY_PROVE_HYP n_th th0' in - if flag then - let th1 = AP_THM (AP_TERM (rator ltm) m_norm) e_tm in - let m_tm, me_tm = (dest_comb o rand o concl) m_norm in - let th2 = INST[me_tm, e1_var_num; e_tm, e2_var_num; rand m_tm, n_var_num] NUM_EXP_EXP' in - let th3 = TRANS th1 th2 in - let ltm, rtm = (dest_comb o rand o concl) th3 in - let th_add = raw_add_conv_hash rtm in - let th4 = TRANS th3 (AP_TERM ltm th_add) in - EQ_MP (AP_TERM (rator (concl th0)) th4) th0 - else - th0;; - -(***************************************) -(* num_exp_lt, num_exp_le *) - -let transform = UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL;; - -let NUM_EXP_LT1_EQ' = transform NUM_EXP_LT1_EQ and - NUM_EXP_LT2_EQ' = transform NUM_EXP_LT2_EQ;; - -let NUM_EXP_LE1_EQ' = transform NUM_EXP_LE1_EQ and - NUM_EXP_LE2_EQ' = transform NUM_EXP_LE2_EQ;; - -let num_exp_lt tm1 tm2 = - let n1_tm, e1_tm = dest_num_exp tm1 in - let n2_tm, e2_tm = dest_num_exp tm2 in - let sub_th, le_th = raw_sub_and_le_hash_conv e1_tm e2_tm in - let r_tm = rand(concl sub_th) in - if (rand(concl le_th) = e1_tm) then - let x_expr = mk_num_exp n1_tm r_tm in - let x_th = denormalize x_expr in - let x_tm = rand(concl x_th) in - - let th0 = INST[e2_tm, e2_var_num; e1_tm, e1_var_num; - r_tm, r_var_num; x_tm, x_var_num; - n1_tm, n1_var_num; n2_tm, n2_var_num] NUM_EXP_LT1_EQ' in - let th1 = MY_PROVE_HYP x_th (MY_PROVE_HYP sub_th (MY_PROVE_HYP le_th th0)) in - let lt_th = raw_lt_hash_conv (rand(concl th1)) in - TRANS th1 lt_th - else - let x_expr = mk_num_exp n2_tm r_tm in - let x_th = denormalize x_expr in - let x_tm = rand(concl x_th) in - - let th0 = INST[e2_tm, e2_var_num; e1_tm, e1_var_num; - r_tm, r_var_num; x_tm, x_var_num; - n1_tm, n1_var_num; n2_tm, n2_var_num] NUM_EXP_LT2_EQ' in - let th1 = MY_PROVE_HYP x_th (MY_PROVE_HYP sub_th (MY_PROVE_HYP le_th th0)) in - let lt_th = raw_lt_hash_conv (rand(concl th1)) in - TRANS th1 lt_th;; - - -let num_exp_le tm1 tm2 = - let n1_tm, e1_tm = dest_num_exp tm1 in - let n2_tm, e2_tm = dest_num_exp tm2 in - let sub_th, le_th = raw_sub_and_le_hash_conv e1_tm e2_tm in - let r_tm = rand(concl sub_th) in - if (rand(concl le_th) = e1_tm) then - let x_expr = mk_num_exp n1_tm r_tm in - let x_th = denormalize x_expr in - let x_tm = rand(concl x_th) in - - let th0 = INST[e2_tm, e2_var_num; e1_tm, e1_var_num; - r_tm, r_var_num; x_tm, x_var_num; - n1_tm, n1_var_num; n2_tm, n2_var_num] NUM_EXP_LE1_EQ' in - let th1 = MY_PROVE_HYP x_th (MY_PROVE_HYP sub_th (MY_PROVE_HYP le_th th0)) in - let le_th = raw_le_hash_conv (rand(concl th1)) in - TRANS th1 le_th - else - let x_expr = mk_num_exp n2_tm r_tm in - let x_th = denormalize x_expr in - let x_tm = rand(concl x_th) in - - let th0 = INST[e2_tm, e2_var_num; e1_tm, e1_var_num; - r_tm, r_var_num; x_tm, x_var_num; - n1_tm, n1_var_num; n2_tm, n2_var_num] NUM_EXP_LE2_EQ' in - let th1 = MY_PROVE_HYP x_th (MY_PROVE_HYP sub_th (MY_PROVE_HYP le_th th0)) in - let le_th = raw_le_hash_conv (rand(concl th1)) in - TRANS th1 le_th;; - -(***************************************) -(* num_exp_mul *) - -let NUM_EXP_MUL' = SPEC_ALL NUM_EXP_MUL;; - -let num_exp_mul tm1 tm2 = - let n1_tm, e1_tm = dest_comb tm1 in - let n1_tm = rand n1_tm in - let n2_tm, e2_tm = dest_comb tm2 in - let n2_tm = rand n2_tm in - let th0 = INST[n1_tm, n1_var_num; e1_tm, e1_var_num; - n2_tm, n2_var_num; e2_tm, e2_var_num] NUM_EXP_MUL' in - let ltm, tm_add = dest_comb (rand (concl th0)) in - let tm_mul = rand ltm in - let th_mul = raw_mul_conv_hash tm_mul in - let th_add = raw_add_conv_hash tm_add in - TRANS th0 (MK_COMB (AP_TERM (rator ltm) th_mul, th_add));; - - -(**********************************) -(* num_exp_add *) - -let NUM_EXP_ADD' = (UNDISCH_ALL o SPEC_ALL) NUM_EXP_ADD;; -let ADD_COMM = ARITH_RULE `m + n = n + m:num`;; - -let num_exp_add tm1 tm2 = - let n1_tm, e1_tm = dest_comb tm1 in - let n1_tm = rand n1_tm in - let n2_tm, e2_tm = dest_comb tm2 in - let n2_tm = rand n2_tm in - let e_sub, e_le = raw_sub_and_le_hash_conv e1_tm e2_tm in - - let flag = (rand(concl e_le) = e2_tm) in - - let th0' = - if flag then - INST[n1_tm, n1_var_num; e1_tm, e1_var_num; - n2_tm, n2_var_num; e2_tm, e2_var_num] NUM_EXP_ADD' - else - INST[n2_tm, n1_var_num; e2_tm, e1_var_num; - n1_tm, n2_var_num; e1_tm, e2_var_num] NUM_EXP_ADD' in - - let th0 = MY_PROVE_HYP e_le th0' in - let ltm, e0_tm = dest_comb(rand(concl th0)) in - let exp_tm, add_tm = dest_comb ltm in - let ltm, d_tm = dest_comb add_tm in - let th1 = AP_TERM (rator d_tm) e_sub in - let th2 = denormalize (rand(concl th1)) in - let th3 = AP_TERM ltm (TRANS th1 th2) in - let th4 = raw_add_conv_hash (rand(concl th3)) in - let th5 = AP_THM (AP_TERM exp_tm (TRANS th3 th4)) e0_tm in - let th = TRANS th0 th5 in - if flag then th else - TRANS (INST[tm1, m_var_num; tm2, n_var_num] ADD_COMM) th;; - -(****************************************) -(* num_exp_sub *) - -let NUM_EXP_SUB1' = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL) NUM_EXP_SUB1 and - NUM_EXP_SUB2' = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL) NUM_EXP_SUB2 and - NUM_EXP_LE1' = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL) NUM_EXP_LE1 and - NUM_EXP_LE2' = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL) NUM_EXP_LE2;; - -(* Returns two theorems: |- tm1 - tm2 = tm, |- tm2 <= tm1 or - |- tm2 - tm1 = tm, |- tm1 <= tm2 *) -let num_exp_sub tm1 tm2 = - let n1_tm, e1_tm = dest_num_exp tm1 in - let n2_tm, e2_tm = dest_num_exp tm2 in - let e_sub, e_le = raw_sub_and_le_hash_conv e1_tm e2_tm in - - if rand(concl e_le) = e1_tm then - (* e2 <= e1 *) - let e1_sub_e2 = rand(concl e_sub) in - let a0 = mk_num_exp n1_tm e1_sub_e2 in - let b = n2_tm in - let a_th = denormalize a0 in - let a = rand(concl a_th) in - - let th_sub, th_le = raw_sub_and_le_hash_conv a b in - if rand(concl th_le) = a then - (* b <= a *) - let a_sub_b = TRANS (AP_THM (AP_TERM sub_op_num a_th) b) th_sub in - let b_le_a = EQ_MP (SYM (AP_TERM (rator(concl th_le)) a_th)) th_le in - let th0 = AP_THM (AP_TERM num_exp_const a_sub_b) e2_tm in - - let inst = INST[n1_tm, n1_var_num; e1_tm, e1_var_num; - n2_tm, n2_var_num; e2_tm, e2_var_num; e1_sub_e2, r_var_num] in - let th1_sub = inst NUM_EXP_SUB1' in - let th1_le = inst NUM_EXP_LE1' in - let th2_sub = MY_PROVE_HYP e_sub (MY_PROVE_HYP e_le th1_sub) in - let th2_le = MY_PROVE_HYP e_sub (MY_PROVE_HYP b_le_a (MY_PROVE_HYP e_le th1_le)) in - TRANS th2_sub th0, th2_le - - else - (* a <= b *) - let b_sub_a = TRANS (AP_TERM (rator(lhand(concl th_sub))) a_th) th_sub in - let a_le_b = EQ_MP (SYM (AP_THM (AP_TERM le_op_num a_th) b)) th_le in - let th0 = AP_THM (AP_TERM num_exp_const b_sub_a) e2_tm in - let inst = INST[n2_tm, n1_var_num; e2_tm, e1_var_num; - n1_tm, n2_var_num; e1_tm, e2_var_num; e1_sub_e2, r_var_num] in - let th1_sub = inst NUM_EXP_SUB2' in - let th1_le = inst NUM_EXP_LE2' in - let th2_sub = MY_PROVE_HYP e_sub (MY_PROVE_HYP e_le th1_sub) in - let th2_le = MY_PROVE_HYP e_sub (MY_PROVE_HYP a_le_b (MY_PROVE_HYP e_le th1_le)) in - TRANS th2_sub th0, th2_le - - else - (* e1 <= e2 *) - let e2_sub_e1 = rand(concl e_sub) in - let b0 = mk_num_exp n2_tm e2_sub_e1 in - let a = n1_tm in - let b_th = denormalize b0 in - let b = rand(concl b_th) in - - let th_sub, th_le = raw_sub_and_le_hash_conv a b in - if rand(concl th_le) = a then - (* b <= a *) - let a_sub_b = TRANS (AP_TERM (rator(lhand(concl th_sub))) b_th) th_sub in - let b_le_a = EQ_MP (SYM (AP_THM (AP_TERM le_op_num b_th) a)) th_le in - let th0 = AP_THM (AP_TERM num_exp_const a_sub_b) e1_tm in - let inst = INST[n1_tm, n1_var_num; e1_tm, e1_var_num; - n2_tm, n2_var_num; e2_tm, e2_var_num; e2_sub_e1, r_var_num] in - let th1_sub = inst NUM_EXP_SUB2' in - let th1_le = inst NUM_EXP_LE2' in - let th2_sub = MY_PROVE_HYP e_sub (MY_PROVE_HYP e_le th1_sub) in - let th2_le = MY_PROVE_HYP e_sub (MY_PROVE_HYP b_le_a (MY_PROVE_HYP e_le th1_le)) in - TRANS th2_sub th0, th2_le - - else - (* a <= b *) - let b_sub_a = TRANS (AP_THM (AP_TERM sub_op_num b_th) a) th_sub in - let a_le_b = EQ_MP (SYM (AP_TERM (rator(concl th_le)) b_th)) th_le in - let th0 = AP_THM (AP_TERM num_exp_const b_sub_a) e1_tm in - let inst = INST[n2_tm, n1_var_num; e2_tm, e1_var_num; - n1_tm, n2_var_num; e1_tm, e2_var_num; e2_sub_e1, r_var_num] in - let th1_sub = inst NUM_EXP_SUB1' in - let th1_le = inst NUM_EXP_LE1' in - let th2_sub = MY_PROVE_HYP e_sub (MY_PROVE_HYP e_le th1_sub) in - let th2_le = MY_PROVE_HYP e_sub (MY_PROVE_HYP a_le_b (MY_PROVE_HYP e_le th1_le)) in - TRANS th2_sub th0, th2_le;; - - -(*************************************) -(* division *) - -let NUM_EXP_DIV1' = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o - PURE_ONCE_REWRITE_RULE[ARITH_RULE `~(x = 0) <=> (x = 0 <=> F)`] o - REWRITE_RULE[GSYM IMP_IMP]) NUM_EXP_DIV1;; -let NUM_EXP_DIV2' = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o - PURE_ONCE_REWRITE_RULE[ARITH_RULE `~(x = 0) <=> (x = 0 <=> F)`] o - REWRITE_RULE[GSYM IMP_IMP]) NUM_EXP_DIV2;; - -let num_exp_div tm1 tm2 = - let n1_tm, e1_tm = dest_comb tm1 in - let n1_tm = rand n1_tm in - let n2_tm, e2_tm = dest_comb tm2 in - let n2_tm = rand n2_tm in - let e_sub, e_le = raw_sub_and_le_hash_conv e1_tm e2_tm in - - let inst = INST[n1_tm, n1_var_num; e1_tm, e1_var_num; - n2_tm, n2_var_num; e2_tm, e2_var_num] in - - let n2_not_0 = raw_eq0_hash_conv n2_tm in - if ((fst o dest_const o rand o concl) n2_not_0 = "T") then - failwith "num_exp_div: n2 = 0" - else - if (rand(concl e_le) = e1_tm) then - let th0' = inst NUM_EXP_DIV1' in - let th0 = MY_PROVE_HYP n2_not_0 (MY_PROVE_HYP e_le th0') in - - let ltm, rtm = dest_comb(rand(concl th0)) in - let div_tm, rtm2 = dest_comb ltm in - let num_exp_tm = rator rtm2 in - - let th1 = AP_THM (AP_TERM div_tm (AP_TERM num_exp_tm e_sub)) rtm in - let ltm, rtm = dest_comb(rand(concl th1)) in - let tm1 = rand ltm in - - let th2 = AP_THM (AP_TERM div_tm (denormalize tm1)) rtm in - let th3 = raw_div_hash_conv (rand(concl th2)) in - let th = TRANS th0 (TRANS th1 (TRANS th2 th3)) in - TRANS th (INST[rand(concl th), n_var_num] NUM_EXP_0') - - else - let th0' = inst NUM_EXP_DIV2' in - let th0 = MY_PROVE_HYP n2_not_0 (MY_PROVE_HYP e_le th0') in - - let ltm, rtm = dest_comb(rand(concl th0)) in - let num_exp_tm = rator rtm in - let th1 = AP_TERM ltm (AP_TERM num_exp_tm e_sub) in - - let ltm, rtm = dest_comb(rand(concl th1)) in - let th2 = AP_TERM ltm (denormalize rtm) in - let th3 = raw_div_hash_conv (rand(concl th2)) in - let th = TRANS th0 (TRANS th1 (TRANS th2 th3)) in - TRANS th (INST[rand(concl th), n_var_num] NUM_EXP_0');; - - -(*****************************) -(* Computes a lower bound for (op tm1 tm2) with p significant digits *) -let num_exp_op_lo p op tm1 tm2 = - let op_th = op tm1 tm2 in - let rtm = rand (concl op_th) in - let lo_th = num_exp_lo p rtm in - let ltm = rator (concl lo_th) in - let th0 = AP_TERM ltm op_th in - EQ_MP (SYM th0) lo_th;; - -(* Computes an upper bound for (op tm1 tm2) with p significant digits *) -let num_exp_op_hi p op tm1 tm2 = - let op_th = op tm1 tm2 in - let rtm = rand (concl op_th) in - let hi_th = num_exp_hi p rtm in - let tm = rand (concl hi_th) in - let th0 = AP_THM (AP_TERM le_op_num op_th) tm in - EQ_MP (SYM th0) hi_th;; - -(* Computes a strict upper bound for (op tm1 tm2) with p significant digits *) -let num_exp_op_hi_lt p op tm1 tm2 = - let op_th = op tm1 tm2 in - let rtm = rand (concl op_th) in - let hi_lt_th = num_exp_hi_lt p rtm in - let tm = rand (concl hi_lt_th) in - let th0 = AP_THM (AP_TERM lt_op_num op_th) tm in - EQ_MP (SYM th0) hi_lt_th;; - -(******************************************) -(* float *) - -let mod_plus = new_definition `mod_plus s1 s2 = (~(s1 /\ s2) /\ (s1 \/ s2))`;; - -(********************) -(* Float operations *) -(********************) - -module Float_ops = struct - -(**********************************) - -(* FLOAT_LT *) - -let FLOAT_LT_FF = prove(`float_num F n1 e1 < float_num F n2 e2 <=> num_exp n1 e1 < num_exp n2 e2`, - REWRITE_TAC[float; GSYM REAL_OF_NUM_LT; REAL_MUL_LID; real_div] THEN - MATCH_MP_TAC REAL_LT_RMUL_EQ THEN - MATCH_MP_TAC REAL_LT_INV THEN - REWRITE_TAC[REAL_OF_NUM_LT; LT_NZ; NUM_EXP_EQ_0] THEN - ARITH_TAC);; - -let FLOAT_LT_TT = prove(`float_num T n1 e1 < float_num T n2 e2 <=> num_exp n2 e2 < num_exp n1 e1`, - REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `--a < --b <=> b < a`] THEN - REWRITE_TAC[FLOAT_LT_FF]);; - -let FLOAT_LT_FT = prove(`float_num F n1 e1 < float_num T n2 e2 <=> F`, - MP_TAC (SPECL [`n2:num`; `e2:num`] FLOAT_T_NEG) THEN - MP_TAC (SPECL [`n1:num`; `e1:num`] FLOAT_F_POS) THEN - REAL_ARITH_TAC);; - -let FLOAT_LT_TF_00 = (PURE_REWRITE_RULE[NUMERAL] o prove) - (`float_num T 0 e1 < float_num F 0 e2 <=> F`, - MP_TAC (SPECL [`T`; `0`; `e1:num`] FLOAT_EQ_0) THEN - MP_TAC (SPECL [`F`; `0`; `e2:num`] FLOAT_EQ_0) THEN - REWRITE_TAC[] THEN - REPLICATE_TAC 2 (DISCH_THEN (fun th -> REWRITE_TAC[th])) THEN - REAL_ARITH_TAC);; - -let FLOAT_LT_TF_1 = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o prove) - (`(n1 = 0 <=> F) ==> (float_num T n1 e1 < float_num F n2 e2 <=> T)`, - DISCH_TAC THEN - MATCH_MP_TAC (REAL_ARITH `a < &0 /\ &0 <= b ==> (a < b <=> T)`) THEN - REWRITE_TAC[FLOAT_F_POS] THEN - MATCH_MP_TAC (REAL_ARITH `~(a = &0) /\ a <= &0 ==> a < &0`) THEN - ASM_REWRITE_TAC[FLOAT_T_NEG; FLOAT_EQ_0]);; - -let FLOAT_LT_TF_2 = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o prove) - (`(n2 = 0 <=> F) ==> (float_num T n1 e1 < float_num F n2 e2 <=> T)`, - DISCH_TAC THEN - MATCH_MP_TAC (REAL_ARITH `a <= &0 /\ &0 < b ==> (a < b <=> T)`) THEN - REWRITE_TAC[FLOAT_T_NEG] THEN - MATCH_MP_TAC (REAL_ARITH `~(a = &0) /\ &0 <= a ==> &0 < a`) THEN - ASM_REWRITE_TAC[FLOAT_F_POS; FLOAT_EQ_0]);; - -let FLOAT_F_LT_0 = prove(`float_num F n e < &0 <=> F`, - MP_TAC (SPEC_ALL FLOAT_F_POS) THEN - REAL_ARITH_TAC);; - -let FLOAT_T_LT_0 = (CONV_RULE (RAND_CONV (REWRITE_CONV[NUMERAL])) o prove) - (`float_num T n e < &0 <=> (0 < n)`, - REWRITE_TAC[REAL_ARITH `a < &0 <=> a <= &0 /\ ~(a = &0)`] THEN - REWRITE_TAC[FLOAT_T_NEG; FLOAT_EQ_0] THEN - ARITH_TAC);; - -let FLOAT_F_GT_0 = (CONV_RULE (RAND_CONV (REWRITE_CONV[NUMERAL])) o prove) - (`&0 < float_num F n e <=> 0 < n`, - REWRITE_TAC[REAL_ARITH `&0 < a <=> &0 <= a /\ ~(a = &0)`] THEN - REWRITE_TAC[FLOAT_F_POS; FLOAT_EQ_0] THEN - ARITH_TAC);; - -let FLOAT_T_GT_0 = prove(`&0 < float_num T n e <=> F`, - MP_TAC (SPEC_ALL FLOAT_T_NEG) THEN - REAL_ARITH_TAC);; - - -(* float_lt0, float_gt0 *) - -let float_lt0 f1 = - let s, n_tm, e_tm = dest_float f1 in - let inst = INST[n_tm, n_var_num; e_tm, e_var_num] in - if s = "F" then - inst FLOAT_F_LT_0 - else - let gt_th = raw_gt0_hash_conv n_tm in - TRANS (inst FLOAT_T_LT_0) gt_th;; - -let float_gt0 f1 = - let s, n_tm, e_tm = dest_float f1 in - let inst = INST[n_tm, n_var_num; e_tm, e_var_num] in - if s = "F" then - let gt_th = raw_gt0_hash_conv n_tm in - TRANS (inst FLOAT_F_GT_0) gt_th - else - inst FLOAT_T_GT_0;; - -(* float_lt *) -let float_lt f1 f2 = - let s1, n1, e1 = dest_float f1 in - let s2, n2, e2 = dest_float f2 in - let inst = INST[n1, n1_var_num; e1, e1_var_num; - n2, n2_var_num; e2, e2_var_num] in - if s1 = "F" then - if s2 = "F" then - (* FF *) - let th0 = inst FLOAT_LT_FF in - let ltm, tm2 = dest_comb (rand (concl th0)) in - let lt_th = num_exp_lt (rand ltm) tm2 in - TRANS th0 lt_th - else - (* FT *) - inst FLOAT_LT_FT - else - if s2 = "F" then - (* TF *) - if (is_const n1 && is_const n2) then - (* n1 = _0 and n2 = _0 *) - inst FLOAT_LT_TF_00 - else - let n1_0 = raw_eq0_hash_conv n1 in - if (fst o dest_const o rand o concl) n1_0 = "F" then - (* n1 <> _0 *) - MY_PROVE_HYP n1_0 (inst FLOAT_LT_TF_1) - else - let n2_0 = raw_eq0_hash_conv n2 in - if (fst o dest_const o rand o concl) n2_0 = "F" then - (* n2 <> _0 *) - MY_PROVE_HYP n2_0 (inst FLOAT_LT_TF_2) - else - failwith "float_lt: D0 _0 exception" - else - (* TT *) - let th0 = inst FLOAT_LT_TT in - let ltm, tm2 = dest_comb (rand (concl th0)) in - let lt_th = num_exp_lt (rand ltm) tm2 in - TRANS th0 lt_th;; - -(**********************************) -(* FLOAT_LE *) - -let FLOAT_LE_FF = prove(`float_num F n1 e1 <= float_num F n2 e2 <=> num_exp n1 e1 <= num_exp n2 e2`, - REWRITE_TAC[float; GSYM REAL_OF_NUM_LE; REAL_MUL_LID; real_div] THEN - MATCH_MP_TAC REAL_LE_RMUL_EQ THEN - MATCH_MP_TAC REAL_LT_INV THEN - REWRITE_TAC[REAL_OF_NUM_LT; LT_NZ; NUM_EXP_EQ_0] THEN - ARITH_TAC);; - -let FLOAT_LE_TT = prove(`float_num T n1 e1 <= float_num T n2 e2 <=> num_exp n2 e2 <= num_exp n1 e1`, - REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `--a <= --b <=> b <= a`] THEN - REWRITE_TAC[FLOAT_LE_FF]);; - -let FLOAT_LE_TF = prove(`float_num T n1 e1 <= float_num F n2 e2 <=> T`, - MP_TAC (SPECL [`n1:num`; `e1:num`] FLOAT_T_NEG) THEN - MP_TAC (SPECL [`n2:num`; `e2:num`] FLOAT_F_POS) THEN - REAL_ARITH_TAC);; - -let FLOAT_LE_FT = prove(`float_num F n1 e1 <= float_num T n2 e2 <=> n1 = 0 /\ n2 = 0`, - REWRITE_TAC[REAL_LE_LT; FLOAT_LT_FT] THEN EQ_TAC THENL - [ - DISCH_TAC THEN SUBGOAL_THEN `float_num F n1 e1 = &0 /\ float_num T n2 e2 = &0` MP_TAC THENL - [ - MP_TAC (SPECL [`n2:num`; `e2:num`] FLOAT_T_NEG) THEN - MP_TAC (SPECL [`n1:num`; `e1:num`] FLOAT_F_POS) THEN - ASM_REWRITE_TAC[] THEN REAL_ARITH_TAC; - ALL_TAC - ] THEN - REWRITE_TAC[FLOAT_EQ_0]; - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - REWRITE_TAC[float; NUM_EXP_n0; real_div; REAL_MUL_LZERO; REAL_MUL_RZERO] - ]);; - -let FLOAT_LE_FT_00 = (PURE_REWRITE_RULE[NUMERAL] o prove) - (`float_num F 0 e1 <= float_num T 0 e2 <=> T`, REWRITE_TAC[FLOAT_LE_FT]);; - -let FLOAT_LE_FT_1 = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o prove) - (`(n1 = 0 <=> F) ==> (float_num F n1 e1 <= float_num T n2 e2 <=> F)`, - DISCH_TAC THEN ASM_REWRITE_TAC[FLOAT_LE_FT]);; - -let FLOAT_LE_FT_2 = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o prove) - (`(n2 = 0 <=> F) ==> (float_num F n1 e1 <= float_num T n2 e2 <=> F)`, - DISCH_TAC THEN ASM_REWRITE_TAC[FLOAT_LE_FT]);; - -let FLOAT_F_LE_0 = (CONV_RULE (RAND_CONV (REWRITE_CONV[NUMERAL])) o prove) - (`float_num F n e <= &0 <=> n = 0`, - REWRITE_TAC[GSYM (SPEC `F` FLOAT_EQ_0)] THEN - MP_TAC (SPEC_ALL FLOAT_F_POS) THEN - REAL_ARITH_TAC);; - -let FLOAT_T_LE_0 = prove(`float_num T n e <= &0 <=> T`, REWRITE_TAC[FLOAT_T_NEG]);; - -let FLOAT_F_GE_0 = prove(`&0 <= float_num F n e <=> T`, REWRITE_TAC[FLOAT_F_POS]);; - -let FLOAT_T_GE_0 = (CONV_RULE (RAND_CONV (REWRITE_CONV[NUMERAL])) o prove) - (`&0 <= float_num T n e <=> n = 0`, - REWRITE_TAC[GSYM (SPEC `T` FLOAT_EQ_0)] THEN - MP_TAC (SPEC_ALL FLOAT_T_NEG) THEN - REAL_ARITH_TAC);; - -(* float_le0, float_ge0 *) -let float_le0 f1 = - let s, n_tm, e_tm = dest_float f1 in - let inst = INST[n_tm, n_var_num; e_tm, e_var_num] in - if s = "T" then - inst FLOAT_T_LE_0 - else - let eq_th = raw_eq0_hash_conv n_tm in - TRANS (inst FLOAT_F_LE_0) eq_th;; - -let float_ge0 f1 = - let s, n_tm, e_tm = dest_float f1 in - let inst = INST[n_tm, n_var_num; e_tm, e_var_num] in - if s = "T" then - let eq_th = raw_eq0_hash_conv n_tm in - TRANS (inst FLOAT_T_GE_0) eq_th - else - inst FLOAT_F_GE_0;; - -(* float_le *) -let float_le f1 f2 = - let s1, n1, e1 = dest_float f1 in - let s2, n2, e2 = dest_float f2 in - let inst = INST[n1, n1_var_num; e1, e1_var_num; - n2, n2_var_num; e2, e2_var_num] in - if s2 = "F" then - if s1 = "F" then - (* FF *) - let th0 = inst FLOAT_LE_FF in - let ltm, tm2 = dest_comb (rand (concl th0)) in - let le_th = num_exp_le (rand ltm) tm2 in - TRANS th0 le_th - else - (* TF *) - inst FLOAT_LE_TF - else - if s1 = "F" then - (* FT *) - if (is_const n1 && is_const n2) then - (* n1 = _0 and n2 = _0 *) - inst FLOAT_LE_FT_00 - else - let n1_0 = raw_eq0_hash_conv n1 in - if (fst o dest_const o rand o concl) n1_0 = "F" then - (* n1 <> _0 *) - MY_PROVE_HYP n1_0 (inst FLOAT_LE_FT_1) - else - let n2_0 = raw_eq0_hash_conv n2 in - if (fst o dest_const o rand o concl) n2_0 = "F" then - (* n2 <> _0 *) - MY_PROVE_HYP n2_0 (inst FLOAT_LE_FT_2) - else - failwith "float_lt: D0 _0 exception" - else - (* TT *) - let th0 = inst FLOAT_LE_TT in - let ltm, tm2 = dest_comb (rand (concl th0)) in - let le_th = num_exp_le (rand ltm) tm2 in - TRANS th0 le_th;; - - -(*************************************) -(* float_max, float_min *) - -let FLOAT_MIN_1 = (UNDISCH_ALL o prove)(`(f1 <= f2 <=> T) ==> min f1 f2 = f1`, REAL_ARITH_TAC);; -let FLOAT_MIN_2 = (UNDISCH_ALL o prove)(`(f1 <= f2 <=> F) ==> min f1 f2 = f2`, REAL_ARITH_TAC);; - -let FLOAT_MAX_1 = (UNDISCH_ALL o prove)(`(f1 <= f2 <=> T) ==> max f1 f2 = f2`, REAL_ARITH_TAC);; -let FLOAT_MAX_2 = (UNDISCH_ALL o prove)(`(f1 <= f2 <=> F) ==> max f1 f2 = f1`, REAL_ARITH_TAC);; - -let float_min f1 f2 = - let inst = INST[f1, f1_var_real; f2, f2_var_real] in - let le_th = float_le f1 f2 in - let th0 = - if (fst o dest_const o rand o concl) le_th = "T" then - inst FLOAT_MIN_1 - else - inst FLOAT_MIN_2 in - MY_PROVE_HYP le_th th0;; - -let float_max f1 f2 = - let inst = INST[f1, f1_var_real; f2, f2_var_real] in - let le_th = float_le f1 f2 in - let th0 = - if (fst o dest_const o rand o concl) le_th = "T" then - inst FLOAT_MAX_1 - else - inst FLOAT_MAX_2 in - MY_PROVE_HYP le_th th0;; - -let float_min_max f1 f2 = - let inst = INST[f1, f1_var_real; f2, f2_var_real] in - let le_th = float_le f1 f2 in - let th_min, th_max = - if (fst o dest_const o rand o concl) le_th = "T" then - inst FLOAT_MIN_1, inst FLOAT_MAX_1 - else - inst FLOAT_MIN_2, inst FLOAT_MAX_2 in - MY_PROVE_HYP le_th th_min, MY_PROVE_HYP le_th th_max;; - - -(*************************************) -(* FLOAT_MUL *) - -let FLOAT_MUL = prove(`!s1 s2. min_exp <= e /\ num_exp n1 e1 * num_exp n2 e2 = num_exp n e - ==> float_num s1 n1 e1 * float_num s2 n2 e2 = - float_num (mod_plus s1 s2) n (e - min_exp)`, - REPEAT STRIP_TAC THEN - REWRITE_TAC[float] THEN - ONCE_REWRITE_TAC[REAL_ARITH `(a * b / c) * (d * e / f) = (a * d) * (b * e) / c / f`] THEN - - SUBGOAL_THEN `(if s1 then -- &1 else &1) * (if s2 then -- &1 else &1) = if mod_plus s1 s2 then -- &1 else &1` MP_TAC THENL - [ - REWRITE_TAC[mod_plus] THEN - COND_CASES_TAC THEN COND_CASES_TAC THEN - REWRITE_TAC[REAL_ARITH `-- &1 * -- &1 = &1`; REAL_MUL_LID; REAL_MUL_RID]; - ALL_TAC - ] THEN - - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - REWRITE_TAC[real_div] THEN - REWRITE_TAC[GSYM REAL_MUL_ASSOC] THEN - REWRITE_TAC[REAL_EQ_MUL_LCANCEL] THEN - DISJ2_TAC THEN - - MP_TAC (SPECL[`n:num`; `e:num`; `min_exp`] NUM_EXP_SUB_lemma) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - REWRITE_TAC[REAL_MUL_ASSOC] THEN - ASM_REWRITE_TAC[REAL_OF_NUM_MUL]);; - -let FLOAT_MUL_FF = prove(`min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 = num_exp n e ==> - float_num F n1 e1 * float_num F n2 e2 = float_num F n r`, - SIMP_TAC[FLOAT_MUL; mod_plus]);; -let FLOAT_MUL_FT = prove(`min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 = num_exp n e ==> - float_num F n1 e1 * float_num T n2 e2 = float_num T n r`, - SIMP_TAC[FLOAT_MUL; mod_plus]);; -let FLOAT_MUL_TF = prove(`min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 = num_exp n e ==> - float_num T n1 e1 * float_num F n2 e2 = float_num T n r`, - SIMP_TAC[FLOAT_MUL; mod_plus]);; -let FLOAT_MUL_TT = prove(`min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 = num_exp n e ==> - float_num T n1 e1 * float_num T n2 e2 = float_num F n r`, - SIMP_TAC[FLOAT_MUL; mod_plus]);; - - -let FLOAT_MUL_0x_hi, FLOAT_MUL_0x_lo, FLOAT_MUL_x0_hi, FLOAT_MUL_x0_lo = - let mul_0x_hi = `(n1 = 0 <=> T) ==> float_num s1 n1 e1 * f2 <= float_num F 0 min_exp` in - let mul_0x_lo = `(n1 = 0 <=> T) ==> float_num F 0 min_exp <= float_num s1 n1 e1 * f2` in - let mul_x0_hi = `(n2 = 0 <=> T) ==> f1 * float_num s2 n2 e2 <= float_num F 0 min_exp` in - let mul_x0_lo = `(n2 = 0 <=> T) ==> float_num F 0 min_exp <= f1 * float_num s2 n2 e2` in - let proof = MP_TAC (GEN_ALL (SPECL [`s:bool`; `0`] FLOAT_EQ_0)) THEN - SIMP_TAC[REAL_MUL_LZERO; REAL_MUL_RZERO; REAL_LE_REFL] in - prove(mul_0x_hi, proof), prove(mul_0x_lo, proof), - prove(mul_x0_hi, proof), prove(mul_x0_lo, proof);; - - -let FLOAT_MUL_FF_hi, FLOAT_MUL_FF_lo = - let ff_hi = `min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 <= num_exp n e - ==> float_num F n1 e1 * float_num F n2 e2 <= float_num F n r` in - let ff_lo = `min_exp <= e /\ e - min_exp = r /\ num_exp n e <= num_exp n1 e1 * num_exp n2 e2 - ==> float_num F n r <= float_num F n1 e1 * float_num F n2 e2` in - let proof = - REPEAT STRIP_TAC THEN - POP_ASSUM MP_TAC THEN - POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN - REWRITE_TAC[GSYM REAL_OF_NUM_LE; GSYM REAL_OF_NUM_MUL] THEN - DISCH_TAC THEN - MAP_EVERY ABBREV_TAC [`z = &(num_exp n e)`; `x = &(num_exp n1 e1)`; `y = &(num_exp n2 e2)`] THEN - ASM_REWRITE_TAC[float; REAL_MUL_LID] THEN - REWRITE_TAC[REAL_ARITH `a / b * c / d = (a * c) / b / d`] THEN - REWRITE_TAC[real_div] THEN - REWRITE_TAC[REAL_MUL_ASSOC] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN - - MP_TAC (SPECL [`n:num`; `e:num`; `min_exp`] NUM_EXP_SUB_lemma) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - ASM_REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] in - prove(ff_hi, proof), prove(ff_lo, proof);; - -let FLOAT_MUL_TT_hi, FLOAT_MUL_TT_lo = - let tt_hi = `min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 <= num_exp n e - ==> float_num T n1 e1 * float_num T n2 e2 <= float_num F n r` in - let tt_lo = `min_exp <= e /\ e - min_exp = r /\ num_exp n e <= num_exp n1 e1 * num_exp n2 e2 - ==> float_num F n r <= float_num T n1 e1 * float_num T n2 e2` in - let proof = - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[REAL_ARITH `--a * --b = a * b`] THEN - REWRITE_TAC[FLOAT_MUL_FF_hi; FLOAT_MUL_FF_lo] in - prove(tt_hi, proof), prove(tt_lo, proof);; - -let FLOAT_MUL_FT_hi, FLOAT_MUL_FT_lo = - let ft_hi = `min_exp <= e /\ e - min_exp = r /\ num_exp n e <= num_exp n1 e1 * num_exp n2 e2 - ==> float_num F n1 e1 * float_num T n2 e2 <= float_num T n r` in - let ft_lo = `min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 <= num_exp n e - ==> float_num T n r <= float_num F n1 e1 * float_num T n2 e2` in - let proof = - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[REAL_ARITH `a * --b <= --c <=> c <= a * b`] THEN - REWRITE_TAC[REAL_ARITH `--c <= a * --b <=> a * b <= c`] THEN - REWRITE_TAC[FLOAT_MUL_FF_hi; FLOAT_MUL_FF_lo] in - prove(ft_hi, proof), prove(ft_lo, proof);; - -let FLOAT_MUL_TF_hi, FLOAT_MUL_TF_lo = - let ft_hi = `min_exp <= e /\ e - min_exp = r /\ num_exp n e <= num_exp n1 e1 * num_exp n2 e2 - ==> float_num T n1 e1 * float_num F n2 e2 <= float_num T n r` in - let ft_lo = `min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 <= num_exp n e - ==> float_num T n r <= float_num T n1 e1 * float_num F n2 e2` in - let proof = - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[REAL_ARITH `--a * b <= --c <=> c <= a * b`] THEN - REWRITE_TAC[REAL_ARITH `--c <= --a * b <=> a * b <= c`] THEN - REWRITE_TAC[FLOAT_MUL_FF_hi; FLOAT_MUL_FF_lo] in - prove(ft_hi, proof), prove(ft_lo, proof);; - - -(*********************************************) -(* float_mul_lo, float_mul_hi *) - -let transform = UNDISCH_ALL o NUMERALS_TO_NUM o PURE_REWRITE_RULE[min_exp_def; GSYM IMP_IMP];; -let FLOAT_MUL_FF_hi' = transform FLOAT_MUL_FF_hi and - FLOAT_MUL_FF_lo' = transform FLOAT_MUL_FF_lo and - FLOAT_MUL_TT_hi' = transform FLOAT_MUL_TT_hi and - FLOAT_MUL_TT_lo' = transform FLOAT_MUL_TT_lo and - FLOAT_MUL_FT_hi' = transform FLOAT_MUL_FT_hi and - FLOAT_MUL_FT_lo' = transform FLOAT_MUL_FT_lo and - FLOAT_MUL_TF_hi' = transform FLOAT_MUL_TF_hi and - FLOAT_MUL_TF_lo' = transform FLOAT_MUL_TF_lo and - FLOAT_MUL_0x_hi' = transform FLOAT_MUL_0x_hi and - FLOAT_MUL_0x_lo' = transform FLOAT_MUL_0x_lo and - FLOAT_MUL_x0_hi' = transform FLOAT_MUL_x0_hi and - FLOAT_MUL_x0_lo' = transform FLOAT_MUL_x0_lo;; - -let FLOAT_MUL_FF' = transform FLOAT_MUL_FF and - FLOAT_MUL_TT' = transform FLOAT_MUL_TT and - FLOAT_MUL_FT' = transform FLOAT_MUL_FT and - FLOAT_MUL_TF' = transform FLOAT_MUL_TF;; - -let float_mul_eq f1 f2 = - let s1, n1, e1 = dest_float f1 and - s2, n2, e2 = dest_float f2 in - let flag = s1 = s2 in - let num_exp1 = mk_num_exp n1 e1 and - num_exp2 = mk_num_exp n2 e2 in - - let mul_th = num_exp_mul num_exp1 num_exp2 in - let n_tm, e_tm = dest_num_exp (rand (concl mul_th)) in - - let sub_th, le_th = raw_sub_and_le_hash_conv e_tm min_exp_num_const in - if (rand(concl le_th) <> e_tm) then - failwith "float_mul_eq: underflow" - else - let r_tm = rand(concl sub_th) in - let inst = INST[e_tm, e_var_num; r_tm, r_var_num; n_tm, n_var_num; - n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] in - let th0 = inst - (if flag then - if s1 = "F" then FLOAT_MUL_FF' else FLOAT_MUL_TT' - else - if s1 = "F" then FLOAT_MUL_FT' else FLOAT_MUL_TF') in - MY_PROVE_HYP sub_th (MY_PROVE_HYP mul_th (MY_PROVE_HYP le_th th0));; - - -let float_mul_lo pp f1 f2 = - let s1, n1, e1 = dest_float f1 and - s2, n2, e2 = dest_float f2 in - (* Multiplication by zero *) - let n1_eq0_th = raw_eq0_hash_conv n1 in - if (rand o concl) n1_eq0_th = t_const then - (MY_PROVE_HYP n1_eq0_th o - INST[e1, e1_var_num; f2, f2_var_real; n1, n1_var_num; - (if s1 = "T" then t_const else f_const), s1_var_bool]) FLOAT_MUL_0x_lo' - else - let n2_eq0_th = raw_eq0_hash_conv n2 in - if (rand o concl) n2_eq0_th = t_const then - (MY_PROVE_HYP n2_eq0_th o - INST[e2, e2_var_num; f1, f1_var_real; n2, n2_var_num; - (if s2 = "T" then t_const else f_const), s2_var_bool]) FLOAT_MUL_x0_lo' - else - let flag = s1 = s2 in - let num_exp1 = mk_num_exp n1 e1 and - num_exp2 = mk_num_exp n2 e2 in - - let mul_th, n_tm, e_tm = - if flag then - let th = num_exp_op_lo pp num_exp_mul num_exp1 num_exp2 in - let n_tm, e_tm = dest_num_exp (lhand (concl th)) in - th, n_tm, e_tm - else - let th = num_exp_op_hi pp num_exp_mul num_exp1 num_exp2 in - let n_tm, e_tm = dest_num_exp (rand (concl th)) in - th, n_tm, e_tm in - - let sub_th, le_th = raw_sub_and_le_hash_conv e_tm min_exp_num_const in - if (rand(concl le_th) <> e_tm) then - failwith "float_mul_lo: underflow" - else - let r_tm = rand(concl sub_th) in - let inst = INST[e_tm, e_var_num; r_tm, r_var_num; n_tm, n_var_num; - n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] in - let th0 = inst - (if flag then - if s1 = "F" then FLOAT_MUL_FF_lo' else FLOAT_MUL_TT_lo' - else - if s1 = "F" then FLOAT_MUL_FT_lo' else FLOAT_MUL_TF_lo') in - MY_PROVE_HYP sub_th (MY_PROVE_HYP mul_th (MY_PROVE_HYP le_th th0));; - - -let float_mul_hi pp f1 f2 = - let s1, n1, e1 = dest_float f1 and - s2, n2, e2 = dest_float f2 in - (* Multiplication by zero *) - let n1_eq0_th = raw_eq0_hash_conv n1 in - if (rand o concl) n1_eq0_th = t_const then - (MY_PROVE_HYP n1_eq0_th o - INST[e1, e1_var_num; f2, f2_var_real; n1, n1_var_num; - (if s1 = "T" then t_const else f_const), s1_var_bool]) FLOAT_MUL_0x_hi' - else - let n2_eq0_th = raw_eq0_hash_conv n2 in - if (rand o concl) n2_eq0_th = t_const then - (MY_PROVE_HYP n2_eq0_th o - INST[e2, e2_var_num; f1, f1_var_real; n2, n2_var_num; - (if s2 = "T" then t_const else f_const), s2_var_bool]) FLOAT_MUL_x0_hi' - else - let flag = s1 = s2 in - let num_exp1 = mk_num_exp n1 e1 and - num_exp2 = mk_num_exp n2 e2 in - - let mul_th, n_tm, e_tm = - if flag then - let th = num_exp_op_hi pp num_exp_mul num_exp1 num_exp2 in - let n_tm, e_tm = dest_num_exp (rand (concl th)) in - th, n_tm, e_tm - else - let th = num_exp_op_lo pp num_exp_mul num_exp1 num_exp2 in - let n_tm, e_tm = dest_num_exp (lhand (concl th)) in - th, n_tm, e_tm in - - let sub_th, le_th = raw_sub_and_le_hash_conv e_tm min_exp_num_const in - if (rand(concl le_th) <> e_tm) then - failwith "float_mul_hi: underflow" - else - let r_tm = rand(concl sub_th) in - let inst = INST[e_tm, e_var_num; r_tm, r_var_num; n_tm, n_var_num; - n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] in - let th0 = inst - (if flag then - if s1 = "F" then FLOAT_MUL_FF_hi' else FLOAT_MUL_TT_hi' - else - if s1 = "F" then FLOAT_MUL_FT_hi' else FLOAT_MUL_TF_hi') in - MY_PROVE_HYP sub_th (MY_PROVE_HYP mul_th (MY_PROVE_HYP le_th th0));; - - - -(*********************************************) -(* FLOAT_DIV *) - -let DIV_lemma = prove(`!x y. ~(y = 0) ==> &(x DIV y) <= &x / &y /\ &x / &y <= &(x DIV y + 1)`, - REPEAT GEN_TAC THEN DISCH_TAC THEN - MP_TAC (SPECL [`y:num`; `x:num`] FLOOR_DIV_DIV) THEN - ASM_REWRITE_TAC[GSYM REAL_OF_NUM_ADD] THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN - SIMP_TAC[FLOOR; REAL_LT_IMP_LE]);; - -let FLOAT_DIV_FF = prove(`e2 + k <= min_exp + e + e1 /\ ~(n2 = 0) /\ - num_exp n1 k DIV num_exp n2 0 = num_exp n e - ==> float_num F n ((min_exp + e + e1) - (e2 + k)) <= float_num F n1 e1 / float_num F n2 e2`, - MAP_EVERY ABBREV_TAC [`z = num_exp n e`; `x = num_exp n1 k`; `y = num_exp n2 0`] THEN - REPEAT STRIP_TAC THEN - REWRITE_TAC[float; REAL_MUL_LID] THEN - REWRITE_TAC[real_div; REAL_INV_MUL; REAL_INV_INV] THEN - REWRITE_TAC[REAL_ARITH `(a * b) * c * d = (b * d) * (a * c)`] THEN - SUBGOAL_THEN `~(&(num_exp 1 min_exp) = &0)` ASSUME_TAC THENL - [ - REWRITE_TAC[num_exp; REAL_OF_NUM_EQ; MULT_CLAUSES; EXP_EQ_0] THEN - ARITH_TAC; - ALL_TAC - ] THEN - - ASM_SIMP_TAC[REAL_MUL_LINV; REAL_MUL_LID] THEN - - ASM_SIMP_TAC[NUM_EXP_SUB_lemma] THEN - SUBGOAL_THEN `&(num_exp n1 e1) * inv(&(num_exp n2 e2)) = (&x / &y) * &(num_exp 1 e1) * inv(&(num_exp 1 (e2 + k)))` MP_TAC THENL - [ - EXPAND_TAC "x" THEN EXPAND_TAC "y" THEN - REWRITE_TAC[real_div] THEN - REWRITE_TAC[num_exp; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW] THEN - REWRITE_TAC[REAL_MUL_LID; REAL_INV_MUL; REAL_INV_1; real_pow; REAL_MUL_RID] THEN - REWRITE_TAC[REAL_POW_ADD; REAL_INV_MUL] THEN - REWRITE_TAC[REAL_ARITH `((a * b) * c) * d * e * f = (b * f) * (a * c * d * e)`] THEN - - SUBGOAL_THEN (mk_comb(`(~)`, mk_eq(mk_binop `pow` (mk_comb (`&`, base_const)) `k:num`, `&0`))) ASSUME_TAC THENL - [ - REWRITE_TAC[REAL_POW_EQ_0] THEN - REAL_ARITH_TAC; - ALL_TAC - ] THEN - - ASM_SIMP_TAC[REAL_MUL_RINV; REAL_MUL_LID] THEN - REAL_ARITH_TAC; - ALL_TAC - ] THEN - - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - ONCE_REWRITE_TAC[REAL_ARITH `(a * b) * c = (a * c) * b`] THEN - REWRITE_TAC[REAL_MUL_ASSOC] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN - ONCE_REWRITE_TAC[NUM_EXP_SUM1] THEN - REWRITE_TAC[NUM_EXP_SUM] THEN - REWRITE_TAC[GSYM REAL_OF_NUM_MUL] THEN - ASM_REWRITE_TAC[REAL_ARITH `(a * b * c) * d = (d * a) * b * c`] THEN - ASM_SIMP_TAC[REAL_MUL_LINV; REAL_MUL_LID] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN - MP_TAC (SPEC_ALL DIV_lemma) THEN - ANTS_TAC THENL - [ - EXPAND_TAC "y" THEN - REWRITE_TAC[num_exp; MULT_EQ_0; DE_MORGAN_THM] THEN - ASM_REWRITE_TAC[EXP] THEN - ARITH_TAC; - ALL_TAC - ] THEN - - ASM_SIMP_TAC[]);; - -let FLOAT_DIV_0x_lo = prove(`(n1 = 0 <=> T) ==> float_num F 0 min_exp <= float_num s1 n1 e1 / f2`, - SIMP_TAC[real_div; FLOAT_MUL_0x_lo]);; - -let FLOAT_DIV_0x_hi = prove(`(n1 = 0 <=> T) ==> float_num s1 n1 e1 / f2 <= float_num F 0 min_exp`, - SIMP_TAC[real_div; FLOAT_MUL_0x_hi]);; - -let FLOAT_DIV_FF_lo = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ r2 - r1 = r /\ - r1 <= r2 /\ ~(n2 = 0) /\ - num_exp n e <= num_exp n1 k DIV num_exp n2 0 - ==> float_num F n r <= float_num F n1 e1 / float_num F n2 e2`, - MAP_EVERY ABBREV_TAC [`z = num_exp n e`; `x = num_exp n1 k`; `y = num_exp n2 0`] THEN - REPEAT STRIP_TAC THEN - REPLICATE_TAC 3 (POP_ASSUM MP_TAC) THEN - REPLICATE_TAC 3 (POP_ASSUM (fun th -> REWRITE_TAC[SYM th])) THEN - REPEAT STRIP_TAC THEN - REWRITE_TAC[float; REAL_MUL_LID] THEN - REWRITE_TAC[real_div; REAL_INV_MUL; REAL_INV_INV] THEN - REWRITE_TAC[REAL_ARITH `(a * b) * c * d = (b * d) * (a * c)`] THEN - SUBGOAL_THEN `~(&(num_exp 1 min_exp) = &0)` ASSUME_TAC THENL - [ - REWRITE_TAC[num_exp; REAL_OF_NUM_EQ; MULT_CLAUSES; EXP_EQ_0] THEN - ARITH_TAC; - ALL_TAC - ] THEN - - ASM_SIMP_TAC[REAL_MUL_LINV; REAL_MUL_LID] THEN - - ASM_SIMP_TAC[NUM_EXP_SUB_lemma] THEN - SUBGOAL_THEN `&(num_exp n1 e1) * inv(&(num_exp n2 e2)) = (&x / &y) * &(num_exp 1 e1) * inv(&(num_exp 1 (e2 + k)))` MP_TAC THENL - [ - EXPAND_TAC "x" THEN EXPAND_TAC "y" THEN - REWRITE_TAC[real_div] THEN - REWRITE_TAC[num_exp; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW] THEN - REWRITE_TAC[REAL_MUL_LID; REAL_INV_MUL; REAL_INV_1; real_pow; REAL_MUL_RID] THEN - REWRITE_TAC[REAL_POW_ADD; REAL_INV_MUL] THEN - REWRITE_TAC[REAL_ARITH `((a * b) * c) * d * e * f = (b * f) * (a * c * d * e)`] THEN - SUBGOAL_THEN - (mk_comb(`(~)`, mk_eq(mk_binop `pow` (mk_comb (`&`, base_const)) `k:num`, `&0`))) ASSUME_TAC THENL - [ - REWRITE_TAC[REAL_POW_EQ_0] THEN - REAL_ARITH_TAC; - ALL_TAC - ] THEN - - ASM_SIMP_TAC[REAL_MUL_RINV; REAL_MUL_LID] THEN - REAL_ARITH_TAC; - ALL_TAC - ] THEN - - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - ONCE_REWRITE_TAC[REAL_ARITH `(a * b) * c = (a * c) * b`] THEN - REWRITE_TAC[REAL_MUL_ASSOC] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN - ONCE_REWRITE_TAC[NUM_EXP_SUM1] THEN - REWRITE_TAC[NUM_EXP_SUM] THEN - REWRITE_TAC[GSYM REAL_OF_NUM_MUL] THEN - ASM_REWRITE_TAC[REAL_ARITH `(a * b * c) * d = (d * a) * b * c`] THEN - ASM_SIMP_TAC[REAL_MUL_LINV; REAL_MUL_LID] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN - MP_TAC (SPEC_ALL DIV_lemma) THEN - ANTS_TAC THENL - [ - EXPAND_TAC "y" THEN - REWRITE_TAC[num_exp; MULT_EQ_0; DE_MORGAN_THM] THEN - ASM_REWRITE_TAC[EXP] THEN - ARITH_TAC; - ALL_TAC - ] THEN - - STRIP_TAC THEN - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `&(x DIV y)` THEN - ASM_REWRITE_TAC[REAL_OF_NUM_LE]);; - -let FLOAT_DIV_FF_hi = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ r2 - r1 = r /\ - r1 <= r2 /\ ~(n2 = 0) /\ - num_exp n1 k DIV num_exp n2 0 < num_exp n e - ==> float_num F n1 e1 / float_num F n2 e2 <= float_num F n r`, - MAP_EVERY ABBREV_TAC [`z = num_exp n e`; `x = num_exp n1 k`; `y = num_exp n2 0`] THEN - REPEAT STRIP_TAC THEN - REPLICATE_TAC 3 (POP_ASSUM MP_TAC) THEN - REPLICATE_TAC 3 (POP_ASSUM (fun th -> REWRITE_TAC[SYM th])) THEN - REPEAT STRIP_TAC THEN - REWRITE_TAC[float; REAL_MUL_LID] THEN - REWRITE_TAC[real_div; REAL_INV_MUL; REAL_INV_INV] THEN - REWRITE_TAC[REAL_ARITH `(a * b) * c * d = (b * d) * (a * c)`] THEN - SUBGOAL_THEN `~(&(num_exp 1 min_exp) = &0)` ASSUME_TAC THENL - [ - REWRITE_TAC[num_exp; REAL_OF_NUM_EQ; MULT_CLAUSES; EXP_EQ_0] THEN - ARITH_TAC; - ALL_TAC - ] THEN - - ASM_SIMP_TAC[REAL_MUL_LINV; REAL_MUL_LID] THEN - - ASM_SIMP_TAC[NUM_EXP_SUB_lemma] THEN - SUBGOAL_THEN `&(num_exp n1 e1) * inv(&(num_exp n2 e2)) = (&x / &y) * &(num_exp 1 e1) * inv(&(num_exp 1 (e2 + k)))` MP_TAC THENL - [ - EXPAND_TAC "x" THEN EXPAND_TAC "y" THEN - REWRITE_TAC[real_div] THEN - REWRITE_TAC[num_exp; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW] THEN - REWRITE_TAC[REAL_MUL_LID; REAL_INV_MUL; REAL_INV_1; real_pow; REAL_MUL_RID] THEN - REWRITE_TAC[REAL_POW_ADD; REAL_INV_MUL] THEN - REWRITE_TAC[REAL_ARITH `((a * b) * c) * d * e * f = (b * f) * (a * c * d * e)`] THEN - SUBGOAL_THEN - (mk_comb(`(~)`, mk_eq(mk_binop `pow` (mk_comb (`&`, base_const)) `k:num`, `&0`))) ASSUME_TAC THENL - [ - REWRITE_TAC[REAL_POW_EQ_0] THEN - REAL_ARITH_TAC; - ALL_TAC - ] THEN - - ASM_SIMP_TAC[REAL_MUL_RINV; REAL_MUL_LID] THEN - REAL_ARITH_TAC; - ALL_TAC - ] THEN - - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - ONCE_REWRITE_TAC[REAL_ARITH `(a * b) * c = (a * c) * b`] THEN - REWRITE_TAC[REAL_MUL_ASSOC] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN - ONCE_REWRITE_TAC[NUM_EXP_SUM1] THEN - REWRITE_TAC[NUM_EXP_SUM] THEN - REWRITE_TAC[GSYM REAL_OF_NUM_MUL] THEN - ASM_REWRITE_TAC[REAL_ARITH `(a * b * c) * d = (d * a) * b * c`] THEN - ASM_SIMP_TAC[REAL_MUL_LINV; REAL_MUL_LID] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN - MP_TAC (SPEC_ALL DIV_lemma) THEN - ANTS_TAC THENL - [ - EXPAND_TAC "y" THEN - REWRITE_TAC[num_exp; MULT_EQ_0; DE_MORGAN_THM] THEN - ASM_REWRITE_TAC[EXP] THEN - ARITH_TAC; - ALL_TAC - ] THEN - - STRIP_TAC THEN - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `&(x DIV y + 1)` THEN - ASM_REWRITE_TAC[REAL_OF_NUM_LE] THEN - UNDISCH_TAC `x DIV y < z` THEN - ARITH_TAC);; - -let FLOAT_DIV_TT_lo = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ r2 - r1 = r /\ - r1 <= r2 /\ ~(n2 = 0) /\ - num_exp n e <= num_exp n1 k DIV num_exp n2 0 - ==> float_num F n r <= float_num T n1 e1 / float_num T n2 e2`, - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[real_div; REAL_INV_NEG; REAL_NEG_MUL2] THEN - REWRITE_TAC[GSYM real_div] THEN - REWRITE_TAC[FLOAT_DIV_FF_lo]);; - -let FLOAT_DIV_TT_hi = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ - r2 - r1 = r /\ r1 <= r2 /\ ~(n2 = 0) /\ - num_exp n1 k DIV num_exp n2 0 < num_exp n e - ==> float_num T n1 e1 / float_num T n2 e2 <= float_num F n r`, - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[real_div; REAL_INV_NEG; REAL_NEG_MUL2] THEN - REWRITE_TAC[GSYM real_div] THEN - REWRITE_TAC[FLOAT_DIV_FF_hi]);; - - -let FLOAT_DIV_FT_lo = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ - r2 - r1 = r /\ r1 <= r2 /\ ~(n2 = 0) /\ - num_exp n1 k DIV num_exp n2 0 < num_exp n e - ==> float_num T n r <= float_num F n1 e1 / float_num T n2 e2`, - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[real_div; REAL_INV_NEG] THEN - REWRITE_TAC[REAL_ARITH `--a <= b * --c <=> b * c <= a`] THEN - REWRITE_TAC[GSYM real_div] THEN - REWRITE_TAC[FLOAT_DIV_FF_hi]);; - -let FLOAT_DIV_FT_hi = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ - r2 - r1 = r /\ r1 <= r2 /\ ~(n2 = 0) /\ - num_exp n e <= num_exp n1 k DIV num_exp n2 0 - ==> float_num F n1 e1 / float_num T n2 e2 <= float_num T n r`, - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[real_div; REAL_INV_NEG] THEN - REWRITE_TAC[REAL_ARITH `a * --b <= --c <=> c <= a * b`] THEN - REWRITE_TAC[GSYM real_div] THEN - REWRITE_TAC[FLOAT_DIV_FF_lo]);; - - -let FLOAT_DIV_TF_lo = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ - r2 - r1 = r /\ r1 <= r2 /\ ~(n2 = 0) /\ - num_exp n1 k DIV num_exp n2 0 < num_exp n e - ==> float_num T n r <= float_num T n1 e1 / float_num F n2 e2`, - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[real_div; REAL_INV_NEG] THEN - REWRITE_TAC[REAL_ARITH `--a <= --b * c <=> b * c <= a`] THEN - REWRITE_TAC[GSYM real_div] THEN - REWRITE_TAC[FLOAT_DIV_FF_hi]);; - -let FLOAT_DIV_TF_hi = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ - r2 - r1 = r /\ r1 <= r2 /\ ~(n2 = 0) /\ - num_exp n e <= num_exp n1 k DIV num_exp n2 0 - ==> float_num T n1 e1 / float_num F n2 e2 <= float_num T n r`, - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[real_div; REAL_INV_NEG] THEN - REWRITE_TAC[REAL_ARITH `--a * b <= --c <=> c <= a * b`] THEN - REWRITE_TAC[GSYM real_div] THEN - REWRITE_TAC[FLOAT_DIV_FF_lo]);; - - -(******************************************) -(* float_div_lo, float_div_hi *) - -let transform = UNDISCH_ALL o PURE_REWRITE_RULE[TAUT `~P <=> (P <=> F)`] o - NUMERALS_TO_NUM o PURE_REWRITE_RULE[GSYM IMP_IMP; min_exp_def];; - -let FLOAT_DIV_FF_hi' = transform FLOAT_DIV_FF_hi and - FLOAT_DIV_FF_lo' = transform FLOAT_DIV_FF_lo and - FLOAT_DIV_TT_hi' = transform FLOAT_DIV_TT_hi and - FLOAT_DIV_TT_lo' = transform FLOAT_DIV_TT_lo and - FLOAT_DIV_FT_hi' = transform FLOAT_DIV_FT_hi and - FLOAT_DIV_FT_lo' = transform FLOAT_DIV_FT_lo and - FLOAT_DIV_TF_hi' = transform FLOAT_DIV_TF_hi and - FLOAT_DIV_TF_lo' = transform FLOAT_DIV_TF_lo and - FLOAT_DIV_0x_hi' = transform FLOAT_DIV_0x_hi and - FLOAT_DIV_0x_lo' = transform FLOAT_DIV_0x_lo;; - -let float_div_lo pp f1 f2 = - let s1, n1, e1 = dest_float f1 and - s2, n2, e2 = dest_float f2 in - let n1_eq0_th = raw_eq0_hash_conv n1 in - if (rand o concl) n1_eq0_th = t_const then - (MY_PROVE_HYP n1_eq0_th o - INST[e1, e1_var_num; f2, f2_var_real; n1, n1_var_num; - (if s1 = "T" then t_const else f_const), s1_var_bool]) FLOAT_DIV_0x_lo' - else - let flag = s1 = s2 in - - let k_tm = rand (mk_small_numeral_array (2 * pp)) in - let num_exp1 = mk_num_exp n1 k_tm and - num_exp2 = mk_num_exp n2 zero_const in - let div_th, n_tm, e_tm = - if flag then - let th = num_exp_op_lo pp num_exp_div num_exp1 num_exp2 in - let n_tm, e_tm = dest_num_exp (lhand(concl th)) in - th, n_tm, e_tm - else - let th = num_exp_op_hi_lt pp num_exp_div num_exp1 num_exp2 in - let n_tm, e_tm = dest_num_exp (rand(concl th)) in - th, n_tm, e_tm in - - let r1_th = raw_add_conv_hash (mk_binop add_op_num e2 k_tm) in - let r1_tm = rand(concl r1_th) in - let e_plus_e1 = raw_add_conv_hash (mk_binop add_op_num e_tm e1) in - let ltm, rtm = dest_comb(concl e_plus_e1) in - let r2_th' = raw_add_conv_hash (mk_binop add_op_num min_exp_num_const rtm) in - let r2_th = TRANS (AP_TERM (mk_comb (add_op_num, min_exp_num_const)) e_plus_e1) r2_th' in - let r2_tm = rand(concl r2_th) in - let sub_th, le_th = raw_sub_and_le_hash_conv r2_tm r1_tm in - if rand(concl le_th) <> r2_tm then - failwith "float_div_lo: underflow" - else - let r_tm = rand(concl sub_th) in - let n2_not_zero = raw_eq0_hash_conv n2 in - let inst = INST[r1_tm, r1_var_num; r2_tm, r2_var_num; - n1, n1_var_num; e1, e1_var_num; - e_tm, e_var_num; k_tm, k_var_num; - n2, n2_var_num; e2, e2_var_num; - n_tm, n_var_num; r_tm, r_var_num] in - let th0 = inst - (if flag then - if s1 = "F" then FLOAT_DIV_FF_lo' else FLOAT_DIV_TT_lo' - else - if s1 = "F" then FLOAT_DIV_FT_lo' else FLOAT_DIV_TF_lo') in - let th1 = MY_PROVE_HYP n2_not_zero (MY_PROVE_HYP div_th (MY_PROVE_HYP le_th th0)) in - MY_PROVE_HYP sub_th (MY_PROVE_HYP r2_th (MY_PROVE_HYP r1_th th1));; - - -let float_div_hi pp f1 f2 = - let s1, n1, e1 = dest_float f1 and - s2, n2, e2 = dest_float f2 in - let n1_eq0_th = raw_eq0_hash_conv n1 in - if (rand o concl) n1_eq0_th = t_const then - (MY_PROVE_HYP n1_eq0_th o - INST[e1, e1_var_num; f2, f2_var_real; n1, n1_var_num; - (if s1 = "T" then t_const else f_const), s1_var_bool]) FLOAT_DIV_0x_hi' - else - let flag = s1 = s2 in - - let k_tm = rand (mk_small_numeral_array (2 * pp)) in - let num_exp1 = mk_num_exp n1 k_tm and - num_exp2 = mk_num_exp n2 zero_const in - let div_th, n_tm, e_tm = - if flag then - let th = num_exp_op_hi_lt pp num_exp_div num_exp1 num_exp2 in - let n_tm, e_tm = dest_num_exp (rand(concl th)) in - th, n_tm, e_tm - else - let th = num_exp_op_lo pp num_exp_div num_exp1 num_exp2 in - let n_tm, e_tm = dest_num_exp (lhand(concl th)) in - th, n_tm, e_tm in - - let r1_th = raw_add_conv_hash (mk_binop add_op_num e2 k_tm) in - let r1_tm = rand(concl r1_th) in - let e_plus_e1 = raw_add_conv_hash (mk_binop add_op_num e_tm e1) in - let ltm, rtm = dest_comb(concl e_plus_e1) in - let r2_th' = raw_add_conv_hash (mk_binop add_op_num min_exp_num_const rtm) in - let r2_th = TRANS (AP_TERM (mk_comb (add_op_num, min_exp_num_const)) e_plus_e1) r2_th' in - let r2_tm = rand(concl r2_th) in - let sub_th, le_th = raw_sub_and_le_hash_conv r2_tm r1_tm in - if rand(concl le_th) <> r2_tm then - failwith "float_div_hi: underflow" - else - let r_tm = rand(concl sub_th) in - let n2_not_zero = raw_eq0_hash_conv n2 in - let inst = INST[r1_tm, r1_var_num; r2_tm, r2_var_num; - n1, n1_var_num; e1, e1_var_num; - e_tm, e_var_num; k_tm, k_var_num; - n2, n2_var_num; e2, e2_var_num; - n_tm, n_var_num; r_tm, r_var_num] in - let th0 = inst - (if flag then - if s1 = "F" then FLOAT_DIV_FF_hi' else FLOAT_DIV_TT_hi' - else - if s1 = "F" then FLOAT_DIV_FT_hi' else FLOAT_DIV_TF_hi') in - let th1 = MY_PROVE_HYP n2_not_zero (MY_PROVE_HYP div_th (MY_PROVE_HYP le_th th0)) in - MY_PROVE_HYP sub_th (MY_PROVE_HYP r2_th (MY_PROVE_HYP r1_th th1));; - - -(***********************************) -(* FLOAT_ADD *) - -let FLOAT_ADD_FF = prove(`num_exp n1 e1 + num_exp n2 e2 = num_exp n e - ==> float_num F n1 e1 + float_num F n2 e2 = float_num F n e`, - REPEAT STRIP_TAC THEN - REWRITE_TAC[float; REAL_MUL_LID] THEN - REWRITE_TAC[REAL_ARITH `a / b + c / b = (a + c) / b`] THEN - ASM_REWRITE_TAC[REAL_OF_NUM_ADD]);; - -let FLOAT_ADD_TT = prove(`num_exp n1 e1 + num_exp n2 e2 = num_exp n e - ==> float_num T n1 e1 + float_num T n2 e2 = float_num T n e`, - REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `--a + --b = --c <=> a + b = c`] THEN - REWRITE_TAC[FLOAT_ADD_FF]);; - -let FLOAT_ADD_FF_lo = prove(`num_exp n e <= num_exp n1 e1 + num_exp n2 e2 - ==> float_num F n e <= float_num F n1 e1 + float_num F n2 e2`, - REWRITE_TAC[GSYM REAL_OF_NUM_LE; GSYM REAL_OF_NUM_ADD] THEN - REPEAT STRIP_TAC THEN - MAP_EVERY ABBREV_TAC [`z = &(num_exp n e)`; `x = &(num_exp n1 e1)`; `y = &(num_exp n2 e2)`] THEN - ASM_REWRITE_TAC[float; REAL_MUL_LID] THEN - REWRITE_TAC[REAL_ARITH `a / b + c / b = (a + c) / b`] THEN - REWRITE_TAC[real_div] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - ASM_REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS]);; - -let FLOAT_ADD_FF_hi = prove(`num_exp n1 e1 + num_exp n2 e2 <= num_exp n e - ==> float_num F n1 e1 + float_num F n2 e2 <= float_num F n e`, - REWRITE_TAC[GSYM REAL_OF_NUM_LE; GSYM REAL_OF_NUM_ADD] THEN - REPEAT STRIP_TAC THEN - MAP_EVERY ABBREV_TAC [`z = &(num_exp n e)`; `x = &(num_exp n1 e1)`; `y = &(num_exp n2 e2)`] THEN - ASM_REWRITE_TAC[float; REAL_MUL_LID] THEN - REWRITE_TAC[REAL_ARITH `a / b + c / b = (a + c) / b`] THEN - REWRITE_TAC[real_div] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - ASM_REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS]);; - -let FLOAT_ADD_TT_lo = prove(`num_exp n1 e1 + num_exp n2 e2 <= num_exp n e - ==> float_num T n e <= float_num T n1 e1 + float_num T n2 e2`, - REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `--a <= --b + --c <=> b + c <= a`] THEN - REWRITE_TAC[FLOAT_ADD_FF_hi]);; - -let FLOAT_ADD_TT_hi = prove(`num_exp n e <= num_exp n1 e1 + num_exp n2 e2 - ==> float_num T n1 e1 + float_num T n2 e2 <= float_num T n e`, - REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `--b + --c <= --a <=> a <= b + c`] THEN - REWRITE_TAC[FLOAT_ADD_FF_lo]);; - -let FLOAT_ADD_FT_F_lo = prove(`num_exp n2 e2 <= num_exp n1 e1 ==> - num_exp n e <= num_exp n1 e1 - num_exp n2 e2 - ==> float_num F n e <= float_num F n1 e1 + float_num T n2 e2`, - MAP_EVERY ABBREV_TAC[`z = num_exp n e`; `x = num_exp n1 e1`; `y = num_exp n2 e2`] THEN - ASM_REWRITE_TAC[FLOAT_NEG_T; float; REAL_MUL_LID] THEN - DISCH_TAC THEN - ASM_SIMP_TAC[GSYM REAL_OF_NUM_LE; GSYM REAL_OF_NUM_SUB] THEN - REWRITE_TAC[num_exp; min_exp_def; MULT_CLAUSES; GSYM REAL_OF_NUM_POW] THEN - REAL_ARITH_TAC);; - -let FLOAT_ADD_FT_T_lo = prove(`num_exp n1 e1 <= num_exp n2 e2 ==> - num_exp n2 e2 - num_exp n1 e1 <= num_exp n e - ==> float_num T n e <= float_num F n1 e1 + float_num T n2 e2`, - MAP_EVERY ABBREV_TAC[`z = num_exp n e`; `x = num_exp n1 e1`; `y = num_exp n2 e2`] THEN - ASM_REWRITE_TAC[FLOAT_NEG_T; float; REAL_MUL_LID] THEN - DISCH_TAC THEN - ASM_SIMP_TAC[GSYM REAL_OF_NUM_LE; GSYM REAL_OF_NUM_SUB] THEN - REWRITE_TAC[num_exp; min_exp_def; MULT_CLAUSES; GSYM REAL_OF_NUM_POW] THEN - REAL_ARITH_TAC);; - -let FLOAT_ADD_FT_F_hi = prove(`num_exp n2 e2 <= num_exp n1 e1 ==> - num_exp n1 e1 - num_exp n2 e2 <= num_exp n e - ==> float_num F n1 e1 + float_num T n2 e2 <= float_num F n e`, - REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `a + --b <= c <=> --c <= b + --a`] THEN - REWRITE_TAC[GSYM FLOAT_NEG_T; FLOAT_ADD_FT_T_lo]);; - -let FLOAT_ADD_FT_T_hi = prove(`num_exp n1 e1 <= num_exp n2 e2 ==> - num_exp n e <= num_exp n2 e2 - num_exp n1 e1 - ==> float_num F n1 e1 + float_num T n2 e2 <= float_num T n e`, - REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `a + --b <= --c <=> c <= b + --a`] THEN - REWRITE_TAC[GSYM FLOAT_NEG_T; FLOAT_ADD_FT_F_lo]);; - - -(******************************************) -(* float_add_lo, float_add_hi *) - -let REAL_ADD_COMM = CONJUNCT1 REAL_ADD_AC;; - -let transform = UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o NUMERALS_TO_NUM;; -let FLOAT_ADD_FF_hi' = transform FLOAT_ADD_FF_hi and - FLOAT_ADD_FF_lo' = transform FLOAT_ADD_FF_lo and - FLOAT_ADD_TT_hi' = transform FLOAT_ADD_TT_hi and - FLOAT_ADD_TT_lo' = transform FLOAT_ADD_TT_lo and - FLOAT_ADD_FT_F_lo' = transform FLOAT_ADD_FT_F_lo and - FLOAT_ADD_FT_T_lo' = transform FLOAT_ADD_FT_T_lo and - FLOAT_ADD_FT_F_hi' = transform FLOAT_ADD_FT_F_hi and - FLOAT_ADD_FT_T_hi' = transform FLOAT_ADD_FT_T_hi;; - -let float_add_lo pp f1 f2 = - let s1, n1, e1 = dest_float f1 in - let s2, n2, e2 = dest_float f2 in - if s1 = s2 then - let num_exp1 = mk_num_exp n1 e1 in - let num_exp2 = mk_num_exp n2 e2 in - - if s1 = "F" then - (* F + F *) - let add_th = num_exp_op_lo pp num_exp_add num_exp1 num_exp2 in - let n_tm, e_tm = dest_num_exp (lhand(concl add_th)) in - let th0 = INST[e_tm, e_var_num; n_tm, n_var_num; n1, n1_var_num; - e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] FLOAT_ADD_FF_lo' in - MY_PROVE_HYP add_th th0 - else - (* T + T *) - let add_th = num_exp_op_hi pp num_exp_add num_exp1 num_exp2 in - let n_tm, e_tm = dest_num_exp (rand(concl add_th)) in - let th0 = INST[e_tm, e_var_num; n_tm, n_var_num; n1, n1_var_num; - e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] FLOAT_ADD_TT_lo' in - MY_PROVE_HYP add_th th0 - else - (* F + T or T + F *) - let th0, n1, e1, n2, e2 = - if s1 = "T" then - INST[f2, m_var_real; f1, n_var_real] REAL_ADD_COMM, n2, e2, n1, e1 - else - REFL(mk_binop add_op_real f1 f2), n1, e1, n2, e2 in - - let num_exp1 = mk_num_exp n1 e1 in - let num_exp2 = mk_num_exp n2 e2 in - - let sub_th, le_th = num_exp_sub num_exp1 num_exp2 in - let sub_tm = rand(concl sub_th) in - - if rand(concl le_th) = num_exp1 then - let lo_th = num_exp_lo pp sub_tm in - let n_tm, e_tm = dest_num_exp (lhand(concl lo_th)) in - let lo_sub_th = EQ_MP (AP_TERM (rator(concl lo_th)) (SYM sub_th)) lo_th in - - let th1 = INST[n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num; - n_tm, n_var_num; e_tm, e_var_num] FLOAT_ADD_FT_F_lo' in - let th2 = MY_PROVE_HYP lo_sub_th (MY_PROVE_HYP le_th th1) in - EQ_MP (AP_TERM (rator(concl th2)) th0) th2 - - else - let hi_th = num_exp_hi pp sub_tm in - let n_tm, e_tm = dest_num_exp(rand(concl hi_th)) in - let hi_sub_th = EQ_MP (SYM (AP_THM (AP_TERM le_op_num sub_th) (rand(concl hi_th)))) hi_th in - - let th1 = INST[n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num; - n_tm, n_var_num; e_tm, e_var_num] FLOAT_ADD_FT_T_lo' in - let th2 = MY_PROVE_HYP hi_sub_th (MY_PROVE_HYP le_th th1) in - EQ_MP (AP_TERM (rator(concl th2)) th0) th2;; - - -let float_add_hi pp f1 f2 = - let s1, n1, e1 = dest_float f1 in - let s2, n2, e2 = dest_float f2 in - if s1 = s2 then - let num_exp1 = mk_num_exp n1 e1 in - let num_exp2 = mk_num_exp n2 e2 in - - if s1 = "F" then - (* F + F *) - let add_th = num_exp_op_hi pp num_exp_add num_exp1 num_exp2 in - let n_tm, e_tm = dest_num_exp (rand(concl add_th)) in - let th0 = INST[e_tm, e_var_num; n_tm, n_var_num; n1, n1_var_num; - e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] FLOAT_ADD_FF_hi' in - MY_PROVE_HYP add_th th0 - else - (* T + T *) - let add_th = num_exp_op_lo pp num_exp_add num_exp1 num_exp2 in - let n_tm, e_tm = dest_num_exp (lhand(concl add_th)) in - let th0 = INST[e_tm, e_var_num; n_tm, n_var_num; n1, n1_var_num; - e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] FLOAT_ADD_TT_hi' in - MY_PROVE_HYP add_th th0 - else - (* F + T or T + F *) - let th0, n1, e1, n2, e2 = - if s1 = "T" then - INST[f2, m_var_real; f1, n_var_real] REAL_ADD_COMM, n2, e2, n1, e1 - else - REFL(mk_binop add_op_real f1 f2), n1, e1, n2, e2 in - - let num_exp1 = mk_num_exp n1 e1 in - let num_exp2 = mk_num_exp n2 e2 in - - let sub_th, le_th = num_exp_sub num_exp1 num_exp2 in - let sub_tm = rand(concl sub_th) in - - if rand(concl le_th) = num_exp1 then - let hi_th = num_exp_hi pp sub_tm in - let n_tm, e_tm = dest_num_exp (rand(concl hi_th)) in - let hi_sub_th = EQ_MP (SYM (AP_THM (AP_TERM le_op_num sub_th) (rand(concl hi_th)))) hi_th in - - let th1 = INST[n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num; - n_tm, n_var_num; e_tm, e_var_num] FLOAT_ADD_FT_F_hi' in - let th2 = MY_PROVE_HYP hi_sub_th (MY_PROVE_HYP le_th th1) in - EQ_MP (AP_THM (AP_TERM le_op_real th0) (rand(concl th2))) th2 - - else - let lo_th = num_exp_lo pp sub_tm in - let n_tm, e_tm = dest_num_exp(lhand(concl lo_th)) in - let lo_sub_th = EQ_MP (AP_TERM (rator(concl lo_th)) (SYM sub_th)) lo_th in - - let th1 = INST[n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num; - n_tm, n_var_num; e_tm, e_var_num] FLOAT_ADD_FT_T_hi' in - let th2 = MY_PROVE_HYP lo_sub_th (MY_PROVE_HYP le_th th1) in - EQ_MP (AP_THM (AP_TERM le_op_real th0) (rand(concl th2))) th2;; - - -(******************************************) -(* float_sub_lo, float_sub_hi *) - -let FLOAT_SUB_F_EQ_ADD = (SYM o prove)(`f1 - float_num F n2 e2 = f1 + float_num T n2 e2`, - REWRITE_TAC[FLOAT_NEG_T] THEN REAL_ARITH_TAC);; - -let FLOAT_SUB_T_EQ_ADD = (SYM o prove)(`f1 - float_num T n2 e2 = f1 + float_num F n2 e2`, - REWRITE_TAC[FLOAT_NEG_T] THEN REAL_ARITH_TAC);; - -let float_sub_lo pp f1 f2 = - let s2, n2, e2 = dest_float f2 in - let th0 = - INST[f1, f1_var_real; n2, n2_var_num; e2, e2_var_num] - (if s2 = "F" then FLOAT_SUB_F_EQ_ADD else FLOAT_SUB_T_EQ_ADD) in - let ltm,f2_tm = dest_comb(lhand(concl th0)) in - let f1_tm = rand ltm in - let lo_th = float_add_lo pp f1_tm f2_tm in - EQ_MP (AP_TERM (rator(concl lo_th)) th0) lo_th;; - - -let float_sub_hi pp f1 f2 = - let s2, n2, e2 = dest_float f2 in - let th0 = - INST[f1, f1_var_real; n2, n2_var_num; e2, e2_var_num] - (if s2 = "F" then FLOAT_SUB_F_EQ_ADD else FLOAT_SUB_T_EQ_ADD) in - let ltm, f2_tm = dest_comb(lhand(concl th0)) in - let f1_tm = rand ltm in - let hi_th = float_add_hi pp f1_tm f2_tm in - EQ_MP (AP_THM (AP_TERM le_op_real th0) (rand(concl hi_th))) hi_th;; - - -(*******************************************) -(* FLOAT_SQRT *) - -(* float_num F m e = float_num F (B0 m) (PRE e) *) -let FLOAT_PRE_EXP = prove(mk_imp(`~(e = 0) /\ PRE e = e1`, - mk_eq(`float_num F m e`, - mk_comb(mk_comb(`float_num F`, mk_comb(b0_const, m_var_num)), `e1:num`))), - STRIP_TAC THEN POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN - REWRITE_TAC[float; REAL_MUL_LID; real_div; REAL_EQ_MUL_RCANCEL] THEN - DISJ1_TAC THEN - REWRITE_TAC[num_exp; b0_thm; REAL_OF_NUM_EQ] THEN - SUBGOAL_THEN `e = SUC (PRE e)` MP_TAC THENL - [ - POP_ASSUM MP_TAC THEN ARITH_TAC; - ALL_TAC - ] THEN - DISCH_THEN (fun th -> CONV_TAC (LAND_CONV (ONCE_REWRITE_CONV[th]))) THEN - REWRITE_TAC[EXP] THEN - ARITH_TAC);; - -let DIV2_EVEN_lemma = prove(`!n. EVEN n ==> 2 * (n DIV 2) = n`, - GEN_TAC THEN - REWRITE_TAC[EVEN_EXISTS] THEN - STRIP_TAC THEN - ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC (ARITH_RULE `x = y ==> 2 * x = 2 * y`) THEN - MATCH_MP_TAC DIV_MULT THEN - ARITH_TAC);; - -let FLOAT_SQRT_EVEN_lo = prove(`f1 * f1 = f2 /\ f2 <= x /\ - num_exp m (2 * p) = x /\ f1 = num_exp n1 e1 - /\ EVEN e /\ e DIV 2 = e2 /\ - e1 + e2 + (min_exp DIV 2) = r /\ - p <= r /\ r - p = r2 - ==> float_num F n1 r2 <= sqrt (float_num F m e)`, - STRIP_TAC THEN - UNDISCH_TAC `f2 <= x:num` THEN - UNDISCH_TAC `num_exp m (2 * p) = x` THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN - UNDISCH_TAC `f1 * f1 = f2:num` THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN - UNDISCH_TAC `e1 + e2 + min_exp DIV 2 = r` THEN - UNDISCH_TAC `e DIV 2 = e2` THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN - REPEAT (POP_ASSUM MP_TAC) THEN - REWRITE_TAC[num_exp; float; REAL_MUL_LID; GSYM REAL_OF_NUM_MUL] THEN - REPEAT STRIP_TAC THEN - UNDISCH_TAC `r - p = r2:num` THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN - - MATCH_MP_TAC REAL_LE_RSQRT THEN - REWRITE_TAC[GSYM REAL_OF_NUM_POW; REAL_POW_DIV] THEN - REWRITE_TAC[REAL_POW_2; real_div; REAL_INV_MUL] THEN - REWRITE_TAC[REAL_MUL_ASSOC] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - CONJ_TAC THENL - [ - REWRITE_TAC[REAL_ARITH `(((a * b) * a) * b) * c = (a * a) * (b * b) * c:real`] THEN - REWRITE_TAC[GSYM REAL_POW_ADD] THEN - REWRITE_TAC[ARITH_RULE `r - p + r - p = 2 * r - 2 * p`] THEN - MP_TAC (SPECL[mk_comb(amp_op_real, base_const); `2 * r`; `2 * p`] REAL_DIV_POW2) THEN - ANTS_TAC THENL [ REAL_ARITH_TAC; ALL_TAC ] THEN - ASM_SIMP_TAC[ARITH_RULE `p <= r ==> 2 * p <= 2 * r`] THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th; real_div]) THEN - - SUBGOAL_THEN `2 * r = (e1 + e1) + min_exp + e` (fun th -> REWRITE_TAC[th]) THENL - [ - EXPAND_TAC "r" THEN - REWRITE_TAC[ARITH_RULE `2 * (e1 + b + c) = (e1 + e1) + 2 * c + 2 * b`] THEN - MATCH_MP_TAC (ARITH_RULE `b1 = b2 /\ c1 = c2 ==> a + b1 + c1 = a + b2 + c2:num`) THEN - SUBGOAL_THEN `EVEN min_exp` ASSUME_TAC THENL - [ - REWRITE_TAC[min_exp_def] THEN ARITH_TAC; - ALL_TAC - ] THEN - ASM_SIMP_TAC[DIV2_EVEN_lemma]; - ALL_TAC - ] THEN - - REWRITE_TAC[REAL_POW_ADD] THEN - REWRITE_TAC[REAL_ARITH `(n * n) * (((e * e) * x * y) * z) * u = (n * e) * (n * e) * (x * u) * z * y:real`] THEN - SUBGOAL_THEN `~(&(num_exp 1 min_exp) = &0)` MP_TAC THENL - [ - REWRITE_TAC[REAL_OF_NUM_EQ; NUM_EXP_EQ_0] THEN ARITH_TAC; - ALL_TAC - ] THEN - REWRITE_TAC[num_exp; REAL_MUL_LID; GSYM REAL_OF_NUM_POW; GSYM REAL_OF_NUM_MUL] THEN - DISCH_THEN (fun th -> SIMP_TAC[th; REAL_MUL_RINV; REAL_MUL_LID]) THEN - FIRST_X_ASSUM (MP_TAC o check(fun th -> (fst o dest_var o lhand o concl) th = "f1")) THEN - REWRITE_TAC[GSYM REAL_OF_NUM_EQ; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW] THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN - - REWRITE_TAC[REAL_MUL_ASSOC] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - CONJ_TAC THENL - [ - SUBGOAL_THEN `!x y z. &0 < x /\ y <= z * x ==> y * inv x <= z` MP_TAC THENL - [ - REPEAT STRIP_TAC THEN - MP_TAC (SPECL [`y * inv x`; `z:real`; `x:real`] REAL_LE_RMUL_EQ) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th; GSYM REAL_MUL_ASSOC]) THEN - ASM_SIMP_TAC[REAL_ARITH `&0 < x ==> ~(x = &0)`; REAL_MUL_LINV; REAL_MUL_RID]; - ALL_TAC - ] THEN - - DISCH_THEN (MP_TAC o SPECL[`&(num_exp 1 (2 * p))`; `&(f1 * f1)`; `&m`]) THEN - REWRITE_TAC[num_exp; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW; REAL_MUL_LID] THEN - DISCH_THEN MATCH_MP_TAC THEN - ASM_REWRITE_TAC[REAL_OF_NUM_MUL; REAL_OF_NUM_POW; REAL_OF_NUM_LE] THEN - REWRITE_TAC[REAL_OF_NUM_LT; EXP_LT_0] THEN - ARITH_TAC; - ALL_TAC - ] THEN - - MATCH_MP_TAC REAL_POW_LE THEN - ARITH_TAC; - ALL_TAC - ] THEN - - MATCH_MP_TAC REAL_LE_INV THEN - MATCH_MP_TAC REAL_POW_LE THEN - ARITH_TAC);; - - -let FLOAT_SQRT_EVEN_hi = prove(`f1 * f1 = f2 /\ x <= f2 /\ - num_exp m (2 * p) = x /\ f1 = num_exp n1 e1 - /\ EVEN e /\ e DIV 2 = e2 /\ - e1 + e2 + (min_exp DIV 2) = r /\ - p <= r /\ r - p = r2 - ==> sqrt (float_num F m e) <= float_num F n1 r2`, - STRIP_TAC THEN - UNDISCH_TAC `x <= f2:num` THEN - UNDISCH_TAC `num_exp m (2 * p) = x` THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN - UNDISCH_TAC `f1 * f1 = f2:num` THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN - UNDISCH_TAC `e1 + e2 + min_exp DIV 2 = r` THEN - UNDISCH_TAC `e DIV 2 = e2` THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN - REPEAT (POP_ASSUM MP_TAC) THEN - REWRITE_TAC[num_exp; float; REAL_MUL_LID; GSYM REAL_OF_NUM_MUL] THEN - REPEAT STRIP_TAC THEN - UNDISCH_TAC `r - p = r2:num` THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN - - MATCH_MP_TAC REAL_LE_LSQRT THEN - REPEAT CONJ_TAC THENL - [ - REWRITE_TAC[REAL_OF_NUM_MUL; real_div] THEN - MATCH_MP_TAC REAL_LE_MUL THEN - REWRITE_TAC[REAL_POS] THEN - MATCH_MP_TAC REAL_LE_INV THEN - REWRITE_TAC[REAL_POS]; - ALL_TAC - ] THEN - - REWRITE_TAC[GSYM REAL_OF_NUM_POW; REAL_POW_DIV] THEN - REWRITE_TAC[REAL_POW_2; real_div; REAL_INV_MUL] THEN - REWRITE_TAC[REAL_MUL_ASSOC] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - CONJ_TAC THENL - [ - REWRITE_TAC[REAL_ARITH `(((a * b) * a) * b) * c = (a * a) * (b * b) * c:real`] THEN - REWRITE_TAC[GSYM REAL_POW_ADD] THEN - REWRITE_TAC[ARITH_RULE `r - p + r - p = 2 * r - 2 * p`] THEN - MP_TAC (SPECL[mk_comb(amp_op_real, base_const); `2 * r`; `2 * p`] REAL_DIV_POW2) THEN - ANTS_TAC THENL [ REAL_ARITH_TAC; ALL_TAC ] THEN - ASM_SIMP_TAC[ARITH_RULE `p <= r ==> 2 * p <= 2 * r`] THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th; real_div]) THEN - - SUBGOAL_THEN `2 * r = (e1 + e1) + min_exp + e` (fun th -> REWRITE_TAC[th]) THENL - [ - EXPAND_TAC "r" THEN - REWRITE_TAC[ARITH_RULE `2 * (e1 + b + c) = (e1 + e1) + 2 * c + 2 * b`] THEN - MATCH_MP_TAC (ARITH_RULE `b1 = b2 /\ c1 = c2 ==> a + b1 + c1 = a + b2 + c2:num`) THEN - SUBGOAL_THEN `EVEN min_exp` ASSUME_TAC THENL - [ - REWRITE_TAC[min_exp_def] THEN ARITH_TAC; - ALL_TAC - ] THEN - ASM_SIMP_TAC[DIV2_EVEN_lemma]; - ALL_TAC - ] THEN - - REWRITE_TAC[REAL_POW_ADD] THEN - REWRITE_TAC[REAL_ARITH `(n * n) * (((e * e) * x * y) * z) * u = (n * e) * (n * e) * (x * u) * z * y:real`] THEN - SUBGOAL_THEN `~(&(num_exp 1 min_exp) = &0)` MP_TAC THENL - [ - REWRITE_TAC[REAL_OF_NUM_EQ; NUM_EXP_EQ_0] THEN ARITH_TAC; - ALL_TAC - ] THEN - REWRITE_TAC[num_exp; REAL_MUL_LID; GSYM REAL_OF_NUM_POW; GSYM REAL_OF_NUM_MUL] THEN - DISCH_THEN (fun th -> SIMP_TAC[th; REAL_MUL_RINV; REAL_MUL_LID]) THEN - FIRST_X_ASSUM (MP_TAC o check(fun th -> (fst o dest_var o lhand o concl) th = "f1")) THEN - REWRITE_TAC[GSYM REAL_OF_NUM_EQ; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW] THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN - - REWRITE_TAC[REAL_MUL_ASSOC] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - CONJ_TAC THENL - [ - SUBGOAL_THEN `!x y z. &0 < x /\ z * x <= y ==> z <= y * inv x` MP_TAC THENL - [ - REPEAT STRIP_TAC THEN - MP_TAC (SPECL [`z:real`; `y * inv x`; `x:real`] REAL_LE_RMUL_EQ) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN (fun th -> REWRITE_TAC[SYM th; GSYM REAL_MUL_ASSOC]) THEN - ASM_SIMP_TAC[REAL_ARITH `&0 < x ==> ~(x = &0)`; REAL_MUL_LINV; REAL_MUL_RID]; - ALL_TAC - ] THEN - - DISCH_THEN (MP_TAC o SPECL[`&(num_exp 1 (2 * p))`; `&(f1 * f1)`; `&m`]) THEN - REWRITE_TAC[num_exp; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW; REAL_MUL_LID] THEN - DISCH_THEN MATCH_MP_TAC THEN - ASM_REWRITE_TAC[REAL_OF_NUM_MUL; REAL_OF_NUM_POW; REAL_OF_NUM_LE] THEN - REWRITE_TAC[REAL_OF_NUM_LT; EXP_LT_0] THEN - ARITH_TAC; - ALL_TAC - ] THEN - - MATCH_MP_TAC REAL_POW_LE THEN - ARITH_TAC; - ALL_TAC - ] THEN - - MATCH_MP_TAC REAL_LE_INV THEN - MATCH_MP_TAC REAL_POW_LE THEN - ARITH_TAC);; - - -(******************) -let transform = UNDISCH_ALL o - PURE_ONCE_REWRITE_RULE[TAUT `EVEN e <=> (EVEN e <=> T)`] o - NUMERALS_TO_NUM o - CONV_RULE (DEPTH_CONV NUM_DIV_CONV) o - REWRITE_RULE[GSYM IMP_IMP; min_exp_def];; - -let FLOAT_SQRT_EVEN_lo' = transform FLOAT_SQRT_EVEN_lo and - FLOAT_SQRT_EVEN_hi' = transform FLOAT_SQRT_EVEN_hi and - FLOAT_PRE_EXP' = (UNDISCH_ALL o - PURE_ONCE_REWRITE_RULE[TAUT `~(e = _0) <=> ((e = _0) <=> F)`] o - REWRITE_RULE[GSYM IMP_IMP; NUMERAL]) FLOAT_PRE_EXP;; - -let even_const = `EVEN` and - pre_const = `PRE` and - two_num = rand(mk_small_numeral_array 2) and - min_exp_div2 = rand(mk_small_numeral_array (min_exp / 2)) and - f2_var_num = `f2:num` and - f1_var_num = `f1:num` and - p_var_num = `p:num`;; - -(* Returns the list of digits of the given Big_int n in the base b *) -let rec get_big_int_digits b n = - let bb = big_int_of_int b in - if le_big_int n zero_big_int then [] - else - let q, r = quomod_big_int n bb in - r :: get_big_int_digits b q;; - -(* [1;2;3] -> 123 (base = 10) *) -let rec big_int_from_list b list = - let rec proc acc list = - match list with - [] -> acc - | h::t -> proc (add_big_int h (mult_int_big_int b acc)) t in - proc zero_big_int list;; - -(* Returns n first elements of the list *) -let rec take n list = - match list with - x :: xs -> if n > 0 then x :: take (n - 1) xs else [] - | [] -> [];; - -(* Returns an integer number that contains at most pp significant digits - in the given base b *) -let big_int_round_lo base pp n = - let digits = rev (get_big_int_digits base n) in - let n_digits = length digits in - if n_digits <= pp then - n - else - let m = big_int_from_list base (take pp digits) in - mult_big_int (power_int_positive_int base (n_digits - pp)) m;; - -let big_int_round_hi base pp n = - let digits = rev (get_big_int_digits base n) in - let n_digits = length digits in - if n_digits <= pp then n - else - let l1, l2 = chop_list pp digits in - if forall (eq_big_int zero_big_int) l2 then n - else - let m = succ_big_int (big_int_from_list base l1) in - mult_big_int (power_int_positive_int base (n_digits - pp)) m;; - - -(******************) -let rec float_sqrt_lo pp tm = - let s, m_tm, e_tm = dest_float tm in - let p_tm = rand (mk_small_numeral_array pp) in - if s <> "F" then - failwith "float_sqrt_lo: negative argument" - else - let even_th = raw_even_hash_conv (mk_comb (even_const, e_tm)) in - if (fst o dest_const o rand o concl) even_th <> "T" then - (* ODD e *) - let pre_e = raw_pre_hash_conv (mk_comb (pre_const, e_tm)) in - let e_neq_0 = raw_eq0_hash_conv e_tm in - let e1_tm = rand (concl pre_e) in - let th0 = INST[e1_tm, e1_var_num; e_tm, e_var_num; m_tm, m_var_num] FLOAT_PRE_EXP' in - let th1 = MY_PROVE_HYP pre_e (MY_PROVE_HYP e_neq_0 th0) in - let th2 = float_sqrt_lo pp (rand(concl th1)) in - let ltm, rtm = dest_comb (concl th2) in - EQ_MP (SYM (AP_TERM ltm (AP_TERM (rator rtm) th1))) th2 - else - (* EVEN e *) - let p2_tm = mk_binop mul_op_num two_num p_tm in - let p2_th = raw_mul_conv_hash p2_tm in - let f1_1 = AP_TERM (mk_comb(num_exp_const, m_tm)) p2_th in - let f1_2 = TRANS f1_1 (denormalize (rand (concl f1_1))) in - - let x_tm = rand(concl f1_2) in - let x = raw_dest_hash x_tm in - let f1' = Big_int.sqrt_big_int (big_int_of_num x) in - let f1 = num_of_big_int (big_int_round_lo Arith_num.arith_base pp f1') in - let f1_tm = rand(mk_numeral_array f1) in - let f1_num_exp = to_num_exp f1_tm in - - let n1_tm, e1_tm = dest_num_exp (rand (concl f1_num_exp)) in - let f1f1_eq_f2 = raw_mul_conv_hash (mk_binop mul_op_num f1_tm f1_tm) in - let f2_tm = rand(concl f1f1_eq_f2) in - let f2_le_x = EQT_ELIM (raw_le_hash_conv (mk_binop le_op_num f2_tm x_tm)) in - - let e_div2_eq_e2 = raw_div_hash_conv (mk_binop div_op_num e_tm two_num) in - let e2_tm = rand(concl e_div2_eq_e2) in - let r_th1 = raw_add_conv_hash (mk_binop add_op_num e2_tm min_exp_div2) in - let r_th2 = AP_TERM (mk_comb(add_op_num, e1_tm)) r_th1 in - let r_th = TRANS r_th2 (raw_add_conv_hash (rand (concl r_th2))) in - - let r_tm = rand(concl r_th) in - let r_sub_p, p_le_r = raw_sub_and_le_hash_conv p_tm r_tm in - let r2_tm = rand(concl r_sub_p) in - if (rand(concl p_le_r) <> r_tm) then - failwith "float_sqrt_lo: underflow" - else - let th0 = INST[f2_tm, f2_var_num; x_tm, x_var_num; p_tm, p_var_num; r_tm, r_var_num; - f1_tm, f1_var_num; n1_tm, n1_var_num; e1_tm, e1_var_num; e2_tm, e2_var_num; - e_tm, e_var_num; m_tm, m_var_num; r2_tm, r2_var_num] - FLOAT_SQRT_EVEN_lo' in - MY_PROVE_HYP f1_2 ( - MY_PROVE_HYP e_div2_eq_e2 ( - MY_PROVE_HYP r_sub_p ( - MY_PROVE_HYP r_th ( - MY_PROVE_HYP f1f1_eq_f2 ( - MY_PROVE_HYP f1_num_exp ( - MY_PROVE_HYP even_th ( - MY_PROVE_HYP f2_le_x ( - MY_PROVE_HYP p_le_r th0 - ))))))));; - - -let rec float_sqrt_hi pp tm = - let s, m_tm, e_tm = dest_float tm in - let p_tm = rand (mk_small_numeral_array pp) in - if s <> "F" then - failwith "float_sqrt_lo: negative argument" - else - let even_th = raw_even_hash_conv (mk_comb (even_const, e_tm)) in - if (fst o dest_const o rand o concl) even_th <> "T" then - (* ODD e *) - let pre_e = raw_pre_hash_conv (mk_comb (pre_const, e_tm)) in - let e_neq_0 = raw_eq0_hash_conv e_tm in - let e1_tm = rand (concl pre_e) in - let th0 = INST[e1_tm, e1_var_num; e_tm, e_var_num; m_tm, m_var_num] FLOAT_PRE_EXP' in - let th1 = MY_PROVE_HYP pre_e (MY_PROVE_HYP e_neq_0 th0) in - let th2 = float_sqrt_hi pp (rand(concl th1)) in - let ltm, rtm = dest_comb (concl th2) in - let ltm2, rtm2 = dest_comb ltm in - let th3 = AP_THM (AP_TERM ltm2 (AP_TERM (rator rtm2) th1)) rtm in - EQ_MP (SYM th3) th2 - else - (* EVEN e *) - let p2_tm = mk_binop mul_op_num two_num p_tm in - let p2_th = raw_mul_conv_hash p2_tm in - let f1_1 = AP_TERM (mk_comb(num_exp_const, m_tm)) p2_th in - let f1_2 = TRANS f1_1 (denormalize (rand (concl f1_1))) in - - let x_tm = rand(concl f1_2) in - let x = raw_dest_hash x_tm in - let x' = big_int_of_num x in - let f1' = sqrt_big_int x' in - let f1 = (num_of_big_int o big_int_round_hi Arith_num.arith_base pp) - (if eq_big_int (mult_big_int f1' f1') x' then f1' else succ_big_int f1') in - - let f1_tm = rand(mk_numeral_array f1) in - let f1_num_exp = to_num_exp f1_tm in - - let n1_tm, e1_tm = dest_num_exp (rand (concl f1_num_exp)) in - let f1f1_eq_f2 = raw_mul_conv_hash (mk_binop mul_op_num f1_tm f1_tm) in - let f2_tm = rand(concl f1f1_eq_f2) in - let x_le_f2 = EQT_ELIM (raw_le_hash_conv (mk_binop le_op_num x_tm f2_tm)) in - - let e_div2_eq_e2 = raw_div_hash_conv (mk_binop div_op_num e_tm two_num) in - let e2_tm = rand(concl e_div2_eq_e2) in - let r_th1 = raw_add_conv_hash (mk_binop add_op_num e2_tm min_exp_div2) in - let r_th2 = AP_TERM (mk_comb(add_op_num, e1_tm)) r_th1 in - let r_th = TRANS r_th2 (raw_add_conv_hash (rand (concl r_th2))) in - - let r_tm = rand(concl r_th) in - let r_sub_p, p_le_r = raw_sub_and_le_hash_conv p_tm r_tm in - let r2_tm = rand(concl r_sub_p) in - if (rand(concl p_le_r) <> r_tm) then - failwith "float_sqrt_lo: underflow" - else - let th0 = INST[f2_tm, f2_var_num; x_tm, x_var_num; p_tm, p_var_num; r_tm, r_var_num; - f1_tm, f1_var_num; n1_tm, n1_var_num; e1_tm, e1_var_num; e2_tm, e2_var_num; - e_tm, e_var_num; m_tm, m_var_num; r2_tm, r2_var_num] - FLOAT_SQRT_EVEN_hi' in - MY_PROVE_HYP f1_2 ( - MY_PROVE_HYP e_div2_eq_e2 ( - MY_PROVE_HYP r_sub_p ( - MY_PROVE_HYP r_th ( - MY_PROVE_HYP f1f1_eq_f2 ( - MY_PROVE_HYP f1_num_exp ( - MY_PROVE_HYP even_th ( - MY_PROVE_HYP x_le_f2 ( - MY_PROVE_HYP p_le_r th0 - ))))))));; - -end;; (* Float_ops module *) - -(************************************) -(* Cached floating point operations *) -(************************************) - -let init_logs () = - if !Arith_options.float_log then - begin - Log.open_log "mul_lo"; - Log.open_log "mul_hi"; - Log.open_log "div_lo"; - Log.open_log "div_hi"; - Log.open_log "add_lo"; - Log.open_log "add_hi"; - Log.open_log "sub_lo"; - Log.open_log "sub_hi" - end;; - -let float_to_log_string tm = - let s, n_tm, e_tm = dest_float tm in - let n = (if s = "T" then minus_num else I) (raw_dest_hash n_tm) in - let e = raw_dest_hash e_tm -/ num Float_theory.min_exp in - string_of_num n ^ "," ^ string_of_num e;; - -let log_binop name pp tm1 tm2 = - let str1 = float_to_log_string tm1 and - str2 = float_to_log_string tm2 in - let str = string_of_int pp ^ ":" ^ str1 ^ ";" ^ str2 in - Log.append_to_log name str;; - -(* Counters for collecting stats *) -let lt0_c = ref 0 and - gt0_c = ref 0 and - lt_c = ref 0 and - le0_c = ref 0 and - ge0_c = ref 0 and - le_c = ref 0 and - min_c = ref 0 and - max_c = ref 0 and - min_max_c = ref 0 and - mul_lo_c = ref 0 and - mul_hi_c = ref 0 and - div_lo_c = ref 0 and - div_hi_c = ref 0 and - add_lo_c = ref 0 and - add_hi_c = ref 0 and - sub_lo_c = ref 0 and - sub_hi_c = ref 0 and - sqrt_lo_c = ref 0 and - sqrt_hi_c = ref 0;; - -(* Hash tables *) -let cache_size = if !Arith_options.float_cached then !Arith_options.init_cache_size else 1;; - -let my_add h key v = - if Hashtbl.length h >= !Arith_options.max_cache_size then - Hashtbl.clear h -(* let _ = Hashtbl.clear h in - print_string "Clearing a float hash table" *) - else - (); - Hashtbl.add h key v;; - -let mul_table = Hashtbl.create cache_size and - div_table = Hashtbl.create cache_size and - add_table = Hashtbl.create cache_size and - sub_table = Hashtbl.create cache_size and - sqrt_table = Hashtbl.create cache_size and - le_table = Hashtbl.create cache_size and - max_table = Hashtbl.create cache_size;; - -let reset_cache () = - Hashtbl.clear mul_table; - Hashtbl.clear div_table; - Hashtbl.clear add_table; - Hashtbl.clear sub_table; - Hashtbl.clear sqrt_table; - Hashtbl.clear le_table; - Hashtbl.clear max_table;; - -let reset_stat () = - lt0_c := 0; - gt0_c := 0; - lt_c := 0; - le0_c := 0; - ge0_c := 0; - le_c := 0; - min_c := 0; - max_c := 0; - min_max_c := 0; - mul_lo_c := 0; - mul_hi_c := 0; - div_lo_c := 0; - div_hi_c := 0; - add_lo_c := 0; - add_hi_c := 0; - sub_lo_c := 0; - sub_hi_c := 0; - sqrt_lo_c := 0; - sqrt_hi_c := 0;; - -let print_stat () = - let len = Hashtbl.length in - let cmp_str1 = sprintf "lt0 = %d\ngt0 = %d\nlt = %d\n" !lt0_c !gt0_c !lt_c and - cmp_str2 = sprintf "le0 = %d\nge0 = %d\n" !le0_c !ge0_c and - cmp_str3 = sprintf "min = %d\nmin_max = %d\n" !min_c !min_max_c and - le_str = sprintf "le = %d (le_hash = %d)\n" !le_c (len le_table) and - max_str = sprintf "max = %d (max_hash = %d)\n" !max_c (len max_table) and - mul_str = sprintf "mul_lo = %d, mul_hi = %d (mul_hash = %d)\n" !mul_lo_c !mul_hi_c (len mul_table) and - div_str = sprintf "div_lo = %d, div_hi = %d (div_hash = %d)\n" !div_lo_c !div_hi_c (len div_table) and - add_str = sprintf "add_lo = %d, add_hi = %d (add_hash = %d)\n" !add_lo_c !add_hi_c (len add_table) and - sub_str = sprintf "sub_lo = %d, sub_hi = %d (sub_hash = %d)\n" !sub_lo_c !sub_hi_c (len sub_table) and - sqrt_str = sprintf "sqrt_lo = %d, sqrt_hi = %d (sqrt_hash = %d)\n" !sqrt_lo_c !sqrt_hi_c (len sqrt_table) in - print_string (cmp_str1 ^ cmp_str2 ^ cmp_str3 ^ - le_str ^ max_str ^ mul_str ^ div_str ^ add_str ^ sub_str ^ sqrt_str);; - - -(* lt0 *) -let float_lt0 = - if !Arith_options.float_cached then - fun tm -> - let _ = lt0_c := !lt0_c + 1 in - Float_ops.float_lt0 tm - else - Float_ops.float_lt0;; - -(* gt0 *) -let float_gt0 = - if !Arith_options.float_cached then - fun tm -> - let _ = gt0_c := !gt0_c + 1 in - Float_ops.float_gt0 tm - else - Float_ops.float_gt0;; - -(* lt *) -let float_lt = - if !Arith_options.float_cached then - fun tm1 tm2 -> - let _ = lt_c := !lt_c + 1 in - Float_ops.float_lt tm1 tm2 - else - Float_ops.float_lt;; - -(* le0 *) -let float_le0 = - if !Arith_options.float_cached then - fun tm -> - let _ = le0_c := !le0_c + 1 in - Float_ops.float_le0 tm - else - Float_ops.float_le0;; - -(* ge0 *) -let float_ge0 = - if !Arith_options.float_cached then - fun tm -> - let _ = ge0_c := !ge0_c + 1 in - Float_ops.float_ge0 tm - else - Float_ops.float_ge0;; - -(* min *) -let float_min = - if !Arith_options.float_cached then - fun tm1 tm2 -> - let _ = min_c := !min_c + 1 in - Float_ops.float_min tm1 tm2 - else - Float_ops.float_min;; - -(* min_max *) -let float_min_max = - if !Arith_options.float_cached then - fun tm1 tm2 -> - let _ = min_max_c := !min_max_c + 1 in - Float_ops.float_min_max tm1 tm2 - else - Float_ops.float_min_max;; - - -(***************) -let float_hash tm = - let s, n_tm, e_tm = dest_float tm in - s ^ (Arith_cache.num_tm_hash n_tm) ^ "e" ^ (Arith_cache.num_tm_hash e_tm);; - -let float_op_hash pp tm1 tm2 = - string_of_int pp ^ float_hash tm1 ^ "x" ^ float_hash tm2;; - -let float_op_hash1 pp tm = - string_of_int pp ^ float_hash tm;; - - -(* le *) -let float_le = - if !Arith_options.float_cached then - fun tm1 tm2 -> - let _ = le_c := !le_c + 1 in - let hash = float_op_hash 0 tm1 tm2 in - try - Hashtbl.find le_table hash - with Not_found -> - let result = Float_ops.float_le tm1 tm2 in - let _ = my_add le_table hash result in - result - else - Float_ops.float_le;; - -(* max *) -let float_max = - if !Arith_options.float_cached then - fun tm1 tm2 -> - let _ = max_c := !max_c + 1 in - let hash = float_op_hash 0 tm1 tm2 in - try - Hashtbl.find max_table hash - with Not_found -> - let result = Float_ops.float_max tm1 tm2 in - let _ = my_add max_table hash result in - result - else - Float_ops.float_max;; - -(* mul_eq *) -let float_mul_eq = Float_ops.float_mul_eq;; - -(* mul_lo *) -let float_mul_lo = - if !Arith_options.float_cached then - fun pp tm1 tm2 -> - let _ = mul_lo_c := !mul_lo_c + 1 in - let hash = "lo" ^ float_op_hash pp tm1 tm2 in - try - Hashtbl.find mul_table hash - with Not_found -> - let result = Float_ops.float_mul_lo pp tm1 tm2 in - let _ = my_add mul_table hash result in - let _ = if !Arith_options.float_log then log_binop "mul_lo" pp tm1 tm2 in - result - else - if !Arith_options.float_log then - fun pp tm1 tm2 -> - let _ = log_binop "mul_lo" pp tm1 tm2 in - Float_ops.float_mul_lo pp tm1 tm2 - else - Float_ops.float_mul_lo;; - -(* mul_hi *) -let float_mul_hi = - if !Arith_options.float_cached then - fun pp tm1 tm2 -> - let _ = mul_hi_c := !mul_hi_c + 1 in - let hash = "hi" ^ float_op_hash pp tm1 tm2 in - try - Hashtbl.find mul_table hash - with Not_found -> - let result = Float_ops.float_mul_hi pp tm1 tm2 in - let _ = my_add mul_table hash result in - let _ = if !Arith_options.float_log then log_binop "mul_hi" pp tm1 tm2 in - result - else - if !Arith_options.float_log then - fun pp tm1 tm2 -> - let _ = log_binop "mul_hi" pp tm1 tm2 in - Float_ops.float_mul_hi pp tm1 tm2 - else - Float_ops.float_mul_hi;; - -(* div_lo *) -let float_div_lo = - if !Arith_options.float_cached then - fun pp tm1 tm2 -> - let _ = div_lo_c := !div_lo_c + 1 in - let hash = "lo" ^ float_op_hash pp tm1 tm2 in - try - Hashtbl.find div_table hash - with Not_found -> - let result = Float_ops.float_div_lo pp tm1 tm2 in - let _ = my_add div_table hash result in - let _ = if !Arith_options.float_log then log_binop "div_lo" pp tm1 tm2 in - result - else - if !Arith_options.float_log then - fun pp tm1 tm2 -> - let _ = log_binop "div_lo" pp tm1 tm2 in - Float_ops.float_div_lo pp tm1 tm2 - else - Float_ops.float_div_lo;; - -(* div_hi *) -let float_div_hi = - if !Arith_options.float_cached then - fun pp tm1 tm2 -> - let _ = div_hi_c := !div_hi_c + 1 in - let hash = "hi" ^ float_op_hash pp tm1 tm2 in - try - Hashtbl.find div_table hash - with Not_found -> - let result = Float_ops.float_div_hi pp tm1 tm2 in - let _ = my_add div_table hash result in - let _ = if !Arith_options.float_log then log_binop "div_hi" pp tm1 tm2 in - result - else - if !Arith_options.float_log then - fun pp tm1 tm2 -> - let _ = log_binop "div_hi" pp tm1 tm2 in - Float_ops.float_div_hi pp tm1 tm2 - else - Float_ops.float_div_hi;; - -(* add_lo *) -let float_add_lo = - if !Arith_options.float_cached then - fun pp tm1 tm2 -> - let _ = add_lo_c := !add_lo_c + 1 in - let hash = "lo" ^ float_op_hash pp tm1 tm2 in - try - Hashtbl.find add_table hash - with Not_found -> - let result = Float_ops.float_add_lo pp tm1 tm2 in - let _ = my_add add_table hash result in - let _ = if !Arith_options.float_log then log_binop "add_lo" pp tm1 tm2 in - result - else - if !Arith_options.float_log then - fun pp tm1 tm2 -> - let _ = log_binop "add_lo" pp tm1 tm2 in - Float_ops.float_add_lo pp tm1 tm2 - else - Float_ops.float_add_lo;; - -(* add_hi *) -let float_add_hi = - if !Arith_options.float_cached then - fun pp tm1 tm2 -> - let _ = add_hi_c := !add_hi_c + 1 in - let hash = "hi" ^ float_op_hash pp tm1 tm2 in - try - Hashtbl.find add_table hash - with Not_found -> - let result = Float_ops.float_add_hi pp tm1 tm2 in - let _ = my_add add_table hash result in - let _ = if !Arith_options.float_log then log_binop "add_hi" pp tm1 tm2 in - result - else - if !Arith_options.float_log then - fun pp tm1 tm2 -> - let _ = log_binop "add_hi" pp tm1 tm2 in - Float_ops.float_add_hi pp tm1 tm2 - else - Float_ops.float_add_hi;; - -(* sub_lo *) -let float_sub_lo = - if !Arith_options.float_cached then - fun pp tm1 tm2 -> - let _ = sub_lo_c := !sub_lo_c + 1 in - let hash = "lo" ^ float_op_hash pp tm1 tm2 in - try - Hashtbl.find sub_table hash - with Not_found -> - let result = Float_ops.float_sub_lo pp tm1 tm2 in - let _ = my_add sub_table hash result in - let _ = if !Arith_options.float_log then log_binop "sub_lo" pp tm1 tm2 in - result - else - if !Arith_options.float_log then - fun pp tm1 tm2 -> - let _ = log_binop "sub_lo" pp tm1 tm2 in - Float_ops.float_sub_lo pp tm1 tm2 - else - Float_ops.float_sub_lo;; - -(* sub_hi *) -let float_sub_hi = - if !Arith_options.float_cached then - fun pp tm1 tm2 -> - let _ = sub_hi_c := !sub_hi_c + 1 in - let hash = "hi" ^ float_op_hash pp tm1 tm2 in - try - Hashtbl.find sub_table hash - with Not_found -> - let result = Float_ops.float_sub_hi pp tm1 tm2 in - let _ = my_add sub_table hash result in - let _ = if !Arith_options.float_log then log_binop "sub_hi" pp tm1 tm2 in - result - else - if !Arith_options.float_log then - fun pp tm1 tm2 -> - let _ = log_binop "sub_hi" pp tm1 tm2 in - Float_ops.float_sub_hi pp tm1 tm2 - else - Float_ops.float_sub_hi;; - -(* sqrt_lo *) -let float_sqrt_lo = - if !Arith_options.float_cached then - fun pp tm -> - let _ = sqrt_lo_c := !sqrt_lo_c + 1 in - let hash = "lo" ^ float_op_hash1 pp tm in - try - Hashtbl.find sqrt_table hash - with Not_found -> - let result = Float_ops.float_sqrt_lo pp tm in - let _ = my_add sqrt_table hash result in - result - else - Float_ops.float_sqrt_lo;; - -(* sqrt_hi *) -let float_sqrt_hi = - if !Arith_options.float_cached then - fun pp tm -> - let _ = sqrt_hi_c := !sqrt_hi_c + 1 in - let hash = "hi" ^ float_op_hash1 pp tm in - try - Hashtbl.find sqrt_table hash - with Not_found -> - let result = Float_ops.float_sqrt_hi pp tm in - let _ = my_add sqrt_table hash result in - result - else - Float_ops.float_sqrt_hi;; - - - -(* ------------------------------------------------ *) -(* Special comparison operations *) -(* ------------------------------------------------ *) - -(* For a given floating-point term returns its components as OCaml objects *) -let dest_float_raw f_tm = - let s, n_tm, e_tm = dest_float f_tm in - let sign = if s = "T" then true else false in - sign, raw_dest_hash n_tm, raw_dest_hash e_tm;; - -(* Compares raw representations of two floating-point numbers f1 and f2. *) -(* Returns -1 if f1 < f2, 0 if f1 = f2, and 1 if f1 > f2. *) -let compare_floats_raw = - let denorm n e = n */ (num Arith_num.arith_base **/ e) in - let cmp n1 e1 n2 e2 = - if e1 <= e2 then - compare_num n1 (denorm n2 (e2 -/ e1)) - else - compare_num (denorm n1 (e1 -/ e2)) n2 in - fun f1_tm f2_tm -> - let s1, n1, e1 = dest_float_raw f1_tm and - s2, n2, e2 = dest_float_raw f2_tm in - let n1' = if s1 then minus_num n1 else n1 and - n2' = if s2 then minus_num n2 else n2 in - cmp n1' e1 n2' e2;; - -(* Returns (true, |- f1 <= f2 <=> T) when f1 <= f2; returns (false, |- f1 <= f2 <=> F) otherwise *) -let float_prove_le f1_tm f2_tm = - let le_th = float_le f1_tm f2_tm in - let flag = (fst o dest_const o rand o concl) le_th = "T" in - flag, le_th;; - -(* Returns (true, |- f1 < f2 <=> T) when f1 < f2; returns (false, |- f1 < f2 <=> F) otherwise *) -let float_prove_lt f1_tm f2_tm = - let lt_th = float_lt f1_tm f2_tm in - let flag = (fst o dest_const o rand o concl) lt_th = "T" in - flag, lt_th;; - -(* -(* Arith_num.raw_dest_hash is too slow (need to retest without caching) *) -let float_prove_le f1_tm f2_tm = - let r = compare_floats_raw f1_tm f2_tm in - if r <= 0 then - true, float_le f1_tm f2_tm - else - false, TRUTH;; -*) - -(* Compares y and an interval a <= x <= b (|- interval_arith x (a, b)). *) -(* Returns (true, |- y <= x) if y <= a. *) -(* Returns (false, |- y <= a <=> F) if y > a. *) -let float_prove_le_interval = - let le_interval = (UNDISCH_ALL o prove)(`interval_arith x (a,b) ==> (y <= a <=> T) ==> y <= x`, - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC) in - fun y_tm x_th -> - let x_tm, bounds = dest_interval_arith (concl x_th) in - let ltm, b_tm = dest_comb bounds in - let a_tm = rand ltm in - let flag, le_th = float_prove_le y_tm a_tm in - if flag then - let th0 = INST[x_tm, x_var_real; y_tm, y_var_real; - a_tm, a_var_real; b_tm, b_var_real] le_interval in - true, MY_PROVE_HYP x_th (MY_PROVE_HYP le_th th0) - else - false, le_th;; - -(* Compares y and an interval a <= x <= b (|- interval_arith x (a, b)). *) -(* Returns (true, |- x <= y) if b <= y. *) -(* Returns (false, |- b <= y <=> F) if b > y. *) -let float_prove_ge_interval = - let ge_interval = (UNDISCH_ALL o prove)(`interval_arith x (a,b) ==> (b <= y <=> T) ==> x <= y`, - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC) in - fun y_tm x_th -> - let x_tm, bounds = dest_interval_arith (concl x_th) in - let ltm, b_tm = dest_comb bounds in - let a_tm = rand ltm in - let flag, le_th = float_prove_le b_tm y_tm in - if flag then - let th0 = INST[x_tm, x_var_real; y_tm, y_var_real; - a_tm, a_var_real; b_tm, b_var_real] ge_interval in - true, MY_PROVE_HYP x_th (MY_PROVE_HYP le_th th0) - else - false, le_th;; - -(* Compares y and an interval a <= x <= b (|- interval_arith x (a, b)). *) -(* Returns (true, |- y < x) if y < a. *) -(* Returns (false, |- y < a <=> F) if y >= a. *) -let float_prove_lt_interval = - let lt_interval = (UNDISCH_ALL o prove)(`interval_arith x (a,b) ==> (y < a <=> T) ==> y < x`, - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC) in - fun y_tm x_th -> - let x_tm, bounds = dest_interval_arith (concl x_th) in - let ltm, b_tm = dest_comb bounds in - let a_tm = rand ltm in - let flag, lt_th = float_prove_lt y_tm a_tm in - if flag then - let th0 = INST[x_tm, x_var_real; y_tm, y_var_real; - a_tm, a_var_real; b_tm, b_var_real] lt_interval in - true, MY_PROVE_HYP x_th (MY_PROVE_HYP lt_th th0) - else - false, lt_th;; - -(* Compares y and an interval a <= x <= b (|- interval_arith x (a, b)). *) -(* Returns (true, |- x < y) if b < y. *) -(* Returns (false, |- b < y <=> F) if b >= y. *) -let float_prove_gt_interval = - let gt_interval = (UNDISCH_ALL o prove)(`interval_arith x (a,b) ==> (b < y <=> T) ==> x < y`, - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC) in - fun y_tm x_th -> - let x_tm, bounds = dest_interval_arith (concl x_th) in - let ltm, b_tm = dest_comb bounds in - let a_tm = rand ltm in - let flag, lt_th = float_prove_lt b_tm y_tm in - if flag then - let th0 = INST[x_tm, x_var_real; y_tm, y_var_real; - a_tm, a_var_real; b_tm, b_var_real] gt_interval in - true, MY_PROVE_HYP x_th (MY_PROVE_HYP lt_th th0) - else - false, lt_th;; - -(* Compares y and an interval a <= x <= b (|- interval_arith x (a, b)). *) -(* Returns (-1, |- y <= x) if y <= a. *) -(* Returns (1, |- x <= y) if b <= y. *) -(* Returns (0, |- T ) otherwise. *) -let float_compare_interval = - let le_interval = (UNDISCH_ALL o prove)(`interval_arith x (a,b) ==> (y <= a <=> T) ==> y <= x`, - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC) in - let ge_interval = (UNDISCH_ALL o prove)(`interval_arith x (a,b) ==> (b <= y <=> T) ==> x <= y`, - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC) in - fun y_tm x_th -> - let x_tm, bounds = dest_interval_arith (concl x_th) in - let ltm, b_tm = dest_comb bounds in - let a_tm = rand ltm in - let inst th le_th = - let th0 = INST[x_tm, x_var_real; y_tm, y_var_real; - a_tm, a_var_real; b_tm, b_var_real] th in - MY_PROVE_HYP x_th (MY_PROVE_HYP le_th th0) in - let flag, le_th = float_prove_le y_tm a_tm in - if flag then - -1, inst le_interval le_th - else - let flag, le_th = float_prove_le b_tm y_tm in - if flag then - 1, inst ge_interval le_th - else - 0, TRUTH;; - - -(* ------------------------------------------------ *) -(* Float interval operations *) -(* ------------------------------------------------ *) - -let FLOAT_OF_NUM' = (SPEC_ALL o REWRITE_RULE[min_exp_def]) FLOAT_OF_NUM;; - -let FLOAT_INTERVAL_OF_NUM = (NUMERALS_TO_NUM o REWRITE_RULE[min_exp_def] o prove)(`interval_arith (&n) (float_num F n min_exp, float_num F n min_exp)`, - REWRITE_TAC[FLOAT_OF_NUM; CONST_INTERVAL]);; - -let FLOAT_F_bound' = (UNDISCH_ALL o SPEC_ALL) FLOAT_F_bound;; - -let FLOAT_T_bound' = (UNDISCH_ALL o SPEC_ALL) FLOAT_T_bound;; - -(* interval_arith x (float_num s1 n1 e1, float_num s2 n2 e2) -> x, float_num s1 n1 e1, float_num s2 n2 e2 *) -let dest_float_interval tm = - let ltm, rtm = dest_comb tm in - let f1, f2 = dest_pair rtm in - rand ltm, f1, f2;; - -let mk_float_interval_small_num n = - let n_tm0 = mk_small_numeral n in - let n_th = NUMERAL_TO_NUM_CONV n_tm0 in - let n_tm = rand(rand(concl n_th)) in - let n_th1 = TRANS n_th (INST[n_tm, n_var_num] NUM_REMOVE) in - let th1 = AP_TERM amp_op_real n_th1 in - let int_th = INST[n_tm, n_var_num] FLOAT_INTERVAL_OF_NUM in - let rtm = rand(concl int_th) in - EQ_MP (SYM (AP_THM (AP_TERM interval_const th1) rtm)) int_th;; - -let mk_float_interval_num n = - let n_tm0 = mk_numeral n in - let n_th = NUMERAL_TO_NUM_CONV n_tm0 in - let n_tm = rand(rand(concl n_th)) in - let n_th1 = TRANS n_th (INST[n_tm, n_var_num] NUM_REMOVE) in - let th1 = AP_TERM amp_op_real n_th1 in - let int_th = INST[n_tm, n_var_num] FLOAT_INTERVAL_OF_NUM in - let rtm = rand(concl int_th) in - EQ_MP (SYM (AP_THM (AP_TERM interval_const th1) rtm)) int_th;; - - -(* Returns the lower bound for the given float *) -let float_lo p tm = - let s, n_tm, e_tm = dest_float tm in - if s = "F" then - let num_exp_tm = mk_num_exp n_tm e_tm in - let th0 = num_exp_lo p num_exp_tm in - let ltm, e1_tm = dest_comb(lhand(concl th0)) in - let n1_tm = rand ltm in - let th1 = INST[n1_tm, n1_var_num; e1_tm, e1_var_num; n_tm, n2_var_num; e_tm, e2_var_num] FLOAT_F_bound' in - MY_PROVE_HYP th0 th1 - else - let num_exp_tm = mk_num_exp n_tm e_tm in - let th0 = num_exp_hi p num_exp_tm in - let ltm, e1_tm = dest_comb(rand(concl th0)) in - let n1_tm = rand ltm in - let th1 = INST[n_tm, n1_var_num; e_tm, e1_var_num; n1_tm, n2_var_num; e1_tm, e2_var_num] FLOAT_T_bound' in - MY_PROVE_HYP th0 th1;; - -(* Returns the upper bound for the given float *) -let float_hi p tm = - let s, n_tm, e_tm = dest_float tm in - if s = "F" then - let num_exp_tm = mk_num_exp n_tm e_tm in - let th0 = num_exp_hi p num_exp_tm in - let ltm, e2_tm = dest_comb(rand(concl th0)) in - let n2_tm = rand ltm in - let th1 = INST[n_tm, n1_var_num; e_tm, e1_var_num; n2_tm, n2_var_num; e2_tm, e2_var_num] FLOAT_F_bound' in - MY_PROVE_HYP th0 th1 - else - let num_exp_tm = mk_num_exp n_tm e_tm in - let th0 = num_exp_lo p num_exp_tm in - let ltm, e1_tm = dest_comb(lhand(concl th0)) in - let n1_tm = rand ltm in - let th1 = INST[n1_tm, n1_var_num; e1_tm, e1_var_num; n_tm, n2_var_num; e_tm, e2_var_num] FLOAT_T_bound' in - MY_PROVE_HYP th0 th1;; - - -(* Approximates the given interval with p-digits floating point numbers *) -let float_interval_round p th = - let x_tm, f1, f2 = dest_float_interval (concl th) in - let lo_th = float_lo p f1 in - let hi_th = float_hi p f2 in - let lo_tm = lhand(concl lo_th) in - let hi_tm = rand(concl hi_th) in - let th0 = INST[x_tm, x_var_real; f1, lo_var_real; f2, hi_var_real; lo_tm, a_var_real; hi_tm, b_var_real] APPROX_INTERVAL' in - MY_PROVE_HYP lo_th (MY_PROVE_HYP hi_th (MY_PROVE_HYP th th0));; - - -(* ------------------------------------------------ *) -(* Absolute value of an interval *) -(* ------------------------------------------------ *) - -let RULE = UNDISCH_ALL o NUMERALS_TO_NUM o PURE_REWRITE_RULE[GSYM IMP_IMP; min_exp_def];; - -let float_interval_abs_ft = (RULE o prove) - (`interval_arith x (float_num F n1 e1, float_num T n2 e2) - ==> interval_arith (abs x) (float_num F 0 min_exp, float_num F 0 min_exp)`, - DISCH_THEN (MP_TAC o MATCH_MP FLOAT_INTERVAL_FT_IMP_0) THEN - REWRITE_TAC[FLOAT_0; interval_arith] THEN REAL_ARITH_TAC);; - -let float_interval_abs_ff = (RULE o prove) - (`interval_arith x (float_num F n1 e1, float_num F n2 e2) - ==> interval_arith (abs x) (float_num F n1 e1, float_num F n2 e2)`, - MP_TAC (SPECL[`n1:num`; `e1:num`] FLOAT_F_POS) THEN - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; - -let float_interval_abs_tt = (RULE o prove) - (`interval_arith x (float_num T n1 e1, float_num T n2 e2) - ==> interval_arith (abs x) (float_num F n2 e2, float_num F n1 e1)`, - MP_TAC ((GSYM o SPEC `T`) FLOAT_NEG) THEN SIMP_TAC[] THEN DISCH_TAC THEN - MP_TAC (SPECL[`n2:num`; `e2:num`] FLOAT_T_NEG) THEN - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; - -let float_interval_abs_tf = (RULE o prove) - (`interval_arith x (float_num T n1 e1, float_num F n2 e2) /\ - max (float_num F n1 e1) (float_num F n2 e2) = t - ==> interval_arith (abs x) (float_num F 0 min_exp, t)`, - MP_TAC ((GSYM o SPEC `F`) FLOAT_NEG) THEN SIMP_TAC[] THEN DISCH_TAC THEN - REWRITE_TAC[interval_arith; FLOAT_0] THEN REAL_ARITH_TAC);; - -let float_interval_abs th = - let x, lo, hi = dest_float_interval (concl th) in - let s1, n1, e1 = dest_float lo and - s2, n2, e2 = dest_float hi in - let inst = INST[x, x_var_real; - n1, n1_var_num; e1, e1_var_num; - n2, n2_var_num; e2, e2_var_num] in - if s1 = s2 then - if s1 = "F" then - (* FF *) - MY_PROVE_HYP th (inst float_interval_abs_ff) - else - (* TT *) - MY_PROVE_HYP th (inst float_interval_abs_tt) - else - if s1 = "F" then - (* FT *) - MY_PROVE_HYP th (inst float_interval_abs_ft) - else - let max_th = float_max (make_float "F" n1 e1) hi in - let t_tm = rand (concl max_th) in - let th0 = INST[x, x_var_real; t_tm, t_var_real; - n1, n1_var_num; e1, e1_var_num; - n2, n2_var_num; e2, e2_var_num] float_interval_abs_tf in - MY_PROVE_HYP max_th (MY_PROVE_HYP th th0);; - - -(****************************************) -(* float_interval_lt *) - -let FLOAT_INTERVAL_LT = prove(`interval_arith x (lo1, hi1) /\ interval_arith y (lo2, hi2) /\ hi1 < lo2 - ==> x < y`, - REWRITE_TAC[interval_arith] THEN - REAL_ARITH_TAC);; - - -(****************************************) -(* float_interval_neg *) - -let FLOAT_INTERVAL_NEG = prove(`!s1 s2. interval_arith x (float_num s1 n1 e1, float_num s2 n2 e2) - ==> interval_arith (--x) (float_num (~s2) n2 e2, float_num (~s1) n1 e1)`, - REPEAT GEN_TAC THEN - DISCH_THEN (fun th -> MP_TAC (MATCH_MP INTERVAL_NEG th)) THEN - SIMP_TAC[FLOAT_NEG]);; - -let FLOAT_INTERVAL_NEG_FF = (UNDISCH_ALL o REWRITE_RULE[] o SPECL[`F`; `F`]) FLOAT_INTERVAL_NEG;; -let FLOAT_INTERVAL_NEG_FT = (UNDISCH_ALL o REWRITE_RULE[] o SPECL[`F`; `T`]) FLOAT_INTERVAL_NEG;; -let FLOAT_INTERVAL_NEG_TF = (UNDISCH_ALL o REWRITE_RULE[] o SPECL[`T`; `F`]) FLOAT_INTERVAL_NEG;; -let FLOAT_INTERVAL_NEG_TT = (UNDISCH_ALL o REWRITE_RULE[] o SPECL[`T`; `T`]) FLOAT_INTERVAL_NEG;; - -(* |- interval x (float s1 n1 e1, float s2 n2 e2) -> - |- interval (--x) (float ~s2 n2 e2, float ~s1 n1 e1 *) -let float_interval_neg th = - let x_tm, f1, f2 = dest_float_interval (concl th) in - let s1, n1_tm, e1_tm = dest_float f1 in - let s2, n2_tm, e2_tm = dest_float f2 in - let inst = INST[x_tm, x_var_real; n1_tm, n1_var_num; e1_tm, e1_var_num; - n2_tm, n2_var_num; e2_tm, e2_var_num] in - let th0 = - if s1 = "F" then - if s2 = "F" then - inst FLOAT_INTERVAL_NEG_FF - else - inst FLOAT_INTERVAL_NEG_FT - else - if s2 = "F" then - inst FLOAT_INTERVAL_NEG_TF - else - inst FLOAT_INTERVAL_NEG_TT in - MY_PROVE_HYP th th0;; - - -(***********************************************) -(* float_interval_mul *) - -let f1_1_var = `f1_1:real` and - f1_2_var = `f1_2:real` and - f2_1_var = `f2_1:real` and - f2_2_var = `f2_2:real`;; - - -(* FT_xx *) -let FLOAT_INTERVAL_MUL_FT_xx = (UNDISCH_ALL o NUMERALS_TO_NUM o REWRITE_RULE[GSYM IMP_IMP; min_exp_def] o prove)( - `interval_arith x (float_num F n1 e1, float_num T n2 e2) - ==> interval_arith (x * y) (float_num F 0 min_exp, float_num F 0 min_exp)`, - STRIP_TAC THEN - FIRST_X_ASSUM (fun th -> REWRITE_TAC[MATCH_MP FLOAT_INTERVAL_FT_IMP_0 th]) THEN - REWRITE_TAC[REAL_MUL_LZERO; interval_arith] THEN - MP_TAC (GEN_ALL (SPECL [`s:bool`; `0`] FLOAT_EQ_0)) THEN SIMP_TAC[REAL_LE_REFL]);; - - -(* xx_FT *) -let FLOAT_INTERVAL_MUL_xx_FT = (UNDISCH_ALL o NUMERALS_TO_NUM o REWRITE_RULE[GSYM IMP_IMP; min_exp_def] o prove)( - `interval_arith y (float_num F m1 r1, float_num T m2 r2) - ==> interval_arith (x * y) (float_num F 0 min_exp, float_num F 0 min_exp)`, - STRIP_TAC THEN - FIRST_X_ASSUM (fun th -> REWRITE_TAC[MATCH_MP FLOAT_INTERVAL_FT_IMP_0 th]) THEN - REWRITE_TAC[REAL_MUL_RZERO; interval_arith] THEN - MP_TAC (GEN_ALL (SPECL [`s:bool`; `0`] FLOAT_EQ_0)) THEN SIMP_TAC[REAL_LE_REFL]);; - - - -(* FF_FF *) -let FLOAT_INTERVAL_MUL_FF_FF = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( - `interval_arith x (float_num F n1 e1, float_num F n2 e2) /\ - interval_arith y (float_num F m1 r1, float_num F m2 r2) /\ - f1 <= float_num F n1 e1 * float_num F m1 r1 /\ - float_num F n2 e2 * float_num F m2 r2 <= f2 - ==> interval_arith (x * y) (f1, f2)`, - MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN - SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` MP_TAC THENL - [ - MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN - REWRITE_TAC[FLOAT_F_POS]; - ALL_TAC - ] THEN - REPEAT (POP_ASSUM (fun th -> ALL_TAC)) THEN - REWRITE_TAC[interval_arith] THEN - REPEAT STRIP_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `a * c:real` THEN - ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN - ASM_REWRITE_TAC[]; - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `b * d:real` THEN - ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN - ASM_REWRITE_TAC[] THEN - CONJ_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THENL - [ - EXISTS_TAC `a:real` THEN ASM_REWRITE_TAC[]; - EXISTS_TAC `c:real` THEN ASM_REWRITE_TAC[] - ] - ]);; - -(* TT_TT *) -let FLOAT_INTERVAL_MUL_TT_TT = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( - `interval_arith x (float_num T n1 e1, float_num T n2 e2) /\ - interval_arith y (float_num T m1 r1, float_num T m2 r2) /\ - f1 <= float_num T n2 e2 * float_num T m2 r2 /\ - float_num T n1 e1 * float_num T m1 r1 <= f2 - ==> interval_arith (x * y) (f1, f2)`, - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[interval_arith] THEN - MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN - SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` MP_TAC THENL - [ - MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN - REWRITE_TAC[FLOAT_F_POS]; - ALL_TAC - ] THEN - REWRITE_TAC[REAL_NEG_MUL2] THEN - REPEAT STRIP_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `b * d:real` THEN - ASM_REWRITE_TAC[] THEN - ONCE_REWRITE_TAC[REAL_ARITH `a <= x * y <=> a <= --x * --y`] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN - ONCE_REWRITE_TAC[REAL_ARITH `b <= --x <=> x <= --b`] THEN - ASM_REWRITE_TAC[]; - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `a * c:real` THEN - ASM_REWRITE_TAC[] THEN - ONCE_REWRITE_TAC[REAL_ARITH `x * y <= a <=> --x * --y <= a`] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN - ONCE_REWRITE_TAC[REAL_ARITH `--x <= c <=> --c <= x`] THEN - ASM_REWRITE_TAC[] THEN - ASSUME_TAC (REAL_ARITH `!b x. &0 <= b /\ x <= --b ==> &0 <= --x`) THEN - CONJ_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THENL - [ - EXISTS_TAC `b:real` THEN ASM_REWRITE_TAC[]; - EXISTS_TAC `d:real` THEN ASM_REWRITE_TAC[] - ] - ]);; - -(* FF_TT *) -let FLOAT_INTERVAL_MUL_FF_TT = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( - `interval_arith x (float_num F n1 e1, float_num F n2 e2) /\ - interval_arith y (float_num T m1 r1, float_num T m2 r2) /\ - f1 <= float_num F n2 e2 * float_num T m1 r1 /\ - float_num F n1 e1 * float_num T m2 r2 <= f2 - ==> interval_arith (x * y) (f1, f2)`, - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[interval_arith] THEN - MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN - SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` MP_TAC THENL - [ - MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN - REWRITE_TAC[FLOAT_F_POS]; - ALL_TAC - ] THEN - - REPEAT STRIP_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `b * --c` THEN - ASM_REWRITE_TAC[] THEN - ONCE_REWRITE_TAC[REAL_ARITH `b * --c <= x * y <=> x * --y <= b * c`] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN - ONCE_REWRITE_TAC[REAL_ARITH `--y <= c <=> --c <= y`] THEN - ASM_REWRITE_TAC[] THEN - CONJ_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THENL - [ - EXISTS_TAC `a:real` THEN ASM_REWRITE_TAC[]; - EXISTS_TAC `d:real` THEN - ONCE_REWRITE_TAC[REAL_ARITH `d <= --y <=> y <= --d`] THEN - ASM_REWRITE_TAC[] - ]; - - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `a * --d` THEN - ASM_REWRITE_TAC[] THEN - ONCE_REWRITE_TAC[REAL_ARITH `x * y <= a * --d <=> a * d <= x * --y`] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN - ONCE_REWRITE_TAC[REAL_ARITH `d <= --y <=> y <= --d`] THEN - ASM_REWRITE_TAC[] - ]);; - -(* TT_FF *) -let FLOAT_INTERVAL_MUL_TT_FF = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( - `interval_arith x (float_num T n1 e1, float_num T n2 e2) /\ - interval_arith y (float_num F m1 r1, float_num F m2 r2) /\ - f1 <= float_num T n1 e1 * float_num F m2 r2 /\ - float_num T n2 e2 * float_num F m1 r1 <= f2 - ==> interval_arith (x * y) (f1, f2)`, - STRIP_TAC THEN - MP_TAC ((GEN_ALL o DISCH_ALL) FLOAT_INTERVAL_MUL_FF_TT) THEN - DISCH_THEN (MP_TAC o SPECL[`n1:num`; `e1:num`; `n2:num`; `e2:num`; `m1:num`; `r1:num`; `m2:num`; `r2:num`]) THEN - DISCH_THEN (MP_TAC o SPECL[`y:real`; `x:real`; `f1:real`; `f2:real`]) THEN - REPEAT (POP_ASSUM MP_TAC) THEN - SIMP_TAC[REAL_MUL_AC]);; - -(* TF_FF *) -let FLOAT_INTERVAL_MUL_TF_FF = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( - `interval_arith x (float_num T n1 e1, float_num F n2 e2) /\ - interval_arith y (float_num F m1 r1, float_num F m2 r2) /\ - f1 <= float_num T n1 e1 * float_num F m2 r2 /\ - float_num F n2 e2 * float_num F m2 r2 <= f2 - ==> interval_arith (x * y) (f1, f2)`, - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[interval_arith] THEN - MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN - STRIP_TAC THEN - SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` ASSUME_TAC THENL - [ - MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN - REWRITE_TAC[FLOAT_F_POS]; - ALL_TAC - ] THEN - - SUBGOAL_THEN `&0 <= y` ASSUME_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `c:real` THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - - CONJ_TAC THENL - [ - DISJ_CASES_TAC (REAL_ARITH `&0 <= x \/ &0 <= --x`) THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN - CONJ_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * d` THEN - ASM_REWRITE_TAC[REAL_ARITH `--a * d <= &0 <=> &0 <= a * d`] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[] - ]; - ALL_TAC - ] THEN - - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * d` THEN ASM_REWRITE_TAC[] THEN - ONCE_REWRITE_TAC[REAL_ARITH `--a * d <= x * y <=> --x * y <= a * d`] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN - ONCE_REWRITE_TAC[REAL_ARITH `--x <= a <=> --a <= x`] THEN ASM_REWRITE_TAC[]; - - DISJ_CASES_TAC (REAL_ARITH `&0 <= --x \/ &0 <= x`) THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN - CONJ_TAC THENL - [ - REWRITE_TAC[REAL_ARITH `x * y <= &0 <=> &0 <= --x * y`] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * d` THEN ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[] - ]; - ALL_TAC - ] THEN - - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * d` THEN ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_REWRITE_TAC[] - ]);; - -(* TF_TT *) -let FLOAT_INTERVAL_MUL_TF_TT = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( - `interval_arith x (float_num T n1 e1, float_num F n2 e2) /\ - interval_arith y (float_num T m1 r1, float_num T m2 r2) /\ - f1 <= float_num F n2 e2 * float_num T m1 r1 /\ - float_num T n1 e1 * float_num T m1 r1 <= f2 - ==> interval_arith (x * y) (f1, f2)`, - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[interval_arith] THEN - MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN - STRIP_TAC THEN - SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` ASSUME_TAC THENL - [ - MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN - REWRITE_TAC[FLOAT_F_POS]; - ALL_TAC - ] THEN - - SUBGOAL_THEN `&0 <= --y` ASSUME_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `d:real` THEN - ONCE_REWRITE_TAC[REAL_ARITH `d <= --y <=> y <= --d`] THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - - CONJ_TAC THENL - [ - DISJ_CASES_TAC (REAL_ARITH `&0 <= --x \/ &0 <= x`) THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN - CONJ_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * --c` THEN - ASM_REWRITE_TAC[REAL_ARITH `b * --c <= &0 <=> &0 <= b * c`] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; - ONCE_REWRITE_TAC[REAL_ARITH `x * y = --x * --y`] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[] - ]; - ALL_TAC - ] THEN - - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * --c` THEN ASM_REWRITE_TAC[] THEN - ONCE_REWRITE_TAC[REAL_ARITH `b * --c <= x * y <=> x * --y <= b * c`] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN ONCE_REWRITE_TAC[REAL_ARITH `--y <= c <=> --c <= y`] THEN - ASM_REWRITE_TAC[]; - - DISJ_CASES_TAC (REAL_ARITH `&0 <= x \/ &0 <= --x`) THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN - CONJ_TAC THENL - [ - REWRITE_TAC[REAL_ARITH `x * y <= &0 <=> &0 <= x * --y`] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * --c` THEN ASM_REWRITE_TAC[REAL_NEG_MUL2] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[] - ]; - ALL_TAC - ] THEN - - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * --c` THEN ASM_REWRITE_TAC[REAL_NEG_MUL2] THEN - ONCE_REWRITE_TAC[REAL_ARITH `x * y <= a * c <=> --x * --y <= a * c`] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN ONCE_REWRITE_TAC[REAL_ARITH `--x <= a <=> --a <= x`] THEN - ASM_REWRITE_TAC[] - ]);; - -(* FF_TF *) -let FLOAT_INTERVAL_MUL_FF_TF = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( - `interval_arith x (float_num F n1 e1, float_num F n2 e2) /\ - interval_arith y (float_num T m1 r1, float_num F m2 r2) /\ - f1 <= float_num F n2 e2 * float_num T m1 r1 /\ - float_num F n2 e2 * float_num F m2 r2 <= f2 - ==> interval_arith (x * y) (f1, f2)`, - STRIP_TAC THEN - MP_TAC ((SPECL [`n1:num`; `e1:num`; `n2:num`; `e2:num`; `m1:num`; `r1:num`; `m2:num`; `r2:num`; `y:real`; `x:real`; `f1:real`; `f2:real`] o GEN_ALL o DISCH_ALL) FLOAT_INTERVAL_MUL_TF_FF) THEN - ASM_REWRITE_TAC[REAL_MUL_SYM] THEN DISCH_THEN MATCH_MP_TAC THEN - POP_ASSUM MP_TAC THEN REAL_ARITH_TAC);; - -(* TT_TF *) -let FLOAT_INTERVAL_MUL_TT_TF = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( - `interval_arith x (float_num T n1 e1, float_num T n2 e2) /\ - interval_arith y (float_num T m1 r1, float_num F m2 r2) /\ - f1 <= float_num T n1 e1 * float_num F m2 r2 /\ - float_num T n1 e1 * float_num T m1 r1 <= f2 - ==> interval_arith (x * y) (f1, f2)`, - STRIP_TAC THEN - MP_TAC ((SPECL [`n1:num`; `e1:num`; `n2:num`; `e2:num`; `m1:num`; `r1:num`; `m2:num`; `r2:num`; `y:real`; `x:real`; `f1:real`; `f2:real`] o GEN_ALL o DISCH_ALL) FLOAT_INTERVAL_MUL_TF_TT) THEN - ASM_REWRITE_TAC[REAL_MUL_SYM] THEN POP_ASSUM MP_TAC THEN POP_ASSUM MP_TAC THEN - SIMP_TAC[REAL_MUL_SYM]);; - -(* TF_TF *) -let FLOAT_INTERVAL_MUL_TF_TF = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( - `interval_arith x (float_num T n1 e1, float_num F n2 e2) /\ - interval_arith y (float_num T m1 r1, float_num F m2 r2) /\ - f1_1 <= float_num T n1 e1 * float_num F m2 r2 /\ - f1_2 <= float_num F n2 e2 * float_num T m1 r1 /\ - min f1_1 f1_2 = f1 /\ - float_num T n1 e1 * float_num T m1 r1 <= f2_1 /\ - float_num F n2 e2 * float_num F m2 r2 <= f2_2 /\ - max f2_1 f2_2 = f2 - ==> interval_arith (x * y) (f1, f2)`, - REWRITE_TAC[EQ_SYM_EQ; FLOAT_NEG_T] THEN - REWRITE_TAC[interval_arith] THEN - MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN - STRIP_TAC THEN - SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` ASSUME_TAC THENL - [ - MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN - REWRITE_TAC[FLOAT_F_POS]; - ALL_TAC - ] THEN - - DISJ_CASES_TAC (REAL_ARITH `&0 <= x \/ &0 <= --x`) THENL - [ - DISJ_CASES_TAC (REAL_ARITH `&0 <= y \/ &0 <= --y`) THENL - [ - CONJ_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN - CONJ_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f1_1:real` THEN - ASM_REWRITE_TAC[REAL_MIN_MIN] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * d` THEN - ASM_REWRITE_TAC[REAL_ARITH `--a * d <= &0 <=> &0 <= a * d`] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * d` THEN - CONJ_TAC THENL - [ - MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f2_2:real` THEN - ASM_REWRITE_TAC[REAL_MAX_MAX]; - ALL_TAC - ] THEN - - CONJ_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f1_2:real` THEN - ASM_REWRITE_TAC[REAL_MIN_MIN] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * --c` THEN - ONCE_REWRITE_TAC[REAL_ARITH `b * --c <= x * y <=> x * --y <= b * c`] THEN ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN ONCE_REWRITE_TAC[REAL_ARITH `--y <= c <=> --c <= y`] THEN - ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN - CONJ_TAC THENL - [ - REWRITE_TAC[REAL_ARITH `x * y <= &0 <=> &0 <= x * --y`] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f2_2:real` THEN - ASM_REWRITE_TAC[REAL_MAX_MAX] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * d` THEN ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - - DISJ_CASES_TAC (REAL_ARITH `&0 <= y \/ &0 <= --y`) THENL - [ - CONJ_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f1_1:real` THEN - ASM_REWRITE_TAC[REAL_MIN_MIN] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * d` THEN - ASM_REWRITE_TAC[] THEN ONCE_REWRITE_TAC[REAL_ARITH `--a * d <= x * y <=> --x * y <= a * d`] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN ONCE_REWRITE_TAC[REAL_ARITH `--x <= a <=> --a <= x`] THEN - ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN - CONJ_TAC THENL - [ - REWRITE_TAC[REAL_ARITH `x * y <= &0 <=> &0 <= --x * y`] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f2_2:real` THEN - ASM_REWRITE_TAC[REAL_MAX_MAX] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * d` THEN ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - - CONJ_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN - CONJ_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f1_1:real` THEN - ASM_REWRITE_TAC[REAL_MIN_MIN] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * d` THEN - ASM_REWRITE_TAC[REAL_ARITH `--a * d <= &0 <=> &0 <= a * d`] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - ONCE_REWRITE_TAC[GSYM REAL_NEG_MUL2] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * --c` THEN - CONJ_TAC THENL - [ - ONCE_REWRITE_TAC[GSYM REAL_NEG_MUL2] THEN REWRITE_TAC[REAL_NEG_NEG] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN ONCE_REWRITE_TAC[REAL_ARITH `--x <= a <=> --a <= x`] THEN - ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f2_1:real` THEN - ASM_REWRITE_TAC[REAL_MAX_MAX]);; - - -(****************************) -let float_interval_mul = - let mul_ft_xx th1 x y n1 e1 n2 e2 = - let th0 = INST[x, x_var_real; y, y_var_real; - n1, n1_var_num; e1, e1_var_num; - n2, n2_var_num; e2, e2_var_num] FLOAT_INTERVAL_MUL_FT_xx in - MY_PROVE_HYP th1 th0 in - let mul_xx_ft th2 x y m1 r1 m2 r2 = - let th0 = INST[x, x_var_real; y, y_var_real; - m1, m1_var_num; r1, r1_var_num; - m2, m2_var_num; r2, r2_var_num] FLOAT_INTERVAL_MUL_xx_FT in - MY_PROVE_HYP th2 th0 in - - fun pp th1 th2 -> - let x, l_lo, l_hi = dest_float_interval (concl th1) and - y, r_lo, r_hi = dest_float_interval (concl th2) in - let s1, n1, e1 = dest_float l_lo and - s2, n2, e2 = dest_float l_hi and - s3, m1, r1 = dest_float r_lo and - s4, m2, r2 = dest_float r_hi in - - (* Special case 1 *) - if s1 <> s2 && s1 = "F" then - mul_ft_xx th1 x y n1 e1 n2 e2 - else if s3 <> s4 && s3 = "F" then - mul_xx_ft th2 x y m1 r1 m2 r2 - else - (* Special case 2 *) - if s1 <> s2 && s3 <> s4 then - let lo1, lo2 = float_mul_lo pp l_lo r_hi, float_mul_lo pp l_hi r_lo and - hi1, hi2 = float_mul_hi pp l_lo r_lo, float_mul_hi pp l_hi r_hi in - let f1_1 = (lhand o concl) lo1 and - f1_2 = (lhand o concl) lo2 and - f2_1 = (rand o concl) hi1 and - f2_2 = (rand o concl) hi2 in - let min_th = float_min f1_1 f1_2 and - max_th = float_max f2_1 f2_2 in - let f1_tm = (rand o concl) min_th and - f2_tm = (rand o concl) max_th in - let th0 = INST[x, x_var_real; n1, n1_var_num; e1, e1_var_num; - y, y_var_real; n2, n2_var_num; e2, e2_var_num; - m1, m1_var_num; r1, r1_var_num; - m2, m2_var_num; r2, r2_var_num; - f1_tm, f1_var_real; f2_tm, f2_var_real; - f1_1, f1_1_var; f1_2, f1_2_var; - f2_1, f2_1_var; f2_2, f2_2_var] FLOAT_INTERVAL_MUL_TF_TF in - (MY_PROVE_HYP min_th o MY_PROVE_HYP max_th o MY_PROVE_HYP lo1 o MY_PROVE_HYP lo2 o - MY_PROVE_HYP hi1 o MY_PROVE_HYP hi2 o MY_PROVE_HYP th1 o MY_PROVE_HYP th2) th0 - else - let lo_th, hi_th, th0 = - if s1 <> s2 then - if s3 = "F" then - float_mul_lo pp l_lo r_hi, float_mul_hi pp l_hi r_hi, FLOAT_INTERVAL_MUL_TF_FF - else - float_mul_lo pp l_hi r_lo, float_mul_hi pp l_lo r_lo, FLOAT_INTERVAL_MUL_TF_TT - else - if s3 <> s4 then - if s1 = "F" then - float_mul_lo pp l_hi r_lo, float_mul_hi pp l_hi r_hi, FLOAT_INTERVAL_MUL_FF_TF - else - float_mul_lo pp l_lo r_hi, float_mul_hi pp l_lo r_lo, FLOAT_INTERVAL_MUL_TT_TF - else - if s1 = "F" then - if s3 = "F" then - float_mul_lo pp l_lo r_lo, float_mul_hi pp l_hi r_hi, FLOAT_INTERVAL_MUL_FF_FF - else - float_mul_lo pp l_hi r_lo, float_mul_hi pp l_lo r_hi, FLOAT_INTERVAL_MUL_FF_TT - else - if s3 = "F" then - float_mul_lo pp l_lo r_hi, float_mul_hi pp l_hi r_lo, FLOAT_INTERVAL_MUL_TT_FF - else - float_mul_lo pp l_hi r_hi, float_mul_hi pp l_lo r_lo, FLOAT_INTERVAL_MUL_TT_TT in - - let f1_tm = lhand(concl lo_th) and - f2_tm = rand(concl hi_th) in - - let th = INST[x, x_var_real; n1, n1_var_num; e1, e1_var_num; - y, y_var_real; n2, n2_var_num; e2, e2_var_num; - m1, m1_var_num; r1, r1_var_num; - m2, m2_var_num; r2, r2_var_num; - f1_tm, f1_var_real; f2_tm, f2_var_real] th0 in - MY_PROVE_HYP lo_th (MY_PROVE_HYP hi_th (MY_PROVE_HYP th1 (MY_PROVE_HYP th2 th)));; - - -(*************************************) -(* float_interval_div *) - -(* FT_xx *) -let FLOAT_INTERVAL_DIV_FT_xx = prove( - `interval_arith x (float_num F n1 e1, float_num T n2 e2) - ==> interval_arith (x / y) (float_num F 0 min_exp, float_num F 0 min_exp)`, - REWRITE_TAC[real_div] THEN DISCH_THEN (MP_TAC o MATCH_MP FLOAT_INTERVAL_FT_IMP_0) THEN - SIMP_TAC[REAL_MUL_LZERO; interval_arith] THEN - MP_TAC (GEN_ALL (SPECL [`s:bool`; `0`] FLOAT_EQ_0)) THEN SIMP_TAC[REAL_LE_REFL]);; - -(* FF_FF *) -let FLOAT_INTERVAL_DIV_FF_FF = prove( - `~(m1 = 0) /\ - interval_arith x (float_num F n1 e1, float_num F n2 e2) /\ - interval_arith y (float_num F m1 r1, float_num F m2 r2) /\ - f1 <= float_num F n1 e1 / float_num F m2 r2 /\ - float_num F n2 e2 / float_num F m1 r1 <= f2 - ==> interval_arith (x / y) (f1, f2)`, - MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN - REWRITE_TAC[real_div] THEN - STRIP_TAC THEN - SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` MP_TAC THENL - [ - MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN - REWRITE_TAC[FLOAT_F_POS]; - ALL_TAC - ] THEN - SUBGOAL_THEN `~(c = &0)` ASSUME_TAC THENL - [ - EXPAND_TAC "c" THEN ASM_REWRITE_TAC[FLOAT_EQ_0]; - ALL_TAC - ] THEN - STRIP_TAC THEN - SUBGOAL_THEN `~(d = &0)` MP_TAC THENL - [ - MATCH_MP_TAC (REAL_ARITH `~(c = &0) /\ &0 <= c /\ c <= d ==> ~(d = &0)`) THEN - ASM_REWRITE_TAC[] THEN - UNDISCH_TAC `interval_arith y (c,d)` THEN - REWRITE_TAC[interval_arith] THEN - REAL_ARITH_TAC; - ALL_TAC - ] THEN - - REPLICATE_TAC 10 (POP_ASSUM MP_TAC) THEN - REPEAT (POP_ASSUM (fun th -> ALL_TAC)) THEN - REWRITE_TAC[interval_arith] THEN - REPEAT STRIP_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `a * inv d` THEN - ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN - ASM_REWRITE_TAC[REAL_LE_INV_EQ] THEN - MATCH_MP_TAC REAL_LE_INV2 THEN - ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LTE_TRANS THEN - EXISTS_TAC `c:real` THEN - ASM_REWRITE_TAC[REAL_ARITH `&0 < c <=> ~(c = &0) /\ &0 <= c`]; - - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `b * inv c` THEN - ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN - ASM_REWRITE_TAC[REAL_LE_INV_EQ] THEN - REPEAT CONJ_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `a:real` THEN ASM_REWRITE_TAC[]; - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `c:real` THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - MATCH_MP_TAC REAL_LE_INV2 THEN - ASM_REWRITE_TAC[REAL_ARITH `&0 < c <=> ~(c = &0) /\ &0 <= c`] - ]);; - -(* TT_TT *) -let FLOAT_INTERVAL_DIV_TT_TT = prove( - `~(m2 = 0) /\ - interval_arith x (float_num T n1 e1, float_num T n2 e2) /\ - interval_arith y (float_num T m1 r1, float_num T m2 r2) /\ - f1 <= float_num T n2 e2 / float_num T m1 r1 /\ - float_num T n1 e1 / float_num T m2 r2 <= f2 - ==> interval_arith (x / y) (f1,f2)`, - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[real_div; REAL_INV_NEG; REAL_NEG_MUL2] THEN - GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV)[interval_arith] THEN - REWRITE_TAC[REAL_ARITH `--a <= x /\ x <= --b <=> b <= --x /\ --x <= a`] THEN - GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV)[GSYM interval_arith] THEN - STRIP_TAC THEN - MP_TAC ((SPECL[`n2:num`; `e2:num`; `m1:num`; `r1:num`; `n1:num`; `e1:num`; `m2:num`; `r2:num`; `--x`; `--y`] o GEN_ALL) FLOAT_INTERVAL_DIV_FF_FF) THEN - REWRITE_TAC[real_div; REAL_INV_NEG; REAL_NEG_MUL2] THEN - DISCH_THEN MATCH_MP_TAC THEN - ASM_REWRITE_TAC[]);; - -(* FF_TT *) -let FLOAT_INTERVAL_DIV_FF_TT = prove( - `~(m2 = 0) /\ - interval_arith x (float_num F n1 e1, float_num F n2 e2) /\ - interval_arith y (float_num T m1 r1, float_num T m2 r2) /\ - f1 <= float_num F n2 e2 / float_num T m2 r2 /\ - float_num F n1 e1 / float_num T m1 r1 <= f2 - ==> interval_arith (x / y) (f1,f2)`, - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[real_div; REAL_INV_NEG] THEN - REWRITE_TAC[REAL_ARITH `a * --b <= c <=> --c <= a * b`] THEN - REWRITE_TAC[REAL_ARITH `c <= a * --b <=> a * b <= --c`] THEN - GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV)[interval_arith] THEN - REWRITE_TAC[REAL_ARITH `--a <= x /\ x <= --b <=> b <= --x /\ --x <= a`] THEN - GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV)[GSYM interval_arith] THEN - STRIP_TAC THEN - MP_TAC ((SPECL[`n1:num`; `e1:num`; `m1:num`; `r1:num`; `n2:num`; `e2:num`; `m2:num`; `r2:num`; `x:real`; `--y`; `--f2`; `--f1`] o GEN_ALL) FLOAT_INTERVAL_DIV_FF_FF) THEN - ANTS_TAC THENL - [ - ASM_REWRITE_TAC[real_div]; - ALL_TAC - ] THEN - REWRITE_TAC[real_div; REAL_INV_NEG; interval_arith] THEN - REAL_ARITH_TAC);; - -(* TT_FF *) -let FLOAT_INTERVAL_DIV_TT_FF = prove( - `~(m1 = 0) /\ - interval_arith x (float_num T n1 e1, float_num T n2 e2) /\ - interval_arith y (float_num F m1 r1, float_num F m2 r2) /\ - f1 <= float_num T n1 e1 / float_num F m1 r1 /\ - float_num T n2 e2 / float_num F m2 r2 <= f2 - ==> interval_arith (x / y) (f1,f2)`, - REWRITE_TAC[FLOAT_NEG_T] THEN - REWRITE_TAC[real_div; REAL_INV_NEG] THEN - REWRITE_TAC[REAL_ARITH `--a * b <= c <=> --c <= a * b`] THEN - REWRITE_TAC[REAL_ARITH `c <= --a * b <=> a * b <= --c`] THEN - GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV)[interval_arith] THEN - REWRITE_TAC[REAL_ARITH `--a <= x /\ x <= --b <=> b <= --x /\ --x <= a`] THEN - GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV)[GSYM interval_arith] THEN - STRIP_TAC THEN - MP_TAC ((SPECL[`n2:num`; `e2:num`; `m2:num`; `r2:num`; `n1:num`; `e1:num`; `m1:num`; `r1:num`; `--x:real`; `y:real`; `--f2`; `--f1`] o GEN_ALL) FLOAT_INTERVAL_DIV_FF_FF) THEN - ANTS_TAC THENL - [ - ASM_REWRITE_TAC[real_div]; - ALL_TAC - ] THEN - REWRITE_TAC[real_div; REAL_INV_NEG; interval_arith] THEN - REAL_ARITH_TAC);; - - -(* TF_FF *) -let FLOAT_INTERVAL_DIV_TF_FF = prove( - `~(m1 = 0) /\ - interval_arith x (float_num T n1 e1, float_num F n2 e2) /\ - interval_arith y (float_num F m1 r1, float_num F m2 r2) /\ - f1 <= float_num T n1 e1 / float_num F m1 r1 /\ - float_num F n2 e2 / float_num F m1 r1 <= f2 - ==> interval_arith (x / y) (f1,f2)`, - REWRITE_TAC[FLOAT_NEG_T] THEN STRIP_TAC THEN - MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN - SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` ASSUME_TAC THENL - [ - MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN REWRITE_TAC[FLOAT_F_POS]; - ALL_TAC - ] THEN - - DISJ_CASES_TAC (REAL_ARITH `&0 <= x \/ x <= &0`) THENL - [ - SUBGOAL_THEN `interval_arith x (float_num F 0 0, b)` ASSUME_TAC THENL - [ - UNDISCH_TAC `interval_arith x (--a, b)` THEN POP_ASSUM MP_TAC THEN - REWRITE_TAC[interval_arith; FLOAT_0] THEN REAL_ARITH_TAC; - ALL_TAC - ] THEN - - MP_TAC ((SPEC_ALL o SPECL [`0`; `0`] o GEN_ALL) FLOAT_INTERVAL_DIV_FF_FF) THEN - ASM_REWRITE_TAC[] THEN DISCH_THEN MATCH_MP_TAC THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN - REWRITE_TAC[real_div] THEN CONJ_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a / c` THEN - ASM_REWRITE_TAC[real_div; ARITH_RULE `--a * b <= &0 <=> &0 <= a * b`] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[REAL_LE_INV_EQ]; - ALL_TAC - ] THEN - - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[REAL_LE_INV_EQ; FLOAT_F_POS]; - ALL_TAC - ] THEN - - SUBGOAL_THEN `interval_arith x (--a, float_num T 0 0)` ASSUME_TAC THENL - [ - UNDISCH_TAC `interval_arith x (--a, b)` THEN POP_ASSUM MP_TAC THEN - REWRITE_TAC[interval_arith; FLOAT_0] THEN REAL_ARITH_TAC; - ALL_TAC - ] THEN - - MP_TAC ((INST[`0`, `n2:num`; `0`, `e2:num`]) FLOAT_INTERVAL_DIV_TT_FF) THEN - ASM_REWRITE_TAC[FLOAT_NEG_T] THEN DISCH_THEN MATCH_MP_TAC THEN - ASM_REWRITE_TAC[GSYM FLOAT_NEG_T] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN - REWRITE_TAC[real_div] THEN CONJ_TAC THENL - [ - REWRITE_TAC[FLOAT_0; REAL_MUL_LZERO; REAL_LE_REFL]; - ALL_TAC - ] THEN - - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b / c` THEN ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_DIV THEN ASM_REWRITE_TAC[]);; - -(* TF_TT *) -let FLOAT_INTERVAL_DIV_TF_TT = prove( - `~(m2 = 0) /\ - interval_arith x (float_num T n1 e1, float_num F n2 e2) /\ - interval_arith y (float_num T m1 r1, float_num T m2 r2) /\ - f1 <= float_num F n2 e2 / float_num T m2 r2 /\ - float_num T n1 e1 / float_num T m2 r2 <= f2 - ==> interval_arith (x / y) (f1,f2)`, - REWRITE_TAC[FLOAT_NEG_T] THEN STRIP_TAC THEN - MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN - SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` ASSUME_TAC THENL - [ - MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN REWRITE_TAC[FLOAT_F_POS]; - ALL_TAC - ] THEN - - DISJ_CASES_TAC (REAL_ARITH `x <= &0 \/ &0 <= x`) THENL - [ - SUBGOAL_THEN `interval_arith x (--a, float_num T 0 0)` ASSUME_TAC THENL - [ - UNDISCH_TAC `interval_arith x (--a, b)` THEN POP_ASSUM MP_TAC THEN - REWRITE_TAC[interval_arith; FLOAT_0] THEN REAL_ARITH_TAC; - ALL_TAC - ] THEN - - MP_TAC ((SPEC_ALL o SPECL [`0`; `0`] o GEN_ALL) FLOAT_INTERVAL_DIV_TT_TT) THEN - ASM_REWRITE_TAC[] THEN DISCH_THEN MATCH_MP_TAC THEN ASM_REWRITE_TAC[FLOAT_NEG_T] THEN - ASM_REWRITE_TAC[GSYM FLOAT_NEG_T] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN - REWRITE_TAC[real_div] THEN CONJ_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b / --d` THEN - ASM_REWRITE_TAC[real_div; REAL_INV_NEG; REAL_ARITH `b * --d <= &0 <=> &0 <= b * d`] THEN - MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[REAL_LE_INV_EQ]; - ALL_TAC - ] THEN - - REWRITE_TAC[FLOAT_0; REAL_MUL_LZERO; REAL_LE_REFL]; - ALL_TAC - ] THEN - - SUBGOAL_THEN `interval_arith x (float_num F 0 0, b)` ASSUME_TAC THENL - [ - UNDISCH_TAC `interval_arith x (--a, b)` THEN POP_ASSUM MP_TAC THEN - REWRITE_TAC[interval_arith; FLOAT_0] THEN REAL_ARITH_TAC; - ALL_TAC - ] THEN - - MP_TAC ((INST[`0`, `n1:num`; `0`, `e1:num`]) FLOAT_INTERVAL_DIV_FF_TT) THEN - ASM_REWRITE_TAC[FLOAT_NEG_T] THEN DISCH_THEN MATCH_MP_TAC THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN - REWRITE_TAC[real_div] THEN CONJ_TAC THENL - [ - REWRITE_TAC[FLOAT_0; REAL_MUL_LZERO; REAL_LE_REFL]; - ALL_TAC - ] THEN - - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a / --d` THEN ASM_REWRITE_TAC[] THEN - REWRITE_TAC[real_div; REAL_INV_NEG; REAL_NEG_MUL2] THEN MATCH_MP_TAC REAL_LE_MUL THEN - ASM_REWRITE_TAC[REAL_LE_INV_EQ]);; - -let transform = UNDISCH_ALL o - PURE_REWRITE_RULE[TAUT `~P <=> (P <=> F)`] o - NUMERALS_TO_NUM o - REWRITE_RULE[GSYM IMP_IMP; min_exp_def];; - -let FLOAT_INTERVAL_DIV_FT_xx' = transform FLOAT_INTERVAL_DIV_FT_xx and - FLOAT_INTERVAL_DIV_FF_FF' = transform FLOAT_INTERVAL_DIV_FF_FF and - FLOAT_INTERVAL_DIV_TT_TT' = transform FLOAT_INTERVAL_DIV_TT_TT and - FLOAT_INTERVAL_DIV_FF_TT' = transform FLOAT_INTERVAL_DIV_FF_TT and - FLOAT_INTERVAL_DIV_TT_FF' = transform FLOAT_INTERVAL_DIV_TT_FF and - FLOAT_INTERVAL_DIV_TF_FF' = transform FLOAT_INTERVAL_DIV_TF_FF and - FLOAT_INTERVAL_DIV_TF_TT' = transform FLOAT_INTERVAL_DIV_TF_TT;; - - -let float_interval_div pp th1 th2 = - let x, l_lo, l_hi = dest_float_interval (concl th1) and - y, r_lo, r_hi = dest_float_interval (concl th2) in - let s1, n1, e1 = dest_float l_lo and - s2, n2, e2 = dest_float l_hi and - s3, m1, r1 = dest_float r_lo and - s4, m2, r2 = dest_float r_hi in - - if s1 <> s2 && s1 = "F" then - let th0 = INST[x, x_var_real; y, y_var_real; - n1, n1_var_num; e1, e1_var_num; - n2, n2_var_num; e2, e2_var_num] FLOAT_INTERVAL_DIV_FT_xx' in - MY_PROVE_HYP th1 th0 - else if s3 <> s4 then - failwith "float_interval_div: division by an interval containing 0" - else - let lo_th, hi_th, th0, zero_th = - if s1 = s2 then - if s1 = "F" then - if s3 = "F" then - float_div_lo pp l_lo r_hi, float_div_hi pp l_hi r_lo, - FLOAT_INTERVAL_DIV_FF_FF', raw_eq0_hash_conv m1 - else - float_div_lo pp l_hi r_hi, float_div_hi pp l_lo r_lo, - FLOAT_INTERVAL_DIV_FF_TT', raw_eq0_hash_conv m2 - else - if s3 = "F" then - float_div_lo pp l_lo r_lo, float_div_hi pp l_hi r_hi, - FLOAT_INTERVAL_DIV_TT_FF', raw_eq0_hash_conv m1 - else - float_div_lo pp l_hi r_lo, float_div_hi pp l_lo r_hi, - FLOAT_INTERVAL_DIV_TT_TT', raw_eq0_hash_conv m2 - else - if s3 = "F" then - float_div_lo pp l_lo r_lo, float_div_hi pp l_hi r_lo, - FLOAT_INTERVAL_DIV_TF_FF', raw_eq0_hash_conv m1 - else - float_div_lo pp l_hi r_hi, float_div_hi pp l_lo r_hi, - FLOAT_INTERVAL_DIV_TF_TT', raw_eq0_hash_conv m2 in - - let f1_tm = lhand(concl lo_th) and - f2_tm = rand(concl hi_th) in - - let th = INST[x, x_var_real; n1, n1_var_num; e1, e1_var_num; - y, y_var_real; n2, n2_var_num; e2, e2_var_num; - m1, m1_var_num; r1, r1_var_num; - m2, m2_var_num; r2, r2_var_num; - f1_tm, f1_var_real; f2_tm, f2_var_real] th0 in - (MY_PROVE_HYP lo_th o MY_PROVE_HYP hi_th o - MY_PROVE_HYP th1 o MY_PROVE_HYP th2 o MY_PROVE_HYP zero_th) th;; - - -(*****************************************) -(* float_interval_add, float_interval_sub *) - -let n1_var_real = `n1:real` and - n2_var_real = `n2:real` and - m1_var_real = `m1:real` and - m2_var_real = `m2:real` and - n_var_real = `n:real` and - m_var_real = `m:real`;; - -let INTERVAL_ADD = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( - `interval_arith x (n1, m1) /\ - interval_arith y (n2, m2) /\ - n <= n1 + n2 /\ m1 + m2 <= m - ==> interval_arith (x + y) (n, m)`, - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; - -let INTERVAL_SUB = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( - `interval_arith x (n1, m1) /\ - interval_arith y (n2, m2) /\ - n <= n1 - m2 /\ m1 - n2 <= m - ==> interval_arith (x - y) (n, m)`, - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; - - - let float_interval_add pp th1 th2 = - let x, n1, m1 = dest_float_interval (concl th1) in - let y, n2, m2 = dest_float_interval (concl th2) in - let lo_th = float_add_lo pp n1 n2 in - let hi_th = float_add_hi pp m1 m2 in - let n_tm = lhand (concl lo_th) in - let m_tm = rand (concl hi_th) in - let th0 = INST[x, x_var_real; n1, n1_var_real; m1, m1_var_real; - y, y_var_real; n2, n2_var_real; m2, m2_var_real; - n_tm, n_var_real; m_tm, m_var_real] INTERVAL_ADD in - MY_PROVE_HYP lo_th (MY_PROVE_HYP hi_th (MY_PROVE_HYP th2 (MY_PROVE_HYP th1 th0)));; - -let float_interval_sub pp th1 th2 = - let x, n1, m1 = dest_float_interval (concl th1) in - let y, n2, m2 = dest_float_interval (concl th2) in - let lo_th = float_sub_lo pp n1 m2 in - let hi_th = float_sub_hi pp m1 n2 in - let n_tm = lhand(concl lo_th) in - let m_tm = rand(concl hi_th) in - let th0 = INST[x, x_var_real; n1, n1_var_real; m1, m1_var_real; - y, y_var_real; n2, n2_var_real; m2, m2_var_real; - n_tm, n_var_real; m_tm, m_var_real] INTERVAL_SUB in - MY_PROVE_HYP lo_th (MY_PROVE_HYP hi_th (MY_PROVE_HYP th2 (MY_PROVE_HYP th1 th0)));; - - -(********************************************) -(* FLOAT_ABS *) - -let s_var_bool = `s:bool`;; - -let float_abs tm = - let ltm, rtm = dest_comb tm in - if ((fst o dest_const) ltm <> "real_abs") then - failwith "float_abs: no abs" - else - let ltm, e_tm = dest_comb rtm in - let ltm, n_tm = dest_comb ltm in - let s_tm = rand ltm in - INST[s_tm, s_var_bool; n_tm, n_var_num; e_tm, e_var_num] FLOAT_ABS;; - - -(*******************************) -(* float_interval_sqrt *) - -let FLOAT_INTERVAL_SQRT = prove(`interval_arith x (float_num F n1 e1, hi) /\ - f1 <= sqrt (float_num F n1 e1) /\ sqrt hi <= f2 - ==> interval_arith (sqrt x) (f1, f2)`, - ABBREV_TAC `lo = float_num F n1 e1` THEN - REWRITE_TAC[interval_arith] THEN - STRIP_TAC THEN - SUBGOAL_THEN `&0 <= lo /\ &0 <= hi` ASSUME_TAC THENL - [ - EXPAND_TAC "lo" THEN - REWRITE_TAC[FLOAT_F_POS] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `x:real` THEN - ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `lo:real` THEN - ASM_REWRITE_TAC[] THEN - EXPAND_TAC "lo" THEN - REWRITE_TAC[FLOAT_F_POS]; - ALL_TAC - ] THEN - SUBGOAL_THEN `&0 <= x` ASSUME_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `lo:real` THEN - ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - CONJ_TAC THENL - [ - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `sqrt lo` THEN - ASM_REWRITE_TAC[SQRT_MONO_LE_EQ]; - MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `sqrt hi` THEN - ASM_REWRITE_TAC[SQRT_MONO_LE_EQ] - ]);; - - -let FLOAT_INTERVAL_SQRT' = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP]) FLOAT_INTERVAL_SQRT;; - -let float_interval_sqrt pp th = - let x_tm, lo_tm, hi_tm = dest_float_interval (concl th) in - let s1, n1_tm, e1_tm = dest_float lo_tm in - if s1 <> "F" then - failwith "float_interval_sqrt: negative low bound" - else - let lo_th = float_sqrt_lo pp lo_tm in - let hi_th = float_sqrt_hi pp hi_tm in - let f1_tm = lhand (concl lo_th) in - let f2_tm = rand (concl hi_th) in - let th0 = INST[x_tm, x_var_real; n1_tm, n1_var_num; e1_tm, e1_var_num; - hi_tm, hi_var_real; f1_tm, f1_var_real; - f2_tm, f2_var_real] FLOAT_INTERVAL_SQRT' in - MY_PROVE_HYP lo_th (MY_PROVE_HYP hi_th (MY_PROVE_HYP th th0));; - - -(******************************************) -(* FLOAT_TO_NUM_CONV *) - -let FLOAT_TO_NUM_CONV tm = - let ltm, e_tm = dest_comb tm in - let f_tm, n_tm = dest_comb ltm in - if (fst o dest_const o rator) f_tm <> "float_num" then - failwith "FLOAT_TO_NUM_CONV" - else - let n_th' = SYM (INST[n_tm, n_var_num] Arith_num.NUM_THM) in - let e_th' = SYM (INST[e_tm, n_var_num] Arith_num.NUM_THM) in - let n_th = TRANS n_th' (NUM_TO_NUMERAL_CONV (mk_comb(Arith_num.num_const, n_tm))) in - let e_th = TRANS e_th' (NUM_TO_NUMERAL_CONV (mk_comb(Arith_num.num_const, e_tm))) in - let th0 = MK_COMB (AP_TERM f_tm n_th, e_th) in - let tm0 = rand(concl th0) in - let th1 = (REWRITE_CONV[float_alt; arith_base_def; min_exp_def] THENC - NUM_REDUCE_CONV THENC - REAL_RAT_REDUCE_CONV) tm0 in - TRANS th0 th1;; - -end;; - - -(**************************************) -(* Printer for floating-point numbers *) -(**************************************) - -let print_float fmt tm = - try - let s, m_tm, e_tm = Arith_float.dest_float tm in - let m = Arith_num.raw_dest_hash m_tm and - e = Arith_num.raw_dest_hash e_tm -/ num Float_theory.min_exp in - let s_str = if s = "T" then "-" else "" in - let m_str = Num.string_of_num m in - let e_str = if e = num_0 then "" - else "*" ^ string_of_int Arith_num.arith_base ^ "^" ^ Num.string_of_num e in - let str = "##" ^ s_str ^ m_str ^ e_str in - Format.pp_print_string fmt str - with _ -> failwith "print_float";; - -install_user_printer ("float_num", print_float);; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal floating point arithmetic *) +(* -------------------------------------------------------------------------- *) + +(* Dependencies *) +needs "Formal_ineqs/arith/arith_nat.hl";; +needs "Formal_ineqs/arith/num_exp_theory.hl";; +needs "Formal_ineqs/arith/float_theory.hl";; +needs "Formal_ineqs/arith/interval_arith.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; + +needs "Formal_ineqs/tests/log.hl";; + +(* FLOOR_DIV_DIV *) +needs "Library/floor.ml";; + +module type Arith_float_sig = + sig + val mk_num_exp : term -> term -> term + val dest_num_exp : term -> (term * term) + val dest_float : term -> (string * term * term) + val make_float : string -> term -> term -> term + val mk_float : int -> int -> term + + (* num_exp normalization/denormalization *) + val normalize : term -> thm * bool + val denormalize : term -> thm + + val lo_num_conv : int -> term -> thm + val hi_num_conv : int -> term -> thm + val num_exp_lo : int -> term -> thm + val num_exp_hi : int -> term -> thm + + val float_lt0 : term -> thm + val float_gt0 : term -> thm + val float_lt : term -> term -> thm + val float_le0 : term -> thm + val float_ge0 : term -> thm + val float_le : term -> term -> thm + val float_min : term -> term -> thm + val float_max : term -> term -> thm + val float_min_max : term -> term -> (thm * thm) + val float_mul_eq : term -> term -> thm + val float_mul_lo : int -> term -> term -> thm + val float_mul_hi : int -> term -> term -> thm + val float_div_lo : int -> term -> term -> thm + val float_div_hi : int -> term -> term -> thm + val float_add_lo : int -> term -> term -> thm + val float_add_hi : int -> term -> term -> thm + val float_sub_lo : int -> term -> term -> thm + val float_sub_hi : int -> term -> term -> thm + val float_sqrt_lo : int -> term -> thm + val float_sqrt_hi : int -> term -> thm + + val float_prove_le : term -> term -> bool * thm + val float_prove_lt : term -> term -> bool * thm + val float_prove_le_interval : term -> thm -> bool * thm + val float_prove_ge_interval : term -> thm -> bool * thm + val float_prove_lt_interval : term -> thm -> bool * thm + val float_prove_gt_interval : term -> thm -> bool * thm + val float_compare_interval : term -> thm -> int * thm + + val init_logs : unit -> unit + val reset_stat : unit -> unit + val reset_cache : unit -> unit + val print_stat : unit -> unit + + val dest_float_interval : term -> term * term * term + val mk_float_interval_small_num : int -> thm + val mk_float_interval_num : num -> thm + + val float_lo : int -> term -> thm + val float_hi : int -> term -> thm + + val float_interval_round : int -> thm -> thm + + val float_interval_abs : thm -> thm + val float_interval_neg : thm -> thm + val float_interval_mul : int -> thm -> thm -> thm + val float_interval_div : int -> thm -> thm -> thm + val float_interval_add : int -> thm -> thm -> thm + val float_interval_sub : int -> thm -> thm -> thm + val float_interval_sqrt : int -> thm -> thm + + val float_abs : term -> thm + + val FLOAT_TO_NUM_CONV : term -> thm +end;; + + +module Arith_float : Arith_float_sig = struct + +open Misc_functions;; +open Arith_nat;; +open Num_exp_theory;; +open Float_theory;; +open Interval_arith;; +open Misc_vars;; + +prioritize_real();; + +(* interval *) + +let APPROX_INTERVAL' = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL) APPROX_INTERVAL;; + +let interval_const = `interval_arith` and + num_exp_const = `num_exp`;; + +let b0_const = (fst o dest_comb o lhand o concl) (Arith_num.def_array.(0));; +let b0_name = (fst o dest_const) b0_const;; +let base_const = mk_small_numeral Arith_num.arith_base;; + +let NUM_REMOVE = prove(mk_eq(mk_comb(Arith_num.num_const, n_var_num), n_var_num), + REWRITE_TAC[Arith_num.num_def; NUMERAL]);; + +(* B0 n = base * n *) +let b0_thm = prove(mk_eq(mk_comb(b0_const, n_var_num), + mk_binop mul_op_num base_const n_var_num), + REWRITE_TAC[Arith_num.def_array.(0)] THEN + TRY ARITH_TAC THEN + GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV) [BIT0] THEN + ARITH_TAC);; + +let dest_num_exp tm = + let ltm, e_tm = dest_comb tm in + rand ltm, e_tm;; + +let num_exp_const = `num_exp`;; +let mk_num_exp n_tm e_tm = mk_binop num_exp_const n_tm e_tm;; + +(* float_num s n e -> "s", n, e *) +let dest_float tm = + let ltm, e_tm = dest_comb tm in + let ltm, n_tm = dest_comb ltm in + let float_tm, s_tm = dest_comb ltm in + if (fst o dest_const) float_tm <> "float_num" then + failwith "dest_float: not float" + else + (fst o dest_const) s_tm, n_tm, e_tm;; + +(* "s", n, e -> float_num s n e *) +let make_float = + let float_const = `float_num` in + fun s n_tm e_tm -> + let s_tm = if s = "T" then t_const else f_const in + mk_comb (mk_comb (mk_comb (float_const, s_tm), n_tm), e_tm);; + +(* Creates a float term with the value (n * base^e) *) +let mk_float = + let float_const = `float_num` in + fun n e -> + let n, s = if n < 0 then -n, t_const else n, f_const in + let n_tm = rand (Arith_nat.mk_small_numeral_array n) in + let e_tm = rand (Arith_nat.mk_small_numeral_array (e + Float_theory.min_exp)) in + mk_comb(mk_comb(mk_comb (float_const, s), n_tm), e_tm);; + + +(************************************) + +let NUM_EXP_EXP' = SPEC_ALL NUM_EXP_EXP;; +let NUM_EXP_0' = (SPEC_ALL o REWRITE_RULE[NUMERAL]) NUM_EXP_0;; +let NUM_EXP_LE' = (UNDISCH_ALL o SPEC_ALL) NUM_EXP_LE;; +let NUM_EXP_LT' = (UNDISCH_ALL o SPEC_ALL) NUM_EXP_LT;; + +(* B0 n = num_exp n bits *) +let normal_lemma1 = prove(mk_eq(mk_comb(b0_const, n_var_num), `num_exp n 1`), + REWRITE_TAC[Arith_num.def_array.(0); num_exp] THEN + TRY ARITH_TAC THEN + GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV) [BIT0] THEN + ARITH_TAC);; + +let NORMAL_LEMMA1 = NUMERALS_TO_NUM normal_lemma1;; + +let normal_lemma2 = prove(mk_eq (mk_comb (b0_const, `num_exp n e`), `num_exp n (SUC e)`), + REWRITE_TAC[normal_lemma1; NUM_EXP_EXP] THEN ARITH_TAC);; + +let rec normalize tm = + if (is_comb tm) then + let ltm, rtm = dest_comb tm in + let lname = (fst o dest_const) ltm in + if (lname = b0_name) then + let lth = INST[rtm, n_var_num] NORMAL_LEMMA1 in + let rth, flag = normalize rtm in + if flag then + let ltm, lexp = (dest_comb o snd o dest_eq o concl) lth in + let ltm, rtm = dest_comb ltm in + let rn, rexp = (dest_comb o snd o dest_eq o concl) rth in + let rn = rand rn in + let th1 = AP_THM (AP_TERM ltm rth) lexp in + let th2 = INST[rexp, e1_var_num; lexp, e2_var_num; rn, n_var_num] NUM_EXP_EXP' in + let th3 = TRANS lth (TRANS th1 th2) in + let ltm, rtm = (dest_comb o snd o dest_eq o concl) th3 in + let add_th = raw_add_conv_hash rtm in + let th4 = AP_TERM ltm add_th in + (TRANS th3 th4, true) + else + (lth, true) + else + (REFL tm, false) + else + (REFL tm, false);; + +(* Converts a raw numeral to a num_exp expression *) +let to_num_exp tm = + let x, flag = normalize tm in + if flag then x + else + INST[tm, n_var_num] NUM_EXP_0';; + +(************************************) + +let SYM_NUM_EXP_0' = SYM NUM_EXP_0';; +let NUM_EXP_n0 = prove(`!e. num_exp 0 e = 0`, REWRITE_TAC[num_exp; MULT_CLAUSES]);; +let NUM_EXP_n0' = (REWRITE_RULE[NUMERAL] o SPEC_ALL) NUM_EXP_n0;; + +let NUM_EXP_DENORM = (UNDISCH_ALL o prove) + (mk_imp(`e = _0 <=> F`, mk_eq(`num_exp n e`, mk_comb (b0_const, `num_exp n (PRE e)`))), + REWRITE_TAC[] THEN ONCE_REWRITE_TAC[SYM (REWRITE_CONV[NUMERAL] `0`)] THEN + REWRITE_TAC[num_exp; b0_thm] THEN + REWRITE_TAC[ARITH_RULE (mk_eq(mk_binop mul_op_num base_const `n * a:num`, + mk_binop mul_op_num `n:num` (mk_binop mul_op_num base_const `a:num`)))] THEN + REWRITE_TAC[GSYM EXP] THEN + SIMP_TAC[ARITH_RULE `~(e = 0) ==> SUC (PRE e) = e`]);; + +(* Converts num_exp n e to a numeral by adding e B0's *) +let rec denormalize tm = + let ltm, etm = dest_comb tm in + let ntm = rand ltm in + if (etm = zero_const) then + INST[ntm, n_var_num] SYM_NUM_EXP_0' + else + if ntm = zero_const then + INST[etm, e_var_num] NUM_EXP_n0' + else + let e_th = raw_eq0_hash_conv etm in + let th0' = INST[etm, e_var_num; ntm, n_var_num] NUM_EXP_DENORM in + let th0 = MY_PROVE_HYP e_th th0' in + let b0_tm, rtm = dest_comb(rand(concl th0)) in + let ltm, pre_tm = dest_comb rtm in + let pre_th = raw_pre_hash_conv pre_tm in + let th1 = AP_TERM ltm pre_th in + let th2 = denormalize (rand(concl th1)) in + TRANS th0 (AP_TERM b0_tm (TRANS th1 th2));; + +(***************************************) + +let rec comb_number tm n = + if (is_comb tm) then comb_number ((snd o dest_comb) tm) (n + 1) else n;; + +let make_lo_thm i = + let th_concl = mk_binop `(<=):num->num->bool` + (mk_comb (Arith_num.const_array.(0), n_var_num)) + (mk_comb (Arith_num.const_array.(i), n_var_num)) in + prove(th_concl, + REWRITE_TAC[Arith_num.def_array.(i); Arith_num.def_array.(0)] THEN + REWRITE_TAC[ARITH_LE; LE_REFL] THEN + ARITH_TAC);; + +let lo_thm_array = Array.init Arith_num.arith_base make_lo_thm;; +let lo_thm_table = Hashtbl.create Arith_num.arith_base;; + +for i = 0 to Arith_num.arith_base - 1 do + Hashtbl.add lo_thm_table Arith_num.const_array.(i) lo_thm_array.(i); +done;; + +let make_lo_thm2 i = + let th_concl = mk_imp (`n <= m:num`, + mk_binop `(<=):num->num->bool` + (mk_comb (Arith_num.const_array.(0), n_var_num)) + (mk_comb (Arith_num.const_array.(i), m_var_num))) in + (UNDISCH_ALL o prove) (th_concl, + REWRITE_TAC[Arith_num.def_array.(i); Arith_num.def_array.(0); ARITH_LE] THEN + ARITH_TAC);; + +let lo_thm2_array = Array.init Arith_num.arith_base make_lo_thm2;; +let lo_thm2_table = Hashtbl.create Arith_num.arith_base;; + +for i = 0 to Arith_num.arith_base - 1 do + Hashtbl.add lo_thm2_table Arith_num.const_array.(i) lo_thm2_array.(i); +done;; + + +let make_hi_thm i = + let th_concl = mk_imp (`n < m:num`, + mk_binop `(<):num->num->bool` + (mk_comb (Arith_num.const_array.(i), n_var_num)) + (mk_comb (Arith_num.const_array.(0), m_var_num))) in + (UNDISCH_ALL o prove) (th_concl, + REWRITE_TAC[Arith_num.def_array.(i); Arith_num.def_array.(0); ARITH_LT] THEN + ARITH_TAC);; + +let hi_thm_array = Array.init Arith_num.arith_base make_hi_thm;; +let hi_thm_table = Hashtbl.create Arith_num.arith_base;; + +for i = 0 to Arith_num.arith_base - 1 do + Hashtbl.add hi_thm_table Arith_num.const_array.(i) hi_thm_array.(i); +done;; + +(***************************************) + +let LE_REFL' = SPEC_ALL LE_REFL;; +let LE_TRANS' = (UNDISCH_ALL o SPEC_ALL o REWRITE_RULE[GSYM IMP_IMP]) LE_TRANS;; + +let lo_num_conv p tm = + let n = comb_number tm 0 in + if (n <= p) then + INST[tm, n_var_num] LE_REFL' + else + let rec lo_bound n tm = + let btm, rtm = dest_comb tm in + let th0 = INST[rtm, n_var_num] (Hashtbl.find lo_thm_table btm) in + if n > 1 then + let rth = lo_bound (n - 1) rtm in + let xtm = rand (rator (concl rth)) in + let th1' = INST[xtm, n_var_num; rtm, m_var_num] (Hashtbl.find lo_thm2_table btm) in + let th1 = MY_PROVE_HYP rth th1' in + th1 + else + th0 in + + lo_bound (n - p) tm;; + +let N_LT_SUC = ARITH_RULE `n < SUC n`;; +let LT_IMP_LE' = (UNDISCH_ALL o SPEC_ALL) LT_IMP_LE;; +let N_LT_SUC = ARITH_RULE `n < SUC n`;; +let LT_LE_TRANS = (UNDISCH_ALL o ARITH_RULE) `n < e ==> e <= m ==> n < m:num`;; + +(* Generates a theorem |- n <= m such that m contains at most p non-zero digits *) +let hi_num_conv p tm = + let n = comb_number tm 0 in + if (n <= p) then + INST[tm, n_var_num] LE_REFL' + else + let k = n - p in + + let rec check_b0s n tm = + let btm, rtm = dest_comb tm in + if ((fst o dest_const) btm = b0_name) then + if n > 1 then check_b0s (n - 1) rtm else true + else + false in + + if (check_b0s k tm) then + INST[tm, n_var_num] LE_REFL' + else + let rec hi_bound n tm = + if n > 0 then + let btm, rtm = dest_comb tm in + let r_th = hi_bound (n - 1) rtm in + let xtm = rand (concl r_th) in + let th0 = INST[rtm, n_var_num; xtm, m_var_num] (Hashtbl.find hi_thm_table btm) in + MY_PROVE_HYP r_th th0 + else + let th0 = INST[tm, n_var_num] N_LT_SUC in + let ltm, suc_tm = dest_comb (concl th0) in + let suc_th = raw_suc_conv_hash suc_tm in + EQ_MP (AP_TERM ltm suc_th) th0 in + + let th = hi_bound k tm in + let m_tm, l_tm = dest_comb (concl th) in + MY_PROVE_HYP th (INST[rand m_tm, m_var_num; l_tm, n_var_num] LT_IMP_LE');; + +(* Generates a theorem |- n < m such that m contains at most p non-zero digits *) +let hi_lt_num_conv p tm = + let n = comb_number tm 0 in + if (n <= p) then + let th0 = INST[tm, n_var_num] N_LT_SUC in + let ltm, rtm = dest_comb(concl th0) in + let suc_th = raw_suc_conv_hash rtm in + EQ_MP (AP_TERM ltm suc_th) th0 + else + let k = n - p in + + let rec check_b0s n tm = + let btm, rtm = dest_comb tm in + if ((fst o dest_const) btm = b0_name) then + if n > 1 then check_b0s (n - 1) rtm else true + else + false in + + if (check_b0s k tm) then + let th0 = INST[tm, n_var_num] N_LT_SUC in + let ltm, rtm = dest_comb (concl th0) in + let suc_th = raw_suc_conv_hash rtm in + let suc_tm = rand(concl suc_th) in + let th1 = hi_num_conv p suc_tm in + let th2 = EQ_MP (AP_TERM ltm suc_th) th0 in + let th = INST[tm, n_var_num; suc_tm, e_var_num; rand(concl th1), m_var_num] LT_LE_TRANS in + MY_PROVE_HYP th1 (MY_PROVE_HYP th2 th) + + else + let rec hi_bound n tm = + if n > 0 then + let btm, rtm = dest_comb tm in + let r_th = hi_bound (n - 1) rtm in + let xtm = rand (concl r_th) in + let th0 = INST[rtm, n_var_num; xtm, m_var_num] (Hashtbl.find hi_thm_table btm) in + MY_PROVE_HYP r_th th0 + else + let th0 = INST[tm, n_var_num] N_LT_SUC in + let ltm, suc_tm = dest_comb (concl th0) in + let suc_th = raw_suc_conv_hash suc_tm in + EQ_MP (AP_TERM ltm suc_th) th0 in + hi_bound k tm;; + +(*****************************************) + +let num_exp_lo p tm = + let ltm, e_tm = dest_comb tm in + let n_tm = rand ltm in + let n_th = lo_num_conv p n_tm in + let m_tm = rand (rator (concl n_th)) in + let m_norm, flag = normalize m_tm in + + let th0' = INST[m_tm, m_var_num; n_tm, n_var_num; e_tm, e_var_num] NUM_EXP_LE' in + let th0 = MY_PROVE_HYP n_th th0' in + if flag then + let th1 = AP_THM (AP_TERM (rator ltm) m_norm) e_tm in + let m_tm, me_tm = (dest_comb o rand o concl) m_norm in + let th2 = INST[me_tm, e1_var_num; e_tm, e2_var_num; rand m_tm, n_var_num] NUM_EXP_EXP' in + let th3 = TRANS th1 th2 in + let ltm, rtm = (dest_comb o rand o concl) th3 in + let th_add = raw_add_conv_hash rtm in + let th4 = TRANS th3 (AP_TERM ltm th_add) in + EQ_MP (AP_THM (AP_TERM le_op_num th4) tm) th0 + else + th0;; + +let num_exp_hi p tm = + let ltm, e_tm = dest_comb tm in + let n_tm = rand ltm in + let n_th = hi_num_conv p n_tm in + let m_tm = rand (concl n_th) in + let m_norm, flag = normalize m_tm in + + let th0' = INST[m_tm, n_var_num; n_tm, m_var_num; e_tm, e_var_num] NUM_EXP_LE' in + let th0 = MY_PROVE_HYP n_th th0' in + if flag then + let th1 = AP_THM (AP_TERM (rator ltm) m_norm) e_tm in + let m_tm, me_tm = (dest_comb o rand o concl) m_norm in + let th2 = INST[me_tm, e1_var_num; e_tm, e2_var_num; rand m_tm, n_var_num] NUM_EXP_EXP' in + let th3 = TRANS th1 th2 in + let ltm, rtm = (dest_comb o rand o concl) th3 in + let th_add = raw_add_conv_hash rtm in + let th4 = TRANS th3 (AP_TERM ltm th_add) in + EQ_MP (AP_TERM (rator (concl th0)) th4) th0 + else + th0;; + +let num_exp_hi_lt p tm = + let ltm, e_tm = dest_comb tm in + let n_tm = rand ltm in + let n_th = hi_lt_num_conv p n_tm in + let m_tm = rand (concl n_th) in + let m_norm, flag = normalize m_tm in + + let th0' = INST[m_tm, n_var_num; n_tm, m_var_num; e_tm, e_var_num] NUM_EXP_LT' in + let th0 = MY_PROVE_HYP n_th th0' in + if flag then + let th1 = AP_THM (AP_TERM (rator ltm) m_norm) e_tm in + let m_tm, me_tm = (dest_comb o rand o concl) m_norm in + let th2 = INST[me_tm, e1_var_num; e_tm, e2_var_num; rand m_tm, n_var_num] NUM_EXP_EXP' in + let th3 = TRANS th1 th2 in + let ltm, rtm = (dest_comb o rand o concl) th3 in + let th_add = raw_add_conv_hash rtm in + let th4 = TRANS th3 (AP_TERM ltm th_add) in + EQ_MP (AP_TERM (rator (concl th0)) th4) th0 + else + th0;; + +(***************************************) +(* num_exp_lt, num_exp_le *) + +let transform = UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL;; + +let NUM_EXP_LT1_EQ' = transform NUM_EXP_LT1_EQ and + NUM_EXP_LT2_EQ' = transform NUM_EXP_LT2_EQ;; + +let NUM_EXP_LE1_EQ' = transform NUM_EXP_LE1_EQ and + NUM_EXP_LE2_EQ' = transform NUM_EXP_LE2_EQ;; + +let num_exp_lt tm1 tm2 = + let n1_tm, e1_tm = dest_num_exp tm1 in + let n2_tm, e2_tm = dest_num_exp tm2 in + let sub_th, le_th = raw_sub_and_le_hash_conv e1_tm e2_tm in + let r_tm = rand(concl sub_th) in + if (rand(concl le_th) = e1_tm) then + let x_expr = mk_num_exp n1_tm r_tm in + let x_th = denormalize x_expr in + let x_tm = rand(concl x_th) in + + let th0 = INST[e2_tm, e2_var_num; e1_tm, e1_var_num; + r_tm, r_var_num; x_tm, x_var_num; + n1_tm, n1_var_num; n2_tm, n2_var_num] NUM_EXP_LT1_EQ' in + let th1 = MY_PROVE_HYP x_th (MY_PROVE_HYP sub_th (MY_PROVE_HYP le_th th0)) in + let lt_th = raw_lt_hash_conv (rand(concl th1)) in + TRANS th1 lt_th + else + let x_expr = mk_num_exp n2_tm r_tm in + let x_th = denormalize x_expr in + let x_tm = rand(concl x_th) in + + let th0 = INST[e2_tm, e2_var_num; e1_tm, e1_var_num; + r_tm, r_var_num; x_tm, x_var_num; + n1_tm, n1_var_num; n2_tm, n2_var_num] NUM_EXP_LT2_EQ' in + let th1 = MY_PROVE_HYP x_th (MY_PROVE_HYP sub_th (MY_PROVE_HYP le_th th0)) in + let lt_th = raw_lt_hash_conv (rand(concl th1)) in + TRANS th1 lt_th;; + + +let num_exp_le tm1 tm2 = + let n1_tm, e1_tm = dest_num_exp tm1 in + let n2_tm, e2_tm = dest_num_exp tm2 in + let sub_th, le_th = raw_sub_and_le_hash_conv e1_tm e2_tm in + let r_tm = rand(concl sub_th) in + if (rand(concl le_th) = e1_tm) then + let x_expr = mk_num_exp n1_tm r_tm in + let x_th = denormalize x_expr in + let x_tm = rand(concl x_th) in + + let th0 = INST[e2_tm, e2_var_num; e1_tm, e1_var_num; + r_tm, r_var_num; x_tm, x_var_num; + n1_tm, n1_var_num; n2_tm, n2_var_num] NUM_EXP_LE1_EQ' in + let th1 = MY_PROVE_HYP x_th (MY_PROVE_HYP sub_th (MY_PROVE_HYP le_th th0)) in + let le_th = raw_le_hash_conv (rand(concl th1)) in + TRANS th1 le_th + else + let x_expr = mk_num_exp n2_tm r_tm in + let x_th = denormalize x_expr in + let x_tm = rand(concl x_th) in + + let th0 = INST[e2_tm, e2_var_num; e1_tm, e1_var_num; + r_tm, r_var_num; x_tm, x_var_num; + n1_tm, n1_var_num; n2_tm, n2_var_num] NUM_EXP_LE2_EQ' in + let th1 = MY_PROVE_HYP x_th (MY_PROVE_HYP sub_th (MY_PROVE_HYP le_th th0)) in + let le_th = raw_le_hash_conv (rand(concl th1)) in + TRANS th1 le_th;; + +(***************************************) +(* num_exp_mul *) + +let NUM_EXP_MUL' = SPEC_ALL NUM_EXP_MUL;; + +let num_exp_mul tm1 tm2 = + let n1_tm, e1_tm = dest_comb tm1 in + let n1_tm = rand n1_tm in + let n2_tm, e2_tm = dest_comb tm2 in + let n2_tm = rand n2_tm in + let th0 = INST[n1_tm, n1_var_num; e1_tm, e1_var_num; + n2_tm, n2_var_num; e2_tm, e2_var_num] NUM_EXP_MUL' in + let ltm, tm_add = dest_comb (rand (concl th0)) in + let tm_mul = rand ltm in + let th_mul = raw_mul_conv_hash tm_mul in + let th_add = raw_add_conv_hash tm_add in + TRANS th0 (MK_COMB (AP_TERM (rator ltm) th_mul, th_add));; + + +(**********************************) +(* num_exp_add *) + +let NUM_EXP_ADD' = (UNDISCH_ALL o SPEC_ALL) NUM_EXP_ADD;; +let ADD_COMM = ARITH_RULE `m + n = n + m:num`;; + +let num_exp_add tm1 tm2 = + let n1_tm, e1_tm = dest_comb tm1 in + let n1_tm = rand n1_tm in + let n2_tm, e2_tm = dest_comb tm2 in + let n2_tm = rand n2_tm in + let e_sub, e_le = raw_sub_and_le_hash_conv e1_tm e2_tm in + + let flag = (rand(concl e_le) = e2_tm) in + + let th0' = + if flag then + INST[n1_tm, n1_var_num; e1_tm, e1_var_num; + n2_tm, n2_var_num; e2_tm, e2_var_num] NUM_EXP_ADD' + else + INST[n2_tm, n1_var_num; e2_tm, e1_var_num; + n1_tm, n2_var_num; e1_tm, e2_var_num] NUM_EXP_ADD' in + + let th0 = MY_PROVE_HYP e_le th0' in + let ltm, e0_tm = dest_comb(rand(concl th0)) in + let exp_tm, add_tm = dest_comb ltm in + let ltm, d_tm = dest_comb add_tm in + let th1 = AP_TERM (rator d_tm) e_sub in + let th2 = denormalize (rand(concl th1)) in + let th3 = AP_TERM ltm (TRANS th1 th2) in + let th4 = raw_add_conv_hash (rand(concl th3)) in + let th5 = AP_THM (AP_TERM exp_tm (TRANS th3 th4)) e0_tm in + let th = TRANS th0 th5 in + if flag then th else + TRANS (INST[tm1, m_var_num; tm2, n_var_num] ADD_COMM) th;; + +(****************************************) +(* num_exp_sub *) + +let NUM_EXP_SUB1' = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL) NUM_EXP_SUB1 and + NUM_EXP_SUB2' = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL) NUM_EXP_SUB2 and + NUM_EXP_LE1' = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL) NUM_EXP_LE1 and + NUM_EXP_LE2' = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL) NUM_EXP_LE2;; + +(* Returns two theorems: |- tm1 - tm2 = tm, |- tm2 <= tm1 or + |- tm2 - tm1 = tm, |- tm1 <= tm2 *) +let num_exp_sub tm1 tm2 = + let n1_tm, e1_tm = dest_num_exp tm1 in + let n2_tm, e2_tm = dest_num_exp tm2 in + let e_sub, e_le = raw_sub_and_le_hash_conv e1_tm e2_tm in + + if rand(concl e_le) = e1_tm then + (* e2 <= e1 *) + let e1_sub_e2 = rand(concl e_sub) in + let a0 = mk_num_exp n1_tm e1_sub_e2 in + let b = n2_tm in + let a_th = denormalize a0 in + let a = rand(concl a_th) in + + let th_sub, th_le = raw_sub_and_le_hash_conv a b in + if rand(concl th_le) = a then + (* b <= a *) + let a_sub_b = TRANS (AP_THM (AP_TERM sub_op_num a_th) b) th_sub in + let b_le_a = EQ_MP (SYM (AP_TERM (rator(concl th_le)) a_th)) th_le in + let th0 = AP_THM (AP_TERM num_exp_const a_sub_b) e2_tm in + + let inst = INST[n1_tm, n1_var_num; e1_tm, e1_var_num; + n2_tm, n2_var_num; e2_tm, e2_var_num; e1_sub_e2, r_var_num] in + let th1_sub = inst NUM_EXP_SUB1' in + let th1_le = inst NUM_EXP_LE1' in + let th2_sub = MY_PROVE_HYP e_sub (MY_PROVE_HYP e_le th1_sub) in + let th2_le = MY_PROVE_HYP e_sub (MY_PROVE_HYP b_le_a (MY_PROVE_HYP e_le th1_le)) in + TRANS th2_sub th0, th2_le + + else + (* a <= b *) + let b_sub_a = TRANS (AP_TERM (rator(lhand(concl th_sub))) a_th) th_sub in + let a_le_b = EQ_MP (SYM (AP_THM (AP_TERM le_op_num a_th) b)) th_le in + let th0 = AP_THM (AP_TERM num_exp_const b_sub_a) e2_tm in + let inst = INST[n2_tm, n1_var_num; e2_tm, e1_var_num; + n1_tm, n2_var_num; e1_tm, e2_var_num; e1_sub_e2, r_var_num] in + let th1_sub = inst NUM_EXP_SUB2' in + let th1_le = inst NUM_EXP_LE2' in + let th2_sub = MY_PROVE_HYP e_sub (MY_PROVE_HYP e_le th1_sub) in + let th2_le = MY_PROVE_HYP e_sub (MY_PROVE_HYP a_le_b (MY_PROVE_HYP e_le th1_le)) in + TRANS th2_sub th0, th2_le + + else + (* e1 <= e2 *) + let e2_sub_e1 = rand(concl e_sub) in + let b0 = mk_num_exp n2_tm e2_sub_e1 in + let a = n1_tm in + let b_th = denormalize b0 in + let b = rand(concl b_th) in + + let th_sub, th_le = raw_sub_and_le_hash_conv a b in + if rand(concl th_le) = a then + (* b <= a *) + let a_sub_b = TRANS (AP_TERM (rator(lhand(concl th_sub))) b_th) th_sub in + let b_le_a = EQ_MP (SYM (AP_THM (AP_TERM le_op_num b_th) a)) th_le in + let th0 = AP_THM (AP_TERM num_exp_const a_sub_b) e1_tm in + let inst = INST[n1_tm, n1_var_num; e1_tm, e1_var_num; + n2_tm, n2_var_num; e2_tm, e2_var_num; e2_sub_e1, r_var_num] in + let th1_sub = inst NUM_EXP_SUB2' in + let th1_le = inst NUM_EXP_LE2' in + let th2_sub = MY_PROVE_HYP e_sub (MY_PROVE_HYP e_le th1_sub) in + let th2_le = MY_PROVE_HYP e_sub (MY_PROVE_HYP b_le_a (MY_PROVE_HYP e_le th1_le)) in + TRANS th2_sub th0, th2_le + + else + (* a <= b *) + let b_sub_a = TRANS (AP_THM (AP_TERM sub_op_num b_th) a) th_sub in + let a_le_b = EQ_MP (SYM (AP_TERM (rator(concl th_le)) b_th)) th_le in + let th0 = AP_THM (AP_TERM num_exp_const b_sub_a) e1_tm in + let inst = INST[n2_tm, n1_var_num; e2_tm, e1_var_num; + n1_tm, n2_var_num; e1_tm, e2_var_num; e2_sub_e1, r_var_num] in + let th1_sub = inst NUM_EXP_SUB1' in + let th1_le = inst NUM_EXP_LE1' in + let th2_sub = MY_PROVE_HYP e_sub (MY_PROVE_HYP e_le th1_sub) in + let th2_le = MY_PROVE_HYP e_sub (MY_PROVE_HYP a_le_b (MY_PROVE_HYP e_le th1_le)) in + TRANS th2_sub th0, th2_le;; + + +(*************************************) +(* division *) + +let NUM_EXP_DIV1' = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o + PURE_ONCE_REWRITE_RULE[ARITH_RULE `~(x = 0) <=> (x = 0 <=> F)`] o + REWRITE_RULE[GSYM IMP_IMP]) NUM_EXP_DIV1;; +let NUM_EXP_DIV2' = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o + PURE_ONCE_REWRITE_RULE[ARITH_RULE `~(x = 0) <=> (x = 0 <=> F)`] o + REWRITE_RULE[GSYM IMP_IMP]) NUM_EXP_DIV2;; + +let num_exp_div tm1 tm2 = + let n1_tm, e1_tm = dest_comb tm1 in + let n1_tm = rand n1_tm in + let n2_tm, e2_tm = dest_comb tm2 in + let n2_tm = rand n2_tm in + let e_sub, e_le = raw_sub_and_le_hash_conv e1_tm e2_tm in + + let inst = INST[n1_tm, n1_var_num; e1_tm, e1_var_num; + n2_tm, n2_var_num; e2_tm, e2_var_num] in + + let n2_not_0 = raw_eq0_hash_conv n2_tm in + if ((fst o dest_const o rand o concl) n2_not_0 = "T") then + failwith "num_exp_div: n2 = 0" + else + if (rand(concl e_le) = e1_tm) then + let th0' = inst NUM_EXP_DIV1' in + let th0 = MY_PROVE_HYP n2_not_0 (MY_PROVE_HYP e_le th0') in + + let ltm, rtm = dest_comb(rand(concl th0)) in + let div_tm, rtm2 = dest_comb ltm in + let num_exp_tm = rator rtm2 in + + let th1 = AP_THM (AP_TERM div_tm (AP_TERM num_exp_tm e_sub)) rtm in + let ltm, rtm = dest_comb(rand(concl th1)) in + let tm1 = rand ltm in + + let th2 = AP_THM (AP_TERM div_tm (denormalize tm1)) rtm in + let th3 = raw_div_hash_conv (rand(concl th2)) in + let th = TRANS th0 (TRANS th1 (TRANS th2 th3)) in + TRANS th (INST[rand(concl th), n_var_num] NUM_EXP_0') + + else + let th0' = inst NUM_EXP_DIV2' in + let th0 = MY_PROVE_HYP n2_not_0 (MY_PROVE_HYP e_le th0') in + + let ltm, rtm = dest_comb(rand(concl th0)) in + let num_exp_tm = rator rtm in + let th1 = AP_TERM ltm (AP_TERM num_exp_tm e_sub) in + + let ltm, rtm = dest_comb(rand(concl th1)) in + let th2 = AP_TERM ltm (denormalize rtm) in + let th3 = raw_div_hash_conv (rand(concl th2)) in + let th = TRANS th0 (TRANS th1 (TRANS th2 th3)) in + TRANS th (INST[rand(concl th), n_var_num] NUM_EXP_0');; + + +(*****************************) +(* Computes a lower bound for (op tm1 tm2) with p significant digits *) +let num_exp_op_lo p op tm1 tm2 = + let op_th = op tm1 tm2 in + let rtm = rand (concl op_th) in + let lo_th = num_exp_lo p rtm in + let ltm = rator (concl lo_th) in + let th0 = AP_TERM ltm op_th in + EQ_MP (SYM th0) lo_th;; + +(* Computes an upper bound for (op tm1 tm2) with p significant digits *) +let num_exp_op_hi p op tm1 tm2 = + let op_th = op tm1 tm2 in + let rtm = rand (concl op_th) in + let hi_th = num_exp_hi p rtm in + let tm = rand (concl hi_th) in + let th0 = AP_THM (AP_TERM le_op_num op_th) tm in + EQ_MP (SYM th0) hi_th;; + +(* Computes a strict upper bound for (op tm1 tm2) with p significant digits *) +let num_exp_op_hi_lt p op tm1 tm2 = + let op_th = op tm1 tm2 in + let rtm = rand (concl op_th) in + let hi_lt_th = num_exp_hi_lt p rtm in + let tm = rand (concl hi_lt_th) in + let th0 = AP_THM (AP_TERM lt_op_num op_th) tm in + EQ_MP (SYM th0) hi_lt_th;; + +(******************************************) +(* float *) + +let mod_plus = new_definition `mod_plus s1 s2 = (~(s1 /\ s2) /\ (s1 \/ s2))`;; + +(********************) +(* Float operations *) +(********************) + +module Float_ops = struct + +(**********************************) + +(* FLOAT_LT *) + +let FLOAT_LT_FF = prove(`float_num F n1 e1 < float_num F n2 e2 <=> num_exp n1 e1 < num_exp n2 e2`, + REWRITE_TAC[float; GSYM REAL_OF_NUM_LT; REAL_MUL_LID; real_div] THEN + MATCH_MP_TAC REAL_LT_RMUL_EQ THEN + MATCH_MP_TAC REAL_LT_INV THEN + REWRITE_TAC[REAL_OF_NUM_LT; LT_NZ; NUM_EXP_EQ_0] THEN + ARITH_TAC);; + +let FLOAT_LT_TT = prove(`float_num T n1 e1 < float_num T n2 e2 <=> num_exp n2 e2 < num_exp n1 e1`, + REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `--a < --b <=> b < a`] THEN + REWRITE_TAC[FLOAT_LT_FF]);; + +let FLOAT_LT_FT = prove(`float_num F n1 e1 < float_num T n2 e2 <=> F`, + MP_TAC (SPECL [`n2:num`; `e2:num`] FLOAT_T_NEG) THEN + MP_TAC (SPECL [`n1:num`; `e1:num`] FLOAT_F_POS) THEN + REAL_ARITH_TAC);; + +let FLOAT_LT_TF_00 = (PURE_REWRITE_RULE[NUMERAL] o prove) + (`float_num T 0 e1 < float_num F 0 e2 <=> F`, + MP_TAC (SPECL [`T`; `0`; `e1:num`] FLOAT_EQ_0) THEN + MP_TAC (SPECL [`F`; `0`; `e2:num`] FLOAT_EQ_0) THEN + REWRITE_TAC[] THEN + REPLICATE_TAC 2 (DISCH_THEN (fun th -> REWRITE_TAC[th])) THEN + REAL_ARITH_TAC);; + +let FLOAT_LT_TF_1 = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o prove) + (`(n1 = 0 <=> F) ==> (float_num T n1 e1 < float_num F n2 e2 <=> T)`, + DISCH_TAC THEN + MATCH_MP_TAC (REAL_ARITH `a < &0 /\ &0 <= b ==> (a < b <=> T)`) THEN + REWRITE_TAC[FLOAT_F_POS] THEN + MATCH_MP_TAC (REAL_ARITH `~(a = &0) /\ a <= &0 ==> a < &0`) THEN + ASM_REWRITE_TAC[FLOAT_T_NEG; FLOAT_EQ_0]);; + +let FLOAT_LT_TF_2 = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o prove) + (`(n2 = 0 <=> F) ==> (float_num T n1 e1 < float_num F n2 e2 <=> T)`, + DISCH_TAC THEN + MATCH_MP_TAC (REAL_ARITH `a <= &0 /\ &0 < b ==> (a < b <=> T)`) THEN + REWRITE_TAC[FLOAT_T_NEG] THEN + MATCH_MP_TAC (REAL_ARITH `~(a = &0) /\ &0 <= a ==> &0 < a`) THEN + ASM_REWRITE_TAC[FLOAT_F_POS; FLOAT_EQ_0]);; + +let FLOAT_F_LT_0 = prove(`float_num F n e < &0 <=> F`, + MP_TAC (SPEC_ALL FLOAT_F_POS) THEN + REAL_ARITH_TAC);; + +let FLOAT_T_LT_0 = (CONV_RULE (RAND_CONV (REWRITE_CONV[NUMERAL])) o prove) + (`float_num T n e < &0 <=> (0 < n)`, + REWRITE_TAC[REAL_ARITH `a < &0 <=> a <= &0 /\ ~(a = &0)`] THEN + REWRITE_TAC[FLOAT_T_NEG; FLOAT_EQ_0] THEN + ARITH_TAC);; + +let FLOAT_F_GT_0 = (CONV_RULE (RAND_CONV (REWRITE_CONV[NUMERAL])) o prove) + (`&0 < float_num F n e <=> 0 < n`, + REWRITE_TAC[REAL_ARITH `&0 < a <=> &0 <= a /\ ~(a = &0)`] THEN + REWRITE_TAC[FLOAT_F_POS; FLOAT_EQ_0] THEN + ARITH_TAC);; + +let FLOAT_T_GT_0 = prove(`&0 < float_num T n e <=> F`, + MP_TAC (SPEC_ALL FLOAT_T_NEG) THEN + REAL_ARITH_TAC);; + + +(* float_lt0, float_gt0 *) + +let float_lt0 f1 = + let s, n_tm, e_tm = dest_float f1 in + let inst = INST[n_tm, n_var_num; e_tm, e_var_num] in + if s = "F" then + inst FLOAT_F_LT_0 + else + let gt_th = raw_gt0_hash_conv n_tm in + TRANS (inst FLOAT_T_LT_0) gt_th;; + +let float_gt0 f1 = + let s, n_tm, e_tm = dest_float f1 in + let inst = INST[n_tm, n_var_num; e_tm, e_var_num] in + if s = "F" then + let gt_th = raw_gt0_hash_conv n_tm in + TRANS (inst FLOAT_F_GT_0) gt_th + else + inst FLOAT_T_GT_0;; + +(* float_lt *) +let float_lt f1 f2 = + let s1, n1, e1 = dest_float f1 in + let s2, n2, e2 = dest_float f2 in + let inst = INST[n1, n1_var_num; e1, e1_var_num; + n2, n2_var_num; e2, e2_var_num] in + if s1 = "F" then + if s2 = "F" then + (* FF *) + let th0 = inst FLOAT_LT_FF in + let ltm, tm2 = dest_comb (rand (concl th0)) in + let lt_th = num_exp_lt (rand ltm) tm2 in + TRANS th0 lt_th + else + (* FT *) + inst FLOAT_LT_FT + else + if s2 = "F" then + (* TF *) + if (is_const n1 && is_const n2) then + (* n1 = _0 and n2 = _0 *) + inst FLOAT_LT_TF_00 + else + let n1_0 = raw_eq0_hash_conv n1 in + if (fst o dest_const o rand o concl) n1_0 = "F" then + (* n1 <> _0 *) + MY_PROVE_HYP n1_0 (inst FLOAT_LT_TF_1) + else + let n2_0 = raw_eq0_hash_conv n2 in + if (fst o dest_const o rand o concl) n2_0 = "F" then + (* n2 <> _0 *) + MY_PROVE_HYP n2_0 (inst FLOAT_LT_TF_2) + else + failwith "float_lt: D0 _0 exception" + else + (* TT *) + let th0 = inst FLOAT_LT_TT in + let ltm, tm2 = dest_comb (rand (concl th0)) in + let lt_th = num_exp_lt (rand ltm) tm2 in + TRANS th0 lt_th;; + +(**********************************) +(* FLOAT_LE *) + +let FLOAT_LE_FF = prove(`float_num F n1 e1 <= float_num F n2 e2 <=> num_exp n1 e1 <= num_exp n2 e2`, + REWRITE_TAC[float; GSYM REAL_OF_NUM_LE; REAL_MUL_LID; real_div] THEN + MATCH_MP_TAC REAL_LE_RMUL_EQ THEN + MATCH_MP_TAC REAL_LT_INV THEN + REWRITE_TAC[REAL_OF_NUM_LT; LT_NZ; NUM_EXP_EQ_0] THEN + ARITH_TAC);; + +let FLOAT_LE_TT = prove(`float_num T n1 e1 <= float_num T n2 e2 <=> num_exp n2 e2 <= num_exp n1 e1`, + REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `--a <= --b <=> b <= a`] THEN + REWRITE_TAC[FLOAT_LE_FF]);; + +let FLOAT_LE_TF = prove(`float_num T n1 e1 <= float_num F n2 e2 <=> T`, + MP_TAC (SPECL [`n1:num`; `e1:num`] FLOAT_T_NEG) THEN + MP_TAC (SPECL [`n2:num`; `e2:num`] FLOAT_F_POS) THEN + REAL_ARITH_TAC);; + +let FLOAT_LE_FT = prove(`float_num F n1 e1 <= float_num T n2 e2 <=> n1 = 0 /\ n2 = 0`, + REWRITE_TAC[REAL_LE_LT; FLOAT_LT_FT] THEN EQ_TAC THENL + [ + DISCH_TAC THEN SUBGOAL_THEN `float_num F n1 e1 = &0 /\ float_num T n2 e2 = &0` MP_TAC THENL + [ + MP_TAC (SPECL [`n2:num`; `e2:num`] FLOAT_T_NEG) THEN + MP_TAC (SPECL [`n1:num`; `e1:num`] FLOAT_F_POS) THEN + ASM_REWRITE_TAC[] THEN REAL_ARITH_TAC; + ALL_TAC + ] THEN + REWRITE_TAC[FLOAT_EQ_0]; + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + REWRITE_TAC[float; NUM_EXP_n0; real_div; REAL_MUL_LZERO; REAL_MUL_RZERO] + ]);; + +let FLOAT_LE_FT_00 = (PURE_REWRITE_RULE[NUMERAL] o prove) + (`float_num F 0 e1 <= float_num T 0 e2 <=> T`, REWRITE_TAC[FLOAT_LE_FT]);; + +let FLOAT_LE_FT_1 = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o prove) + (`(n1 = 0 <=> F) ==> (float_num F n1 e1 <= float_num T n2 e2 <=> F)`, + DISCH_TAC THEN ASM_REWRITE_TAC[FLOAT_LE_FT]);; + +let FLOAT_LE_FT_2 = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o prove) + (`(n2 = 0 <=> F) ==> (float_num F n1 e1 <= float_num T n2 e2 <=> F)`, + DISCH_TAC THEN ASM_REWRITE_TAC[FLOAT_LE_FT]);; + +let FLOAT_F_LE_0 = (CONV_RULE (RAND_CONV (REWRITE_CONV[NUMERAL])) o prove) + (`float_num F n e <= &0 <=> n = 0`, + REWRITE_TAC[GSYM (SPEC `F` FLOAT_EQ_0)] THEN + MP_TAC (SPEC_ALL FLOAT_F_POS) THEN + REAL_ARITH_TAC);; + +let FLOAT_T_LE_0 = prove(`float_num T n e <= &0 <=> T`, REWRITE_TAC[FLOAT_T_NEG]);; + +let FLOAT_F_GE_0 = prove(`&0 <= float_num F n e <=> T`, REWRITE_TAC[FLOAT_F_POS]);; + +let FLOAT_T_GE_0 = (CONV_RULE (RAND_CONV (REWRITE_CONV[NUMERAL])) o prove) + (`&0 <= float_num T n e <=> n = 0`, + REWRITE_TAC[GSYM (SPEC `T` FLOAT_EQ_0)] THEN + MP_TAC (SPEC_ALL FLOAT_T_NEG) THEN + REAL_ARITH_TAC);; + +(* float_le0, float_ge0 *) +let float_le0 f1 = + let s, n_tm, e_tm = dest_float f1 in + let inst = INST[n_tm, n_var_num; e_tm, e_var_num] in + if s = "T" then + inst FLOAT_T_LE_0 + else + let eq_th = raw_eq0_hash_conv n_tm in + TRANS (inst FLOAT_F_LE_0) eq_th;; + +let float_ge0 f1 = + let s, n_tm, e_tm = dest_float f1 in + let inst = INST[n_tm, n_var_num; e_tm, e_var_num] in + if s = "T" then + let eq_th = raw_eq0_hash_conv n_tm in + TRANS (inst FLOAT_T_GE_0) eq_th + else + inst FLOAT_F_GE_0;; + +(* float_le *) +let float_le f1 f2 = + let s1, n1, e1 = dest_float f1 in + let s2, n2, e2 = dest_float f2 in + let inst = INST[n1, n1_var_num; e1, e1_var_num; + n2, n2_var_num; e2, e2_var_num] in + if s2 = "F" then + if s1 = "F" then + (* FF *) + let th0 = inst FLOAT_LE_FF in + let ltm, tm2 = dest_comb (rand (concl th0)) in + let le_th = num_exp_le (rand ltm) tm2 in + TRANS th0 le_th + else + (* TF *) + inst FLOAT_LE_TF + else + if s1 = "F" then + (* FT *) + if (is_const n1 && is_const n2) then + (* n1 = _0 and n2 = _0 *) + inst FLOAT_LE_FT_00 + else + let n1_0 = raw_eq0_hash_conv n1 in + if (fst o dest_const o rand o concl) n1_0 = "F" then + (* n1 <> _0 *) + MY_PROVE_HYP n1_0 (inst FLOAT_LE_FT_1) + else + let n2_0 = raw_eq0_hash_conv n2 in + if (fst o dest_const o rand o concl) n2_0 = "F" then + (* n2 <> _0 *) + MY_PROVE_HYP n2_0 (inst FLOAT_LE_FT_2) + else + failwith "float_lt: D0 _0 exception" + else + (* TT *) + let th0 = inst FLOAT_LE_TT in + let ltm, tm2 = dest_comb (rand (concl th0)) in + let le_th = num_exp_le (rand ltm) tm2 in + TRANS th0 le_th;; + + +(*************************************) +(* float_max, float_min *) + +let FLOAT_MIN_1 = (UNDISCH_ALL o prove)(`(f1 <= f2 <=> T) ==> min f1 f2 = f1`, REAL_ARITH_TAC);; +let FLOAT_MIN_2 = (UNDISCH_ALL o prove)(`(f1 <= f2 <=> F) ==> min f1 f2 = f2`, REAL_ARITH_TAC);; + +let FLOAT_MAX_1 = (UNDISCH_ALL o prove)(`(f1 <= f2 <=> T) ==> max f1 f2 = f2`, REAL_ARITH_TAC);; +let FLOAT_MAX_2 = (UNDISCH_ALL o prove)(`(f1 <= f2 <=> F) ==> max f1 f2 = f1`, REAL_ARITH_TAC);; + +let float_min f1 f2 = + let inst = INST[f1, f1_var_real; f2, f2_var_real] in + let le_th = float_le f1 f2 in + let th0 = + if (fst o dest_const o rand o concl) le_th = "T" then + inst FLOAT_MIN_1 + else + inst FLOAT_MIN_2 in + MY_PROVE_HYP le_th th0;; + +let float_max f1 f2 = + let inst = INST[f1, f1_var_real; f2, f2_var_real] in + let le_th = float_le f1 f2 in + let th0 = + if (fst o dest_const o rand o concl) le_th = "T" then + inst FLOAT_MAX_1 + else + inst FLOAT_MAX_2 in + MY_PROVE_HYP le_th th0;; + +let float_min_max f1 f2 = + let inst = INST[f1, f1_var_real; f2, f2_var_real] in + let le_th = float_le f1 f2 in + let th_min, th_max = + if (fst o dest_const o rand o concl) le_th = "T" then + inst FLOAT_MIN_1, inst FLOAT_MAX_1 + else + inst FLOAT_MIN_2, inst FLOAT_MAX_2 in + MY_PROVE_HYP le_th th_min, MY_PROVE_HYP le_th th_max;; + + +(*************************************) +(* FLOAT_MUL *) + +let FLOAT_MUL = prove(`!s1 s2. min_exp <= e /\ num_exp n1 e1 * num_exp n2 e2 = num_exp n e + ==> float_num s1 n1 e1 * float_num s2 n2 e2 = + float_num (mod_plus s1 s2) n (e - min_exp)`, + REPEAT STRIP_TAC THEN + REWRITE_TAC[float] THEN + ONCE_REWRITE_TAC[REAL_ARITH `(a * b / c) * (d * e / f) = (a * d) * (b * e) / c / f`] THEN + + SUBGOAL_THEN `(if s1 then -- &1 else &1) * (if s2 then -- &1 else &1) = if mod_plus s1 s2 then -- &1 else &1` MP_TAC THENL + [ + REWRITE_TAC[mod_plus] THEN + COND_CASES_TAC THEN COND_CASES_TAC THEN + REWRITE_TAC[REAL_ARITH `-- &1 * -- &1 = &1`; REAL_MUL_LID; REAL_MUL_RID]; + ALL_TAC + ] THEN + + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + REWRITE_TAC[real_div] THEN + REWRITE_TAC[GSYM REAL_MUL_ASSOC] THEN + REWRITE_TAC[REAL_EQ_MUL_LCANCEL] THEN + DISJ2_TAC THEN + + MP_TAC (SPECL[`n:num`; `e:num`; `min_exp`] NUM_EXP_SUB_lemma) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + REWRITE_TAC[REAL_MUL_ASSOC] THEN + ASM_REWRITE_TAC[REAL_OF_NUM_MUL]);; + +let FLOAT_MUL_FF = prove(`min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 = num_exp n e ==> + float_num F n1 e1 * float_num F n2 e2 = float_num F n r`, + SIMP_TAC[FLOAT_MUL; mod_plus]);; +let FLOAT_MUL_FT = prove(`min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 = num_exp n e ==> + float_num F n1 e1 * float_num T n2 e2 = float_num T n r`, + SIMP_TAC[FLOAT_MUL; mod_plus]);; +let FLOAT_MUL_TF = prove(`min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 = num_exp n e ==> + float_num T n1 e1 * float_num F n2 e2 = float_num T n r`, + SIMP_TAC[FLOAT_MUL; mod_plus]);; +let FLOAT_MUL_TT = prove(`min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 = num_exp n e ==> + float_num T n1 e1 * float_num T n2 e2 = float_num F n r`, + SIMP_TAC[FLOAT_MUL; mod_plus]);; + + +let FLOAT_MUL_0x_hi, FLOAT_MUL_0x_lo, FLOAT_MUL_x0_hi, FLOAT_MUL_x0_lo = + let mul_0x_hi = `(n1 = 0 <=> T) ==> float_num s1 n1 e1 * f2 <= float_num F 0 min_exp` in + let mul_0x_lo = `(n1 = 0 <=> T) ==> float_num F 0 min_exp <= float_num s1 n1 e1 * f2` in + let mul_x0_hi = `(n2 = 0 <=> T) ==> f1 * float_num s2 n2 e2 <= float_num F 0 min_exp` in + let mul_x0_lo = `(n2 = 0 <=> T) ==> float_num F 0 min_exp <= f1 * float_num s2 n2 e2` in + let proof = MP_TAC (GEN_ALL (SPECL [`s:bool`; `0`] FLOAT_EQ_0)) THEN + SIMP_TAC[REAL_MUL_LZERO; REAL_MUL_RZERO; REAL_LE_REFL] in + prove(mul_0x_hi, proof), prove(mul_0x_lo, proof), + prove(mul_x0_hi, proof), prove(mul_x0_lo, proof);; + + +let FLOAT_MUL_FF_hi, FLOAT_MUL_FF_lo = + let ff_hi = `min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 <= num_exp n e + ==> float_num F n1 e1 * float_num F n2 e2 <= float_num F n r` in + let ff_lo = `min_exp <= e /\ e - min_exp = r /\ num_exp n e <= num_exp n1 e1 * num_exp n2 e2 + ==> float_num F n r <= float_num F n1 e1 * float_num F n2 e2` in + let proof = + REPEAT STRIP_TAC THEN + POP_ASSUM MP_TAC THEN + POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN + REWRITE_TAC[GSYM REAL_OF_NUM_LE; GSYM REAL_OF_NUM_MUL] THEN + DISCH_TAC THEN + MAP_EVERY ABBREV_TAC [`z = &(num_exp n e)`; `x = &(num_exp n1 e1)`; `y = &(num_exp n2 e2)`] THEN + ASM_REWRITE_TAC[float; REAL_MUL_LID] THEN + REWRITE_TAC[REAL_ARITH `a / b * c / d = (a * c) / b / d`] THEN + REWRITE_TAC[real_div] THEN + REWRITE_TAC[REAL_MUL_ASSOC] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN + + MP_TAC (SPECL [`n:num`; `e:num`; `min_exp`] NUM_EXP_SUB_lemma) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + ASM_REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] in + prove(ff_hi, proof), prove(ff_lo, proof);; + +let FLOAT_MUL_TT_hi, FLOAT_MUL_TT_lo = + let tt_hi = `min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 <= num_exp n e + ==> float_num T n1 e1 * float_num T n2 e2 <= float_num F n r` in + let tt_lo = `min_exp <= e /\ e - min_exp = r /\ num_exp n e <= num_exp n1 e1 * num_exp n2 e2 + ==> float_num F n r <= float_num T n1 e1 * float_num T n2 e2` in + let proof = + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[REAL_ARITH `--a * --b = a * b`] THEN + REWRITE_TAC[FLOAT_MUL_FF_hi; FLOAT_MUL_FF_lo] in + prove(tt_hi, proof), prove(tt_lo, proof);; + +let FLOAT_MUL_FT_hi, FLOAT_MUL_FT_lo = + let ft_hi = `min_exp <= e /\ e - min_exp = r /\ num_exp n e <= num_exp n1 e1 * num_exp n2 e2 + ==> float_num F n1 e1 * float_num T n2 e2 <= float_num T n r` in + let ft_lo = `min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 <= num_exp n e + ==> float_num T n r <= float_num F n1 e1 * float_num T n2 e2` in + let proof = + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[REAL_ARITH `a * --b <= --c <=> c <= a * b`] THEN + REWRITE_TAC[REAL_ARITH `--c <= a * --b <=> a * b <= c`] THEN + REWRITE_TAC[FLOAT_MUL_FF_hi; FLOAT_MUL_FF_lo] in + prove(ft_hi, proof), prove(ft_lo, proof);; + +let FLOAT_MUL_TF_hi, FLOAT_MUL_TF_lo = + let ft_hi = `min_exp <= e /\ e - min_exp = r /\ num_exp n e <= num_exp n1 e1 * num_exp n2 e2 + ==> float_num T n1 e1 * float_num F n2 e2 <= float_num T n r` in + let ft_lo = `min_exp <= e /\ e - min_exp = r /\ num_exp n1 e1 * num_exp n2 e2 <= num_exp n e + ==> float_num T n r <= float_num T n1 e1 * float_num F n2 e2` in + let proof = + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[REAL_ARITH `--a * b <= --c <=> c <= a * b`] THEN + REWRITE_TAC[REAL_ARITH `--c <= --a * b <=> a * b <= c`] THEN + REWRITE_TAC[FLOAT_MUL_FF_hi; FLOAT_MUL_FF_lo] in + prove(ft_hi, proof), prove(ft_lo, proof);; + + +(*********************************************) +(* float_mul_lo, float_mul_hi *) + +let transform = UNDISCH_ALL o NUMERALS_TO_NUM o PURE_REWRITE_RULE[min_exp_def; GSYM IMP_IMP];; +let FLOAT_MUL_FF_hi' = transform FLOAT_MUL_FF_hi and + FLOAT_MUL_FF_lo' = transform FLOAT_MUL_FF_lo and + FLOAT_MUL_TT_hi' = transform FLOAT_MUL_TT_hi and + FLOAT_MUL_TT_lo' = transform FLOAT_MUL_TT_lo and + FLOAT_MUL_FT_hi' = transform FLOAT_MUL_FT_hi and + FLOAT_MUL_FT_lo' = transform FLOAT_MUL_FT_lo and + FLOAT_MUL_TF_hi' = transform FLOAT_MUL_TF_hi and + FLOAT_MUL_TF_lo' = transform FLOAT_MUL_TF_lo and + FLOAT_MUL_0x_hi' = transform FLOAT_MUL_0x_hi and + FLOAT_MUL_0x_lo' = transform FLOAT_MUL_0x_lo and + FLOAT_MUL_x0_hi' = transform FLOAT_MUL_x0_hi and + FLOAT_MUL_x0_lo' = transform FLOAT_MUL_x0_lo;; + +let FLOAT_MUL_FF' = transform FLOAT_MUL_FF and + FLOAT_MUL_TT' = transform FLOAT_MUL_TT and + FLOAT_MUL_FT' = transform FLOAT_MUL_FT and + FLOAT_MUL_TF' = transform FLOAT_MUL_TF;; + +let float_mul_eq f1 f2 = + let s1, n1, e1 = dest_float f1 and + s2, n2, e2 = dest_float f2 in + let flag = s1 = s2 in + let num_exp1 = mk_num_exp n1 e1 and + num_exp2 = mk_num_exp n2 e2 in + + let mul_th = num_exp_mul num_exp1 num_exp2 in + let n_tm, e_tm = dest_num_exp (rand (concl mul_th)) in + + let sub_th, le_th = raw_sub_and_le_hash_conv e_tm min_exp_num_const in + if (rand(concl le_th) <> e_tm) then + failwith "float_mul_eq: underflow" + else + let r_tm = rand(concl sub_th) in + let inst = INST[e_tm, e_var_num; r_tm, r_var_num; n_tm, n_var_num; + n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] in + let th0 = inst + (if flag then + if s1 = "F" then FLOAT_MUL_FF' else FLOAT_MUL_TT' + else + if s1 = "F" then FLOAT_MUL_FT' else FLOAT_MUL_TF') in + MY_PROVE_HYP sub_th (MY_PROVE_HYP mul_th (MY_PROVE_HYP le_th th0));; + + +let float_mul_lo pp f1 f2 = + let s1, n1, e1 = dest_float f1 and + s2, n2, e2 = dest_float f2 in + (* Multiplication by zero *) + let n1_eq0_th = raw_eq0_hash_conv n1 in + if (rand o concl) n1_eq0_th = t_const then + (MY_PROVE_HYP n1_eq0_th o + INST[e1, e1_var_num; f2, f2_var_real; n1, n1_var_num; + (if s1 = "T" then t_const else f_const), s1_var_bool]) FLOAT_MUL_0x_lo' + else + let n2_eq0_th = raw_eq0_hash_conv n2 in + if (rand o concl) n2_eq0_th = t_const then + (MY_PROVE_HYP n2_eq0_th o + INST[e2, e2_var_num; f1, f1_var_real; n2, n2_var_num; + (if s2 = "T" then t_const else f_const), s2_var_bool]) FLOAT_MUL_x0_lo' + else + let flag = s1 = s2 in + let num_exp1 = mk_num_exp n1 e1 and + num_exp2 = mk_num_exp n2 e2 in + + let mul_th, n_tm, e_tm = + if flag then + let th = num_exp_op_lo pp num_exp_mul num_exp1 num_exp2 in + let n_tm, e_tm = dest_num_exp (lhand (concl th)) in + th, n_tm, e_tm + else + let th = num_exp_op_hi pp num_exp_mul num_exp1 num_exp2 in + let n_tm, e_tm = dest_num_exp (rand (concl th)) in + th, n_tm, e_tm in + + let sub_th, le_th = raw_sub_and_le_hash_conv e_tm min_exp_num_const in + if (rand(concl le_th) <> e_tm) then + failwith "float_mul_lo: underflow" + else + let r_tm = rand(concl sub_th) in + let inst = INST[e_tm, e_var_num; r_tm, r_var_num; n_tm, n_var_num; + n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] in + let th0 = inst + (if flag then + if s1 = "F" then FLOAT_MUL_FF_lo' else FLOAT_MUL_TT_lo' + else + if s1 = "F" then FLOAT_MUL_FT_lo' else FLOAT_MUL_TF_lo') in + MY_PROVE_HYP sub_th (MY_PROVE_HYP mul_th (MY_PROVE_HYP le_th th0));; + + +let float_mul_hi pp f1 f2 = + let s1, n1, e1 = dest_float f1 and + s2, n2, e2 = dest_float f2 in + (* Multiplication by zero *) + let n1_eq0_th = raw_eq0_hash_conv n1 in + if (rand o concl) n1_eq0_th = t_const then + (MY_PROVE_HYP n1_eq0_th o + INST[e1, e1_var_num; f2, f2_var_real; n1, n1_var_num; + (if s1 = "T" then t_const else f_const), s1_var_bool]) FLOAT_MUL_0x_hi' + else + let n2_eq0_th = raw_eq0_hash_conv n2 in + if (rand o concl) n2_eq0_th = t_const then + (MY_PROVE_HYP n2_eq0_th o + INST[e2, e2_var_num; f1, f1_var_real; n2, n2_var_num; + (if s2 = "T" then t_const else f_const), s2_var_bool]) FLOAT_MUL_x0_hi' + else + let flag = s1 = s2 in + let num_exp1 = mk_num_exp n1 e1 and + num_exp2 = mk_num_exp n2 e2 in + + let mul_th, n_tm, e_tm = + if flag then + let th = num_exp_op_hi pp num_exp_mul num_exp1 num_exp2 in + let n_tm, e_tm = dest_num_exp (rand (concl th)) in + th, n_tm, e_tm + else + let th = num_exp_op_lo pp num_exp_mul num_exp1 num_exp2 in + let n_tm, e_tm = dest_num_exp (lhand (concl th)) in + th, n_tm, e_tm in + + let sub_th, le_th = raw_sub_and_le_hash_conv e_tm min_exp_num_const in + if (rand(concl le_th) <> e_tm) then + failwith "float_mul_hi: underflow" + else + let r_tm = rand(concl sub_th) in + let inst = INST[e_tm, e_var_num; r_tm, r_var_num; n_tm, n_var_num; + n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] in + let th0 = inst + (if flag then + if s1 = "F" then FLOAT_MUL_FF_hi' else FLOAT_MUL_TT_hi' + else + if s1 = "F" then FLOAT_MUL_FT_hi' else FLOAT_MUL_TF_hi') in + MY_PROVE_HYP sub_th (MY_PROVE_HYP mul_th (MY_PROVE_HYP le_th th0));; + + + +(*********************************************) +(* FLOAT_DIV *) + +let DIV_lemma = prove(`!x y. ~(y = 0) ==> &(x DIV y) <= &x / &y /\ &x / &y <= &(x DIV y + 1)`, + REPEAT GEN_TAC THEN DISCH_TAC THEN + MP_TAC (SPECL [`y:num`; `x:num`] FLOOR_DIV_DIV) THEN + ASM_REWRITE_TAC[GSYM REAL_OF_NUM_ADD] THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN + SIMP_TAC[FLOOR; REAL_LT_IMP_LE]);; + +let FLOAT_DIV_FF = prove(`e2 + k <= min_exp + e + e1 /\ ~(n2 = 0) /\ + num_exp n1 k DIV num_exp n2 0 = num_exp n e + ==> float_num F n ((min_exp + e + e1) - (e2 + k)) <= float_num F n1 e1 / float_num F n2 e2`, + MAP_EVERY ABBREV_TAC [`z = num_exp n e`; `x = num_exp n1 k`; `y = num_exp n2 0`] THEN + REPEAT STRIP_TAC THEN + REWRITE_TAC[float; REAL_MUL_LID] THEN + REWRITE_TAC[real_div; REAL_INV_MUL; REAL_INV_INV] THEN + REWRITE_TAC[REAL_ARITH `(a * b) * c * d = (b * d) * (a * c)`] THEN + SUBGOAL_THEN `~(&(num_exp 1 min_exp) = &0)` ASSUME_TAC THENL + [ + REWRITE_TAC[num_exp; REAL_OF_NUM_EQ; MULT_CLAUSES; EXP_EQ_0] THEN + ARITH_TAC; + ALL_TAC + ] THEN + + ASM_SIMP_TAC[REAL_MUL_LINV; REAL_MUL_LID] THEN + + ASM_SIMP_TAC[NUM_EXP_SUB_lemma] THEN + SUBGOAL_THEN `&(num_exp n1 e1) * inv(&(num_exp n2 e2)) = (&x / &y) * &(num_exp 1 e1) * inv(&(num_exp 1 (e2 + k)))` MP_TAC THENL + [ + EXPAND_TAC "x" THEN EXPAND_TAC "y" THEN + REWRITE_TAC[real_div] THEN + REWRITE_TAC[num_exp; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW] THEN + REWRITE_TAC[REAL_MUL_LID; REAL_INV_MUL; REAL_INV_1; real_pow; REAL_MUL_RID] THEN + REWRITE_TAC[REAL_POW_ADD; REAL_INV_MUL] THEN + REWRITE_TAC[REAL_ARITH `((a * b) * c) * d * e * f = (b * f) * (a * c * d * e)`] THEN + + SUBGOAL_THEN (mk_comb(`(~)`, mk_eq(mk_binop `pow` (mk_comb (`&`, base_const)) `k:num`, `&0`))) ASSUME_TAC THENL + [ + REWRITE_TAC[REAL_POW_EQ_0] THEN + REAL_ARITH_TAC; + ALL_TAC + ] THEN + + ASM_SIMP_TAC[REAL_MUL_RINV; REAL_MUL_LID] THEN + REAL_ARITH_TAC; + ALL_TAC + ] THEN + + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + ONCE_REWRITE_TAC[REAL_ARITH `(a * b) * c = (a * c) * b`] THEN + REWRITE_TAC[REAL_MUL_ASSOC] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN + ONCE_REWRITE_TAC[NUM_EXP_SUM1] THEN + REWRITE_TAC[NUM_EXP_SUM] THEN + REWRITE_TAC[GSYM REAL_OF_NUM_MUL] THEN + ASM_REWRITE_TAC[REAL_ARITH `(a * b * c) * d = (d * a) * b * c`] THEN + ASM_SIMP_TAC[REAL_MUL_LINV; REAL_MUL_LID] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN + MP_TAC (SPEC_ALL DIV_lemma) THEN + ANTS_TAC THENL + [ + EXPAND_TAC "y" THEN + REWRITE_TAC[num_exp; MULT_EQ_0; DE_MORGAN_THM] THEN + ASM_REWRITE_TAC[EXP] THEN + ARITH_TAC; + ALL_TAC + ] THEN + + ASM_SIMP_TAC[]);; + +let FLOAT_DIV_0x_lo = prove(`(n1 = 0 <=> T) ==> float_num F 0 min_exp <= float_num s1 n1 e1 / f2`, + SIMP_TAC[real_div; FLOAT_MUL_0x_lo]);; + +let FLOAT_DIV_0x_hi = prove(`(n1 = 0 <=> T) ==> float_num s1 n1 e1 / f2 <= float_num F 0 min_exp`, + SIMP_TAC[real_div; FLOAT_MUL_0x_hi]);; + +let FLOAT_DIV_FF_lo = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ r2 - r1 = r /\ + r1 <= r2 /\ ~(n2 = 0) /\ + num_exp n e <= num_exp n1 k DIV num_exp n2 0 + ==> float_num F n r <= float_num F n1 e1 / float_num F n2 e2`, + MAP_EVERY ABBREV_TAC [`z = num_exp n e`; `x = num_exp n1 k`; `y = num_exp n2 0`] THEN + REPEAT STRIP_TAC THEN + REPLICATE_TAC 3 (POP_ASSUM MP_TAC) THEN + REPLICATE_TAC 3 (POP_ASSUM (fun th -> REWRITE_TAC[SYM th])) THEN + REPEAT STRIP_TAC THEN + REWRITE_TAC[float; REAL_MUL_LID] THEN + REWRITE_TAC[real_div; REAL_INV_MUL; REAL_INV_INV] THEN + REWRITE_TAC[REAL_ARITH `(a * b) * c * d = (b * d) * (a * c)`] THEN + SUBGOAL_THEN `~(&(num_exp 1 min_exp) = &0)` ASSUME_TAC THENL + [ + REWRITE_TAC[num_exp; REAL_OF_NUM_EQ; MULT_CLAUSES; EXP_EQ_0] THEN + ARITH_TAC; + ALL_TAC + ] THEN + + ASM_SIMP_TAC[REAL_MUL_LINV; REAL_MUL_LID] THEN + + ASM_SIMP_TAC[NUM_EXP_SUB_lemma] THEN + SUBGOAL_THEN `&(num_exp n1 e1) * inv(&(num_exp n2 e2)) = (&x / &y) * &(num_exp 1 e1) * inv(&(num_exp 1 (e2 + k)))` MP_TAC THENL + [ + EXPAND_TAC "x" THEN EXPAND_TAC "y" THEN + REWRITE_TAC[real_div] THEN + REWRITE_TAC[num_exp; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW] THEN + REWRITE_TAC[REAL_MUL_LID; REAL_INV_MUL; REAL_INV_1; real_pow; REAL_MUL_RID] THEN + REWRITE_TAC[REAL_POW_ADD; REAL_INV_MUL] THEN + REWRITE_TAC[REAL_ARITH `((a * b) * c) * d * e * f = (b * f) * (a * c * d * e)`] THEN + SUBGOAL_THEN + (mk_comb(`(~)`, mk_eq(mk_binop `pow` (mk_comb (`&`, base_const)) `k:num`, `&0`))) ASSUME_TAC THENL + [ + REWRITE_TAC[REAL_POW_EQ_0] THEN + REAL_ARITH_TAC; + ALL_TAC + ] THEN + + ASM_SIMP_TAC[REAL_MUL_RINV; REAL_MUL_LID] THEN + REAL_ARITH_TAC; + ALL_TAC + ] THEN + + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + ONCE_REWRITE_TAC[REAL_ARITH `(a * b) * c = (a * c) * b`] THEN + REWRITE_TAC[REAL_MUL_ASSOC] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN + ONCE_REWRITE_TAC[NUM_EXP_SUM1] THEN + REWRITE_TAC[NUM_EXP_SUM] THEN + REWRITE_TAC[GSYM REAL_OF_NUM_MUL] THEN + ASM_REWRITE_TAC[REAL_ARITH `(a * b * c) * d = (d * a) * b * c`] THEN + ASM_SIMP_TAC[REAL_MUL_LINV; REAL_MUL_LID] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN + MP_TAC (SPEC_ALL DIV_lemma) THEN + ANTS_TAC THENL + [ + EXPAND_TAC "y" THEN + REWRITE_TAC[num_exp; MULT_EQ_0; DE_MORGAN_THM] THEN + ASM_REWRITE_TAC[EXP] THEN + ARITH_TAC; + ALL_TAC + ] THEN + + STRIP_TAC THEN + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `&(x DIV y)` THEN + ASM_REWRITE_TAC[REAL_OF_NUM_LE]);; + +let FLOAT_DIV_FF_hi = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ r2 - r1 = r /\ + r1 <= r2 /\ ~(n2 = 0) /\ + num_exp n1 k DIV num_exp n2 0 < num_exp n e + ==> float_num F n1 e1 / float_num F n2 e2 <= float_num F n r`, + MAP_EVERY ABBREV_TAC [`z = num_exp n e`; `x = num_exp n1 k`; `y = num_exp n2 0`] THEN + REPEAT STRIP_TAC THEN + REPLICATE_TAC 3 (POP_ASSUM MP_TAC) THEN + REPLICATE_TAC 3 (POP_ASSUM (fun th -> REWRITE_TAC[SYM th])) THEN + REPEAT STRIP_TAC THEN + REWRITE_TAC[float; REAL_MUL_LID] THEN + REWRITE_TAC[real_div; REAL_INV_MUL; REAL_INV_INV] THEN + REWRITE_TAC[REAL_ARITH `(a * b) * c * d = (b * d) * (a * c)`] THEN + SUBGOAL_THEN `~(&(num_exp 1 min_exp) = &0)` ASSUME_TAC THENL + [ + REWRITE_TAC[num_exp; REAL_OF_NUM_EQ; MULT_CLAUSES; EXP_EQ_0] THEN + ARITH_TAC; + ALL_TAC + ] THEN + + ASM_SIMP_TAC[REAL_MUL_LINV; REAL_MUL_LID] THEN + + ASM_SIMP_TAC[NUM_EXP_SUB_lemma] THEN + SUBGOAL_THEN `&(num_exp n1 e1) * inv(&(num_exp n2 e2)) = (&x / &y) * &(num_exp 1 e1) * inv(&(num_exp 1 (e2 + k)))` MP_TAC THENL + [ + EXPAND_TAC "x" THEN EXPAND_TAC "y" THEN + REWRITE_TAC[real_div] THEN + REWRITE_TAC[num_exp; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW] THEN + REWRITE_TAC[REAL_MUL_LID; REAL_INV_MUL; REAL_INV_1; real_pow; REAL_MUL_RID] THEN + REWRITE_TAC[REAL_POW_ADD; REAL_INV_MUL] THEN + REWRITE_TAC[REAL_ARITH `((a * b) * c) * d * e * f = (b * f) * (a * c * d * e)`] THEN + SUBGOAL_THEN + (mk_comb(`(~)`, mk_eq(mk_binop `pow` (mk_comb (`&`, base_const)) `k:num`, `&0`))) ASSUME_TAC THENL + [ + REWRITE_TAC[REAL_POW_EQ_0] THEN + REAL_ARITH_TAC; + ALL_TAC + ] THEN + + ASM_SIMP_TAC[REAL_MUL_RINV; REAL_MUL_LID] THEN + REAL_ARITH_TAC; + ALL_TAC + ] THEN + + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + ONCE_REWRITE_TAC[REAL_ARITH `(a * b) * c = (a * c) * b`] THEN + REWRITE_TAC[REAL_MUL_ASSOC] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN + ONCE_REWRITE_TAC[NUM_EXP_SUM1] THEN + REWRITE_TAC[NUM_EXP_SUM] THEN + REWRITE_TAC[GSYM REAL_OF_NUM_MUL] THEN + ASM_REWRITE_TAC[REAL_ARITH `(a * b * c) * d = (d * a) * b * c`] THEN + ASM_SIMP_TAC[REAL_MUL_LINV; REAL_MUL_LID] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS] THEN + MP_TAC (SPEC_ALL DIV_lemma) THEN + ANTS_TAC THENL + [ + EXPAND_TAC "y" THEN + REWRITE_TAC[num_exp; MULT_EQ_0; DE_MORGAN_THM] THEN + ASM_REWRITE_TAC[EXP] THEN + ARITH_TAC; + ALL_TAC + ] THEN + + STRIP_TAC THEN + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `&(x DIV y + 1)` THEN + ASM_REWRITE_TAC[REAL_OF_NUM_LE] THEN + UNDISCH_TAC `x DIV y < z` THEN + ARITH_TAC);; + +let FLOAT_DIV_TT_lo = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ r2 - r1 = r /\ + r1 <= r2 /\ ~(n2 = 0) /\ + num_exp n e <= num_exp n1 k DIV num_exp n2 0 + ==> float_num F n r <= float_num T n1 e1 / float_num T n2 e2`, + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[real_div; REAL_INV_NEG; REAL_NEG_MUL2] THEN + REWRITE_TAC[GSYM real_div] THEN + REWRITE_TAC[FLOAT_DIV_FF_lo]);; + +let FLOAT_DIV_TT_hi = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ + r2 - r1 = r /\ r1 <= r2 /\ ~(n2 = 0) /\ + num_exp n1 k DIV num_exp n2 0 < num_exp n e + ==> float_num T n1 e1 / float_num T n2 e2 <= float_num F n r`, + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[real_div; REAL_INV_NEG; REAL_NEG_MUL2] THEN + REWRITE_TAC[GSYM real_div] THEN + REWRITE_TAC[FLOAT_DIV_FF_hi]);; + + +let FLOAT_DIV_FT_lo = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ + r2 - r1 = r /\ r1 <= r2 /\ ~(n2 = 0) /\ + num_exp n1 k DIV num_exp n2 0 < num_exp n e + ==> float_num T n r <= float_num F n1 e1 / float_num T n2 e2`, + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[real_div; REAL_INV_NEG] THEN + REWRITE_TAC[REAL_ARITH `--a <= b * --c <=> b * c <= a`] THEN + REWRITE_TAC[GSYM real_div] THEN + REWRITE_TAC[FLOAT_DIV_FF_hi]);; + +let FLOAT_DIV_FT_hi = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ + r2 - r1 = r /\ r1 <= r2 /\ ~(n2 = 0) /\ + num_exp n e <= num_exp n1 k DIV num_exp n2 0 + ==> float_num F n1 e1 / float_num T n2 e2 <= float_num T n r`, + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[real_div; REAL_INV_NEG] THEN + REWRITE_TAC[REAL_ARITH `a * --b <= --c <=> c <= a * b`] THEN + REWRITE_TAC[GSYM real_div] THEN + REWRITE_TAC[FLOAT_DIV_FF_lo]);; + + +let FLOAT_DIV_TF_lo = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ + r2 - r1 = r /\ r1 <= r2 /\ ~(n2 = 0) /\ + num_exp n1 k DIV num_exp n2 0 < num_exp n e + ==> float_num T n r <= float_num T n1 e1 / float_num F n2 e2`, + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[real_div; REAL_INV_NEG] THEN + REWRITE_TAC[REAL_ARITH `--a <= --b * c <=> b * c <= a`] THEN + REWRITE_TAC[GSYM real_div] THEN + REWRITE_TAC[FLOAT_DIV_FF_hi]);; + +let FLOAT_DIV_TF_hi = prove(`e2 + k = r1 /\ min_exp + e + e1 = r2 /\ + r2 - r1 = r /\ r1 <= r2 /\ ~(n2 = 0) /\ + num_exp n e <= num_exp n1 k DIV num_exp n2 0 + ==> float_num T n1 e1 / float_num F n2 e2 <= float_num T n r`, + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[real_div; REAL_INV_NEG] THEN + REWRITE_TAC[REAL_ARITH `--a * b <= --c <=> c <= a * b`] THEN + REWRITE_TAC[GSYM real_div] THEN + REWRITE_TAC[FLOAT_DIV_FF_lo]);; + + +(******************************************) +(* float_div_lo, float_div_hi *) + +let transform = UNDISCH_ALL o PURE_REWRITE_RULE[TAUT `~P <=> (P <=> F)`] o + NUMERALS_TO_NUM o PURE_REWRITE_RULE[GSYM IMP_IMP; min_exp_def];; + +let FLOAT_DIV_FF_hi' = transform FLOAT_DIV_FF_hi and + FLOAT_DIV_FF_lo' = transform FLOAT_DIV_FF_lo and + FLOAT_DIV_TT_hi' = transform FLOAT_DIV_TT_hi and + FLOAT_DIV_TT_lo' = transform FLOAT_DIV_TT_lo and + FLOAT_DIV_FT_hi' = transform FLOAT_DIV_FT_hi and + FLOAT_DIV_FT_lo' = transform FLOAT_DIV_FT_lo and + FLOAT_DIV_TF_hi' = transform FLOAT_DIV_TF_hi and + FLOAT_DIV_TF_lo' = transform FLOAT_DIV_TF_lo and + FLOAT_DIV_0x_hi' = transform FLOAT_DIV_0x_hi and + FLOAT_DIV_0x_lo' = transform FLOAT_DIV_0x_lo;; + +let float_div_lo pp f1 f2 = + let s1, n1, e1 = dest_float f1 and + s2, n2, e2 = dest_float f2 in + let n1_eq0_th = raw_eq0_hash_conv n1 in + if (rand o concl) n1_eq0_th = t_const then + (MY_PROVE_HYP n1_eq0_th o + INST[e1, e1_var_num; f2, f2_var_real; n1, n1_var_num; + (if s1 = "T" then t_const else f_const), s1_var_bool]) FLOAT_DIV_0x_lo' + else + let flag = s1 = s2 in + + let k_tm = rand (mk_small_numeral_array (2 * pp)) in + let num_exp1 = mk_num_exp n1 k_tm and + num_exp2 = mk_num_exp n2 zero_const in + let div_th, n_tm, e_tm = + if flag then + let th = num_exp_op_lo pp num_exp_div num_exp1 num_exp2 in + let n_tm, e_tm = dest_num_exp (lhand(concl th)) in + th, n_tm, e_tm + else + let th = num_exp_op_hi_lt pp num_exp_div num_exp1 num_exp2 in + let n_tm, e_tm = dest_num_exp (rand(concl th)) in + th, n_tm, e_tm in + + let r1_th = raw_add_conv_hash (mk_binop add_op_num e2 k_tm) in + let r1_tm = rand(concl r1_th) in + let e_plus_e1 = raw_add_conv_hash (mk_binop add_op_num e_tm e1) in + let ltm, rtm = dest_comb(concl e_plus_e1) in + let r2_th' = raw_add_conv_hash (mk_binop add_op_num min_exp_num_const rtm) in + let r2_th = TRANS (AP_TERM (mk_comb (add_op_num, min_exp_num_const)) e_plus_e1) r2_th' in + let r2_tm = rand(concl r2_th) in + let sub_th, le_th = raw_sub_and_le_hash_conv r2_tm r1_tm in + if rand(concl le_th) <> r2_tm then + failwith "float_div_lo: underflow" + else + let r_tm = rand(concl sub_th) in + let n2_not_zero = raw_eq0_hash_conv n2 in + let inst = INST[r1_tm, r1_var_num; r2_tm, r2_var_num; + n1, n1_var_num; e1, e1_var_num; + e_tm, e_var_num; k_tm, k_var_num; + n2, n2_var_num; e2, e2_var_num; + n_tm, n_var_num; r_tm, r_var_num] in + let th0 = inst + (if flag then + if s1 = "F" then FLOAT_DIV_FF_lo' else FLOAT_DIV_TT_lo' + else + if s1 = "F" then FLOAT_DIV_FT_lo' else FLOAT_DIV_TF_lo') in + let th1 = MY_PROVE_HYP n2_not_zero (MY_PROVE_HYP div_th (MY_PROVE_HYP le_th th0)) in + MY_PROVE_HYP sub_th (MY_PROVE_HYP r2_th (MY_PROVE_HYP r1_th th1));; + + +let float_div_hi pp f1 f2 = + let s1, n1, e1 = dest_float f1 and + s2, n2, e2 = dest_float f2 in + let n1_eq0_th = raw_eq0_hash_conv n1 in + if (rand o concl) n1_eq0_th = t_const then + (MY_PROVE_HYP n1_eq0_th o + INST[e1, e1_var_num; f2, f2_var_real; n1, n1_var_num; + (if s1 = "T" then t_const else f_const), s1_var_bool]) FLOAT_DIV_0x_hi' + else + let flag = s1 = s2 in + + let k_tm = rand (mk_small_numeral_array (2 * pp)) in + let num_exp1 = mk_num_exp n1 k_tm and + num_exp2 = mk_num_exp n2 zero_const in + let div_th, n_tm, e_tm = + if flag then + let th = num_exp_op_hi_lt pp num_exp_div num_exp1 num_exp2 in + let n_tm, e_tm = dest_num_exp (rand(concl th)) in + th, n_tm, e_tm + else + let th = num_exp_op_lo pp num_exp_div num_exp1 num_exp2 in + let n_tm, e_tm = dest_num_exp (lhand(concl th)) in + th, n_tm, e_tm in + + let r1_th = raw_add_conv_hash (mk_binop add_op_num e2 k_tm) in + let r1_tm = rand(concl r1_th) in + let e_plus_e1 = raw_add_conv_hash (mk_binop add_op_num e_tm e1) in + let ltm, rtm = dest_comb(concl e_plus_e1) in + let r2_th' = raw_add_conv_hash (mk_binop add_op_num min_exp_num_const rtm) in + let r2_th = TRANS (AP_TERM (mk_comb (add_op_num, min_exp_num_const)) e_plus_e1) r2_th' in + let r2_tm = rand(concl r2_th) in + let sub_th, le_th = raw_sub_and_le_hash_conv r2_tm r1_tm in + if rand(concl le_th) <> r2_tm then + failwith "float_div_hi: underflow" + else + let r_tm = rand(concl sub_th) in + let n2_not_zero = raw_eq0_hash_conv n2 in + let inst = INST[r1_tm, r1_var_num; r2_tm, r2_var_num; + n1, n1_var_num; e1, e1_var_num; + e_tm, e_var_num; k_tm, k_var_num; + n2, n2_var_num; e2, e2_var_num; + n_tm, n_var_num; r_tm, r_var_num] in + let th0 = inst + (if flag then + if s1 = "F" then FLOAT_DIV_FF_hi' else FLOAT_DIV_TT_hi' + else + if s1 = "F" then FLOAT_DIV_FT_hi' else FLOAT_DIV_TF_hi') in + let th1 = MY_PROVE_HYP n2_not_zero (MY_PROVE_HYP div_th (MY_PROVE_HYP le_th th0)) in + MY_PROVE_HYP sub_th (MY_PROVE_HYP r2_th (MY_PROVE_HYP r1_th th1));; + + +(***********************************) +(* FLOAT_ADD *) + +let FLOAT_ADD_FF = prove(`num_exp n1 e1 + num_exp n2 e2 = num_exp n e + ==> float_num F n1 e1 + float_num F n2 e2 = float_num F n e`, + REPEAT STRIP_TAC THEN + REWRITE_TAC[float; REAL_MUL_LID] THEN + REWRITE_TAC[REAL_ARITH `a / b + c / b = (a + c) / b`] THEN + ASM_REWRITE_TAC[REAL_OF_NUM_ADD]);; + +let FLOAT_ADD_TT = prove(`num_exp n1 e1 + num_exp n2 e2 = num_exp n e + ==> float_num T n1 e1 + float_num T n2 e2 = float_num T n e`, + REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `--a + --b = --c <=> a + b = c`] THEN + REWRITE_TAC[FLOAT_ADD_FF]);; + +let FLOAT_ADD_FF_lo = prove(`num_exp n e <= num_exp n1 e1 + num_exp n2 e2 + ==> float_num F n e <= float_num F n1 e1 + float_num F n2 e2`, + REWRITE_TAC[GSYM REAL_OF_NUM_LE; GSYM REAL_OF_NUM_ADD] THEN + REPEAT STRIP_TAC THEN + MAP_EVERY ABBREV_TAC [`z = &(num_exp n e)`; `x = &(num_exp n1 e1)`; `y = &(num_exp n2 e2)`] THEN + ASM_REWRITE_TAC[float; REAL_MUL_LID] THEN + REWRITE_TAC[REAL_ARITH `a / b + c / b = (a + c) / b`] THEN + REWRITE_TAC[real_div] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + ASM_REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS]);; + +let FLOAT_ADD_FF_hi = prove(`num_exp n1 e1 + num_exp n2 e2 <= num_exp n e + ==> float_num F n1 e1 + float_num F n2 e2 <= float_num F n e`, + REWRITE_TAC[GSYM REAL_OF_NUM_LE; GSYM REAL_OF_NUM_ADD] THEN + REPEAT STRIP_TAC THEN + MAP_EVERY ABBREV_TAC [`z = &(num_exp n e)`; `x = &(num_exp n1 e1)`; `y = &(num_exp n2 e2)`] THEN + ASM_REWRITE_TAC[float; REAL_MUL_LID] THEN + REWRITE_TAC[REAL_ARITH `a / b + c / b = (a + c) / b`] THEN + REWRITE_TAC[real_div] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + ASM_REWRITE_TAC[REAL_LE_INV_EQ; REAL_POS]);; + +let FLOAT_ADD_TT_lo = prove(`num_exp n1 e1 + num_exp n2 e2 <= num_exp n e + ==> float_num T n e <= float_num T n1 e1 + float_num T n2 e2`, + REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `--a <= --b + --c <=> b + c <= a`] THEN + REWRITE_TAC[FLOAT_ADD_FF_hi]);; + +let FLOAT_ADD_TT_hi = prove(`num_exp n e <= num_exp n1 e1 + num_exp n2 e2 + ==> float_num T n1 e1 + float_num T n2 e2 <= float_num T n e`, + REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `--b + --c <= --a <=> a <= b + c`] THEN + REWRITE_TAC[FLOAT_ADD_FF_lo]);; + +let FLOAT_ADD_FT_F_lo = prove(`num_exp n2 e2 <= num_exp n1 e1 ==> + num_exp n e <= num_exp n1 e1 - num_exp n2 e2 + ==> float_num F n e <= float_num F n1 e1 + float_num T n2 e2`, + MAP_EVERY ABBREV_TAC[`z = num_exp n e`; `x = num_exp n1 e1`; `y = num_exp n2 e2`] THEN + ASM_REWRITE_TAC[FLOAT_NEG_T; float; REAL_MUL_LID] THEN + DISCH_TAC THEN + ASM_SIMP_TAC[GSYM REAL_OF_NUM_LE; GSYM REAL_OF_NUM_SUB] THEN + REWRITE_TAC[num_exp; min_exp_def; MULT_CLAUSES; GSYM REAL_OF_NUM_POW] THEN + REAL_ARITH_TAC);; + +let FLOAT_ADD_FT_T_lo = prove(`num_exp n1 e1 <= num_exp n2 e2 ==> + num_exp n2 e2 - num_exp n1 e1 <= num_exp n e + ==> float_num T n e <= float_num F n1 e1 + float_num T n2 e2`, + MAP_EVERY ABBREV_TAC[`z = num_exp n e`; `x = num_exp n1 e1`; `y = num_exp n2 e2`] THEN + ASM_REWRITE_TAC[FLOAT_NEG_T; float; REAL_MUL_LID] THEN + DISCH_TAC THEN + ASM_SIMP_TAC[GSYM REAL_OF_NUM_LE; GSYM REAL_OF_NUM_SUB] THEN + REWRITE_TAC[num_exp; min_exp_def; MULT_CLAUSES; GSYM REAL_OF_NUM_POW] THEN + REAL_ARITH_TAC);; + +let FLOAT_ADD_FT_F_hi = prove(`num_exp n2 e2 <= num_exp n1 e1 ==> + num_exp n1 e1 - num_exp n2 e2 <= num_exp n e + ==> float_num F n1 e1 + float_num T n2 e2 <= float_num F n e`, + REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `a + --b <= c <=> --c <= b + --a`] THEN + REWRITE_TAC[GSYM FLOAT_NEG_T; FLOAT_ADD_FT_T_lo]);; + +let FLOAT_ADD_FT_T_hi = prove(`num_exp n1 e1 <= num_exp n2 e2 ==> + num_exp n e <= num_exp n2 e2 - num_exp n1 e1 + ==> float_num F n1 e1 + float_num T n2 e2 <= float_num T n e`, + REWRITE_TAC[FLOAT_NEG_T; REAL_ARITH `a + --b <= --c <=> c <= b + --a`] THEN + REWRITE_TAC[GSYM FLOAT_NEG_T; FLOAT_ADD_FT_F_lo]);; + + +(******************************************) +(* float_add_lo, float_add_hi *) + +let REAL_ADD_COMM = CONJUNCT1 REAL_ADD_AC;; + +let transform = UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o NUMERALS_TO_NUM;; +let FLOAT_ADD_FF_hi' = transform FLOAT_ADD_FF_hi and + FLOAT_ADD_FF_lo' = transform FLOAT_ADD_FF_lo and + FLOAT_ADD_TT_hi' = transform FLOAT_ADD_TT_hi and + FLOAT_ADD_TT_lo' = transform FLOAT_ADD_TT_lo and + FLOAT_ADD_FT_F_lo' = transform FLOAT_ADD_FT_F_lo and + FLOAT_ADD_FT_T_lo' = transform FLOAT_ADD_FT_T_lo and + FLOAT_ADD_FT_F_hi' = transform FLOAT_ADD_FT_F_hi and + FLOAT_ADD_FT_T_hi' = transform FLOAT_ADD_FT_T_hi;; + +let float_add_lo pp f1 f2 = + let s1, n1, e1 = dest_float f1 in + let s2, n2, e2 = dest_float f2 in + if s1 = s2 then + let num_exp1 = mk_num_exp n1 e1 in + let num_exp2 = mk_num_exp n2 e2 in + + if s1 = "F" then + (* F + F *) + let add_th = num_exp_op_lo pp num_exp_add num_exp1 num_exp2 in + let n_tm, e_tm = dest_num_exp (lhand(concl add_th)) in + let th0 = INST[e_tm, e_var_num; n_tm, n_var_num; n1, n1_var_num; + e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] FLOAT_ADD_FF_lo' in + MY_PROVE_HYP add_th th0 + else + (* T + T *) + let add_th = num_exp_op_hi pp num_exp_add num_exp1 num_exp2 in + let n_tm, e_tm = dest_num_exp (rand(concl add_th)) in + let th0 = INST[e_tm, e_var_num; n_tm, n_var_num; n1, n1_var_num; + e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] FLOAT_ADD_TT_lo' in + MY_PROVE_HYP add_th th0 + else + (* F + T or T + F *) + let th0, n1, e1, n2, e2 = + if s1 = "T" then + INST[f2, m_var_real; f1, n_var_real] REAL_ADD_COMM, n2, e2, n1, e1 + else + REFL(mk_binop add_op_real f1 f2), n1, e1, n2, e2 in + + let num_exp1 = mk_num_exp n1 e1 in + let num_exp2 = mk_num_exp n2 e2 in + + let sub_th, le_th = num_exp_sub num_exp1 num_exp2 in + let sub_tm = rand(concl sub_th) in + + if rand(concl le_th) = num_exp1 then + let lo_th = num_exp_lo pp sub_tm in + let n_tm, e_tm = dest_num_exp (lhand(concl lo_th)) in + let lo_sub_th = EQ_MP (AP_TERM (rator(concl lo_th)) (SYM sub_th)) lo_th in + + let th1 = INST[n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num; + n_tm, n_var_num; e_tm, e_var_num] FLOAT_ADD_FT_F_lo' in + let th2 = MY_PROVE_HYP lo_sub_th (MY_PROVE_HYP le_th th1) in + EQ_MP (AP_TERM (rator(concl th2)) th0) th2 + + else + let hi_th = num_exp_hi pp sub_tm in + let n_tm, e_tm = dest_num_exp(rand(concl hi_th)) in + let hi_sub_th = EQ_MP (SYM (AP_THM (AP_TERM le_op_num sub_th) (rand(concl hi_th)))) hi_th in + + let th1 = INST[n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num; + n_tm, n_var_num; e_tm, e_var_num] FLOAT_ADD_FT_T_lo' in + let th2 = MY_PROVE_HYP hi_sub_th (MY_PROVE_HYP le_th th1) in + EQ_MP (AP_TERM (rator(concl th2)) th0) th2;; + + +let float_add_hi pp f1 f2 = + let s1, n1, e1 = dest_float f1 in + let s2, n2, e2 = dest_float f2 in + if s1 = s2 then + let num_exp1 = mk_num_exp n1 e1 in + let num_exp2 = mk_num_exp n2 e2 in + + if s1 = "F" then + (* F + F *) + let add_th = num_exp_op_hi pp num_exp_add num_exp1 num_exp2 in + let n_tm, e_tm = dest_num_exp (rand(concl add_th)) in + let th0 = INST[e_tm, e_var_num; n_tm, n_var_num; n1, n1_var_num; + e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] FLOAT_ADD_FF_hi' in + MY_PROVE_HYP add_th th0 + else + (* T + T *) + let add_th = num_exp_op_lo pp num_exp_add num_exp1 num_exp2 in + let n_tm, e_tm = dest_num_exp (lhand(concl add_th)) in + let th0 = INST[e_tm, e_var_num; n_tm, n_var_num; n1, n1_var_num; + e1, e1_var_num; n2, n2_var_num; e2, e2_var_num] FLOAT_ADD_TT_hi' in + MY_PROVE_HYP add_th th0 + else + (* F + T or T + F *) + let th0, n1, e1, n2, e2 = + if s1 = "T" then + INST[f2, m_var_real; f1, n_var_real] REAL_ADD_COMM, n2, e2, n1, e1 + else + REFL(mk_binop add_op_real f1 f2), n1, e1, n2, e2 in + + let num_exp1 = mk_num_exp n1 e1 in + let num_exp2 = mk_num_exp n2 e2 in + + let sub_th, le_th = num_exp_sub num_exp1 num_exp2 in + let sub_tm = rand(concl sub_th) in + + if rand(concl le_th) = num_exp1 then + let hi_th = num_exp_hi pp sub_tm in + let n_tm, e_tm = dest_num_exp (rand(concl hi_th)) in + let hi_sub_th = EQ_MP (SYM (AP_THM (AP_TERM le_op_num sub_th) (rand(concl hi_th)))) hi_th in + + let th1 = INST[n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num; + n_tm, n_var_num; e_tm, e_var_num] FLOAT_ADD_FT_F_hi' in + let th2 = MY_PROVE_HYP hi_sub_th (MY_PROVE_HYP le_th th1) in + EQ_MP (AP_THM (AP_TERM le_op_real th0) (rand(concl th2))) th2 + + else + let lo_th = num_exp_lo pp sub_tm in + let n_tm, e_tm = dest_num_exp(lhand(concl lo_th)) in + let lo_sub_th = EQ_MP (AP_TERM (rator(concl lo_th)) (SYM sub_th)) lo_th in + + let th1 = INST[n1, n1_var_num; e1, e1_var_num; n2, n2_var_num; e2, e2_var_num; + n_tm, n_var_num; e_tm, e_var_num] FLOAT_ADD_FT_T_hi' in + let th2 = MY_PROVE_HYP lo_sub_th (MY_PROVE_HYP le_th th1) in + EQ_MP (AP_THM (AP_TERM le_op_real th0) (rand(concl th2))) th2;; + + +(******************************************) +(* float_sub_lo, float_sub_hi *) + +let FLOAT_SUB_F_EQ_ADD = (SYM o prove)(`f1 - float_num F n2 e2 = f1 + float_num T n2 e2`, + REWRITE_TAC[FLOAT_NEG_T] THEN REAL_ARITH_TAC);; + +let FLOAT_SUB_T_EQ_ADD = (SYM o prove)(`f1 - float_num T n2 e2 = f1 + float_num F n2 e2`, + REWRITE_TAC[FLOAT_NEG_T] THEN REAL_ARITH_TAC);; + +let float_sub_lo pp f1 f2 = + let s2, n2, e2 = dest_float f2 in + let th0 = + INST[f1, f1_var_real; n2, n2_var_num; e2, e2_var_num] + (if s2 = "F" then FLOAT_SUB_F_EQ_ADD else FLOAT_SUB_T_EQ_ADD) in + let ltm,f2_tm = dest_comb(lhand(concl th0)) in + let f1_tm = rand ltm in + let lo_th = float_add_lo pp f1_tm f2_tm in + EQ_MP (AP_TERM (rator(concl lo_th)) th0) lo_th;; + + +let float_sub_hi pp f1 f2 = + let s2, n2, e2 = dest_float f2 in + let th0 = + INST[f1, f1_var_real; n2, n2_var_num; e2, e2_var_num] + (if s2 = "F" then FLOAT_SUB_F_EQ_ADD else FLOAT_SUB_T_EQ_ADD) in + let ltm, f2_tm = dest_comb(lhand(concl th0)) in + let f1_tm = rand ltm in + let hi_th = float_add_hi pp f1_tm f2_tm in + EQ_MP (AP_THM (AP_TERM le_op_real th0) (rand(concl hi_th))) hi_th;; + + +(*******************************************) +(* FLOAT_SQRT *) + +(* float_num F m e = float_num F (B0 m) (PRE e) *) +let FLOAT_PRE_EXP = prove(mk_imp(`~(e = 0) /\ PRE e = e1`, + mk_eq(`float_num F m e`, + mk_comb(mk_comb(`float_num F`, mk_comb(b0_const, m_var_num)), `e1:num`))), + STRIP_TAC THEN POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN + REWRITE_TAC[float; REAL_MUL_LID; real_div; REAL_EQ_MUL_RCANCEL] THEN + DISJ1_TAC THEN + REWRITE_TAC[num_exp; b0_thm; REAL_OF_NUM_EQ] THEN + SUBGOAL_THEN `e = SUC (PRE e)` MP_TAC THENL + [ + POP_ASSUM MP_TAC THEN ARITH_TAC; + ALL_TAC + ] THEN + DISCH_THEN (fun th -> CONV_TAC (LAND_CONV (ONCE_REWRITE_CONV[th]))) THEN + REWRITE_TAC[EXP] THEN + ARITH_TAC);; + +let DIV2_EVEN_lemma = prove(`!n. EVEN n ==> 2 * (n DIV 2) = n`, + GEN_TAC THEN + REWRITE_TAC[EVEN_EXISTS] THEN + STRIP_TAC THEN + ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC (ARITH_RULE `x = y ==> 2 * x = 2 * y`) THEN + MATCH_MP_TAC DIV_MULT THEN + ARITH_TAC);; + +let FLOAT_SQRT_EVEN_lo = prove(`f1 * f1 = f2 /\ f2 <= x /\ + num_exp m (2 * p) = x /\ f1 = num_exp n1 e1 + /\ EVEN e /\ e DIV 2 = e2 /\ + e1 + e2 + (min_exp DIV 2) = r /\ + p <= r /\ r - p = r2 + ==> float_num F n1 r2 <= sqrt (float_num F m e)`, + STRIP_TAC THEN + UNDISCH_TAC `f2 <= x:num` THEN + UNDISCH_TAC `num_exp m (2 * p) = x` THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN + UNDISCH_TAC `f1 * f1 = f2:num` THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN + UNDISCH_TAC `e1 + e2 + min_exp DIV 2 = r` THEN + UNDISCH_TAC `e DIV 2 = e2` THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN + REPEAT (POP_ASSUM MP_TAC) THEN + REWRITE_TAC[num_exp; float; REAL_MUL_LID; GSYM REAL_OF_NUM_MUL] THEN + REPEAT STRIP_TAC THEN + UNDISCH_TAC `r - p = r2:num` THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN + + MATCH_MP_TAC REAL_LE_RSQRT THEN + REWRITE_TAC[GSYM REAL_OF_NUM_POW; REAL_POW_DIV] THEN + REWRITE_TAC[REAL_POW_2; real_div; REAL_INV_MUL] THEN + REWRITE_TAC[REAL_MUL_ASSOC] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + CONJ_TAC THENL + [ + REWRITE_TAC[REAL_ARITH `(((a * b) * a) * b) * c = (a * a) * (b * b) * c:real`] THEN + REWRITE_TAC[GSYM REAL_POW_ADD] THEN + REWRITE_TAC[ARITH_RULE `r - p + r - p = 2 * r - 2 * p`] THEN + MP_TAC (SPECL[mk_comb(amp_op_real, base_const); `2 * r`; `2 * p`] REAL_DIV_POW2) THEN + ANTS_TAC THENL [ REAL_ARITH_TAC; ALL_TAC ] THEN + ASM_SIMP_TAC[ARITH_RULE `p <= r ==> 2 * p <= 2 * r`] THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th; real_div]) THEN + + SUBGOAL_THEN `2 * r = (e1 + e1) + min_exp + e` (fun th -> REWRITE_TAC[th]) THENL + [ + EXPAND_TAC "r" THEN + REWRITE_TAC[ARITH_RULE `2 * (e1 + b + c) = (e1 + e1) + 2 * c + 2 * b`] THEN + MATCH_MP_TAC (ARITH_RULE `b1 = b2 /\ c1 = c2 ==> a + b1 + c1 = a + b2 + c2:num`) THEN + SUBGOAL_THEN `EVEN min_exp` ASSUME_TAC THENL + [ + REWRITE_TAC[min_exp_def] THEN ARITH_TAC; + ALL_TAC + ] THEN + ASM_SIMP_TAC[DIV2_EVEN_lemma]; + ALL_TAC + ] THEN + + REWRITE_TAC[REAL_POW_ADD] THEN + REWRITE_TAC[REAL_ARITH `(n * n) * (((e * e) * x * y) * z) * u = (n * e) * (n * e) * (x * u) * z * y:real`] THEN + SUBGOAL_THEN `~(&(num_exp 1 min_exp) = &0)` MP_TAC THENL + [ + REWRITE_TAC[REAL_OF_NUM_EQ; NUM_EXP_EQ_0] THEN ARITH_TAC; + ALL_TAC + ] THEN + REWRITE_TAC[num_exp; REAL_MUL_LID; GSYM REAL_OF_NUM_POW; GSYM REAL_OF_NUM_MUL] THEN + DISCH_THEN (fun th -> SIMP_TAC[th; REAL_MUL_RINV; REAL_MUL_LID]) THEN + FIRST_X_ASSUM (MP_TAC o check(fun th -> (fst o dest_var o lhand o concl) th = "f1")) THEN + REWRITE_TAC[GSYM REAL_OF_NUM_EQ; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW] THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN + + REWRITE_TAC[REAL_MUL_ASSOC] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + CONJ_TAC THENL + [ + SUBGOAL_THEN `!x y z. &0 < x /\ y <= z * x ==> y * inv x <= z` MP_TAC THENL + [ + REPEAT STRIP_TAC THEN + MP_TAC (SPECL [`y * inv x`; `z:real`; `x:real`] REAL_LE_RMUL_EQ) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th; GSYM REAL_MUL_ASSOC]) THEN + ASM_SIMP_TAC[REAL_ARITH `&0 < x ==> ~(x = &0)`; REAL_MUL_LINV; REAL_MUL_RID]; + ALL_TAC + ] THEN + + DISCH_THEN (MP_TAC o SPECL[`&(num_exp 1 (2 * p))`; `&(f1 * f1)`; `&m`]) THEN + REWRITE_TAC[num_exp; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW; REAL_MUL_LID] THEN + DISCH_THEN MATCH_MP_TAC THEN + ASM_REWRITE_TAC[REAL_OF_NUM_MUL; REAL_OF_NUM_POW; REAL_OF_NUM_LE] THEN + REWRITE_TAC[REAL_OF_NUM_LT; EXP_LT_0] THEN + ARITH_TAC; + ALL_TAC + ] THEN + + MATCH_MP_TAC REAL_POW_LE THEN + ARITH_TAC; + ALL_TAC + ] THEN + + MATCH_MP_TAC REAL_LE_INV THEN + MATCH_MP_TAC REAL_POW_LE THEN + ARITH_TAC);; + + +let FLOAT_SQRT_EVEN_hi = prove(`f1 * f1 = f2 /\ x <= f2 /\ + num_exp m (2 * p) = x /\ f1 = num_exp n1 e1 + /\ EVEN e /\ e DIV 2 = e2 /\ + e1 + e2 + (min_exp DIV 2) = r /\ + p <= r /\ r - p = r2 + ==> sqrt (float_num F m e) <= float_num F n1 r2`, + STRIP_TAC THEN + UNDISCH_TAC `x <= f2:num` THEN + UNDISCH_TAC `num_exp m (2 * p) = x` THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN + UNDISCH_TAC `f1 * f1 = f2:num` THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN + UNDISCH_TAC `e1 + e2 + min_exp DIV 2 = r` THEN + UNDISCH_TAC `e DIV 2 = e2` THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN + REPEAT (POP_ASSUM MP_TAC) THEN + REWRITE_TAC[num_exp; float; REAL_MUL_LID; GSYM REAL_OF_NUM_MUL] THEN + REPEAT STRIP_TAC THEN + UNDISCH_TAC `r - p = r2:num` THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN + + MATCH_MP_TAC REAL_LE_LSQRT THEN + REPEAT CONJ_TAC THENL + [ + REWRITE_TAC[REAL_OF_NUM_MUL; real_div] THEN + MATCH_MP_TAC REAL_LE_MUL THEN + REWRITE_TAC[REAL_POS] THEN + MATCH_MP_TAC REAL_LE_INV THEN + REWRITE_TAC[REAL_POS]; + ALL_TAC + ] THEN + + REWRITE_TAC[GSYM REAL_OF_NUM_POW; REAL_POW_DIV] THEN + REWRITE_TAC[REAL_POW_2; real_div; REAL_INV_MUL] THEN + REWRITE_TAC[REAL_MUL_ASSOC] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + CONJ_TAC THENL + [ + REWRITE_TAC[REAL_ARITH `(((a * b) * a) * b) * c = (a * a) * (b * b) * c:real`] THEN + REWRITE_TAC[GSYM REAL_POW_ADD] THEN + REWRITE_TAC[ARITH_RULE `r - p + r - p = 2 * r - 2 * p`] THEN + MP_TAC (SPECL[mk_comb(amp_op_real, base_const); `2 * r`; `2 * p`] REAL_DIV_POW2) THEN + ANTS_TAC THENL [ REAL_ARITH_TAC; ALL_TAC ] THEN + ASM_SIMP_TAC[ARITH_RULE `p <= r ==> 2 * p <= 2 * r`] THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th; real_div]) THEN + + SUBGOAL_THEN `2 * r = (e1 + e1) + min_exp + e` (fun th -> REWRITE_TAC[th]) THENL + [ + EXPAND_TAC "r" THEN + REWRITE_TAC[ARITH_RULE `2 * (e1 + b + c) = (e1 + e1) + 2 * c + 2 * b`] THEN + MATCH_MP_TAC (ARITH_RULE `b1 = b2 /\ c1 = c2 ==> a + b1 + c1 = a + b2 + c2:num`) THEN + SUBGOAL_THEN `EVEN min_exp` ASSUME_TAC THENL + [ + REWRITE_TAC[min_exp_def] THEN ARITH_TAC; + ALL_TAC + ] THEN + ASM_SIMP_TAC[DIV2_EVEN_lemma]; + ALL_TAC + ] THEN + + REWRITE_TAC[REAL_POW_ADD] THEN + REWRITE_TAC[REAL_ARITH `(n * n) * (((e * e) * x * y) * z) * u = (n * e) * (n * e) * (x * u) * z * y:real`] THEN + SUBGOAL_THEN `~(&(num_exp 1 min_exp) = &0)` MP_TAC THENL + [ + REWRITE_TAC[REAL_OF_NUM_EQ; NUM_EXP_EQ_0] THEN ARITH_TAC; + ALL_TAC + ] THEN + REWRITE_TAC[num_exp; REAL_MUL_LID; GSYM REAL_OF_NUM_POW; GSYM REAL_OF_NUM_MUL] THEN + DISCH_THEN (fun th -> SIMP_TAC[th; REAL_MUL_RINV; REAL_MUL_LID]) THEN + FIRST_X_ASSUM (MP_TAC o check(fun th -> (fst o dest_var o lhand o concl) th = "f1")) THEN + REWRITE_TAC[GSYM REAL_OF_NUM_EQ; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW] THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th]) THEN + + REWRITE_TAC[REAL_MUL_ASSOC] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + CONJ_TAC THENL + [ + SUBGOAL_THEN `!x y z. &0 < x /\ z * x <= y ==> z <= y * inv x` MP_TAC THENL + [ + REPEAT STRIP_TAC THEN + MP_TAC (SPECL [`z:real`; `y * inv x`; `x:real`] REAL_LE_RMUL_EQ) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN (fun th -> REWRITE_TAC[SYM th; GSYM REAL_MUL_ASSOC]) THEN + ASM_SIMP_TAC[REAL_ARITH `&0 < x ==> ~(x = &0)`; REAL_MUL_LINV; REAL_MUL_RID]; + ALL_TAC + ] THEN + + DISCH_THEN (MP_TAC o SPECL[`&(num_exp 1 (2 * p))`; `&(f1 * f1)`; `&m`]) THEN + REWRITE_TAC[num_exp; GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW; REAL_MUL_LID] THEN + DISCH_THEN MATCH_MP_TAC THEN + ASM_REWRITE_TAC[REAL_OF_NUM_MUL; REAL_OF_NUM_POW; REAL_OF_NUM_LE] THEN + REWRITE_TAC[REAL_OF_NUM_LT; EXP_LT_0] THEN + ARITH_TAC; + ALL_TAC + ] THEN + + MATCH_MP_TAC REAL_POW_LE THEN + ARITH_TAC; + ALL_TAC + ] THEN + + MATCH_MP_TAC REAL_LE_INV THEN + MATCH_MP_TAC REAL_POW_LE THEN + ARITH_TAC);; + + +(******************) +let transform = UNDISCH_ALL o + PURE_ONCE_REWRITE_RULE[TAUT `EVEN e <=> (EVEN e <=> T)`] o + NUMERALS_TO_NUM o + CONV_RULE (DEPTH_CONV NUM_DIV_CONV) o + REWRITE_RULE[GSYM IMP_IMP; min_exp_def];; + +let FLOAT_SQRT_EVEN_lo' = transform FLOAT_SQRT_EVEN_lo and + FLOAT_SQRT_EVEN_hi' = transform FLOAT_SQRT_EVEN_hi and + FLOAT_PRE_EXP' = (UNDISCH_ALL o + PURE_ONCE_REWRITE_RULE[TAUT `~(e = _0) <=> ((e = _0) <=> F)`] o + REWRITE_RULE[GSYM IMP_IMP; NUMERAL]) FLOAT_PRE_EXP;; + +let even_const = `EVEN` and + pre_const = `PRE` and + two_num = rand(mk_small_numeral_array 2) and + min_exp_div2 = rand(mk_small_numeral_array (min_exp / 2)) and + f2_var_num = `f2:num` and + f1_var_num = `f1:num` and + p_var_num = `p:num`;; + +open Big_int;; + +(* Returns the list of digits of the given Big_int n in the base b *) +let rec get_big_int_digits b n = + let bb = big_int_of_int b in + if le_big_int n zero_big_int then [] + else + let q, r = quomod_big_int n bb in + r :: get_big_int_digits b q;; + +(* [1;2;3] -> 123 (base = 10) *) +let rec big_int_from_list b list = + let rec proc acc list = + match list with + [] -> acc + | h::t -> proc (add_big_int h (mult_int_big_int b acc)) t in + proc zero_big_int list;; + +(* Returns n first elements of the list *) +let rec take n list = + match list with + x :: xs -> if n > 0 then x :: take (n - 1) xs else [] + | [] -> [];; + +(* Returns an integer number that contains at most pp significant digits + in the given base b *) +let big_int_round_lo base pp n = + let digits = rev (get_big_int_digits base n) in + let n_digits = length digits in + if n_digits <= pp then + n + else + let m = big_int_from_list base (take pp digits) in + mult_big_int (power_int_positive_int base (n_digits - pp)) m;; + +let big_int_round_hi base pp n = + let digits = rev (get_big_int_digits base n) in + let n_digits = length digits in + if n_digits <= pp then n + else + let l1, l2 = chop_list pp digits in + if forall (eq_big_int zero_big_int) l2 then n + else + let m = succ_big_int (big_int_from_list base l1) in + mult_big_int (power_int_positive_int base (n_digits - pp)) m;; + + +(******************) +let rec float_sqrt_lo pp tm = + let s, m_tm, e_tm = dest_float tm in + let p_tm = rand (mk_small_numeral_array pp) in + if s <> "F" then + failwith "float_sqrt_lo: negative argument" + else + let even_th = raw_even_hash_conv (mk_comb (even_const, e_tm)) in + if (fst o dest_const o rand o concl) even_th <> "T" then + (* ODD e *) + let pre_e = raw_pre_hash_conv (mk_comb (pre_const, e_tm)) in + let e_neq_0 = raw_eq0_hash_conv e_tm in + let e1_tm = rand (concl pre_e) in + let th0 = INST[e1_tm, e1_var_num; e_tm, e_var_num; m_tm, m_var_num] FLOAT_PRE_EXP' in + let th1 = MY_PROVE_HYP pre_e (MY_PROVE_HYP e_neq_0 th0) in + let th2 = float_sqrt_lo pp (rand(concl th1)) in + let ltm, rtm = dest_comb (concl th2) in + EQ_MP (SYM (AP_TERM ltm (AP_TERM (rator rtm) th1))) th2 + else + (* EVEN e *) + let p2_tm = mk_binop mul_op_num two_num p_tm in + let p2_th = raw_mul_conv_hash p2_tm in + let f1_1 = AP_TERM (mk_comb(num_exp_const, m_tm)) p2_th in + let f1_2 = TRANS f1_1 (denormalize (rand (concl f1_1))) in + + let x_tm = rand(concl f1_2) in + let x = raw_dest_hash x_tm in + let f1' = Big_int.sqrt_big_int (big_int_of_num x) in + let f1 = num_of_big_int (big_int_round_lo Arith_num.arith_base pp f1') in + let f1_tm = rand(mk_numeral_array f1) in + let f1_num_exp = to_num_exp f1_tm in + + let n1_tm, e1_tm = dest_num_exp (rand (concl f1_num_exp)) in + let f1f1_eq_f2 = raw_mul_conv_hash (mk_binop mul_op_num f1_tm f1_tm) in + let f2_tm = rand(concl f1f1_eq_f2) in + let f2_le_x = EQT_ELIM (raw_le_hash_conv (mk_binop le_op_num f2_tm x_tm)) in + + let e_div2_eq_e2 = raw_div_hash_conv (mk_binop div_op_num e_tm two_num) in + let e2_tm = rand(concl e_div2_eq_e2) in + let r_th1 = raw_add_conv_hash (mk_binop add_op_num e2_tm min_exp_div2) in + let r_th2 = AP_TERM (mk_comb(add_op_num, e1_tm)) r_th1 in + let r_th = TRANS r_th2 (raw_add_conv_hash (rand (concl r_th2))) in + + let r_tm = rand(concl r_th) in + let r_sub_p, p_le_r = raw_sub_and_le_hash_conv p_tm r_tm in + let r2_tm = rand(concl r_sub_p) in + if (rand(concl p_le_r) <> r_tm) then + failwith "float_sqrt_lo: underflow" + else + let th0 = INST[f2_tm, f2_var_num; x_tm, x_var_num; p_tm, p_var_num; r_tm, r_var_num; + f1_tm, f1_var_num; n1_tm, n1_var_num; e1_tm, e1_var_num; e2_tm, e2_var_num; + e_tm, e_var_num; m_tm, m_var_num; r2_tm, r2_var_num] + FLOAT_SQRT_EVEN_lo' in + MY_PROVE_HYP f1_2 ( + MY_PROVE_HYP e_div2_eq_e2 ( + MY_PROVE_HYP r_sub_p ( + MY_PROVE_HYP r_th ( + MY_PROVE_HYP f1f1_eq_f2 ( + MY_PROVE_HYP f1_num_exp ( + MY_PROVE_HYP even_th ( + MY_PROVE_HYP f2_le_x ( + MY_PROVE_HYP p_le_r th0 + ))))))));; + + +let rec float_sqrt_hi pp tm = + let s, m_tm, e_tm = dest_float tm in + let p_tm = rand (mk_small_numeral_array pp) in + if s <> "F" then + failwith "float_sqrt_lo: negative argument" + else + let even_th = raw_even_hash_conv (mk_comb (even_const, e_tm)) in + if (fst o dest_const o rand o concl) even_th <> "T" then + (* ODD e *) + let pre_e = raw_pre_hash_conv (mk_comb (pre_const, e_tm)) in + let e_neq_0 = raw_eq0_hash_conv e_tm in + let e1_tm = rand (concl pre_e) in + let th0 = INST[e1_tm, e1_var_num; e_tm, e_var_num; m_tm, m_var_num] FLOAT_PRE_EXP' in + let th1 = MY_PROVE_HYP pre_e (MY_PROVE_HYP e_neq_0 th0) in + let th2 = float_sqrt_hi pp (rand(concl th1)) in + let ltm, rtm = dest_comb (concl th2) in + let ltm2, rtm2 = dest_comb ltm in + let th3 = AP_THM (AP_TERM ltm2 (AP_TERM (rator rtm2) th1)) rtm in + EQ_MP (SYM th3) th2 + else + (* EVEN e *) + let p2_tm = mk_binop mul_op_num two_num p_tm in + let p2_th = raw_mul_conv_hash p2_tm in + let f1_1 = AP_TERM (mk_comb(num_exp_const, m_tm)) p2_th in + let f1_2 = TRANS f1_1 (denormalize (rand (concl f1_1))) in + + let x_tm = rand(concl f1_2) in + let x = raw_dest_hash x_tm in + let x' = big_int_of_num x in + let f1' = sqrt_big_int x' in + let f1 = (num_of_big_int o big_int_round_hi Arith_num.arith_base pp) + (if eq_big_int (mult_big_int f1' f1') x' then f1' else succ_big_int f1') in + + let f1_tm = rand(mk_numeral_array f1) in + let f1_num_exp = to_num_exp f1_tm in + + let n1_tm, e1_tm = dest_num_exp (rand (concl f1_num_exp)) in + let f1f1_eq_f2 = raw_mul_conv_hash (mk_binop mul_op_num f1_tm f1_tm) in + let f2_tm = rand(concl f1f1_eq_f2) in + let x_le_f2 = EQT_ELIM (raw_le_hash_conv (mk_binop le_op_num x_tm f2_tm)) in + + let e_div2_eq_e2 = raw_div_hash_conv (mk_binop div_op_num e_tm two_num) in + let e2_tm = rand(concl e_div2_eq_e2) in + let r_th1 = raw_add_conv_hash (mk_binop add_op_num e2_tm min_exp_div2) in + let r_th2 = AP_TERM (mk_comb(add_op_num, e1_tm)) r_th1 in + let r_th = TRANS r_th2 (raw_add_conv_hash (rand (concl r_th2))) in + + let r_tm = rand(concl r_th) in + let r_sub_p, p_le_r = raw_sub_and_le_hash_conv p_tm r_tm in + let r2_tm = rand(concl r_sub_p) in + if (rand(concl p_le_r) <> r_tm) then + failwith "float_sqrt_lo: underflow" + else + let th0 = INST[f2_tm, f2_var_num; x_tm, x_var_num; p_tm, p_var_num; r_tm, r_var_num; + f1_tm, f1_var_num; n1_tm, n1_var_num; e1_tm, e1_var_num; e2_tm, e2_var_num; + e_tm, e_var_num; m_tm, m_var_num; r2_tm, r2_var_num] + FLOAT_SQRT_EVEN_hi' in + MY_PROVE_HYP f1_2 ( + MY_PROVE_HYP e_div2_eq_e2 ( + MY_PROVE_HYP r_sub_p ( + MY_PROVE_HYP r_th ( + MY_PROVE_HYP f1f1_eq_f2 ( + MY_PROVE_HYP f1_num_exp ( + MY_PROVE_HYP even_th ( + MY_PROVE_HYP x_le_f2 ( + MY_PROVE_HYP p_le_r th0 + ))))))));; + +end;; (* Float_ops module *) + +(************************************) +(* Cached floating point operations *) +(************************************) + +let init_logs () = + if !Arith_options.float_log then + begin + Log.open_log "mul_lo"; + Log.open_log "mul_hi"; + Log.open_log "div_lo"; + Log.open_log "div_hi"; + Log.open_log "add_lo"; + Log.open_log "add_hi"; + Log.open_log "sub_lo"; + Log.open_log "sub_hi" + end;; + +let float_to_log_string tm = + let s, n_tm, e_tm = dest_float tm in + let n = (if s = "T" then minus_num else I) (raw_dest_hash n_tm) in + let e = raw_dest_hash e_tm -/ Num.num_of_int Float_theory.min_exp in + string_of_num n ^ "," ^ string_of_num e;; + +let log_binop name pp tm1 tm2 = + let str1 = float_to_log_string tm1 and + str2 = float_to_log_string tm2 in + let str = string_of_int pp ^ ":" ^ str1 ^ ";" ^ str2 in + Log.append_to_log name str;; + +(* Counters for collecting stats *) +let lt0_c = ref 0 and + gt0_c = ref 0 and + lt_c = ref 0 and + le0_c = ref 0 and + ge0_c = ref 0 and + le_c = ref 0 and + min_c = ref 0 and + max_c = ref 0 and + min_max_c = ref 0 and + mul_lo_c = ref 0 and + mul_hi_c = ref 0 and + div_lo_c = ref 0 and + div_hi_c = ref 0 and + add_lo_c = ref 0 and + add_hi_c = ref 0 and + sub_lo_c = ref 0 and + sub_hi_c = ref 0 and + sqrt_lo_c = ref 0 and + sqrt_hi_c = ref 0;; + +(* Hash tables *) +let cache_size = if !Arith_options.float_cached then !Arith_options.init_cache_size else 1;; + +let my_add h key v = + if Hashtbl.length h >= !Arith_options.max_cache_size then + Hashtbl.clear h +(* let _ = Hashtbl.clear h in + print_string "Clearing a float hash table" *) + else + (); + Hashtbl.add h key v;; + +let mul_table = Hashtbl.create cache_size and + div_table = Hashtbl.create cache_size and + add_table = Hashtbl.create cache_size and + sub_table = Hashtbl.create cache_size and + sqrt_table = Hashtbl.create cache_size and + le_table = Hashtbl.create cache_size and + max_table = Hashtbl.create cache_size;; + +let reset_cache () = + Hashtbl.clear mul_table; + Hashtbl.clear div_table; + Hashtbl.clear add_table; + Hashtbl.clear sub_table; + Hashtbl.clear sqrt_table; + Hashtbl.clear le_table; + Hashtbl.clear max_table;; + +let reset_stat () = + lt0_c := 0; + gt0_c := 0; + lt_c := 0; + le0_c := 0; + ge0_c := 0; + le_c := 0; + min_c := 0; + max_c := 0; + min_max_c := 0; + mul_lo_c := 0; + mul_hi_c := 0; + div_lo_c := 0; + div_hi_c := 0; + add_lo_c := 0; + add_hi_c := 0; + sub_lo_c := 0; + sub_hi_c := 0; + sqrt_lo_c := 0; + sqrt_hi_c := 0;; + +let print_stat () = + let len = Hashtbl.length in + let cmp_str1 = sprintf "lt0 = %d\ngt0 = %d\nlt = %d\n" !lt0_c !gt0_c !lt_c and + cmp_str2 = sprintf "le0 = %d\nge0 = %d\n" !le0_c !ge0_c and + cmp_str3 = sprintf "min = %d\nmin_max = %d\n" !min_c !min_max_c and + le_str = sprintf "le = %d (le_hash = %d)\n" !le_c (len le_table) and + max_str = sprintf "max = %d (max_hash = %d)\n" !max_c (len max_table) and + mul_str = sprintf "mul_lo = %d, mul_hi = %d (mul_hash = %d)\n" !mul_lo_c !mul_hi_c (len mul_table) and + div_str = sprintf "div_lo = %d, div_hi = %d (div_hash = %d)\n" !div_lo_c !div_hi_c (len div_table) and + add_str = sprintf "add_lo = %d, add_hi = %d (add_hash = %d)\n" !add_lo_c !add_hi_c (len add_table) and + sub_str = sprintf "sub_lo = %d, sub_hi = %d (sub_hash = %d)\n" !sub_lo_c !sub_hi_c (len sub_table) and + sqrt_str = sprintf "sqrt_lo = %d, sqrt_hi = %d (sqrt_hash = %d)\n" !sqrt_lo_c !sqrt_hi_c (len sqrt_table) in + print_string (cmp_str1 ^ cmp_str2 ^ cmp_str3 ^ + le_str ^ max_str ^ mul_str ^ div_str ^ add_str ^ sub_str ^ sqrt_str);; + + +(* lt0 *) +let float_lt0 = + if !Arith_options.float_cached then + fun tm -> + let _ = lt0_c := !lt0_c + 1 in + Float_ops.float_lt0 tm + else + Float_ops.float_lt0;; + +(* gt0 *) +let float_gt0 = + if !Arith_options.float_cached then + fun tm -> + let _ = gt0_c := !gt0_c + 1 in + Float_ops.float_gt0 tm + else + Float_ops.float_gt0;; + +(* lt *) +let float_lt = + if !Arith_options.float_cached then + fun tm1 tm2 -> + let _ = lt_c := !lt_c + 1 in + Float_ops.float_lt tm1 tm2 + else + Float_ops.float_lt;; + +(* le0 *) +let float_le0 = + if !Arith_options.float_cached then + fun tm -> + let _ = le0_c := !le0_c + 1 in + Float_ops.float_le0 tm + else + Float_ops.float_le0;; + +(* ge0 *) +let float_ge0 = + if !Arith_options.float_cached then + fun tm -> + let _ = ge0_c := !ge0_c + 1 in + Float_ops.float_ge0 tm + else + Float_ops.float_ge0;; + +(* min *) +let float_min = + if !Arith_options.float_cached then + fun tm1 tm2 -> + let _ = min_c := !min_c + 1 in + Float_ops.float_min tm1 tm2 + else + Float_ops.float_min;; + +(* min_max *) +let float_min_max = + if !Arith_options.float_cached then + fun tm1 tm2 -> + let _ = min_max_c := !min_max_c + 1 in + Float_ops.float_min_max tm1 tm2 + else + Float_ops.float_min_max;; + + +(***************) +let float_hash tm = + let s, n_tm, e_tm = dest_float tm in + s ^ (Arith_cache.num_tm_hash n_tm) ^ "e" ^ (Arith_cache.num_tm_hash e_tm);; + +let float_op_hash pp tm1 tm2 = + string_of_int pp ^ float_hash tm1 ^ "x" ^ float_hash tm2;; + +let float_op_hash1 pp tm = + string_of_int pp ^ float_hash tm;; + + +(* le *) +let float_le = + if !Arith_options.float_cached then + fun tm1 tm2 -> + let _ = le_c := !le_c + 1 in + let hash = float_op_hash 0 tm1 tm2 in + try + Hashtbl.find le_table hash + with Not_found -> + let result = Float_ops.float_le tm1 tm2 in + let _ = my_add le_table hash result in + result + else + Float_ops.float_le;; + +(* max *) +let float_max = + if !Arith_options.float_cached then + fun tm1 tm2 -> + let _ = max_c := !max_c + 1 in + let hash = float_op_hash 0 tm1 tm2 in + try + Hashtbl.find max_table hash + with Not_found -> + let result = Float_ops.float_max tm1 tm2 in + let _ = my_add max_table hash result in + result + else + Float_ops.float_max;; + +(* mul_eq *) +let float_mul_eq = Float_ops.float_mul_eq;; + +(* mul_lo *) +let float_mul_lo = + if !Arith_options.float_cached then + fun pp tm1 tm2 -> + let _ = mul_lo_c := !mul_lo_c + 1 in + let hash = "lo" ^ float_op_hash pp tm1 tm2 in + try + Hashtbl.find mul_table hash + with Not_found -> + let result = Float_ops.float_mul_lo pp tm1 tm2 in + let _ = my_add mul_table hash result in + let _ = if !Arith_options.float_log then log_binop "mul_lo" pp tm1 tm2 in + result + else + if !Arith_options.float_log then + fun pp tm1 tm2 -> + let _ = log_binop "mul_lo" pp tm1 tm2 in + Float_ops.float_mul_lo pp tm1 tm2 + else + Float_ops.float_mul_lo;; + +(* mul_hi *) +let float_mul_hi = + if !Arith_options.float_cached then + fun pp tm1 tm2 -> + let _ = mul_hi_c := !mul_hi_c + 1 in + let hash = "hi" ^ float_op_hash pp tm1 tm2 in + try + Hashtbl.find mul_table hash + with Not_found -> + let result = Float_ops.float_mul_hi pp tm1 tm2 in + let _ = my_add mul_table hash result in + let _ = if !Arith_options.float_log then log_binop "mul_hi" pp tm1 tm2 in + result + else + if !Arith_options.float_log then + fun pp tm1 tm2 -> + let _ = log_binop "mul_hi" pp tm1 tm2 in + Float_ops.float_mul_hi pp tm1 tm2 + else + Float_ops.float_mul_hi;; + +(* div_lo *) +let float_div_lo = + if !Arith_options.float_cached then + fun pp tm1 tm2 -> + let _ = div_lo_c := !div_lo_c + 1 in + let hash = "lo" ^ float_op_hash pp tm1 tm2 in + try + Hashtbl.find div_table hash + with Not_found -> + let result = Float_ops.float_div_lo pp tm1 tm2 in + let _ = my_add div_table hash result in + let _ = if !Arith_options.float_log then log_binop "div_lo" pp tm1 tm2 in + result + else + if !Arith_options.float_log then + fun pp tm1 tm2 -> + let _ = log_binop "div_lo" pp tm1 tm2 in + Float_ops.float_div_lo pp tm1 tm2 + else + Float_ops.float_div_lo;; + +(* div_hi *) +let float_div_hi = + if !Arith_options.float_cached then + fun pp tm1 tm2 -> + let _ = div_hi_c := !div_hi_c + 1 in + let hash = "hi" ^ float_op_hash pp tm1 tm2 in + try + Hashtbl.find div_table hash + with Not_found -> + let result = Float_ops.float_div_hi pp tm1 tm2 in + let _ = my_add div_table hash result in + let _ = if !Arith_options.float_log then log_binop "div_hi" pp tm1 tm2 in + result + else + if !Arith_options.float_log then + fun pp tm1 tm2 -> + let _ = log_binop "div_hi" pp tm1 tm2 in + Float_ops.float_div_hi pp tm1 tm2 + else + Float_ops.float_div_hi;; + +(* add_lo *) +let float_add_lo = + if !Arith_options.float_cached then + fun pp tm1 tm2 -> + let _ = add_lo_c := !add_lo_c + 1 in + let hash = "lo" ^ float_op_hash pp tm1 tm2 in + try + Hashtbl.find add_table hash + with Not_found -> + let result = Float_ops.float_add_lo pp tm1 tm2 in + let _ = my_add add_table hash result in + let _ = if !Arith_options.float_log then log_binop "add_lo" pp tm1 tm2 in + result + else + if !Arith_options.float_log then + fun pp tm1 tm2 -> + let _ = log_binop "add_lo" pp tm1 tm2 in + Float_ops.float_add_lo pp tm1 tm2 + else + Float_ops.float_add_lo;; + +(* add_hi *) +let float_add_hi = + if !Arith_options.float_cached then + fun pp tm1 tm2 -> + let _ = add_hi_c := !add_hi_c + 1 in + let hash = "hi" ^ float_op_hash pp tm1 tm2 in + try + Hashtbl.find add_table hash + with Not_found -> + let result = Float_ops.float_add_hi pp tm1 tm2 in + let _ = my_add add_table hash result in + let _ = if !Arith_options.float_log then log_binop "add_hi" pp tm1 tm2 in + result + else + if !Arith_options.float_log then + fun pp tm1 tm2 -> + let _ = log_binop "add_hi" pp tm1 tm2 in + Float_ops.float_add_hi pp tm1 tm2 + else + Float_ops.float_add_hi;; + +(* sub_lo *) +let float_sub_lo = + if !Arith_options.float_cached then + fun pp tm1 tm2 -> + let _ = sub_lo_c := !sub_lo_c + 1 in + let hash = "lo" ^ float_op_hash pp tm1 tm2 in + try + Hashtbl.find sub_table hash + with Not_found -> + let result = Float_ops.float_sub_lo pp tm1 tm2 in + let _ = my_add sub_table hash result in + let _ = if !Arith_options.float_log then log_binop "sub_lo" pp tm1 tm2 in + result + else + if !Arith_options.float_log then + fun pp tm1 tm2 -> + let _ = log_binop "sub_lo" pp tm1 tm2 in + Float_ops.float_sub_lo pp tm1 tm2 + else + Float_ops.float_sub_lo;; + +(* sub_hi *) +let float_sub_hi = + if !Arith_options.float_cached then + fun pp tm1 tm2 -> + let _ = sub_hi_c := !sub_hi_c + 1 in + let hash = "hi" ^ float_op_hash pp tm1 tm2 in + try + Hashtbl.find sub_table hash + with Not_found -> + let result = Float_ops.float_sub_hi pp tm1 tm2 in + let _ = my_add sub_table hash result in + let _ = if !Arith_options.float_log then log_binop "sub_hi" pp tm1 tm2 in + result + else + if !Arith_options.float_log then + fun pp tm1 tm2 -> + let _ = log_binop "sub_hi" pp tm1 tm2 in + Float_ops.float_sub_hi pp tm1 tm2 + else + Float_ops.float_sub_hi;; + +(* sqrt_lo *) +let float_sqrt_lo = + if !Arith_options.float_cached then + fun pp tm -> + let _ = sqrt_lo_c := !sqrt_lo_c + 1 in + let hash = "lo" ^ float_op_hash1 pp tm in + try + Hashtbl.find sqrt_table hash + with Not_found -> + let result = Float_ops.float_sqrt_lo pp tm in + let _ = my_add sqrt_table hash result in + result + else + Float_ops.float_sqrt_lo;; + +(* sqrt_hi *) +let float_sqrt_hi = + if !Arith_options.float_cached then + fun pp tm -> + let _ = sqrt_hi_c := !sqrt_hi_c + 1 in + let hash = "hi" ^ float_op_hash1 pp tm in + try + Hashtbl.find sqrt_table hash + with Not_found -> + let result = Float_ops.float_sqrt_hi pp tm in + let _ = my_add sqrt_table hash result in + result + else + Float_ops.float_sqrt_hi;; + + + +(* ------------------------------------------------ *) +(* Special comparison operations *) +(* ------------------------------------------------ *) + +(* For a given floating-point term returns its components as OCaml objects *) +let dest_float_raw f_tm = + let s, n_tm, e_tm = dest_float f_tm in + let sign = if s = "T" then true else false in + sign, raw_dest_hash n_tm, raw_dest_hash e_tm;; + +(* Compares raw representations of two floating-point numbers f1 and f2. *) +(* Returns -1 if f1 < f2, 0 if f1 = f2, and 1 if f1 > f2. *) +let compare_floats_raw = + let denorm n e = n */ (Num.num_of_int Arith_num.arith_base **/ e) in + let cmp n1 e1 n2 e2 = + if e1 <= e2 then + compare_num n1 (denorm n2 (e2 -/ e1)) + else + compare_num (denorm n1 (e1 -/ e2)) n2 in + fun f1_tm f2_tm -> + let s1, n1, e1 = dest_float_raw f1_tm and + s2, n2, e2 = dest_float_raw f2_tm in + let n1' = if s1 then minus_num n1 else n1 and + n2' = if s2 then minus_num n2 else n2 in + cmp n1' e1 n2' e2;; + +(* Returns (true, |- f1 <= f2 <=> T) when f1 <= f2; returns (false, |- f1 <= f2 <=> F) otherwise *) +let float_prove_le f1_tm f2_tm = + let le_th = float_le f1_tm f2_tm in + let flag = (fst o dest_const o rand o concl) le_th = "T" in + flag, le_th;; + +(* Returns (true, |- f1 < f2 <=> T) when f1 < f2; returns (false, |- f1 < f2 <=> F) otherwise *) +let float_prove_lt f1_tm f2_tm = + let lt_th = float_lt f1_tm f2_tm in + let flag = (fst o dest_const o rand o concl) lt_th = "T" in + flag, lt_th;; + +(* +(* Arith_num.raw_dest_hash is too slow (need to retest without caching) *) +let float_prove_le f1_tm f2_tm = + let r = compare_floats_raw f1_tm f2_tm in + if r <= 0 then + true, float_le f1_tm f2_tm + else + false, TRUTH;; +*) + +(* Compares y and an interval a <= x <= b (|- interval_arith x (a, b)). *) +(* Returns (true, |- y <= x) if y <= a. *) +(* Returns (false, |- y <= a <=> F) if y > a. *) +let float_prove_le_interval = + let le_interval = (UNDISCH_ALL o prove)(`interval_arith x (a,b) ==> (y <= a <=> T) ==> y <= x`, + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC) in + fun y_tm x_th -> + let x_tm, bounds = dest_interval_arith (concl x_th) in + let ltm, b_tm = dest_comb bounds in + let a_tm = rand ltm in + let flag, le_th = float_prove_le y_tm a_tm in + if flag then + let th0 = INST[x_tm, x_var_real; y_tm, y_var_real; + a_tm, a_var_real; b_tm, b_var_real] le_interval in + true, MY_PROVE_HYP x_th (MY_PROVE_HYP le_th th0) + else + false, le_th;; + +(* Compares y and an interval a <= x <= b (|- interval_arith x (a, b)). *) +(* Returns (true, |- x <= y) if b <= y. *) +(* Returns (false, |- b <= y <=> F) if b > y. *) +let float_prove_ge_interval = + let ge_interval = (UNDISCH_ALL o prove)(`interval_arith x (a,b) ==> (b <= y <=> T) ==> x <= y`, + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC) in + fun y_tm x_th -> + let x_tm, bounds = dest_interval_arith (concl x_th) in + let ltm, b_tm = dest_comb bounds in + let a_tm = rand ltm in + let flag, le_th = float_prove_le b_tm y_tm in + if flag then + let th0 = INST[x_tm, x_var_real; y_tm, y_var_real; + a_tm, a_var_real; b_tm, b_var_real] ge_interval in + true, MY_PROVE_HYP x_th (MY_PROVE_HYP le_th th0) + else + false, le_th;; + +(* Compares y and an interval a <= x <= b (|- interval_arith x (a, b)). *) +(* Returns (true, |- y < x) if y < a. *) +(* Returns (false, |- y < a <=> F) if y >= a. *) +let float_prove_lt_interval = + let lt_interval = (UNDISCH_ALL o prove)(`interval_arith x (a,b) ==> (y < a <=> T) ==> y < x`, + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC) in + fun y_tm x_th -> + let x_tm, bounds = dest_interval_arith (concl x_th) in + let ltm, b_tm = dest_comb bounds in + let a_tm = rand ltm in + let flag, lt_th = float_prove_lt y_tm a_tm in + if flag then + let th0 = INST[x_tm, x_var_real; y_tm, y_var_real; + a_tm, a_var_real; b_tm, b_var_real] lt_interval in + true, MY_PROVE_HYP x_th (MY_PROVE_HYP lt_th th0) + else + false, lt_th;; + +(* Compares y and an interval a <= x <= b (|- interval_arith x (a, b)). *) +(* Returns (true, |- x < y) if b < y. *) +(* Returns (false, |- b < y <=> F) if b >= y. *) +let float_prove_gt_interval = + let gt_interval = (UNDISCH_ALL o prove)(`interval_arith x (a,b) ==> (b < y <=> T) ==> x < y`, + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC) in + fun y_tm x_th -> + let x_tm, bounds = dest_interval_arith (concl x_th) in + let ltm, b_tm = dest_comb bounds in + let a_tm = rand ltm in + let flag, lt_th = float_prove_lt b_tm y_tm in + if flag then + let th0 = INST[x_tm, x_var_real; y_tm, y_var_real; + a_tm, a_var_real; b_tm, b_var_real] gt_interval in + true, MY_PROVE_HYP x_th (MY_PROVE_HYP lt_th th0) + else + false, lt_th;; + +(* Compares y and an interval a <= x <= b (|- interval_arith x (a, b)). *) +(* Returns (-1, |- y <= x) if y <= a. *) +(* Returns (1, |- x <= y) if b <= y. *) +(* Returns (0, |- T ) otherwise. *) +let float_compare_interval = + let le_interval = (UNDISCH_ALL o prove)(`interval_arith x (a,b) ==> (y <= a <=> T) ==> y <= x`, + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC) in + let ge_interval = (UNDISCH_ALL o prove)(`interval_arith x (a,b) ==> (b <= y <=> T) ==> x <= y`, + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC) in + fun y_tm x_th -> + let x_tm, bounds = dest_interval_arith (concl x_th) in + let ltm, b_tm = dest_comb bounds in + let a_tm = rand ltm in + let inst th le_th = + let th0 = INST[x_tm, x_var_real; y_tm, y_var_real; + a_tm, a_var_real; b_tm, b_var_real] th in + MY_PROVE_HYP x_th (MY_PROVE_HYP le_th th0) in + let flag, le_th = float_prove_le y_tm a_tm in + if flag then + -1, inst le_interval le_th + else + let flag, le_th = float_prove_le b_tm y_tm in + if flag then + 1, inst ge_interval le_th + else + 0, TRUTH;; + + +(* ------------------------------------------------ *) +(* Float interval operations *) +(* ------------------------------------------------ *) + +let FLOAT_OF_NUM' = (SPEC_ALL o REWRITE_RULE[min_exp_def]) FLOAT_OF_NUM;; + +let FLOAT_INTERVAL_OF_NUM = (NUMERALS_TO_NUM o REWRITE_RULE[min_exp_def] o prove)(`interval_arith (&n) (float_num F n min_exp, float_num F n min_exp)`, + REWRITE_TAC[FLOAT_OF_NUM; CONST_INTERVAL]);; + +let FLOAT_F_bound' = (UNDISCH_ALL o SPEC_ALL) FLOAT_F_bound;; + +let FLOAT_T_bound' = (UNDISCH_ALL o SPEC_ALL) FLOAT_T_bound;; + +(* interval_arith x (float_num s1 n1 e1, float_num s2 n2 e2) -> x, float_num s1 n1 e1, float_num s2 n2 e2 *) +let dest_float_interval tm = + let ltm, rtm = dest_comb tm in + let f1, f2 = dest_pair rtm in + rand ltm, f1, f2;; + +let mk_float_interval_small_num n = + let n_tm0 = mk_small_numeral n in + let n_th = NUMERAL_TO_NUM_CONV n_tm0 in + let n_tm = rand(rand(concl n_th)) in + let n_th1 = TRANS n_th (INST[n_tm, n_var_num] NUM_REMOVE) in + let th1 = AP_TERM amp_op_real n_th1 in + let int_th = INST[n_tm, n_var_num] FLOAT_INTERVAL_OF_NUM in + let rtm = rand(concl int_th) in + EQ_MP (SYM (AP_THM (AP_TERM interval_const th1) rtm)) int_th;; + +let mk_float_interval_num n = + let n_tm0 = mk_numeral n in + let n_th = NUMERAL_TO_NUM_CONV n_tm0 in + let n_tm = rand(rand(concl n_th)) in + let n_th1 = TRANS n_th (INST[n_tm, n_var_num] NUM_REMOVE) in + let th1 = AP_TERM amp_op_real n_th1 in + let int_th = INST[n_tm, n_var_num] FLOAT_INTERVAL_OF_NUM in + let rtm = rand(concl int_th) in + EQ_MP (SYM (AP_THM (AP_TERM interval_const th1) rtm)) int_th;; + + +(* Returns the lower bound for the given float *) +let float_lo p tm = + let s, n_tm, e_tm = dest_float tm in + if s = "F" then + let num_exp_tm = mk_num_exp n_tm e_tm in + let th0 = num_exp_lo p num_exp_tm in + let ltm, e1_tm = dest_comb(lhand(concl th0)) in + let n1_tm = rand ltm in + let th1 = INST[n1_tm, n1_var_num; e1_tm, e1_var_num; n_tm, n2_var_num; e_tm, e2_var_num] FLOAT_F_bound' in + MY_PROVE_HYP th0 th1 + else + let num_exp_tm = mk_num_exp n_tm e_tm in + let th0 = num_exp_hi p num_exp_tm in + let ltm, e1_tm = dest_comb(rand(concl th0)) in + let n1_tm = rand ltm in + let th1 = INST[n_tm, n1_var_num; e_tm, e1_var_num; n1_tm, n2_var_num; e1_tm, e2_var_num] FLOAT_T_bound' in + MY_PROVE_HYP th0 th1;; + +(* Returns the upper bound for the given float *) +let float_hi p tm = + let s, n_tm, e_tm = dest_float tm in + if s = "F" then + let num_exp_tm = mk_num_exp n_tm e_tm in + let th0 = num_exp_hi p num_exp_tm in + let ltm, e2_tm = dest_comb(rand(concl th0)) in + let n2_tm = rand ltm in + let th1 = INST[n_tm, n1_var_num; e_tm, e1_var_num; n2_tm, n2_var_num; e2_tm, e2_var_num] FLOAT_F_bound' in + MY_PROVE_HYP th0 th1 + else + let num_exp_tm = mk_num_exp n_tm e_tm in + let th0 = num_exp_lo p num_exp_tm in + let ltm, e1_tm = dest_comb(lhand(concl th0)) in + let n1_tm = rand ltm in + let th1 = INST[n1_tm, n1_var_num; e1_tm, e1_var_num; n_tm, n2_var_num; e_tm, e2_var_num] FLOAT_T_bound' in + MY_PROVE_HYP th0 th1;; + + +(* Approximates the given interval with p-digits floating point numbers *) +let float_interval_round p th = + let x_tm, f1, f2 = dest_float_interval (concl th) in + let lo_th = float_lo p f1 in + let hi_th = float_hi p f2 in + let lo_tm = lhand(concl lo_th) in + let hi_tm = rand(concl hi_th) in + let th0 = INST[x_tm, x_var_real; f1, lo_var_real; f2, hi_var_real; lo_tm, a_var_real; hi_tm, b_var_real] APPROX_INTERVAL' in + MY_PROVE_HYP lo_th (MY_PROVE_HYP hi_th (MY_PROVE_HYP th th0));; + + +(* ------------------------------------------------ *) +(* Absolute value of an interval *) +(* ------------------------------------------------ *) + +let RULE = UNDISCH_ALL o NUMERALS_TO_NUM o PURE_REWRITE_RULE[GSYM IMP_IMP; min_exp_def];; + +let float_interval_abs_ft = (RULE o prove) + (`interval_arith x (float_num F n1 e1, float_num T n2 e2) + ==> interval_arith (abs x) (float_num F 0 min_exp, float_num F 0 min_exp)`, + DISCH_THEN (MP_TAC o MATCH_MP FLOAT_INTERVAL_FT_IMP_0) THEN + REWRITE_TAC[FLOAT_0; interval_arith] THEN REAL_ARITH_TAC);; + +let float_interval_abs_ff = (RULE o prove) + (`interval_arith x (float_num F n1 e1, float_num F n2 e2) + ==> interval_arith (abs x) (float_num F n1 e1, float_num F n2 e2)`, + MP_TAC (SPECL[`n1:num`; `e1:num`] FLOAT_F_POS) THEN + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; + +let float_interval_abs_tt = (RULE o prove) + (`interval_arith x (float_num T n1 e1, float_num T n2 e2) + ==> interval_arith (abs x) (float_num F n2 e2, float_num F n1 e1)`, + MP_TAC ((GSYM o SPEC `T`) FLOAT_NEG) THEN SIMP_TAC[] THEN DISCH_TAC THEN + MP_TAC (SPECL[`n2:num`; `e2:num`] FLOAT_T_NEG) THEN + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; + +let float_interval_abs_tf = (RULE o prove) + (`interval_arith x (float_num T n1 e1, float_num F n2 e2) /\ + max (float_num F n1 e1) (float_num F n2 e2) = t + ==> interval_arith (abs x) (float_num F 0 min_exp, t)`, + MP_TAC ((GSYM o SPEC `F`) FLOAT_NEG) THEN SIMP_TAC[] THEN DISCH_TAC THEN + REWRITE_TAC[interval_arith; FLOAT_0] THEN REAL_ARITH_TAC);; + +let float_interval_abs th = + let x, lo, hi = dest_float_interval (concl th) in + let s1, n1, e1 = dest_float lo and + s2, n2, e2 = dest_float hi in + let inst = INST[x, x_var_real; + n1, n1_var_num; e1, e1_var_num; + n2, n2_var_num; e2, e2_var_num] in + if s1 = s2 then + if s1 = "F" then + (* FF *) + MY_PROVE_HYP th (inst float_interval_abs_ff) + else + (* TT *) + MY_PROVE_HYP th (inst float_interval_abs_tt) + else + if s1 = "F" then + (* FT *) + MY_PROVE_HYP th (inst float_interval_abs_ft) + else + let max_th = float_max (make_float "F" n1 e1) hi in + let t_tm = rand (concl max_th) in + let th0 = INST[x, x_var_real; t_tm, t_var_real; + n1, n1_var_num; e1, e1_var_num; + n2, n2_var_num; e2, e2_var_num] float_interval_abs_tf in + MY_PROVE_HYP max_th (MY_PROVE_HYP th th0);; + + +(****************************************) +(* float_interval_lt *) + +let FLOAT_INTERVAL_LT = prove(`interval_arith x (lo1, hi1) /\ interval_arith y (lo2, hi2) /\ hi1 < lo2 + ==> x < y`, + REWRITE_TAC[interval_arith] THEN + REAL_ARITH_TAC);; + + +(****************************************) +(* float_interval_neg *) + +let FLOAT_INTERVAL_NEG = prove(`!s1 s2. interval_arith x (float_num s1 n1 e1, float_num s2 n2 e2) + ==> interval_arith (--x) (float_num (~s2) n2 e2, float_num (~s1) n1 e1)`, + REPEAT GEN_TAC THEN + DISCH_THEN (fun th -> MP_TAC (MATCH_MP INTERVAL_NEG th)) THEN + SIMP_TAC[FLOAT_NEG]);; + +let FLOAT_INTERVAL_NEG_FF = (UNDISCH_ALL o REWRITE_RULE[] o SPECL[`F`; `F`]) FLOAT_INTERVAL_NEG;; +let FLOAT_INTERVAL_NEG_FT = (UNDISCH_ALL o REWRITE_RULE[] o SPECL[`F`; `T`]) FLOAT_INTERVAL_NEG;; +let FLOAT_INTERVAL_NEG_TF = (UNDISCH_ALL o REWRITE_RULE[] o SPECL[`T`; `F`]) FLOAT_INTERVAL_NEG;; +let FLOAT_INTERVAL_NEG_TT = (UNDISCH_ALL o REWRITE_RULE[] o SPECL[`T`; `T`]) FLOAT_INTERVAL_NEG;; + +(* |- interval x (float s1 n1 e1, float s2 n2 e2) -> + |- interval (--x) (float ~s2 n2 e2, float ~s1 n1 e1 *) +let float_interval_neg th = + let x_tm, f1, f2 = dest_float_interval (concl th) in + let s1, n1_tm, e1_tm = dest_float f1 in + let s2, n2_tm, e2_tm = dest_float f2 in + let inst = INST[x_tm, x_var_real; n1_tm, n1_var_num; e1_tm, e1_var_num; + n2_tm, n2_var_num; e2_tm, e2_var_num] in + let th0 = + if s1 = "F" then + if s2 = "F" then + inst FLOAT_INTERVAL_NEG_FF + else + inst FLOAT_INTERVAL_NEG_FT + else + if s2 = "F" then + inst FLOAT_INTERVAL_NEG_TF + else + inst FLOAT_INTERVAL_NEG_TT in + MY_PROVE_HYP th th0;; + + +(***********************************************) +(* float_interval_mul *) + +let f1_1_var = `f1_1:real` and + f1_2_var = `f1_2:real` and + f2_1_var = `f2_1:real` and + f2_2_var = `f2_2:real`;; + + +(* FT_xx *) +let FLOAT_INTERVAL_MUL_FT_xx = (UNDISCH_ALL o NUMERALS_TO_NUM o REWRITE_RULE[GSYM IMP_IMP; min_exp_def] o prove)( + `interval_arith x (float_num F n1 e1, float_num T n2 e2) + ==> interval_arith (x * y) (float_num F 0 min_exp, float_num F 0 min_exp)`, + STRIP_TAC THEN + FIRST_X_ASSUM (fun th -> REWRITE_TAC[MATCH_MP FLOAT_INTERVAL_FT_IMP_0 th]) THEN + REWRITE_TAC[REAL_MUL_LZERO; interval_arith] THEN + MP_TAC (GEN_ALL (SPECL [`s:bool`; `0`] FLOAT_EQ_0)) THEN SIMP_TAC[REAL_LE_REFL]);; + + +(* xx_FT *) +let FLOAT_INTERVAL_MUL_xx_FT = (UNDISCH_ALL o NUMERALS_TO_NUM o REWRITE_RULE[GSYM IMP_IMP; min_exp_def] o prove)( + `interval_arith y (float_num F m1 r1, float_num T m2 r2) + ==> interval_arith (x * y) (float_num F 0 min_exp, float_num F 0 min_exp)`, + STRIP_TAC THEN + FIRST_X_ASSUM (fun th -> REWRITE_TAC[MATCH_MP FLOAT_INTERVAL_FT_IMP_0 th]) THEN + REWRITE_TAC[REAL_MUL_RZERO; interval_arith] THEN + MP_TAC (GEN_ALL (SPECL [`s:bool`; `0`] FLOAT_EQ_0)) THEN SIMP_TAC[REAL_LE_REFL]);; + + + +(* FF_FF *) +let FLOAT_INTERVAL_MUL_FF_FF = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( + `interval_arith x (float_num F n1 e1, float_num F n2 e2) /\ + interval_arith y (float_num F m1 r1, float_num F m2 r2) /\ + f1 <= float_num F n1 e1 * float_num F m1 r1 /\ + float_num F n2 e2 * float_num F m2 r2 <= f2 + ==> interval_arith (x * y) (f1, f2)`, + MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN + SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` MP_TAC THENL + [ + MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN + REWRITE_TAC[FLOAT_F_POS]; + ALL_TAC + ] THEN + REPEAT (POP_ASSUM (fun th -> ALL_TAC)) THEN + REWRITE_TAC[interval_arith] THEN + REPEAT STRIP_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `a * c:real` THEN + ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN + ASM_REWRITE_TAC[]; + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `b * d:real` THEN + ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN + ASM_REWRITE_TAC[] THEN + CONJ_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THENL + [ + EXISTS_TAC `a:real` THEN ASM_REWRITE_TAC[]; + EXISTS_TAC `c:real` THEN ASM_REWRITE_TAC[] + ] + ]);; + +(* TT_TT *) +let FLOAT_INTERVAL_MUL_TT_TT = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( + `interval_arith x (float_num T n1 e1, float_num T n2 e2) /\ + interval_arith y (float_num T m1 r1, float_num T m2 r2) /\ + f1 <= float_num T n2 e2 * float_num T m2 r2 /\ + float_num T n1 e1 * float_num T m1 r1 <= f2 + ==> interval_arith (x * y) (f1, f2)`, + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[interval_arith] THEN + MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN + SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` MP_TAC THENL + [ + MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN + REWRITE_TAC[FLOAT_F_POS]; + ALL_TAC + ] THEN + REWRITE_TAC[REAL_NEG_MUL2] THEN + REPEAT STRIP_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `b * d:real` THEN + ASM_REWRITE_TAC[] THEN + ONCE_REWRITE_TAC[REAL_ARITH `a <= x * y <=> a <= --x * --y`] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN + ONCE_REWRITE_TAC[REAL_ARITH `b <= --x <=> x <= --b`] THEN + ASM_REWRITE_TAC[]; + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `a * c:real` THEN + ASM_REWRITE_TAC[] THEN + ONCE_REWRITE_TAC[REAL_ARITH `x * y <= a <=> --x * --y <= a`] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN + ONCE_REWRITE_TAC[REAL_ARITH `--x <= c <=> --c <= x`] THEN + ASM_REWRITE_TAC[] THEN + ASSUME_TAC (REAL_ARITH `!b x. &0 <= b /\ x <= --b ==> &0 <= --x`) THEN + CONJ_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THENL + [ + EXISTS_TAC `b:real` THEN ASM_REWRITE_TAC[]; + EXISTS_TAC `d:real` THEN ASM_REWRITE_TAC[] + ] + ]);; + +(* FF_TT *) +let FLOAT_INTERVAL_MUL_FF_TT = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( + `interval_arith x (float_num F n1 e1, float_num F n2 e2) /\ + interval_arith y (float_num T m1 r1, float_num T m2 r2) /\ + f1 <= float_num F n2 e2 * float_num T m1 r1 /\ + float_num F n1 e1 * float_num T m2 r2 <= f2 + ==> interval_arith (x * y) (f1, f2)`, + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[interval_arith] THEN + MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN + SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` MP_TAC THENL + [ + MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN + REWRITE_TAC[FLOAT_F_POS]; + ALL_TAC + ] THEN + + REPEAT STRIP_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `b * --c` THEN + ASM_REWRITE_TAC[] THEN + ONCE_REWRITE_TAC[REAL_ARITH `b * --c <= x * y <=> x * --y <= b * c`] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN + ONCE_REWRITE_TAC[REAL_ARITH `--y <= c <=> --c <= y`] THEN + ASM_REWRITE_TAC[] THEN + CONJ_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THENL + [ + EXISTS_TAC `a:real` THEN ASM_REWRITE_TAC[]; + EXISTS_TAC `d:real` THEN + ONCE_REWRITE_TAC[REAL_ARITH `d <= --y <=> y <= --d`] THEN + ASM_REWRITE_TAC[] + ]; + + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `a * --d` THEN + ASM_REWRITE_TAC[] THEN + ONCE_REWRITE_TAC[REAL_ARITH `x * y <= a * --d <=> a * d <= x * --y`] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN + ONCE_REWRITE_TAC[REAL_ARITH `d <= --y <=> y <= --d`] THEN + ASM_REWRITE_TAC[] + ]);; + +(* TT_FF *) +let FLOAT_INTERVAL_MUL_TT_FF = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( + `interval_arith x (float_num T n1 e1, float_num T n2 e2) /\ + interval_arith y (float_num F m1 r1, float_num F m2 r2) /\ + f1 <= float_num T n1 e1 * float_num F m2 r2 /\ + float_num T n2 e2 * float_num F m1 r1 <= f2 + ==> interval_arith (x * y) (f1, f2)`, + STRIP_TAC THEN + MP_TAC ((GEN_ALL o DISCH_ALL) FLOAT_INTERVAL_MUL_FF_TT) THEN + DISCH_THEN (MP_TAC o SPECL[`n1:num`; `e1:num`; `n2:num`; `e2:num`; `m1:num`; `r1:num`; `m2:num`; `r2:num`]) THEN + DISCH_THEN (MP_TAC o SPECL[`y:real`; `x:real`; `f1:real`; `f2:real`]) THEN + REPEAT (POP_ASSUM MP_TAC) THEN + SIMP_TAC[REAL_MUL_AC]);; + +(* TF_FF *) +let FLOAT_INTERVAL_MUL_TF_FF = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( + `interval_arith x (float_num T n1 e1, float_num F n2 e2) /\ + interval_arith y (float_num F m1 r1, float_num F m2 r2) /\ + f1 <= float_num T n1 e1 * float_num F m2 r2 /\ + float_num F n2 e2 * float_num F m2 r2 <= f2 + ==> interval_arith (x * y) (f1, f2)`, + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[interval_arith] THEN + MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN + STRIP_TAC THEN + SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` ASSUME_TAC THENL + [ + MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN + REWRITE_TAC[FLOAT_F_POS]; + ALL_TAC + ] THEN + + SUBGOAL_THEN `&0 <= y` ASSUME_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `c:real` THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + + CONJ_TAC THENL + [ + DISJ_CASES_TAC (REAL_ARITH `&0 <= x \/ &0 <= --x`) THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN + CONJ_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * d` THEN + ASM_REWRITE_TAC[REAL_ARITH `--a * d <= &0 <=> &0 <= a * d`] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[] + ]; + ALL_TAC + ] THEN + + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * d` THEN ASM_REWRITE_TAC[] THEN + ONCE_REWRITE_TAC[REAL_ARITH `--a * d <= x * y <=> --x * y <= a * d`] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN + ONCE_REWRITE_TAC[REAL_ARITH `--x <= a <=> --a <= x`] THEN ASM_REWRITE_TAC[]; + + DISJ_CASES_TAC (REAL_ARITH `&0 <= --x \/ &0 <= x`) THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN + CONJ_TAC THENL + [ + REWRITE_TAC[REAL_ARITH `x * y <= &0 <=> &0 <= --x * y`] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * d` THEN ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[] + ]; + ALL_TAC + ] THEN + + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * d` THEN ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_REWRITE_TAC[] + ]);; + +(* TF_TT *) +let FLOAT_INTERVAL_MUL_TF_TT = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( + `interval_arith x (float_num T n1 e1, float_num F n2 e2) /\ + interval_arith y (float_num T m1 r1, float_num T m2 r2) /\ + f1 <= float_num F n2 e2 * float_num T m1 r1 /\ + float_num T n1 e1 * float_num T m1 r1 <= f2 + ==> interval_arith (x * y) (f1, f2)`, + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[interval_arith] THEN + MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN + STRIP_TAC THEN + SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` ASSUME_TAC THENL + [ + MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN + REWRITE_TAC[FLOAT_F_POS]; + ALL_TAC + ] THEN + + SUBGOAL_THEN `&0 <= --y` ASSUME_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `d:real` THEN + ONCE_REWRITE_TAC[REAL_ARITH `d <= --y <=> y <= --d`] THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + + CONJ_TAC THENL + [ + DISJ_CASES_TAC (REAL_ARITH `&0 <= --x \/ &0 <= x`) THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN + CONJ_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * --c` THEN + ASM_REWRITE_TAC[REAL_ARITH `b * --c <= &0 <=> &0 <= b * c`] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; + ONCE_REWRITE_TAC[REAL_ARITH `x * y = --x * --y`] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[] + ]; + ALL_TAC + ] THEN + + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * --c` THEN ASM_REWRITE_TAC[] THEN + ONCE_REWRITE_TAC[REAL_ARITH `b * --c <= x * y <=> x * --y <= b * c`] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN ONCE_REWRITE_TAC[REAL_ARITH `--y <= c <=> --c <= y`] THEN + ASM_REWRITE_TAC[]; + + DISJ_CASES_TAC (REAL_ARITH `&0 <= x \/ &0 <= --x`) THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN + CONJ_TAC THENL + [ + REWRITE_TAC[REAL_ARITH `x * y <= &0 <=> &0 <= x * --y`] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * --c` THEN ASM_REWRITE_TAC[REAL_NEG_MUL2] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[] + ]; + ALL_TAC + ] THEN + + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * --c` THEN ASM_REWRITE_TAC[REAL_NEG_MUL2] THEN + ONCE_REWRITE_TAC[REAL_ARITH `x * y <= a * c <=> --x * --y <= a * c`] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN ONCE_REWRITE_TAC[REAL_ARITH `--x <= a <=> --a <= x`] THEN + ASM_REWRITE_TAC[] + ]);; + +(* FF_TF *) +let FLOAT_INTERVAL_MUL_FF_TF = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( + `interval_arith x (float_num F n1 e1, float_num F n2 e2) /\ + interval_arith y (float_num T m1 r1, float_num F m2 r2) /\ + f1 <= float_num F n2 e2 * float_num T m1 r1 /\ + float_num F n2 e2 * float_num F m2 r2 <= f2 + ==> interval_arith (x * y) (f1, f2)`, + STRIP_TAC THEN + MP_TAC ((SPECL [`n1:num`; `e1:num`; `n2:num`; `e2:num`; `m1:num`; `r1:num`; `m2:num`; `r2:num`; `y:real`; `x:real`; `f1:real`; `f2:real`] o GEN_ALL o DISCH_ALL) FLOAT_INTERVAL_MUL_TF_FF) THEN + ASM_REWRITE_TAC[REAL_MUL_SYM] THEN DISCH_THEN MATCH_MP_TAC THEN + POP_ASSUM MP_TAC THEN REAL_ARITH_TAC);; + +(* TT_TF *) +let FLOAT_INTERVAL_MUL_TT_TF = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( + `interval_arith x (float_num T n1 e1, float_num T n2 e2) /\ + interval_arith y (float_num T m1 r1, float_num F m2 r2) /\ + f1 <= float_num T n1 e1 * float_num F m2 r2 /\ + float_num T n1 e1 * float_num T m1 r1 <= f2 + ==> interval_arith (x * y) (f1, f2)`, + STRIP_TAC THEN + MP_TAC ((SPECL [`n1:num`; `e1:num`; `n2:num`; `e2:num`; `m1:num`; `r1:num`; `m2:num`; `r2:num`; `y:real`; `x:real`; `f1:real`; `f2:real`] o GEN_ALL o DISCH_ALL) FLOAT_INTERVAL_MUL_TF_TT) THEN + ASM_REWRITE_TAC[REAL_MUL_SYM] THEN POP_ASSUM MP_TAC THEN POP_ASSUM MP_TAC THEN + SIMP_TAC[REAL_MUL_SYM]);; + +(* TF_TF *) +let FLOAT_INTERVAL_MUL_TF_TF = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( + `interval_arith x (float_num T n1 e1, float_num F n2 e2) /\ + interval_arith y (float_num T m1 r1, float_num F m2 r2) /\ + f1_1 <= float_num T n1 e1 * float_num F m2 r2 /\ + f1_2 <= float_num F n2 e2 * float_num T m1 r1 /\ + min f1_1 f1_2 = f1 /\ + float_num T n1 e1 * float_num T m1 r1 <= f2_1 /\ + float_num F n2 e2 * float_num F m2 r2 <= f2_2 /\ + max f2_1 f2_2 = f2 + ==> interval_arith (x * y) (f1, f2)`, + REWRITE_TAC[EQ_SYM_EQ; FLOAT_NEG_T] THEN + REWRITE_TAC[interval_arith] THEN + MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN + STRIP_TAC THEN + SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` ASSUME_TAC THENL + [ + MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN + REWRITE_TAC[FLOAT_F_POS]; + ALL_TAC + ] THEN + + DISJ_CASES_TAC (REAL_ARITH `&0 <= x \/ &0 <= --x`) THENL + [ + DISJ_CASES_TAC (REAL_ARITH `&0 <= y \/ &0 <= --y`) THENL + [ + CONJ_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN + CONJ_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f1_1:real` THEN + ASM_REWRITE_TAC[REAL_MIN_MIN] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * d` THEN + ASM_REWRITE_TAC[REAL_ARITH `--a * d <= &0 <=> &0 <= a * d`] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * d` THEN + CONJ_TAC THENL + [ + MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f2_2:real` THEN + ASM_REWRITE_TAC[REAL_MAX_MAX]; + ALL_TAC + ] THEN + + CONJ_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f1_2:real` THEN + ASM_REWRITE_TAC[REAL_MIN_MIN] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * --c` THEN + ONCE_REWRITE_TAC[REAL_ARITH `b * --c <= x * y <=> x * --y <= b * c`] THEN ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN ONCE_REWRITE_TAC[REAL_ARITH `--y <= c <=> --c <= y`] THEN + ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN + CONJ_TAC THENL + [ + REWRITE_TAC[REAL_ARITH `x * y <= &0 <=> &0 <= x * --y`] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f2_2:real` THEN + ASM_REWRITE_TAC[REAL_MAX_MAX] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * d` THEN ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + + DISJ_CASES_TAC (REAL_ARITH `&0 <= y \/ &0 <= --y`) THENL + [ + CONJ_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f1_1:real` THEN + ASM_REWRITE_TAC[REAL_MIN_MIN] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * d` THEN + ASM_REWRITE_TAC[] THEN ONCE_REWRITE_TAC[REAL_ARITH `--a * d <= x * y <=> --x * y <= a * d`] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN ONCE_REWRITE_TAC[REAL_ARITH `--x <= a <=> --a <= x`] THEN + ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN + CONJ_TAC THENL + [ + REWRITE_TAC[REAL_ARITH `x * y <= &0 <=> &0 <= --x * y`] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f2_2:real` THEN + ASM_REWRITE_TAC[REAL_MAX_MAX] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b * d` THEN ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + + CONJ_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN + CONJ_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f1_1:real` THEN + ASM_REWRITE_TAC[REAL_MIN_MIN] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * d` THEN + ASM_REWRITE_TAC[REAL_ARITH `--a * d <= &0 <=> &0 <= a * d`] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + ONCE_REWRITE_TAC[GSYM REAL_NEG_MUL2] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a * --c` THEN + CONJ_TAC THENL + [ + ONCE_REWRITE_TAC[GSYM REAL_NEG_MUL2] THEN REWRITE_TAC[REAL_NEG_NEG] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN ONCE_REWRITE_TAC[REAL_ARITH `--x <= a <=> --a <= x`] THEN + ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `f2_1:real` THEN + ASM_REWRITE_TAC[REAL_MAX_MAX]);; + + +(****************************) +let float_interval_mul = + let mul_ft_xx th1 x y n1 e1 n2 e2 = + let th0 = INST[x, x_var_real; y, y_var_real; + n1, n1_var_num; e1, e1_var_num; + n2, n2_var_num; e2, e2_var_num] FLOAT_INTERVAL_MUL_FT_xx in + MY_PROVE_HYP th1 th0 in + let mul_xx_ft th2 x y m1 r1 m2 r2 = + let th0 = INST[x, x_var_real; y, y_var_real; + m1, m1_var_num; r1, r1_var_num; + m2, m2_var_num; r2, r2_var_num] FLOAT_INTERVAL_MUL_xx_FT in + MY_PROVE_HYP th2 th0 in + + fun pp th1 th2 -> + let x, l_lo, l_hi = dest_float_interval (concl th1) and + y, r_lo, r_hi = dest_float_interval (concl th2) in + let s1, n1, e1 = dest_float l_lo and + s2, n2, e2 = dest_float l_hi and + s3, m1, r1 = dest_float r_lo and + s4, m2, r2 = dest_float r_hi in + + (* Special case 1 *) + if s1 <> s2 && s1 = "F" then + mul_ft_xx th1 x y n1 e1 n2 e2 + else if s3 <> s4 && s3 = "F" then + mul_xx_ft th2 x y m1 r1 m2 r2 + else + (* Special case 2 *) + if s1 <> s2 && s3 <> s4 then + let lo1, lo2 = float_mul_lo pp l_lo r_hi, float_mul_lo pp l_hi r_lo and + hi1, hi2 = float_mul_hi pp l_lo r_lo, float_mul_hi pp l_hi r_hi in + let f1_1 = (lhand o concl) lo1 and + f1_2 = (lhand o concl) lo2 and + f2_1 = (rand o concl) hi1 and + f2_2 = (rand o concl) hi2 in + let min_th = float_min f1_1 f1_2 and + max_th = float_max f2_1 f2_2 in + let f1_tm = (rand o concl) min_th and + f2_tm = (rand o concl) max_th in + let th0 = INST[x, x_var_real; n1, n1_var_num; e1, e1_var_num; + y, y_var_real; n2, n2_var_num; e2, e2_var_num; + m1, m1_var_num; r1, r1_var_num; + m2, m2_var_num; r2, r2_var_num; + f1_tm, f1_var_real; f2_tm, f2_var_real; + f1_1, f1_1_var; f1_2, f1_2_var; + f2_1, f2_1_var; f2_2, f2_2_var] FLOAT_INTERVAL_MUL_TF_TF in + (MY_PROVE_HYP min_th o MY_PROVE_HYP max_th o MY_PROVE_HYP lo1 o MY_PROVE_HYP lo2 o + MY_PROVE_HYP hi1 o MY_PROVE_HYP hi2 o MY_PROVE_HYP th1 o MY_PROVE_HYP th2) th0 + else + let lo_th, hi_th, th0 = + if s1 <> s2 then + if s3 = "F" then + float_mul_lo pp l_lo r_hi, float_mul_hi pp l_hi r_hi, FLOAT_INTERVAL_MUL_TF_FF + else + float_mul_lo pp l_hi r_lo, float_mul_hi pp l_lo r_lo, FLOAT_INTERVAL_MUL_TF_TT + else + if s3 <> s4 then + if s1 = "F" then + float_mul_lo pp l_hi r_lo, float_mul_hi pp l_hi r_hi, FLOAT_INTERVAL_MUL_FF_TF + else + float_mul_lo pp l_lo r_hi, float_mul_hi pp l_lo r_lo, FLOAT_INTERVAL_MUL_TT_TF + else + if s1 = "F" then + if s3 = "F" then + float_mul_lo pp l_lo r_lo, float_mul_hi pp l_hi r_hi, FLOAT_INTERVAL_MUL_FF_FF + else + float_mul_lo pp l_hi r_lo, float_mul_hi pp l_lo r_hi, FLOAT_INTERVAL_MUL_FF_TT + else + if s3 = "F" then + float_mul_lo pp l_lo r_hi, float_mul_hi pp l_hi r_lo, FLOAT_INTERVAL_MUL_TT_FF + else + float_mul_lo pp l_hi r_hi, float_mul_hi pp l_lo r_lo, FLOAT_INTERVAL_MUL_TT_TT in + + let f1_tm = lhand(concl lo_th) and + f2_tm = rand(concl hi_th) in + + let th = INST[x, x_var_real; n1, n1_var_num; e1, e1_var_num; + y, y_var_real; n2, n2_var_num; e2, e2_var_num; + m1, m1_var_num; r1, r1_var_num; + m2, m2_var_num; r2, r2_var_num; + f1_tm, f1_var_real; f2_tm, f2_var_real] th0 in + MY_PROVE_HYP lo_th (MY_PROVE_HYP hi_th (MY_PROVE_HYP th1 (MY_PROVE_HYP th2 th)));; + + +(*************************************) +(* float_interval_div *) + +(* FT_xx *) +let FLOAT_INTERVAL_DIV_FT_xx = prove( + `interval_arith x (float_num F n1 e1, float_num T n2 e2) + ==> interval_arith (x / y) (float_num F 0 min_exp, float_num F 0 min_exp)`, + REWRITE_TAC[real_div] THEN DISCH_THEN (MP_TAC o MATCH_MP FLOAT_INTERVAL_FT_IMP_0) THEN + SIMP_TAC[REAL_MUL_LZERO; interval_arith] THEN + MP_TAC (GEN_ALL (SPECL [`s:bool`; `0`] FLOAT_EQ_0)) THEN SIMP_TAC[REAL_LE_REFL]);; + +(* FF_FF *) +let FLOAT_INTERVAL_DIV_FF_FF = prove( + `~(m1 = 0) /\ + interval_arith x (float_num F n1 e1, float_num F n2 e2) /\ + interval_arith y (float_num F m1 r1, float_num F m2 r2) /\ + f1 <= float_num F n1 e1 / float_num F m2 r2 /\ + float_num F n2 e2 / float_num F m1 r1 <= f2 + ==> interval_arith (x / y) (f1, f2)`, + MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN + REWRITE_TAC[real_div] THEN + STRIP_TAC THEN + SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` MP_TAC THENL + [ + MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN + REWRITE_TAC[FLOAT_F_POS]; + ALL_TAC + ] THEN + SUBGOAL_THEN `~(c = &0)` ASSUME_TAC THENL + [ + EXPAND_TAC "c" THEN ASM_REWRITE_TAC[FLOAT_EQ_0]; + ALL_TAC + ] THEN + STRIP_TAC THEN + SUBGOAL_THEN `~(d = &0)` MP_TAC THENL + [ + MATCH_MP_TAC (REAL_ARITH `~(c = &0) /\ &0 <= c /\ c <= d ==> ~(d = &0)`) THEN + ASM_REWRITE_TAC[] THEN + UNDISCH_TAC `interval_arith y (c,d)` THEN + REWRITE_TAC[interval_arith] THEN + REAL_ARITH_TAC; + ALL_TAC + ] THEN + + REPLICATE_TAC 10 (POP_ASSUM MP_TAC) THEN + REPEAT (POP_ASSUM (fun th -> ALL_TAC)) THEN + REWRITE_TAC[interval_arith] THEN + REPEAT STRIP_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `a * inv d` THEN + ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN + ASM_REWRITE_TAC[REAL_LE_INV_EQ] THEN + MATCH_MP_TAC REAL_LE_INV2 THEN + ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LTE_TRANS THEN + EXISTS_TAC `c:real` THEN + ASM_REWRITE_TAC[REAL_ARITH `&0 < c <=> ~(c = &0) /\ &0 <= c`]; + + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `b * inv c` THEN + ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN + ASM_REWRITE_TAC[REAL_LE_INV_EQ] THEN + REPEAT CONJ_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `a:real` THEN ASM_REWRITE_TAC[]; + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `c:real` THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + MATCH_MP_TAC REAL_LE_INV2 THEN + ASM_REWRITE_TAC[REAL_ARITH `&0 < c <=> ~(c = &0) /\ &0 <= c`] + ]);; + +(* TT_TT *) +let FLOAT_INTERVAL_DIV_TT_TT = prove( + `~(m2 = 0) /\ + interval_arith x (float_num T n1 e1, float_num T n2 e2) /\ + interval_arith y (float_num T m1 r1, float_num T m2 r2) /\ + f1 <= float_num T n2 e2 / float_num T m1 r1 /\ + float_num T n1 e1 / float_num T m2 r2 <= f2 + ==> interval_arith (x / y) (f1,f2)`, + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[real_div; REAL_INV_NEG; REAL_NEG_MUL2] THEN + GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV)[interval_arith] THEN + REWRITE_TAC[REAL_ARITH `--a <= x /\ x <= --b <=> b <= --x /\ --x <= a`] THEN + GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV)[GSYM interval_arith] THEN + STRIP_TAC THEN + MP_TAC ((SPECL[`n2:num`; `e2:num`; `m1:num`; `r1:num`; `n1:num`; `e1:num`; `m2:num`; `r2:num`; `--x`; `--y`] o GEN_ALL) FLOAT_INTERVAL_DIV_FF_FF) THEN + REWRITE_TAC[real_div; REAL_INV_NEG; REAL_NEG_MUL2] THEN + DISCH_THEN MATCH_MP_TAC THEN + ASM_REWRITE_TAC[]);; + +(* FF_TT *) +let FLOAT_INTERVAL_DIV_FF_TT = prove( + `~(m2 = 0) /\ + interval_arith x (float_num F n1 e1, float_num F n2 e2) /\ + interval_arith y (float_num T m1 r1, float_num T m2 r2) /\ + f1 <= float_num F n2 e2 / float_num T m2 r2 /\ + float_num F n1 e1 / float_num T m1 r1 <= f2 + ==> interval_arith (x / y) (f1,f2)`, + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[real_div; REAL_INV_NEG] THEN + REWRITE_TAC[REAL_ARITH `a * --b <= c <=> --c <= a * b`] THEN + REWRITE_TAC[REAL_ARITH `c <= a * --b <=> a * b <= --c`] THEN + GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV)[interval_arith] THEN + REWRITE_TAC[REAL_ARITH `--a <= x /\ x <= --b <=> b <= --x /\ --x <= a`] THEN + GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV)[GSYM interval_arith] THEN + STRIP_TAC THEN + MP_TAC ((SPECL[`n1:num`; `e1:num`; `m1:num`; `r1:num`; `n2:num`; `e2:num`; `m2:num`; `r2:num`; `x:real`; `--y`; `--f2`; `--f1`] o GEN_ALL) FLOAT_INTERVAL_DIV_FF_FF) THEN + ANTS_TAC THENL + [ + ASM_REWRITE_TAC[real_div]; + ALL_TAC + ] THEN + REWRITE_TAC[real_div; REAL_INV_NEG; interval_arith] THEN + REAL_ARITH_TAC);; + +(* TT_FF *) +let FLOAT_INTERVAL_DIV_TT_FF = prove( + `~(m1 = 0) /\ + interval_arith x (float_num T n1 e1, float_num T n2 e2) /\ + interval_arith y (float_num F m1 r1, float_num F m2 r2) /\ + f1 <= float_num T n1 e1 / float_num F m1 r1 /\ + float_num T n2 e2 / float_num F m2 r2 <= f2 + ==> interval_arith (x / y) (f1,f2)`, + REWRITE_TAC[FLOAT_NEG_T] THEN + REWRITE_TAC[real_div; REAL_INV_NEG] THEN + REWRITE_TAC[REAL_ARITH `--a * b <= c <=> --c <= a * b`] THEN + REWRITE_TAC[REAL_ARITH `c <= --a * b <=> a * b <= --c`] THEN + GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV)[interval_arith] THEN + REWRITE_TAC[REAL_ARITH `--a <= x /\ x <= --b <=> b <= --x /\ --x <= a`] THEN + GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV)[GSYM interval_arith] THEN + STRIP_TAC THEN + MP_TAC ((SPECL[`n2:num`; `e2:num`; `m2:num`; `r2:num`; `n1:num`; `e1:num`; `m1:num`; `r1:num`; `--x:real`; `y:real`; `--f2`; `--f1`] o GEN_ALL) FLOAT_INTERVAL_DIV_FF_FF) THEN + ANTS_TAC THENL + [ + ASM_REWRITE_TAC[real_div]; + ALL_TAC + ] THEN + REWRITE_TAC[real_div; REAL_INV_NEG; interval_arith] THEN + REAL_ARITH_TAC);; + + +(* TF_FF *) +let FLOAT_INTERVAL_DIV_TF_FF = prove( + `~(m1 = 0) /\ + interval_arith x (float_num T n1 e1, float_num F n2 e2) /\ + interval_arith y (float_num F m1 r1, float_num F m2 r2) /\ + f1 <= float_num T n1 e1 / float_num F m1 r1 /\ + float_num F n2 e2 / float_num F m1 r1 <= f2 + ==> interval_arith (x / y) (f1,f2)`, + REWRITE_TAC[FLOAT_NEG_T] THEN STRIP_TAC THEN + MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN + SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` ASSUME_TAC THENL + [ + MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN REWRITE_TAC[FLOAT_F_POS]; + ALL_TAC + ] THEN + + DISJ_CASES_TAC (REAL_ARITH `&0 <= x \/ x <= &0`) THENL + [ + SUBGOAL_THEN `interval_arith x (float_num F 0 0, b)` ASSUME_TAC THENL + [ + UNDISCH_TAC `interval_arith x (--a, b)` THEN POP_ASSUM MP_TAC THEN + REWRITE_TAC[interval_arith; FLOAT_0] THEN REAL_ARITH_TAC; + ALL_TAC + ] THEN + + MP_TAC ((SPEC_ALL o SPECL [`0`; `0`] o GEN_ALL) FLOAT_INTERVAL_DIV_FF_FF) THEN + ASM_REWRITE_TAC[] THEN DISCH_THEN MATCH_MP_TAC THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN + REWRITE_TAC[real_div] THEN CONJ_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a / c` THEN + ASM_REWRITE_TAC[real_div; ARITH_RULE `--a * b <= &0 <=> &0 <= a * b`] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[REAL_LE_INV_EQ]; + ALL_TAC + ] THEN + + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[REAL_LE_INV_EQ; FLOAT_F_POS]; + ALL_TAC + ] THEN + + SUBGOAL_THEN `interval_arith x (--a, float_num T 0 0)` ASSUME_TAC THENL + [ + UNDISCH_TAC `interval_arith x (--a, b)` THEN POP_ASSUM MP_TAC THEN + REWRITE_TAC[interval_arith; FLOAT_0] THEN REAL_ARITH_TAC; + ALL_TAC + ] THEN + + MP_TAC ((INST[`0`, `n2:num`; `0`, `e2:num`]) FLOAT_INTERVAL_DIV_TT_FF) THEN + ASM_REWRITE_TAC[FLOAT_NEG_T] THEN DISCH_THEN MATCH_MP_TAC THEN + ASM_REWRITE_TAC[GSYM FLOAT_NEG_T] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN + REWRITE_TAC[real_div] THEN CONJ_TAC THENL + [ + REWRITE_TAC[FLOAT_0; REAL_MUL_LZERO; REAL_LE_REFL]; + ALL_TAC + ] THEN + + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b / c` THEN ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LE_DIV THEN ASM_REWRITE_TAC[]);; + +(* TF_TT *) +let FLOAT_INTERVAL_DIV_TF_TT = prove( + `~(m2 = 0) /\ + interval_arith x (float_num T n1 e1, float_num F n2 e2) /\ + interval_arith y (float_num T m1 r1, float_num T m2 r2) /\ + f1 <= float_num F n2 e2 / float_num T m2 r2 /\ + float_num T n1 e1 / float_num T m2 r2 <= f2 + ==> interval_arith (x / y) (f1,f2)`, + REWRITE_TAC[FLOAT_NEG_T] THEN STRIP_TAC THEN + MAP_EVERY ABBREV_TAC [`a = float_num F n1 e1`; `b = float_num F n2 e2`; `c = float_num F m1 r1`; `d = float_num F m2 r2`] THEN + SUBGOAL_THEN `&0 <= a /\ &0 <= b /\ &0 <= c /\ &0 <= d` ASSUME_TAC THENL + [ + MAP_EVERY EXPAND_TAC ["a"; "b"; "c"; "d"] THEN REWRITE_TAC[FLOAT_F_POS]; + ALL_TAC + ] THEN + + DISJ_CASES_TAC (REAL_ARITH `x <= &0 \/ &0 <= x`) THENL + [ + SUBGOAL_THEN `interval_arith x (--a, float_num T 0 0)` ASSUME_TAC THENL + [ + UNDISCH_TAC `interval_arith x (--a, b)` THEN POP_ASSUM MP_TAC THEN + REWRITE_TAC[interval_arith; FLOAT_0] THEN REAL_ARITH_TAC; + ALL_TAC + ] THEN + + MP_TAC ((SPEC_ALL o SPECL [`0`; `0`] o GEN_ALL) FLOAT_INTERVAL_DIV_TT_TT) THEN + ASM_REWRITE_TAC[] THEN DISCH_THEN MATCH_MP_TAC THEN ASM_REWRITE_TAC[FLOAT_NEG_T] THEN + ASM_REWRITE_TAC[GSYM FLOAT_NEG_T] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN + REWRITE_TAC[real_div] THEN CONJ_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b / --d` THEN + ASM_REWRITE_TAC[real_div; REAL_INV_NEG; REAL_ARITH `b * --d <= &0 <=> &0 <= b * d`] THEN + MATCH_MP_TAC REAL_LE_MUL THEN ASM_REWRITE_TAC[REAL_LE_INV_EQ]; + ALL_TAC + ] THEN + + REWRITE_TAC[FLOAT_0; REAL_MUL_LZERO; REAL_LE_REFL]; + ALL_TAC + ] THEN + + SUBGOAL_THEN `interval_arith x (float_num F 0 0, b)` ASSUME_TAC THENL + [ + UNDISCH_TAC `interval_arith x (--a, b)` THEN POP_ASSUM MP_TAC THEN + REWRITE_TAC[interval_arith; FLOAT_0] THEN REAL_ARITH_TAC; + ALL_TAC + ] THEN + + MP_TAC ((INST[`0`, `n1:num`; `0`, `e1:num`]) FLOAT_INTERVAL_DIV_FF_TT) THEN + ASM_REWRITE_TAC[FLOAT_NEG_T] THEN DISCH_THEN MATCH_MP_TAC THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&0` THEN + REWRITE_TAC[real_div] THEN CONJ_TAC THENL + [ + REWRITE_TAC[FLOAT_0; REAL_MUL_LZERO; REAL_LE_REFL]; + ALL_TAC + ] THEN + + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--a / --d` THEN ASM_REWRITE_TAC[] THEN + REWRITE_TAC[real_div; REAL_INV_NEG; REAL_NEG_MUL2] THEN MATCH_MP_TAC REAL_LE_MUL THEN + ASM_REWRITE_TAC[REAL_LE_INV_EQ]);; + +let transform = UNDISCH_ALL o + PURE_REWRITE_RULE[TAUT `~P <=> (P <=> F)`] o + NUMERALS_TO_NUM o + REWRITE_RULE[GSYM IMP_IMP; min_exp_def];; + +let FLOAT_INTERVAL_DIV_FT_xx' = transform FLOAT_INTERVAL_DIV_FT_xx and + FLOAT_INTERVAL_DIV_FF_FF' = transform FLOAT_INTERVAL_DIV_FF_FF and + FLOAT_INTERVAL_DIV_TT_TT' = transform FLOAT_INTERVAL_DIV_TT_TT and + FLOAT_INTERVAL_DIV_FF_TT' = transform FLOAT_INTERVAL_DIV_FF_TT and + FLOAT_INTERVAL_DIV_TT_FF' = transform FLOAT_INTERVAL_DIV_TT_FF and + FLOAT_INTERVAL_DIV_TF_FF' = transform FLOAT_INTERVAL_DIV_TF_FF and + FLOAT_INTERVAL_DIV_TF_TT' = transform FLOAT_INTERVAL_DIV_TF_TT;; + + +let float_interval_div pp th1 th2 = + let x, l_lo, l_hi = dest_float_interval (concl th1) and + y, r_lo, r_hi = dest_float_interval (concl th2) in + let s1, n1, e1 = dest_float l_lo and + s2, n2, e2 = dest_float l_hi and + s3, m1, r1 = dest_float r_lo and + s4, m2, r2 = dest_float r_hi in + + if s1 <> s2 && s1 = "F" then + let th0 = INST[x, x_var_real; y, y_var_real; + n1, n1_var_num; e1, e1_var_num; + n2, n2_var_num; e2, e2_var_num] FLOAT_INTERVAL_DIV_FT_xx' in + MY_PROVE_HYP th1 th0 + else if s3 <> s4 then + failwith "float_interval_div: division by an interval containing 0" + else + let lo_th, hi_th, th0, zero_th = + if s1 = s2 then + if s1 = "F" then + if s3 = "F" then + float_div_lo pp l_lo r_hi, float_div_hi pp l_hi r_lo, + FLOAT_INTERVAL_DIV_FF_FF', raw_eq0_hash_conv m1 + else + float_div_lo pp l_hi r_hi, float_div_hi pp l_lo r_lo, + FLOAT_INTERVAL_DIV_FF_TT', raw_eq0_hash_conv m2 + else + if s3 = "F" then + float_div_lo pp l_lo r_lo, float_div_hi pp l_hi r_hi, + FLOAT_INTERVAL_DIV_TT_FF', raw_eq0_hash_conv m1 + else + float_div_lo pp l_hi r_lo, float_div_hi pp l_lo r_hi, + FLOAT_INTERVAL_DIV_TT_TT', raw_eq0_hash_conv m2 + else + if s3 = "F" then + float_div_lo pp l_lo r_lo, float_div_hi pp l_hi r_lo, + FLOAT_INTERVAL_DIV_TF_FF', raw_eq0_hash_conv m1 + else + float_div_lo pp l_hi r_hi, float_div_hi pp l_lo r_hi, + FLOAT_INTERVAL_DIV_TF_TT', raw_eq0_hash_conv m2 in + + let f1_tm = lhand(concl lo_th) and + f2_tm = rand(concl hi_th) in + + let th = INST[x, x_var_real; n1, n1_var_num; e1, e1_var_num; + y, y_var_real; n2, n2_var_num; e2, e2_var_num; + m1, m1_var_num; r1, r1_var_num; + m2, m2_var_num; r2, r2_var_num; + f1_tm, f1_var_real; f2_tm, f2_var_real] th0 in + (MY_PROVE_HYP lo_th o MY_PROVE_HYP hi_th o + MY_PROVE_HYP th1 o MY_PROVE_HYP th2 o MY_PROVE_HYP zero_th) th;; + + +(*****************************************) +(* float_interval_add, float_interval_sub *) + +let n1_var_real = `n1:real` and + n2_var_real = `n2:real` and + m1_var_real = `m1:real` and + m2_var_real = `m2:real` and + n_var_real = `n:real` and + m_var_real = `m:real`;; + +let INTERVAL_ADD = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( + `interval_arith x (n1, m1) /\ + interval_arith y (n2, m2) /\ + n <= n1 + n2 /\ m1 + m2 <= m + ==> interval_arith (x + y) (n, m)`, + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; + +let INTERVAL_SUB = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove)( + `interval_arith x (n1, m1) /\ + interval_arith y (n2, m2) /\ + n <= n1 - m2 /\ m1 - n2 <= m + ==> interval_arith (x - y) (n, m)`, + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; + + + let float_interval_add pp th1 th2 = + let x, n1, m1 = dest_float_interval (concl th1) in + let y, n2, m2 = dest_float_interval (concl th2) in + let lo_th = float_add_lo pp n1 n2 in + let hi_th = float_add_hi pp m1 m2 in + let n_tm = lhand (concl lo_th) in + let m_tm = rand (concl hi_th) in + let th0 = INST[x, x_var_real; n1, n1_var_real; m1, m1_var_real; + y, y_var_real; n2, n2_var_real; m2, m2_var_real; + n_tm, n_var_real; m_tm, m_var_real] INTERVAL_ADD in + MY_PROVE_HYP lo_th (MY_PROVE_HYP hi_th (MY_PROVE_HYP th2 (MY_PROVE_HYP th1 th0)));; + +let float_interval_sub pp th1 th2 = + let x, n1, m1 = dest_float_interval (concl th1) in + let y, n2, m2 = dest_float_interval (concl th2) in + let lo_th = float_sub_lo pp n1 m2 in + let hi_th = float_sub_hi pp m1 n2 in + let n_tm = lhand(concl lo_th) in + let m_tm = rand(concl hi_th) in + let th0 = INST[x, x_var_real; n1, n1_var_real; m1, m1_var_real; + y, y_var_real; n2, n2_var_real; m2, m2_var_real; + n_tm, n_var_real; m_tm, m_var_real] INTERVAL_SUB in + MY_PROVE_HYP lo_th (MY_PROVE_HYP hi_th (MY_PROVE_HYP th2 (MY_PROVE_HYP th1 th0)));; + + +(********************************************) +(* FLOAT_ABS *) + +let s_var_bool = `s:bool`;; + +let float_abs tm = + let ltm, rtm = dest_comb tm in + if ((fst o dest_const) ltm <> "real_abs") then + failwith "float_abs: no abs" + else + let ltm, e_tm = dest_comb rtm in + let ltm, n_tm = dest_comb ltm in + let s_tm = rand ltm in + INST[s_tm, s_var_bool; n_tm, n_var_num; e_tm, e_var_num] FLOAT_ABS;; + + +(*******************************) +(* float_interval_sqrt *) + +let FLOAT_INTERVAL_SQRT = prove(`interval_arith x (float_num F n1 e1, hi) /\ + f1 <= sqrt (float_num F n1 e1) /\ sqrt hi <= f2 + ==> interval_arith (sqrt x) (f1, f2)`, + ABBREV_TAC `lo = float_num F n1 e1` THEN + REWRITE_TAC[interval_arith] THEN + STRIP_TAC THEN + SUBGOAL_THEN `&0 <= lo /\ &0 <= hi` ASSUME_TAC THENL + [ + EXPAND_TAC "lo" THEN + REWRITE_TAC[FLOAT_F_POS] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `x:real` THEN + ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `lo:real` THEN + ASM_REWRITE_TAC[] THEN + EXPAND_TAC "lo" THEN + REWRITE_TAC[FLOAT_F_POS]; + ALL_TAC + ] THEN + SUBGOAL_THEN `&0 <= x` ASSUME_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `lo:real` THEN + ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + CONJ_TAC THENL + [ + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `sqrt lo` THEN + ASM_REWRITE_TAC[SQRT_MONO_LE_EQ]; + MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `sqrt hi` THEN + ASM_REWRITE_TAC[SQRT_MONO_LE_EQ] + ]);; + + +let FLOAT_INTERVAL_SQRT' = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP]) FLOAT_INTERVAL_SQRT;; + +let float_interval_sqrt pp th = + let x_tm, lo_tm, hi_tm = dest_float_interval (concl th) in + let s1, n1_tm, e1_tm = dest_float lo_tm in + if s1 <> "F" then + failwith "float_interval_sqrt: negative low bound" + else + let lo_th = float_sqrt_lo pp lo_tm in + let hi_th = float_sqrt_hi pp hi_tm in + let f1_tm = lhand (concl lo_th) in + let f2_tm = rand (concl hi_th) in + let th0 = INST[x_tm, x_var_real; n1_tm, n1_var_num; e1_tm, e1_var_num; + hi_tm, hi_var_real; f1_tm, f1_var_real; + f2_tm, f2_var_real] FLOAT_INTERVAL_SQRT' in + MY_PROVE_HYP lo_th (MY_PROVE_HYP hi_th (MY_PROVE_HYP th th0));; + + +(******************************************) +(* FLOAT_TO_NUM_CONV *) + +let FLOAT_TO_NUM_CONV tm = + let ltm, e_tm = dest_comb tm in + let f_tm, n_tm = dest_comb ltm in + if (fst o dest_const o rator) f_tm <> "float_num" then + failwith "FLOAT_TO_NUM_CONV" + else + let n_th' = SYM (INST[n_tm, n_var_num] Arith_num.NUM_THM) in + let e_th' = SYM (INST[e_tm, n_var_num] Arith_num.NUM_THM) in + let n_th = TRANS n_th' (NUM_TO_NUMERAL_CONV (mk_comb(Arith_num.num_const, n_tm))) in + let e_th = TRANS e_th' (NUM_TO_NUMERAL_CONV (mk_comb(Arith_num.num_const, e_tm))) in + let th0 = MK_COMB (AP_TERM f_tm n_th, e_th) in + let tm0 = rand(concl th0) in + let th1 = (REWRITE_CONV[float_alt; arith_base_def; min_exp_def] THENC + NUM_REDUCE_CONV THENC + REAL_RAT_REDUCE_CONV) tm0 in + TRANS th0 th1;; + +end;; + + +(**************************************) +(* Printer for floating-point numbers *) +(**************************************) + +let print_float fmt tm = + try + let s, m_tm, e_tm = Arith_float.dest_float tm in + let m = Arith_num.raw_dest_hash m_tm and + e = Arith_num.raw_dest_hash e_tm -/ Num.num_of_int Float_theory.min_exp in + let s_str = if s = "T" then "-" else "" in + let m_str = Num.string_of_num m in + let e_str = if e = num_0 then "" + else "*" ^ string_of_int Arith_num.arith_base ^ "^" ^ Num.string_of_num e in + let str = "##" ^ s_str ^ m_str ^ e_str in + Format.pp_print_string fmt str + with _ -> failwith "print_float";; + +install_user_printer ("float_num", print_float);; diff --git a/Formal_ineqs/arith/arith_nat.hl b/Formal_ineqs/arith/arith_nat.hl index a44f8d24..e13c6ade 100644 --- a/Formal_ineqs/arith/arith_nat.hl +++ b/Formal_ineqs/arith/arith_nat.hl @@ -1,106 +1,112 @@ -(* =========================================================== *) -(* Formal natural number arithmetic *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -(* Dependencies *) -needs "arith_options.hl";; -needs "arith/arith_cache.hl";; - -module Arith_nat = struct - -open Arith_options;; - -(* mk *) -let mk_small_numeral_array = Arith_num.mk_small_numeral_array;; -let mk_numeral_array = Arith_num.mk_numeral_array;; -let NUMERAL_TO_NUM_CONV = Arith_num.NUMERAL_TO_NUM_CONV;; -let NUM_TO_NUMERAL_CONV = Arith_num.NUM_TO_NUMERAL_CONV;; - -(* dest *) -let raw_dest_hash = Arith_num.raw_dest_hash;; - -(* SUC *) -let raw_suc_conv_hash = - if !cached then Arith_cache.raw_suc_conv_hash else Arith_num.raw_suc_conv_hash;; - -let NUM_SUC_HASH_CONV = Arith_num.NUM_SUC_HASH_CONV;; - -(* x = 0 *) -let raw_eq0_hash_conv = - if !cached then Arith_cache.raw_eq0_hash_conv else Arith_num.raw_eq0_hash_conv;; - -let NUM_EQ0_HASH_CONV = Arith_num.NUM_EQ0_HASH_CONV;; - -(* PRE *) -let raw_pre_hash_conv = - if !cached then Arith_cache.raw_pre_hash_conv else Arith_num.raw_pre_hash_conv;; - -let NUM_PRE_HASH_CONV = Arith_num.NUM_PRE_HASH_CONV;; - -(* x > 0 *) -let raw_gt0_hash_conv = - if !cached then Arith_cache.raw_gt0_hash_conv else Arith_num.raw_gt0_hash_conv;; - -let NUM_GT0_HASH_CONV = Arith_num.NUM_GT0_HASH_CONV;; - -(* x = y *) -let raw_eq_hash_conv = Arith_num.raw_eq_hash_conv;; -let NUM_EQ_HASH_CONV = Arith_num.NUM_EQ_HASH_CONV;; - -(* x < y, x <= y *) -let raw_lt_hash_conv = - if !cached then Arith_cache.raw_lt_hash_conv else Arith_num.raw_lt_hash_conv;; - -let raw_le_hash_conv = - if !cached then Arith_cache.raw_le_hash_conv else Arith_num.raw_le_hash_conv;; - -let NUM_LT_HASH_CONV = Arith_num.NUM_LT_HASH_CONV;; - -let NUM_LE_HASH_CONV = Arith_num.NUM_LE_HASH_CONV;; - -(* x + y *) -let raw_add_conv_hash = - if !cached then Arith_cache.raw_add_conv_hash else Arith_num.raw_add_conv_hash;; - -let NUM_ADD_HASH_CONV = Arith_num.NUM_ADD_HASH_CONV;; - -(* x - y *) -let raw_sub_hash_conv = - if !cached then Arith_cache.raw_sub_hash_conv else Arith_num.raw_sub_hash_conv;; - -let raw_sub_and_le_hash_conv = - if !cached then Arith_cache.raw_sub_and_le_hash_conv else Arith_num.raw_sub_and_le_hash_conv;; - -let NUM_SUB_HASH_CONV = Arith_num.NUM_SUB_HASH_CONV;; - -(* x * y *) -let raw_mul_conv_hash = - if !cached then Arith_cache.raw_mul_conv_hash else Arith_num.raw_mul_conv_hash;; - -let NUM_MULT_HASH_CONV = Arith_num.NUM_MULT_HASH_CONV;; - -(* x / y *) -let raw_div_hash_conv = - if !cached then Arith_cache.raw_div_hash_conv else Arith_num.raw_div_hash_conv;; - -let NUM_DIV_HASH_CONV = Arith_num.NUM_DIV_HASH_CONV;; - -(* EVEN, ODD *) -let raw_even_hash_conv = - if !cached then Arith_cache.raw_even_hash_conv else Arith_num.raw_even_hash_conv;; - -let raw_odd_hash_conv = - if !cached then Arith_cache.raw_odd_hash_conv else Arith_num.raw_odd_hash_conv;; - -let NUM_EVEN_HASH_CONV = Arith_num.NUM_EVEN_HASH_CONV;; - -let NUM_ODD_HASH_CONV = Arith_num.NUM_ODD_HASH_CONV;; - - -let NUMERALS_TO_NUM = - PURE_REWRITE_RULE[Arith_num.NUM_THM] o CONV_RULE (DEPTH_CONV NUMERAL_TO_NUM_CONV);; - - -end;; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal natural number arithmetic *) +(* -------------------------------------------------------------------------- *) + +(* Dependencies *) +needs "Formal_ineqs/arith_options.hl";; +needs "Formal_ineqs/arith/arith_cache.hl";; + +module Arith_nat = struct + +open Arith_options;; + +(* mk *) +let mk_small_numeral_array = Arith_num.mk_small_numeral_array;; +let mk_numeral_array = Arith_num.mk_numeral_array;; +let NUMERAL_TO_NUM_CONV = Arith_num.NUMERAL_TO_NUM_CONV;; +let NUM_TO_NUMERAL_CONV = Arith_num.NUM_TO_NUMERAL_CONV;; + +(* dest *) +let raw_dest_hash = Arith_num.raw_dest_hash;; + +(* SUC *) +let raw_suc_conv_hash = + if !cached then Arith_cache.raw_suc_conv_hash else Arith_num.raw_suc_conv_hash;; + +let NUM_SUC_HASH_CONV = Arith_num.NUM_SUC_HASH_CONV;; + +(* x = 0 *) +let raw_eq0_hash_conv = + if !cached then Arith_cache.raw_eq0_hash_conv else Arith_num.raw_eq0_hash_conv;; + +let NUM_EQ0_HASH_CONV = Arith_num.NUM_EQ0_HASH_CONV;; + +(* PRE *) +let raw_pre_hash_conv = + if !cached then Arith_cache.raw_pre_hash_conv else Arith_num.raw_pre_hash_conv;; + +let NUM_PRE_HASH_CONV = Arith_num.NUM_PRE_HASH_CONV;; + +(* x > 0 *) +let raw_gt0_hash_conv = + if !cached then Arith_cache.raw_gt0_hash_conv else Arith_num.raw_gt0_hash_conv;; + +let NUM_GT0_HASH_CONV = Arith_num.NUM_GT0_HASH_CONV;; + +(* x = y *) +let raw_eq_hash_conv = Arith_num.raw_eq_hash_conv;; +let NUM_EQ_HASH_CONV = Arith_num.NUM_EQ_HASH_CONV;; + +(* x < y, x <= y *) +let raw_lt_hash_conv = + if !cached then Arith_cache.raw_lt_hash_conv else Arith_num.raw_lt_hash_conv;; + +let raw_le_hash_conv = + if !cached then Arith_cache.raw_le_hash_conv else Arith_num.raw_le_hash_conv;; + +let NUM_LT_HASH_CONV = Arith_num.NUM_LT_HASH_CONV;; + +let NUM_LE_HASH_CONV = Arith_num.NUM_LE_HASH_CONV;; + +(* x + y *) +let raw_add_conv_hash = + if !cached then Arith_cache.raw_add_conv_hash else Arith_num.raw_add_conv_hash;; + +let NUM_ADD_HASH_CONV = Arith_num.NUM_ADD_HASH_CONV;; + +(* x - y *) +let raw_sub_hash_conv = + if !cached then Arith_cache.raw_sub_hash_conv else Arith_num.raw_sub_hash_conv;; + +let raw_sub_and_le_hash_conv = + if !cached then Arith_cache.raw_sub_and_le_hash_conv else Arith_num.raw_sub_and_le_hash_conv;; + +let NUM_SUB_HASH_CONV = Arith_num.NUM_SUB_HASH_CONV;; + +(* x * y *) +let raw_mul_conv_hash = + if !cached then Arith_cache.raw_mul_conv_hash else Arith_num.raw_mul_conv_hash;; + +let NUM_MULT_HASH_CONV = Arith_num.NUM_MULT_HASH_CONV;; + +(* x / y *) +let raw_div_hash_conv = + if !cached then Arith_cache.raw_div_hash_conv else Arith_num.raw_div_hash_conv;; + +let NUM_DIV_HASH_CONV = Arith_num.NUM_DIV_HASH_CONV;; + +(* EVEN, ODD *) +let raw_even_hash_conv = + if !cached then Arith_cache.raw_even_hash_conv else Arith_num.raw_even_hash_conv;; + +let raw_odd_hash_conv = + if !cached then Arith_cache.raw_odd_hash_conv else Arith_num.raw_odd_hash_conv;; + +let NUM_EVEN_HASH_CONV = Arith_num.NUM_EVEN_HASH_CONV;; + +let NUM_ODD_HASH_CONV = Arith_num.NUM_ODD_HASH_CONV;; + + +let NUMERALS_TO_NUM = + PURE_REWRITE_RULE[Arith_num.NUM_THM] o CONV_RULE (DEPTH_CONV NUMERAL_TO_NUM_CONV);; + + +end;; diff --git a/Formal_ineqs/arith/arith_num.hl b/Formal_ineqs/arith/arith_num.hl index 5b5068b9..ddf4c641 100644 --- a/Formal_ineqs/arith/arith_num.hl +++ b/Formal_ineqs/arith/arith_num.hl @@ -1,12 +1,18 @@ -(* =========================================================== *) -(* Formal natural arithmetic with an arbitrary base *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal natural arithmetic with an arbitrary base *) +(* -------------------------------------------------------------------------- *) module type Arith_num_sig = sig - val arith_base : int + val arith_base : int val num_def : thm val NUM_THM : thm val num_const : term @@ -20,40 +26,40 @@ module type Arith_num_sig = val dest_numeral_hash : term -> num val NUMERAL_TO_NUM_CONV : term -> thm val NUM_TO_NUMERAL_CONV : term -> thm - (* SUC *) + (* SUC *) val raw_suc_conv_hash : term -> thm val NUM_SUC_HASH_CONV : term -> thm - (* eq0 *) + (* eq0 *) val raw_eq0_hash_conv : term -> thm val NUM_EQ0_HASH_CONV : term -> thm - (* PRE *) + (* PRE *) val raw_pre_hash_conv : term -> thm val NUM_PRE_HASH_CONV : term -> thm - (* gt0 *) + (* gt0 *) val raw_gt0_hash_conv : term -> thm val NUM_GT0_HASH_CONV : term -> thm - (* eq *) - val raw_eq_hash_conv : term -> thm - val NUM_EQ_HASH_CONV : term -> thm - (* lt, le *) + (* eq *) + val raw_eq_hash_conv : term -> thm + val NUM_EQ_HASH_CONV : term -> thm + (* lt, le *) val raw_lt_hash_conv : term -> thm val raw_le_hash_conv : term -> thm val NUM_LT_HASH_CONV : term -> thm val NUM_LE_HASH_CONV : term -> thm - (* add *) + (* add *) val raw_add_conv_hash : term -> thm val NUM_ADD_HASH_CONV : term -> thm - (* sub *) + (* sub *) val raw_sub_hash_conv : term -> thm val raw_sub_and_le_hash_conv : term -> term -> thm * thm val NUM_SUB_HASH_CONV : term -> thm - (* mul *) + (* mul *) val raw_mul_conv_hash : term -> thm val NUM_MULT_HASH_CONV : term -> thm - (* DIV *) + (* DIV *) val raw_div_hash_conv : term -> thm val NUM_DIV_HASH_CONV : term -> thm - (* EVEN, ODD *) + (* EVEN, ODD *) val raw_even_hash_conv : term -> thm val raw_odd_hash_conv : term -> thm val NUM_EVEN_HASH_CONV : term -> thm @@ -61,9 +67,9 @@ module type Arith_num_sig = end;; (* Dependencies *) -needs "misc/misc_functions.hl";; -needs "misc/misc_vars.hl";; -needs "arith_options.hl";; +needs "Formal_ineqs/misc/misc_functions.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; +needs "Formal_ineqs/arith_options.hl";; module Arith_num : Arith_num_sig = struct @@ -136,11 +142,11 @@ let ADD_n_0 = prove(`n + _0 = n`, REWRITE_TAC[GSYM ARITH_ADD; ADD_CLAUSES]);; let MUL_n_0 = prove(`n * _0 = 0`, - REWRITE_TAC[NUMERAL] THEN - SUBGOAL_THEN `_0 = 0` MP_TAC THENL [ REWRITE_TAC[NUMERAL]; ALL_TAC ] THEN + REWRITE_TAC[NUMERAL] THEN + SUBGOAL_THEN `_0 = 0` MP_TAC THENL [ REWRITE_TAC[NUMERAL]; ALL_TAC ] THEN DISCH_THEN (fun th -> ONCE_REWRITE_TAC[th]) THEN - ARITH_TAC);; - + ARITH_TAC);; + (* D_i(n) = i + D_0(n) *) let def_thm i = let bin = mk_comb(const_array.(i), n_var_num) in @@ -158,7 +164,7 @@ let B0_0 = prove(mk_eq(mk_comb(b0_const, zero_const), zero_const), let B0_EXPLICIT = prove(mk_eq(mk_comb(b0_const, n_var_num), mk_binop mul_op_num max_const n_var_num), REWRITE_TAC[b0_def; ADD_CLAUSES]);; - + (******************************) (* mk_numeral and dest_numeral *) @@ -167,11 +173,11 @@ let B0_EXPLICIT = prove(mk_eq(mk_comb(b0_const, n_var_num), let mk_table = Hashtbl.create maximum;; for i = 0 to maximum - 1 do - Hashtbl.add mk_table (num i) const_array.(i) + Hashtbl.add mk_table (Num.num_of_int i) const_array.(i) done;; (* mk_numeral *) -let max_num = num maximum;; +let max_num = Num.num_of_int maximum;; let mk_numeral_hash = let rec mk_num n = @@ -209,12 +215,12 @@ let mk_small_numeral_array = let dest_table_num = Hashtbl.create maximum;; for i = 0 to maximum - 1 do - Hashtbl.add dest_table_num names_array.(i) (num i) + Hashtbl.add dest_table_num names_array.(i) (Num.num_of_int i) done;; (* dest_numeral *) -let max_num = num maximum;; +let max_num = Num.num_of_int maximum;; let rec raw_dest_hash tm = if tm = zero_const then @@ -235,7 +241,7 @@ let dest_numeral_hash tm = raw_dest_hash (rand tm);; let th_num_conv = Array.init maximum (fun i -> (SYM o SPEC_ALL) def_array.(i));; let mod_op_num = `MOD`;; -let DIV_BASE = +let DIV_BASE = let h1 = mk_eq(mk_binop div_op_num m_var_num max_const, q_var_num) in let h2 = mk_eq(mk_binop mod_op_num m_var_num max_const, r_var_num) in let c = mk_eq(m_var_num, mk_binop add_op_num (mk_binop mul_op_num max_const q_var_num) r_var_num) in @@ -261,14 +267,14 @@ let NUMERAL_TO_NUM_CONV tm = let th = TRANS th1 th2 in let ltm, rtm = dest_comb(rand(concl th)) in let r_th = raw_conv rtm in - TRANS th (AP_TERM ltm r_th) in + TRANS th (AP_TERM ltm r_th) in if (fst o dest_const o rator) tm <> "NUMERAL" then failwith "NUMERAL_TO_NUM_CONV" else let th0 = raw_conv tm in let n_tm = rand(concl th0) in - TRANS th0 (INST[n_tm, n_var_num] SYM_NUM_THM);; + TRANS th0 (INST[n_tm, n_var_num] SYM_NUM_THM);; let replace_numerals = rand o concl o DEPTH_CONV NUMERAL_TO_NUM_CONV;; let REPLACE_NUMERALS = CONV_RULE (DEPTH_CONV NUMERAL_TO_NUM_CONV);; @@ -290,14 +296,14 @@ let NUM_TO_NUMERAL_CONV tm = let mul_th = NUM_MULT_CONV (rand ltm) in let add_th0 = AP_THM (AP_TERM add_op_num mul_th) rtm in let add_th = TRANS add_th0 (NUM_ADD_CONV (rand(concl add_th0))) in - TRANS th2 add_th in + TRANS th2 add_th in let ltm, rtm = dest_comb tm in if (fst o dest_const) ltm <> num_name then failwith "NUM_TO_NUMERAL_CONV" else let num_th = INST[rtm, n_var_num] NUM_THM in let th0 = raw_conv rtm in - TRANS num_th th0;; + TRANS num_th th0;; (*************************) (* SUC_CONV *) @@ -357,10 +363,10 @@ let NUM_SUC_HASH_CONV tm = (**************************************) (* EQ_0_CONV *) let EQ_0_NUM = prove(mk_eq(mk_eq(mk_comb(num_const, n_var_num), `_0`), `n = _0`), - REWRITE_TAC[num_def; NUMERAL]);; + REWRITE_TAC[num_def; NUMERAL]);; let EQ_B0_0 = prove(mk_eq(mk_eq(mk_comb(b0_const, n_var_num), `_0`), `n = _0`), - REWRITE_TAC[b0_def; ADD_CLAUSES; NUMERAL; REWRITE_RULE[NUMERAL] MULT_EQ_0; ARITH_EQ]);; + REWRITE_TAC[b0_def; ADD_CLAUSES; NUMERAL; REWRITE_RULE[NUMERAL] MULT_EQ_0; ARITH_EQ]);; let EQ_0_0 = prove(`_0 = _0 <=> T`, REWRITE_TAC[ARITH_EQ]);; @@ -385,11 +391,11 @@ let rec raw_eq0_hash_conv rtm = let b_tm, n_tm = dest_comb rtm in let cn = (fst o dest_const) b_tm in if (cn = b0_name) then - let th0 = INST[n_tm, n_var_num] EQ_B0_0 in - let th1 = raw_eq0_hash_conv n_tm in - TRANS th0 th1 + let th0 = INST[n_tm, n_var_num] EQ_B0_0 in + let th1 = raw_eq0_hash_conv n_tm in + TRANS th0 th1 else - INST[n_tm, n_var_num] (Hashtbl.find th_eq0_table cn);; + INST[n_tm, n_var_num] (Hashtbl.find th_eq0_table cn);; let NUM_EQ0_HASH_CONV rtm = let n_tm = rand rtm in @@ -404,23 +410,23 @@ let PRE_NUM = prove(mk_eq(mk_comb(pre_op_num, mk_comb (num_const, n_var_num)), mk_comb(num_const, mk_comb (pre_op_num, n_var_num))), REWRITE_TAC[num_def; NUMERAL]);; -let PRE_0 = prove(`PRE _0 = _0`, - MP_TAC (CONJUNCT1 PRE) THEN SIMP_TAC[NUMERAL]);; +let PRE_0 = prove(`PRE _0 = _0`, + MP_TAC (CONJUNCT1 PRE) THEN SIMP_TAC[NUMERAL]);; let PRE_B1_0 = prove(mk_eq(mk_comb(`PRE`, mk_comb(const_array.(1), `_0`)), `_0`), - REWRITE_TAC[def_array.(1); MUL_n_0; ARITH_ADD; NUMERAL; ARITH_PRE; ARITH_EQ]);; + REWRITE_TAC[def_array.(1); MUL_n_0; ARITH_ADD; NUMERAL; ARITH_PRE; ARITH_EQ]);; -let PRE_B0_n0 = (UNDISCH_ALL o prove)(mk_imp(`n = _0 <=> T`, - mk_eq(mk_comb(`PRE`, mk_comb(b0_const, `n:num`)), `_0`)), - REWRITE_TAC[B0_EXPLICIT] THEN - DISCH_THEN (fun th -> REWRITE_TAC[th; MUL_n_0]) THEN - REWRITE_TAC[NUMERAL; ARITH_PRE]);; +let PRE_B0_n0 = (UNDISCH_ALL o prove)(mk_imp(`n = _0 <=> T`, + mk_eq(mk_comb(`PRE`, mk_comb(b0_const, `n:num`)), `_0`)), + REWRITE_TAC[B0_EXPLICIT] THEN + DISCH_THEN (fun th -> REWRITE_TAC[th; MUL_n_0]) THEN + REWRITE_TAC[NUMERAL; ARITH_PRE]);; let PRE_B0_n1 = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o prove)(mk_imp(`n = 0 <=> F`, - mk_eq(mk_comb(`PRE`, mk_comb(b0_const, `n:num`)), - mk_comb(const_array.(maximum - 1), `PRE n`))), - REWRITE_TAC[B0_EXPLICIT; def_array.(maximum - 1)] THEN ARITH_TAC);; + mk_eq(mk_comb(`PRE`, mk_comb(b0_const, `n:num`)), + mk_comb(const_array.(maximum - 1), `PRE n`))), + REWRITE_TAC[B0_EXPLICIT; def_array.(maximum - 1)] THEN ARITH_TAC);; let PRE_lemma = (UNDISCH_ALL o PURE_REWRITE_RULE[NUMERAL] o ARITH_RULE) `((n = 0) <=> F) ==> (SUC m = n <=> PRE n = m)`;; @@ -455,24 +461,24 @@ let rec raw_pre_hash_conv tm = let btm, ntm = dest_comb otm in let cn = fst(dest_const btm) in if (cn = b0_name) then - let n_th = raw_eq0_hash_conv ntm in - if (rand(concl n_th) = f_const) then - let th0 = INST[ntm, n_var_num] PRE_B0_n1 in - let th1 = MY_PROVE_HYP n_th th0 in - let ltm, rtm = dest_comb(rand(concl th1)) in - let th2 = raw_pre_hash_conv rtm in - TRANS th1 (AP_TERM ltm th2) - else - let th = INST[ntm, n_var_num] PRE_B0_n0 in - MY_PROVE_HYP n_th th + let n_th = raw_eq0_hash_conv ntm in + if (rand(concl n_th) = f_const) then + let th0 = INST[ntm, n_var_num] PRE_B0_n1 in + let th1 = MY_PROVE_HYP n_th th0 in + let ltm, rtm = dest_comb(rand(concl th1)) in + let th2 = raw_pre_hash_conv rtm in + TRANS th1 (AP_TERM ltm th2) + else + let th = INST[ntm, n_var_num] PRE_B0_n0 in + MY_PROVE_HYP n_th th else - if (cn = b1_name) then - if (ntm = zero_const) then - PRE_B1_0 - else - INST[ntm, n_var_num] b1_pre_thm - else - INST [ntm, n_var_num] (Hashtbl.find th_pre_table cn);; + if (cn = b1_name) then + if (ntm = zero_const) then + PRE_B1_0 + else + INST[ntm, n_var_num] b1_pre_thm + else + INST [ntm, n_var_num] (Hashtbl.find th_pre_table cn);; let NUM_PRE_HASH_CONV tm = let ntm = rand (rand tm) in @@ -492,7 +498,7 @@ let GT0_NUM = (REWRITE_RULE[GSYM num_def] o prove)(`0 < NUMERAL n <=> _0 < n`, R let gt0_0 = prove(`_0 < _0 <=> F`, REWRITE_TAC[ARITH_LT]);; let gt0_b0 = (REWRITE_RULE[NUMERAL] o prove)(mk_eq (mk_binop lt_op_num `0` (mk_comb(b0_const, n_var_num)), `0 < n`), - REWRITE_TAC[b0_def] THEN ARITH_TAC);; + REWRITE_TAC[b0_def] THEN ARITH_TAC);; let gt0_th i = let bi = const_array.(i) in @@ -511,12 +517,12 @@ let rec raw_gt0_hash_conv rtm = let b_tm, n_tm = dest_comb rtm in let cn = (fst o dest_const) b_tm in if (cn = b0_name) then - let th0 = INST[n_tm, n_var_num] gt0_b0 in - let th1 = raw_gt0_hash_conv n_tm in - TRANS th0 th1 + let th0 = INST[n_tm, n_var_num] gt0_b0 in + let th1 = raw_gt0_hash_conv n_tm in + TRANS th0 th1 else - INST[n_tm, n_var_num] (Hashtbl.find gt0_table cn);; - + INST[n_tm, n_var_num] (Hashtbl.find gt0_table cn);; + let NUM_GT0_HASH_CONV rtm = let n_tm = rand rtm in let th = INST [n_tm, n_var_num] GT0_NUM in @@ -533,30 +539,30 @@ let EQ_NUM = (REWRITE_RULE[SYM num_def] o prove)(`NUMERAL m = NUMERAL n <=> m = (* Generates the theorem |- D_i(m) = D_j(n) <=> F if i <> j and D_i(m) = D_i(n) <=> m = n *) let gen_bi_eq_bj = let aux_th = prove(`i < b /\ j < b /\ ~(i = j) ==> ~(b * m + i = b * n + j:num)`, - ONCE_REWRITE_TAC[ARITH_RULE `b * m + i = b * n + j <=> m * b + i = n * b + j:num`] THEN - STRIP_TAC THEN POP_ASSUM MP_TAC THEN - SUBGOAL_THEN `~(b = 0)` ASSUME_TAC THENL [POP_ASSUM MP_TAC THEN ARITH_TAC; ALL_TAC] THEN - REWRITE_TAC[CONTRAPOS_THM] THEN - DISCH_TAC THEN FIRST_ASSUM (MP_TAC o AP_TERM `\x. x DIV b`) THEN REWRITE_TAC[] THEN - FIRST_ASSUM (MP_TAC o MATCH_MP(CONJUNCT1 DIV_MULT_ADD)) THEN - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - SUBGOAL_THEN `i DIV b = 0 /\ j DIV b = 0` (fun th -> REWRITE_TAC[th]) THENL - [ - FIRST_ASSUM (ASSUME_TAC o MATCH_MP DIV_EQ_0) THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - REWRITE_TAC[ADD_0] THEN DISCH_TAC THEN - UNDISCH_TAC `m * b + i = n * b + j:num` THEN - ASM_REWRITE_TAC[] THEN ARITH_TAC) in + ONCE_REWRITE_TAC[ARITH_RULE `b * m + i = b * n + j <=> m * b + i = n * b + j:num`] THEN + STRIP_TAC THEN POP_ASSUM MP_TAC THEN + SUBGOAL_THEN `~(b = 0)` ASSUME_TAC THENL [POP_ASSUM MP_TAC THEN ARITH_TAC; ALL_TAC] THEN + REWRITE_TAC[CONTRAPOS_THM] THEN + DISCH_TAC THEN FIRST_ASSUM (MP_TAC o AP_TERM `\x. x DIV b`) THEN REWRITE_TAC[] THEN + FIRST_ASSUM (MP_TAC o MATCH_MP (CONJUNCT1 DIV_MULT_ADD)) THEN + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + SUBGOAL_THEN `i DIV b = 0 /\ j DIV b = 0` (fun th -> REWRITE_TAC[th]) THENL + [ + FIRST_ASSUM (ASSUME_TAC o MATCH_MP DIV_EQ_0) THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + REWRITE_TAC[ADD_0] THEN DISCH_TAC THEN + UNDISCH_TAC `m * b + i = n * b + j:num` THEN + ASM_REWRITE_TAC[] THEN ARITH_TAC) in fun i j -> let bi_m = mk_comb (const_array.(i), m_var_num) in let bj_n = mk_comb (const_array.(j), n_var_num) in let eq_tm = mk_binop eq_op_num bi_m bj_n in - if i = j then - prove (mk_eq (eq_tm, mk_binop eq_op_num m_var_num n_var_num), REWRITE_TAC[def_array.(i)] THEN ARITH_TAC) - else - prove (mk_eq (eq_tm, f_const), REWRITE_TAC[def_array.(i); def_array.(j)] THEN - MATCH_MP_TAC aux_th THEN CONV_TAC NUM_REDUCE_CONV);; + if i = j then + prove (mk_eq (eq_tm, mk_binop eq_op_num m_var_num n_var_num), REWRITE_TAC[def_array.(i)] THEN ARITH_TAC) + else + prove (mk_eq (eq_tm, f_const), REWRITE_TAC[def_array.(i); def_array.(j)] THEN + MATCH_MP_TAC aux_th THEN CONV_TAC NUM_REDUCE_CONV);; let gen_next_eq_thm = @@ -568,7 +574,7 @@ let gen_next_eq_thm = let suc_m = raw_suc_conv_hash (mk_comb (suc_op_num, m_tm)) in let suc_n = raw_suc_conv_hash (mk_comb (suc_op_num, n_tm)) in let th1 = SYM (MK_COMB ((AP_TERM eq_op_num suc_m), suc_n)) in - TRANS (TRANS th1 th0) th;; + TRANS (TRANS th1 th0) th;; let th_eq_table = Hashtbl.create (maximum * maximum);; @@ -586,13 +592,13 @@ for i = 0 to maximum - 1 do for k = 1 to maximum - i - 1 do let x = k and y = i + k in let name_left = names_array.(x) and - name_right = names_array.(y) in + name_right = names_array.(y) in let _ = th := gen_next_eq_thm (!th) in let _ = Hashtbl.add th_eq_table (name_left ^ name_right) !th in - if y > x then - Hashtbl.add th_eq_table (name_right ^ name_left) (sym !th) - else - () + if y > x then + Hashtbl.add th_eq_table (name_right ^ name_left) (sym !th) + else + () done; done;; @@ -610,22 +616,22 @@ let rec raw_eq_hash_conv tm = else (* n <> _0 *) if is_const ltm then - (* m = _0 *) - let th0 = raw_eq0_hash_conv rtm in - let th1 = INST[rtm, n_var_num] eq_0_sym in - TRANS th1 th0 + (* m = _0 *) + let th0 = raw_eq0_hash_conv rtm in + let th1 = INST[rtm, n_var_num] eq_0_sym in + TRANS th1 th0 else - (* m <> _0 *) - let bm_tm, m_tm = dest_comb ltm in - let bn_tm, n_tm = dest_comb rtm in - let cbm = (fst o dest_const) bm_tm in - let cbn = (fst o dest_const) bn_tm in - let th0 = INST[m_tm, m_var_num; n_tm, n_var_num] (Hashtbl.find th_eq_table (cbm^cbn)) in - if cbm <> cbn then - th0 - else - let th1 = raw_eq_hash_conv (rand (concl th0)) in - TRANS th0 th1;; + (* m <> _0 *) + let bm_tm, m_tm = dest_comb ltm in + let bn_tm, n_tm = dest_comb rtm in + let cbm = (fst o dest_const) bm_tm in + let cbn = (fst o dest_const) bn_tm in + let th0 = INST[m_tm, m_var_num; n_tm, n_var_num] (Hashtbl.find th_eq_table (cbm^cbn)) in + if cbm <> cbn then + th0 + else + let th1 = raw_eq_hash_conv (rand (concl th0)) in + TRANS th0 th1;; let NUM_EQ_HASH_CONV tm = @@ -643,15 +649,15 @@ let NUM_EQ_HASH_CONV tm = let LT_NUM = (REWRITE_RULE[SYM num_def] o prove)(`NUMERAL m < NUMERAL n <=> m < n`, REWRITE_TAC[NUMERAL]);; let LE_NUM = (REWRITE_RULE[SYM num_def] o prove)(`NUMERAL m <= NUMERAL n <=> m <= n`, REWRITE_TAC[NUMERAL]);; -let LT_n_0 = prove(`n < _0 <=> F`, - SUBGOAL_THEN `_0 = 0` MP_TAC THENL [ REWRITE_TAC[NUMERAL]; ALL_TAC ] THEN - DISCH_THEN (fun th -> PURE_ONCE_REWRITE_TAC[th]) THEN - ARITH_TAC);; +let LT_n_0 = prove(`n < _0 <=> F`, + SUBGOAL_THEN `_0 = 0` MP_TAC THENL [ REWRITE_TAC[NUMERAL]; ALL_TAC ] THEN + DISCH_THEN (fun th -> PURE_ONCE_REWRITE_TAC[th]) THEN + ARITH_TAC);; let LE_0_n = prove(`_0 <= n <=> T`, - SUBGOAL_THEN `_0 = 0` MP_TAC THENL [ REWRITE_TAC[NUMERAL]; ALL_TAC ] THEN - DISCH_THEN (fun th -> PURE_ONCE_REWRITE_TAC[th]) THEN - ARITH_TAC);; + SUBGOAL_THEN `_0 = 0` MP_TAC THENL [ REWRITE_TAC[NUMERAL]; ALL_TAC ] THEN + DISCH_THEN (fun th -> PURE_ONCE_REWRITE_TAC[th]) THEN + ARITH_TAC);; let SUC_LT_THM = ARITH_RULE `SUC m < SUC n <=> m < n`;; let SUC_LE_THM = ARITH_RULE `SUC m <= SUC n <=> m <= n`;; @@ -680,9 +686,9 @@ let gen_bi_lt_bj i j = let bim = mk_comb (const_array.(i), m_var_num) in let bjn = mk_comb (const_array.(j), n_var_num) in let lt_tm = mk_binop lt_op_num bim bjn in - let rhs = + let rhs = if i >= j then - mk_binop lt_op_num m_var_num n_var_num + mk_binop lt_op_num m_var_num n_var_num else mk_binop le_op_num m_var_num n_var_num in prove(mk_eq(lt_tm, rhs), REWRITE_TAC[def_array.(i); def_array.(j)] THEN ARITH_TAC);; @@ -709,9 +715,9 @@ for i = 0 to maximum - 1 do for k = 1 to maximum - i - 1 do let x = k and y = i + k in let name_left = names_array.(x) and - name_right = names_array.(y) in - th := gen_next_lt_thm (!th); - Hashtbl.add th_lt_table (name_left ^ name_right) !th + name_right = names_array.(y) in + th := gen_next_lt_thm (!th); + Hashtbl.add th_lt_table (name_left ^ name_right) !th done; done;; @@ -724,9 +730,9 @@ for i = 1 to maximum - 1 do for k = 1 to maximum - i - 1 do let x = i + k and y = k in let name_left = names_array.(x) and - name_right = names_array.(y) in - th := gen_next_lt_thm (!th); - Hashtbl.add th_lt_table (name_left ^ name_right) !th + name_right = names_array.(y) in + th := gen_next_lt_thm (!th); + Hashtbl.add th_lt_table (name_left ^ name_right) !th done; done;; @@ -754,9 +760,9 @@ let gen_bi_le_bj i j = let bim = mk_comb (const_array.(i), m_var_num) in let bjn = mk_comb (const_array.(j), n_var_num) in let lt_tm = mk_binop le_op_num bim bjn in - let rhs = + let rhs = if i > j then - mk_binop lt_op_num m_var_num n_var_num + mk_binop lt_op_num m_var_num n_var_num else mk_binop le_op_num m_var_num n_var_num in prove(mk_eq(lt_tm, rhs), REWRITE_TAC[def_array.(i); def_array.(j)] THEN ARITH_TAC);; @@ -771,7 +777,7 @@ let gen_next_le_thm th = let suc_n = raw_suc_conv_hash (mk_comb (suc_op_num, n_tm)) in let th1 = SYM (MK_COMB ((AP_TERM le_op_num suc_m), suc_n)) in TRANS (TRANS th1 th0) th;; - + let th_le_table = Hashtbl.create (maximum * maximum);; for i = 0 to maximum - 1 do @@ -783,9 +789,9 @@ for i = 0 to maximum - 1 do for k = 1 to maximum - i - 1 do let x = k and y = i + k in let name_left = names_array.(x) and - name_right = names_array.(y) in - th := gen_next_le_thm (!th); - Hashtbl.add th_le_table (name_left ^ name_right) !th + name_right = names_array.(y) in + th := gen_next_le_thm (!th); + Hashtbl.add th_le_table (name_left ^ name_right) !th done; done;; @@ -798,9 +804,9 @@ for i = 1 to maximum - 1 do for k = 1 to maximum - i - 1 do let x = i + k and y = k in let name_left = names_array.(x) and - name_right = names_array.(y) in - th := gen_next_le_thm (!th); - Hashtbl.add th_le_table (name_left ^ name_right) !th + name_right = names_array.(y) in + th := gen_next_le_thm (!th); + Hashtbl.add th_le_table (name_left ^ name_right) !th done; done;; @@ -814,30 +820,30 @@ let rec raw_lt_hash_conv tm = INST[ltm, n_var_num] LT_n_0 else if is_const ltm then - (* _0 < Bi(n) *) - let bn_tm, n_tm = dest_comb rtm in - let cbn = (fst o dest_const) bn_tm in - let th0 = INST[n_tm, n_var_num] (Hashtbl.find th_lt0_table cbn) in - if cbn = b0_name then - let th1 = raw_lt_hash_conv (rand (concl th0)) in - TRANS th0 th1 - else - th0 + (* _0 < Bi(n) *) + let bn_tm, n_tm = dest_comb rtm in + let cbn = (fst o dest_const) bn_tm in + let th0 = INST[n_tm, n_var_num] (Hashtbl.find th_lt0_table cbn) in + if cbn = b0_name then + let th1 = raw_lt_hash_conv (rand (concl th0)) in + TRANS th0 th1 + else + th0 else - (* Bi(n) < Bj(m) *) - let bm_tm, m_tm = dest_comb ltm in - let bn_tm, n_tm = dest_comb rtm in - let cbm = (fst o dest_const) bm_tm in - let cbn = (fst o dest_const) bn_tm in - let th0 = INST[m_tm, m_var_num; n_tm, n_var_num] (Hashtbl.find th_lt_table (cbm^cbn)) in - let op = (fst o dest_const o rator o rator o rand o concl) th0 in - let th1 = - if op = "<" then - raw_lt_hash_conv (rand (concl th0)) - else - raw_le_hash_conv (rand (concl th0)) in - TRANS th0 th1 and - + (* Bi(n) < Bj(m) *) + let bm_tm, m_tm = dest_comb ltm in + let bn_tm, n_tm = dest_comb rtm in + let cbm = (fst o dest_const) bm_tm in + let cbn = (fst o dest_const) bn_tm in + let th0 = INST[m_tm, m_var_num; n_tm, n_var_num] (Hashtbl.find th_lt_table (cbm^cbn)) in + let op = (fst o dest_const o rator o rator o rand o concl) th0 in + let th1 = + if op = "<" then + raw_lt_hash_conv (rand (concl th0)) + else + raw_le_hash_conv (rand (concl th0)) in + TRANS th0 th1 and + raw_le_hash_conv tm = let ltm, rtm = dest_comb tm in let ltm = rand ltm in @@ -846,29 +852,29 @@ let rec raw_lt_hash_conv tm = INST[rtm, n_var_num] LE_0_n else if is_const rtm then - (* Bi(n) <= _0 *) - let bn_tm, n_tm = dest_comb ltm in - let cbn = (fst o dest_const) bn_tm in - let th0 = INST[n_tm, n_var_num] (Hashtbl.find th_le0_table cbn) in - if cbn = b0_name then - let th1 = raw_le_hash_conv (rand (concl th0)) in - TRANS th0 th1 - else - th0 + (* Bi(n) <= _0 *) + let bn_tm, n_tm = dest_comb ltm in + let cbn = (fst o dest_const) bn_tm in + let th0 = INST[n_tm, n_var_num] (Hashtbl.find th_le0_table cbn) in + if cbn = b0_name then + let th1 = raw_le_hash_conv (rand (concl th0)) in + TRANS th0 th1 + else + th0 else - (* Bi(n) <= Bj(m) *) - let bm_tm, m_tm = dest_comb ltm in - let bn_tm, n_tm = dest_comb rtm in - let cbm = (fst o dest_const) bm_tm in - let cbn = (fst o dest_const) bn_tm in - let th0 = INST[m_tm, m_var_num; n_tm, n_var_num] (Hashtbl.find th_le_table (cbm^cbn)) in - let op = (fst o dest_const o rator o rator o rand o concl) th0 in - let th1 = - if op = "<" then - raw_lt_hash_conv (rand (concl th0)) - else - raw_le_hash_conv (rand (concl th0)) in - TRANS th0 th1;; + (* Bi(n) <= Bj(m) *) + let bm_tm, m_tm = dest_comb ltm in + let bn_tm, n_tm = dest_comb rtm in + let cbm = (fst o dest_const) bm_tm in + let cbn = (fst o dest_const) bn_tm in + let th0 = INST[m_tm, m_var_num; n_tm, n_var_num] (Hashtbl.find th_le_table (cbm^cbn)) in + let op = (fst o dest_const o rator o rator o rand o concl) th0 in + let th1 = + if op = "<" then + raw_lt_hash_conv (rand (concl th0)) + else + raw_le_hash_conv (rand (concl th0)) in + TRANS th0 th1;; let NUM_LT_HASH_CONV tm = let atm, rtm = dest_comb tm in @@ -1048,12 +1054,12 @@ let raw_sub_hash_conv tm = let t_tm = rand (mk_numeral_array t) in let th0 = INST[n_tm, n_var_num; t_tm, t_var_num; m_tm, m_var_num] SUB_lemma1 in let th_add = raw_add_conv_hash (mk_binop add_op_num n_tm t_tm) in - MY_PROVE_HYP th_add th0 + MY_PROVE_HYP th_add th0 else let t_tm = rand (mk_numeral_array (Num.abs_num t)) in let th0 = INST[m_tm, m_var_num; t_tm, t_var_num; n_tm, n_var_num] SUB_lemma2 in let th_add = raw_add_conv_hash (mk_binop add_op_num m_tm t_tm) in - MY_PROVE_HYP th_add th0;; + MY_PROVE_HYP th_add th0;; (* Returns either (tm1 - tm2, tm2 <= tm1) or (tm2 - tm1, tm1 <= tm2) *) let raw_sub_and_le_hash_conv tm1 tm2 = @@ -1066,14 +1072,14 @@ let raw_sub_and_le_hash_conv tm1 tm2 = let th_sub = inst SUB_lemma1 in let th_le = inst LE_lemma in let th_add = raw_add_conv_hash (mk_binop add_op_num tm2 t_tm) in - (MY_PROVE_HYP th_add th_sub, MY_PROVE_HYP th_add th_le) + (MY_PROVE_HYP th_add th_sub, MY_PROVE_HYP th_add th_le) else let t_tm = rand (mk_numeral_array (Num.abs_num t)) in let inst = INST[tm2, m_var_num; t_tm, t_var_num; tm1, n_var_num] in let th_sub = inst SUB_lemma1 in let th_le = inst LE_lemma in let th_add = raw_add_conv_hash (mk_binop add_op_num tm1 t_tm) in - (MY_PROVE_HYP th_add th_sub, MY_PROVE_HYP th_add th_le);; + (MY_PROVE_HYP th_add th_sub, MY_PROVE_HYP th_add th_le);; let NUM_SUB_HASH_CONV tm = let atm, rtm = dest_comb tm in @@ -1392,10 +1398,10 @@ let DIV_NUM = prove(mk_eq(mk_binop div_op_num (mk_comb(num_const, m_var_num)) (m mk_comb(num_const, mk_binop div_op_num m_var_num n_var_num)), REWRITE_TAC[num_def; NUMERAL]);; -let DIV_UNIQ' = (UNDISCH_ALL o - PURE_ONCE_REWRITE_RULE[ARITH_RULE `a < b <=> (a < b:num <=> T)`] o - ONCE_REWRITE_RULE[ARITH_RULE `m = q * n + r <=> q * n + r = m:num`] o - REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL) DIV_UNIQ;; +let DIV_UNIQ' = (UNDISCH_ALL o + PURE_ONCE_REWRITE_RULE[ARITH_RULE `a < b <=> (a < b:num <=> T)`] o + ONCE_REWRITE_RULE[ARITH_RULE `m = q * n + r <=> q * n + r = m:num`] o + REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL) DIV_UNIQ;; (* Computes m DIV n *) let raw_div_hash_conv tm = @@ -1447,13 +1453,13 @@ let ODD_ZERO = prove(`ODD _0 <=> F`, REWRITE_TAC[ARITH_ODD]);; let EVEN_B0 = prove(mk_eq(mk_comb(`EVEN`, mk_comb(b0_const, `n:num`)), `T`), - REWRITE_TAC[B0_EXPLICIT; EVEN_MULT] THEN - DISJ1_TAC THEN CONV_TAC NUM_EVEN_CONV);; + REWRITE_TAC[B0_EXPLICIT; EVEN_MULT] THEN + DISJ1_TAC THEN CONV_TAC NUM_EVEN_CONV);; let ODD_B0 = prove(mk_eq(mk_comb(`ODD`, mk_comb(b0_const, `n:num`)), `F`), - REWRITE_TAC[NOT_ODD; EVEN_B0]);; - + REWRITE_TAC[NOT_ODD; EVEN_B0]);; + let EVEN_SUC_T = prove(`(EVEN (SUC n) <=> T) <=> (EVEN n <=> F)`, REWRITE_TAC[EVEN]);; let EVEN_SUC_F = prove(`(EVEN (SUC n) <=> F) <=> (EVEN n <=> T)`, REWRITE_TAC[EVEN]);; @@ -1503,11 +1509,11 @@ let raw_even_hash_conv tm = failwith "raw_even_hash_conv: no EVEN" else if (is_const rtm) then - EVEN_ZERO + EVEN_ZERO else - let b_tm, n_tm = dest_comb rtm in - let th0 = Hashtbl.find even_thm_table ((fst o dest_const) b_tm) in - INST[n_tm, n_var_num] th0;; + let b_tm, n_tm = dest_comb rtm in + let th0 = Hashtbl.find even_thm_table ((fst o dest_const) b_tm) in + INST[n_tm, n_var_num] th0;; let raw_odd_hash_conv tm = let ltm, rtm = dest_comb tm in @@ -1515,11 +1521,11 @@ let raw_odd_hash_conv tm = failwith "raw_odd_hash_conv: no ODD" else if (is_const rtm) then - ODD_ZERO + ODD_ZERO else - let b_tm, n_tm = dest_comb rtm in - let th0 = Hashtbl.find odd_thm_table ((fst o dest_const) b_tm) in - INST[n_tm, n_var_num] th0;; + let b_tm, n_tm = dest_comb rtm in + let th0 = Hashtbl.find odd_thm_table ((fst o dest_const) b_tm) in + INST[n_tm, n_var_num] th0;; let NUM_EVEN_HASH_CONV tm = let ltm, rtm = dest_comb tm in @@ -1529,7 +1535,7 @@ let NUM_EVEN_HASH_CONV tm = failwith "NUM_EVEN_HASH_CONV" else let th1 = raw_even_hash_conv rtm in - TRANS th0 th1;; + TRANS th0 th1;; let NUM_ODD_HASH_CONV tm = let ltm, rtm = dest_comb tm in @@ -1539,6 +1545,6 @@ let NUM_ODD_HASH_CONV tm = failwith "NUM_ODD_HASH_CONV" else let th1 = raw_odd_hash_conv rtm in - TRANS th0 th1;; + TRANS th0 th1;; end;; diff --git a/Formal_ineqs/arith/eval_interval.hl b/Formal_ineqs/arith/eval_interval.hl index 74026e22..ce6175ea 100644 --- a/Formal_ineqs/arith/eval_interval.hl +++ b/Formal_ineqs/arith/eval_interval.hl @@ -1,20 +1,26 @@ -(* =========================================================== *) -(* Formal interval evaluation of arithmetic expressions *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -needs "arith/float_pow.hl";; -needs "arith/more_float.hl";; -needs "trig/exp_eval.hl";; -needs "trig/log_eval.hl";; -needs "trig/atn_eval.hl";; -needs "trig/asn_acs_eval.hl";; -needs "trig/cos_eval.hl";; -needs "trig/sin_eval.hl";; -needs "trig/matan_eval.hl";; -needs "misc/misc_functions.hl";; -needs "misc/misc_vars.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of arithmetic expressions *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/arith/float_pow.hl";; +needs "Formal_ineqs/arith/more_float.hl";; +needs "Formal_ineqs/trig/exp_eval.hl";; +needs "Formal_ineqs/trig/log_eval.hl";; +needs "Formal_ineqs/trig/atn_eval.hl";; +needs "Formal_ineqs/trig/asn_acs_eval.hl";; +needs "Formal_ineqs/trig/cos_eval.hl";; +needs "Formal_ineqs/trig/sin_eval.hl";; +needs "Formal_ineqs/trig/matan_eval.hl";; +needs "Formal_ineqs/misc/misc_functions.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; module Eval_interval = struct @@ -45,7 +51,7 @@ let mk_float_interval_decimal = (* Unary interval operations *) -let unary_interval_operations = +let unary_interval_operations = let table = Hashtbl.create 10 in let add = Hashtbl.add in add table `--` (fun pp -> float_interval_neg); @@ -66,7 +72,7 @@ let unary_interval_operations = (* Binary interval operations *) -let binary_interval_operations = +let binary_interval_operations = let table = Hashtbl.create 10 in let add = Hashtbl.add in add table `+` float_interval_add; @@ -125,7 +131,7 @@ let eval_constants pp ifun = match f with | Int_decimal_const tm -> Int_const (mk_float_interval_decimal pp tm) | Int_named_const tm -> Int_const (c_find tm pp) - | Int_pow (n,f1) -> + | Int_pow (n,f1) -> (let f1_val = rec_eval f1 in match f1_val with | Int_const th -> Int_const (float_interval_pow pp n th) @@ -166,7 +172,7 @@ let rec build_interval_fun expr_tm = if ltm = amp_op_real then let n = dest_numeral r_tm in Int_const (mk_float_interval_num n) - else + else let r_fun = build_interval_fun r_tm in Int_unary (ltm, r_fun) else @@ -237,7 +243,7 @@ let find_and_replace_all f_list acc = if c1 > 1 then expr, (c1, fs) else find_and_replace f2 i f_list | _ -> f, (0, f_list) in if c > 1 then expr, (c, fs) else f, replace_subexpr f i f_list in - + let rec iterate fs acc = let i = length acc in let expr, (c, fs') = find_and_replace (hd fs) i fs in @@ -268,7 +274,7 @@ let eval_interval_fun_list pp (f_list, refs) vars = (* Approximate the bounds of the given interval with floating point numbers *) let interval_to_float_interval = - let th = (UNDISCH_ALL o prove)(`interval_arith x (lo, hi) ==> + let th = (UNDISCH_ALL o prove)(`interval_arith x (lo, hi) ==> interval_arith lo (a, y) ==> interval_arith hi (z, b) ==> interval_arith x (a, b)`, @@ -286,7 +292,7 @@ let interval_to_float_interval = a_tm, a_var_real; y_tm, y_var_real; z_tm, z_var_real; b_tm, b_var_real] th in (MY_PROVE_HYP int_th o MY_PROVE_HYP th_lo o MY_PROVE_HYP th_hi) th1;; - + (* Adds a new constant approximation to the table of constants *) let add_constant_interval int_th = diff --git a/Formal_ineqs/arith/float_pow.hl b/Formal_ineqs/arith/float_pow.hl index 0fc178da..22d6f0c4 100644 --- a/Formal_ineqs/arith/float_pow.hl +++ b/Formal_ineqs/arith/float_pow.hl @@ -1,4 +1,16 @@ -needs "arith/arith_float.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of integer exponents *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/arith/arith_float.hl";; module type Float_pow_sig = sig @@ -25,7 +37,7 @@ prioritize_real();; (* Power of a float *) (* ------------------------------------------------ *) -let RULE = UNDISCH_ALL o Arith_nat.NUMERALS_TO_NUM o +let RULE = UNDISCH_ALL o Arith_nat.NUMERALS_TO_NUM o PURE_REWRITE_RULE[FLOAT_OF_NUM; min_exp_def; GSYM IMP_IMP] o SPEC_ALL;; let float_eq n = SYM (FLOAT_TO_NUM_CONV (mk_float n 0));; @@ -82,19 +94,19 @@ let float_pow_pos_double_hi_th = (th_rule o prove) MATCH_MP_TAC REAL_POW_LE THEN REWRITE_TAC[FLOAT_F_POS]);; let float_pow0_hi = (REWRITE_RULE[float_eq 1] o prove) - (`x pow 0 <= &1`, + (`x pow 0 <= &1`, REWRITE_TAC[real_pow; REAL_LE_REFL]);; let float_pow1_hi = (th_rule o prove) - (`x pow 1 <= x`, + (`x pow 1 <= x`, REWRITE_TAC[REAL_POW_1; REAL_LE_REFL]);; let float_pow2_hi = (th_rule o prove) - (`x * x <= hi ==> x pow 2 <= hi`, + (`x * x <= hi ==> x pow 2 <= hi`, REWRITE_TAC[REAL_POW_2]);; let float_pow_pos_suc_lo_th = (th_rule o prove) - (`float_num F n1 e1 <= float_num F n e pow j /\ + (`float_num F n1 e1 <= float_num F n e pow j /\ lo <= float_num F n e * float_num F n1 e1 /\ SUC j = i ==> lo <= float_num F n e pow i`, STRIP_TAC THEN (POP_ASSUM (fun th -> REWRITE_TAC[SYM th])) THEN @@ -112,15 +124,15 @@ let float_pow_pos_double_lo_th = (th_rule o prove) MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_REWRITE_TAC[FLOAT_F_POS; REAL_LE_REFL]);; let float_pow0_lo = (th_rule o REWRITE_RULE[float_eq 1] o prove) - (`&1 <= x pow 0`, + (`&1 <= x pow 0`, REWRITE_TAC[real_pow; REAL_LE_REFL]);; let float_pow1_lo = (th_rule o prove) - (`x <= x pow 1`, + (`x <= x pow 1`, REWRITE_TAC[REAL_POW_1; REAL_LE_REFL]);; let float_pow2_lo = (th_rule o prove) - (`lo <= x * x ==> lo <= x pow 2`, + (`lo <= x * x ==> lo <= x pow 2`, REWRITE_TAC[REAL_POW_2]);; let float_pow_neg_even_hi_th = (th_rule o prove) @@ -206,7 +218,7 @@ let float_pow_pos_double_lo pp x_tm t_th = let i_tm = rand (concl double_j) in let mul_lo = float_mul_lo pp t_tm t_tm in let lo_tm = lhand (concl mul_lo) in - let th0 = INST[x_tm, x_var_real; n1_tm, n1_var_num; e1_tm, e1_var_num; + let th0 = INST[x_tm, x_var_real; n1_tm, n1_var_num; e1_tm, e1_var_num; j_tm, j_var_num; lo_tm, lo_var_real; i_tm, i_var_num] float_pow_pos_double_lo_th in MY_PROVE_HYP t_th (MY_PROVE_HYP mul_lo (MY_PROVE_HYP double_j th0));; @@ -435,7 +447,7 @@ let float_interval_pow pp n x_th = let x, low, high = dest_float_interval (concl x_th) in match n with | 0 -> INST[x, x_var_real] float_interval_pow0 - | 1 -> + | 1 -> let th0 = INST[x, x_var_real; low, lo_var_real; high, hi_var_real] float_interval_pow1 in MY_PROVE_HYP x_th th0 | _ -> @@ -447,7 +459,7 @@ let float_interval_pow pp n x_th = let hi = rand (concl hi_th) and lo = lhand (concl lo_th) and n_tm = rand (rand (concl lo_th)) in - let even_n = eval_even n_tm in + let even_n = eval_even n_tm in let th0 = INST[x, x_var_real; low, low_var_real; high, high_var_real; hi, hi_var_real; lo, lo_var_real; n_tm, n_var_num] float_interval_pow_odd in MY_PROVE_HYP even_n (MY_PROVE_HYP lo_th (MY_PROVE_HYP hi_th (MY_PROVE_HYP x_th th0))) diff --git a/Formal_ineqs/arith/float_theory.hl b/Formal_ineqs/arith/float_theory.hl index e55ad40a..2d0048a7 100644 --- a/Formal_ineqs/arith/float_theory.hl +++ b/Formal_ineqs/arith/float_theory.hl @@ -1,119 +1,124 @@ -(* =========================================================== *) -(* Theoretical results for floating point arithmetic *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - - -(* Dependencies *) -needs "arith/arith_nat.hl";; -needs "arith/num_exp_theory.hl";; -needs "arith/interval_arith.hl";; - -module Float_theory = struct - -open Num_exp_theory;; -open Arith_nat;; -open Interval_arith;; - -(* Fix the minimal exponent *) -let min_exp = !Arith_options.min_exp;; - -(* The main definition *) -let min_exp_num_const = rand (mk_small_numeral_array min_exp);; -let min_exp_const = mk_small_numeral min_exp;; - -let min_exp_def = new_definition (mk_eq(`min_exp:num`, min_exp_const));; - -let float_tm = `float_num s n e = (if s then (-- &1) else &1) * &(num_exp n e) / &(num_exp 1 min_exp)`;; -let float = new_definition float_tm;; - -let arith_base_def = new_definition - (mk_eq (`arith_base:num`, mk_small_numeral Arith_num.arith_base));; - -let arith_base_neq0 = prove - (`~(arith_base = 0)`, - REWRITE_TAC[arith_base_def] THEN ARITH_TAC);; - -let float_alt = prove - (`!s n e. float_num s n e = (if s then -- &1 else &1) * &n * - if (min_exp <= e) then &arith_base pow (e - min_exp) else inv (&arith_base pow (min_exp - e))`, - REPEAT GEN_TAC THEN REWRITE_TAC[float; num_exp; GSYM arith_base_def] THEN - REWRITE_TAC[GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW; REAL_MUL_LID] THEN - REWRITE_TAC[REAL_ARITH `(&n * b) / c = &n * (b / c)`] THEN - SIMP_TAC[arith_base_neq0; REAL_OF_NUM_EQ; REAL_DIV_POW2]);; - -let FLOAT_OF_NUM = (GEN_ALL o prove)(`&n = float_num F n min_exp`, - REWRITE_TAC[float; num_exp; REAL_MUL_LID] THEN - REWRITE_TAC[GSYM REAL_OF_NUM_MUL; REAL_MUL_LID; real_div] THEN - SUBGOAL_THEN (mk_comb(`(~)`, mk_eq(mk_comb(`&`, mk_binop `EXP` base_const `min_exp`), `&0`))) ASSUME_TAC THENL - [ - REWRITE_TAC[REAL_OF_NUM_EQ; EXP_EQ_0] THEN ARITH_TAC; - ALL_TAC - ] THEN - ASM_SIMP_TAC[GSYM REAL_MUL_ASSOC; REAL_MUL_RINV; REAL_MUL_RID]);; - -let FLOAT_NEG = prove(`!s n e. --float_num s n e = float_num (~s) n e`, - REWRITE_TAC[float] THEN REAL_ARITH_TAC);; - -let FLOAT_NEG_F = (GSYM o REWRITE_RULE[] o SPEC `T`) FLOAT_NEG;; -let FLOAT_NEG_T = (GSYM o REWRITE_RULE[] o SPEC `F`) FLOAT_NEG;; - -let FLOAT_F_POS = prove(`!n e. &0 <= float_num F n e`, - REPEAT GEN_TAC THEN REWRITE_TAC[float; REAL_MUL_LID; real_div] THEN - MATCH_MP_TAC REAL_LE_MUL THEN - REWRITE_TAC[REAL_POS; REAL_LE_INV_EQ]);; - -let FLOAT_T_NEG = prove(`!n e. float_num T n e <= &0`, - REPEAT GEN_TAC THEN REWRITE_TAC[float; real_div] THEN - REWRITE_TAC[REAL_ARITH `-- &1 * a * b <= &0 <=> &0 <= a * b`] THEN - MATCH_MP_TAC REAL_LE_MUL THEN - REWRITE_TAC[REAL_POS; REAL_LE_INV_EQ]);; - -let FLOAT_EQ_0 = prove(`!s n e. float_num s n e = &0 <=> n = 0`, - REPEAT GEN_TAC THEN REWRITE_TAC[float; real_div] THEN - REWRITE_TAC[REAL_ENTIRE] THEN - EQ_TAC THENL - [ - STRIP_TAC THEN POP_ASSUM MP_TAC THENL - [ - COND_CASES_TAC THEN REAL_ARITH_TAC; - REWRITE_TAC[REAL_OF_NUM_EQ; NUM_EXP_EQ_0]; - REWRITE_TAC[REAL_INV_EQ_0; REAL_OF_NUM_EQ; NUM_EXP_EQ_0] THEN - ARITH_TAC - ]; - - DISCH_TAC THEN - DISJ2_TAC THEN DISJ1_TAC THEN - ASM_REWRITE_TAC[REAL_OF_NUM_EQ; NUM_EXP_EQ_0] - ]);; - -let FLOAT_F_bound = (GEN_ALL o prove)(`num_exp n1 e1 <= num_exp n2 e2 - ==> float_num F n1 e1 <= float_num F n2 e2`, - DISCH_TAC THEN - REWRITE_TAC[float; REAL_MUL_LID; real_div] THEN - MATCH_MP_TAC REAL_LE_RMUL THEN - ASM_REWRITE_TAC[REAL_OF_NUM_LE; REAL_LE_INV_EQ; REAL_POS]);; - -let FLOAT_T_bound = (GEN_ALL o prove)(`num_exp n1 e1 <= num_exp n2 e2 - ==> float_num T n2 e2 <= float_num T n1 e1`, - REWRITE_TAC[FLOAT_NEG_T; REAL_LE_NEG2; FLOAT_F_bound]);; - - -let FLOAT_INTERVAL_FT_IMP_0 = prove - (`interval_arith x (float_num F n1 e1, float_num T n2 e2) ==> x = &0`, - REWRITE_TAC[interval_arith] THEN STRIP_TAC THEN - REWRITE_TAC[GSYM REAL_LE_ANTISYM] THEN - CONJ_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THENL [ - EXISTS_TAC `float_num T n2 e2` THEN ASM_REWRITE_TAC[FLOAT_T_NEG]; - EXISTS_TAC `float_num F n1 e1` THEN ASM_REWRITE_TAC[FLOAT_F_POS] - ]);; - -let FLOAT_0 = prove(`!s e. float_num s 0 e = &0`, REWRITE_TAC[FLOAT_EQ_0]);; - -let FLOAT_ABS = prove(`abs (float_num s n e) = float_num F n e`, - BOOL_CASES_TAC `s:bool` THEN - REWRITE_TAC[FLOAT_NEG_T; REAL_ABS_NEG; REAL_ABS_REFL; FLOAT_F_POS]);; - - -end;; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Theoretical results for floating point arithmetic *) +(* -------------------------------------------------------------------------- *) + +(* Dependencies *) +needs "Formal_ineqs/arith/arith_nat.hl";; +needs "Formal_ineqs/arith/num_exp_theory.hl";; +needs "Formal_ineqs/arith/interval_arith.hl";; + +module Float_theory = struct + +open Num_exp_theory;; +open Arith_nat;; +open Interval_arith;; + +(* Fix the minimal exponent *) +let min_exp = !Arith_options.min_exp;; + +(* The main definition *) +let min_exp_num_const = rand (mk_small_numeral_array min_exp);; +let min_exp_const = mk_small_numeral min_exp;; + +let min_exp_def = new_definition (mk_eq(`min_exp:num`, min_exp_const));; + +let float_tm = `float_num s n e = (if s then (-- &1) else &1) * &(num_exp n e) / &(num_exp 1 min_exp)`;; +let float = new_definition float_tm;; + +let arith_base_def = new_definition + (mk_eq (`arith_base:num`, mk_small_numeral Arith_num.arith_base));; + +let arith_base_neq0 = prove + (`~(arith_base = 0)`, + REWRITE_TAC[arith_base_def] THEN ARITH_TAC);; + +let float_alt = prove + (`!s n e. float_num s n e = (if s then -- &1 else &1) * &n * + if (min_exp <= e) then &arith_base pow (e - min_exp) else inv (&arith_base pow (min_exp - e))`, + REPEAT GEN_TAC THEN REWRITE_TAC[float; num_exp; GSYM arith_base_def] THEN + REWRITE_TAC[GSYM REAL_OF_NUM_MUL; GSYM REAL_OF_NUM_POW; REAL_MUL_LID] THEN + REWRITE_TAC[REAL_ARITH `(&n * b) / c = &n * (b / c)`] THEN + SIMP_TAC[arith_base_neq0; REAL_OF_NUM_EQ; REAL_DIV_POW2]);; + +let FLOAT_OF_NUM = (GEN_ALL o prove)(`&n = float_num F n min_exp`, + REWRITE_TAC[float; num_exp; REAL_MUL_LID] THEN + REWRITE_TAC[GSYM REAL_OF_NUM_MUL; REAL_MUL_LID; real_div] THEN + SUBGOAL_THEN (mk_comb(`(~)`, mk_eq(mk_comb(`&`, mk_binop `EXP` base_const `min_exp`), `&0`))) ASSUME_TAC THENL + [ + REWRITE_TAC[REAL_OF_NUM_EQ; EXP_EQ_0] THEN ARITH_TAC; + ALL_TAC + ] THEN + ASM_SIMP_TAC[GSYM REAL_MUL_ASSOC; REAL_MUL_RINV; REAL_MUL_RID]);; + +let FLOAT_NEG = prove(`!s n e. --float_num s n e = float_num (~s) n e`, + REWRITE_TAC[float] THEN REAL_ARITH_TAC);; + +let FLOAT_NEG_F = (GSYM o REWRITE_RULE[] o SPEC `T`) FLOAT_NEG;; +let FLOAT_NEG_T = (GSYM o REWRITE_RULE[] o SPEC `F`) FLOAT_NEG;; + +let FLOAT_F_POS = prove(`!n e. &0 <= float_num F n e`, + REPEAT GEN_TAC THEN REWRITE_TAC[float; REAL_MUL_LID; real_div] THEN + MATCH_MP_TAC REAL_LE_MUL THEN + REWRITE_TAC[REAL_POS; REAL_LE_INV_EQ]);; + +let FLOAT_T_NEG = prove(`!n e. float_num T n e <= &0`, + REPEAT GEN_TAC THEN REWRITE_TAC[float; real_div] THEN + REWRITE_TAC[REAL_ARITH `-- &1 * a * b <= &0 <=> &0 <= a * b`] THEN + MATCH_MP_TAC REAL_LE_MUL THEN + REWRITE_TAC[REAL_POS; REAL_LE_INV_EQ]);; + +let FLOAT_EQ_0 = prove(`!s n e. float_num s n e = &0 <=> n = 0`, + REPEAT GEN_TAC THEN REWRITE_TAC[float; real_div] THEN + REWRITE_TAC[REAL_ENTIRE] THEN + EQ_TAC THENL + [ + STRIP_TAC THEN POP_ASSUM MP_TAC THENL + [ + COND_CASES_TAC THEN REAL_ARITH_TAC; + REWRITE_TAC[REAL_OF_NUM_EQ; NUM_EXP_EQ_0]; + REWRITE_TAC[REAL_INV_EQ_0; REAL_OF_NUM_EQ; NUM_EXP_EQ_0] THEN + ARITH_TAC + ]; + + DISCH_TAC THEN + DISJ2_TAC THEN DISJ1_TAC THEN + ASM_REWRITE_TAC[REAL_OF_NUM_EQ; NUM_EXP_EQ_0] + ]);; + +let FLOAT_F_bound = (GEN_ALL o prove)(`num_exp n1 e1 <= num_exp n2 e2 + ==> float_num F n1 e1 <= float_num F n2 e2`, + DISCH_TAC THEN + REWRITE_TAC[float; REAL_MUL_LID; real_div] THEN + MATCH_MP_TAC REAL_LE_RMUL THEN + ASM_REWRITE_TAC[REAL_OF_NUM_LE; REAL_LE_INV_EQ; REAL_POS]);; + +let FLOAT_T_bound = (GEN_ALL o prove)(`num_exp n1 e1 <= num_exp n2 e2 + ==> float_num T n2 e2 <= float_num T n1 e1`, + REWRITE_TAC[FLOAT_NEG_T; REAL_LE_NEG2; FLOAT_F_bound]);; + + +let FLOAT_INTERVAL_FT_IMP_0 = prove + (`interval_arith x (float_num F n1 e1, float_num T n2 e2) ==> x = &0`, + REWRITE_TAC[interval_arith] THEN STRIP_TAC THEN + REWRITE_TAC[GSYM REAL_LE_ANTISYM] THEN + CONJ_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THENL [ + EXISTS_TAC `float_num T n2 e2` THEN ASM_REWRITE_TAC[FLOAT_T_NEG]; + EXISTS_TAC `float_num F n1 e1` THEN ASM_REWRITE_TAC[FLOAT_F_POS] + ]);; + +let FLOAT_0 = prove(`!s e. float_num s 0 e = &0`, REWRITE_TAC[FLOAT_EQ_0]);; + +let FLOAT_ABS = prove(`abs (float_num s n e) = float_num F n e`, + BOOL_CASES_TAC `s:bool` THEN + REWRITE_TAC[FLOAT_NEG_T; REAL_ABS_NEG; REAL_ABS_REFL; FLOAT_F_POS]);; + + +end;; diff --git a/Formal_ineqs/arith/interval_arith.hl b/Formal_ineqs/arith/interval_arith.hl index 2f1f1dd9..264ad0ca 100644 --- a/Formal_ineqs/arith/interval_arith.hl +++ b/Formal_ineqs/arith/interval_arith.hl @@ -1,68 +1,74 @@ -(* =========================================================== *) -(* Theoretical results for interval arithmetic *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -needs "misc/misc_vars.hl";; - -module Interval_arith = struct - -prioritize_real();; - -(*******************************) -(* The main definition *) -let interval_arith = - new_definition `interval_arith (x:real) (lo, hi) <=> lo <= x /\ x <= hi`;; - -(* Additional definitions *) -let bounded_on = new_definition `bounded_on f s f_bounds <=> - !x. x IN s ==> interval_arith (f x) f_bounds`;; - -let bounded_on_int = new_definition `bounded_on_int f int f_bounds <=> - !x. interval_arith x int ==> interval_arith (f x) f_bounds`;; - -let iabs = new_definition `iabs (x_lo, x_hi) = max x_hi (-- x_lo)`;; -let interval_not_zero = new_definition `interval_not_zero (lo, hi) <=> &0 < lo \/ hi < &0`;; -let interval_pos = new_definition `interval_pos (lo, hi) <=> &0 < lo`;; -let interval_neg = new_definition `interval_neg (lo, hi) <=> hi < &0`;; -let interval_gt = new_definition `interval_gt f (lo, hi) <=> f < lo`;; - - -(********************************) -(* Lemmas *) -let CONST_INTERVAL = prove(`!x. interval_arith x (x,x)`, - REWRITE_TAC[interval_arith; REAL_LE_REFL]);; - -let APPROX_INTERVAL = (GEN_ALL o prove)(`(a <= lo /\ hi <= b) /\ interval_arith x (lo, hi) - ==> interval_arith x (a,b)`, - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; - -let INTERVAL_NEG = (GEN_ALL o prove)(`interval_arith x (a, b) ==> - interval_arith (--x) (--b, --a)`, - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; - -let INTERVAL_NEG_EQ = prove(`!x a b. interval_arith (--x) (--b, --a) <=> - interval_arith x (a, b)`, - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; - - - -(**************************************) -(* Conversions *) -open Misc_vars;; - -let interval_tm = `interval_arith`;; - -let dest_interval_arith tm = - let lhs, int_tm = dest_comb tm in - rand lhs, int_tm;; - -let mk_interval tm bounds = - mk_comb (mk_comb (interval_tm, tm), bounds);; - -let mk_const_interval = - let lemma = SPEC_ALL CONST_INTERVAL in - fun tm -> INST[tm, x_var_real] lemma;; - -end;; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Theoretical results for interval arithmetic *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/misc/misc_vars.hl";; + +module Interval_arith = struct + +prioritize_real();; + +(*******************************) +(* The main definition *) +let interval_arith = + new_definition `interval_arith (x:real) (lo, hi) <=> lo <= x /\ x <= hi`;; + +(* Additional definitions *) +let bounded_on = new_definition `bounded_on f s f_bounds <=> + !x. x IN s ==> interval_arith (f x) f_bounds`;; + +let bounded_on_int = new_definition `bounded_on_int f int f_bounds <=> + !x. interval_arith x int ==> interval_arith (f x) f_bounds`;; + +let iabs = new_definition `iabs (x_lo, x_hi) = max x_hi (-- x_lo)`;; +let interval_not_zero = new_definition `interval_not_zero (lo, hi) <=> &0 < lo \/ hi < &0`;; +let interval_pos = new_definition `interval_pos (lo, hi) <=> &0 < lo`;; +let interval_neg = new_definition `interval_neg (lo, hi) <=> hi < &0`;; +let interval_gt = new_definition `interval_gt f (lo, hi) <=> f < lo`;; + + +(********************************) +(* Lemmas *) +let CONST_INTERVAL = prove(`!x. interval_arith x (x,x)`, + REWRITE_TAC[interval_arith; REAL_LE_REFL]);; + +let APPROX_INTERVAL = (GEN_ALL o prove)(`(a <= lo /\ hi <= b) /\ interval_arith x (lo, hi) + ==> interval_arith x (a,b)`, + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; + +let INTERVAL_NEG = (GEN_ALL o prove)(`interval_arith x (a, b) ==> + interval_arith (--x) (--b, --a)`, + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; + +let INTERVAL_NEG_EQ = prove(`!x a b. interval_arith (--x) (--b, --a) <=> + interval_arith x (a, b)`, + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; + + + +(**************************************) +(* Conversions *) +open Misc_vars;; + +let interval_tm = `interval_arith`;; + +let dest_interval_arith tm = + let lhs, int_tm = dest_comb tm in + rand lhs, int_tm;; + +let mk_interval tm bounds = + mk_comb (mk_comb (interval_tm, tm), bounds);; + +let mk_const_interval = + let lemma = SPEC_ALL CONST_INTERVAL in + fun tm -> INST[tm, x_var_real] lemma;; + +end;; diff --git a/Formal_ineqs/arith/more_float.hl b/Formal_ineqs/arith/more_float.hl index 008f1609..74eef748 100644 --- a/Formal_ineqs/arith/more_float.hl +++ b/Formal_ineqs/arith/more_float.hl @@ -1,596 +1,602 @@ -(* =========================================================== *) -(* Additional floating point procedures *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -needs "arith/arith_float.hl";; -needs "misc/misc_vars.hl";; - -module More_float = struct - -open Misc_functions;; -open Float_theory;; -open Interval_arith;; -open Arith_float;; -open Misc_vars;; - -prioritize_real();; - -let RULE = UNDISCH_ALL o Arith_nat.NUMERALS_TO_NUM o REWRITE_RULE[FLOAT_OF_NUM; min_exp_def; GSYM IMP_IMP] o SPEC_ALL;; - -(*************************************) -(* More float *) - - -(* Converts a float term to the corresponding rational number *) -let num_of_float_tm tm = - let s, n_tm, e_tm = dest_float tm in - let b = num Arith_num.arith_base in - let m = num Float_theory.min_exp in - let ( * ), (^), (-), (!) = ( */ ), ( **/ ), (-/), Arith_nat.raw_dest_hash in - let r = !n_tm * (b ^ (!e_tm - m)) in - if s = "T" then minus_num r else r;; - -(* Converts a float term to a floating-point number *) -(* Note: float_of_num gives a very bad approximation in some cases *) -let float_of_float_tm tm = - (float_of_string o approx_num_exp 30 o num_of_float_tm) tm;; - -(* Converts a floating-point number into a float term *) -let float_tm_of_float = - let split = - let b = float_of_int Arith_num.arith_base in - let log_b = log b in - let rec fix t k = - if t = 0.0 || t = infinity || t = nan then - t, k - else if t < 1.0 then - fix (t *. b) (k - 1) - else if t >= b then - fix (t /. b) (k + 1) - else - t, k in - fun f -> - if f = 0.0 then - false, f, 0 - else - let s, x = if f < 0.0 then true, -.f else false, f in - let e = -.floor (log x /. log_b) in - let k = -int_of_float e in - let t, n = fix (x *. (b ** e)) k in - s, t, n - and extract = - let b = float_of_int Arith_num.arith_base in - let nb = num Arith_num.arith_base in - let rec step k f acc = - if k <= 0 then acc - else - let d = int_of_float f in - step (k - 1) ((f -. float_of_int d) *. b) (num d +/ (nb */ acc)) in - fun pp f -> - step pp f (num 0) - in - fun pp f -> - let s, x, k = split f in - let n = extract pp x in - let n_tm = rand (Arith_num.mk_numeral_array n) in - let e_tm = rand (Arith_num.mk_small_numeral_array (min_exp + k - pp + 1)) in - let sign = if s then "T" else "F" in - make_float sign n_tm e_tm;; - -(* Creates a (not normalized) floating-point number from a given natural number. *) -(* Returns the theorem |- &n = float F n min_exp and the term n *) -let float_eq_th_of_num = - let conv_th = (Arith_nat.NUMERALS_TO_NUM o REWRITE_RULE[min_exp_def] o SPEC_ALL) FLOAT_OF_NUM in - fun n -> - let _ = if sign_num n < 0 then failwith "float_eq_th_of_num: negative number" in - let n_tm = rand (Arith_nat.mk_numeral_array n) in - INST[n_tm, n_var_num] conv_th, n_tm;; - -(* |- ##0 = &0, |- ##1 = &1, |- ##2 = &2, |- ##3 = &3, |- ##4 = &4 *) -let float0_eq = FLOAT_TO_NUM_CONV (mk_float 0 0) and - float1_eq = FLOAT_TO_NUM_CONV (mk_float 1 0) and - float2_eq = FLOAT_TO_NUM_CONV (mk_float 2 0) and - float3_eq = FLOAT_TO_NUM_CONV (mk_float 3 0) and - float4_eq = FLOAT_TO_NUM_CONV (mk_float 4 0);; - -(* |- D_k _0 = k for k = 1, 2, 3 *) -let num1_eq, num2_eq, num3_eq = - let conv = SYM o REWRITE_RULE[Arith_num.NUM_THM] o Arith_nat.NUMERAL_TO_NUM_CONV in - conv `1`, conv `2`, conv `3`;; - -(*********************) - -let float_F_const = `float_num F`;; - -let mk_float_small n = - let n_tm0 = mk_small_numeral n in - let n_th = Arith_nat.NUMERAL_TO_NUM_CONV n_tm0 in - let n_tm = rand(rand(concl n_th)) in - mk_comb(mk_comb(float_F_const, n_tm), min_exp_num_const);; - -(* Small float constants and intervals *) -let one_float = mk_float_small 1 and - two_float = mk_float_small 2 and - one_interval = mk_float_interval_small_num 1 and - two_interval = mk_float_interval_small_num 2;; -let neg_two_interval = float_interval_neg two_interval;; - -(* ----------------------------------------- *) -(* float_eq0 *) -(* ----------------------------------------- *) - -let FLOAT_EQ_0' = (GEN_REWRITE_RULE (RAND_CONV o RAND_CONV) [NUMERAL] o SPEC_ALL) - FLOAT_EQ_0;; - - -let float_eq0 f_tm = - let lhs, e_tm = dest_comb f_tm in - let lhs2, n_tm = dest_comb lhs in - let th0 = INST[rand lhs2, s_var_bool; n_tm, n_var_num; e_tm, e_var_num] FLOAT_EQ_0' in - let eq_th = Arith_nat.raw_eq0_hash_conv n_tm in - TRANS th0 eq_th;; - - - -(* ----------------------------------------- *) -(* float_interval_scale *) -(* ----------------------------------------- *) - - -let float_interval_scale pp c_tm th = - let c_th = mk_const_interval c_tm in - float_interval_mul pp c_th th;; - - -(* ----------------------------------------- *) -(* float_interval_lt0 *) -(* ----------------------------------------- *) - - -let FLOAT_INTERVAL_LT0' = (UNDISCH_ALL o prove) - (`interval_arith x (lo, hi) ==> (hi < &0 <=> T) ==> x < &0`, - REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; - -let float_interval_lt0 th = - let x_tm, bounds = dest_interval_arith (concl th) in - let lo_tm, hi_tm = dest_pair bounds in - let lt0_th = float_lt0 hi_tm in - let _ = ((rand o concl) lt0_th = t_const) || failwith "float_interval_lt0: &0 <= hi" in - let th0 = INST[x_tm, x_var_real; lo_tm, lo_var_real; hi_tm, hi_var_real] FLOAT_INTERVAL_LT0' in - (MY_PROVE_HYP th o MY_PROVE_HYP lt0_th) th0;; - - -(* ----------------------------------------- *) -(* float_pos *) -(* ----------------------------------------- *) - -let FLOAT_F_POS' = SPEC_ALL FLOAT_F_POS;; - -(* Returns &0 <= float F n e *) -let float_pos tm = - let _, n_tm, e_tm = dest_float tm in - INST[n_tm, n_var_num; e_tm, e_var_num] FLOAT_F_POS';; - - - -(* ----------------------------------------- *) -(* float_iabs *) -(* ----------------------------------------- *) - -let FLOAT_NEG_F' = RULE FLOAT_NEG_T;; -let FLOAT_NEG_T' = RULE FLOAT_NEG_F;; - -let float_neg tm = - let sign, n_tm, e_tm = dest_float tm in - if sign = "T" then - INST[n_tm, n_var_num; e_tm, e_var_num] FLOAT_NEG_T' - else - INST[n_tm, n_var_num; e_tm, e_var_num] FLOAT_NEG_F';; - - -let IABS' = RULE iabs;; - - -let float_iabs int_tm = - let lo_tm, hi_tm = dest_pair int_tm in - let neg_lo_th = float_neg lo_tm in - let max_th = SYM (float_max hi_tm ((rand o rator o concl) neg_lo_th)) in - let lhs, rhs = dest_comb (concl max_th) in - let th0 = SYM (EQ_MP (AP_TERM lhs (AP_TERM (rator rhs) neg_lo_th)) max_th) in - let th1 = INST[lo_tm, x_lo_var; hi_tm, x_hi_var] IABS' in - TRANS th1 th0;; - - -let FLOAT_IABS_FF = prove(`iabs (float_num F n1 e1, float_num F n2 e2) = float_num F n2 e2`, - REWRITE_TAC[iabs] THEN - MP_TAC (SPECL [`n1:num`; `e1:num`] FLOAT_F_POS) THEN - MP_TAC (SPECL [`n2:num`; `e2:num`] FLOAT_F_POS) THEN - REAL_ARITH_TAC);; - - -let FLOAT_IABS_TT = prove(`iabs (float_num T n1 e1, float_num T n2 e2) = float_num F n1 e1`, - REWRITE_TAC[iabs; GSYM FLOAT_NEG_F] THEN - MP_TAC (SPECL [`n1:num`; `e1:num`] FLOAT_F_POS) THEN - MP_TAC (SPECL [`n2:num`; `e2:num`] FLOAT_T_NEG) THEN - REAL_ARITH_TAC);; - - -(* ----------------------------------------- *) -(* interval_not_zero *) -(* ----------------------------------------- *) - -let INTERVAL_NOT_ZERO1' = (UNDISCH_ALL o prove) - (`(&0 < lo <=> T) ==> interval_not_zero (lo, hi)`, SIMP_TAC[interval_not_zero]);; -let INTERVAL_NOT_ZERO2' = (UNDISCH_ALL o prove) - (`(hi < &0 <=> T) ==> interval_not_zero (lo, hi)`, SIMP_TAC[interval_not_zero]);; - - -let check_interval_not_zero int_tm = - let lo, hi = dest_pair int_tm in - let inst = INST[lo, lo_var_real; hi, hi_var_real] in - let s1, _, _ = dest_float lo in - if s1 = "F" then - let gt_th = float_gt0 lo in - if (fst o dest_const o rand o concl) gt_th <> "T" then - failwith "check_interval_not_zero: &0 < lo <=> F" - else - (MY_PROVE_HYP gt_th o inst) INTERVAL_NOT_ZERO1' - else - let lt_th = float_lt0 hi in - if (fst o dest_const o rand o concl) lt_th <> "T" then - failwith "check_interval_not_zero: hi < &0 <=> F" - else - (MY_PROVE_HYP lt_th o inst) INTERVAL_NOT_ZERO2';; - - - -(* ----------------------------------------- *) -(* interval_pos *) -(* ----------------------------------------- *) - -let INTERVAL_POS' = (UNDISCH_ALL o prove) - (`(&0 < lo <=> T) ==> interval_pos (lo, hi:real)`, SIMP_TAC[interval_pos]);; - - -let check_interval_pos int_tm = - let lo, hi = dest_pair int_tm in - let gt_th = float_gt0 lo in - if (fst o dest_const o rand o concl) gt_th <> "T" then - failwith "check_interval_pos: &0 < lo <=> F" - else - (MY_PROVE_HYP gt_th o INST[lo, lo_var_real; hi, hi_var_real]) INTERVAL_POS';; - - -(* ----------------------------------------- *) -(* interval_neg *) -(* ----------------------------------------- *) - -let INTERVAL_NEG' = (UNDISCH_ALL o prove) - (`(hi < &0 <=> T) ==> interval_neg (lo:real, hi:real)`, SIMP_TAC[interval_neg]);; - - -let check_interval_neg int_tm = - let lo, hi = dest_pair int_tm in - let cmp_th = float_lt0 hi in - if (fst o dest_const o rand o concl) cmp_th <> "T" then - failwith "check_interval_neg: hi < &0 <=> F" - else - (MY_PROVE_HYP cmp_th o INST[lo, lo_var_real; hi, hi_var_real]) INTERVAL_NEG';; - -(* ----------------------------------------- *) -(* interval_gt *) -(* ----------------------------------------- *) - -let check_interval_gt = - let gt_th0 = (UNDISCH_ALL o prove)(`(y < lo <=> T) ==> interval_gt y (lo, hi:real)`, - REWRITE_TAC[interval_gt]) in - fun y_tm int_tm -> - let lo, hi = dest_pair int_tm in - let lt_th = float_lt y_tm lo in - if (fst o dest_const o rand o concl) lt_th <> "T" then - failwith "check_interval_gt: y < lo <=> F" - else - (MY_PROVE_HYP lt_th o - INST[lo, lo_var_real; hi, hi_var_real; y_tm, y_var_real]) gt_th0;; - - -(* ----------------------------------------- *) -(* check_interval_iabs *) -(* ----------------------------------------- *) - - -(* proves |- iabs int < rhs <=> T *) -let check_interval_iabs int_tm rhs_tm = - let iabs_eq = float_iabs int_tm in - let lt_th = float_lt (rand (concl iabs_eq)) rhs_tm in - if (fst o dest_const o rand o concl) lt_th <> "T" then - failwith "check_interval_iabs: iabs < rhs <=> F" - else - let th0 = AP_THM (AP_TERM lt_op_real iabs_eq) rhs_tm in - TRANS th0 lt_th;; - - - -(* ----------------------------------------- *) -(* inv *) -(* ----------------------------------------- *) - -let float_inv_hi = - let div_lemma = (RULE o prove) - (`&1 / x <= hi ==> inv x <= hi`, - REWRITE_TAC[real_div; REAL_MUL_LID]) in - fun pp tm -> - let div_th = float_div_hi pp one_float tm in - let hi_tm = rand (concl div_th) in - let th0 = INST[tm, x_var_real; hi_tm, hi_var_real] div_lemma in - MY_PROVE_HYP div_th th0;; - -let float_inv_lo = - let div_lemma = (RULE o prove) - (`lo <= &1 / x ==> lo <= inv x`, - REWRITE_TAC[real_div; REAL_MUL_LID]) in - fun pp tm -> - let div_th = float_div_lo pp one_float tm in - let lo_tm = rand (rator (concl div_th)) in - let th0 = INST[tm, x_var_real; lo_tm, lo_var_real] div_lemma in - MY_PROVE_HYP div_th th0;; - -let float_interval_inv = - let INV_EQ_DIV_LEMMA = prove(`&1 / x = inv x`, REWRITE_TAC[real_div; REAL_MUL_LID]) in - fun pp th -> - let x_tm = (rand o rator o concl) th in - let div_th = INST[x_tm, x_var_real] INV_EQ_DIV_LEMMA in - let th0 = float_interval_div pp one_interval th in - let lhs, rhs = dest_comb (concl th0) in - let lhs2, rhs2 = dest_comb lhs in - EQ_MP (AP_THM (AP_TERM lhs2 div_th) rhs) th0;; - - -(* Explicit representation of inv(&2) *) -let float_inv2_th = - let one_float_eq_one = FLOAT_TO_NUM_CONV one_float in - let inv2_eq_lemma = prove(`interval_arith (&2 * x) (&1, &1) ==> inv (&2) = x`, - REWRITE_TAC[interval_arith] THEN CONV_TAC REAL_FIELD) in - let half_tm = (fst o dest_pair o rand o concl) (float_interval_inv 1 two_interval) in - let half_interval = mk_const_interval half_tm in - let mul_th = REWRITE_RULE[one_float_eq_one] (float_interval_mul 2 two_interval half_interval) in - MATCH_MP inv2_eq_lemma mul_th;; - -let float_inv2 = rand (concl float_inv2_th);; -let inv2_interval = mk_const_interval float_inv2;; - - - -(*****************************************) -(* bounded_on_int *) - -let norm_derivative d_th eq_th = - let lhs, rhs = (dest_eq o concl) d_th in - let lhs2, rhs2 = dest_comb lhs in - let th0 = AP_THM (AP_TERM (rator lhs2) eq_th) rhs2 in - TRANS (SYM th0) d_th;; - -let norm_diff d_th eq_th = - let lhs, rhs = (dest_comb o concl) d_th in - let th0 = AP_THM (AP_TERM (rator lhs) eq_th) rhs in - EQ_MP th0 d_th;; - -let norm_interval int_th eq_th = - let lhs, rhs = (dest_comb o concl) int_th in - let th0 = AP_THM (AP_TERM (rator lhs) eq_th) rhs in - EQ_MP (SYM th0) int_th;; - -let norm_second_derivative th eq_th = - let lhs, dd_bounds = dest_comb (concl th) in - let lhs2, int_tm = dest_comb lhs in - let th0 = AP_THM (AP_THM (AP_TERM (rator lhs2) eq_th) int_tm) dd_bounds in - EQ_MP th0 th;; - -let norm_lin_approx th eq_th = - let lhs, df_bounds = dest_comb (concl th) in - let lhs2, f_bounds = dest_comb lhs in - let lhs3, x_tm = dest_comb lhs2 in - let th0 = AP_THM (AP_THM (AP_THM (AP_TERM (rator lhs3) eq_th) x_tm) f_bounds) df_bounds in - EQ_MP th0 th;; - - - -let BOUNDED_ON_INT = (UNDISCH_ALL o prove)(`(!x. interval_arith x int ==> - interval_arith (f x) f_bounds) - ==> bounded_on_int f int f_bounds`, - REWRITE_TAC[bounded_on_int; interval_arith]);; - - -let BOUNDED_ON_INT_DEST = (UNDISCH_ALL o prove)(`bounded_on_int f int f_bounds ==> - (!x. interval_arith x int ==> interval_arith (f x) f_bounds)`, - REWRITE_TAC[bounded_on_int; interval_arith]);; - - -(* Given a theorem (interval_arith x int |- interval_arith (f x) f_bounds), yields - |- bounded_on_int (\x. f x) int f_bounds *) -let mk_bounded_on_int th = - let int_tm = (rand o hd o hyp) th in - let lhs, f_bounds_tm = dest_comb (concl th) in - let lhs2, rhs2 = dest_comb lhs in - let f_tm = mk_abs (x_var_real, rhs2) in - let b_th0 = (SYM o BETA_CONV) (mk_comb (f_tm , x_var_real)) in - let b_th1 = AP_THM (AP_TERM lhs2 b_th0) f_bounds_tm in - let th2 = EQ_MP b_th1 th in - let th3 = DISCH_ALL th2 in - let th4 = GEN x_var_real th3 in - let th_int = INST[int_tm, int_var; f_bounds_tm, f_bounds_var; - f_tm, f_var_fun] BOUNDED_ON_INT in - MY_PROVE_HYP th4 th_int;; - - -let dest_bounded_on_int th = - let lhs, f_bounds = dest_comb (concl th) in - let lhs2, int_tm = dest_comb lhs in - let f_tm = rand lhs2 in - let th0 = INST[f_tm, f_var_fun; int_tm, int_var; f_bounds, f_bounds_var] BOUNDED_ON_INT_DEST in - let th1 = UNDISCH_ALL (SPEC x_var_real (MY_PROVE_HYP th th0)) in - if is_abs f_tm then - let f_tm = (rand o rator o concl) th1 in - let eq_th = BETA_CONV f_tm in - norm_interval th1 (SYM eq_th) - else - th1;; - - -let dest_bounded_on_int_raw th = - let lhs, f_bounds = dest_comb (concl th) in - let lhs2, int_tm = dest_comb lhs in - let f_tm = rand lhs2 in - let th0 = INST[f_tm, f_var_fun; int_tm, int_var; f_bounds, f_bounds_var] BOUNDED_ON_INT_DEST in - UNDISCH_ALL (SPEC x_var_real (MY_PROVE_HYP th th0));; - - -(***********************************) -(* bounded_on_int arithmetic *) - -let bounded_on_int_scale pp c_tm th = - let i_th = dest_bounded_on_int th in - let th0 = float_interval_scale pp c_tm i_th in - mk_bounded_on_int th0;; - - -let bounded_on_int_mul_int pp int_th th = - let i_th = dest_bounded_on_int th in - let th0 = float_interval_mul pp int_th i_th in - mk_bounded_on_int th0;; - - -let bounded_on_int_neg th1 = - let i_th = dest_bounded_on_int th1 in - let th0 = float_interval_neg i_th in - mk_bounded_on_int th0;; - - -let bounded_on_int_add pp th1 th2 = - let i_th1, i_th2 = dest_bounded_on_int th1, dest_bounded_on_int th2 in - let th0 = float_interval_add pp i_th1 i_th2 in - mk_bounded_on_int th0;; - - -let bounded_on_int_sub pp th1 th2 = - let i_th1, i_th2 = dest_bounded_on_int th1, dest_bounded_on_int th2 in - let th0 = float_interval_sub pp i_th1 i_th2 in - mk_bounded_on_int th0;; - - -let bounded_on_int_mul pp th1 th2 = - let i_th1, i_th2 = dest_bounded_on_int th1, dest_bounded_on_int th2 in - let th0 = float_interval_mul pp i_th1 i_th2 in - mk_bounded_on_int th0;; - - -let bounded_on_int_mul_raw pp th1 th2 = - let i_th1, i_th2 = dest_bounded_on_int_raw th1, dest_bounded_on_int_raw th2 in - let th0 = float_interval_mul pp i_th1 i_th2 in - mk_bounded_on_int th0;; - - - -let bounded_on_int_div pp th1 th2 = - let i_th1, i_th2 = dest_bounded_on_int th1, dest_bounded_on_int th2 in - let th0 = float_interval_div pp i_th1 i_th2 in - mk_bounded_on_int th0;; - - -(************************************) -let ADD_INEQ_HI = (RULE o REAL_ARITH) `x1 <= y1 /\ x2 <= y2 /\ y1 + y2 <= y ==> x1 + x2 <= y`;; -let ADD_INEQ_LO = (RULE o REAL_ARITH) `x1 <= y1 /\ x2 <= y2 /\ x <= x1 + x2 ==> x <= y1 + y2`;; -let SUB_INEQ_HI = (RULE o REAL_ARITH) `x1 <= y1 /\ y2 <= x2 /\ y1 - y2 <= y ==> x1 - x2 <= y`;; -let SUB_INEQ_LO = (RULE o REAL_ARITH) `x1 <= y1 /\ y2 <= x2 /\ x <= x1 - x2 ==> x <= y1 - y2`;; -let MUL_INEQ_HI = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove) - (`&0 <= x1 /\ &0 <= x2 /\ x1 <= y1 /\ x2 <= y2 /\ y1 * y2 <= y ==> x1 * x2 <= y`, - DISCH_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN - EXISTS_TAC `y1 * y2` THEN ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_REWRITE_TAC[]);; - -let MUL_INEQ_POS_CONST_HI = (UNDISCH_ALL o prove) - (`(&0 <= x <=> T) ==> y1 <= y2 ==> x * y2 <= z ==> x * y1 <= z`, - REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `x * y2` THEN - ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_LMUL THEN ASM_REWRITE_TAC[]);; - - -let mk_refl_ineq = - let REAL_LE_REFL' = RULE REAL_LE_REFL in - fun tm -> INST[tm, x_var_real] REAL_LE_REFL';; - -let dest_le_op ineq = - let lhs, y_tm = dest_comb ineq in - (rand lhs, y_tm);; - - -let mul_ineq_pos_const_hi pp c_tm ineq = - let y1_tm, y2_tm = dest_le_op (concl ineq) in - let ge0_th = float_ge0 c_tm in - let mul_hi_th = float_mul_hi pp c_tm y2_tm in - let z_tm = (rand o concl) mul_hi_th in - (MY_PROVE_HYP ge0_th o MY_PROVE_HYP ineq o MY_PROVE_HYP mul_hi_th o - INST[c_tm, x_var_real; y1_tm, y1_var_real; y2_tm, y2_var_real; z_tm, z_var_real]) - MUL_INEQ_POS_CONST_HI;; - - -let mul_ineq_hi pp ineq1 ineq2 = - let x1_tm, y1_tm = dest_le_op (concl ineq1) in - let x2_tm, y2_tm = dest_le_op (concl ineq2) in - let x1_pos, x2_pos = float_pos x1_tm, float_pos x2_tm in - let rhs_mul = float_mul_hi pp y1_tm y2_tm in - let y_tm = (rand o concl) rhs_mul in - let th0 = INST[x1_tm, x1_var_real; y1_tm, y1_var_real; - x2_tm, x2_var_real; y2_tm, y2_var_real; - y_tm, y_var_real] MUL_INEQ_HI in - (MY_PROVE_HYP x1_pos o MY_PROVE_HYP x2_pos o MY_PROVE_HYP ineq1 o - MY_PROVE_HYP ineq2 o MY_PROVE_HYP rhs_mul) th0;; - - - -let sub_ineq_hi pp ineq1 ineq2 = - let x1_tm, y1_tm = dest_le_op (concl ineq1) in - let y2_tm, x2_tm = dest_le_op (concl ineq2) in - let rhs_sub = float_sub_hi pp y1_tm y2_tm in - let y_tm = (rand o concl) rhs_sub in - let th0 = INST[x1_tm, x1_var_real; y1_tm, y1_var_real; - x2_tm, x2_var_real; y2_tm, y2_var_real; - y_tm, y_var_real] SUB_INEQ_HI in - MY_PROVE_HYP ineq1 (MY_PROVE_HYP ineq2 (MY_PROVE_HYP rhs_sub th0));; - - -let sub_ineq_lo pp ineq1 ineq2 = - let x1_tm, y1_tm = dest_le_op (concl ineq1) in - let y2_tm, x2_tm = dest_le_op (concl ineq2) in - let lhs_sub = float_sub_lo pp x1_tm x2_tm in - let x_tm = (lhand o concl) lhs_sub in - let th0 = INST[x1_tm, x1_var_real; y1_tm, y1_var_real; - x2_tm, x2_var_real; y2_tm, y2_var_real; - x_tm, x_var_real] SUB_INEQ_LO in - MY_PROVE_HYP ineq1 (MY_PROVE_HYP ineq2 (MY_PROVE_HYP lhs_sub th0));; - - -let add_ineq_hi pp ineq1 ineq2 = - let x1_tm, y1_tm = dest_le_op (concl ineq1) in - let x2_tm, y2_tm = dest_le_op (concl ineq2) in - let rhs_sum = float_add_hi pp y1_tm y2_tm in - let y_tm = (rand o concl) rhs_sum in - let th0 = INST[x1_tm, x1_var_real; y1_tm, y1_var_real; - x2_tm, x2_var_real; y2_tm, y2_var_real; - y_tm, y_var_real] ADD_INEQ_HI in - MY_PROVE_HYP ineq1 (MY_PROVE_HYP ineq2 (MY_PROVE_HYP rhs_sum th0));; - - -let add_ineq_lo pp ineq1 ineq2 = - let x1_tm, y1_tm = dest_le_op (concl ineq1) in - let x2_tm, y2_tm = dest_le_op (concl ineq2) in - let lhs_sum = float_add_lo pp x1_tm x2_tm in - let x_tm = (lhand o concl) lhs_sum in - let th0 = INST[x1_tm, x1_var_real; y1_tm, y1_var_real; - x2_tm, x2_var_real; y2_tm, y2_var_real; - x_tm, x_var_real] ADD_INEQ_LO in - MY_PROVE_HYP ineq1 (MY_PROVE_HYP ineq2 (MY_PROVE_HYP lhs_sum th0));; - - -end;; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Additional floating point procedures *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/arith/arith_float.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; + +module More_float = struct + +open Misc_functions;; +open Float_theory;; +open Interval_arith;; +open Arith_float;; +open Misc_vars;; + +prioritize_real();; + +let RULE = UNDISCH_ALL o Arith_nat.NUMERALS_TO_NUM o REWRITE_RULE[FLOAT_OF_NUM; min_exp_def; GSYM IMP_IMP] o SPEC_ALL;; + +(*************************************) +(* More float *) + + +(* Converts a float term to the corresponding rational number *) +let num_of_float_tm tm = + let s, n_tm, e_tm = dest_float tm in + let b = Num.num_of_int Arith_num.arith_base in + let m = Num.num_of_int Float_theory.min_exp in + let ( * ), (^), (-), (!) = ( */ ), ( **/ ), (-/), Arith_nat.raw_dest_hash in + let r = !n_tm * (b ^ (!e_tm - m)) in + if s = "T" then minus_num r else r;; + +(* Converts a float term to a floating-point number *) +let float_of_float_tm tm = + (Num.float_of_num o num_of_float_tm) tm;; + +(* Converts a floating-point number into a float term *) +let float_tm_of_float = + let split = + let b = float_of_int Arith_num.arith_base in + let log_b = log b in + let rec fix t k = + if t = 0.0 || t = infinity || t = nan then + t, k + else if t < 1.0 then + fix (t *. b) (k - 1) + else if t >= b then + fix (t /. b) (k + 1) + else + t, k in + fun f -> + if f = 0.0 then + false, f, 0 + else + let s, x = if f < 0.0 then true, -.f else false, f in + let e = -.floor (log x /. log_b) in + let k = -int_of_float e in + let t, n = fix (x *. (b ** e)) k in + s, t, n + and extract = + let b = float_of_int Arith_num.arith_base in + let nb = Num.num_of_int Arith_num.arith_base in + let rec step k f acc = + if k <= 0 then + acc + else + let d = int_of_float f in + step (k - 1) ((f -. float_of_int d) *. b) (Num.num_of_int d +/ (nb */ acc)) in + fun pp f -> + step pp f (Num.num_of_int 0) + in + fun pp f -> + let s, x, k = split f in + let n = extract pp x in + let n_tm = rand (Arith_num.mk_numeral_array n) in + let e_tm = rand (Arith_num.mk_small_numeral_array (min_exp + k - pp + 1)) in + let sign = if s then "T" else "F" in + make_float sign n_tm e_tm;; + +(* Creates a (not normalized) floating-point number from a given natural number. *) +(* Returns the theorem |- &n = float F n min_exp and the term n *) +let float_eq_th_of_num = + let conv_th = (Arith_nat.NUMERALS_TO_NUM o REWRITE_RULE[min_exp_def] o SPEC_ALL) FLOAT_OF_NUM in + fun n -> + let _ = if sign_num n < 0 then failwith "float_eq_th_of_num: negative number" in + let n_tm = rand (Arith_nat.mk_numeral_array n) in + INST[n_tm, n_var_num] conv_th, n_tm;; + +(* |- ##0 = &0, |- ##1 = &1, |- ##2 = &2, |- ##3 = &3, |- ##4 = &4 *) +let float0_eq = FLOAT_TO_NUM_CONV (mk_float 0 0) and + float1_eq = FLOAT_TO_NUM_CONV (mk_float 1 0) and + float2_eq = FLOAT_TO_NUM_CONV (mk_float 2 0) and + float3_eq = FLOAT_TO_NUM_CONV (mk_float 3 0) and + float4_eq = FLOAT_TO_NUM_CONV (mk_float 4 0);; + +(* |- D_k _0 = k for k = 1, 2, 3 *) +let num1_eq, num2_eq, num3_eq = + let conv = SYM o REWRITE_RULE[Arith_num.NUM_THM] o Arith_nat.NUMERAL_TO_NUM_CONV in + conv `1`, conv `2`, conv `3`;; + +(*********************) + +let float_F_const = `float_num F`;; + +let mk_float_small n = + let n_tm0 = mk_small_numeral n in + let n_th = Arith_nat.NUMERAL_TO_NUM_CONV n_tm0 in + let n_tm = rand(rand(concl n_th)) in + mk_comb(mk_comb(float_F_const, n_tm), min_exp_num_const);; + +(* Small float constants and intervals *) +let one_float = mk_float_small 1 and + two_float = mk_float_small 2 and + one_interval = mk_float_interval_small_num 1 and + two_interval = mk_float_interval_small_num 2;; +let neg_two_interval = float_interval_neg two_interval;; + +(* ----------------------------------------- *) +(* float_eq0 *) +(* ----------------------------------------- *) + +let FLOAT_EQ_0' = (GEN_REWRITE_RULE (RAND_CONV o RAND_CONV) [NUMERAL] o SPEC_ALL) + FLOAT_EQ_0;; + + +let float_eq0 f_tm = + let lhs, e_tm = dest_comb f_tm in + let lhs2, n_tm = dest_comb lhs in + let th0 = INST[rand lhs2, s_var_bool; n_tm, n_var_num; e_tm, e_var_num] FLOAT_EQ_0' in + let eq_th = Arith_nat.raw_eq0_hash_conv n_tm in + TRANS th0 eq_th;; + + + +(* ----------------------------------------- *) +(* float_interval_scale *) +(* ----------------------------------------- *) + + +let float_interval_scale pp c_tm th = + let c_th = mk_const_interval c_tm in + float_interval_mul pp c_th th;; + + +(* ----------------------------------------- *) +(* float_interval_lt0 *) +(* ----------------------------------------- *) + + +let FLOAT_INTERVAL_LT0' = (UNDISCH_ALL o prove) + (`interval_arith x (lo, hi) ==> (hi < &0 <=> T) ==> x < &0`, + REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; + +let float_interval_lt0 th = + let x_tm, bounds = dest_interval_arith (concl th) in + let lo_tm, hi_tm = dest_pair bounds in + let lt0_th = float_lt0 hi_tm in + let _ = ((rand o concl) lt0_th = t_const) || failwith "float_interval_lt0: &0 <= hi" in + let th0 = INST[x_tm, x_var_real; lo_tm, lo_var_real; hi_tm, hi_var_real] FLOAT_INTERVAL_LT0' in + (MY_PROVE_HYP th o MY_PROVE_HYP lt0_th) th0;; + + +(* ----------------------------------------- *) +(* float_pos *) +(* ----------------------------------------- *) + +let FLOAT_F_POS' = SPEC_ALL FLOAT_F_POS;; + +(* Returns &0 <= float F n e *) +let float_pos tm = + let _, n_tm, e_tm = dest_float tm in + INST[n_tm, n_var_num; e_tm, e_var_num] FLOAT_F_POS';; + + + +(* ----------------------------------------- *) +(* float_iabs *) +(* ----------------------------------------- *) + +let FLOAT_NEG_F' = RULE FLOAT_NEG_T;; +let FLOAT_NEG_T' = RULE FLOAT_NEG_F;; + +let float_neg tm = + let sign, n_tm, e_tm = dest_float tm in + if sign = "T" then + INST[n_tm, n_var_num; e_tm, e_var_num] FLOAT_NEG_T' + else + INST[n_tm, n_var_num; e_tm, e_var_num] FLOAT_NEG_F';; + + +let IABS' = RULE iabs;; + + +let float_iabs int_tm = + let lo_tm, hi_tm = dest_pair int_tm in + let neg_lo_th = float_neg lo_tm in + let max_th = SYM (float_max hi_tm ((rand o rator o concl) neg_lo_th)) in + let lhs, rhs = dest_comb (concl max_th) in + let th0 = SYM (EQ_MP (AP_TERM lhs (AP_TERM (rator rhs) neg_lo_th)) max_th) in + let th1 = INST[lo_tm, x_lo_var; hi_tm, x_hi_var] IABS' in + TRANS th1 th0;; + + +let FLOAT_IABS_FF = prove(`iabs (float_num F n1 e1, float_num F n2 e2) = float_num F n2 e2`, + REWRITE_TAC[iabs] THEN + MP_TAC (SPECL [`n1:num`; `e1:num`] FLOAT_F_POS) THEN + MP_TAC (SPECL [`n2:num`; `e2:num`] FLOAT_F_POS) THEN + REAL_ARITH_TAC);; + + +let FLOAT_IABS_TT = prove(`iabs (float_num T n1 e1, float_num T n2 e2) = float_num F n1 e1`, + REWRITE_TAC[iabs; GSYM FLOAT_NEG_F] THEN + MP_TAC (SPECL [`n1:num`; `e1:num`] FLOAT_F_POS) THEN + MP_TAC (SPECL [`n2:num`; `e2:num`] FLOAT_T_NEG) THEN + REAL_ARITH_TAC);; + + +(* ----------------------------------------- *) +(* interval_not_zero *) +(* ----------------------------------------- *) + +let INTERVAL_NOT_ZERO1' = (UNDISCH_ALL o prove) + (`(&0 < lo <=> T) ==> interval_not_zero (lo, hi)`, SIMP_TAC[interval_not_zero]);; +let INTERVAL_NOT_ZERO2' = (UNDISCH_ALL o prove) + (`(hi < &0 <=> T) ==> interval_not_zero (lo, hi)`, SIMP_TAC[interval_not_zero]);; + + +let check_interval_not_zero int_tm = + let lo, hi = dest_pair int_tm in + let inst = INST[lo, lo_var_real; hi, hi_var_real] in + let s1, _, _ = dest_float lo in + if s1 = "F" then + let gt_th = float_gt0 lo in + if (fst o dest_const o rand o concl) gt_th <> "T" then + failwith "check_interval_not_zero: &0 < lo <=> F" + else + (MY_PROVE_HYP gt_th o inst) INTERVAL_NOT_ZERO1' + else + let lt_th = float_lt0 hi in + if (fst o dest_const o rand o concl) lt_th <> "T" then + failwith "check_interval_not_zero: hi < &0 <=> F" + else + (MY_PROVE_HYP lt_th o inst) INTERVAL_NOT_ZERO2';; + + + +(* ----------------------------------------- *) +(* interval_pos *) +(* ----------------------------------------- *) + +let INTERVAL_POS' = (UNDISCH_ALL o prove) + (`(&0 < lo <=> T) ==> interval_pos (lo, hi:real)`, SIMP_TAC[interval_pos]);; + + +let check_interval_pos int_tm = + let lo, hi = dest_pair int_tm in + let gt_th = float_gt0 lo in + if (fst o dest_const o rand o concl) gt_th <> "T" then + failwith "check_interval_pos: &0 < lo <=> F" + else + (MY_PROVE_HYP gt_th o INST[lo, lo_var_real; hi, hi_var_real]) INTERVAL_POS';; + + +(* ----------------------------------------- *) +(* interval_neg *) +(* ----------------------------------------- *) + +let INTERVAL_NEG' = (UNDISCH_ALL o prove) + (`(hi < &0 <=> T) ==> interval_neg (lo:real, hi:real)`, SIMP_TAC[interval_neg]);; + + +let check_interval_neg int_tm = + let lo, hi = dest_pair int_tm in + let cmp_th = float_lt0 hi in + if (fst o dest_const o rand o concl) cmp_th <> "T" then + failwith "check_interval_neg: hi < &0 <=> F" + else + (MY_PROVE_HYP cmp_th o INST[lo, lo_var_real; hi, hi_var_real]) INTERVAL_NEG';; + +(* ----------------------------------------- *) +(* interval_gt *) +(* ----------------------------------------- *) + +let check_interval_gt = + let gt_th0 = (UNDISCH_ALL o prove)(`(y < lo <=> T) ==> interval_gt y (lo, hi:real)`, + REWRITE_TAC[interval_gt]) in + fun y_tm int_tm -> + let lo, hi = dest_pair int_tm in + let lt_th = float_lt y_tm lo in + if (fst o dest_const o rand o concl) lt_th <> "T" then + failwith "check_interval_gt: y < lo <=> F" + else + (MY_PROVE_HYP lt_th o + INST[lo, lo_var_real; hi, hi_var_real; y_tm, y_var_real]) gt_th0;; + + +(* ----------------------------------------- *) +(* check_interval_iabs *) +(* ----------------------------------------- *) + + +(* proves |- iabs int < rhs <=> T *) +let check_interval_iabs int_tm rhs_tm = + let iabs_eq = float_iabs int_tm in + let lt_th = float_lt (rand (concl iabs_eq)) rhs_tm in + if (fst o dest_const o rand o concl) lt_th <> "T" then + failwith "check_interval_iabs: iabs < rhs <=> F" + else + let th0 = AP_THM (AP_TERM lt_op_real iabs_eq) rhs_tm in + TRANS th0 lt_th;; + + + +(* ----------------------------------------- *) +(* inv *) +(* ----------------------------------------- *) + +let float_inv_hi = + let div_lemma = (RULE o prove) + (`&1 / x <= hi ==> inv x <= hi`, + REWRITE_TAC[real_div; REAL_MUL_LID]) in + fun pp tm -> + let div_th = float_div_hi pp one_float tm in + let hi_tm = rand (concl div_th) in + let th0 = INST[tm, x_var_real; hi_tm, hi_var_real] div_lemma in + MY_PROVE_HYP div_th th0;; + +let float_inv_lo = + let div_lemma = (RULE o prove) + (`lo <= &1 / x ==> lo <= inv x`, + REWRITE_TAC[real_div; REAL_MUL_LID]) in + fun pp tm -> + let div_th = float_div_lo pp one_float tm in + let lo_tm = rand (rator (concl div_th)) in + let th0 = INST[tm, x_var_real; lo_tm, lo_var_real] div_lemma in + MY_PROVE_HYP div_th th0;; + +let float_interval_inv = + let INV_EQ_DIV_LEMMA = prove(`&1 / x = inv x`, REWRITE_TAC[real_div; REAL_MUL_LID]) in + fun pp th -> + let x_tm = (rand o rator o concl) th in + let div_th = INST[x_tm, x_var_real] INV_EQ_DIV_LEMMA in + let th0 = float_interval_div pp one_interval th in + let lhs, rhs = dest_comb (concl th0) in + let lhs2, rhs2 = dest_comb lhs in + EQ_MP (AP_THM (AP_TERM lhs2 div_th) rhs) th0;; + + +(* Explicit representation of inv(&2) *) +let float_inv2_th = + let one_float_eq_one = FLOAT_TO_NUM_CONV one_float in + let inv2_eq_lemma = prove(`interval_arith (&2 * x) (&1, &1) ==> inv (&2) = x`, + REWRITE_TAC[interval_arith] THEN CONV_TAC REAL_FIELD) in + let half_tm = (fst o dest_pair o rand o concl) (float_interval_inv 1 two_interval) in + let half_interval = mk_const_interval half_tm in + let mul_th = REWRITE_RULE[one_float_eq_one] (float_interval_mul 2 two_interval half_interval) in + MATCH_MP inv2_eq_lemma mul_th;; + +let float_inv2 = rand (concl float_inv2_th);; +let inv2_interval = mk_const_interval float_inv2;; + + + +(*****************************************) +(* bounded_on_int *) + +let norm_derivative d_th eq_th = + let lhs, rhs = (dest_eq o concl) d_th in + let lhs2, rhs2 = dest_comb lhs in + let th0 = AP_THM (AP_TERM (rator lhs2) eq_th) rhs2 in + TRANS (SYM th0) d_th;; + +let norm_diff d_th eq_th = + let lhs, rhs = (dest_comb o concl) d_th in + let th0 = AP_THM (AP_TERM (rator lhs) eq_th) rhs in + EQ_MP th0 d_th;; + +let norm_interval int_th eq_th = + let lhs, rhs = (dest_comb o concl) int_th in + let th0 = AP_THM (AP_TERM (rator lhs) eq_th) rhs in + EQ_MP (SYM th0) int_th;; + +let norm_second_derivative th eq_th = + let lhs, dd_bounds = dest_comb (concl th) in + let lhs2, int_tm = dest_comb lhs in + let th0 = AP_THM (AP_THM (AP_TERM (rator lhs2) eq_th) int_tm) dd_bounds in + EQ_MP th0 th;; + +let norm_lin_approx th eq_th = + let lhs, df_bounds = dest_comb (concl th) in + let lhs2, f_bounds = dest_comb lhs in + let lhs3, x_tm = dest_comb lhs2 in + let th0 = AP_THM (AP_THM (AP_THM (AP_TERM (rator lhs3) eq_th) x_tm) f_bounds) df_bounds in + EQ_MP th0 th;; + + + +let BOUNDED_ON_INT = (UNDISCH_ALL o prove)(`(!x. interval_arith x int ==> + interval_arith (f x) f_bounds) + ==> bounded_on_int f int f_bounds`, + REWRITE_TAC[bounded_on_int; interval_arith]);; + + +let BOUNDED_ON_INT_DEST = (UNDISCH_ALL o prove)(`bounded_on_int f int f_bounds ==> + (!x. interval_arith x int ==> interval_arith (f x) f_bounds)`, + REWRITE_TAC[bounded_on_int; interval_arith]);; + + +(* Given a theorem (interval_arith x int |- interval_arith (f x) f_bounds), yields + |- bounded_on_int (\x. f x) int f_bounds *) +let mk_bounded_on_int th = + let int_tm = (rand o hd o hyp) th in + let lhs, f_bounds_tm = dest_comb (concl th) in + let lhs2, rhs2 = dest_comb lhs in + let f_tm = mk_abs (x_var_real, rhs2) in + let b_th0 = (SYM o BETA_CONV) (mk_comb (f_tm , x_var_real)) in + let b_th1 = AP_THM (AP_TERM lhs2 b_th0) f_bounds_tm in + let th2 = EQ_MP b_th1 th in + let th3 = DISCH_ALL th2 in + let th4 = GEN x_var_real th3 in + let th_int = INST[int_tm, int_var; f_bounds_tm, f_bounds_var; + f_tm, f_var_fun] BOUNDED_ON_INT in + MY_PROVE_HYP th4 th_int;; + + +let dest_bounded_on_int th = + let lhs, f_bounds = dest_comb (concl th) in + let lhs2, int_tm = dest_comb lhs in + let f_tm = rand lhs2 in + let th0 = INST[f_tm, f_var_fun; int_tm, int_var; f_bounds, f_bounds_var] BOUNDED_ON_INT_DEST in + let th1 = UNDISCH_ALL (SPEC x_var_real (MY_PROVE_HYP th th0)) in + if is_abs f_tm then + let f_tm = (rand o rator o concl) th1 in + let eq_th = BETA_CONV f_tm in + norm_interval th1 (SYM eq_th) + else + th1;; + + +let dest_bounded_on_int_raw th = + let lhs, f_bounds = dest_comb (concl th) in + let lhs2, int_tm = dest_comb lhs in + let f_tm = rand lhs2 in + let th0 = INST[f_tm, f_var_fun; int_tm, int_var; f_bounds, f_bounds_var] BOUNDED_ON_INT_DEST in + UNDISCH_ALL (SPEC x_var_real (MY_PROVE_HYP th th0));; + + +(***********************************) +(* bounded_on_int arithmetic *) + +let bounded_on_int_scale pp c_tm th = + let i_th = dest_bounded_on_int th in + let th0 = float_interval_scale pp c_tm i_th in + mk_bounded_on_int th0;; + + +let bounded_on_int_mul_int pp int_th th = + let i_th = dest_bounded_on_int th in + let th0 = float_interval_mul pp int_th i_th in + mk_bounded_on_int th0;; + + +let bounded_on_int_neg th1 = + let i_th = dest_bounded_on_int th1 in + let th0 = float_interval_neg i_th in + mk_bounded_on_int th0;; + + +let bounded_on_int_add pp th1 th2 = + let i_th1, i_th2 = dest_bounded_on_int th1, dest_bounded_on_int th2 in + let th0 = float_interval_add pp i_th1 i_th2 in + mk_bounded_on_int th0;; + + +let bounded_on_int_sub pp th1 th2 = + let i_th1, i_th2 = dest_bounded_on_int th1, dest_bounded_on_int th2 in + let th0 = float_interval_sub pp i_th1 i_th2 in + mk_bounded_on_int th0;; + + +let bounded_on_int_mul pp th1 th2 = + let i_th1, i_th2 = dest_bounded_on_int th1, dest_bounded_on_int th2 in + let th0 = float_interval_mul pp i_th1 i_th2 in + mk_bounded_on_int th0;; + + +let bounded_on_int_mul_raw pp th1 th2 = + let i_th1, i_th2 = dest_bounded_on_int_raw th1, dest_bounded_on_int_raw th2 in + let th0 = float_interval_mul pp i_th1 i_th2 in + mk_bounded_on_int th0;; + + + +let bounded_on_int_div pp th1 th2 = + let i_th1, i_th2 = dest_bounded_on_int th1, dest_bounded_on_int th2 in + let th0 = float_interval_div pp i_th1 i_th2 in + mk_bounded_on_int th0;; + + +(************************************) +let ADD_INEQ_HI = (RULE o REAL_ARITH) `x1 <= y1 /\ x2 <= y2 /\ y1 + y2 <= y ==> x1 + x2 <= y`;; +let ADD_INEQ_LO = (RULE o REAL_ARITH) `x1 <= y1 /\ x2 <= y2 /\ x <= x1 + x2 ==> x <= y1 + y2`;; +let SUB_INEQ_HI = (RULE o REAL_ARITH) `x1 <= y1 /\ y2 <= x2 /\ y1 - y2 <= y ==> x1 - x2 <= y`;; +let SUB_INEQ_LO = (RULE o REAL_ARITH) `x1 <= y1 /\ y2 <= x2 /\ x <= x1 - x2 ==> x <= y1 - y2`;; +let MUL_INEQ_HI = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o prove) + (`&0 <= x1 /\ &0 <= x2 /\ x1 <= y1 /\ x2 <= y2 /\ y1 * y2 <= y ==> x1 * x2 <= y`, + DISCH_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN + EXISTS_TAC `y1 * y2` THEN ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_REWRITE_TAC[]);; + +let MUL_INEQ_POS_CONST_HI = (UNDISCH_ALL o prove) + (`(&0 <= x <=> T) ==> y1 <= y2 ==> x * y2 <= z ==> x * y1 <= z`, + REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `x * y2` THEN + ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_LMUL THEN ASM_REWRITE_TAC[]);; + + +let mk_refl_ineq = + let REAL_LE_REFL' = RULE REAL_LE_REFL in + fun tm -> INST[tm, x_var_real] REAL_LE_REFL';; + +let dest_le_op ineq = + let lhs, y_tm = dest_comb ineq in + (rand lhs, y_tm);; + + +let mul_ineq_pos_const_hi pp c_tm ineq = + let y1_tm, y2_tm = dest_le_op (concl ineq) in + let ge0_th = float_ge0 c_tm in + let mul_hi_th = float_mul_hi pp c_tm y2_tm in + let z_tm = (rand o concl) mul_hi_th in + (MY_PROVE_HYP ge0_th o MY_PROVE_HYP ineq o MY_PROVE_HYP mul_hi_th o + INST[c_tm, x_var_real; y1_tm, y1_var_real; y2_tm, y2_var_real; z_tm, z_var_real]) + MUL_INEQ_POS_CONST_HI;; + + +let mul_ineq_hi pp ineq1 ineq2 = + let x1_tm, y1_tm = dest_le_op (concl ineq1) in + let x2_tm, y2_tm = dest_le_op (concl ineq2) in + let x1_pos, x2_pos = float_pos x1_tm, float_pos x2_tm in + let rhs_mul = float_mul_hi pp y1_tm y2_tm in + let y_tm = (rand o concl) rhs_mul in + let th0 = INST[x1_tm, x1_var_real; y1_tm, y1_var_real; + x2_tm, x2_var_real; y2_tm, y2_var_real; + y_tm, y_var_real] MUL_INEQ_HI in + (MY_PROVE_HYP x1_pos o MY_PROVE_HYP x2_pos o MY_PROVE_HYP ineq1 o + MY_PROVE_HYP ineq2 o MY_PROVE_HYP rhs_mul) th0;; + + + +let sub_ineq_hi pp ineq1 ineq2 = + let x1_tm, y1_tm = dest_le_op (concl ineq1) in + let y2_tm, x2_tm = dest_le_op (concl ineq2) in + let rhs_sub = float_sub_hi pp y1_tm y2_tm in + let y_tm = (rand o concl) rhs_sub in + let th0 = INST[x1_tm, x1_var_real; y1_tm, y1_var_real; + x2_tm, x2_var_real; y2_tm, y2_var_real; + y_tm, y_var_real] SUB_INEQ_HI in + MY_PROVE_HYP ineq1 (MY_PROVE_HYP ineq2 (MY_PROVE_HYP rhs_sub th0));; + + +let sub_ineq_lo pp ineq1 ineq2 = + let x1_tm, y1_tm = dest_le_op (concl ineq1) in + let y2_tm, x2_tm = dest_le_op (concl ineq2) in + let lhs_sub = float_sub_lo pp x1_tm x2_tm in + let x_tm = (lhand o concl) lhs_sub in + let th0 = INST[x1_tm, x1_var_real; y1_tm, y1_var_real; + x2_tm, x2_var_real; y2_tm, y2_var_real; + x_tm, x_var_real] SUB_INEQ_LO in + MY_PROVE_HYP ineq1 (MY_PROVE_HYP ineq2 (MY_PROVE_HYP lhs_sub th0));; + + +let add_ineq_hi pp ineq1 ineq2 = + let x1_tm, y1_tm = dest_le_op (concl ineq1) in + let x2_tm, y2_tm = dest_le_op (concl ineq2) in + let rhs_sum = float_add_hi pp y1_tm y2_tm in + let y_tm = (rand o concl) rhs_sum in + let th0 = INST[x1_tm, x1_var_real; y1_tm, y1_var_real; + x2_tm, x2_var_real; y2_tm, y2_var_real; + y_tm, y_var_real] ADD_INEQ_HI in + MY_PROVE_HYP ineq1 (MY_PROVE_HYP ineq2 (MY_PROVE_HYP rhs_sum th0));; + + +let add_ineq_lo pp ineq1 ineq2 = + let x1_tm, y1_tm = dest_le_op (concl ineq1) in + let x2_tm, y2_tm = dest_le_op (concl ineq2) in + let lhs_sum = float_add_lo pp x1_tm x2_tm in + let x_tm = (lhand o concl) lhs_sum in + let th0 = INST[x1_tm, x1_var_real; y1_tm, y1_var_real; + x2_tm, x2_var_real; y2_tm, y2_var_real; + x_tm, x_var_real] ADD_INEQ_LO in + MY_PROVE_HYP ineq1 (MY_PROVE_HYP ineq2 (MY_PROVE_HYP lhs_sum th0));; + + +end;; diff --git a/Formal_ineqs/arith/num_exp_theory.hl b/Formal_ineqs/arith/num_exp_theory.hl index c058b112..6e388cc2 100644 --- a/Formal_ineqs/arith/num_exp_theory.hl +++ b/Formal_ineqs/arith/num_exp_theory.hl @@ -1,251 +1,256 @@ -(* =========================================================== *) -(* Exponential representation of natural numbers *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -(* Dependencies *) -needs "arith/arith_nat.hl";; - - -module Num_exp_theory = struct - -let base_const = mk_small_numeral Arith_num.arith_base;; - -(* num_exp definition *) -let num_exp_tm = mk_eq (`(num_exp:num->num->num) n e`, - mk_binop `( * ):num->num->num` `n:num` (mk_binop `EXP` base_const `e:num`));; -(* let num_exp = new_definition `num_exp n e = n * 2 EXP e`;; *) -let num_exp = new_definition num_exp_tm;; - -(**********************************) -(* Theorems *) - -let NUM_EXP_EXP = prove(`!n e1 e2. num_exp (num_exp n e1) e2 = num_exp n (e1 + e2)`, - REPEAT GEN_TAC THEN - REWRITE_TAC[num_exp; EXP_ADD] THEN - ARITH_TAC);; - -let NUM_EXP_SUM = prove(`!n e1 e2. num_exp n (e1 + e2) = num_exp n e1 * num_exp 1 e2`, - REPEAT GEN_TAC THEN - REWRITE_TAC[num_exp; EXP_ADD] THEN - ARITH_TAC);; - -let NUM_EXP_SUM1 = prove(`!n e1 e2. num_exp n (e1 + e2) = num_exp 1 e1 * num_exp n e2`, - REPEAT GEN_TAC THEN REWRITE_TAC[num_exp; EXP_ADD] THEN ARITH_TAC);; - -let NUM_EXP_0 = prove(`!n. n = num_exp n 0`, - GEN_TAC THEN REWRITE_TAC[num_exp; EXP; MULT_CLAUSES]);; - -let NUM_EXP_LE = prove(`!m n e. m <= n ==> num_exp m e <= num_exp n e`, - SIMP_TAC[num_exp; LE_MULT_RCANCEL]);; - -let NUM_EXP_LT = prove(`!m n e. m < n ==> num_exp m e < num_exp n e`, - SIMP_TAC[num_exp; LT_MULT_RCANCEL; EXP_EQ_0] THEN - ARITH_TAC);; - -let NUM_EXP_EQ_0 = prove(`!n e. num_exp n e = 0 <=> n = 0`, - REPEAT STRIP_TAC THEN - ASM_REWRITE_TAC[num_exp; MULT_EQ_0; EXP_EQ_0] THEN - ARITH_TAC);; - -let NUM_EXP_MUL = prove(`!n1 e1 n2 e2. num_exp n1 e1 * num_exp n2 e2 = num_exp (n1 * n2) (e1 + e2)`, - REWRITE_TAC[num_exp; EXP_ADD] THEN ARITH_TAC);; - -let NUM_EXP_ADD = prove(`!n1 e1 n2 e2. e1 <= e2 ==> - num_exp n1 e1 + num_exp n2 e2 = num_exp (n1 + num_exp n2 (e2 - e1)) e1`, - REPEAT STRIP_TAC THEN - REWRITE_TAC[num_exp] THEN - REWRITE_TAC[ARITH_RULE `(a + b * c) * d = a * d + b * (c * d):num`] THEN - REWRITE_TAC[GSYM EXP_ADD] THEN - ASM_SIMP_TAC[ARITH_RULE `e1 <= e2 ==> e2 - e1 + e1 = e2:num`]);; - -let NUM_EXP_SUB2 = prove(`!n1 e1 n2 e2 r. e1 <= e2 /\ e2 - e1 = r ==> - num_exp n1 e1 - num_exp n2 e2 = num_exp (n1 - num_exp n2 r) e1`, - REPEAT STRIP_TAC THEN - POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN - REWRITE_TAC[num_exp] THEN - MP_TAC (ARITH_RULE `e1 <= e2 ==> e2 = (e2 - e1) + e1:num`) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN (fun th -> GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [th]) THEN - REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN - REWRITE_TAC[GSYM RIGHT_SUB_DISTRIB]);; - -let NUM_EXP_SUB1 = prove(`!n1 e1 n2 e2 r. e2 <= e1 /\ e1 - e2 = r ==> - num_exp n1 e1 - num_exp n2 e2 = num_exp (num_exp n1 r - n2) e2`, - REPEAT STRIP_TAC THEN - POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN - REWRITE_TAC[num_exp] THEN - MP_TAC (ARITH_RULE `e2 <= e1 ==> e1 = (e1 - e2) + e2:num`) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN (fun th -> GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [th]) THEN - REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN - REWRITE_TAC[GSYM RIGHT_SUB_DISTRIB]);; - -(* NUM_EXP_LE *) - -let NUM_EXP_LE1 = prove(`!n1 e1 n2 e2 r. e2 <= e1 /\ e1 - e2 = r /\ n2 <= num_exp n1 r - ==> num_exp n2 e2 <= num_exp n1 e1`, - REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN - STRIP_TAC THEN - POP_ASSUM MP_TAC THEN POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN DISCH_TAC THEN - MP_TAC (ARITH_RULE `e2 <= e1 ==> e1 = (e1 - e2) + e2:num`) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN(fun th -> ONCE_REWRITE_TAC[th]) THEN - REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN - ASM_REWRITE_TAC[LE_MULT_RCANCEL]);; - -let NUM_EXP_LE2 = prove(`!n1 e1 n2 e2 r. e1 <= e2 /\ e2 - e1 = r /\ num_exp n2 r <= n1 - ==> num_exp n2 e2 <= num_exp n1 e1`, - REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN - STRIP_TAC THEN - POP_ASSUM MP_TAC THEN POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN DISCH_TAC THEN - MP_TAC (ARITH_RULE `e1 <= e2 ==> e2 = (e2 - e1) + e1:num`) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN(fun th -> ONCE_REWRITE_TAC[th]) THEN - REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN - ASM_REWRITE_TAC[LE_MULT_RCANCEL]);; - -let NUM_EXP_LE1_EQ = prove(`!n1 e1 n2 e2 r x. e2 <= e1 /\ e1 - e2 = r /\ num_exp n1 r = x ==> - (num_exp n1 e1 <= num_exp n2 e2 <=> x <= n2)`, - REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN - STRIP_TAC THEN - REPLICATE_TAC 2 (POP_ASSUM (fun th -> REWRITE_TAC[SYM th])) THEN - MP_TAC (ARITH_RULE `e2 <= e1 ==> e1 = (e1 - e2) + e2:num`) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN (fun th -> GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV) [th]) THEN - REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN - ASM_REWRITE_TAC[LE_MULT_RCANCEL; EXP_EQ_0; ARITH_EQ]);; - -let NUM_EXP_LE2_EQ = prove(`!n1 e1 n2 e2 r x. e1 <= e2 /\ e2 - e1 = r /\ num_exp n2 r = x ==> - (num_exp n1 e1 <= num_exp n2 e2 <=> n1 <= x)`, - REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN - STRIP_TAC THEN - REPLICATE_TAC 2 (POP_ASSUM (fun th -> REWRITE_TAC[SYM th])) THEN - MP_TAC (ARITH_RULE `e1 <= e2 ==> e2 = (e2 - e1) + e1:num`) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN (fun th -> GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV) [th]) THEN - REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN - ASM_REWRITE_TAC[LE_MULT_RCANCEL; EXP_EQ_0; ARITH_EQ]);; - -(* NUM_EXP_LT *) - -let NUM_EXP_LT1 = prove(`!n1 e1 n2 e2 r. e2 <= e1 /\ e1 - e2 = r /\ n2 < num_exp n1 r - ==> num_exp n2 e2 < num_exp n1 e1`, - REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN - STRIP_TAC THEN - POP_ASSUM MP_TAC THEN POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN DISCH_TAC THEN - MP_TAC (ARITH_RULE `e2 <= e1 ==> e1 = (e1 - e2) + e2:num`) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN (fun th -> ONCE_REWRITE_TAC[th]) THEN - REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN - ASM_REWRITE_TAC[LT_MULT_RCANCEL; EXP_EQ_0; ARITH_EQ]);; - -let NUM_EXP_LT2 = prove(`!n1 e1 n2 e2 r. e1 <= e2 /\ e2 - e1 = r /\ num_exp n2 r < n1 - ==> num_exp n2 e2 < num_exp n1 e1`, - REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN - STRIP_TAC THEN - POP_ASSUM MP_TAC THEN POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN DISCH_TAC THEN - MP_TAC (ARITH_RULE `e1 <= e2 ==> e2 = (e2 - e1) + e1:num`) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN (fun th -> ONCE_REWRITE_TAC[th]) THEN - REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN - ASM_REWRITE_TAC[LT_MULT_RCANCEL; EXP_EQ_0; ARITH_EQ]);; - -let NUM_EXP_LT1_EQ = prove(`!n1 e1 n2 e2 r x. e2 <= e1 /\ e1 - e2 = r /\ num_exp n1 r = x ==> - (num_exp n1 e1 < num_exp n2 e2 <=> x < n2)`, - REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN - STRIP_TAC THEN - REPLICATE_TAC 2 (POP_ASSUM (fun th -> REWRITE_TAC[SYM th])) THEN - MP_TAC (ARITH_RULE `e2 <= e1 ==> e1 = (e1 - e2) + e2:num`) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN (fun th -> GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV) [th]) THEN - REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN - ASM_REWRITE_TAC[LT_MULT_RCANCEL; EXP_EQ_0; ARITH_EQ]);; - -let NUM_EXP_LT2_EQ = prove(`!n1 e1 n2 e2 r x. e1 <= e2 /\ e2 - e1 = r /\ num_exp n2 r = x ==> - (num_exp n1 e1 < num_exp n2 e2 <=> n1 < x)`, - REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN - STRIP_TAC THEN - REPLICATE_TAC 2 (POP_ASSUM (fun th -> REWRITE_TAC[SYM th])) THEN - MP_TAC (ARITH_RULE `e1 <= e2 ==> e2 = (e2 - e1) + e1:num`) THEN - ASM_REWRITE_TAC[] THEN - DISCH_THEN (fun th -> GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV) [th]) THEN - REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN - ASM_REWRITE_TAC[LT_MULT_RCANCEL; EXP_EQ_0; ARITH_EQ]);; - -(* NUM_EXP_DIV *) - -let mul_op_num = `( * ):num->num->num`;; - -let NUM_EXP_DIV1 = prove(`~(n2 = 0) /\ e2 <= e1 ==> - num_exp n1 e1 DIV num_exp n2 e2 = num_exp n1 (e1 - e2) DIV n2`, - STRIP_TAC THEN - (*`num_exp n1 e1 = 16 EXP e2 * num_exp n1 (e1 - e2)` MP_TAC THENL*) - SUBGOAL_THEN (mk_eq(`num_exp n1 e1`, mk_binop mul_op_num (mk_binop `EXP` base_const `e2:num`) `num_exp n1 (e1 - e2)`)) MP_TAC THENL - [ - REWRITE_TAC[num_exp] THEN - ONCE_REWRITE_TAC[ARITH_RULE `a * b * c = b * (a * c:num)`] THEN - REWRITE_TAC[GSYM EXP_ADD] THEN - ASM_SIMP_TAC[ARITH_RULE `e2 <= e1 ==> e2 + e1 - e2 = e1:num`]; - ALL_TAC - ] THEN - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - SUBGOAL_THEN (mk_eq(`num_exp n2 e2`, mk_binop mul_op_num (mk_binop `EXP` base_const `e2:num`) `n2:num`)) MP_TAC THENL - [ - REWRITE_TAC[num_exp; MULT_AC]; - ALL_TAC - ] THEN - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - MATCH_MP_TAC DIV_MULT2 THEN - ASM_REWRITE_TAC[MULT_EQ_0; DE_MORGAN_THM; EXP_EQ_0] THEN - ARITH_TAC);; - -let NUM_EXP_DIV2 = prove(`~(n2 = 0) /\ e1 <= e2 ==> - num_exp n1 e1 DIV num_exp n2 e2 = n1 DIV num_exp n2 (e2 - e1)`, - STRIP_TAC THEN - (*`num_exp n2 e2 = 16 EXP e1 * num_exp n2 (e2 - e1)` MP_TAC THENL*) - SUBGOAL_THEN (mk_eq(`num_exp n2 e2`, mk_binop mul_op_num (mk_binop `EXP` base_const `e1:num`) `num_exp n2 (e2 - e1)`)) MP_TAC THENL - [ - REWRITE_TAC[num_exp] THEN - ONCE_REWRITE_TAC[ARITH_RULE `a * b * c = b * (a * c:num)`] THEN - REWRITE_TAC[GSYM EXP_ADD] THEN - ASM_SIMP_TAC[ARITH_RULE `e1 <= e2 ==> e1 + e2 - e1 = e2:num`]; - ALL_TAC - ] THEN - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - SUBGOAL_THEN (mk_eq(`num_exp n1 e1`, mk_binop mul_op_num (mk_binop `EXP` base_const `e1:num`) `n1:num`)) MP_TAC THENL - [ - REWRITE_TAC[num_exp; MULT_AC]; - ALL_TAC - ] THEN - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - MATCH_MP_TAC DIV_MULT2 THEN - ASM_REWRITE_TAC[num_exp; MULT_EQ_0; DE_MORGAN_THM; EXP_EQ_0] THEN - ARITH_TAC);; - - -let EXP_INV_lemma = prove(`!n e1 e2. ~(n = 0) /\ e2 <= e1 ==> &(n EXP (e1 - e2)) = - &(n EXP e1) * inv(&(n EXP e2))`, - REPEAT STRIP_TAC THEN - REWRITE_TAC[GSYM REAL_OF_NUM_POW] THEN - MP_TAC (SPECL [`&n`; `e2:num`; `e1:num`] REAL_POW_SUB) THEN - ASM_REWRITE_TAC[REAL_OF_NUM_EQ; real_div]);; - -let NUM_EXP_SUB_lemma = prove(`!n e1 e2. e2 <= e1 ==> &(num_exp n (e1 - e2)) = - &(num_exp n e1) * inv(&(num_exp 1 e2))`, - REPEAT STRIP_TAC THEN - REWRITE_TAC[num_exp] THEN - REWRITE_TAC[GSYM REAL_OF_NUM_MUL] THEN - MP_TAC (SPECL [base_const; `e1:num`; `e2:num`] EXP_INV_lemma) THEN - ANTS_TAC THENL - [ - ASM_REWRITE_TAC[] THEN ARITH_TAC; - ALL_TAC - ] THEN - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - REAL_ARITH_TAC);; - - - -end;; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Exponential representation of natural numbers *) +(* -------------------------------------------------------------------------- *) + +(* Dependencies *) +needs "Formal_ineqs/arith/arith_nat.hl";; + +module Num_exp_theory = struct + +let base_const = mk_small_numeral Arith_num.arith_base;; + +(* num_exp definition *) +let num_exp_tm = mk_eq (`(num_exp:num->num->num) n e`, + mk_binop `( * ):num->num->num` `n:num` (mk_binop `EXP` base_const `e:num`));; +(* let num_exp = new_definition `num_exp n e = n * 2 EXP e`;; *) +let num_exp = new_definition num_exp_tm;; + +(**********************************) +(* Theorems *) + +let NUM_EXP_EXP = prove(`!n e1 e2. num_exp (num_exp n e1) e2 = num_exp n (e1 + e2)`, + REPEAT GEN_TAC THEN + REWRITE_TAC[num_exp; EXP_ADD] THEN + ARITH_TAC);; + +let NUM_EXP_SUM = prove(`!n e1 e2. num_exp n (e1 + e2) = num_exp n e1 * num_exp 1 e2`, + REPEAT GEN_TAC THEN + REWRITE_TAC[num_exp; EXP_ADD] THEN + ARITH_TAC);; + +let NUM_EXP_SUM1 = prove(`!n e1 e2. num_exp n (e1 + e2) = num_exp 1 e1 * num_exp n e2`, + REPEAT GEN_TAC THEN REWRITE_TAC[num_exp; EXP_ADD] THEN ARITH_TAC);; + +let NUM_EXP_0 = prove(`!n. n = num_exp n 0`, + GEN_TAC THEN REWRITE_TAC[num_exp; EXP; MULT_CLAUSES]);; + +let NUM_EXP_LE = prove(`!m n e. m <= n ==> num_exp m e <= num_exp n e`, + SIMP_TAC[num_exp; LE_MULT_RCANCEL]);; + +let NUM_EXP_LT = prove(`!m n e. m < n ==> num_exp m e < num_exp n e`, + SIMP_TAC[num_exp; LT_MULT_RCANCEL; EXP_EQ_0] THEN + ARITH_TAC);; + +let NUM_EXP_EQ_0 = prove(`!n e. num_exp n e = 0 <=> n = 0`, + REPEAT STRIP_TAC THEN + ASM_REWRITE_TAC[num_exp; MULT_EQ_0; EXP_EQ_0] THEN + ARITH_TAC);; + +let NUM_EXP_MUL = prove(`!n1 e1 n2 e2. num_exp n1 e1 * num_exp n2 e2 = num_exp (n1 * n2) (e1 + e2)`, + REWRITE_TAC[num_exp; EXP_ADD] THEN ARITH_TAC);; + +let NUM_EXP_ADD = prove(`!n1 e1 n2 e2. e1 <= e2 ==> + num_exp n1 e1 + num_exp n2 e2 = num_exp (n1 + num_exp n2 (e2 - e1)) e1`, + REPEAT STRIP_TAC THEN + REWRITE_TAC[num_exp] THEN + REWRITE_TAC[ARITH_RULE `(a + b * c) * d = a * d + b * (c * d):num`] THEN + REWRITE_TAC[GSYM EXP_ADD] THEN + ASM_SIMP_TAC[ARITH_RULE `e1 <= e2 ==> e2 - e1 + e1 = e2:num`]);; + +let NUM_EXP_SUB2 = prove(`!n1 e1 n2 e2 r. e1 <= e2 /\ e2 - e1 = r ==> + num_exp n1 e1 - num_exp n2 e2 = num_exp (n1 - num_exp n2 r) e1`, + REPEAT STRIP_TAC THEN + POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN + REWRITE_TAC[num_exp] THEN + MP_TAC (ARITH_RULE `e1 <= e2 ==> e2 = (e2 - e1) + e1:num`) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN (fun th -> GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [th]) THEN + REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN + REWRITE_TAC[GSYM RIGHT_SUB_DISTRIB]);; + +let NUM_EXP_SUB1 = prove(`!n1 e1 n2 e2 r. e2 <= e1 /\ e1 - e2 = r ==> + num_exp n1 e1 - num_exp n2 e2 = num_exp (num_exp n1 r - n2) e2`, + REPEAT STRIP_TAC THEN + POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN + REWRITE_TAC[num_exp] THEN + MP_TAC (ARITH_RULE `e2 <= e1 ==> e1 = (e1 - e2) + e2:num`) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN (fun th -> GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [th]) THEN + REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN + REWRITE_TAC[GSYM RIGHT_SUB_DISTRIB]);; + +(* NUM_EXP_LE *) + +let NUM_EXP_LE1 = prove(`!n1 e1 n2 e2 r. e2 <= e1 /\ e1 - e2 = r /\ n2 <= num_exp n1 r + ==> num_exp n2 e2 <= num_exp n1 e1`, + REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN + STRIP_TAC THEN + POP_ASSUM MP_TAC THEN POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN DISCH_TAC THEN + MP_TAC (ARITH_RULE `e2 <= e1 ==> e1 = (e1 - e2) + e2:num`) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN(fun th -> ONCE_REWRITE_TAC[th]) THEN + REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN + ASM_REWRITE_TAC[LE_MULT_RCANCEL]);; + +let NUM_EXP_LE2 = prove(`!n1 e1 n2 e2 r. e1 <= e2 /\ e2 - e1 = r /\ num_exp n2 r <= n1 + ==> num_exp n2 e2 <= num_exp n1 e1`, + REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN + STRIP_TAC THEN + POP_ASSUM MP_TAC THEN POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN DISCH_TAC THEN + MP_TAC (ARITH_RULE `e1 <= e2 ==> e2 = (e2 - e1) + e1:num`) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN(fun th -> ONCE_REWRITE_TAC[th]) THEN + REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN + ASM_REWRITE_TAC[LE_MULT_RCANCEL]);; + +let NUM_EXP_LE1_EQ = prove(`!n1 e1 n2 e2 r x. e2 <= e1 /\ e1 - e2 = r /\ num_exp n1 r = x ==> + (num_exp n1 e1 <= num_exp n2 e2 <=> x <= n2)`, + REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN + STRIP_TAC THEN + REPLICATE_TAC 2 (POP_ASSUM (fun th -> REWRITE_TAC[SYM th])) THEN + MP_TAC (ARITH_RULE `e2 <= e1 ==> e1 = (e1 - e2) + e2:num`) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN (fun th -> GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV) [th]) THEN + REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN + ASM_REWRITE_TAC[LE_MULT_RCANCEL; EXP_EQ_0; ARITH_EQ]);; + +let NUM_EXP_LE2_EQ = prove(`!n1 e1 n2 e2 r x. e1 <= e2 /\ e2 - e1 = r /\ num_exp n2 r = x ==> + (num_exp n1 e1 <= num_exp n2 e2 <=> n1 <= x)`, + REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN + STRIP_TAC THEN + REPLICATE_TAC 2 (POP_ASSUM (fun th -> REWRITE_TAC[SYM th])) THEN + MP_TAC (ARITH_RULE `e1 <= e2 ==> e2 = (e2 - e1) + e1:num`) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN (fun th -> GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV) [th]) THEN + REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN + ASM_REWRITE_TAC[LE_MULT_RCANCEL; EXP_EQ_0; ARITH_EQ]);; + +(* NUM_EXP_LT *) + +let NUM_EXP_LT1 = prove(`!n1 e1 n2 e2 r. e2 <= e1 /\ e1 - e2 = r /\ n2 < num_exp n1 r + ==> num_exp n2 e2 < num_exp n1 e1`, + REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN + STRIP_TAC THEN + POP_ASSUM MP_TAC THEN POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN DISCH_TAC THEN + MP_TAC (ARITH_RULE `e2 <= e1 ==> e1 = (e1 - e2) + e2:num`) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN (fun th -> ONCE_REWRITE_TAC[th]) THEN + REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN + ASM_REWRITE_TAC[LT_MULT_RCANCEL; EXP_EQ_0; ARITH_EQ]);; + +let NUM_EXP_LT2 = prove(`!n1 e1 n2 e2 r. e1 <= e2 /\ e2 - e1 = r /\ num_exp n2 r < n1 + ==> num_exp n2 e2 < num_exp n1 e1`, + REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN + STRIP_TAC THEN + POP_ASSUM MP_TAC THEN POP_ASSUM (fun th -> REWRITE_TAC[SYM th]) THEN DISCH_TAC THEN + MP_TAC (ARITH_RULE `e1 <= e2 ==> e2 = (e2 - e1) + e1:num`) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN (fun th -> ONCE_REWRITE_TAC[th]) THEN + REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN + ASM_REWRITE_TAC[LT_MULT_RCANCEL; EXP_EQ_0; ARITH_EQ]);; + +let NUM_EXP_LT1_EQ = prove(`!n1 e1 n2 e2 r x. e2 <= e1 /\ e1 - e2 = r /\ num_exp n1 r = x ==> + (num_exp n1 e1 < num_exp n2 e2 <=> x < n2)`, + REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN + STRIP_TAC THEN + REPLICATE_TAC 2 (POP_ASSUM (fun th -> REWRITE_TAC[SYM th])) THEN + MP_TAC (ARITH_RULE `e2 <= e1 ==> e1 = (e1 - e2) + e2:num`) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN (fun th -> GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV) [th]) THEN + REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN + ASM_REWRITE_TAC[LT_MULT_RCANCEL; EXP_EQ_0; ARITH_EQ]);; + +let NUM_EXP_LT2_EQ = prove(`!n1 e1 n2 e2 r x. e1 <= e2 /\ e2 - e1 = r /\ num_exp n2 r = x ==> + (num_exp n1 e1 < num_exp n2 e2 <=> n1 < x)`, + REPEAT GEN_TAC THEN REWRITE_TAC[num_exp] THEN + STRIP_TAC THEN + REPLICATE_TAC 2 (POP_ASSUM (fun th -> REWRITE_TAC[SYM th])) THEN + MP_TAC (ARITH_RULE `e1 <= e2 ==> e2 = (e2 - e1) + e1:num`) THEN + ASM_REWRITE_TAC[] THEN + DISCH_THEN (fun th -> GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV) [th]) THEN + REWRITE_TAC[EXP_ADD; MULT_ASSOC] THEN + ASM_REWRITE_TAC[LT_MULT_RCANCEL; EXP_EQ_0; ARITH_EQ]);; + +(* NUM_EXP_DIV *) + +let mul_op_num = `( * ):num->num->num`;; + +let NUM_EXP_DIV1 = prove(`~(n2 = 0) /\ e2 <= e1 ==> + num_exp n1 e1 DIV num_exp n2 e2 = num_exp n1 (e1 - e2) DIV n2`, + STRIP_TAC THEN + (*`num_exp n1 e1 = 16 EXP e2 * num_exp n1 (e1 - e2)` MP_TAC THENL*) + SUBGOAL_THEN (mk_eq(`num_exp n1 e1`, mk_binop mul_op_num (mk_binop `EXP` base_const `e2:num`) `num_exp n1 (e1 - e2)`)) MP_TAC THENL + [ + REWRITE_TAC[num_exp] THEN + ONCE_REWRITE_TAC[ARITH_RULE `a * b * c = b * (a * c:num)`] THEN + REWRITE_TAC[GSYM EXP_ADD] THEN + ASM_SIMP_TAC[ARITH_RULE `e2 <= e1 ==> e2 + e1 - e2 = e1:num`]; + ALL_TAC + ] THEN + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + SUBGOAL_THEN (mk_eq(`num_exp n2 e2`, mk_binop mul_op_num (mk_binop `EXP` base_const `e2:num`) `n2:num`)) MP_TAC THENL + [ + REWRITE_TAC[num_exp; MULT_AC]; + ALL_TAC + ] THEN + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + MATCH_MP_TAC DIV_MULT2 THEN + ASM_REWRITE_TAC[MULT_EQ_0; DE_MORGAN_THM; EXP_EQ_0] THEN + ARITH_TAC);; + +let NUM_EXP_DIV2 = prove(`~(n2 = 0) /\ e1 <= e2 ==> + num_exp n1 e1 DIV num_exp n2 e2 = n1 DIV num_exp n2 (e2 - e1)`, + STRIP_TAC THEN + (*`num_exp n2 e2 = 16 EXP e1 * num_exp n2 (e2 - e1)` MP_TAC THENL*) + SUBGOAL_THEN (mk_eq(`num_exp n2 e2`, mk_binop mul_op_num (mk_binop `EXP` base_const `e1:num`) `num_exp n2 (e2 - e1)`)) MP_TAC THENL + [ + REWRITE_TAC[num_exp] THEN + ONCE_REWRITE_TAC[ARITH_RULE `a * b * c = b * (a * c:num)`] THEN + REWRITE_TAC[GSYM EXP_ADD] THEN + ASM_SIMP_TAC[ARITH_RULE `e1 <= e2 ==> e1 + e2 - e1 = e2:num`]; + ALL_TAC + ] THEN + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + SUBGOAL_THEN (mk_eq(`num_exp n1 e1`, mk_binop mul_op_num (mk_binop `EXP` base_const `e1:num`) `n1:num`)) MP_TAC THENL + [ + REWRITE_TAC[num_exp; MULT_AC]; + ALL_TAC + ] THEN + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + MATCH_MP_TAC DIV_MULT2 THEN + ASM_REWRITE_TAC[num_exp; MULT_EQ_0; DE_MORGAN_THM; EXP_EQ_0] THEN + ARITH_TAC);; + + +let EXP_INV_lemma = prove(`!n e1 e2. ~(n = 0) /\ e2 <= e1 ==> &(n EXP (e1 - e2)) = + &(n EXP e1) * inv(&(n EXP e2))`, + REPEAT STRIP_TAC THEN + REWRITE_TAC[GSYM REAL_OF_NUM_POW] THEN + MP_TAC (SPECL [`&n`; `e2:num`; `e1:num`] REAL_POW_SUB) THEN + ASM_REWRITE_TAC[REAL_OF_NUM_EQ; real_div]);; + +let NUM_EXP_SUB_lemma = prove(`!n e1 e2. e2 <= e1 ==> &(num_exp n (e1 - e2)) = + &(num_exp n e1) * inv(&(num_exp 1 e2))`, + REPEAT STRIP_TAC THEN + REWRITE_TAC[num_exp] THEN + REWRITE_TAC[GSYM REAL_OF_NUM_MUL] THEN + MP_TAC (SPECL [base_const; `e1:num`; `e2:num`] EXP_INV_lemma) THEN + ANTS_TAC THENL + [ + ASM_REWRITE_TAC[] THEN ARITH_TAC; + ALL_TAC + ] THEN + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + REAL_ARITH_TAC);; + + + +end;; diff --git a/Formal_ineqs/arith_options.hl b/Formal_ineqs/arith_options.hl index a7df310a..41868b58 100644 --- a/Formal_ineqs/arith_options.hl +++ b/Formal_ineqs/arith_options.hl @@ -1,28 +1,34 @@ -(* =========================================================== *) -(* Options of the arithmetic library *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -module Arith_options = struct - -(* Base of arithmetic operations with natural numbers *) -(* The base should be even in order to represent inv(2) exactly *) -let base = ref 200;; -(* If true then results of natural number operations are cached *) -let cached = ref true;; -(* Initial size of the cache *) -let init_cache_size = ref 10000;; -(* Maximal size of the cache *) -let max_cache_size = ref 20000;; - -(* Minimal exponent value for floating point numbers *) -(* (should be even for the square root operation) *) -let min_exp = ref 50;; -(* If true, then arithmetic operations with floating point numbers are cached *) -let float_cached = ref true;; - -(* If true, then all floating point operations are logged *) -let float_log = ref false;; - -end;; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Options of the arithmetic library *) +(* -------------------------------------------------------------------------- *) + +module Arith_options = struct + +(* Base of arithmetic operations with natural numbers *) +(* The base should be even in order to represent inv(2) exactly *) +let base = ref 200;; +(* If true then results of natural number operations are cached *) +let cached = ref true;; +(* Initial size of the cache *) +let init_cache_size = ref 10000;; +(* Maximal size of the cache *) +let max_cache_size = ref 20000;; + +(* Minimal exponent value for floating point numbers *) +(* (should be even for the square root operation) *) +let min_exp = ref 50;; +(* If true, then arithmetic operations with floating point numbers are cached *) +let float_cached = ref true;; + +(* If true, then all floating point operations are logged *) +let float_log = ref false;; + +end;; diff --git a/Formal_ineqs/examples.hl b/Formal_ineqs/examples.hl index d0a0f470..f7ab53e5 100644 --- a/Formal_ineqs/examples.hl +++ b/Formal_ineqs/examples.hl @@ -1,4 +1,14 @@ -(* Several simple examples *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Several simple examples *) +(* -------------------------------------------------------------------------- *) (* Set up the loading path: load_path := "path to the formal_ineqs directory" :: !load_path;; @@ -6,7 +16,7 @@ load_path := "path to the formal_ineqs directory" :: !load_path;; (* Change default arithmetic options before loading other libraries *) (* (arithmetic options cannot be changed later) *) -needs "arith_options.hl";; +needs "Formal_ineqs/arith_options.hl";; (* Set the base of natural number arithmetic to 200 *) Arith_options.base := 200;; @@ -15,7 +25,7 @@ Arith_options.base := 200;; (* Note: the verification library loads Multivariate/realanalysis.ml, so it is recommended to use a checkpointed version of HOL Light with preloaded realanalysis.ml *) -needs "verifier/m_verifier_main.hl";; +needs "Formal_ineqs/verifier/m_verifier_main.hl";; (* Set the level of info/debug printing: @@ -23,7 +33,7 @@ needs "verifier/m_verifier_main.hl";; 1 - report important steps (default) 2 - report everything *) -needs "verifier_options.hl";; +needs "Formal_ineqs/verifier_options.hl";; Verifier_options.info_print_level := 1;; (* Open the main verification module *) @@ -54,8 +64,8 @@ let test4 () = (* A polynomial approximation of atn *) (* Taken from: *) -(* Marc Daumas, David Lester, and César Muñoz, - Verified real number calculations: A library for interval arithmetic, +(* Marc Daumas, David Lester, and César Muñoz, + Verified real number calculations: A library for interval arithmetic, IEEE Transactions on Computers, Volume 58, Number 2, 2009. *) let test5 () = let ineq1 = `-- &1 / &30 <= x /\ x <= &1 / &30 ==> x * (&1 - (x * x) * (&11184811 / &33554432 - (x * x) * (&13421773 / &67108864))) - atn x < #0.1 pow 7` in @@ -64,7 +74,7 @@ let test5 () = verify_ineq default_params 5 ineq2];; let test5_abs () = - let ineq_abs = `-- &1 / &30 <= x /\ x <= &1 / &30 + let ineq_abs = `-- &1 / &30 <= x /\ x <= &1 / &30 ==> abs (x * (&1 - (x * x) * (&11184811 / &33554432 - (x * x) * (&13421773 / &67108864))) - atn x) < #0.1 pow 7` in [verify_ineq default_params 10 ineq_abs];; diff --git a/Formal_ineqs/examples_flyspeck.hl b/Formal_ineqs/examples_flyspeck.hl index 33d2a8ef..3cb14d6c 100644 --- a/Formal_ineqs/examples_flyspeck.hl +++ b/Formal_ineqs/examples_flyspeck.hl @@ -1,5 +1,15 @@ -(* Some inequalities from the Flyspeck project *) -(* https://github.com/flyspeck/flyspeck *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Some inequalities from the Flyspeck project *) +(* https://github.com/flyspeck/flyspeck *) +(* -------------------------------------------------------------------------- *) (* Set up the loading path: load_path := "path to the formal_ineqs directory" :: !load_path;; @@ -7,7 +17,7 @@ load_path := "path to the formal_ineqs directory" :: !load_path;; (* Change default arithmetic options before loading other libraries *) (* (arithmetic options cannot be changed later) *) -needs "arith_options.hl";; +needs "Formal_ineqs/arith_options.hl";; (* Set the base of natural number arithmetic to 200 *) Arith_options.base := 200;; @@ -16,7 +26,7 @@ Arith_options.base := 200;; (* Note: the verification library loads Multivariate/realanalysis.ml, so it is recommended to use a checkpointed version of HOL Light with preloaded realanalysis.ml *) -needs "verifier/m_verifier_main.hl";; +needs "Formal_ineqs/verifier/m_verifier_main.hl";; (* Set the level of info/debug printing: @@ -24,7 +34,7 @@ needs "verifier/m_verifier_main.hl";; 1 - report important steps (default) 2 - report everything *) -needs "verifier_options.hl";; +needs "Formal_ineqs/verifier_options.hl";; Verifier_options.info_print_level := 1;; (* Open the main verification module *) @@ -44,7 +54,7 @@ let ineq = define (* A modified (only one case is considered, x > 0) definition of atn2 *) (* Add ' to some definitions to avoid conflicts with original Flyspeck definitions *) -let atn2' = new_definition `atn2'(x,y) = atn(y / x)` +let atn2' = new_definition `atn2'(x,y) = atn(y / x)`;; (* delta_x *) @@ -113,39 +123,39 @@ let h0 = new_definition `h0 = #1.26`;; let lfun = new_definition `lfun h = (h0 - h)/(h0 - &1)`;; (* lfun_y1 *) -let lfun_y1 = new_definition `lfun_y1 (y1:real) (y2:real) (y3:real) +let lfun_y1 = new_definition `lfun_y1 (y1:real) (y2:real) (y3:real) (y4:real) (y5:real) (y6:real) = lfun y1`;; (* num1 *) let num1 = new_definition `num1 e1 e2 e3 a2 b2 c2 = - -- &4*((a2 pow 2) *e1 + &8*(b2 - c2)*(e2 - e3) - + -- &4*((a2 pow 2) *e1 + &8*(b2 - c2)*(e2 - e3) - a2*(&16*e1 + ( b2 - &8 )*e2 + (c2 - &8)*e3))`;; (* unit6 *) let unit6 = define `unit6 x1 x2 x3 x4 x5 x6 = &1`;; (* arc_hhn *) -let arc_hhn' = new_definition `arc_hhn' = +let arc_hhn' = new_definition `arc_hhn' = arclength' (&2 * h0) (&2 * h0) (&2)`;; (* arclength_y1 *) -let arclength_y1' = new_definition - `arclength_y1' a b +let arclength_y1' = new_definition + `arclength_y1' a b (y1:real) (y2:real) (y3:real) (y4:real) (y5:real) (y6:real) = arclength' y1 a b`;; (* arclength_x1 *) -let arclength_x1' = new_definition - `arclength_x1' a b x1 x2 x3 x4 x5 x6 = +let arclength_x1' = new_definition + `arclength_x1' a b x1 x2 x3 x4 x5 x6 = arclength_y1' a b (sqrt x1) (sqrt x2) (sqrt x3) (sqrt x4) (sqrt x5) (sqrt x6)`;; (* arclength_x_123 *) -let arclength_x_123' = new_definition `arclength_x_123' (x1:real) (x2:real) (x3:real) (x4:real) (x5:real) (x6:real) = +let arclength_x_123' = new_definition `arclength_x_123' (x1:real) (x2:real) (x3:real) (x4:real) (x5:real) (x6:real) = arclength' (sqrt x1) (sqrt x2) (sqrt x3)`;; - -(* acs_sqrt_x1_d4 *) -let acs_sqrt_x1_d4 = new_definition `acs_sqrt_x1_d4 (x1:real) (x2:real) (x3:real) (x4:real) (x5:real) (x6:real) = + +(* acs_sqrt_x1_d4 *) +let acs_sqrt_x1_d4 = new_definition `acs_sqrt_x1_d4 (x1:real) (x2:real) (x3:real) (x4:real) (x5:real) (x6:real) = acs (sqrt(x1)/ &4)`;; let sqrt_x1 = define `sqrt_x1 x1 x2 x3 x4 x5 x6 = sqrt x1`;; @@ -198,7 +208,7 @@ let add_example ex = add_example {id = "2485876245a"; difficulty = Easy; ineq_tm = `ineq - [ #4.0,x1, #6.3504; #4.0,x2, #6.3504; #4.0,x3, #6.3504; #4.0,x4, #6.3504; + [ #4.0,x1, #6.3504; #4.0,x2, #6.3504; #4.0,x3, #6.3504; #4.0,x4, #6.3504; #3.0 * #3.0, x5, #2.0 * #2.52 * #2.0 * #2.52; #4.0,x6, #6.3504] (delta_x4 x1 x2 x3 x4 x5 x6 * -- &1 < &0)`};; @@ -207,7 +217,7 @@ add_example {id = "2485876245a"; add_example {id = "4559601669b"; difficulty = Easy; ineq_tm = `ineq - [ #4.0,x1, #6.3504; #4.0,x2, #4.0; #4.0,x3, #6.3504; + [ #4.0,x1, #6.3504; #4.0,x2, #4.0; #4.0,x3, #6.3504; #3.01 * #3.01, x4, #3.01 * #3.01; #4.0, x5, #6.3504; #4.0,x6, #4.0] (delta_x4 x1 x2 x3 x4 x5 x6 < &0)`};; @@ -215,8 +225,8 @@ add_example {id = "4559601669b"; (* 5512912661 *) add_example {id = "5512912661"; difficulty = Easy; - ineq_tm = `ineq [&1,x1,&1 + (pi * const1') / pi; &1,x2,&1 + (pi * const1') / pi; - &1, x3, &1 + (pi * const1') / pi; #2.38 * #2.38, x4, #3.01 * #3.01; + ineq_tm = `ineq [&1,x1,&1 + (pi * const1') / pi; &1,x2,&1 + (pi * const1') / pi; + &1, x3, &1 + (pi * const1') / pi; #2.38 * #2.38, x4, #3.01 * #3.01; &2 * &2, x5, #2.52 * #2.52; #3.15 / #1.26 * #3.15 / #1.26,x6, #15.53] (num1 x1 x2 x3 x4 x5 x6 * -- &1 < &0)`};; @@ -224,8 +234,8 @@ add_example {id = "5512912661"; (* 6843920790 *) add_example {id = "6843920790"; difficulty = Easy; - ineq_tm = `ineq [&1,x1,&1 + (pi * const1') / pi; &1,x2,&1 + (pi * const1') / pi; - &1, x3, &1 + (pi * const1') / pi; &2 / #1.26 * &2 / #1.26, x4, #3.01 * #3.01; + ineq_tm = `ineq [&1,x1,&1 + (pi * const1') / pi; &1,x2,&1 + (pi * const1') / pi; + &1, x3, &1 + (pi * const1') / pi; &2 / #1.26 * &2 / #1.26, x4, #3.01 * #3.01; #2.38 * #2.38, x5, #15.53; #2.38 * #2.38,x6, #15.53] (num1 x1 x2 x3 x4 x5 x6 * -- &1 < &0)`};; @@ -245,8 +255,8 @@ add_example {id = "6096597438a"; add_example {id = "4717061266"; difficulty = Easy; ineq_tm = `ineq - [ #4.0,x1, #2.0 * #1.26 * #2.0 * #1.26; #4.0, x2, #2.0 * #1.26 * #2.0 * #1.26; - #4.0, x3, #2.0 * #1.26 * #2.0 * #1.26; #4.0,x4, #2.0 * #1.26 * #2.0 * #1.26; + [ #4.0,x1, #2.0 * #1.26 * #2.0 * #1.26; #4.0, x2, #2.0 * #1.26 * #2.0 * #1.26; + #4.0, x3, #2.0 * #1.26 * #2.0 * #1.26; #4.0,x4, #2.0 * #1.26 * #2.0 * #1.26; #4.0, x5, #2.0 * #1.26 * #2.0 * #1.26; #4.0,x6, #2.0 * #1.26 * #2.0 * #1.26] (delta_x x1 x2 x3 x4 x5 x6 * -- &1 < &0)`};; @@ -254,7 +264,7 @@ add_example {id = "4717061266"; (* SDCCMGA b *) add_example {id = "SDCCMGA b"; difficulty = Easy; - ineq_tm = `ineq [ #4.0,x1, #6.3504; &1 * &1,x2,&1 * &1; &1 * &1,x3,&1 * &1; + ineq_tm = `ineq [ #4.0,x1, #6.3504; &1 * &1,x2,&1 * &1; &1 * &1,x3,&1 * &1; &1 * &1, x4, &1 * &1; &1 * &1, x5, &1 * &1; &1 * &1,x6,&1 * &1] (arclength_x1' #2.0 ( #2.0 * #1.26) x1 x2 x3 x4 x5 x6 + arclength_x1' #2.0 ( #2.0 * #1.26) x1 x2 x3 x4 x5 x6 + @@ -267,7 +277,7 @@ add_example {id = "SDCCMGA b"; add_example {id = "TSKAJXY-TADIAMB"; difficulty = Medium; ineq_tm = `ineq - [ #2.0 * #1.3254 * #2.0 * #1.3254,x1, #8.0; #2.0 * #1.3254 * #2.0 * #1.3254, x2, #8.0; + [ #2.0 * #1.3254 * #2.0 * #1.3254,x1, #8.0; #2.0 * #1.3254 * #2.0 * #1.3254, x2, #8.0; #4.0,x3, #8.0; #4.0, x4, #8.0; #4.0,x5, #8.0; #4.0,x6, #8.0] ((unit6 x1 x2 x3 x4 x5 x6 * #2.0) * (delta_x x1 x2 x3 x4 x5 x6 * &4) < rho_x x1 x2 x3 x4 x5 x6)`};; @@ -275,7 +285,7 @@ add_example {id = "TSKAJXY-TADIAMB"; (* 7067938795 *) add_example {id = "7067938795"; difficulty = Medium; - ineq_tm = `ineq [ #4.0,x1, #6.3504; #4.0,x2, #6.3504; #4.0,x3, #6.3504; #4.0,x4, #4.0; + ineq_tm = `ineq [ #4.0,x1, #6.3504; #4.0,x2, #6.3504; #4.0,x3, #6.3504; #4.0,x4, #4.0; #3.01 * #3.01, x5, #3.24 * #3.24; #3.01 * #3.01,x6, #3.24 * #3.24] (dih_x' x1 x2 x3 x4 x5 x6 + unit6 x1 x2 x3 x4 x5 x6 * pi * --(&1 / #2.0) + @@ -286,7 +296,7 @@ add_example {id = "7067938795"; add_example { id = "5490182221"; difficulty = Medium; ineq_tm = `ineq - [ #4.0,x1, #6.3504; #4.0,x2, #6.3504; #4.0,x3, #6.3504; #4.0,x4, #6.3504; + [ #4.0,x1, #6.3504; #4.0,x2, #6.3504; #4.0,x3, #6.3504; #4.0,x4, #6.3504; #4.0, x5, #6.3504; #4.0,x6, #6.3504] (dih_x' x1 x2 x3 x4 x5 x6 + unit6 x1 x2 x3 x4 x5 x6 * -- #1.893 < &0)`};; @@ -298,9 +308,9 @@ add_example { id = "3318775219"; ineq_tm = `ineq [&2, y1, #2.52; &2, y2, #2.52; &2, y3, #2.52; #2.52, y4, sqrt(&8); &2, y5, #2.52; &2, y6, #2.52] - ( ((dih_y' y1 y2 y3 y4 y5 y6) - #1.629 + + ( ((dih_y' y1 y2 y3 y4 y5 y6) - #1.629 + (#0.414 * (y2 + y3 + y5 + y6 - #8.0)) - - (#0.763 * (y4 - #2.52)) - + (#0.763 * (y4 - #2.52)) - (#0.315 * (y1 - #2.0))) * (-- &1) < &0)`};; diff --git a/Formal_ineqs/examples_other.hl b/Formal_ineqs/examples_other.hl index 0001d7c9..d6422662 100644 --- a/Formal_ineqs/examples_other.hl +++ b/Formal_ineqs/examples_other.hl @@ -1,16 +1,28 @@ -needs "arith_options.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Additional test inequalities *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/arith_options.hl";; (* Arith_options.base := 100;; *) -needs "verifier/m_verifier_main.hl";; +needs "Formal_ineqs/verifier/m_verifier_main.hl";; -needs "verifier_options.hl";; +needs "Formal_ineqs/verifier_options.hl";; Verifier_options.info_print_level := 2;; open M_verifier_main;; -verify_ineq default_params 5 +verify_ineq default_params 5 `&0 <= x /\ x <= &1 ==> exp (x pow 3) < &3`;; verify_ineq default_params 5 @@ -19,7 +31,7 @@ verify_ineq default_params 5 (* MetiTarski home page *) verify_ineq default_params 5 - `&0 <= x /\ x <= #1.46 / &10 pow 6 ==> + `&0 <= x /\ x <= #1.46 / &10 pow 6 ==> (#64.42 * sin(#1.71 * &10 pow 6 * x) - #21.08 * cos(#1.71 * &10 pow 6 * x)) * exp(#9.05 * &10 pow 5 * x) + #24.24 * exp(-- #1.86 * &10 pow 6 * x) > &0`;; @@ -30,6 +42,20 @@ verify_ineq default_params 5 verify_ineq default_params 5 `#0.35 <= t /\ t <= &10 /\ &0 <= v /\ v <= &10 ==> - ((#1.565 + #0.313 * v) * cos(#1.16 * v) + + ((#1.565 + #0.313 * v) * cos(#1.16 * v) + (#0.01340 + #0.00268 * v) * sin(#1.16 * t)) * exp(-- #1.34 * t) - (#6.55 + #1.31 * v) * exp(-- #0.318 * t) + v + &10 > &0`;; + +(* Approximations of abs(x) from + "A certificate-based approach to formally verified approximations" *) +verify_ineq default_params 5 + `&0 <= x /\ x <= &1 ==> abs(sqrt(#0.01 + x * x) - abs(x)) < #0.1001`;; + +verify_ineq default_params 5 + `&0 <= x /\ x <= &1 ==> abs(sqrt(#0.001 * #0.001 + x * x) - abs(x)) < #0.0011`;; + +verify_ineq default_params 5 + `&0 <= x /\ x <= &1 ==> sqrt(#0.001 * #0.001 + x * x) - abs(x) < #0.0011`;; + +verify_ineq default_params 5 + `&0 <= x /\ x <= &1 ==> sqrt(#0.001 * #0.001 + x * x) - abs(x) > -- #0.0011`;; diff --git a/Formal_ineqs/examples_poly.hl b/Formal_ineqs/examples_poly.hl index df076201..a30668d6 100644 --- a/Formal_ineqs/examples_poly.hl +++ b/Formal_ineqs/examples_poly.hl @@ -1,10 +1,20 @@ -(* Multivariate polynomial inequalities *) -(* Examples are taken from the paper: - César Muñoz and Anthony Narkawicz, - Formalization of a Representation of Bernstein Polynomials and Applications to Global Optimization, - Journal of Automated Reasoning, DOI: 10.1007/s10817-012-9256-3 - http://shemesh.larc.nasa.gov/people/cam/Bernstein/ *) - +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Multivariate polynomial inequalities *) +(* Examples are taken from the paper: *) +(* César Muñoz and Anthony Narkawicz, *) +(* Formalization of a Representation of Bernstein Polynomials and *) +(* Applications to Global Optimization, *) +(* Journal of Automated Reasoning, DOI: 10.1007/s10817-012-9256-3 *) +(* http://shemesh.larc.nasa.gov/people/cam/Bernstein/ *) +(* -------------------------------------------------------------------------- *) (* Set up the loading path: load_path := "path to the formal_ineqs directory" :: !load_path;; @@ -12,7 +22,7 @@ load_path := "path to the formal_ineqs directory" :: !load_path;; (* Change default arithmetic options before loading other libraries *) (* (arithmetic options cannot be changed later) *) -needs "arith_options.hl";; +needs "Formal_ineqs/arith_options.hl";; (* Set the base of natural number arithmetic to 200 *) Arith_options.base := 200;; @@ -21,7 +31,7 @@ Arith_options.base := 200;; (* Note: the verification library loads Multivariate/realanalysis.ml, so it is recommended to use a checkpointed version of HOL Light with preloaded realanalysis.ml *) -needs "verifier/m_verifier_main.hl";; +needs "Formal_ineqs/verifier/m_verifier_main.hl";; (* Set the level of info/debug printing: @@ -29,7 +39,7 @@ needs "verifier/m_verifier_main.hl";; 1 - report important steps (default) 2 - report everything *) -needs "verifier_options.hl";; +needs "Formal_ineqs/verifier_options.hl";; Verifier_options.info_print_level := 1;; (* Open the main verification module *) @@ -39,13 +49,13 @@ open M_verifier_main;; (* Data *) (* Polynomials *) -let schwefel_poly = `(x1 - x2 pow 2) pow 2 + (x2 - &1) pow 2 + +let schwefel_poly = `(x1 - x2 pow 2) pow 2 + (x2 - &1) pow 2 + (x1 - x3 pow 2) pow 2 + (x3 - &1) pow 2` and rd_poly = `-- x1 + &2 * x2 - x3 - #0.835634534 * x2 * (&1 + x2)` and - caprasse_poly = `-- x1 * x3 pow 3 + &4 * x2 * x3 pow 2 * x4 + - &4 * x1 * x3 * x4 pow 2 + &2 * x2 * x4 pow 3 + &4 * x1 * x3 + &4 * x3 pow 2 - + caprasse_poly = `-- x1 * x3 pow 3 + &4 * x2 * x3 pow 2 * x4 + + &4 * x1 * x3 * x4 pow 2 + &2 * x2 * x4 pow 3 + &4 * x1 * x3 + &4 * x3 pow 2 - &10 * x2 * x4 - &10 * x4 pow 2 + &2` and lv_poly = `x1 * x2 pow 2 + x1 * x3 pow 2 + x1 * x4 pow 2 - #1.1 * x1 + &1` and @@ -57,7 +67,7 @@ let schwefel_poly = `(x1 - x2 pow 2) pow 2 + (x2 - &1) pow 2 + &2 * x5 pow 2 + &2 * x6 pow 2 + &2 * x7 pow 2 - x1` and heart_poly = `-- x1 * x6 pow 3 + &3 * x1 * x6 * x7 pow 2 - x3 * x7 pow 3 + - &3 * x3 * x7 * x6 pow 2 - x2 * x5 pow 3 + &3 * x2 * x5 * x8 pow 2 - x4 * x8 pow 3 + + &3 * x3 * x7 * x6 pow 2 - x2 * x5 pow 3 + &3 * x2 * x5 * x8 pow 2 - x4 * x8 pow 3 + &3 * x4 * x8 * x5 pow 2 - #0.9563453`;; (* Minimal values *) @@ -130,7 +140,7 @@ let test_heart () = verify_ineq {default_params with eps = 1e-10} 5 heart_ineq;; let run_tests () = - [test_schwefel(); test_rd(); test_caprasse(); test_lv(); + [test_schwefel(); test_rd(); test_caprasse(); test_lv(); test_butcher(); test_magnetism(); test_heart()];; let results () = diff --git a/Formal_ineqs/informal/informal_asn_acs.hl b/Formal_ineqs/informal/informal_asn_acs.hl index 6943afbf..0c5810c1 100644 --- a/Formal_ineqs/informal/informal_asn_acs.hl +++ b/Formal_ineqs/informal/informal_asn_acs.hl @@ -1,11 +1,17 @@ -(* =========================================================== *) -(* Informal arithmetic procedures: asn, acs *) -(* Author: Alexey Solovyev *) -(* Date: 2014-10-01 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal arithmetic procedures: asn, acs *) +(* -------------------------------------------------------------------------- *) (* Dependencies *) -needs "informal/informal_atn.hl";; +needs "Formal_ineqs/informal/informal_atn.hl";; module type Informal_asn_acs_sig = sig @@ -120,7 +126,7 @@ end;; (* -needs "trig/asn_acs_eval.hl";; +needs "Formal_ineqs/trig/asn_acs_eval.hl";; open Informal_asn_acs;; diff --git a/Formal_ineqs/informal/informal_atn.hl b/Formal_ineqs/informal/informal_atn.hl index 02bf8c3a..0319749d 100644 --- a/Formal_ineqs/informal/informal_atn.hl +++ b/Formal_ineqs/informal/informal_atn.hl @@ -1,12 +1,18 @@ -(* =========================================================== *) -(* Informal arithmetic procedures: atn *) -(* Author: Alexey Solovyev *) -(* Date: 2014-09-26 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal arithmetic procedures: atn *) +(* -------------------------------------------------------------------------- *) (* Dependencies *) -needs "informal/informal_interval.hl";; -needs "informal/informal_poly.hl";; +needs "Formal_ineqs/informal/informal_interval.hl";; +needs "Formal_ineqs/informal/informal_poly.hl";; module type Informal_atn_sig = sig @@ -54,13 +60,13 @@ let halfatn4_pos_lo pp t = f (f (f (f t)));; let rec fact n = - if sign_num n <= 0 then num 1 else - n */ fact (pred_num n);; + if sign_num n <= 0 then Num.num_of_int 1 else + n */ fact (n -/ num_1);; let atn_tables pp n = let t1 = map (fun i -> 4 * i + 1) (0 -- (n / 2)) and t2 = map (fun i -> 4 * i + 3) (0 -- ((n + 1) / 2 - 1)) in - let f n = if n = 1 then mk_num_interval (num n) else inv_interval pp (mk_num_interval (num n)) in + let f n = if n = 1 then mk_num_interval (Num.num_of_int n) else inv_interval pp (mk_num_interval (Num.num_of_int n)) in map f t1, map f t2;; (* Computes x^k / k! *) @@ -69,14 +75,14 @@ let rec x_pow_over_fact x k = x /. (float_of_int k) *. x_pow_over_fact x (k - 1);; (* Computes i such that x^(2 * i + 1) / (2 * i + 1) <= base^(-(p + 1)) and cond(i) *) -let n_of_p_atn x pp cond = +let n_of_p_atn x pp cond = let t = (float_of_int Informal_nat.arith_base) ** (float_of_int (-pp - 1)) in let rec try_i i = let _ = if i > 50 then failwith "n_of_p_atn: cannot find i" else () in if cond i then let d = float_of_int (2 * i + 1) in let r = (x ** d) /. d in - if r <= t then i else try_i (i + 1) + if r <= t then i else try_i (i + 1) else try_i (i + 1) in @@ -175,7 +181,7 @@ end;; (* -needs "trig/atn_eval.hl";; +needs "Formal_ineqs/trig/atn_eval.hl";; #install_printer print_ifloat;; #install_printer print_interval;; diff --git a/Formal_ineqs/informal/informal_eval_interval.hl b/Formal_ineqs/informal/informal_eval_interval.hl index 5a70ccca..8e9c62a6 100644 --- a/Formal_ineqs/informal/informal_eval_interval.hl +++ b/Formal_ineqs/informal/informal_eval_interval.hl @@ -1,15 +1,21 @@ -(* =========================================================== *) -(* Informal interval evaluation of arithmetic expressions *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -needs "informal/informal_exp.hl";; -needs "informal/informal_log.hl";; -needs "informal/informal_atn.hl";; -needs "informal/informal_asn_acs.hl";; -needs "informal/informal_sin_cos.hl";; -needs "informal/informal_matan.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal interval evaluation of arithmetic expressions *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/informal/informal_exp.hl";; +needs "Formal_ineqs/informal/informal_log.hl";; +needs "Formal_ineqs/informal/informal_atn.hl";; +needs "Formal_ineqs/informal/informal_asn_acs.hl";; +needs "Formal_ineqs/informal/informal_sin_cos.hl";; +needs "Formal_ineqs/informal/informal_matan.hl";; module Informal_eval_interval = struct @@ -32,7 +38,7 @@ let mk_float_interval_decimal pp decimal_tm = (* Unary interval operations *) -let unary_interval_operations = +let unary_interval_operations = let table = Hashtbl.create 10 in let add = Hashtbl.add in add table "real_neg" (fun pp -> neg_interval); @@ -53,7 +59,7 @@ let unary_interval_operations = (* Binary interval operations *) -let binary_interval_operations = +let binary_interval_operations = let table = Hashtbl.create 10 in let add = Hashtbl.add in add table "real_add" add_interval; @@ -131,7 +137,7 @@ let eval_constants = match f with | Int_decimal_const tm -> Int_const (mk_float_interval_decimal pp tm) | Int_named_const name -> Int_const (c_find name pp) - | Int_pow (n, f1) -> + | Int_pow (n, f1) -> (let f1_val = rec_eval f1 in match f1_val with | Int_const int -> Int_const (pow_interval pp n int) @@ -174,7 +180,7 @@ let build_interval_fun = if ltm = amp_op_real then let n = dest_numeral r_tm in Int_const (mk_num_interval n) - else + else let r_fun = rec_build r_tm in Int_unary ((fst o dest_const) ltm, r_fun) else @@ -244,7 +250,7 @@ let find_and_replace_all f_list acc = if c1 > 1 then expr, (c1, fs) else find_and_replace f2 i f_list | _ -> f, (0, f_list) in if c > 1 then expr, (c, fs) else f, replace_subexpr f i f_list in - + let rec iterate fs acc = let i = length acc in let expr, (c, fs') = find_and_replace (hd fs) i fs in @@ -281,7 +287,7 @@ let interval_to_float_interval pp int_th = let a, _ = dest_interval int_lo and _, b = dest_interval int_hi in mk_interval (a, b);; - + (* Adds a new constant approximation to the table of constants *) let add_constant_interval int_th = diff --git a/Formal_ineqs/informal/informal_exp.hl b/Formal_ineqs/informal/informal_exp.hl index d863bd06..475e1fb6 100644 --- a/Formal_ineqs/informal/informal_exp.hl +++ b/Formal_ineqs/informal/informal_exp.hl @@ -1,12 +1,18 @@ -(* =========================================================== *) -(* Informal arithmetic procedures: exp *) -(* Author: Alexey Solovyev *) -(* Date: 2014-12-10 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal arithmetic procedures: exp *) +(* -------------------------------------------------------------------------- *) (* Dependencies *) -needs "informal/informal_interval.hl";; -needs "informal/informal_poly.hl";; +needs "Formal_ineqs/informal/informal_interval.hl";; +needs "Formal_ineqs/informal/informal_poly.hl";; module type Informal_exp_sig = sig @@ -25,19 +31,20 @@ open Informal_poly;; (* TODO: Make sure that this constant is the same as in trig/exp_eval.hl *) let exp_max_x = 1.0;; +let num_1 = Num.num_of_int 1;; let rec fact n = - if sign_num n <= 0 then num 1 else - n */ fact (pred_num n);; + if sign_num n <= 0 then Num.num_of_int 1 else + n */ fact (n -/ num_1);; let exp_pos_tables pp n = - let t1 = map (fun i -> fact (num i)) (0 -- n) in + let t1 = map (fun i -> fact (Num.num_of_int i)) (0 -- n) in let f n = inv_interval pp (mk_num_interval n) in map f t1;; let exp_neg_tables pp n = - let t1 = map (fun i -> fact (num (2 * i))) (0 -- (n / 2)) and - t2 = map (fun i -> fact (num (2 * i + 1))) (0 -- ((n + 1) / 2 - 1)) in + let t1 = map (fun i -> fact (Num.num_of_int (2 * i))) (0 -- (n / 2)) and + t2 = map (fun i -> fact (Num.num_of_int (2 * i + 1))) (0 -- ((n + 1) / 2 - 1)) in let f n = inv_interval pp (mk_num_interval n) in map f t1, map f t2;; @@ -47,13 +54,13 @@ let rec x_pow_over_fact x k = x /. (float_of_int k) *. x_pow_over_fact x (k - 1);; (* Computes i such that x^i / i! <= base^(-(p + 1)) and cond(i) *) -let n_of_p_exp x pp cond = +let n_of_p_exp x pp cond = let t = (float_of_int Informal_nat.arith_base) ** (float_of_int (-pp - 1)) in let rec try_i i = let _ = if i > 50 then failwith "n_of_p_exp: cannot find i" else () in if cond i then let r = x_pow_over_fact x i in - if r <= t then i else try_i (i + 1) + if r <= t then i else try_i (i + 1) else try_i (i + 1) in @@ -172,7 +179,7 @@ end;; (* -needs "trig/exp_eval.hl";; +needs "Formal_ineqs/trig/exp_eval.hl";; open Informal_float;; open Informal_interval;; diff --git a/Formal_ineqs/informal/informal_float.hl b/Formal_ineqs/informal/informal_float.hl index 3fa679dd..81afbc60 100644 --- a/Formal_ineqs/informal/informal_float.hl +++ b/Formal_ineqs/informal/informal_float.hl @@ -1,11 +1,17 @@ -(* =========================================================== *) -(* Informal arithmetic procedures: floating-point numbers *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal arithmetic procedures: floating-point numbers *) +(* -------------------------------------------------------------------------- *) (* Dependencies *) -needs "informal/informal_nat.hl";; +needs "Formal_ineqs/informal/informal_nat.hl";; (* Floating-point numbers *) @@ -73,7 +79,7 @@ let print_ifloat_fmt fmt (s, n, e) = let k = e - min_exp in let n_str = string_of_num (dest_nat n) in let s_str = if s then "-" else "" in - let str = + let str = if k = 0 then Printf.sprintf "%s%s" s_str n_str else @@ -84,11 +90,11 @@ let print_ifloat = print_ifloat_fmt Format.std_formatter;; (* Creates a floating-point value *) -let mk_float n e : ifloat = +let mk_float n e : ifloat = if n < 0 then - true, mk_nat (minus_num (num n)), e + min_exp + true, mk_nat (minus_num (Num.num_of_int n)), e + min_exp else - false, mk_nat (num n), e + min_exp;; + false, mk_nat (Num.num_of_int n), e + min_exp;; let mk_num_float n = false, mk_nat n, min_exp;; @@ -107,13 +113,12 @@ let make_float s n e : ifloat = s, mk_nat n, e;; let num_of_ifloat = let b = num_of_int arith_base in fun ((s, n, e) : ifloat) -> - let r = (dest_nat n) */ (b **/ num (e - min_exp)) in + let r = (dest_nat n) */ (b **/ Num.num_of_int (e - min_exp)) in if s then minus_num r else r;; -let float_of_ifloat f = +let float_of_ifloat f = let n = num_of_ifloat f in - let a = approx_num_exp 30 n in - float_of_string a;; + Num.float_of_num n;; let ifloat_of_float = let split = @@ -139,14 +144,14 @@ let ifloat_of_float = s, t, n and extract = let b = float_of_int arith_base in - let nb = num arith_base in + let nb = Num.num_of_int arith_base in let rec step k f acc = if k <= 0 then acc else let d = int_of_float f in - step (k - 1) ((f -. float_of_int d) *. b) (num d +/ (nb */ acc)) in + step (k - 1) ((f -. float_of_int d) *. b) (Num.num_of_int d +/ (nb */ acc)) in fun pp f -> - step pp f (num 0) + step pp f (Num.num_of_int 0) in fun pp f -> let s, x, k = split f in @@ -220,7 +225,7 @@ let neg_float (s,n,e) = (not s, n, e);; (* abs *) let abs_float (_,n,e) = (false, n, e);; - + (* lt0, gt0 *) @@ -247,7 +252,7 @@ let lt_float (s1,n1,e1) (s2,n2,e2) = if not s1 then if s2 then false else num_exp_lt (n1,e1) (n2,e2) else - if s2 then num_exp_lt (n2,e2) (n1,e1) + if s2 then num_exp_lt (n2,e2) (n1,e1) else (* TF *) if eq0_nat n1 then gt0_nat n2 else true;; @@ -262,8 +267,8 @@ let le_float (s1,n1,e1) (s2,n2,e2) = (* FT *) if eq0_nat n2 then eq0_nat n1 else false;; - - + + (* min, max *) let min_float f1 f2 = @@ -406,7 +411,7 @@ let add_float_hi pp (s1,n1,e1) (s2,n2,e2) = else let n, e'' = lo_nat pp n' in (true, n, e' + e'');; - + (* sub *) @@ -445,7 +450,7 @@ let rec sqrt_float_hi pp (s,n1,e1) = let x = (big_int_of_num o dest_nat o denormalize_nat) (n1, p2) in let f1' = Big_int.sqrt_big_int x in let f1 = (mk_nat o num_of_big_int) f1' in - let n, e' = + let n, e' = let ( * ) = Big_int.mult_big_int and (==) = Big_int.eq_big_int in hi_nat pp (if f1' * f1' == x then f1 else suc_nat f1) in @@ -488,7 +493,7 @@ let pow_float_pos_lo pp n (x : ifloat) = mul_float_lo pp x t in let _ = assert (n >= 0) in pow n;; - + let pow_float_hi pp n x = match n with | 0 -> float1 diff --git a/Formal_ineqs/informal/informal_interval.hl b/Formal_ineqs/informal/informal_interval.hl index 8d890ad5..50996452 100644 --- a/Formal_ineqs/informal/informal_interval.hl +++ b/Formal_ineqs/informal/informal_interval.hl @@ -1,11 +1,17 @@ -(* =========================================================== *) -(* Informal arithmetic procedures: interval arithmetic *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal arithmetic procedures: interval arithmetic *) +(* -------------------------------------------------------------------------- *) (* Dependencies *) -needs "informal/informal_float.hl";; +needs "Formal_ineqs/informal/informal_float.hl";; (* Interval arithmetic *) @@ -107,7 +113,7 @@ let sqrt_interval pp (lo,hi) = (sqrt_float_lo pp lo, sqrt_float_hi pp hi);; (* mul *) -let mul_interval pp (l_lo,l_hi) (r_lo,r_hi) = +let mul_interval pp (l_lo,l_hi) (r_lo,r_hi) = let s1 = sign_float l_lo and s2 = sign_float l_hi and s3 = sign_float r_lo and @@ -148,7 +154,7 @@ let mul_interval pp (l_lo,l_hi) (r_lo,r_hi) = (mul_float_lo pp lo1 lo2, mul_float_hi pp hi1 hi2);; (* div *) -let div_interval pp (l_lo,l_hi) (r_lo,r_hi) = +let div_interval pp (l_lo,l_hi) (r_lo,r_hi) = let s1 = sign_float l_lo and s2 = sign_float l_hi and s3 = sign_float r_lo and @@ -216,7 +222,7 @@ let le_interval x (lo, hi) = le_float x lo;; let ge_interval x (lo, hi) = le_float hi x;; (* compare_interval *) -let compare_interval x (lo, hi) = +let compare_interval x (lo, hi) = if le_float x lo then -1 else if le_float hi x then 1 else 0;; @@ -235,7 +241,7 @@ open Informal_interval;; #install_printer print_float;; #install_printer print_interval;; - + let pp = 3;; let n = 41;; let mk_test a b = mk_interval (mk_float a 0, mk_float b 0);; diff --git a/Formal_ineqs/informal/informal_log.hl b/Formal_ineqs/informal/informal_log.hl index 52789ea2..10b7d879 100644 --- a/Formal_ineqs/informal/informal_log.hl +++ b/Formal_ineqs/informal/informal_log.hl @@ -1,12 +1,18 @@ -(* =========================================================== *) -(* Informal arithmetic procedures: log *) -(* Author: Alexey Solovyev *) -(* Date: 2014-12-12 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal arithmetic procedures: log *) +(* -------------------------------------------------------------------------- *) (* Dependencies *) -needs "informal/informal_interval.hl";; -needs "informal/informal_exp.hl";; +needs "Formal_ineqs/informal/informal_interval.hl";; +needs "Formal_ineqs/informal/informal_exp.hl";; module type Informal_log_sig = sig @@ -55,7 +61,7 @@ end;; (* -needs "trig/log_eval.hl";; +needs "Formal_ineqs/trig/log_eval.hl";; open Informal_float;; open Informal_interval;; diff --git a/Formal_ineqs/informal/informal_matan.hl b/Formal_ineqs/informal/informal_matan.hl index 0d173902..14ed607a 100644 --- a/Formal_ineqs/informal/informal_matan.hl +++ b/Formal_ineqs/informal/informal_matan.hl @@ -1,11 +1,17 @@ -(* =========================================================== *) -(* Informal arithmetic procedures: matan *) -(* Author: Alexey Solovyev *) -(* Date: 2014-07-29 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal arithmetic procedures: matan *) +(* -------------------------------------------------------------------------- *) (* Dependencies *) -needs "informal/informal_atn.hl";; +needs "Formal_ineqs/informal/informal_atn.hl";; (* matan *) module type Informal_matan_sig = @@ -117,7 +123,7 @@ let dmatan_interval = let ddmatan_interval = let v = mk_float_interval_decimal 20 `#0.65` in let v0 = mk_small_num_interval 1 in - let vs = Array.init 20 + let vs = Array.init 20 (fun i -> if i = 0 then v0 else round_interval i v) in fun pp x -> let lo, _ = dest_interval x in @@ -154,7 +160,7 @@ let f = Eval_interval.build_interval_fun `&1 + &3`;; Eval_interval.eval_interval_fun 3 f [] [];; let mk_intervals pp a b = - let inf_mk tm = + let inf_mk tm = let f = Informal_eval_interval.build_interval_fun tm in Informal_eval_interval.eval_interval_fun pp f [] [] in let mk tm = @@ -200,7 +206,7 @@ let ddcheck pp (x, x_th) = ddcheck 3 (List.nth test_ints2 4);; - - + + *) diff --git a/Formal_ineqs/informal/informal_nat.hl b/Formal_ineqs/informal/informal_nat.hl index 0c6a967b..a30d1871 100644 --- a/Formal_ineqs/informal/informal_nat.hl +++ b/Formal_ineqs/informal/informal_nat.hl @@ -1,13 +1,19 @@ -(* =========================================================== *) -(* Informal arithmetic procedures: natural numbers *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal arithmetic procedures: natural numbers *) +(* -------------------------------------------------------------------------- *) (* Dependencies *) -needs "misc/misc_functions.hl";; -needs "misc/misc_vars.hl";; -needs "arith_options.hl";; +needs "Formal_ineqs/misc/misc_functions.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; +needs "Formal_ineqs/arith_options.hl";; (* Natural numbers *) @@ -56,11 +62,11 @@ type nat = big_int;; let arith_base = !Arith_options.base;; -let mk_nat n = +let mk_nat n = let result = big_int_of_num n in if sign_big_int result < 0 then zero_big_int else result;; -let mk_small_nat n = +let mk_small_nat n = if n < 0 then zero_big_int else big_int_of_int n;; let dest_nat = num_of_big_int;; @@ -69,7 +75,7 @@ let eq_nat = eq_big_int;; let suc_nat = succ_big_int;; -let pre_nat n = +let pre_nat n = let result = pred_big_int n in if sign_big_int result < 0 then zero_big_int else result;; @@ -114,7 +120,7 @@ let normalize_nat = (n, e) else normalize q (succ e) in - fun n -> + fun n -> if sign_big_int n = 0 then (n, 0) else normalize n 0;; @@ -125,7 +131,7 @@ let denormalize_nat (n, e) = let lo_nat pp = let max = power_int_positive_int arith_base pp in let rec lo m e = - if lt_big_int m max then + if lt_big_int m max then (m, e) else let q = div_big_int m base_nat in diff --git a/Formal_ineqs/informal/informal_poly.hl b/Formal_ineqs/informal/informal_poly.hl index 276afb7e..90f5d27d 100644 --- a/Formal_ineqs/informal/informal_poly.hl +++ b/Formal_ineqs/informal/informal_poly.hl @@ -1,11 +1,17 @@ -(* =========================================================== *) -(* Informal arithmetic procedures: polynomials *) -(* Author: Alexey Solovyev *) -(* Date: 2014-07-16 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal arithmetic procedures: polynomials *) +(* -------------------------------------------------------------------------- *) (* Dependencies *) -needs "informal/informal_interval.hl";; +needs "Formal_ineqs/informal/informal_interval.hl";; (* poly *) module type Informal_poly_sig = @@ -35,10 +41,10 @@ let rec eval_interval_poly_f pp cs x_int = let r = eval_interval_poly_f pp rest x_int in add_interval pp first (mul_interval pp x_int r);; -let rec eval_high_poly_f_pos_pos = +let rec eval_high_poly_f_pos_pos = let zero = mk_small_num_float 0 in - let check_pos c = - if le_interval zero c then () + let check_pos c = + if le_interval zero c then () else failwith "Informal_poly.eval_high_poly_f_pos_pos: negative coefficient" in fun pp cs x -> let _ = check_pos (mk_interval (x, x)) in @@ -59,14 +65,14 @@ let rec eval_high_poly_f_pos_pos = let rec eval_low_poly_f_pos_pos = let zero = mk_small_num_float 0 in let check_pos c = - if le_interval zero c then () + if le_interval zero c then () else failwith "Informal_poly.eval_high_poly_f_pos_pos: negative coefficient" in fun pp cs x -> let _ = check_pos (mk_interval (x, x)) in let rec eval cs = match cs with | [] -> zero - | [first] -> + | [first] -> let _ = check_pos first in fst (dest_interval first) | first :: rest -> diff --git a/Formal_ineqs/informal/informal_search.hl b/Formal_ineqs/informal/informal_search.hl index dab17fb4..47baad3b 100644 --- a/Formal_ineqs/informal/informal_search.hl +++ b/Formal_ineqs/informal/informal_search.hl @@ -1,12 +1,18 @@ -(* =========================================================== *) -(* Informal certificate search *) -(* Author: Alexey Solovyev *) -(* Date: 2014-07-31 *) -(* =========================================================== *) - -needs "informal/informal_taylor.hl";; -needs "verifier/certificate.hl";; -needs "verifier_options.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal certificate search *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/informal/informal_taylor.hl";; +needs "Formal_ineqs/verifier/certificate.hl";; +needs "Formal_ineqs/verifier_options.hl";; module Informal_search = struct @@ -39,7 +45,7 @@ let find_max s = find (hd s) 0 1 (tl s);; (* split_domain *) -let split_domain pp j domain = +let split_domain pp j domain = let n = length domain.w in let t = nth domain.y (j - 1) in let vv = map (fun i -> if i = j then t else nth domain.hi (i - 1)) (1--n) in @@ -64,7 +70,7 @@ type function_info = { index : int; };; -type result = +type result = | Cell_inconclusive | Cell_false | Cell_result of result_tree @@ -133,7 +139,7 @@ let rec check_cell opt dom = let lower = eval_m_taylor_lower_bound pp ti in if ge0_float lower then check flag rest - else + else (if opt.mono_depth > 0 then try try_mono {opt with max_depth = opt.mono_depth} dom f ti @@ -186,10 +192,10 @@ and construct_certificate0 opt = else let result = check_cell opt dom fs in match result with - | Cell_false -> + | Cell_false -> error "False result" dom fs | Cell_result r -> r - | Cell_pass (raw_flag, f) -> + | Cell_pass (raw_flag, f) -> let _ = update_verified_vol dom in Result_pass (f.index, raw_flag) | Cell_inconclusive -> @@ -207,7 +213,7 @@ and construct_certificate0 opt = let construct_certificate opt dom fs_informal = let mk_f (f0, ft) i = { eval0 = f0; - taylor = ft; + taylor = ft; index = i } in let fs = map2 mk_f fs_informal (0--(length fs_informal - 1)) in @@ -220,8 +226,8 @@ end;; (* (*******************) -needs "informal/informal_taylor.hl";; -needs "verifier/m_verifier_build.hl";; +needs "Formal_ineqs/informal/informal_taylor.hl";; +needs "Formal_ineqs/verifier/m_verifier_build.hl";; open M_verifier_build;; #install_printer Informal_float.print_ifloat;; @@ -243,13 +249,13 @@ let opt = { let x_list = mk_list ([mk_float (-10) 0; mk_float (-10) 0; mk_float (-10) 0], `:real`);; let z_list = mk_list ([mk_float 10 0; mk_float 10 0; mk_float 10 0], `:real`);; -let x_inf = [Informal_float.mk_float (num (-10)) 0; Informal_float.mk_float (num (-10)) 0; Informal_float.mk_float (num (-10)) 0];; -let z_inf = [Informal_float.mk_float (num 10) 0; Informal_float.mk_float (num 10) 0; Informal_float.mk_float (num 10) 0];; +let x_inf = [Informal_float.mk_float (Num.num_of_int (-10)) 0; Informal_float.mk_float (Num.num_of_int (-10)) 0; Informal_float.mk_float (Num.num_of_int (-10)) 0];; +let z_inf = [Informal_float.mk_float (Num.num_of_int 10) 0; Informal_float.mk_float (Num.num_of_int 10) 0; Informal_float.mk_float (Num.num_of_int 10) 0];; let dom_th = mk_m_center_domain 3 5 x_list z_list;; let dom_inf = Informal_taylor.mk_m_center_domain 5 x_inf z_inf;; -let schwefel_poly = `\x:real^3. -- #0.0000001 - ((x$1 - x$2 pow 2) pow 2 + (x$2 - &1) pow 2 + +let schwefel_poly = `\x:real^3. -- #0.0000001 - ((x$1 - x$2 pow 2) pow 2 + (x$2 - &1) pow 2 + (x$1 - x$3 pow 2) pow 2 + (x$3 - &1) pow 2)`;; let f1, f2 = mk_verification_functions_poly 5 schwefel_poly;; @@ -268,8 +274,8 @@ result_stats c;; (***) -let x_inf = [Informal_float.mk_float (num 1) 0];; -let y_inf = [Informal_float.mk_float (num 3) 0];; +let x_inf = [Informal_float.mk_float (Num.num_of_int 1) 0];; +let y_inf = [Informal_float.mk_float (Num.num_of_int 3) 0];; let dom_inf = Informal_taylor.mk_m_center_domain 5 x_inf y_inf;; diff --git a/Formal_ineqs/informal/informal_sin_cos.hl b/Formal_ineqs/informal/informal_sin_cos.hl index cfd05e7c..20e67b8b 100644 --- a/Formal_ineqs/informal/informal_sin_cos.hl +++ b/Formal_ineqs/informal/informal_sin_cos.hl @@ -1,12 +1,18 @@ -(* =========================================================== *) -(* Informal arithmetic procedures: sin, cos *) -(* Author: Alexey Solovyev *) -(* Date: 2014-07-16 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal arithmetic procedures: sin, cos *) +(* -------------------------------------------------------------------------- *) (* Dependencies *) -needs "informal/informal_poly.hl";; -needs "informal/informal_atn.hl";; +needs "Formal_ineqs/informal/informal_poly.hl";; +needs "Formal_ineqs/informal/informal_atn.hl";; (* sin, cos *) module type Informal_sin_cos_sig = @@ -24,15 +30,15 @@ open Informal_poly;; open Informal_atn;; let rec fact n = - if sign_num n <= 0 then num 1 else - n */ fact (pred_num n);; + if sign_num n <= 0 then Num.num_of_int 1 else + n */ fact (n -/ num_1);; let one = mk_small_num_float 1 and neg_one = neg_float (mk_small_num_float 1);; let cos_tables pp n = - let t1 = map (fun i -> fact (num (4 * i))) (0 -- (n / 2)) and - t2 = map (fun i -> fact (num (4 * i + 2))) (0 -- ((n + 1) / 2 - 1)) in + let t1 = map (fun i -> fact (Num.num_of_int (4 * i))) (0 -- (n / 2)) and + t2 = map (fun i -> fact (Num.num_of_int (4 * i + 2))) (0 -- ((n + 1) / 2 - 1)) in let f n = inv_interval pp (mk_num_interval n) in map f t1, map f t2;; @@ -42,13 +48,13 @@ let rec x_pow_over_fact x k = x /. (float_of_int k) *. x_pow_over_fact x (k - 1);; (* Computes i such that x^(2(i + 1))/(2(i + 1))! <= base^(-(p + 1)) and cond(i) *) -let n_of_p_cos x pp cond = +let n_of_p_cos x pp cond = let t = (float_of_int Informal_nat.arith_base) ** (float_of_int (-pp - 1)) in let rec try_i i = let _ = if i > 50 then failwith "n_of_p_cos: cannot find i" else () in if cond i then let r = x_pow_over_fact x (2 * (i + 1)) in - if r <= t then i else try_i (i + 1) + if r <= t then i else try_i (i + 1) else try_i (i + 1) in @@ -107,7 +113,7 @@ let pi32_array = Array.init (n - 1) (fun i -> if i = 0 then zero_interval else round_interval i pi32);; (* -pi *) -let neg_pi_array = +let neg_pi_array = let n = Array.length pi_approx_array in Array.init n (fun i -> neg_interval pi_approx_array.(i));; @@ -151,7 +157,7 @@ let eval_high_neg_pi_0 pp x = let eval_low_neg_pi_0 pp x = let y = neg_float x in eval_low_0_pi pp y;; - + let eval_high_pi_2pi pp x = let pi32_lo = fst (dest_interval pi32_array.(pp)) in if le_float x pi32_lo then @@ -207,7 +213,7 @@ let get_i = let k0 = -int_of_float (x /. f_2_pi) in let y = x +. float_of_int k0 *. f_2_pi in if y < -.f_pi then k0 + 1 - else if y > f_pi then k0 - 1 + else if y > f_pi then k0 - 1 else k0;; let reduction_zero a b = a, b;; @@ -331,7 +337,7 @@ let cos_interval pp x_int = mk_interval (low, one) else mk_interval (neg_one, one) - with Correction_failed -> + with Correction_failed -> let _ = warn true (Printf.sprintf "cos_interval: reduction failed") in mk_interval (neg_one, one);; @@ -366,17 +372,17 @@ let n = -10000;; let a = 3 and b = 5;; -let x = mk_interval (mk_float (num a) 0, mk_float (num b) 0);; -let y = mk_interval (mk_float (num (-a)) 0, mk_float (num b) 0);; -let z = mk_interval (mk_float (num (-b)) 0, mk_float (num a) 0);; -let u = mk_interval (mk_float (num (-3)) 0, mk_float (num (-2)) 0);; +let x = mk_interval (mk_float (Num.num_of_int a) 0, mk_float (Num.num_of_int b) 0);; +let y = mk_interval (mk_float (Num.num_of_int (-a)) 0, mk_float (Num.num_of_int b) 0);; +let z = mk_interval (mk_float (Num.num_of_int (-b)) 0, mk_float (Num.num_of_int a) 0);; +let u = mk_interval (mk_float (Num.num_of_int (-3)) 0, mk_float (Num.num_of_int (-2)) 0);; abs_interval x;; abs_interval y;; abs_interval z;; abs_interval u;; -let x = mk_interval (mk_float (num a) 0, mk_float (num b) 0);; +let x = mk_interval (mk_float (Num.num_of_int a) 0, mk_float (Num.num_of_int b) 0);; let x_th = (ASSUME o Interval_arith.mk_interval `x:real` o mk_pair) (More_float.mk_float a 0, More_float.mk_float b 0);; diff --git a/Formal_ineqs/informal/informal_taylor.hl b/Formal_ineqs/informal/informal_taylor.hl index cbccb318..bc70765e 100644 --- a/Formal_ineqs/informal/informal_taylor.hl +++ b/Formal_ineqs/informal/informal_taylor.hl @@ -1,12 +1,18 @@ -(* =========================================================== *) -(* Informal taylor intervals *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal taylor intervals *) +(* -------------------------------------------------------------------------- *) (* Dependencies *) -needs "misc/misc_functions.hl";; -needs "informal/informal_eval_interval.hl";; +needs "Formal_ineqs/misc/misc_functions.hl";; +needs "Formal_ineqs/informal/informal_eval_interval.hl";; module Informal_taylor = struct @@ -23,7 +29,7 @@ open Informal_matan;; open Informal_eval_interval;; -type m_cell_domain = +type m_cell_domain = { lo : ifloat list; hi : ifloat list; @@ -38,7 +44,7 @@ type m_taylor_interval = domain : m_cell_domain; f : interval; df : interval list; - ddf : interval list list; + ddf : interval list list; };; @@ -83,7 +89,7 @@ let eval_m_taylor pp0 f_tm partials partials2 = let f = build f_tm in let n = length partials in (* Verify that the list of second partial derivatives is correct *) - let _ = map2 (fun i list -> if length list <> i then + let _ = map2 (fun i list -> if length list <> i then failwith "eval_m_taylor: incorrect partials2" else ()) (1--n) partials2 in let dfs = map (build o rand o concl) partials in let d2fs = map (build o rand o concl) (List.flatten partials2) in @@ -129,7 +135,7 @@ let eval_m_taylor_error pp ti = let sums2 = (hd o hd) mul_wdd :: map2 (fun list t1 -> last list + float_2 * t1) (tl mul_wdd) sums1 in let sums = map2 ( * ) w sums2 in end_itlist ( + ) sums;; - + (* eval_m_taylor_upper_bound *) let eval_m_taylor_upper_bound pp ti = @@ -172,10 +178,10 @@ let eval_m_taylor_bound pp ti = let eval_m_taylor_partial_upper pp i ti = let df_hi = (snd o dest_interval o List.nth ti.df) (i - 1) in let dd_list = map (fun j -> if j <= i then - List.nth (List.nth ti.ddf (i - 1)) (j - 1) + List.nth (List.nth ti.ddf (i - 1)) (j - 1) else List.nth (List.nth ti.ddf (j - 1)) (i - 1)) (1--ti.n) in - let sum2 = + let sum2 = let mul_dd = map2 (error_mul_f2_hi pp) ti.domain.w dd_list in end_itlist (add_float_hi pp) mul_dd in add_float_hi pp df_hi sum2;; @@ -185,10 +191,10 @@ let eval_m_taylor_partial_upper pp i ti = let eval_m_taylor_partial_lower pp i ti = let df_lo = (fst o dest_interval o List.nth ti.df) (i - 1) in let dd_list = map (fun j -> if j <= i then - List.nth (List.nth ti.ddf (i - 1)) (j - 1) + List.nth (List.nth ti.ddf (i - 1)) (j - 1) else List.nth (List.nth ti.ddf (j - 1)) (i - 1)) (1--ti.n) in - let sum2 = + let sum2 = let mul_dd = map2 (error_mul_f2_hi pp) ti.domain.w dd_list in end_itlist (add_float_hi pp) mul_dd in sub_float_lo pp df_lo sum2;; @@ -198,10 +204,10 @@ let eval_m_taylor_partial_lower pp i ti = let eval_m_taylor_partial_bound pp i ti = let df_lo, df_hi = (dest_interval o List.nth ti.df) (i - 1) in let dd_list = map (fun j -> if j <= i then - List.nth (List.nth ti.ddf (i - 1)) (j - 1) + List.nth (List.nth ti.ddf (i - 1)) (j - 1) else List.nth (List.nth ti.ddf (j - 1)) (i - 1)) (1--ti.n) in - let sum2 = + let sum2 = let mul_dd = map2 (error_mul_f2_hi pp) ti.domain.w dd_list in end_itlist (add_float_hi pp) mul_dd in let lo = sub_float_lo pp df_lo sum2 in @@ -231,7 +237,7 @@ let eval_m_taylor_sub p_lin p_second taylor1 taylor2 = df = map2 (-) taylor1.df taylor2.df; ddf = map2 (map2 (--)) taylor1.ddf taylor2.ddf };; - + (* mul *) let eval_m_taylor_mul p_lin p_second ti1 ti2 = @@ -245,7 +251,7 @@ let eval_m_taylor_mul p_lin p_second ti1 ti2 = let d2_bounds = map (fun i -> eval_m_taylor_partial_bound p_second i ti2) ns in let f1_bound = eval_m_taylor_bound p_second ti1 in let f2_bound = eval_m_taylor_bound p_second ti2 in - let ddf = + let ddf = let ( * ), ( + ) = mul_interval p_second, add_interval p_second in map2 (fun (list1, list2) i -> let di1 = List.nth d1_bounds (i - 1) in @@ -295,7 +301,7 @@ let eval_m_taylor_pow2 p_lin p_second ti = let ( * ) = mul_interval p_lin in map (fun d -> u_bounds * d) ti.df in let d1_bounds = map (fun i -> eval_m_taylor_partial_bound p_second i ti) ns in - let ddf = + let ddf = let ( * ), ( + ) = mul_interval p_second, add_interval p_second in map2 (fun dd_list di1 -> my_map2 (fun dd dj1 -> two_interval * (dj1 * di1 + f1_bound * dd)) @@ -320,7 +326,7 @@ let eval_m_taylor_pow k = let ns = 1--n in let f1_bound = eval_m_taylor_bound p_second ti in let bounds_pow_k1 = pow_interval p_lin (k - 1) ti.f in - let bounds = + let bounds = let ( * ) = mul_interval p_lin in ti.f * bounds_pow_k1 in let u_bounds = @@ -334,7 +340,7 @@ let eval_m_taylor_pow k = let ( * ) = mul_interval p_second in let pow_k2 = pow_interval p_second (k - 2) in k_interval * pow_k2 f1_bound in - let ddf = + let ddf = let ( * ), ( + ) = mul_interval p_second, add_interval p_second in map2 (fun dd_list di1 -> my_map2 (fun dd dj1 -> d2 * ((k1_interval * dj1) * di1 + f1_bound * dd)) @@ -364,7 +370,7 @@ let eval_m_taylor_inv p_lin p_second ti = let inv, ( * ) = inv_interval p_second, mul_interval p_second in let ff = f1_bound * f1_bound in inv ff, two_interval * inv (f1_bound * ff) in - let ddf = + let ddf = let ( * ), ( - ) = mul_interval p_second, sub_interval p_second in map2 (fun dd_list di1 -> my_map2 (fun dd dj1 -> @@ -392,11 +398,11 @@ let eval_m_taylor_sqrt p_lin p_second ti = map (fun d -> u_bounds * d) ti.df in let d1_bounds = map (fun i -> eval_m_taylor_partial_bound p_second i ti) ns in let d1, d2 = - let neg, sqrt, inv, ( * ) = neg_interval, sqrt_interval p_second, + let neg, sqrt, inv, ( * ) = neg_interval, sqrt_interval p_second, inv_interval p_second, mul_interval p_second in let two_sqrt_f = two_interval * sqrt f1_bound in inv two_sqrt_f, neg (inv (two_sqrt_f * (two_interval * f1_bound))) in - let ddf = + let ddf = let ( * ), ( + ) = mul_interval p_second, add_interval p_second in map2 (fun dd_list di1 -> my_map2 (fun dd dj1 -> @@ -426,7 +432,7 @@ let eval_m_taylor_exp p_lin p_second ti = let exp = exp_interval p_second in let exp_f = exp f1_bound in exp_f, exp_f in - let ddf = + let ddf = let ( * ), ( + ) = mul_interval p_second, add_interval p_second in map2 (fun dd_list di1 -> my_map2 (fun dd dj1 -> @@ -453,11 +459,11 @@ let eval_m_taylor_log p_lin p_second ti = map (fun d -> u_bounds * d) ti.df in let d1_bounds = map (fun i -> eval_m_taylor_partial_bound p_second i ti) ns in let d1, d2 = - let neg, pow2, inv = + let neg, pow2, inv = neg_interval, pow_interval p_second 2, inv_interval p_second in let inv_f = inv f1_bound in inv_f, neg (pow2 inv_f) in - let ddf = + let ddf = let ( * ), ( + ) = mul_interval p_second, add_interval p_second in map2 (fun dd_list di1 -> my_map2 (fun dd dj1 -> @@ -487,7 +493,7 @@ let eval_m_taylor_matan p_lin p_second ti = let d1, d2 = let dmatan, ddmatan = dmatan_interval p_second, ddmatan_interval p_second in dmatan f1_bound, ddmatan f1_bound in - let ddf = + let ddf = let ( * ), ( + ) = mul_interval p_second, add_interval p_second in map2 (fun dd_list di1 -> my_map2 (fun dd dj1 -> @@ -517,12 +523,12 @@ let eval_m_taylor_atn = map (fun d -> u_bounds * d) ti.df in let d1_bounds = map (fun i -> eval_m_taylor_partial_bound p_second i ti) ns in let d1, d2 = - let neg, inv, ( + ), ( * ) = neg_interval, inv_interval p_second, + let neg, inv, ( + ), ( * ) = neg_interval, inv_interval p_second, add_interval p_second, mul_interval p_second in let pow2 = pow_interval p_second 2 in let inv_one_ff = inv (one_interval + f1_bound * f1_bound) in inv_one_ff, (neg_two_interval * f1_bound) * pow2 inv_one_ff in - let ddf = + let ddf = let ( * ), ( + ) = mul_interval p_second, add_interval p_second in map2 (fun dd_list di1 -> my_map2 (fun dd dj1 -> @@ -551,7 +557,7 @@ let eval_m_taylor_cos p_lin p_second ti = let d1, d2 = let cos, sin = cos_interval p_second, sin_interval p_second in sin f1_bound, cos f1_bound in - let ddf = + let ddf = let ( * ), ( + ) = mul_interval p_second, add_interval p_second in let neg = neg_interval in map2 (fun dd_list di1 -> @@ -582,7 +588,7 @@ let eval_m_taylor_sin p_lin p_second ti = let cos, sin = cos_interval p_second, sin_interval p_second in let neg = neg_interval in cos f1_bound, neg (sin f1_bound) in - let ddf = + let ddf = let ( * ), ( + ) = mul_interval p_second, add_interval p_second in map2 (fun dd_list di1 -> my_map2 (fun dd dj1 -> @@ -616,7 +622,7 @@ let eval_m_taylor_asn p_lin p_second ti = let pow3 = pow_interval p_second 3 in let ff_1 = one_interval - f1_bound * f1_bound in inv (sqrt ff_1), f1_bound / sqrt (pow3 ff_1) in - let ddf = + let ddf = let ( * ), ( + ) = mul_interval p_second, add_interval p_second in map2 (fun dd_list di1 -> my_map2 (fun dd dj1 -> @@ -650,7 +656,7 @@ let eval_m_taylor_acs p_lin p_second ti = let pow3 = pow_interval p_second 3 in let ff_1 = one_interval - f1_bound * f1_bound in inv (sqrt ff_1), neg (f1_bound / sqrt (pow3 ff_1)) in - let ddf = + let ddf = let ( * ), ( - ) = mul_interval p_second, sub_interval p_second in map2 (fun dd_list di1 -> my_map2 (fun dd dj1 -> @@ -662,6 +668,6 @@ let eval_m_taylor_acs p_lin p_second ti = df = df; ddf = ddf; };; - + end;; diff --git a/Formal_ineqs/informal/informal_verifier.hl b/Formal_ineqs/informal/informal_verifier.hl index 6feb8bf5..79c47308 100644 --- a/Formal_ineqs/informal/informal_verifier.hl +++ b/Formal_ineqs/informal/informal_verifier.hl @@ -1,13 +1,19 @@ -(* =========================================================== *) -(* Informal verification procedures *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Informal verification procedures *) +(* -------------------------------------------------------------------------- *) (* Dependencies *) -needs "informal/informal_taylor.hl";; -needs "verifier/certificate.hl";; -needs "verifier_options.hl";; +needs "Formal_ineqs/informal/informal_taylor.hl";; +needs "Formal_ineqs/verifier/certificate.hl";; +needs "Formal_ineqs/verifier_options.hl";; module Informal_verifier = struct @@ -77,7 +83,7 @@ let mk_verification_functions_poly pp0 f partials partials2 = let eval0 = mk_eval_function pp0 f in let eval1 = map (fun i -> mk_eval_function pp0 ((rand o concl o List.nth partials) (i - 1))) (1--n) in let eval2 = map (fun i -> - map (fun j -> + map (fun j -> let d2 = List.nth (List.nth partials2 (i - 1)) (j - 1) in mk_eval_function pp0 ((rand o concl) d2)) (1--i)) (1--n) in { @@ -89,13 +95,13 @@ let mk_verification_functions_poly pp0 f partials partials2 = (* split_domain *) -let split_domain pp j domain = +let split_domain pp j domain = let n = length domain.w in let t = List.nth domain.y (j - 1) in let vv = map (fun i -> if i = j then t else List.nth domain.hi (i - 1)) (1--n) in let uu = map (fun i -> if i = j then t else List.nth domain.lo (i - 1)) (1--n) in mk_m_center_domain pp domain.lo vv, mk_m_center_domain pp uu domain.hi;; - + (* restrict_domain *) let restrict_domain j left_flag domain = @@ -134,24 +140,24 @@ let m_verify_raw (report_start, total_size) p_split p_min p_max fs_list certific let rec find_p p_fun p_list = match p_list with | [] -> failwith "find_p: no good p found" - | p :: ps -> + | p :: ps -> let _ = if !info_print_level >= 2 then report (sprintf "Testing p = %d (other: %d)" p (length ps)) else () in - let flag = - (try p_fun p - with - | Failure msg -> - let _ = if !info_print_level >= 2 then + let flag = + (try p_fun p + with + | Failure msg -> + let _ = if !info_print_level >= 2 then report (sprintf "Failure at p = %d: %s" p msg) else () in - false - | Division_by_zero -> + false + | Division_by_zero -> let _ = if !info_print_level >= 2 then report (sprintf "Failure at p = %d: Division_by_zero" p) else () in false) in - if flag then + if flag then let _ = if !info_print_level >= 2 then report (sprintf "p = %d" p) else () in p else find_p p_fun ps in @@ -188,7 +194,7 @@ let m_verify_raw (report_start, total_size) p_split p_min p_max fs_list certific else map (eval_m_taylor_partial_lower pp m.variable) taylors in let monos = map gen_mono mono in - rev_itlist (fun (m, bounds) pass -> + rev_itlist (fun (m, bounds) pass -> let flag = m.decr_flag in forall (m_mono_pass_gen flag) bounds && pass) (rev (zip mono monos)) true in @@ -205,10 +211,10 @@ let m_verify_raw (report_start, total_size) p_split p_min p_max fs_list certific let rec rec_verify domain certificate = match certificate with | Result_mono (mono, r1) -> - let _ = + let _ = if !info_print_level >= 2 then - let mono_strs = - map (fun m -> sprintf "%s%d (%b)" (if m.decr_flag then "-" else "") + let mono_strs = + map (fun m -> sprintf "%s%d (%b)" (if m.decr_flag then "-" else "") m.variable m.df0_flag) mono in report (sprintf "Mono: [%s]" (String.concat ";" mono_strs)) else () in @@ -219,9 +225,9 @@ let m_verify_raw (report_start, total_size) p_split p_min p_max fs_list certific P_result_mono ({pp = pp}, mono, tree1), inds with Failure _ -> failwith "mono: failed") - | Result_pass (j, f0_flag) -> + | Result_pass (j, f0_flag) -> let _ = k := !k + 1; kk := !kk + 1 in - let _ = + let _ = if !info_print_level >= 2 then report (sprintf "Verifying: %d/%d (f0_flag = %b)" !k r_size f0_flag) else () in @@ -265,7 +271,7 @@ let m_verify_raw (report_start, total_size) p_split p_min p_max fs_list certific P_result_ref i, inds | _ -> failwith "False result" in - + rec_verify domain0 certificate;; @@ -276,11 +282,11 @@ let m_verify_raw (report_start, total_size) p_split p_min p_max fs_list certific let m_verify_raw0 p_split p_min p_max fs_list certificate xx zz = m_verify_raw (0, 0) p_split p_min p_max fs_list certificate (mk_m_center_domain p_split xx zz) [];; - + (* m_verify_list *) let m_verify_list p_split p_min p_max fs_list certificate_list xx zz = let domain_hash = Hashtbl.create (length certificate_list * 10) in - let mem, find, add = Hashtbl.mem domain_hash, + let mem, find, add = Hashtbl.mem domain_hash, Hashtbl.find domain_hash, Hashtbl.add domain_hash in let get_m_cell_domain pp domain0 path = @@ -289,7 +295,7 @@ let m_verify_list p_split p_min p_max fs_list certificate_list xx zz = | [] -> domain | (s, j) :: ps -> let hash' = hash^s^(string_of_int j) in - if mem hash' then + if mem hash' then get_rec (find hash') ps hash' else if s = "l" || s = "r" then @@ -313,7 +319,7 @@ let m_verify_list p_split p_min p_max fs_list certificate_list xx zz = let k = ref 0 in let kk = ref 0 in let total_size = end_itlist (+) (map (result_size o snd) certificate_list) in - + let rec rec_verify certificate_list dom_list tree_list = match certificate_list with | [] -> rev tree_list diff --git a/Formal_ineqs/lib/ipow.hl b/Formal_ineqs/lib/ipow.hl index f6ac65f1..c365f351 100644 --- a/Formal_ineqs/lib/ipow.hl +++ b/Formal_ineqs/lib/ipow.hl @@ -1,7 +1,7 @@ (* ========================================================================== *) (* Formal verification of FPTaylor certificates *) (* *) -(* Author: Alexey Solovyev, University of Utah *) +(* Copyright (c) 2015 Alexey Solovyev, University of Utah *) (* *) (* This file is distributed under the terms of the MIT licence *) (* ========================================================================== *) @@ -18,7 +18,7 @@ prioritize_real();; unparse_as_infix("ipow");; let ipow = new_definition - `ipow (x:real) (e:int) = + `ipow (x:real) (e:int) = (if (&0 <= e) then (x pow (num_of_int e)) else (inv (x pow (num_of_int (--e)))))`;; @@ -57,7 +57,7 @@ let IPOW_INV = prove ASM_REWRITE_TAC[IPOW_0; INT_NEG_0; REAL_INV_1]; ALL_TAC ] THEN - REWRITE_TAC[ipow] THEN REPEAT COND_CASES_TAC THEN + REWRITE_TAC[ipow] THEN REPEAT COND_CASES_TAC THEN REWRITE_TAC[REAL_INV_POW; REAL_INV_INV] THEN ASM_ARITH_TAC);; let INV_IPOW = prove @@ -71,7 +71,7 @@ let IPOW_NEG = prove ASM_REWRITE_TAC[IPOW_0; INT_NEG_0; REAL_INV_1]; ALL_TAC ] THEN - REWRITE_TAC[ipow] THEN REPEAT COND_CASES_TAC THEN + REWRITE_TAC[ipow] THEN REPEAT COND_CASES_TAC THEN ASM_REWRITE_TAC[INT_NEG_NEG; REAL_INV_INV] THEN ASM_ARITH_TAC);; let IPOW_NEG_NUM = prove @@ -126,7 +126,7 @@ let IPOW_ADD = prove let IPOW_MUL = prove (`!x y u. (x * y) ipow u = x ipow u * y ipow u`, - REPEAT GEN_TAC THEN REWRITE_TAC[ipow] THEN + REPEAT GEN_TAC THEN REWRITE_TAC[ipow] THEN COND_CASES_TAC THEN REWRITE_TAC[REAL_POW_MUL; REAL_INV_MUL]);; let IPOW_ZERO = prove @@ -150,9 +150,9 @@ let IPOW_ONE = prove let IPOW_IPOW = prove (`!x u v. (x ipow u) ipow v = x ipow (u * v)`, REPEAT STRIP_TAC THEN ASM_CASES_TAC `x = &0` THENL [ - ASM_REWRITE_TAC[IPOW_ZERO] THEN + ASM_REWRITE_TAC[IPOW_ZERO] THEN COND_CASES_TAC THEN ASM_REWRITE_TAC[INT_MUL_LZERO; IPOW_ONE] THEN - REWRITE_TAC[IPOW_ZERO] THEN + REWRITE_TAC[IPOW_ZERO] THEN COND_CASES_TAC THEN ASM_REWRITE_TAC[INT_ENTIRE; INT_MUL_RZERO]; ALL_TAC ] THEN @@ -349,9 +349,9 @@ let IPOW_EQ_1_IMP = prove ]);; let IPOW_OF_NUM = prove - (`!x i. &x ipow i = + (`!x i. &x ipow i = if &0 <= i then &(x EXP (num_of_int i)) else inv (&(x EXP (num_of_int (--i))))`, - REPEAT STRIP_TAC THEN REWRITE_TAC[ipow] THEN + REPEAT STRIP_TAC THEN REWRITE_TAC[ipow] THEN COND_CASES_TAC THEN REWRITE_TAC[REAL_OF_NUM_POW]);; let IPOW_EQ_EXP = prove @@ -365,7 +365,7 @@ let IPOW_EQ_EXP_P = prove REWRITE_TAC[NUM_OF_INT_OF_NUM]);; let IPOW_BETWEEN = prove - (`!(x:real) (y:num) (z:num) (e:int). + (`!(x:real) (y:num) (z:num) (e:int). &0 < x /\ &y * x ipow e <= &z * x ipow e /\ &z * x ipow e <= (&y + &1) * x ipow e ==> z = y \/ z = y + 1`, @@ -438,7 +438,7 @@ let IPOW_EXP_MONO_INV_LT = prove MATCH_MP_TAC IPOW_MONO_INV THEN ASM_ARITH_TAC);; let IPOW_MONOTONE = prove - (`!(x:num) (e1:int) (e2:int). + (`!(x:num) (e1:int) (e2:int). 2 <= x /\ &x ipow e1 <= &x ipow e2 ==> e1 <= e2`, REPEAT STRIP_TAC THEN MATCH_MP_TAC IPOW_EXP_MONO THEN EXISTS_TAC `&x:real` THEN ASM_REWRITE_TAC[REAL_OF_NUM_LT] THEN diff --git a/Formal_ineqs/lib/ssrbool.hl b/Formal_ineqs/lib/ssrbool.hl index b73ad476..dc100d04 100644 --- a/Formal_ineqs/lib/ssrbool.hl +++ b/Formal_ineqs/lib/ssrbool.hl @@ -1,4 +1,4 @@ -needs "lib/ssrfun.hl";; +needs "Formal_ineqs/lib/ssrfun.hl";; (* Module Ssrbool*) module Ssrbool = struct diff --git a/Formal_ineqs/lib/ssreflect/sections.hl b/Formal_ineqs/lib/ssreflect/sections.hl index f2134737..21e1c1c9 100644 --- a/Formal_ineqs/lib/ssreflect/sections.hl +++ b/Formal_ineqs/lib/ssreflect/sections.hl @@ -2,7 +2,7 @@ module Sections = struct (* Basic commands for working with the goal stack *) (* b() from tactics.ml *) -let revert_proof_step() = +let revert_proof_step() = let l = !current_goalstack in if length l = 1 then failwith "Can't back up any more" else current_goalstack := tl l; @@ -13,7 +13,7 @@ let revert_proof_step() = let fast_load_flag = ref false;; (* Section variables, hypotheses (with labels), implicit types, and auxiliary lemmas *) -type section_info = +type section_info = { vars : term list; hyps : (string * term) list; @@ -109,8 +109,8 @@ let inst_section_vars tm = let inst_var (name, ty) tm = let ty_dst, ty_src = find_var (name, ty) in try (inst (type_match ty_src ty_dst []) tm) - with Failure _ -> - failwith ("Section variable " ^ name ^ + with Failure _ -> + failwith ("Section variable " ^ name ^ " has type " ^ string_of_type ty_dst) in let f_vars = map dest_var (frees tm) in itlist inst_var f_vars tm;; @@ -118,7 +118,7 @@ let inst_section_vars tm = (* Instantiates implicit types in the given term *) (* (free variables and top generalized variables are considered in the term) *) -let inst_section_types tm = +let inst_section_types tm = let s_types = section_types() in let find_type tm = let name, ty = dest_var tm in @@ -130,7 +130,7 @@ let inst_section_types tm = inst ty_inst tm;; -(* Checks if the term contains any free variables +(* Checks if the term contains any free variables which are not section variables *) let check_section_term tm = let f_vars = frees tm in @@ -146,7 +146,7 @@ let check_section_term tm = let str = String.concat ", " (map string_of_term vars) in failwith ("Free variables: " ^ str) else ();; - + (* Adds the given variable to the active section *) let add_section_var var = @@ -269,7 +269,7 @@ let prepare_goal_term tm = (* Prepares a goal term and an initial tactic *) let prepare_section_proof names tm = let f_vars = map dest_var (frees tm) in - let find_type var_name = + let find_type var_name = try assoc var_name f_vars with Failure _ -> failwith ("Unused variable: " ^ var_name) in let g_vars = map (fun name -> mk_var (name, find_type name)) names in let g_tm = list_mk_forall (g_vars, tm) in @@ -312,7 +312,7 @@ let section_proof names tm tac_list = let hyps = section_hyps() in itlist (fun _ th -> UNDISCH th) hyps th0;; - + (* Discharges all assumptions and generalizes all section variables *) let finalize_theorem th = let hyps = map snd (current_section_hyps()) in @@ -323,5 +323,5 @@ let finalize_theorem th = let f_vars = frees (concl th1) in let vars = intersect f_vars s_vars in itlist (fun var th -> GEN var th) vars th1;; - + end;; diff --git a/Formal_ineqs/lib/ssreflect/ssreflect.hl b/Formal_ineqs/lib/ssreflect/ssreflect.hl index 98e174ad..c2076eba 100644 --- a/Formal_ineqs/lib/ssreflect/ssreflect.hl +++ b/Formal_ineqs/lib/ssreflect/ssreflect.hl @@ -82,14 +82,14 @@ let (THENL_FIRST),(THENL_LAST) = fun tac1 tac2 g -> let _,gls,_ as gstate = tac1 g in if gls = [] then failwith "No subgoals" - else + else let tac_list = tac2 :: (replicate ALL_TAC (length gls - 1)) in tacsequence gstate tac_list and (thenl_last: tactic -> tactic -> tactic) = fun tac1 tac2 g -> let _,gls,_ as gstate = tac1 g in if gls = [] then failwith "No subgoals" - else + else let tac_list = (replicate ALL_TAC (length gls - 1)) @ [tac2] in tacsequence gstate tac_list in thenl_first, thenl_last;; @@ -138,7 +138,7 @@ let move labels = ALL_TAC with Failure _ -> ALL_TAC in tac g in - + let move1 name (g:goal) = let g_tm = snd g in let tac = @@ -159,7 +159,7 @@ let move labels = failwith "move: not (!) or (==>)" in tac g in fun g -> - let tac = itlist + let tac = itlist (fun name tac -> move_eq THEN move1 name THEN tac) labels ALL_TAC in tac g;; @@ -171,15 +171,15 @@ let in_tac a_list in_goal tac (g:goal) = let tmp_goal_var = mk_var (tmp_goal_name, bool_ty) in let tmp_goal = mk_eq (tmp_goal_var, goal_tm) in let tmp_goal_sym = mk_eq (goal_tm, tmp_goal_var) in - let disch_tac = + let disch_tac = rev_itlist (fun name tac -> REMOVE_THEN name MP_TAC THEN tac) a_list ALL_TAC in let intro_tac = move a_list in - let hide_goal, unfold_goal = - if in_goal then + let hide_goal, unfold_goal = + if in_goal then ALL_TAC, ALL_TAC else - ABBREV_TAC tmp_goal, - EXPAND_TAC tmp_goal_name THEN + ABBREV_TAC tmp_goal, + EXPAND_TAC tmp_goal_name THEN UNDISCH_TAC tmp_goal_sym THEN DISCH_THEN (fun th -> ALL_TAC) in (hide_goal THEN disch_tac THEN tac THEN TRY intro_tac THEN unfold_goal) g;; @@ -187,7 +187,7 @@ let in_tac a_list in_goal tac (g:goal) = (* Finds a subterm in the given term which matches against the given - pattern; local_consts is a list of variable which must be fixed in + pattern; local_consts is a list of variable which must be fixed in the pattern. This function returns the path to the first matched subterm *) let match_subterm local_consts pat tm = @@ -196,7 +196,7 @@ let match_subterm local_consts pat tm = let inst = term_match local_consts pat tm in if instantiate inst pat = tm then path else failwith "Bad instantiation" with x -> - try + try match tm with | Abs(_, b_tm) -> find b_tm (path^"b") | Comb(l_tm, r_tm) -> @@ -224,12 +224,12 @@ let find_all_paths p tm = (* Instantiates types of the given context variables in the given term.*) -let inst_context_vars vars tm_vars tm = +let inst_context_vars vars tm_vars tm = let find_type var = let name, ty = dest_var var in try (ty, type_of (assoc name vars)) - with Failure _ -> + with Failure _ -> failwith (name^" is free in the term `"^(string_of_term tm)^"` and in the context") in let ty_src, ty_dst = unzip (map find_type tm_vars) in let ty_inst = itlist2 type_match ty_src ty_dst [] in @@ -261,7 +261,7 @@ let match_subterm_in_context pat tm (g : goal) = let rec break_conjuncts th : thm list = (* Convert P ==> (!x. Q x) to !x. P ==> Q x and P ==> Q ==> R to P /\ Q ==> R *) let th0 = PURE_REWRITE_RULE[GSYM RIGHT_FORALL_IMP_THM; IMP_IMP] th in - let th1 = SPEC_ALL th0 in + let th1 = SPEC_ALL th0 in (* Break top level conjunctions *) let th_list = CONJUNCTS th1 in if length th_list > 1 then @@ -280,7 +280,7 @@ let rec break_conjuncts th : thm list = [PURE_ONCE_REWRITE_RULE[TAUT `~P <=> (P <=> F)`] th1] else [EQT_INTRO th1];; - + (* Finds an instantination for the given term inside another term *) let rec find_term_inst local_consts tm src_tm path = @@ -336,7 +336,7 @@ let new_rewrite occ pat th g = (* Free variables in the given theorem will not be matched *) let local_consts = frees (concl th) in (* Apply the pattern *) - let goal_subterm_path = + let goal_subterm_path = if pat = [] then "" else match_subterm_in_context (hd pat) goal_tm g in let goal_subterm = follow_path goal_subterm_path goal_tm in @@ -382,7 +382,7 @@ let new_rewrite occ pat th g = (* Try to rewrite with all given theorems *) let th_list = break_conjuncts th in - let rec my_first th_list = + let rec my_first th_list = if length th_list = 1 then rewrite (hd th_list) g else @@ -436,7 +436,7 @@ let rewrite occ pat th g = let match_fun = (if is_eq eq_tm then lhand else I) o (if cond_flag then rand else I) in (* Apply the pattern *) - let goal_subterm_path = + let goal_subterm_path = if pat = [] then "" else match_subterm_in_context (hd pat) goal_tm g in let goal_subterm = follow_path goal_subterm_path goal_tm in @@ -492,8 +492,8 @@ let split_tac = FIRST [CONJ_TAC; EQ_TAC];; (* Creates an abbreviation for the given term with the given name *) let set_tac name tm (g : goal) = let goal_tm = snd g in - let tm0 = - try + let tm0 = + try follow_path (match_subterm_in_context tm goal_tm g) goal_tm with Failure _ -> tm in let tm1 = inst_all_free_vars tm0 g in @@ -503,7 +503,7 @@ let set_tac name tm (g : goal) = (* Generates a fresh name for the given term *) (* taking into account names of the provided variables *) let generate_fresh_name names tm = - let rec find_name prefix n = + let rec find_name prefix n = let name = prefix ^ (if n = 0 then "" else string_of_int n) in if can (find (fun str -> str = name)) names then find_name prefix (n + 1) @@ -533,12 +533,12 @@ let disch_tm_tac occs tm (g : goal) = let tm0 = prepare_term tm g in let name = generate_fresh_name ((fst o unzip) (get_context_vars g)) tm in let new_tm = mk_var (name, type_of tm0) in - let new_tm1 = - if occs = [] && is_var tm then + let new_tm1 = + if occs = [] && is_var tm then mk_var ((fst o dest_var) tm, type_of tm0) - else new_tm in + else new_tm in let abbrev_tm = mk_eq (new_tm, tm0) in - (ABBREV_TAC abbrev_tm THEN + (ABBREV_TAC abbrev_tm THEN EXPAND_TAC name THEN POP_ASSUM (fun th -> TRY (new_rewrite occs [] th)) THEN SPEC_TAC (new_tm, new_tm1)) g;; @@ -639,14 +639,14 @@ let match_mp_th ith n th = let rec rec_match n list head = match list with | ("undisch", (_ as tm0)) :: t -> - (try + (try let ii = term_match lconsts tm0 tm in if n <= 1 then let th1 = INSTANTIATE_ALL ii th0 in let th2 = PROVE_HYP th th1 in let list0 = head @ (("undisch", `T`) :: t) in let f_vars = frees tm0 in - let list1 = filter + let list1 = filter (fun s, tm -> not (s = "spec" && mem tm f_vars)) list0 in let list = map (fun s, tm -> s, instantiate ii tm) list1 in list, th2 @@ -708,7 +708,7 @@ let combine_args arg1 arg2 = let th1 = get_arg_thm arg1 in let th0 = match arg2 with - | Arg_theorem th2 -> + | Arg_theorem th2 -> (try MATCH_MP th1 th2 with Failure _ -> match_mp_th th1 1 th2) | Arg_term tm2 -> (try ISPEC tm2 th1 with Failure _ -> spec_var_th th1 1 tm2) @@ -722,7 +722,7 @@ let use_arg_then_result = ref TRUTH;; (* (* Tests if the given id defines a theorem *) let test_id_thm id = - let lexbuf = + let lexbuf = Lexing.from_string ("use_arg_then_result := " ^ id ^ ";;") in let ast = (!Toploop.parse_toplevel_phrase) lexbuf in try @@ -745,7 +745,7 @@ let use_arg_then id (arg_tac:arg_type->tactic) (g:goal) = let var = assoc id vars in Arg_term var with Failure _ -> - let lexbuf = + let lexbuf = Lexing.from_string ("use_arg_then_result := " ^ id ^ ";;") in let ast = (!Toploop.parse_toplevel_phrase) lexbuf in let _ = @@ -780,15 +780,15 @@ let combine_args_then (tac:arg_type->tactic) arg1 arg2 (g:goal) = let th1 = get_arg_thm arg1 in let th0 = match arg2 with - | Arg_theorem th2 -> + | Arg_theorem th2 -> (try MATCH_MP th1 th2 with Failure _ -> match_mp_th th1 1 th2) | Arg_term tm2 -> let tm0 = prepare_term tm2 g in (try ISPEC tm0 th1 with Failure _ -> spec_var_th th1 1 tm0) | Arg_type ty2 -> inst_first_type th1 ty2 in tac (Arg_theorem th0) g;; - - + + (* Specializes a variable and applies the next tactic *) @@ -832,23 +832,23 @@ let apply_tac th g = if th = th0 then failwith "apply_tac: no match" else try_match th0 in - + try MATCH_ACCEPT_TAC th g with Failure _ -> try_match th;; -(*let apply_tac th = +(*let apply_tac th = FIRST [MATCH_ACCEPT_TAC th; MATCH_MP_TAC th];; *) (* The 'exact' tactic *) -(* TODO: do [done | by move => top; apply top], here apply top +(* TODO: do [done | by move => top; apply top], here apply top works as ACCEPT_TAC with matching (rewriting) in some cases *) let exact_tac = FIRST [done_tac; DISCH_THEN (fun th -> apply_tac th) THEN done_tac];; (* Specializes the theorem using the given set of variables *) -let spec0 names vars = +let spec0 names vars = let find name = try (assoc name vars, true) with Failure _ -> (parse_term name, false) in @@ -857,7 +857,7 @@ let spec0 names vars = let t, flag = find name in if flag then (ty, type_of t) - else + else (`:bool`, `:bool`) in let inst_term tm = let ty_src, ty_dst = unzip (map find_type (frees tm)) in @@ -877,13 +877,13 @@ let spec_mp names th g = MP_TAC (spec0 names (get_context_vars g) th) g;; (* Case theorems *) let bool_cases = ONCE_REWRITE_RULE[CONJ_ACI] bool_INDUCT;; let list_cases = prove(`!P. P [] /\ (!(h:A) t. P (CONS h t)) ==> (!l. P l)`, - REPEAT STRIP_TAC THEN + REPEAT STRIP_TAC THEN MP_TAC (SPEC `l:(A)list` list_CASES) THEN DISCH_THEN DISJ_CASES_TAC THEN ASM_REWRITE_TAC[] THEN FIRST_X_ASSUM (CHOOSE_THEN MP_TAC) THEN DISCH_THEN (CHOOSE_THEN MP_TAC) THEN DISCH_THEN (fun th -> ASM_REWRITE_TAC[th]));; let pair_cases = pair_INDUCT;; let num_cases = prove(`!P. P 0 /\ (!n. P (SUC n)) ==> (!m. P m)`, - REPEAT STRIP_TAC THEN + REPEAT STRIP_TAC THEN MP_TAC (SPEC `m:num` num_CASES) THEN DISCH_THEN DISJ_CASES_TAC THEN ASM_REWRITE_TAC[] THEN FIRST_X_ASSUM (CHOOSE_THEN (fun th -> ASM_REWRITE_TAC[th])));; let option_cases = option_INDUCT;; @@ -914,7 +914,7 @@ Hashtbl.add elim_table "option" option_elim;; (* case: works only for (A /\ B) -> C; (A \/ B) -> C; (?x. P) -> Q; !(n:num). P; !(l:list(A)). P *) -let case (g:goal) = +let case (g:goal) = let goal_tm = snd g in if not (is_imp goal_tm) then (* !a. P *) @@ -942,7 +942,7 @@ let case (g:goal) = (* elim: works only for num and list *) -let elim (g:goal) = +let elim (g:goal) = let goal_tm = snd g in (* !a. P *) if is_forall goal_tm then @@ -954,7 +954,7 @@ let elim (g:goal) = failwith "elim: not forall";; - + (* Instantiates the first type variable in the given theorem *) let INST_FIRST_TYPE_THEN ty (then_tac:thm_tactic) th = let ty_vars = type_vars_in_term (concl th) in @@ -969,10 +969,10 @@ let transform_pattern pat_tm = let names = ref (map (fst o dest_var) (frees pat_tm)) in let rec transform tm = match tm with - | Abs(x_tm, b_tm) -> + | Abs(x_tm, b_tm) -> let _ = names := (fst o dest_var) x_tm :: !names in mk_abs (x_tm, transform b_tm) - | Comb(l_tm, r_tm) -> + | Comb(l_tm, r_tm) -> mk_comb (transform l_tm, transform r_tm) | Var ("_", ty) -> let name = generate_fresh_name !names tm in @@ -1002,7 +1002,7 @@ let congr_tac pat_tm goal = let lhs, rhs = dest_eq goal_tm in let lm, rm = term_match const_pat pattern lhs, term_match const_pat pattern rhs in - let eq_tms = map + let eq_tms = map (fun tm -> mk_eq (instantiate lm tm, instantiate rm tm)) wild_pat in let eq_tm = itlist (curry mk_imp) eq_tms goal_tm in let eq_thm = EQT_ELIM (SIMP_CONV[] eq_tm) in @@ -1050,7 +1050,7 @@ let disch_tm_eq_tac eq_name occs tm (g : goal) = let eq_var = mk_var (eq_name, aty) in let new_tm = mk_var (name, type_of tm0) in let abbrev_tm = mk_eq (new_tm, tm0) in - (ABBREV_TAC abbrev_tm THEN + (ABBREV_TAC abbrev_tm THEN EXPAND_TAC name THEN FIRST_ASSUM (fun th -> TRY (new_rewrite occs [] th)) THEN POP_ASSUM (MP_TAC o PURE_ONCE_REWRITE_RULE[GSYM (SPEC eq_var ssreflect_eq_def)]) THEN diff --git a/Formal_ineqs/lib/ssrfun.hl b/Formal_ineqs/lib/ssrfun.hl index c4cb6734..6aff8bcd 100644 --- a/Formal_ineqs/lib/ssrfun.hl +++ b/Formal_ineqs/lib/ssrfun.hl @@ -1,5 +1,5 @@ -needs "lib/ssreflect/ssreflect.hl";; -needs "lib/ssreflect/sections.hl";; +needs "Formal_ineqs/lib/ssreflect/ssreflect.hl";; +needs "Formal_ineqs/lib/ssreflect/sections.hl";; (* Module Ssrfun*) module Ssrfun = struct @@ -324,7 +324,7 @@ Sections.end_section "SopTisR";; Sections.begin_section "SopTisS";; let right_id = new_definition `right_id e op = !x. op x e = x`;; let left_zero = new_definition `left_zero z op = !x. op z x = z`;; -let right_commutative = new_definition +let right_commutative = new_definition `right_commutative op = !x y z. op (op x y) z = op (op x z) y`;; let left_distributive = new_definition `left_distributive op add = !x y z. op (add x y) z = add (op x z) (op y z)`;; @@ -340,13 +340,13 @@ Sections.end_section "SopTisS";; Sections.begin_section "SopTisT";; let left_id = new_definition `left_id e op = !x. op e x = x`;; let right_zero = new_definition `right_zero z op = !x. op x z = z`;; -let left_commutative = new_definition +let left_commutative = new_definition `left_commutative op = !x y z. op x (op y z) = op y (op x z)`;; let right_distributive = new_definition `right_distributive op add = !x y z. op x (add y z) = add (op x y) (op x z)`;; -let left_loop = new_definition +let left_loop = new_definition `left_loop inv op = !x. cancel (op x) (op (inv x))`;; -let rev_left_loop = new_definition +let rev_left_loop = new_definition `rev_left_loop inv op = !x. cancel (op (inv x)) (op x)`;; (* Finalization of the section SopTisT *) diff --git a/Formal_ineqs/lib/ssrnat.hl b/Formal_ineqs/lib/ssrnat.hl index eb1975d4..73ed3b50 100644 --- a/Formal_ineqs/lib/ssrnat.hl +++ b/Formal_ineqs/lib/ssrnat.hl @@ -1,4 +1,4 @@ -needs "lib/ssrbool.hl";; +needs "Formal_ineqs/lib/ssrbool.hl";; (* Module Ssrnat*) module Ssrnat = struct diff --git a/Formal_ineqs/list/list_conversions.hl b/Formal_ineqs/list/list_conversions.hl index b03e7ec5..9ac372b6 100644 --- a/Formal_ineqs/list/list_conversions.hl +++ b/Formal_ineqs/list/list_conversions.hl @@ -1,11 +1,17 @@ -(* =========================================================== *) -(* Efficient formal list conversions *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) -needs "arith/arith_nat.hl";; -needs "misc/misc_vars.hl";; +(* -------------------------------------------------------------------------- *) +(* Efficient formal list conversions *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/arith/arith_nat.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; module type List_conversions_sig = sig @@ -67,7 +73,7 @@ let hd_conv hd_tm = let EL_0' = (MY_RULE_NUM o prove)(`EL 0 (CONS (h:A) t) = h`, REWRITE_TAC[EL; HD]);; let EL_n' = (MY_RULE_NUM o prove)(`0 < n /\ PRE n = m ==> EL n (CONS (h:A) t) = EL m t`, - STRIP_TAC THEN SUBGOAL_THEN `n = SUC m` ASSUME_TAC THENL + STRIP_TAC THEN SUBGOAL_THEN `n = SUC m` ASSUME_TAC THENL [ REPEAT (POP_ASSUM MP_TAC) THEN ARITH_TAC; ALL_TAC ] THEN ASM_REWRITE_TAC[EL; TL]);; @@ -88,7 +94,7 @@ let eval_el n_tm list_tm = let n_gt0 = (EQT_ELIM o raw_gt0_hash_conv) n_tm in let pre_n = raw_pre_hash_conv (mk_comb (pre_op_num, n_tm)) in let m_tm = (rand o concl) pre_n in - let th0 = (MY_PROVE_HYP pre_n o MY_PROVE_HYP n_gt0 o + let th0 = (MY_PROVE_HYP pre_n o MY_PROVE_HYP n_gt0 o INST[n_tm, n_var_num; m_tm, m_var_num] o inst0) el_n in let th1 = el_conv_raw m_tm t_tm in TRANS th0 th1 in @@ -198,7 +204,7 @@ let eval_zip list1_tm list2_tm = (******************) let ALL_0' = prove(`ALL P ([]:(A)list) <=> T`, REWRITE_TAC[ALL]) and - ALL_CONS_T' = (MY_RULE o prove)(`(P h <=> T) /\ (ALL P t <=> T) ==> (ALL P (CONS (h:A) t) <=> T)`, + ALL_CONS_T' = (MY_RULE o prove)(`(P h <=> T) /\ (ALL P t <=> T) ==> (ALL P (CONS (h:A) t) <=> T)`, REWRITE_TAC[ALL]) and ALL_CONS_F2' = (MY_RULE o prove)(`(ALL P t <=> F) ==> (ALL P (CONS (h:A) t) <=> F)`, SIMP_TAC[ALL]) and @@ -216,7 +222,7 @@ let all_conv_univ p_conv tm = let ty = (hd o snd o dest_type) list_ty in let inst_t = INST_TYPE[ty, aty] in - let all_0, all_t, all_f1, all_f2 = inst_t ALL_0', inst_t ALL_CONS_T', + let all_0, all_t, all_f1, all_f2 = inst_t ALL_0', inst_t ALL_CONS_T', inst_t ALL_CONS_F1', inst_t ALL_CONS_F2' in let h_var, t_var = mk_var("h", ty), mk_var("t", list_ty) and p_var = mk_var("P", p_ty) in @@ -246,13 +252,13 @@ let all_conv_univ p_conv tm = (*******************) let ALL2_0' = prove(`ALL2 P ([]:(A)list) ([]:(B)list) <=> T`, REWRITE_TAC[ALL2]) and - ALL2_CONS_T' = (MY_RULE o prove)(`(P h1 h2 <=> T) /\ (ALL2 P t1 t2 <=> T) ==> + ALL2_CONS_T' = (MY_RULE o prove)(`(P h1 h2 <=> T) /\ (ALL2 P t1 t2 <=> T) ==> (ALL2 P (CONS (h1:A) t1) (CONS (h2:B) t2) <=> T)`, REWRITE_TAC[ALL2]) and - ALL2_CONS_F2' = (MY_RULE o prove)(`(ALL2 P t1 t2 <=> F) ==> + ALL2_CONS_F2' = (MY_RULE o prove)(`(ALL2 P t1 t2 <=> F) ==> (ALL2 P (CONS (h1:A) t1) (CONS (h2:B) t2) <=> F)`, SIMP_TAC[ALL2]) and - ALL2_CONS_F1' = (MY_RULE o prove)(`(P h1 h2 <=> F) ==> + ALL2_CONS_F1' = (MY_RULE o prove)(`(P h1 h2 <=> F) ==> (ALL2 P (CONS (h1:A) t1) (CONS (h2:B) t2) <=> F)`, SIMP_TAC[ALL2]);; @@ -270,7 +276,7 @@ let all2_conv_univ p_conv tm = ty2 = (hd o snd o dest_type) list2_ty in let inst_t = INST_TYPE[ty1, aty; ty2, bty] in - let all2_0, all2_t, all2_f1, all2_f2 = inst_t ALL2_0', inst_t ALL2_CONS_T', + let all2_0, all2_t, all2_f1, all2_f2 = inst_t ALL2_0', inst_t ALL2_CONS_T', inst_t ALL2_CONS_F1', inst_t ALL2_CONS_F2' in let h1_var, t1_var = mk_var("h1", ty1), mk_var("t1", list1_ty) and h2_var, t2_var = mk_var("h2", ty2), mk_var("t2", list2_ty) and @@ -344,9 +350,9 @@ let mem_conv_univ eq_conv mem_tm = (* FILTER conversions *) let FILTER_A_EMPTY = prove(`FILTER (P:A->bool) [] = []`, REWRITE_TAC[FILTER]) and - FILTER_A_HD = (MY_RULE o prove)(`(P h <=> T) ==> FILTER (P:A->bool) (CONS h t) = CONS h (FILTER P t)`, + FILTER_A_HD = (MY_RULE o prove)(`(P h <=> T) ==> FILTER (P:A->bool) (CONS h t) = CONS h (FILTER P t)`, SIMP_TAC[FILTER]) and - FILTER_A_TL = (MY_RULE o prove)(`(P h <=> F) ==> FILTER (P:A->bool) (CONS h t) = FILTER P t`, + FILTER_A_TL = (MY_RULE o prove)(`(P h <=> F) ==> FILTER (P:A->bool) (CONS h t) = FILTER P t`, SIMP_TAC[FILTER]);; @@ -356,12 +362,12 @@ let filter_conv_univ p_conv tm = let p_ty = type_of p_tm in let ty = (hd o snd o dest_type) p_ty in let inst_t = INST_TYPE[ty, aty] in - let filter_empty, filter_hd, filter_tl = + let filter_empty, filter_hd, filter_tl = inst_t FILTER_A_EMPTY, inst_t FILTER_A_HD, inst_t FILTER_A_TL in let p_var = mk_var("P", p_ty) in let h_var = mk_var("h", ty) in let t_var = mk_var("t", mk_type("list",[ty])) in - + let rec filter_conv_raw = fun list_tm -> if (is_comb list_tm) then let ltm, t_tm = dest_comb list_tm in @@ -381,9 +387,9 @@ let filter_conv_univ p_conv tm = else INST[p_tm, p_var] filter_empty in filter_conv_raw list_tm;; - - - + + + (***************************) (* MAP conversions *) @@ -451,7 +457,7 @@ let get_all th = else [] in get_all_raw th list_tm;; - + (* Given a theorem `ALL P list`, returns (P x_i1),..., (P x_in) @@ -485,7 +491,7 @@ let select_all th indices = else get_all_raw th_tl t_tm (i::is) (n + 1) in get_all_raw th list_tm indices 0;; - + (*****************************************) (* set_of_list conversions *) diff --git a/Formal_ineqs/list/list_float.hl b/Formal_ineqs/list/list_float.hl index 40d27ae0..66579294 100644 --- a/Formal_ineqs/list/list_float.hl +++ b/Formal_ineqs/list/list_float.hl @@ -1,202 +1,208 @@ -(* =========================================================== *) -(* Special list conversions *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -needs "arith/more_float.hl";; -needs "list/list_conversions.hl";; -needs "misc/misc_vars.hl";; - - -module type List_float_sig = sig - val list_sum : thm - val list_sum2 : thm - val error_mul_f2 : thm - val error_mul_f1 : thm - val list_sum_conv : (term -> thm) -> term -> thm - val list_sum2_le_conv : int -> (int -> term -> term -> thm) -> term -> thm - val error_mul_f2_le_conv : int -> term -> term -> thm - val error_mul_f2_le_conv2 : int -> term -> term -> thm - val error_mul_f1_le_conv : term -> int -> term -> term -> thm -end;; - - -module List_float : List_float_sig = struct - -open Misc_functions;; -open Arith_nat;; -open Arith_float;; -open More_float;; -open Float_theory;; -open List_conversions;; -open Misc_vars;; - -let MY_RULE_FLOAT = UNDISCH_ALL o NUMERALS_TO_NUM o - PURE_REWRITE_RULE[FLOAT_OF_NUM; min_exp_def; GSYM IMP_IMP] o SPEC_ALL;; - - -(****************************) -(* new definitions *) - -let list_sum = new_definition `list_sum list f = ITLIST (\t1 t2. f t1 + t2) list (&0)`;; -let list_sum2 = new_definition `list_sum2 f l1 l2 = ITLIST2 (\a b c. f a b + c) l1 l2 (&0)`;; - -let error_mul_f2 = new_definition `error_mul_f2 a int = a * iabs int`;; -let error_mul_f1 = new_definition `error_mul_f1 w x list = x * list_sum2 error_mul_f2 w list`;; - -(*************************************) -(* list_sum conversions *) - -let LIST_SUM_A_EMPTY = prove(`list_sum [] (f:A->real) = &0`, REWRITE_TAC[list_sum; ITLIST]) and - LIST_SUM_A_H = prove(`list_sum [h:A] f = f h`, REWRITE_TAC[list_sum; ITLIST; REAL_ADD_RID]) and - LIST_SUM_A_CONS = prove(`list_sum (CONS (h:A) t) f = f h + list_sum t f`, REWRITE_TAC[list_sum; ITLIST]);; - - -let list_sum_conv f_conv tm = - let ltm, f_tm = dest_comb tm in - let list_tm = rand ltm in - let list_ty = type_of list_tm in - let f_ty = type_of f_tm in - let ty = (hd o snd o dest_type) list_ty in - let f_var = mk_var("f", f_ty) and - h_var = mk_var("h", ty) and - t_var = mk_var("t", list_ty) in - let inst_t = INST[f_tm, f_var] o INST_TYPE[ty, aty] in - let list_sum_h = inst_t LIST_SUM_A_H and - list_sum_cons = inst_t LIST_SUM_A_CONS in - - let rec list_sum_conv_raw = fun h_tm t_tm -> - if (is_comb t_tm) then - let h_tm', t_tm' = dest_comb t_tm in - let th0 = INST[h_tm, h_var; t_tm, t_var] list_sum_cons in - let ltm, rtm = dest_comb(rand(concl th0)) in - let plus_op, fh_tm = dest_comb ltm in - let f_th = f_conv fh_tm in - let th1 = list_sum_conv_raw (rand h_tm') t_tm' in - let th2 = MK_COMB(AP_TERM plus_op f_th, th1) in - TRANS th0 th2 - else - let th0 = INST[h_tm, h_var] list_sum_h in - let f_th = f_conv (rand(concl th0)) in - TRANS th0 f_th in - - if (is_comb list_tm) then - let h_tm, t_tm = dest_comb list_tm in - list_sum_conv_raw (rand h_tm) t_tm - else - inst_t LIST_SUM_A_EMPTY;; - - - -(*************************************) -(* list_sum2 evaluation *) - -let LIST_SUM2_0_LE' = (MY_RULE_FLOAT o prove)(`list_sum2 (f:A->B->real) [] [] <= &0`, - REWRITE_TAC[list_sum2; ITLIST2; REAL_LE_REFL]);; -let LIST_SUM2_1_LE' = (MY_RULE_FLOAT o prove)(`f h1 h2 <= x ==> list_sum2 (f:A->B->real) [h1] [h2] <= x`, - REWRITE_TAC[list_sum2; ITLIST2; REAL_ADD_RID]);; -let LIST_SUM2_LE' = (MY_RULE_FLOAT o prove)(`f h1 h2 <= x /\ list_sum2 f t1 t2 <= y /\ x + y <= z ==> - list_sum2 (f:A->B->real) (CONS h1 t1) (CONS h2 t2) <= z`, - REWRITE_TAC[list_sum2; ITLIST2] THEN STRIP_TAC THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `x + y:real` THEN - ASM_SIMP_TAC[REAL_LE_ADD2]);; - - -let list_sum2_le_conv pp f_le_conv tm = - let ltm, list2_tm = dest_comb tm in - let ltm2, list1_tm = dest_comb ltm in - let f_tm = rand ltm2 in - let list1_ty = type_of list1_tm and - list2_ty = type_of list2_tm and - f_ty = type_of f_tm in - let ty1 = (hd o snd o dest_type) list1_ty and - ty2 = (hd o snd o dest_type) list2_ty in - let f_var = mk_var ("f", f_ty) and - h1_var, t1_var = mk_var ("h1", ty1), mk_var ("t1", list1_ty) and - h2_var, t2_var = mk_var ("h2", ty2), mk_var ("t2", list2_ty) in - let inst_t = INST[f_tm, f_var] o INST_TYPE[ty1, aty; ty2, bty] in - let list2_0, list2_1, list2_le = inst_t LIST_SUM2_0_LE', inst_t LIST_SUM2_1_LE', inst_t LIST_SUM2_LE' in - - let rec rec_conv = fun list1_tm list2_tm -> - if (is_comb list1_tm) then - let h1_tm, t1_tm = dest_cons list1_tm and - h2_tm, t2_tm = dest_cons list2_tm in - let f_le_th = f_le_conv pp h1_tm h2_tm in - let x_tm = (rand o concl) f_le_th in - let inst0 = INST[h1_tm, h1_var; h2_tm, h2_var; x_tm, x_var_real] in - if is_comb t1_tm then - let sum2_t_th = rec_conv t1_tm t2_tm in - let y_tm = (rand o concl) sum2_t_th in - let xy_th = float_add_hi pp x_tm y_tm in - let z_tm = (rand o concl) xy_th in - (MY_PROVE_HYP xy_th o MY_PROVE_HYP sum2_t_th o MY_PROVE_HYP f_le_th o - INST[y_tm, y_var_real; z_tm, z_var_real; t1_tm, t1_var; t2_tm, t2_var] o - inst0) list2_le - else - if is_comb t2_tm then failwith ("sum2_le_conv: t1 = []; t2 = "^string_of_term t2_tm) else - (MY_PROVE_HYP f_le_th o inst0) list2_1 - else - if is_comb list2_tm then failwith ("sum2_le_conv: list1 = []; list2 = "^string_of_term list2_tm) else - list2_0 in - - rec_conv list1_tm list2_tm;; - - - -(**************************) -(* \a b c. a * iabs b + c *) - -let ERROR_MUL_F2' = (SYM o MY_RULE_FLOAT) error_mul_f2;; - - -(* |- x = a, |- P x y -> P a y *) -let rewrite_lhs eq_th th = - let ltm, rhs = dest_comb (concl th) in - let th0 = AP_THM (AP_TERM (rator ltm) eq_th) rhs in - EQ_MP th0 th;; - -let error_mul_f2_le_conv pp tm1 tm2 = - let eq_th = INST[tm1, a_var_real; tm2, int_var] ERROR_MUL_F2' in - let iabs_th = float_iabs tm2 in - let iabs_tm = (rand o concl) iabs_th in - let mul_th = float_mul_hi pp tm1 iabs_tm in - let th0 = AP_TERM (mk_comb (mul_op_real, tm1)) iabs_th in - let th1 = AP_THM (AP_TERM le_op_real th0) (rand (concl mul_th)) in - let le_th = EQ_MP (SYM th1) mul_th in - rewrite_lhs eq_th le_th;; - -let ERROR_MUL_F2_LEMMA' = (MY_RULE_FLOAT o prove)(`iabs int = x /\ a * x <= y ==> error_mul_f2 a int <= y`, - REWRITE_TAC[error_mul_f2] THEN STRIP_TAC THEN ASM_REWRITE_TAC[]);; - -let error_mul_f2_le_conv2 pp tm1 tm2 = - let iabs_th = float_iabs tm2 in - let x_tm = (rand o concl) iabs_th in - let mul_th = float_mul_hi pp tm1 x_tm in - let y_tm = (rand o concl) mul_th in - (MY_PROVE_HYP iabs_th o MY_PROVE_HYP mul_th o - INST[tm2, int_var; tm1, a_var_real; x_tm, x_var_real; y_tm, y_var_real]) ERROR_MUL_F2_LEMMA';; - - - -(**************************) -(* \a b c. a * iabs b + c *) - -let ERROR_MUL_F1_LEMMA' = (MY_RULE_FLOAT o prove)(`x * list_sum2 error_mul_f2 w list <= z ==> - error_mul_f1 w x list <= z`, REWRITE_TAC[error_mul_f1]);; - -let list_sum2_error2_const = `list_sum2 error_mul_f2` and - w_var_list = `w:(real)list` and - list_var = `list:(real#real)list`;; - -let error_mul_f1_le_conv w_tm pp x_tm list_tm = - (* TODO: if x = 0 then do not need to compute the sum *) - let sum2_tm = mk_binop list_sum2_error2_const w_tm list_tm in - let sum2_le_th = list_sum2_le_conv pp error_mul_f2_le_conv2 sum2_tm in - let ineq_th = mul_ineq_pos_const_hi pp x_tm sum2_le_th in - let z_tm = (rand o concl) ineq_th in - (MY_PROVE_HYP ineq_th o - INST[x_tm, x_var_real; z_tm, z_var_real; w_tm, w_var_list; list_tm, list_var]) ERROR_MUL_F1_LEMMA';; - - -end;; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Special list conversions *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/arith/more_float.hl";; +needs "Formal_ineqs/list/list_conversions.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; + + +module type List_float_sig = sig + val list_sum : thm + val list_sum2 : thm + val error_mul_f2 : thm + val error_mul_f1 : thm + val list_sum_conv : (term -> thm) -> term -> thm + val list_sum2_le_conv : int -> (int -> term -> term -> thm) -> term -> thm + val error_mul_f2_le_conv : int -> term -> term -> thm + val error_mul_f2_le_conv2 : int -> term -> term -> thm + val error_mul_f1_le_conv : term -> int -> term -> term -> thm +end;; + + +module List_float : List_float_sig = struct + +open Misc_functions;; +open Arith_nat;; +open Arith_float;; +open More_float;; +open Float_theory;; +open List_conversions;; +open Misc_vars;; + +let MY_RULE_FLOAT = UNDISCH_ALL o NUMERALS_TO_NUM o + PURE_REWRITE_RULE[FLOAT_OF_NUM; min_exp_def; GSYM IMP_IMP] o SPEC_ALL;; + + +(****************************) +(* new definitions *) + +let list_sum = new_definition `list_sum list f = ITLIST (\t1 t2. f t1 + t2) list (&0)`;; +let list_sum2 = new_definition `list_sum2 f l1 l2 = ITLIST2 (\a b c. f a b + c) l1 l2 (&0)`;; + +let error_mul_f2 = new_definition `error_mul_f2 a int = a * iabs int`;; +let error_mul_f1 = new_definition `error_mul_f1 w x list = x * list_sum2 error_mul_f2 w list`;; + +(*************************************) +(* list_sum conversions *) + +let LIST_SUM_A_EMPTY = prove(`list_sum [] (f:A->real) = &0`, REWRITE_TAC[list_sum; ITLIST]) and + LIST_SUM_A_H = prove(`list_sum [h:A] f = f h`, REWRITE_TAC[list_sum; ITLIST; REAL_ADD_RID]) and + LIST_SUM_A_CONS = prove(`list_sum (CONS (h:A) t) f = f h + list_sum t f`, REWRITE_TAC[list_sum; ITLIST]);; + + +let list_sum_conv f_conv tm = + let ltm, f_tm = dest_comb tm in + let list_tm = rand ltm in + let list_ty = type_of list_tm in + let f_ty = type_of f_tm in + let ty = (hd o snd o dest_type) list_ty in + let f_var = mk_var("f", f_ty) and + h_var = mk_var("h", ty) and + t_var = mk_var("t", list_ty) in + let inst_t = INST[f_tm, f_var] o INST_TYPE[ty, aty] in + let list_sum_h = inst_t LIST_SUM_A_H and + list_sum_cons = inst_t LIST_SUM_A_CONS in + + let rec list_sum_conv_raw = fun h_tm t_tm -> + if (is_comb t_tm) then + let h_tm', t_tm' = dest_comb t_tm in + let th0 = INST[h_tm, h_var; t_tm, t_var] list_sum_cons in + let ltm, rtm = dest_comb(rand(concl th0)) in + let plus_op, fh_tm = dest_comb ltm in + let f_th = f_conv fh_tm in + let th1 = list_sum_conv_raw (rand h_tm') t_tm' in + let th2 = MK_COMB(AP_TERM plus_op f_th, th1) in + TRANS th0 th2 + else + let th0 = INST[h_tm, h_var] list_sum_h in + let f_th = f_conv (rand(concl th0)) in + TRANS th0 f_th in + + if (is_comb list_tm) then + let h_tm, t_tm = dest_comb list_tm in + list_sum_conv_raw (rand h_tm) t_tm + else + inst_t LIST_SUM_A_EMPTY;; + + + +(*************************************) +(* list_sum2 evaluation *) + +let LIST_SUM2_0_LE' = (MY_RULE_FLOAT o prove)(`list_sum2 (f:A->B->real) [] [] <= &0`, + REWRITE_TAC[list_sum2; ITLIST2; REAL_LE_REFL]);; +let LIST_SUM2_1_LE' = (MY_RULE_FLOAT o prove)(`f h1 h2 <= x ==> list_sum2 (f:A->B->real) [h1] [h2] <= x`, + REWRITE_TAC[list_sum2; ITLIST2; REAL_ADD_RID]);; +let LIST_SUM2_LE' = (MY_RULE_FLOAT o prove)(`f h1 h2 <= x /\ list_sum2 f t1 t2 <= y /\ x + y <= z ==> + list_sum2 (f:A->B->real) (CONS h1 t1) (CONS h2 t2) <= z`, + REWRITE_TAC[list_sum2; ITLIST2] THEN STRIP_TAC THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `x + y:real` THEN + ASM_SIMP_TAC[REAL_LE_ADD2]);; + + +let list_sum2_le_conv pp f_le_conv tm = + let ltm, list2_tm = dest_comb tm in + let ltm2, list1_tm = dest_comb ltm in + let f_tm = rand ltm2 in + let list1_ty = type_of list1_tm and + list2_ty = type_of list2_tm and + f_ty = type_of f_tm in + let ty1 = (hd o snd o dest_type) list1_ty and + ty2 = (hd o snd o dest_type) list2_ty in + let f_var = mk_var ("f", f_ty) and + h1_var, t1_var = mk_var ("h1", ty1), mk_var ("t1", list1_ty) and + h2_var, t2_var = mk_var ("h2", ty2), mk_var ("t2", list2_ty) in + let inst_t = INST[f_tm, f_var] o INST_TYPE[ty1, aty; ty2, bty] in + let list2_0, list2_1, list2_le = inst_t LIST_SUM2_0_LE', inst_t LIST_SUM2_1_LE', inst_t LIST_SUM2_LE' in + + let rec rec_conv = fun list1_tm list2_tm -> + if (is_comb list1_tm) then + let h1_tm, t1_tm = dest_cons list1_tm and + h2_tm, t2_tm = dest_cons list2_tm in + let f_le_th = f_le_conv pp h1_tm h2_tm in + let x_tm = (rand o concl) f_le_th in + let inst0 = INST[h1_tm, h1_var; h2_tm, h2_var; x_tm, x_var_real] in + if is_comb t1_tm then + let sum2_t_th = rec_conv t1_tm t2_tm in + let y_tm = (rand o concl) sum2_t_th in + let xy_th = float_add_hi pp x_tm y_tm in + let z_tm = (rand o concl) xy_th in + (MY_PROVE_HYP xy_th o MY_PROVE_HYP sum2_t_th o MY_PROVE_HYP f_le_th o + INST[y_tm, y_var_real; z_tm, z_var_real; t1_tm, t1_var; t2_tm, t2_var] o + inst0) list2_le + else + if is_comb t2_tm then failwith ("sum2_le_conv: t1 = []; t2 = "^string_of_term t2_tm) else + (MY_PROVE_HYP f_le_th o inst0) list2_1 + else + if is_comb list2_tm then failwith ("sum2_le_conv: list1 = []; list2 = "^string_of_term list2_tm) else + list2_0 in + + rec_conv list1_tm list2_tm;; + + + +(**************************) +(* \a b c. a * iabs b + c *) + +let ERROR_MUL_F2' = (SYM o MY_RULE_FLOAT) error_mul_f2;; + + +(* |- x = a, |- P x y -> P a y *) +let rewrite_lhs eq_th th = + let ltm, rhs = dest_comb (concl th) in + let th0 = AP_THM (AP_TERM (rator ltm) eq_th) rhs in + EQ_MP th0 th;; + +let error_mul_f2_le_conv pp tm1 tm2 = + let eq_th = INST[tm1, a_var_real; tm2, int_var] ERROR_MUL_F2' in + let iabs_th = float_iabs tm2 in + let iabs_tm = (rand o concl) iabs_th in + let mul_th = float_mul_hi pp tm1 iabs_tm in + let th0 = AP_TERM (mk_comb (mul_op_real, tm1)) iabs_th in + let th1 = AP_THM (AP_TERM le_op_real th0) (rand (concl mul_th)) in + let le_th = EQ_MP (SYM th1) mul_th in + rewrite_lhs eq_th le_th;; + +let ERROR_MUL_F2_LEMMA' = (MY_RULE_FLOAT o prove)(`iabs int = x /\ a * x <= y ==> error_mul_f2 a int <= y`, + REWRITE_TAC[error_mul_f2] THEN STRIP_TAC THEN ASM_REWRITE_TAC[]);; + +let error_mul_f2_le_conv2 pp tm1 tm2 = + let iabs_th = float_iabs tm2 in + let x_tm = (rand o concl) iabs_th in + let mul_th = float_mul_hi pp tm1 x_tm in + let y_tm = (rand o concl) mul_th in + (MY_PROVE_HYP iabs_th o MY_PROVE_HYP mul_th o + INST[tm2, int_var; tm1, a_var_real; x_tm, x_var_real; y_tm, y_var_real]) ERROR_MUL_F2_LEMMA';; + + + +(**************************) +(* \a b c. a * iabs b + c *) + +let ERROR_MUL_F1_LEMMA' = (MY_RULE_FLOAT o prove)(`x * list_sum2 error_mul_f2 w list <= z ==> + error_mul_f1 w x list <= z`, REWRITE_TAC[error_mul_f1]);; + +let list_sum2_error2_const = `list_sum2 error_mul_f2` and + w_var_list = `w:(real)list` and + list_var = `list:(real#real)list`;; + +let error_mul_f1_le_conv w_tm pp x_tm list_tm = + (* TODO: if x = 0 then do not need to compute the sum *) + let sum2_tm = mk_binop list_sum2_error2_const w_tm list_tm in + let sum2_le_th = list_sum2_le_conv pp error_mul_f2_le_conv2 sum2_tm in + let ineq_th = mul_ineq_pos_const_hi pp x_tm sum2_le_th in + let z_tm = (rand o concl) ineq_th in + (MY_PROVE_HYP ineq_th o + INST[x_tm, x_var_real; z_tm, z_var_real; w_tm, w_var_list; list_tm, list_var]) ERROR_MUL_F1_LEMMA';; + + +end;; diff --git a/Formal_ineqs/list/more_list.hl b/Formal_ineqs/list/more_list.hl index 0450ddf0..0dbb24b9 100644 --- a/Formal_ineqs/list/more_list.hl +++ b/Formal_ineqs/list/more_list.hl @@ -1,134 +1,140 @@ -(* =========================================================== *) -(* Additional list definitions and theorems *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -module More_list = struct - -(* definitions *) -let REVERSE_TABLE = define `(REVERSE_TABLE (f:num->A) 0 = []) /\ - (REVERSE_TABLE f (SUC i) = CONS (f i) ( REVERSE_TABLE f i))`;; - -let TABLE = new_definition `!(f:num->A) k. TABLE f k = REVERSE (REVERSE_TABLE f k)`;; - -let l_seq = new_definition `l_seq n m = TABLE (\i. n + i) ((m + 1) - n)`;; - -(* lemmas *) -let LENGTH_REVERSE_TABLE = prove(`!(f:num->A) n. LENGTH (REVERSE_TABLE f n) = n`, - GEN_TAC THEN INDUCT_TAC THEN ASM_REWRITE_TAC[REVERSE_TABLE; LENGTH]);; - -let LENGTH_REVERSE = prove(`!(l:(A)list). LENGTH (REVERSE l) = LENGTH l`, - MATCH_MP_TAC list_INDUCT THEN REWRITE_TAC[REVERSE] THEN - REPEAT STRIP_TAC THEN - ASM_REWRITE_TAC[LENGTH_APPEND; LENGTH] THEN - ARITH_TAC);; - -let LENGTH_TABLE = prove(`!(f:num->A) n. LENGTH (TABLE f n) = n`, - REWRITE_TAC[TABLE; LENGTH_REVERSE; LENGTH_REVERSE_TABLE]);; - -let EL_TABLE = prove(`!(f:num->A) n i. i < n ==> EL i (TABLE f n) = f i`, - REPEAT GEN_TAC THEN SPEC_TAC (`n:num`, `n:num`) THEN - INDUCT_TAC THENL [ ARITH_TAC; ALL_TAC ] THEN - REWRITE_TAC[TABLE; REVERSE_TABLE; REVERSE; EL_APPEND] THEN - REWRITE_TAC[GSYM TABLE; LENGTH_TABLE] THEN - DISCH_TAC THEN COND_CASES_TAC THENL - [ - FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - SUBGOAL_THEN `i = n:num` (fun th -> REWRITE_TAC[th]) THENL - [ - POP_ASSUM MP_TAC THEN POP_ASSUM MP_TAC THEN - ARITH_TAC; - ALL_TAC - ] THEN - REWRITE_TAC[ARITH_RULE `n - n = 0`; EL; HD]);; - -let TABLE_0 = prove(`!(f:num->A). TABLE f 0 = []`, - REWRITE_TAC[TABLE; REVERSE_TABLE; REVERSE]);; - -let TABLE_SUC = prove(`!(f:num->A) n. TABLE f (SUC n) = APPEND (TABLE f n) [f n]`, - GEN_TAC THEN INDUCT_TAC THEN REWRITE_TAC[TABLE; REVERSE_TABLE; REVERSE]);; - -let LIST_EL_EQ = prove(`!ul vl:(A)list. ul = vl <=> - (LENGTH ul = LENGTH vl /\ (!j. j < LENGTH ul ==> EL j ul = EL j vl))`, - REPEAT STRIP_TAC THEN - EQ_TAC THEN STRIP_TAC THEN ASM_REWRITE_TAC[] THEN - POP_ASSUM MP_TAC THEN POP_ASSUM MP_TAC THEN - SPEC_TAC (`vl:(A)list`, `vl:(A)list`) THEN SPEC_TAC (`ul:(A)list`, `ul:(A)list`) THEN - LIST_INDUCT_TAC THEN LIST_INDUCT_TAC THEN SIMP_TAC[LENGTH_EQ_NIL; EQ_SYM_EQ; LENGTH; ARITH_RULE `~(0 = SUC a)`] THEN - POP_ASSUM (fun th -> ALL_TAC) THEN - REWRITE_TAC[ARITH_RULE `SUC a = SUC b <=> a = b`] THEN - REPEAT STRIP_TAC THEN - FIRST_X_ASSUM (MP_TAC o SPEC `t':(A)list`) THEN - ASM_REWRITE_TAC[] THEN - ANTS_TAC THENL - [ - REPEAT STRIP_TAC THEN - FIRST_X_ASSUM (MP_TAC o SPEC `SUC j`) THEN - ASM_REWRITE_TAC[ARITH_RULE `SUC a < SUC b <=> a < b`; EL; TL]; - ALL_TAC - ] THEN - DISCH_TAC THEN - FIRST_X_ASSUM (MP_TAC o SPEC `0`) THEN - ASM_SIMP_TAC[ARITH_RULE `0 < SUC a`; EL; HD]);; - - -let LENGTH_L_SEQ = prove(`LENGTH (l_seq n m) = (m + 1) - n`, REWRITE_TAC[l_seq; LENGTH_TABLE]);; - -let EL_L_SEQ = prove(`!i m n. i < (m + 1) - n ==> EL i (l_seq n m) = n + i`, - REWRITE_TAC[l_seq] THEN REPEAT STRIP_TAC THEN MATCH_MP_TAC EL_TABLE THEN ASM_REWRITE_TAC[]);; - -let L_SEQ_NIL = prove(`!n m. l_seq n m = [] <=> (m < n)`, - GEN_TAC THEN GEN_TAC THEN REWRITE_TAC[GSYM LENGTH_EQ_NIL; LENGTH_L_SEQ] THEN ARITH_TAC);; - -let L_SEQ_NN = prove(`!n. l_seq n n = [n]`, - GEN_TAC THEN REWRITE_TAC[l_seq; ARITH_RULE `(n + 1) - n = 1`; ONE; TABLE; REVERSE_TABLE; REVERSE] THEN - REWRITE_TAC[APPEND; ADD_0]);; - -let L_SEQ_CONS = prove(`!n m. n <= m ==> l_seq n m = CONS n (l_seq (n + 1) m)`, - REPEAT STRIP_TAC THEN - REWRITE_TAC[LIST_EL_EQ; LENGTH_L_SEQ; LENGTH] THEN CONJ_TAC THENL - [ - ASM_ARITH_TAC; - ALL_TAC - ] THEN - INDUCT_TAC THENL - [ - ASM_SIMP_TAC[EL_L_SEQ] THEN ASM_REWRITE_TAC[EL; HD] THEN ARITH_TAC; - DISCH_TAC THEN - ASM_SIMP_TAC[EL_L_SEQ] THEN ASM_REWRITE_TAC[EL; TL] THEN - MP_TAC (SPECL [`j:num`; `m:num`; `n + 1`] EL_L_SEQ) THEN ANTS_TAC THENL - [ - ASM_ARITH_TAC; - ALL_TAC - ] THEN - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN ARITH_TAC - ]);; - - -let LENGTH_BUTLAST = prove(`!l. LENGTH (BUTLAST l) = LENGTH l - 1`, - MATCH_MP_TAC list_INDUCT THEN REWRITE_TAC[BUTLAST; LENGTH; ARITH] THEN REPEAT STRIP_TAC THEN - COND_CASES_TAC THEN ASM_REWRITE_TAC[LENGTH; ARITH] THEN - POP_ASSUM MP_TAC THEN REWRITE_TAC[GSYM LENGTH_EQ_NIL] THEN - ARITH_TAC);; - -let EL_BUTLAST = prove(`!(l:(A)list) i. i < LENGTH l - 1 ==> EL i (BUTLAST l) = EL i l`, - MATCH_MP_TAC list_INDUCT THEN REWRITE_TAC[BUTLAST; LENGTH] THEN REPEAT STRIP_TAC THEN - COND_CASES_TAC THENL - [ - UNDISCH_TAC `i < SUC (LENGTH (a1:(A)list)) - 1` THEN - ASM_REWRITE_TAC[LENGTH] THEN ARITH_TAC; - ALL_TAC - ] THEN - REWRITE_TAC[EL_CONS] THEN COND_CASES_TAC THEN ASM_REWRITE_TAC[] THEN - FIRST_X_ASSUM (MP_TAC o SPEC `i - 1`) THEN - ANTS_TAC THENL - [ - ASM_ARITH_TAC; - ALL_TAC - ] THEN - REWRITE_TAC[]);; - -end;; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Additional list definitions and theorems *) +(* -------------------------------------------------------------------------- *) + +module More_list = struct + +(* definitions *) +let REVERSE_TABLE = define `(REVERSE_TABLE (f:num->A) 0 = []) /\ + (REVERSE_TABLE f (SUC i) = CONS (f i) ( REVERSE_TABLE f i))`;; + +let TABLE = new_definition `!(f:num->A) k. TABLE f k = REVERSE (REVERSE_TABLE f k)`;; + +let l_seq = new_definition `l_seq n m = TABLE (\i. n + i) ((m + 1) - n)`;; + +(* lemmas *) +let LENGTH_REVERSE_TABLE = prove(`!(f:num->A) n. LENGTH (REVERSE_TABLE f n) = n`, + GEN_TAC THEN INDUCT_TAC THEN ASM_REWRITE_TAC[REVERSE_TABLE; LENGTH]);; + +let LENGTH_REVERSE = prove(`!(l:(A)list). LENGTH (REVERSE l) = LENGTH l`, + MATCH_MP_TAC list_INDUCT THEN REWRITE_TAC[REVERSE] THEN + REPEAT STRIP_TAC THEN + ASM_REWRITE_TAC[LENGTH_APPEND; LENGTH] THEN + ARITH_TAC);; + +let LENGTH_TABLE = prove(`!(f:num->A) n. LENGTH (TABLE f n) = n`, + REWRITE_TAC[TABLE; LENGTH_REVERSE; LENGTH_REVERSE_TABLE]);; + +let EL_TABLE = prove(`!(f:num->A) n i. i < n ==> EL i (TABLE f n) = f i`, + REPEAT GEN_TAC THEN SPEC_TAC (`n:num`, `n:num`) THEN + INDUCT_TAC THENL [ ARITH_TAC; ALL_TAC ] THEN + REWRITE_TAC[TABLE; REVERSE_TABLE; REVERSE; EL_APPEND] THEN + REWRITE_TAC[GSYM TABLE; LENGTH_TABLE] THEN + DISCH_TAC THEN COND_CASES_TAC THENL + [ + FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + SUBGOAL_THEN `i = n:num` (fun th -> REWRITE_TAC[th]) THENL + [ + POP_ASSUM MP_TAC THEN POP_ASSUM MP_TAC THEN + ARITH_TAC; + ALL_TAC + ] THEN + REWRITE_TAC[ARITH_RULE `n - n = 0`; EL; HD]);; + +let TABLE_0 = prove(`!(f:num->A). TABLE f 0 = []`, + REWRITE_TAC[TABLE; REVERSE_TABLE; REVERSE]);; + +let TABLE_SUC = prove(`!(f:num->A) n. TABLE f (SUC n) = APPEND (TABLE f n) [f n]`, + GEN_TAC THEN INDUCT_TAC THEN REWRITE_TAC[TABLE; REVERSE_TABLE; REVERSE]);; + +let LIST_EL_EQ = prove(`!ul vl:(A)list. ul = vl <=> + (LENGTH ul = LENGTH vl /\ (!j. j < LENGTH ul ==> EL j ul = EL j vl))`, + REPEAT STRIP_TAC THEN + EQ_TAC THEN STRIP_TAC THEN ASM_REWRITE_TAC[] THEN + POP_ASSUM MP_TAC THEN POP_ASSUM MP_TAC THEN + SPEC_TAC (`vl:(A)list`, `vl:(A)list`) THEN SPEC_TAC (`ul:(A)list`, `ul:(A)list`) THEN + LIST_INDUCT_TAC THEN LIST_INDUCT_TAC THEN SIMP_TAC[LENGTH_EQ_NIL; EQ_SYM_EQ; LENGTH; ARITH_RULE `~(0 = SUC a)`] THEN + POP_ASSUM (fun th -> ALL_TAC) THEN + REWRITE_TAC[ARITH_RULE `SUC a = SUC b <=> a = b`] THEN + REPEAT STRIP_TAC THEN + FIRST_X_ASSUM (MP_TAC o SPEC `t':(A)list`) THEN + ASM_REWRITE_TAC[] THEN + ANTS_TAC THENL + [ + REPEAT STRIP_TAC THEN + FIRST_X_ASSUM (MP_TAC o SPEC `SUC j`) THEN + ASM_REWRITE_TAC[ARITH_RULE `SUC a < SUC b <=> a < b`; EL; TL]; + ALL_TAC + ] THEN + DISCH_TAC THEN + FIRST_X_ASSUM (MP_TAC o SPEC `0`) THEN + ASM_SIMP_TAC[ARITH_RULE `0 < SUC a`; EL; HD]);; + + +let LENGTH_L_SEQ = prove(`LENGTH (l_seq n m) = (m + 1) - n`, REWRITE_TAC[l_seq; LENGTH_TABLE]);; + +let EL_L_SEQ = prove(`!i m n. i < (m + 1) - n ==> EL i (l_seq n m) = n + i`, + REWRITE_TAC[l_seq] THEN REPEAT STRIP_TAC THEN MATCH_MP_TAC EL_TABLE THEN ASM_REWRITE_TAC[]);; + +let L_SEQ_NIL = prove(`!n m. l_seq n m = [] <=> (m < n)`, + GEN_TAC THEN GEN_TAC THEN REWRITE_TAC[GSYM LENGTH_EQ_NIL; LENGTH_L_SEQ] THEN ARITH_TAC);; + +let L_SEQ_NN = prove(`!n. l_seq n n = [n]`, + GEN_TAC THEN REWRITE_TAC[l_seq; ARITH_RULE `(n + 1) - n = 1`; ONE; TABLE; REVERSE_TABLE; REVERSE] THEN + REWRITE_TAC[APPEND; ADD_0]);; + +let L_SEQ_CONS = prove(`!n m. n <= m ==> l_seq n m = CONS n (l_seq (n + 1) m)`, + REPEAT STRIP_TAC THEN + REWRITE_TAC[LIST_EL_EQ; LENGTH_L_SEQ; LENGTH] THEN CONJ_TAC THENL + [ + ASM_ARITH_TAC; + ALL_TAC + ] THEN + INDUCT_TAC THENL + [ + ASM_SIMP_TAC[EL_L_SEQ] THEN ASM_REWRITE_TAC[EL; HD] THEN ARITH_TAC; + DISCH_TAC THEN + ASM_SIMP_TAC[EL_L_SEQ] THEN ASM_REWRITE_TAC[EL; TL] THEN + MP_TAC (SPECL [`j:num`; `m:num`; `n + 1`] EL_L_SEQ) THEN ANTS_TAC THENL + [ + ASM_ARITH_TAC; + ALL_TAC + ] THEN + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN ARITH_TAC + ]);; + + +let LENGTH_BUTLAST = prove(`!l. LENGTH (BUTLAST l) = LENGTH l - 1`, + MATCH_MP_TAC list_INDUCT THEN REWRITE_TAC[BUTLAST; LENGTH; ARITH] THEN REPEAT STRIP_TAC THEN + COND_CASES_TAC THEN ASM_REWRITE_TAC[LENGTH; ARITH] THEN + POP_ASSUM MP_TAC THEN REWRITE_TAC[GSYM LENGTH_EQ_NIL] THEN + ARITH_TAC);; + +let EL_BUTLAST = prove(`!(l:(A)list) i. i < LENGTH l - 1 ==> EL i (BUTLAST l) = EL i l`, + MATCH_MP_TAC list_INDUCT THEN REWRITE_TAC[BUTLAST; LENGTH] THEN REPEAT STRIP_TAC THEN + COND_CASES_TAC THENL + [ + UNDISCH_TAC `i < SUC (LENGTH (a1:(A)list)) - 1` THEN + ASM_REWRITE_TAC[LENGTH] THEN ARITH_TAC; + ALL_TAC + ] THEN + REWRITE_TAC[EL_CONS] THEN COND_CASES_TAC THEN ASM_REWRITE_TAC[] THEN + FIRST_X_ASSUM (MP_TAC o SPEC `i - 1`) THEN + ANTS_TAC THENL + [ + ASM_ARITH_TAC; + ALL_TAC + ] THEN + REWRITE_TAC[]);; + +end;; diff --git a/Formal_ineqs/misc/misc_functions.hl b/Formal_ineqs/misc/misc_functions.hl index 9c666745..d0fa9fae 100644 --- a/Formal_ineqs/misc/misc_functions.hl +++ b/Formal_ineqs/misc/misc_functions.hl @@ -1,82 +1,88 @@ -(* =========================================================== *) -(* Miscellaneous functions *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -#load "unix.cma";; - -module Misc_functions = struct - -exception Error of string * term list * thm list;; - -let error msg tms ths = raise (Error (msg, tms, ths));; - -let error_msg msg = raise (Error (msg, [], []));; - -let error_fmt str fmt = - let msg = sprintf str fmt in - error_msg msg;; - -(* A little faster version of PROVE_HYP *) -let MY_PROVE_HYP hyp th = EQ_MP (DEDUCT_ANTISYM_RULE hyp th) hyp;; - -(* A faster version of BETA_RULE *) -let MY_BETA_RULE th = - let rec beta tm = - let op, arg = dest_comb tm in - if is_comb op then - let op_th = AP_THM (beta op) arg in - let beta_th = BETA_CONV (rand (concl op_th)) in - TRANS op_th beta_th - else - BETA_CONV tm in - EQ_MP (beta (concl th)) th;; - -(* Converts a list with two elements into a pair *) -let pair_of_list = function - | [a; b] -> a, b - | _ -> failwith "pair_of_list: list does not contain exactly two element";; - -(* Applies f to arg n times and returns the total execution time *) -let test n f arg = - let start = Unix.gettimeofday() in - let result = f arg in - for i = 1 to n - 1 do - let _ = f arg in () - done; - result, Unix.gettimeofday() -. start;; - -(* Do not use when msg is not a fixed string *) -let assert_msg (cond, msg) = - if cond then () else failwith msg;; - -(* Generates a power function for the given binary operation *) -let gen_pow op id = - let ( * ) = op in - let rec pow n x = - if n <= 0 then id - else if n = 1 then x - else if n land 1 = 1 then - x * pow (n - 1) x - else - let t = pow (n lsr 1) x in - t * t in - pow;; - -let rec shape_list n list = - if length list <= n then [list] - else - let l1, l2 = chop_list n list in - l1 :: shape_list n l2;; - -(* map2 which works for lists of any size (no requirement |l1| = |l2|) *) -let rec my_map2 f l1 l2 = - match l1 with - | [] -> [] - | (h1::t1) -> - (match l2 with - | [] -> [] - | (h2::t2) -> (f h1 h2) :: my_map2 f t1 t2);; - -end;; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Miscellaneous functions *) +(* -------------------------------------------------------------------------- *) + +#load "unix.cma";; + +module Misc_functions = struct + +exception Error of string * term list * thm list;; + +let error msg tms ths = raise (Error (msg, tms, ths));; + +let error_msg msg = raise (Error (msg, [], []));; + +let error_fmt str fmt = + let msg = sprintf str fmt in + error_msg msg;; + +(* A little faster version of PROVE_HYP *) +let MY_PROVE_HYP hyp th = EQ_MP (DEDUCT_ANTISYM_RULE hyp th) hyp;; + +(* A faster version of BETA_RULE *) +let MY_BETA_RULE th = + let rec beta tm = + let op, arg = dest_comb tm in + if is_comb op then + let op_th = AP_THM (beta op) arg in + let beta_th = BETA_CONV (rand (concl op_th)) in + TRANS op_th beta_th + else + BETA_CONV tm in + EQ_MP (beta (concl th)) th;; + +(* Converts a list with two elements into a pair *) +let pair_of_list = function + | [a; b] -> a, b + | _ -> failwith "pair_of_list: list does not contain exactly two element";; + +(* Applies f to arg n times and returns the total execution time *) +let test n f arg = + let start = Unix.gettimeofday() in + let result = f arg in + for i = 1 to n - 1 do + let _ = f arg in () + done; + result, Unix.gettimeofday() -. start;; + +(* Do not use when msg is not a fixed string *) +let assert_msg (cond, msg) = + if cond then () else failwith msg;; + +(* Generates a power function for the given binary operation *) +let gen_pow op id = + let ( * ) = op in + let rec pow n x = + if n <= 0 then id + else if n = 1 then x + else if n land 1 = 1 then + x * pow (n - 1) x + else + let t = pow (n lsr 1) x in + t * t in + pow;; + +let rec shape_list n list = + if length list <= n then [list] + else + let l1, l2 = chop_list n list in + l1 :: shape_list n l2;; + +(* map2 which works for lists of any size (no requirement |l1| = |l2|) *) +let rec my_map2 f l1 l2 = + match l1 with + | [] -> [] + | (h1::t1) -> + (match l2 with + | [] -> [] + | (h2::t2) -> (f h1 h2) :: my_map2 f t1 t2);; + +end;; diff --git a/Formal_ineqs/misc/misc_vars.hl b/Formal_ineqs/misc/misc_vars.hl index fe5497f9..bf2e4e8d 100644 --- a/Formal_ineqs/misc/misc_vars.hl +++ b/Formal_ineqs/misc/misc_vars.hl @@ -1,159 +1,188 @@ -(* =========================================================== *) -(* Commonly used variables and constants *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -module Misc_vars = struct - -(* bool variables *) -let s_var_bool = `s:bool` and - s1_var_bool = `s1:bool` and - s2_var_bool = `s2:bool`;; - -(* num variables *) -let n_var_num = `n:num` and - m_var_num = `m:num` and - t_var_num = `t:num` and - p_var_num = `p:num` and - q_var_num = `q:num` and - k_var_num = `k:num` and - e_var_num = `e:num` and - e1_var_num = `e1:num` and - e2_var_num = `e2:num` and - r_var_num = `r:num` and - r1_var_num = `r1:num` and - r2_var_num = `r2:num` and - n1_var_num = `n1:num` and - n2_var_num = `n2:num` and - m1_var_num = `m1:num` and - m2_var_num = `m2:num` and - x_var_num = `x:num` and - y_var_num = `y:num` and - i_var_num = `i:num` and - j_var_num = `j:num`;; - -(* real variables *) -let x_var_real = `x : real` and - y_var_real = `y : real` and - z_var_real = `z : real` and - w_var_real = `w : real` and - a_var_real = `a : real` and - b_var_real = `b : real` and - a1_var_real = `a1 : real` and - a2_var_real = `a2 : real` and - b1_var_real = `b1 : real` and - b2_var_real = `b2 : real` and - c_var_real = `c : real` and - d_var_real = `d : real` and - h_var_real = `h : real` and - m_var_real = `m : real` and - n_var_real = `n : real` and - r_var_real = `r : real` and - x1_var_real = `x1 : real` and - x2_var_real = `x2 : real` and - y1_var_real = `y1 : real` and - y2_var_real = `y2 : real` and - f1_var_real = `f1 : real` and - f2_var_real = `f2 : real` and - f_var_fun = `f : real->real` and - g_var_fun = `g : real->real` and - f1_var_fun = `f1 : real->real` and - f2_var_fun = `f2 : real->real` and - int_var = `int : real#real` and - f_bounds_var = `f_bounds : real#real` and - df_bounds_var = `df_bounds : real#real` and - dd_bounds_var = `dd_bounds : real#real` and - x_lo_var = `x_lo : real` and - x_hi_var = `x_hi : real` and - low_var_real = `low : real` and - high_var_real = `high : real` and - lo_var_real = `lo : real` and - hi_var_real = `hi : real` and - dd_var_real = `dd : real` and - df_lo_var = `df_lo : real` and - df_hi_var = `df_hi : real` and - df_var_real = `df : real` and - f_lo_var = `f_lo : real` and - f_hi_var = `f_hi : real` and - w1_var_real = `w1 : real` and - w2_var_real = `w2 : real` and - t_var_real = `t : real` and - g_bounds_var = `g_bounds : real#real` and - dg_bounds_var = `dg_bounds : real#real` and - bounds_var = `bounds : real#real` and - d_bounds_var = `d_bounds : real#real` and - x0_var_real = `x0 : real` and - z0_var_real = `z0 : real` and - w0_var_real = `w0 : real` and - error_var = `error : real` and - d_bounds_list_var = `d_bounds_list : (real#real)list` and - dd_bounds_list_var = `dd_bounds_list : ((real#real)list)list` and - df_bounds_list_var = `df_bounds_list : (real#real)list` and - dd_list_var = `dd_list : (real#real)list` and - t_var_real_list = `t:(real)list` and - x_var_real_list = `x:(real)list` and - y_var_real_list = `y:(real)list` and - z_var_real_list = `z:(real)list` and - w_var_real_list = `w:(real)list` and - yw_var = `yw : (real#real)list` and - xz_var = `xz : (real#real)list` and - xz_pair_var = `xz : real#real` and - yw_pair_var = `yw : real#real` and - list_var_real_pair = `list : (real#real)list`;; - -(* bool constants *) -let t_const = `T` and - f_const = `F`;; - -(* num constants *) -let zero_const = `_0` and - zero_num = `0`;; - -(* num operations *) -let add_op_num = `(+) : num->num->num` and - sub_op_num = `(-) : num->num->num` and - mul_op_num = `( * ) : num->num->num` and - eq_op_num = `(=) : num->num->bool` and - le_op_num = `(<=) : num->num->bool` and - lt_op_num = `(<) : num->num->bool` and - div_op_num = `(DIV): num->num->num` and - pre_op_num = `PRE: num->num` and - suc_op_num = `SUC : num->num`;; - - -(* real constants *) -let real_empty_list = `[]:(real)list`;; - -(* real operations *) -let add_op_real = `(+) : real->real->real` and - mul_op_real = `( * ) : real->real->real` and - sub_op_real = `(-) : real->real->real` and - div_op_real = `(/) :real->real->real` and - inv_op_real = `inv : real->real` and - neg_op_real = `(--) : real->real` and - eq_op_real = `(=) : real->real->bool` and - lt_op_real = `(<) : real->real->bool` and - le_op_real = `(<=):real->real->bool` and - ge_op_real = `(>=):real->real->bool` and - amp_op_real = `(&) : num->real` and - pow_op_real = `(pow) : real->num->real`;; - -(* types *) -let real_ty = `:real` and - real_list_ty = `:(real)list` and - real_pair_ty = `:real#real` and - real_pair_list_ty = `:(real#real)list` and - num_type = `:num` and - nty = `:N`;; - - -(* Simple operations *) -let mk_real_list tms = mk_list (tms, real_ty);; - -let mk_names n prefix = map (fun i -> prefix^(string_of_int i)) (1--n);; - -let mk_real_vars n prefix = map (C (curry mk_var) real_ty) (mk_names n prefix);; - - -end;; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Commonly used variables and constants *) +(* -------------------------------------------------------------------------- *) + +module Misc_vars = struct + +(* A performance improvement trick from calc_num.ml *) +let standardize, standardize_tm, standardize_list, standardize_reset, standardize_enable = + (* TODO: a custom hash table module is required for good hash values *) + let cache = Hashtbl.create 100 in + let enabled = ref true in + let find t = + try Hashtbl.find cache t + with Not_found -> Hashtbl.add cache t t; t in + let rec replace tm = + match tm with + | Var _ | Const _ -> find tm + | Comb (s, t) -> (* find @@*) mk_comb (replace s, replace t) + | Abs (v, b) -> (* find @@*) mk_abs (replace v, replace b) in + let st th = + let concl' = replace (concl th) in + EQ_MP (REFL concl') th in + let get () = Hashtbl.fold (fun _ t ts -> t :: ts) cache [] in + let reset () = Hashtbl.clear cache in + let enable flag = enabled := flag in + (fun th -> if !enabled then st th else th), + (fun tm -> if !enabled then replace tm else tm), + get, reset, enable;; + +(* bool variables *) +let s_var_bool = standardize_tm `s:bool` and + s1_var_bool = standardize_tm `s1:bool` and + s2_var_bool = standardize_tm `s2:bool`;; + +(* num variables *) +let n_var_num = standardize_tm `n:num` and + m_var_num = standardize_tm `m:num` and + t_var_num = standardize_tm `t:num` and + p_var_num = standardize_tm `p:num` and + q_var_num = standardize_tm `q:num` and + k_var_num = standardize_tm `k:num` and + e_var_num = standardize_tm `e:num` and + e1_var_num = standardize_tm `e1:num` and + e2_var_num = standardize_tm `e2:num` and + r_var_num = standardize_tm `r:num` and + r1_var_num = standardize_tm `r1:num` and + r2_var_num = standardize_tm `r2:num` and + n1_var_num = standardize_tm `n1:num` and + n2_var_num = standardize_tm `n2:num` and + m1_var_num = standardize_tm `m1:num` and + m2_var_num = standardize_tm `m2:num` and + x_var_num = standardize_tm `x:num` and + y_var_num = standardize_tm `y:num` and + i_var_num = standardize_tm `i:num` and + j_var_num = standardize_tm `j:num`;; + +(* real variables *) +let x_var_real = standardize_tm `x : real` and + y_var_real = standardize_tm `y : real` and + z_var_real = standardize_tm `z : real` and + w_var_real = standardize_tm `w : real` and + a_var_real = standardize_tm `a : real` and + b_var_real = standardize_tm `b : real` and + a1_var_real = standardize_tm `a1 : real` and + a2_var_real = standardize_tm `a2 : real` and + b1_var_real = standardize_tm `b1 : real` and + b2_var_real = standardize_tm `b2 : real` and + c_var_real = standardize_tm `c : real` and + d_var_real = standardize_tm `d : real` and + h_var_real = standardize_tm `h : real` and + m_var_real = standardize_tm `m : real` and + n_var_real = standardize_tm `n : real` and + r_var_real = standardize_tm `r : real` and + x1_var_real = standardize_tm `x1 : real` and + x2_var_real = standardize_tm `x2 : real` and + y1_var_real = standardize_tm `y1 : real` and + y2_var_real = standardize_tm `y2 : real` and + f1_var_real = standardize_tm `f1 : real` and + f2_var_real = standardize_tm `f2 : real` and + f_var_fun = standardize_tm `f : real->real` and + g_var_fun = standardize_tm `g : real->real` and + f1_var_fun = standardize_tm `f1 : real->real` and + f2_var_fun = standardize_tm `f2 : real->real` and + int_var = standardize_tm `int : real#real` and + f_bounds_var = standardize_tm `f_bounds : real#real` and + df_bounds_var = standardize_tm `df_bounds : real#real` and + dd_bounds_var = standardize_tm `dd_bounds : real#real` and + x_lo_var = standardize_tm `x_lo : real` and + x_hi_var = standardize_tm `x_hi : real` and + low_var_real = standardize_tm `low : real` and + high_var_real = standardize_tm `high : real` and + lo_var_real = standardize_tm `lo : real` and + hi_var_real = standardize_tm `hi : real` and + dd_var_real = standardize_tm `dd : real` and + df_lo_var = standardize_tm `df_lo : real` and + df_hi_var = standardize_tm `df_hi : real` and + df_var_real = standardize_tm `df : real` and + f_lo_var = standardize_tm `f_lo : real` and + f_hi_var = standardize_tm `f_hi : real` and + w1_var_real = standardize_tm `w1 : real` and + w2_var_real = standardize_tm `w2 : real` and + t_var_real = standardize_tm `t : real` and + g_bounds_var = standardize_tm `g_bounds : real#real` and + dg_bounds_var = standardize_tm `dg_bounds : real#real` and + bounds_var = standardize_tm `bounds : real#real` and + d_bounds_var = standardize_tm `d_bounds : real#real` and + x0_var_real = standardize_tm `x0 : real` and + z0_var_real = standardize_tm `z0 : real` and + w0_var_real = standardize_tm `w0 : real` and + error_var = standardize_tm `error : real` and + d_bounds_list_var = standardize_tm `d_bounds_list : (real#real)list` and + dd_bounds_list_var = standardize_tm `dd_bounds_list : ((real#real)list)list` and + df_bounds_list_var = standardize_tm `df_bounds_list : (real#real)list` and + dd_list_var = standardize_tm `dd_list : (real#real)list` and + t_var_real_list = standardize_tm `t:(real)list` and + x_var_real_list = standardize_tm `x:(real)list` and + y_var_real_list = standardize_tm `y:(real)list` and + z_var_real_list = standardize_tm `z:(real)list` and + w_var_real_list = standardize_tm `w:(real)list` and + yw_var = standardize_tm `yw : (real#real)list` and + xz_var = standardize_tm `xz : (real#real)list` and + xz_pair_var = standardize_tm `xz : real#real` and + yw_pair_var = standardize_tm `yw : real#real` and + list_var_real_pair = standardize_tm `list : (real#real)list`;; + +(* bool constants *) +let t_const = standardize_tm `T` and + f_const = standardize_tm `F`;; + +(* num constants *) +let zero_const = standardize_tm `_0` and + zero_num = standardize_tm `0`;; + +(* num operations *) +let add_op_num = standardize_tm `(+) : num->num->num` and + sub_op_num = standardize_tm `(-) : num->num->num` and + mul_op_num = standardize_tm `( * ) : num->num->num` and + eq_op_num = standardize_tm `(=) : num->num->bool` and + le_op_num = standardize_tm `(<=) : num->num->bool` and + lt_op_num = standardize_tm `(<) : num->num->bool` and + div_op_num = standardize_tm `(DIV): num->num->num` and + pre_op_num = standardize_tm `PRE: num->num` and + suc_op_num = standardize_tm `SUC : num->num`;; + + +(* real constants *) +let real_empty_list = standardize_tm `[]:(real)list`;; + +(* real operations *) +let add_op_real = standardize_tm `(+) : real->real->real` and + mul_op_real = standardize_tm `( * ) : real->real->real` and + sub_op_real = standardize_tm `(-) : real->real->real` and + div_op_real = standardize_tm `(/) :real->real->real` and + inv_op_real = standardize_tm `inv : real->real` and + neg_op_real = standardize_tm `(--) : real->real` and + eq_op_real = standardize_tm `(=) : real->real->bool` and + lt_op_real = standardize_tm `(<) : real->real->bool` and + le_op_real = standardize_tm `(<=):real->real->bool` and + ge_op_real = standardize_tm `(>=):real->real->bool` and + amp_op_real = standardize_tm `(&) : num->real` and + pow_op_real = standardize_tm `(pow) : real->num->real`;; + +(* types *) +let real_ty = `:real` and + real_list_ty = `:(real)list` and + real_pair_ty = `:real#real` and + real_pair_list_ty = `:(real#real)list` and + num_type = `:num` and + nty = `:N`;; + + +(* Simple operations *) +let mk_real_list tms = mk_list (tms, real_ty);; + +let mk_names n prefix = map (fun i -> prefix^(string_of_int i)) (1--n);; + +let mk_real_vars n prefix = map (C (curry mk_var) real_ty) (mk_names n prefix);; + + +end;; diff --git a/Formal_ineqs/misc/report.hl b/Formal_ineqs/misc/report.hl index 2d6f01ae..ea8f7cb8 100644 --- a/Formal_ineqs/misc/report.hl +++ b/Formal_ineqs/misc/report.hl @@ -1,14 +1,17 @@ -(* =========================================================== *) -(* Report functions *) -(* Author: Thomas C. Hales *) -(* Date: 2011-08-21 *) -(* =========================================================== *) - -(* port of error.cc - basic procedures to print messages to the standard output - and to count errors. - -*) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2011 Thomas C. Hales *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Report functions *) +(* port of error.cc *) +(* basic procedures to print messages to the standard output *) +(* and to count errors. *) +(* -------------------------------------------------------------------------- *) module Report = struct @@ -26,7 +29,7 @@ let (get_corner_count,reset_corner_count,inc_corner_count) = ((fun () -> !corner_count),(fun () -> corner_count := 0), (fun () -> corner_count:= !corner_count + 1));; -let diagnostic_string () = +let diagnostic_string () = let d = get_error_count() in if (d>0) then Printf.sprintf "(errors %d)" (get_error_count()) else "(no errors)";; @@ -35,14 +38,14 @@ let report s = let report_timed s = report (s^" "^(time_string()));; -let report_error = +let report_error = let error_max = 25 in (* was 200, recurse.cc had a separate counter limit at 25 *) fun s -> let ec = get_error_count() in - (inc_error_count(); report_timed (Printf.sprintf "error(%d) --\n%s" ec s); + (inc_error_count(); report_timed (Printf.sprintf "error(%d) --\n%s" ec s); ignore(get_error_count() < error_max || raise Fatal));; - -let report_fatal s = + +let report_fatal s = ( inc_error_count(); report_timed ("error --\n"^s); raise Fatal);; end;; diff --git a/Formal_ineqs/taylor/m_taylor.hl b/Formal_ineqs/taylor/m_taylor.hl index c24d65a7..2cfa0c62 100644 --- a/Formal_ineqs/taylor/m_taylor.hl +++ b/Formal_ineqs/taylor/m_taylor.hl @@ -1,1491 +1,1497 @@ -(* =========================================================== *) -(* Formal taylor intervals *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -needs "arith/more_float.hl";; -needs "arith/eval_interval.hl";; -needs "list/list_conversions.hl";; -needs "list/list_float.hl";; -needs "list/more_list.hl";; -needs "misc/misc_vars.hl";; - -needs "taylor/theory/taylor_interval-compiled.hl";; -needs "taylor/theory/multivariate_taylor-compiled.hl";; - - -module M_taylor = struct - -open Ssreflect;; -open Ssrfun;; -open Ssrbool;; -open Ssrnat;; -open Taylor_interval;; -open Multivariate_taylor;; -open Misc_functions;; -open Arith_float;; -open More_float;; -open Float_theory;; -open Eval_interval;; -open List_conversions;; -open List_float;; -open More_list;; -open Interval_arith;; -open Misc_vars;; - -let MY_RULE = UNDISCH_ALL o PURE_REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL;; -let MY_RULE_NUM = UNDISCH_ALL o Arith_nat.NUMERALS_TO_NUM o PURE_REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL;; -let MY_RULE_FLOAT = UNDISCH_ALL o Arith_nat.NUMERALS_TO_NUM o - PURE_REWRITE_RULE[FLOAT_OF_NUM; min_exp_def; GSYM IMP_IMP] o SPEC_ALL;; - - - -let max_dim = 8;; - -let inst_first_type_var ty th = - let ty_vars = type_vars_in_term (concl th) in - if ty_vars = [] then - failwith "inst_first_type: no type variables in the theorem" - else - INST_TYPE [ty, hd ty_vars] th;; - - -let float0 = mk_float 0 0 and - interval0 = mk_float_interval_small_num 0;; - - -let has_size_array = Array.init (max_dim + 1) - (fun i -> match i with - | 0 -> TRUTH - | 1 -> HAS_SIZE_1 - | _ -> HAS_SIZE_DIMINDEX_RULE(mk_finty(num i)));; - -let dimindex_array = Array.init (max_dim + 1) - (fun i -> if i < 1 then TRUTH else MATCH_MP DIMINDEX_UNIQUE has_size_array.(i));; - -let n_type_array = Array.init (max_dim + 1) - (fun i -> if i < 1 then bool_ty else - let dimindex_th = dimindex_array.(i) in - (hd o snd o dest_type o snd o dest_const o rand o lhand o concl) dimindex_th);; - -let n_vector_type_array = Array.init (max_dim + 1) - (fun i -> if i < 1 then bool_ty else mk_type ("cart", [real_ty; n_type_array.(i)]));; - - -let x_var_names = Array.init (max_dim + 1) (fun i -> "x"^(string_of_int i)) and - y_var_names = Array.init (max_dim + 1) (fun i -> "y"^(string_of_int i)) and - z_var_names = Array.init (max_dim + 1) (fun i -> "z"^(string_of_int i)) and - w_var_names = Array.init (max_dim + 1) (fun i -> "w"^(string_of_int i));; - -let x_vars_array = Array.init (max_dim + 1) (fun i -> mk_var(x_var_names.(i), real_ty)) and - y_vars_array = Array.init (max_dim + 1) (fun i -> mk_var(y_var_names.(i), real_ty)) and - z_vars_array = Array.init (max_dim + 1) (fun i -> mk_var(z_var_names.(i), real_ty)) and - w_vars_array = Array.init (max_dim + 1) (fun i -> mk_var(w_var_names.(i), real_ty));; - -let df_vars_array = Array.init (max_dim + 1) (fun i -> mk_var ("df"^(string_of_int i), real_pair_ty));; -let dd_vars_array = Array.init (max_dim + 1) (fun i -> - Array.init (max_dim + 1) (fun j -> mk_var ("dd"^(string_of_int i)^(string_of_int j), real_pair_ty)));; - -let dest_vector = dest_list o rand;; - -let mk_vector list_tm = - let n = (length o dest_list) list_tm in - let ty = (hd o snd o dest_type o type_of) list_tm in - let vec = mk_const ("vector", [ty, aty; n_type_array.(n), nty]) in - mk_comb (vec, list_tm);; - -let mk_vector_list list = - mk_vector (mk_list (list, type_of (hd list)));; - -let el_thms_array = - let el_tm = `EL : num->(A)list->A` in - let gen0 n = - let e_list = mk_list (map (fun i -> mk_var ("e"^(string_of_int i), aty)) (1--n), aty) in - let el0_th = REWRITE_CONV[EL; HD] (mk_binop el_tm `0` e_list) in - Array.make n el0_th in - let array = Array.init (max_dim + 1) gen0 in - let gen_i n i = - let e_list = (rand o lhand o concl) array.(n).(i) in - let prev_thm = array.(n - 1).(i - 1) in - let i_tm = mk_small_numeral i in - let prev_i = num_CONV i_tm in - let el_th = REWRITE_CONV[prev_i; EL; HD; TL; prev_thm] (mk_binop el_tm i_tm e_list) in - array.(n).(i) <- el_th in - let _ = map (fun n -> map (fun i -> gen_i n i) (1--(n - 1))) (2--max_dim) in - array;; - - -let VECTOR_COMPONENT = prove(`!l i. i IN 1..dimindex (:N) ==> - (vector l:A^N)$i = EL (i - 1) l`, -REWRITE_TAC[IN_NUMSEG] THEN REPEAT GEN_TAC THEN DISCH_TAC THEN REWRITE_TAC[vector] THEN - MATCH_MP_TAC LAMBDA_BETA THEN ASM_REWRITE_TAC[]);; - -let gen_comp_thm n i = - let i_tm = mk_small_numeral i and - x_list = mk_list (map (fun i -> mk_var("x"^(string_of_int i), aty)) (1--n), aty) in - let th0 = (ISPECL [x_list; i_tm] o inst_first_type_var (n_type_array.(n))) VECTOR_COMPONENT in - let th1 = (CONV_RULE NUM_REDUCE_CONV o REWRITE_RULE[IN_NUMSEG; dimindex_array.(n)]) th0 in - REWRITE_RULE[el_thms_array.(n).(i - 1)] th1;; - -let comp_thms_array = Array.init (max_dim + 1) - (fun n -> Array.init (n + 1) - (fun i -> if i < 1 || n < 1 then TRUTH else gen_comp_thm n i));; - - -(************************************) -(* m_cell_domain *) - -let ALL2_ALL_ZIP = prove(`!(P:A->B->bool) l1 l2. LENGTH l1 = LENGTH l2 ==> - (ALL2 P l1 l2 <=> ALL (\p. P (FST p) (SND p)) (ZIP l1 l2))`, - GEN_TAC THEN LIST_INDUCT_TAC THENL - [ - GEN_TAC THEN REWRITE_TAC[LENGTH; EQ_SYM_EQ; LENGTH_EQ_NIL] THEN - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN - REWRITE_TAC[ZIP; ALL2; ALL]; - ALL_TAC - ] THEN - - LIST_INDUCT_TAC THEN REWRITE_TAC[LENGTH] THENL [ARITH_TAC; ALL_TAC] THEN - REWRITE_TAC[eqSS] THEN DISCH_TAC THEN - REWRITE_TAC[ALL2; ZIP; ALL] THEN - FIRST_X_ASSUM (new_rewrite [] []) THEN ASM_REWRITE_TAC[]);; - -let EL_ZIP = prove(`!(l1:(A)list) (l2:(B)list) i. LENGTH l1 = LENGTH l2 /\ i < LENGTH l1 ==> - EL i (ZIP l1 l2) = (EL i l1, EL i l2)`, - LIST_INDUCT_TAC THEN LIST_INDUCT_TAC THEN REWRITE_TAC[ZIP; LENGTH] THEN TRY ARITH_TAC THEN - case THEN REWRITE_TAC[EL; HD; TL] THEN GEN_TAC THEN - REWRITE_TAC[eqSS; ARITH_RULE `SUC n < SUC x <=> n < x`] THEN STRIP_TAC THEN - FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_REWRITE_TAC[]);; - -let LENGTH_ZIP = prove(`!l1 l2. LENGTH l1 = LENGTH l2 ==> LENGTH (ZIP l1 l2) = LENGTH l1`, - LIST_INDUCT_TAC THEN LIST_INDUCT_TAC THEN REWRITE_TAC[ZIP; LENGTH] THEN TRY ARITH_TAC THEN - REWRITE_TAC[eqSS] THEN DISCH_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_REWRITE_TAC[]);; - - - -let test_domain_xi = new_definition - `test_domain_xi xz yw <=> FST xz <= FST yw /\ FST yw <= SND xz /\ - FST yw - FST xz <= SND yw /\ SND xz - FST yw <= SND yw`;; - - -let MK_CELL_DOMAIN = prove(`!xz (yw:(real#real)list) x z y w. - LENGTH x = dimindex (:N) /\ LENGTH z = dimindex (:N) /\ - LENGTH y = dimindex (:N) /\ LENGTH w = dimindex (:N) /\ - ZIP y w = yw /\ ZIP x z = xz /\ - ALL2 test_domain_xi xz yw ==> - m_cell_domain (vector x, vector z:real^N) (vector y) (vector w)`, - REPEAT GEN_TAC THEN STRIP_TAC THEN POP_ASSUM MP_TAC THEN - SUBGOAL_THEN `LENGTH (xz:(real#real)list) = dimindex (:N) /\ LENGTH (yw:(real#real)list) = dimindex (:N)` ASSUME_TAC THENL - [ - EXPAND_TAC "yw" THEN EXPAND_TAC "xz" THEN - REPEAT (new_rewrite [] [] LENGTH_ZIP) THEN ASM_REWRITE_TAC[]; - ALL_TAC - ] THEN - rewrite [] [] ALL2_ALL_ZIP THEN ASM_REWRITE_TAC[m_cell_domain; GSYM ALL_EL] THEN DISCH_TAC THEN - REWRITE_TAC[m_cell_domain] THEN GEN_TAC THEN DISCH_TAC THEN - REPEAT (new_rewrite [] [] VECTOR_COMPONENT) THEN ASM_REWRITE_TAC[] THEN - ABBREV_TAC `j = i - 1` THEN - SUBGOAL_THEN `j < dimindex (:N)` ASSUME_TAC THENL - [ - POP_ASSUM MP_TAC THEN POP_ASSUM MP_TAC THEN REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; - ALL_TAC - ] THEN - FIRST_X_ASSUM (MP_TAC o SPEC `j:num`) THEN REWRITE_TAC[test_domain_xi] THEN - rewrite [] [] LENGTH_ZIP THEN ASM_REWRITE_TAC[] THEN - rewrite [] [] EL_ZIP THEN ASM_REWRITE_TAC[] THEN - EXPAND_TAC "xz" THEN EXPAND_TAC "yw" THEN - REPEAT (new_rewrite [] [] EL_ZIP) THEN ASM_REWRITE_TAC[] THEN - ARITH_TAC);; - - - -(* array of theorems *) -let mk_m_domain_array = - let mk_m_domain n = - let dimindex_th = dimindex_array.(n) in - let n_ty = (hd o snd o dest_type o snd o dest_const o rand o lhand o concl) dimindex_th in - let nty = `:N` in - (UNDISCH_ALL o REWRITE_RULE[float0_eq] o DISCH_ALL o RULE o - REWRITE_RULE[dimindex_th] o INST_TYPE[n_ty, nty]) MK_CELL_DOMAIN in - Array.init (max_dim + 1) (fun i -> if i < 1 then TRUTH else mk_m_domain i);; - - -let TEST_DOMAIN_XI' = (EQT_INTRO o RULE o prove)(`xz = (x,z) /\ yw = (y,w) /\ - x <= y /\ y <= z /\ y - x <= w1 /\ z - y <= w2 /\ w1 <= w /\ w2 <= w ==> test_domain_xi xz yw`, - SIMP_TAC[test_domain_xi] THEN REAL_ARITH_TAC);; - - -let eval_test_domain_xi pp test_domain_tm = - let ltm, yw = dest_comb test_domain_tm in - let xz = rand ltm in - let x, z = dest_pair xz and - y, w = dest_pair yw in - let (<=) = (fun t1 t2 -> EQT_ELIM (float_le t1 t2)) and - (-) = float_sub_hi pp in - let x_le_y = x <= y and - y_le_z = y <= z and - yx_le_w1 = y - x and - zy_le_w2 = z - y in - let w1 = (rand o concl) yx_le_w1 and - w2 = (rand o concl) zy_le_w2 in - let w1_le_w = w1 <= w and - w2_le_w = w2 <= w in - (MY_PROVE_HYP (REFL xz) o MY_PROVE_HYP (REFL yw) o - MY_PROVE_HYP x_le_y o MY_PROVE_HYP y_le_z o - MY_PROVE_HYP yx_le_w1 o MY_PROVE_HYP zy_le_w2 o - MY_PROVE_HYP w1_le_w o MY_PROVE_HYP w2_le_w o - INST[x, x_var_real; y, y_var_real; z, z_var_real; w, w_var_real; - w1, w1_var_real; w2, w2_var_real; - xz, xz_pair_var; yw, yw_pair_var]) TEST_DOMAIN_XI';; - - -(* mk_m_center_domain *) -let mk_m_center_domain n pp x_list_tm z_list_tm = - let x_list = dest_list x_list_tm and - z_list = dest_list z_list_tm in - let y_list = - let ( * ) = (fun t1 t2 -> (rand o concl) (float_mul_eq t1 t2)) and - (+) = (fun t1 t2 -> (rand o concl) (float_add_hi pp t1 t2)) in - map2 (fun x y -> if x = y then x else float_inv2 * (x + y)) x_list z_list in - - let w_list = - let (-) = (fun t1 t2 -> (rand o concl) (float_sub_hi pp t1 t2)) and - max = (fun t1 t2 -> (rand o concl) (float_max t1 t2)) in - let w1 = map2 (-) y_list x_list and - w2 = map2 (-) z_list y_list in - map2 max w1 w2 in - - let y_list_tm = mk_list (y_list, real_ty) and - w_list_tm = mk_list (w_list, real_ty) in - - let yw_zip_th = eval_zip y_list_tm w_list_tm and - xz_zip_th = eval_zip x_list_tm z_list_tm in - - let yw_list_tm = (rand o concl) yw_zip_th and - xz_list_tm = (rand o concl) xz_zip_th in - - let len_x_th = eval_length x_list_tm and - len_z_th = eval_length z_list_tm and - len_y_th = eval_length y_list_tm and - len_w_th = eval_length w_list_tm in - let th0 = (MY_PROVE_HYP len_x_th o MY_PROVE_HYP len_z_th o - MY_PROVE_HYP len_y_th o MY_PROVE_HYP len_w_th o - MY_PROVE_HYP yw_zip_th o MY_PROVE_HYP xz_zip_th o - INST[x_list_tm, x_var_real_list; z_list_tm, z_var_real_list; - y_list_tm, y_var_real_list; w_list_tm, w_var_real_list; - yw_list_tm, yw_var; xz_list_tm, xz_var]) mk_m_domain_array.(n) in - let all_th = (EQT_ELIM o all2_conv_univ (eval_test_domain_xi pp) o hd o hyp) th0 in - MY_PROVE_HYP all_th th0;; - - - -(***********************) - -let MK_M_TAYLOR_INTERVAL' = (RULE o MATCH_MP iffRL o SPEC_ALL) m_taylor_interval;; - - -let get_types_and_vars n = - let ty = n_type_array.(n) and - xty = n_vector_type_array.(n) in - let x_var = mk_var ("x", xty) and - f_var = mk_var ("f", mk_fun_ty xty real_ty) and - y_var = mk_var ("y", xty) and - w_var = mk_var ("w", xty) and - domain_var = mk_var ("domain", mk_type ("prod", [xty; xty])) in - ty, xty, x_var, f_var, y_var, w_var, domain_var;; - - - -let dest_m_cell_domain domain_tm = - let lhs, w_tm = dest_comb domain_tm in - let lhs2, y_tm = dest_comb lhs in - rand lhs2, y_tm, w_tm;; - - -(**************************************************) - -(* Given a variable of the type `:real^N`, returns the number N *) -let get_dim = Num.int_of_num o dest_finty o hd o tl o snd o dest_type o type_of;; - - -(**********************) -(* eval_m_taylor_poly *) - -let partial_pow = prove(`!i f n (y:real^N). lift o f differentiable at y ==> - partial i (\x. f x pow n) y = &n * f y pow (n - 1) * partial i f y`, - REPEAT STRIP_TAC THEN - SUBGOAL_THEN `(\x:real^N. f x pow n) = (\t. t pow n) o f` (fun th -> REWRITE_TAC[th]) THENL - [ - ONCE_REWRITE_TAC[GSYM eq_ext] THEN REWRITE_TAC[o_THM]; - ALL_TAC - ] THEN - new_rewrite [] [] partial_uni_compose THENL - [ - ASM_REWRITE_TAC[] THEN - new_rewrite [] [] REAL_DIFFERENTIABLE_POW_ATREAL THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID]; - ALL_TAC - ] THEN - new_rewrite [] [] derivative_pow THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID; derivative_x] THEN - REAL_ARITH_TAC);; - - -let nth_diff2_pow = prove(`!n y. nth_diff_strong 2 (\x. x pow n) y`, - REWRITE_TAC[nth_diff_strong2_eq] THEN REPEAT GEN_TAC THEN - EXISTS_TAC `(:real)` THEN REWRITE_TAC[REAL_OPEN_UNIV; IN_UNIV] THEN GEN_TAC THEN - new_rewrite [] [] REAL_DIFFERENTIABLE_POW_ATREAL THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID] THEN - MATCH_MP_TAC differentiable_local THEN - EXISTS_TAC `\x. &n * x pow (n - 1)` THEN EXISTS_TAC `(:real)` THEN - REWRITE_TAC[REAL_OPEN_UNIV; IN_UNIV] THEN - new_rewrite [] [] REAL_DIFFERENTIABLE_MUL_ATREAL THEN REWRITE_TAC[REAL_DIFFERENTIABLE_CONST] THENL - [ - new_rewrite [] [] REAL_DIFFERENTIABLE_POW_ATREAL THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID]; - ALL_TAC - ] THEN - GEN_TAC THEN new_rewrite [] [] derivative_pow THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID] THEN - REWRITE_TAC[derivative_x; REAL_MUL_RID]);; - - -let diff2c_pow = prove(`!f n (x:real^N). diff2c f x ==> diff2c (\x. f x pow n) x`, - REPEAT STRIP_TAC THEN - SUBGOAL_THEN `(\x:real^N. f x pow n) = (\t. t pow n) o f` (fun th -> REWRITE_TAC[th]) THENL - [ - ONCE_REWRITE_TAC[GSYM eq_ext] THEN REWRITE_TAC[o_THM]; - ALL_TAC - ] THEN - apply_tac diff2c_uni_compose THEN ASM_REWRITE_TAC[nth_diff2_pow] THEN - REWRITE_TAC[nth_derivative2] THEN - SUBGOAL_THEN `!n. derivative (\t. t pow n) = (\t. &n * t pow (n - 1))` ASSUME_TAC THENL - [ - GEN_TAC THEN REWRITE_TAC[FUN_EQ_THM] THEN GEN_TAC THEN - new_rewrite [] [] derivative_pow THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID] THEN - REWRITE_TAC[derivative_x; REAL_MUL_RID]; - ALL_TAC - ] THEN - ASM_REWRITE_TAC[] THEN - SUBGOAL_THEN `!n. derivative (\t. &n * t pow (n - 1)) = (\t. &n * derivative (\t. t pow (n - 1)) t)` ASSUME_TAC THENL - [ - GEN_TAC THEN REWRITE_TAC[FUN_EQ_THM] THEN GEN_TAC THEN - new_rewrite [] [] derivative_scale THEN REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_DIFFERENTIABLE_POW_ATREAL THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID]; - ALL_TAC - ] THEN - ASM_REWRITE_TAC[] THEN - REPEAT (MATCH_MP_TAC REAL_CONTINUOUS_LMUL) THEN - MATCH_MP_TAC REAL_CONTINUOUS_POW THEN - REWRITE_TAC[REAL_CONTINUOUS_AT_ID]);; - - -let diff2c_domain_pow = prove(`!f n domain. diff2c_domain domain f ==> - diff2c_domain domain (\x. f x pow n)`, - REWRITE_TAC[diff2c_domain] THEN REPEAT STRIP_TAC THEN ASM_SIMP_TAC[diff2c_pow]);; - - -let diff2c_domain_tm = `diff2c_domain domain`;; - -let rec gen_diff2c_domain_poly poly_tm = - let x_var, expr = dest_abs poly_tm in - let n = get_dim x_var in - let diff2c_tm = mk_icomb (diff2c_domain_tm, poly_tm) in - if frees expr = [] then - (* const *) - (SPEC_ALL o ISPEC expr o inst_first_type_var (n_type_array.(n))) diff2c_domain_const - else - let lhs, r_tm = dest_comb expr in - if lhs = neg_op_real then - (* -- *) - let r_th = gen_diff2c_domain_poly (mk_abs (x_var, r_tm)) in - prove(diff2c_tm, MATCH_MP_TAC diff2c_domain_neg THEN REWRITE_TAC[r_th]) - else - let op, l_tm = dest_comb lhs in - let name = (fst o dest_const) op in - if name = "$" then - (* x$k *) - let dim_th = dimindex_array.(n) in - prove(diff2c_tm, MATCH_MP_TAC diff2c_domain_x THEN - REWRITE_TAC[IN_NUMSEG; dim_th] THEN ARITH_TAC) - else - let l_th = gen_diff2c_domain_poly (mk_abs (x_var, l_tm)) in - if name = "real_pow" then - (* f pow n *) - prove(diff2c_tm, MATCH_MP_TAC diff2c_domain_pow THEN REWRITE_TAC[l_th]) - else - let r_th = gen_diff2c_domain_poly (mk_abs (x_var, r_tm)) in - prove(diff2c_tm, - MAP_FIRST apply_tac [diff2c_domain_add; diff2c_domain_sub; diff2c_domain_mul] THEN - REWRITE_TAC[l_th; r_th]);; - - -let gen_diff2c_poly = - let th_imp = prove(`!f. (!domain. diff2c_domain domain f) ==> !x:real^N. diff2c f x`, - REWRITE_TAC[diff2c_domain] THEN REPEAT STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN - EXISTS_TAC `(x:real^N, x:real^N)` THEN - REWRITE_TAC[INTERVAL_SING; IN_SING]) in - fun poly_tm -> - (MATCH_MP th_imp o GEN_ALL o gen_diff2c_domain_poly) poly_tm;; - - -let gen_diff_poly = - let th_imp = prove(`!f. (!domain. diff2c_domain domain f) ==> !x:real^N. lift o f differentiable at x`, - REWRITE_TAC[diff2c_domain; diff2c; diff2] THEN REPEAT STRIP_TAC THEN - FIRST_X_ASSUM (MP_TAC o SPECL [`x:real^N, x:real^N`; `x:real^N`]) THEN - REWRITE_TAC[INTERVAL_SING; IN_SING] THEN case THEN REPEAT STRIP_TAC THEN - FIRST_X_ASSUM (MP_TAC o SPEC `x:real^N`) THEN ASM_SIMP_TAC[]) in - fun poly_tm -> - (MATCH_MP th_imp o GEN_ALL o gen_diff2c_domain_poly) poly_tm;; - - -let in_tm = `IN`;; - -let add_to_hash tbl max_size key value = - let _ = if Hashtbl.length tbl >= max_size then Hashtbl.clear tbl else () in - Hashtbl.add tbl key value;; - -(* Formally computes partial derivatives of a polynomial *) -let gen_partial_poly = - let max_hash = 1000 in - let hash = Hashtbl.create max_hash in - fun i poly_tm -> - let key = (i, poly_tm) in - try Hashtbl.find hash (i, poly_tm) - with Not_found -> - let i_tm = mk_small_numeral i in - let rec gen_rec poly_tm = - let x_var, expr = dest_abs poly_tm in - let n = get_dim x_var in - if frees expr = [] then - (* const *) - (SPECL [i_tm; expr] o inst_first_type_var (n_type_array.(n))) partial_const - else - let lhs, r_tm = dest_comb expr in - if lhs = neg_op_real then - (* -- *) - let r_poly = mk_abs (x_var, r_tm) in - let r_diff = (SPEC_ALL o gen_diff_poly) r_poly and - r_partial = gen_rec r_poly in - let th0 = SPEC i_tm (MATCH_MP partial_neg r_diff) in - REWRITE_RULE[r_partial] th0 - else - let op, l_tm = dest_comb lhs in - let name = (fst o dest_const) op in - if name = "$" then - (* comp *) - let dim_th = dimindex_array.(n) in - let dim_tm = (lhand o concl) dim_th in - let i_eq_k = NUM_EQ_CONV (mk_eq (i_tm, r_tm)) in - let int_tm = mk_binop `..` `1` dim_tm in - let k_in_dim = prove(mk_comb (mk_icomb(in_tm, r_tm), int_tm), - REWRITE_TAC[IN_NUMSEG; dim_th] THEN ARITH_TAC) in - (REWRITE_RULE[i_eq_k] o MATCH_MP (SPECL [r_tm; i_tm] partial_x)) k_in_dim - else - let l_poly = mk_abs (x_var, l_tm) in - let l_partial = gen_rec l_poly in - let l_diff = (SPEC_ALL o gen_diff_poly) l_poly in - if name = "real_pow" then - (* f pow n *) - let th0 = SPECL [i_tm; r_tm] (MATCH_MP partial_pow l_diff) in - REWRITE_RULE[l_partial] th0 - else - let r_poly = mk_abs (x_var, r_tm) in - let r_partial = gen_rec r_poly in - let r_diff = (SPEC_ALL o gen_diff_poly) r_poly in - let imp_th = assoc op [add_op_real, partial_add; - sub_op_real, partial_sub; - mul_op_real, partial_mul] in - let th0 = SPEC i_tm (MATCH_MP (MATCH_MP imp_th l_diff) r_diff) in - REWRITE_RULE[l_partial; r_partial] th0 in - - let th1 = gen_rec poly_tm in - let th2 = ((NUM_REDUCE_CONV THENC REWRITE_CONV[DECIMAL] THENC REAL_POLY_CONV) o rand o concl) th1 in - let th3 = (REWRITE_RULE[ETA_AX] o ONCE_REWRITE_RULE[eq_ext] o GEN_ALL) (TRANS th1 th2) in - let _ = add_to_hash hash max_hash key th3 in - th3;; - - -let gen_partial2_poly i j poly_tm = - let partial_j = gen_partial_poly j poly_tm in - let partial_ij = gen_partial_poly i (rand (concl partial_j)) in - let pi = (rator o lhand o concl) partial_ij in - REWRITE_RULE[GSYM partial2] (TRANS (AP_TERM pi partial_j) partial_ij);; - - -(********************************************) - -let eval_diff2_poly diff2_domain_th = - fun xx zz -> - let domain_tm = mk_pair (xx, zz) in - INST[domain_tm, mk_var ("domain", type_of domain_tm)] diff2_domain_th;; - - -(*****************************) -(* m_lin_approx *) - -let CONST_INTERVAL' = RULE CONST_INTERVAL;; - -let dest_lin_approx approx_tm = - let lhs, df_bounds = dest_comb approx_tm in - let lhs2, f_bounds = dest_comb lhs in - let lhs3, x_tm = dest_comb lhs2 in - let f_tm = rand lhs3 in - f_tm, x_tm, f_bounds, df_bounds;; - -let gen_lin_approx_eq_thm n = - let ty = n_type_array.(n) in - let df_vars = map (fun i -> df_vars_array.(i)) (1--n) in - let df_bounds_list = mk_list (df_vars, real_pair_ty) in - let th0 = (SPECL[f_bounds_var; df_bounds_list] o inst_first_type_var ty) m_lin_approx in - let th1 = (CONV_RULE NUM_REDUCE_CONV o REWRITE_RULE[all_n]) th0 in - th1;; - - -let gen_lin_approx_poly_thm poly_tm diff_th partials = - let x_var, _ = dest_abs poly_tm in - let n = get_dim x_var in - let lin_eq = (REWRITE_RULE partials o SPECL [poly_tm]) (gen_lin_approx_eq_thm n) in - let x_vec = mk_vector_list (map (fun i -> x_vars_array.(i)) (1--n)) in - let th1 = (REWRITE_RULE (Array.to_list comp_thms_array.(n)) o SPEC x_vec o REWRITE_RULE[diff_th]) lin_eq in - th1;; - - -let gen_lin_approx_poly_thm0 poly_tm = - let x_var, _ = dest_abs poly_tm in - let n = get_dim x_var in - let partials = map (fun i -> gen_partial_poly i poly_tm) (1--n) in - let diff_th = gen_diff_poly poly_tm in - gen_lin_approx_poly_thm poly_tm diff_th partials;; - - -let eval_lin_approx pp0 lin_approx_th = - let poly_tm, _, _, _ = (dest_lin_approx o lhand o concl) lin_approx_th in - let x_var, _ = dest_abs poly_tm in - let n = get_dim x_var in - let th0 = lin_approx_th in - let th1 = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o MATCH_MP iffRL) th0 in - let build_eval int_hyp = - let expr, b_var = dest_binary "interval_arith" int_hyp in - (eval_constants pp0 o build_interval_fun) expr, b_var in - let int_fs = map build_eval (hyp th1) in - - let rec split_rules i_list = - match i_list with - | [] -> ([], []) - | ((i_fun, var_tm) :: es) -> - let th_list, i_list' = split_rules es in - match i_fun with - | Int_const th -> (var_tm, th) :: th_list, i_list' - | Int_var v -> (var_tm, INST[v, x_var_real] CONST_INTERVAL') :: th_list, i_list' - | _ -> th_list, (var_tm, i_fun) :: i_list' in - - let const_th_list, i_list0 = split_rules int_fs in - let th2 = itlist (fun (var_tm, th) th0 -> - let b_tm = rand (concl th) in - (MY_PROVE_HYP th o INST[b_tm, var_tm]) th0) const_th_list th1 in - let v_list, i_list' = unzip i_list0 in - let i_list = find_and_replace_all i_list' [] in - fun pp vector_tm -> - let x_vals = dest_vector vector_tm in - if length x_vals <> n then failwith (sprintf "Wrong vector size; expected size: %d" n) - else - let x_ints = map mk_const_interval x_vals in - let vars = map (fun i -> x_vars_array.(i)) (1--n) in - let th3 = INST (zip x_vals vars) th2 in - let i_vals = eval_interval_fun_list pp i_list (zip vars x_ints) in - itlist2 (fun var_tm th th0 -> - let b_tm = rand (concl th) in - (MY_PROVE_HYP th o INST[b_tm, var_tm]) th0) v_list i_vals th3;; - - -let eval_lin_approx_poly0 pp0 poly_tm = - eval_lin_approx pp0 (gen_lin_approx_poly_thm0 poly_tm);; - - - -(*************************************) - -(* 1 <= i /\ i <= n <=> i = 1 \/ i = 2 \/ ... \/ i = n *) -let i_int_array = - let i_tm = `i:num` in - let i_th0 = prove(`1 <= i /\ i <= SUC n <=> (1 <= i /\ i <= n) \/ i = SUC n`, ARITH_TAC) in - let th1 = prove(`1 <= i /\ i <= 1 <=> i = 1`, ARITH_TAC) in - let array = Array.make (max_dim + 1) th1 in - let prove_next n = - let n_tm = mk_small_numeral n in - let prev_n = num_CONV n_tm in - let tm = mk_conj (`1 <= i`, mk_binop le_op_num i_tm n_tm) in - let th = REWRITE_CONV[prev_n; i_th0; array.(n - 1)] tm in - array.(n) <- REWRITE_RULE[SYM prev_n; GSYM DISJ_ASSOC] th in - let _ = map prove_next (2--max_dim) in - array;; - - -(* (!i. 1 <= i /\ i <= n ==> P i) <=> P 1 /\ P 2 /\ ... /\ P n *) -let gen_in_interval = - let th0 = prove(`(!i:num. (i = k \/ Q i) ==> P i) <=> (P k /\ (!i. Q i ==> P i))`, MESON_TAC[]) in - let th1 = prove(`(!i:num. (i = k ==> P i)) <=> P k`, MESON_TAC[]) in - fun n -> - let n_tm = mk_small_numeral n and - i_tm = `i:num` in - let lhs1 = mk_conj (`1 <= i`, mk_binop le_op_num i_tm n_tm) in - let lhs = mk_forall (i_tm, mk_imp (lhs1, `(P:num->bool) i`)) in - REWRITE_CONV[i_int_array.(n); th0; th1] lhs;; - - -let gen_second_bounded_eq_thm n = - let ty, _, x_var, _, _, _, domain_var = get_types_and_vars n in - let dd_vars = map (fun i -> map (fun j -> dd_vars_array.(j).(i)) (1--i)) (1--n) in - let dd_bounds_list = mk_list (map (fun l -> mk_list (l, real_pair_ty)) dd_vars, real_pair_list_ty) in - let th0 = (SPECL[domain_var; dd_bounds_list] o inst_first_type_var ty) second_bounded in - let th1 = (CONV_RULE NUM_REDUCE_CONV o REWRITE_RULE[all_n]) th0 in - th1;; - - -let gen_second_bounded_poly_thm poly_tm partials2 = - let x_var, _ = dest_abs poly_tm in - let n = get_dim x_var in - let partials2' = List.flatten partials2 in - let second_th = (REWRITE_RULE partials2' o SPECL [poly_tm]) (gen_second_bounded_eq_thm n) in - second_th;; - - -let gen_second_bounded_poly_thm0 poly_tm = - let x_var, _ = dest_abs poly_tm in - let n = get_dim x_var in - let partials = map (fun i -> gen_partial_poly i poly_tm) (1--n) in - let get_partial i eq_th = - let partial_i = gen_partial_poly i (rand (concl eq_th)) in - let pi = (rator o lhand o concl) partial_i in - REWRITE_RULE[GSYM partial2] (TRANS (AP_TERM pi eq_th) partial_i) in - let partials2 = map (fun th, i -> map (fun j -> get_partial j th) (1--i)) (zip partials (1--n)) in - gen_second_bounded_poly_thm poly_tm partials2;; - -(* let eq_th = TAUT `(P ==> Q /\ R) <=> ((P ==> Q) /\ (P ==> R))` in - REWRITE_RULE[eq_th; FORALL_AND_THM; GSYM m_bounded_on_int] second_th;;*) - - -(* eval_second_bounded *) -let eval_second_bounded pp0 second_bounded_th = - let poly_tm = (lhand o rator o lhand o concl) second_bounded_th in - let th0 = second_bounded_th in - let n = (get_dim o fst o dest_abs) poly_tm in - let x_vector = mk_vector_list (map (fun i -> x_vars_array.(i)) (1--n)) and - z_vector = mk_vector_list (map (fun i -> z_vars_array.(i)) (1--n)) in - let _, _, _, _, _, _, domain_var = get_types_and_vars n in - - let th1 = INST[mk_pair (x_vector, z_vector), domain_var] th0 in - let th2 = REWRITE_RULE[IN_INTERVAL; dimindex_array.(n)] th1 in - let th3 = REWRITE_RULE[gen_in_interval n; GSYM interval_arith] th2 in - let th4 = (REWRITE_RULE[CONJ_ACI] o REWRITE_RULE (Array.to_list comp_thms_array.(n))) th3 in - let final_th0 = (UNDISCH_ALL o MATCH_MP iffRL) th4 in - - let x_var, h_tm = (dest_forall o hd o hyp) final_th0 in - let _, h2 = dest_imp h_tm in - let concl_ints = striplist dest_conj h2 in - - let i_funs = map (fun int -> - let expr, var = dest_interval_arith int in - (eval_constants pp0 o build_interval_fun) expr, var) concl_ints in - - let rec split_rules i_list = - match i_list with - | [] -> ([], []) - | ((i_fun, var_tm) :: es) -> - let th_list, i_list' = split_rules es in - match i_fun with - | Int_const th -> (var_tm, th) :: th_list, i_list' -(* | Int_var v -> (var_tm, INST[v, x_var_real] CONST_INTERVAL') :: th_list, i_list' *) - | _ -> th_list, (var_tm, i_fun) :: i_list' in - - let const_th_list, i_list0 = split_rules i_funs in - let th5 = itlist (fun (var_tm, th) th0 -> - let b_tm = rand (concl th) in - (REWRITE_RULE[th] o INST[b_tm, var_tm]) th0) const_th_list (SYM th4) in - let final_th = REWRITE_RULE[GSYM IMP_IMP] th5 in - let v_list, i_list' = unzip i_list0 in - let i_list = find_and_replace_all i_list' [] in - - fun pp x_vector_tm z_vector_tm -> - let x_vals = dest_vector x_vector_tm and - z_vals = dest_vector z_vector_tm in - if length x_vals <> n || length z_vals <> n then - failwith (sprintf "Wrong vector size; expected size: %d" n) - else - let x_vars = map (fun i -> x_vars_array.(i)) (1--n) and - z_vars = map (fun i -> z_vars_array.(i)) (1--n) in - - let inst_th = (INST (zip x_vals x_vars) o INST (zip z_vals z_vars)) final_th in - if (not o is_eq) (concl inst_th) then inst_th - else - let x_var, lhs = (dest_forall o lhand o concl) inst_th in - let hs = (butlast o striplist dest_imp) lhs in - let vars = map (rand o rator) hs in - let int_vars = zip vars (map ASSUME hs) in - - let dd_ints = eval_interval_fun_list pp i_list int_vars in - let inst_dd = map2 (fun var th -> (rand o concl) th, var) v_list dd_ints in - let inst_th2 = INST inst_dd inst_th in - - let conj_th = end_itlist CONJ dd_ints in - let lhs_th = GEN x_var (itlist DISCH hs conj_th) in - EQ_MP inst_th2 lhs_th;; - - - -let eval_second_bounded_poly0 pp0 poly_tm = - eval_second_bounded pp0 (gen_second_bounded_poly_thm0 poly_tm);; - - - -(*************************************) -(* eval_m_taylor *) - -let eval_m_taylor pp0 diff2c_th lin_th second_th = - let poly_tm = (rand o concl) diff2c_th in - let n = (get_dim o fst o dest_abs) poly_tm in - let eval_lin = eval_lin_approx pp0 lin_th and - eval_second = eval_second_bounded pp0 second_th in - - let ty, _, x_var, f_var, y_var, w_var, domain_var = get_types_and_vars n in - let th0 = (SPEC_ALL o inst_first_type_var ty) m_taylor_interval in - let th1 = INST[poly_tm, f_var] th0 in - let th2 = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o MATCH_MP iffRL o REWRITE_RULE[diff2c_th]) th1 in - - fun p_lin p_second domain_th -> - let domain_tm, y_tm, w_tm = dest_m_cell_domain (concl domain_th) in - let x_tm, z_tm = dest_pair domain_tm in - - let lin_th = eval_lin p_lin y_tm and - second_th = eval_second p_second x_tm z_tm in - - let _, _, f_bounds, df_bounds_list = dest_lin_approx (concl lin_th) in - let dd_bounds_list = (rand o concl) second_th in - let df_var = mk_var ("d_bounds_list", type_of df_bounds_list) and - dd_var = mk_var ("dd_bounds_list", type_of dd_bounds_list) in - - (MY_PROVE_HYP domain_th o MY_PROVE_HYP lin_th o MY_PROVE_HYP second_th o - INST[domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - f_bounds, f_bounds_var; df_bounds_list, df_var; dd_bounds_list, dd_var]) th2;; - - -let eval_m_taylor_poly0 pp0 poly_tm = - let diff2_th = gen_diff2c_domain_poly poly_tm in - let x_var, _ = dest_abs poly_tm in - let n = get_dim x_var in - let partials = map (fun i -> gen_partial_poly i poly_tm) (1--n) in - let get_partial i eq_th = - let partial_i = gen_partial_poly i (rand (concl eq_th)) in - let pi = (rator o lhand o concl) partial_i in - REWRITE_RULE[GSYM partial2] (TRANS (AP_TERM pi eq_th) partial_i) in - let partials2 = map2 (fun th i -> map (fun j -> get_partial j th) (1--i)) partials (1--n) in - let second_th = gen_second_bounded_poly_thm poly_tm partials2 in - let diff_th = gen_diff_poly poly_tm in - let lin_th = gen_lin_approx_poly_thm poly_tm diff_th partials in - eval_m_taylor pp0 diff2_th lin_th second_th;; - - - -(******************************************) -(* mk_eval_function *) - -let mk_eval_function_eq pp0 eq_th = - let expr_tm = (rand o concl) eq_th in - let tm0 = `!x:real^N. x IN interval [domain] ==> interval_arith (f x) f_bounds` in - let n = (get_dim o fst o dest_abs) expr_tm in - let x_vector = mk_vector_list (map (fun i -> x_vars_array.(i)) (1--n)) and - z_vector = mk_vector_list (map (fun i -> z_vars_array.(i)) (1--n)) in - let ty, _, _, _, _, _, domain_var = get_types_and_vars n and - f_var = mk_var ("f", type_of expr_tm) in - let th1 = (REWRITE_CONV[IN_INTERVAL] o subst[mk_pair(x_vector,z_vector), domain_var] o inst[ty, nty]) tm0 in - let th2 = REWRITE_RULE [dimindex_array.(n)] th1 in - let th3 = REWRITE_RULE [gen_in_interval n; GSYM interval_arith] th2 in - let th4 = (REWRITE_RULE[GSYM IMP_IMP; CONJ_ACI] o REWRITE_RULE (Array.to_list comp_thms_array.(n))) th3 in - let final_th0 = (CONV_RULE ((RAND_CONV o ONCE_DEPTH_CONV) BETA_CONV) o INST[expr_tm, f_var]) th4 in - let x_var, h_tm = (dest_forall o rand o concl) final_th0 in - let f_tm = (fst o dest_interval_arith o last o striplist dest_imp) h_tm in - let i_fun = (eval_constants pp0 o build_interval_fun) f_tm in - let i_list = find_and_replace_all [i_fun] [] in - let final_th = (PURE_REWRITE_RULE[SYM eq_th] o SYM) final_th0 in - fun pp x_vector_tm z_vector_tm -> - let x_vals = dest_vector x_vector_tm and - z_vals = dest_vector z_vector_tm in - if length x_vals <> n || length z_vals <> n then - failwith (sprintf "Wrong vector size; expected size: %d" n) - else - let x_vars = map (fun i -> x_vars_array.(i)) (1--n) and - z_vars = map (fun i -> z_vars_array.(i)) (1--n) in - - let inst_th = (INST (zip x_vals x_vars) o INST (zip z_vals z_vars)) final_th in - let x_var, lhs = (dest_forall o lhand o concl) inst_th in - let hs = (butlast o striplist dest_imp) lhs in - let vars = map (rand o rator) hs in - let int_vars = zip vars (map ASSUME hs) in - let eval_th = hd (eval_interval_fun_list pp i_list int_vars) in - let f_bounds = (rand o concl) eval_th in - let inst_th2 = INST[f_bounds, f_bounds_var] inst_th in - let lhs_th = GEN x_var (itlist DISCH hs eval_th) in - EQ_MP inst_th2 lhs_th;; - - -let mk_eval_function pp0 expr_tm = mk_eval_function_eq pp0 (REFL expr_tm);; - - - -(********************************) -(* m_taylor_error *) - -(* Sum of the list elements *) -let ITLIST2_EQ_SUM = prove(`!(f:A->B->real) l1 l2. LENGTH l1 <= LENGTH l2 ==> - ITLIST2 (\x y z. f x y + z) l1 l2 (&0) = - sum (1..(LENGTH l1)) (\i. f (EL (i - 1) l1) (EL (i - 1) l2))`, - GEN_TAC THEN - LIST_INDUCT_TAC THEN LIST_INDUCT_TAC THEN REWRITE_TAC[LENGTH; ITLIST2_DEF] THEN TRY ARITH_TAC THENL - [ - REWRITE_TAC[SUM_CLAUSES_NUMSEG; ARITH]; - REWRITE_TAC[SUM_CLAUSES_NUMSEG; ARITH]; - ALL_TAC - ] THEN - REWRITE_TAC[leqSS] THEN DISCH_TAC THEN - FIRST_X_ASSUM (MP_TAC o SPEC `t':(B)list`) THEN ASM_REWRITE_TAC[] THEN - DISCH_THEN (fun th -> REWRITE_TAC[TL; th]) THEN - REWRITE_TAC[GSYM add1n] THEN - new_rewrite [] [] SUM_ADD_SPLIT THEN REWRITE_TAC[ARITH] THEN - REWRITE_TAC[TWO; add1n; SUM_SING_NUMSEG; subnn; EL; HD] THEN - REWRITE_TAC[GSYM addn1; SUM_OFFSET; REAL_EQ_ADD_LCANCEL] THEN - MATCH_MP_TAC SUM_EQ THEN move ["i"] THEN REWRITE_TAC[IN_NUMSEG] THEN DISCH_TAC THEN - ASM_SIMP_TAC[ARITH_RULE `1 <= i ==> (i + 1) - 1 = SUC (i - 1)`; EL; TL]);; - - -let interval_arith_abs_le = prove(`!x int y. interval_arith x int ==> iabs int <= y ==> abs x <= y`, - GEN_TAC THEN case THEN REWRITE_TAC[interval_arith; IABS'] THEN REAL_ARITH_TAC);; - - -let ALL_N_ALL2 = prove(`!P (l:(A)list) i0. - (all_n i0 l P <=> if l = [] then T else ALL2 P (l_seq i0 ((i0 + LENGTH l) - 1)) l)`, - GEN_TAC THEN LIST_INDUCT_TAC THEN GEN_TAC THEN REWRITE_TAC[all_n; NOT_CONS_NIL] THEN - new_rewrite [] [] L_SEQ_CONS THEN REWRITE_TAC[LENGTH; ALL2] THEN TRY ARITH_TAC THEN - FIRST_X_ASSUM (new_rewrite [] []) THEN TRY ARITH_TAC THEN - REWRITE_TAC[addSn; addnS; addn1] THEN - SPEC_TAC (`t:(A)list`, `t:(A)list`) THEN case THEN SIMP_TAC[NOT_CONS_NIL] THEN - REWRITE_TAC[LENGTH; addn0] THEN - MP_TAC (SPECL [`SUC i0`; `SUC i0 - 1`] L_SEQ_NIL) THEN - REWRITE_TAC[ARITH_RULE `SUC i0 - 1 < SUC i0`] THEN DISCH_THEN (fun th -> REWRITE_TAC[th; ALL2]));; - - -let ALL_N_EL = prove(`!P (l:(A)list) i0. all_n i0 l P <=> (!i. i < LENGTH l ==> P (i0 + i) (EL i l))`, - REPEAT GEN_TAC THEN REWRITE_TAC[ALL_N_ALL2] THEN - SPEC_TAC (`l:(A)list`, `l:(A)list`) THEN case THEN SIMP_TAC[NOT_CONS_NIL; LENGTH; ltn0] THEN - REPEAT GEN_TAC THEN - new_rewrite [] [] ALL2_ALL_ZIP THENL - [ - REWRITE_TAC[LENGTH_L_SEQ; LENGTH] THEN ARITH_TAC; - ALL_TAC - ] THEN - REWRITE_TAC[GSYM ALL_EL] THEN - new_rewrite [] [] LENGTH_ZIP THENL - [ - REWRITE_TAC[LENGTH_L_SEQ; LENGTH] THEN ARITH_TAC; - ALL_TAC - ] THEN - REWRITE_TAC[LENGTH_L_SEQ; ARITH_RULE `((i0 + SUC a) - 1 + 1) - i0 = SUC a`] THEN - EQ_TAC THEN REPEAT STRIP_TAC THENL - [ - FIRST_X_ASSUM (MP_TAC o SPEC `i:num`) THEN ASM_REWRITE_TAC[] THEN - new_rewrite [] [] EL_ZIP THENL - [ - REWRITE_TAC[LENGTH_L_SEQ; LENGTH] THEN ASM_ARITH_TAC; - ALL_TAC - ] THEN - REWRITE_TAC[] THEN new_rewrite [] [] EL_L_SEQ THEN ASM_ARITH_TAC; - ALL_TAC - ] THEN - new_rewrite [] [] EL_ZIP THENL - [ - REWRITE_TAC[LENGTH_L_SEQ; LENGTH] THEN ASM_ARITH_TAC; - ALL_TAC - ] THEN - REWRITE_TAC[] THEN new_rewrite [] [] EL_L_SEQ THEN TRY ASM_ARITH_TAC THEN - FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_REWRITE_TAC[]);; - - -let M_TAYLOR_ERROR_ITLIST2 = prove(`!f domain y w dd_bounds_list error. - m_cell_domain domain y (vector w) ==> - diff2c_domain domain f ==> - second_bounded (f:real^N->real) domain dd_bounds_list ==> - LENGTH w = dimindex (:N) ==> - LENGTH dd_bounds_list = dimindex (:N) ==> - all_n 1 dd_bounds_list (\i list. LENGTH list = i) ==> - ITLIST2 (\list x z. x * (x * iabs (LAST list) - + &2 * ITLIST2 (\a b c. b * iabs a + c) (BUTLAST list) w (&0)) + z) - dd_bounds_list w (&0) <= error ==> - m_taylor_error f domain (vector w) error`, - REPEAT GEN_TAC THEN REWRITE_TAC[second_bounded] THEN - set_tac "s" `ITLIST2 _1 _2 _3 _4` THEN - move ["domain"; "d2f"; "second"; "lw"; "ldd"; "ldd_all"; "s_le"] THEN - ASM_SIMP_TAC[m_taylor_error_eq] THEN move ["x"; "x_in"] THEN - SUBGOAL_THEN `!i. i IN 1..dimindex (:N) ==> &0 <= EL (i - 1) w` (LABEL_TAC "w_ge0") THENL - [ - GEN_TAC THEN DISCH_TAC THEN REMOVE_THEN "domain" MP_TAC THEN new_rewrite [] [] pair_eq THEN - REWRITE_TAC[m_cell_domain] THEN - DISCH_THEN (MP_TAC o SPEC `i:num`) THEN ASM_REWRITE_TAC[] THEN - ASM_SIMP_TAC[VECTOR_COMPONENT] THEN REAL_ARITH_TAC; - ALL_TAC - ] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `s:real` THEN ASM_REWRITE_TAC[] THEN - EXPAND_TAC "s" THEN - new_rewrite [] [] ITLIST2_EQ_SUM THEN ASM_REWRITE_TAC[LE_REFL] THEN - MATCH_MP_TAC SUM_LE THEN REWRITE_TAC[FINITE_NUMSEG] THEN - move ["i"; "i_in"] THEN ASM_SIMP_TAC[VECTOR_COMPONENT] THEN - USE_THEN "i_in" (ASSUME_TAC o REWRITE_RULE[IN_NUMSEG]) THEN - MATCH_MP_TAC REAL_LE_LMUL THEN ASM_SIMP_TAC[] THEN - SUBGOAL_THEN `LENGTH (EL (i - 1) dd_bounds_list:(real#real)list) = i` (LABEL_TAC "len_i") THENL - [ - REMOVE_THEN "ldd_all" MP_TAC THEN - REWRITE_TAC[ALL_N_EL] THEN DISCH_THEN (MP_TAC o SPEC `i - 1`) THEN - ANTS_TAC THENL - [ - ASM_REWRITE_TAC[] THEN POP_ASSUM MP_TAC THEN ARITH_TAC; - ALL_TAC - ] THEN - DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN POP_ASSUM MP_TAC THEN ARITH_TAC; - ALL_TAC - ] THEN - new_rewrite [] [] ITLIST2_EQ_SUM THEN ASM_REWRITE_TAC[] THENL - [ - REWRITE_TAC[LENGTH_BUTLAST] THEN POP_ASSUM MP_TAC THEN POP_ASSUM MP_TAC THEN ARITH_TAC; - ALL_TAC - ] THEN - MATCH_MP_TAC REAL_LE_ADD2 THEN CONJ_TAC THENL - [ - new_rewrite [] [] LAST_EL THENL - [ - ASM_REWRITE_TAC[GSYM LENGTH_EQ_NIL] THEN - REMOVE_THEN "i_in" MP_TAC THEN REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; - ALL_TAC - ] THEN - MATCH_MP_TAC REAL_LE_LMUL THEN ASM_SIMP_TAC[] THEN - FIRST_X_ASSUM (MP_TAC o SPEC `x:real^N`) THEN ASM_REWRITE_TAC[ALL_N_EL] THEN - DISCH_THEN (MP_TAC o SPEC `i - 1`) THEN - ANTS_TAC THENL - [ - REMOVE_THEN "i_in" MP_TAC THEN REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; - ALL_TAC - ] THEN - DISCH_THEN (MP_TAC o SPEC `i - 1`) THEN ANTS_TAC THENL - [ - UNDISCH_TAC `1 <= i /\ i <= dimindex (:N)` THEN ASM_REWRITE_TAC[] THEN ARITH_TAC; - ALL_TAC - ] THEN - SUBGOAL_THEN `1 + i - 1 = i /\ 1 + i - 1 = i` (fun th -> REWRITE_TAC[th]) THENL - [ - REPLICATE_TAC 3 (POP_ASSUM MP_TAC) THEN ARITH_TAC; - ALL_TAC - ] THEN - REWRITE_TAC[partial2] THEN - DISCH_THEN (MP_TAC o MATCH_MP interval_arith_abs_le) THEN - DISCH_THEN MATCH_MP_TAC THEN REWRITE_TAC[REAL_LE_REFL]; - ALL_TAC - ] THEN - MATCH_MP_TAC REAL_LE_LMUL THEN CONJ_TAC THENL [ REAL_ARITH_TAC; ALL_TAC ] THEN - ASM_REWRITE_TAC[LENGTH_BUTLAST] THEN - MATCH_MP_TAC SUM_LE THEN REWRITE_TAC[FINITE_NUMSEG] THEN - move ["j"; "j_in"] THEN - SUBGOAL_THEN `j IN 1..dimindex (:N)` (LABEL_TAC "j_in2") THENL - [ - REMOVE_THEN "i_in" MP_TAC THEN REMOVE_THEN "j_in" MP_TAC THEN - REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; - ALL_TAC - ] THEN - ASM_SIMP_TAC[VECTOR_COMPONENT] THEN - USE_THEN "j_in" (ASSUME_TAC o REWRITE_RULE[IN_NUMSEG]) THEN - MATCH_MP_TAC REAL_LE_LMUL THEN ASM_SIMP_TAC[] THEN - FIRST_X_ASSUM (MP_TAC o SPEC `x:real^N`) THEN ASM_REWRITE_TAC[ALL_N_EL] THEN - DISCH_THEN (MP_TAC o SPEC `i - 1`) THEN ANTS_TAC THENL - [ - REMOVE_THEN "i_in" MP_TAC THEN REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; - ALL_TAC - ] THEN - DISCH_THEN (MP_TAC o SPEC `j - 1`) THEN ANTS_TAC THENL - [ - ASM_REWRITE_TAC[] THEN POP_ASSUM MP_TAC THEN ARITH_TAC; - ALL_TAC - ] THEN - SUBGOAL_THEN `1 + j - 1 = j /\ 1 + i - 1 = i` (fun th -> REWRITE_TAC[th]) THENL - [ - REPLICATE_TAC 3 (POP_ASSUM MP_TAC) THEN ARITH_TAC; - ALL_TAC - ] THEN - REWRITE_TAC[partial2] THEN - DISCH_THEN (MP_TAC o MATCH_MP interval_arith_abs_le) THEN - DISCH_THEN MATCH_MP_TAC THEN - new_rewrite [] [] EL_BUTLAST THENL - [ - ASM_REWRITE_TAC[] THEN REMOVE_THEN "j_in" MP_TAC THEN REMOVE_THEN "i_in" MP_TAC THEN - REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; - ALL_TAC - ] THEN - REWRITE_TAC[REAL_LE_REFL]);; - - -let M_TAYLOR_ERROR_ITLIST2_ALT = prove(`!f domain y w f_lo f_hi d_bounds_list dd_bounds_list error. - m_taylor_interval f domain y (vector w:real^N) (f_lo, f_hi) d_bounds_list dd_bounds_list ==> - LENGTH w = dimindex (:N) ==> - LENGTH dd_bounds_list = dimindex (:N) ==> - all_n 1 dd_bounds_list (\i list. LENGTH list = i) ==> - ITLIST2 (\list x z. x * (x * iabs (LAST list) - + &2 * ITLIST2 (\a b c. b * iabs a + c) (BUTLAST list) w (&0)) + z) - dd_bounds_list w (&0) <= error ==> - m_taylor_error f domain (vector w) error`, - REWRITE_TAC[m_taylor_interval] THEN REPEAT STRIP_TAC THEN - MP_TAC (SPEC_ALL M_TAYLOR_ERROR_ITLIST2) THEN ASM_REWRITE_TAC[]);; - - - -(****************************) - -let M_TAYLOR_INTERVAL' = MY_RULE m_taylor_interval;; - -let dest_m_taylor m_taylor_tm = - let ltm1, dd_bounds_list = dest_comb m_taylor_tm in - let ltm2, d_bounds_list = dest_comb ltm1 in - let ltm3, f_bounds = dest_comb ltm2 in - let ltm4, w = dest_comb ltm3 in - let ltm5, y = dest_comb ltm4 in - let ltm6, domain = dest_comb ltm5 in - rand ltm6, domain, y, w, f_bounds, d_bounds_list, dd_bounds_list;; - -let dest_m_taylor_thms n = - let ty, xty, x_var, f_var, y_var, w_var, domain_var = get_types_and_vars n in - fun m_taylor_th -> - let f, domain, y, w, f_bounds, d_bounds_list, dd_bounds_list = dest_m_taylor (concl m_taylor_th) in - let th0 = (INST[f, f_var; domain, domain_var; y, y_var; w, w_var; f_bounds, f_bounds_var; - d_bounds_list, d_bounds_list_var; dd_bounds_list, dd_bounds_list_var] o - inst_first_type_var ty) M_TAYLOR_INTERVAL' in - let th1 = EQ_MP th0 m_taylor_th in - match (CONJUNCTS th1) with - | [domain_th; d2_th; lin_th; second_th] -> - domain_th, d2_th, lin_th, second_th - | _ -> failwith "dest_m_taylor_thms";; - - -(**********************) - -(* bound *) -let M_TAYLOR_BOUND' = - prove(`m_taylor_interval f domain y (vector w:real^N) (f_lo, f_hi) d_bounds_list dd_bounds_list /\ - m_taylor_error f domain (vector w) error /\ - ITLIST2 (\a b c. b * iabs a + c) d_bounds_list w (&0) <= b /\ - b + inv(&2) * error <= a /\ - lo <= f_lo - a /\ - f_hi + a <= hi /\ - LENGTH w = dimindex (:N) /\ LENGTH d_bounds_list = dimindex (:N) ==> - (!x. x IN interval [domain] ==> interval_arith (f x) (lo, hi))`, - REWRITE_TAC[GSYM m_bounded_on_int; m_taylor_interval; m_lin_approx; ALL_N_EL] THEN - set_tac "s" `ITLIST2 _1 _2 _3 _4` THEN STRIP_TAC THEN - SUBGOAL_THEN `diff2_domain domain (f:real^N->real)` ASSUME_TAC THENL - [ - UNDISCH_TAC `diff2c_domain domain (f:real^N->real)` THEN - SIMP_TAC[diff2_domain; diff2c_domain; diff2c]; - ALL_TAC - ] THEN - apply_tac m_taylor_bounds THEN - MAP_EVERY EXISTS_TAC [`y:real^N`; `vector w:real^N`; `error:real`; `f_lo:real`; `f_hi:real`; `a:real`] THEN - ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b + inv (&2) * error` THEN ASM_REWRITE_TAC[] THEN - REWRITE_TAC[real_div; REAL_MUL_AC] THEN - MATCH_MP_TAC REAL_LE_ADD2 THEN REWRITE_TAC[REAL_LE_REFL] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `s:real` THEN ASM_REWRITE_TAC[] THEN - EXPAND_TAC "s" THEN new_rewrite [] [] ITLIST2_EQ_SUM THEN ASM_REWRITE_TAC[LE_REFL] THEN - MATCH_MP_TAC SUM_LE THEN REWRITE_TAC[FINITE_NUMSEG] THEN REPEAT STRIP_TAC THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_SIMP_TAC[VECTOR_COMPONENT; REAL_LE_REFL; REAL_ABS_POS] THEN - CONJ_TAC THENL - [ - UNDISCH_TAC `m_cell_domain domain (y:real^N) (vector w)` THEN - new_rewrite [] [] pair_eq THEN REWRITE_TAC[m_cell_domain] THEN - DISCH_THEN (MP_TAC o SPEC `x:num`) THEN ASM_REWRITE_TAC[] THEN - ASM_SIMP_TAC[VECTOR_COMPONENT] THEN REAL_ARITH_TAC; - ALL_TAC - ] THEN - FIRST_X_ASSUM (MP_TAC o SPEC `x - 1`) THEN - ANTS_TAC THENL - [ - POP_ASSUM MP_TAC THEN ASM_REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; - ALL_TAC - ] THEN - DISCH_THEN (MP_TAC o MATCH_MP interval_arith_abs_le) THEN - SUBGOAL_THEN `1 + x - 1 = x` (fun th -> REWRITE_TAC[th]) THENL - [ - POP_ASSUM MP_TAC THEN REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; - ALL_TAC - ] THEN - DISCH_THEN MATCH_MP_TAC THEN REWRITE_TAC[REAL_LE_REFL]);; - - -(* upper *) -let M_TAYLOR_UPPER_BOUND' = prove(`m_taylor_interval f domain y (vector w) (f_lo, f_hi) - d_bounds_list dd_bounds_list /\ - m_taylor_error f domain (vector w:real^N) error /\ - ITLIST2 (\a b c. b * iabs a + c) d_bounds_list w (&0) <= b /\ - b + inv(&2) * error <= a /\ - f_hi + a <= hi /\ - LENGTH w = dimindex (:N) /\ LENGTH d_bounds_list = dimindex (:N) ==> - (!p. p IN interval [domain] ==> f p <= hi)`, - STRIP_TAC THEN - MP_TAC (INST[`f_lo - a:real`, `lo:real`] M_TAYLOR_BOUND') THEN - ASM_SIMP_TAC[interval_arith; REAL_LE_REFL]);; - - -(* lower *) -let M_TAYLOR_LOWER_BOUND' = prove(`m_taylor_interval f domain y (vector w:real^N) (f_lo, f_hi) - d_bounds_list dd_bounds_list /\ - m_taylor_error f domain (vector w) error /\ - ITLIST2 (\a b c. b * iabs a + c) d_bounds_list w (&0) <= b /\ - b + inv(&2) * error <= a /\ - lo <= f_lo - a /\ - LENGTH w = dimindex (:N) /\ LENGTH d_bounds_list = dimindex (:N) ==> - (!p. p IN interval [domain] ==> lo <= f p)`, - STRIP_TAC THEN - MP_TAC (INST[`f_hi + a:real`, `hi:real`] M_TAYLOR_BOUND') THEN - ASM_SIMP_TAC[interval_arith; REAL_LE_REFL]);; - - - -(* arrays *) -let gen_taylor_bound_th bound_th n = - let th0 = (DISCH_ALL o MY_RULE o REWRITE_RULE[MY_RULE M_TAYLOR_ERROR_ITLIST2_ALT]) bound_th in - let ns = 1--n in - let mk_list_hd l = mk_list (l, type_of (hd l)) in - let w_list = mk_list_hd (map (fun i -> w_vars_array.(i)) ns) in - let d_bounds_list = mk_list_hd (map (fun i -> df_vars_array.(i)) ns) in - let dd_bounds_list = mk_list_hd (map (fun i -> mk_list_hd (map (fun j -> dd_vars_array.(i).(j)) (1--i))) ns) in - let th1 = (INST[w_list, w_var_real_list; d_bounds_list, d_bounds_list_var; - dd_bounds_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) th0 in - let th2 = REWRITE_RULE[LAST; NOT_CONS_NIL; BUTLAST; all_n; ITLIST2_DEF; LENGTH; ARITH; dimindex_array.(n)] th1 in - let th3 = REWRITE_RULE[HD; TL; REAL_MUL_RZERO; REAL_ADD_RID; GSYM error_mul_f2] th2 in - (MY_RULE o REWRITE_RULE[float_inv2_th; SYM float2_eq]) th3;; - - -let m_taylor_upper_array = Array.init (max_dim + 1) - (fun i -> if i < 1 then TRUTH else gen_taylor_bound_th M_TAYLOR_UPPER_BOUND' i);; - -let m_taylor_lower_array = Array.init (max_dim + 1) - (fun i -> if i < 1 then TRUTH else gen_taylor_bound_th M_TAYLOR_LOWER_BOUND' i);; - -let m_taylor_bound_array = Array.init (max_dim + 1) - (fun i -> if i < 1 then TRUTH else gen_taylor_bound_th M_TAYLOR_BOUND' i);; - - -(***************************) -(* eval_m_taylor_bounds0 *) - -let eval_m_taylor_bounds0 mode n pp m_taylor_th = - let bound_th = - if mode = "upper" then - m_taylor_upper_array.(n) - else if mode = "bound" then - m_taylor_bound_array.(n) - else - m_taylor_lower_array.(n) in - - let f_tm, domain_tm, y_tm, w_tm, f_bounds, d_bounds_list, dd_bounds_list = dest_m_taylor (concl m_taylor_th) in - let f_lo, f_hi = dest_pair f_bounds and - ws = dest_list (rand w_tm) and - dfs = dest_list d_bounds_list and - dds = map dest_list (dest_list dd_bounds_list) in - let ns = 1--n in - let df_vars = map (fun i -> df_vars_array.(i)) ns and - dd_vars = map (fun i -> map (fun j -> dd_vars_array.(i).(j)) (1--i)) ns and - w_vars = map (fun i -> w_vars_array.(i)) ns and - y_var = mk_var ("y", type_of y_tm) and - f_var = mk_var ("f", type_of f_tm) and - domain_var = mk_var ("domain", type_of domain_tm) in - - (* sum of first partials *) - let d_th = - let mul_wd = map2 (error_mul_f2_le_conv2 pp) ws dfs in - end_itlist (add_ineq_hi pp) mul_wd in - - let b_tm = (rand o concl) d_th in - - (* sum of second partials *) - let dd_th = - let ( * ), ( + ) = mul_ineq_pos_const_hi pp, add_ineq_hi pp in - let mul_wdd = map2 (fun list i -> my_map2 (error_mul_f2_le_conv2 pp) ws list) dds ns in - let sums1 = map (end_itlist ( + ) o butlast) (tl mul_wdd) in - let sums2 = (hd o hd) mul_wdd :: map2 (fun list th1 -> last list + two_float * th1) (tl mul_wdd) sums1 in - let sums = map2 ( * ) ws sums2 in - end_itlist ( + ) sums in - - let error_tm = (rand o concl) dd_th in - - (* additional inequalities *) - let ineq1_th = - let ( * ), ( + ) = float_mul_hi pp, add_ineq_hi pp in - mk_refl_ineq b_tm + float_inv2 * error_tm in - - let a_tm = (rand o concl) ineq1_th in - - let prove_ineq2, bounds_inst = - if mode = "upper" then - let ineq2 = float_add_hi pp f_hi a_tm in - MY_PROVE_HYP ineq2, [(rand o concl) ineq2, hi_var_real] - else if mode = "bound" then - let ineq2 = float_add_hi pp f_hi a_tm in - let ineq3 = float_sub_lo pp f_lo a_tm in - MY_PROVE_HYP ineq2 o MY_PROVE_HYP ineq3, - [(rand o concl) ineq2, hi_var_real; (lhand o concl) ineq3, lo_var_real] - else - let ineq2 = float_sub_lo pp f_lo a_tm in - MY_PROVE_HYP ineq2, [(lhand o concl) ineq2, lo_var_real] in - - (* final step *) - let inst_list = - let inst1 = zip dfs df_vars in - let inst2 = zip (List.flatten dds) (List.flatten dd_vars) in - let inst3 = zip ws w_vars in - inst1 @ inst2 @ inst3 in - - (MY_PROVE_HYP m_taylor_th o MY_PROVE_HYP dd_th o MY_PROVE_HYP d_th o - MY_PROVE_HYP ineq1_th o prove_ineq2 o - INST ([f_hi, f_hi_var; f_lo, f_lo_var; error_tm, error_var; - a_tm, a_var_real; b_tm, b_var_real; - y_tm, y_var; domain_tm, domain_var; - f_tm, f_var] @ bounds_inst @ inst_list)) bound_th;; - - -(* upper *) -let eval_m_taylor_upper_bound = eval_m_taylor_bounds0 "upper";; - -(* lower *) -let eval_m_taylor_lower_bound = eval_m_taylor_bounds0 "lower";; - -(* bound *) -let eval_m_taylor_bound = eval_m_taylor_bounds0 "bound";; - - - - -(******************************) -(* taylor_upper_partial_bound *) -(* taylor_lower_partial_bound *) - - -(* bound *) -let M_TAYLOR_PARTIAL_BOUND' = - prove(`m_taylor_interval f domain (y:real^N) (vector w) f_bounds d_bounds_list dd_bounds_list /\ - i IN 1..dimindex (:N) /\ - EL (i - 1) d_bounds_list = (df_lo, df_hi) /\ - (!x. x IN interval [domain] ==> all_n 1 dd_list - (\j int. interval_arith (if j <= i then partial2 j i f x else partial2 i j f x) int)) /\ - LENGTH dd_list = dimindex (:N) /\ - LENGTH d_bounds_list = dimindex (:N) /\ - LENGTH w = dimindex (:N) /\ - ITLIST2 (\a b c. b * iabs a + c) dd_list w (&0) <= error /\ - df_hi + error <= hi ==> - lo <= df_lo - error ==> - (!x. x IN interval [domain] ==> interval_arith (partial i f x) (lo, hi))`, - REWRITE_TAC[m_taylor_interval; m_lin_approx; ALL_N_EL; GSYM m_bounded_on_int] THEN - set_tac "s" `ITLIST2 _1 _2 _3 _4` THEN REPEAT STRIP_TAC THEN - SUBGOAL_THEN `1 <= i /\ i <= dimindex (:N)` (LABEL_TAC "i_in") THENL - [ - ASM_REWRITE_TAC[GSYM IN_NUMSEG]; - ALL_TAC - ] THEN - SUBGOAL_THEN `diff2_domain domain (f:real^N->real)` ASSUME_TAC THENL - [ - UNDISCH_TAC `diff2c_domain domain (f:real^N->real)` THEN - SIMP_TAC[diff2_domain; diff2c_domain; diff2c]; - ALL_TAC - ] THEN - REWRITE_TAC[ETA_AX] THEN apply_tac m_taylor_partial_bounds THEN - MAP_EVERY EXISTS_TAC [`y:real^N`; `vector w:real^N`; `error:real`; `df_lo:real`; `df_hi:real`] THEN - ASM_REWRITE_TAC[] THEN - FIRST_X_ASSUM (MP_TAC o SPEC `i - 1`) THEN - ASM_SIMP_TAC[ARITH_RULE `1 <= i /\ i <= n ==> i - 1 < n`] THEN - ASM_SIMP_TAC[ARITH_RULE `1 <= i ==> 1 + i - 1 = i`; interval_arith] THEN - DISCH_THEN (fun th -> ALL_TAC) THEN - REWRITE_TAC[m_taylor_partial_error] THEN REPEAT STRIP_TAC THEN - MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `s:real` THEN ASM_REWRITE_TAC[] THEN - EXPAND_TAC "s" THEN new_rewrite [] [] ITLIST2_EQ_SUM THEN ASM_REWRITE_TAC[LE_REFL] THEN - MATCH_MP_TAC SUM_LE THEN REWRITE_TAC[FINITE_NUMSEG] THEN - move ["j"; "j_in"] THEN - ASM_SIMP_TAC[VECTOR_COMPONENT] THEN - MATCH_MP_TAC REAL_LE_LMUL THEN CONJ_TAC THENL - [ - UNDISCH_TAC `m_cell_domain domain (y:real^N) (vector w)` THEN - new_rewrite [] [] pair_eq THEN REWRITE_TAC[m_cell_domain] THEN - DISCH_THEN (MP_TAC o SPEC `j:num`) THEN ASM_REWRITE_TAC[] THEN - ASM_SIMP_TAC[VECTOR_COMPONENT] THEN REAL_ARITH_TAC; - ALL_TAC - ] THEN - FIRST_X_ASSUM (MP_TAC o SPEC `x:real^N`) THEN ASM_REWRITE_TAC[] THEN - DISCH_THEN (MP_TAC o SPEC `j - 1`) THEN - POP_ASSUM MP_TAC THEN REWRITE_TAC[IN_NUMSEG] THEN DISCH_TAC THEN - ASM_SIMP_TAC[ARITH_RULE `1 <= j /\ j <= n ==> j - 1 < n`] THEN - ASM_SIMP_TAC[ARITH_RULE `!i. 1 <= i ==> 1 + i - 1 = i`; GSYM partial2] THEN - DISCH_THEN (MP_TAC o MATCH_MP interval_arith_abs_le) THEN - COND_CASES_TAC THEN TRY (DISCH_THEN MATCH_MP_TAC THEN REWRITE_TAC[REAL_LE_REFL]) THEN - new_rewrite [] [] mixed_second_partials THENL - [ - UNDISCH_TAC `diff2c_domain domain (f:real^N->real)` THEN - ASM_SIMP_TAC[diff2c_domain]; - ALL_TAC - ] THEN - DISCH_THEN MATCH_MP_TAC THEN REWRITE_TAC[REAL_LE_REFL]);; - - -(* upper *) -let M_TAYLOR_PARTIAL_UPPER' = - prove(`m_taylor_interval f domain (y:real^N) (vector w) f_bounds d_bounds_list dd_bounds_list /\ - i IN 1..dimindex (:N) /\ - EL (i - 1) d_bounds_list = (df_lo, df_hi) /\ - (!x. x IN interval [domain] ==> all_n 1 dd_list - (\j int. interval_arith (if j <= i then partial2 j i f x else partial2 i j f x) int)) /\ - LENGTH dd_list = dimindex (:N) /\ - LENGTH d_bounds_list = dimindex (:N) /\ - LENGTH w = dimindex (:N) /\ - ITLIST2 (\a b c. b * iabs a + c) dd_list w (&0) <= error /\ - df_hi + error <= hi ==> - (!x. x IN interval [domain] ==> partial i f x <= hi)`, - REPEAT STRIP_TAC THEN - MP_TAC (INST[`df_lo - error`, `lo:real`] M_TAYLOR_PARTIAL_BOUND') THEN ASM_REWRITE_TAC[REAL_LE_REFL] THEN - ASM_SIMP_TAC[interval_arith]);; - - -(* lower *) -let M_TAYLOR_PARTIAL_LOWER' = - prove(`m_taylor_interval f domain (y:real^N) (vector w) f_bounds d_bounds_list dd_bounds_list /\ - i IN 1..dimindex (:N) /\ - EL (i - 1) d_bounds_list = (df_lo, df_hi) /\ - (!x. x IN interval [domain] ==> all_n 1 dd_list - (\j int. interval_arith (if j <= i then partial2 j i f x else partial2 i j f x) int)) /\ - LENGTH dd_list = dimindex (:N) /\ - LENGTH d_bounds_list = dimindex (:N) /\ - LENGTH w = dimindex (:N) /\ - ITLIST2 (\a b c. b * iabs a + c) dd_list w (&0) <= error /\ - lo <= df_lo - error ==> - (!x. x IN interval [domain] ==> lo <= partial i f x)`, - REPEAT STRIP_TAC THEN - MP_TAC (INST[`df_hi + error`, `hi:real`] M_TAYLOR_PARTIAL_BOUND') THEN ASM_REWRITE_TAC[REAL_LE_REFL] THEN - ASM_SIMP_TAC[interval_arith]);; - - - - -(* arrays *) -let gen_taylor_partial_bound_th = - let imp_and_eq = TAUT `((P ==> Q) /\ (P ==> R)) <=> (P ==> Q /\ R)` in - let mk_list_hd l = mk_list (l, type_of (hd l)) in - let dd_list_var = `dd_list : (real#real)list` in - fun bound_th n i -> - let ns = 1--n in - let i_tm = mk_small_numeral i in - let w_list = mk_list_hd (map (fun i -> w_vars_array.(i)) ns) in - let d_bounds_list = mk_list_hd (map (fun i -> df_vars_array.(i)) ns) in - let dd_bounds_list = mk_list_hd - (map (fun i -> mk_list_hd (map (fun j -> dd_vars_array.(i).(j)) (1--i))) ns) in - let dd_list = mk_list_hd - (map (fun j -> if j <= i then dd_vars_array.(i).(j) else dd_vars_array.(j).(i)) ns) in - - let th1 = (INST[w_list, w_var_real_list; d_bounds_list, d_bounds_list_var; dd_list, dd_list_var; i_tm, `i:num`; - dd_bounds_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) bound_th in - let th2 = REWRITE_RULE[REAL_ADD_RID; HD; TL; ITLIST2_DEF; LENGTH; ARITH; dimindex_array.(n)] th1 in - let th3 = REWRITE_RULE[IN_NUMSEG; ARITH; el_thms_array.(n).(i - 1)] th2 in - let th4 = (REWRITE_RULE[] o INST[`df_lo:real, df_hi:real`, df_vars_array.(i)]) th3 in - let th5 = (MY_RULE o REWRITE_RULE[GSYM imp_and_eq; GSYM AND_FORALL_THM; all_n; ARITH]) th4 in - - let m_taylor_hyp = find (can dest_m_taylor) (hyp th5) in - let t_th0 = (REWRITE_RULE[ARITH; all_n; second_bounded; m_taylor_interval] o ASSUME) m_taylor_hyp in - let t_th1 = REWRITE_RULE[GSYM imp_and_eq; GSYM AND_FORALL_THM] t_th0 in - (MY_RULE_NUM o REWRITE_RULE[GSYM error_mul_f2; t_th1] o DISCH_ALL) th5;; - - -(* The (n, i)-th element is the theorem |- i IN 1..dimindex (:n) *) -let i_in_array = Array.init (max_dim + 1) - (fun i -> Array.init (i + 1) - (fun j -> - if j < 1 then TRUTH else - let j_tm = mk_small_numeral j in - let tm0 = `j IN 1..dimindex (:N)` in - let tm1 = (subst [j_tm, `j:num`] o inst [n_type_array.(i), nty]) tm0 in - prove(tm1, REWRITE_TAC[dimindex_array.(i); IN_NUMSEG] THEN ARITH_TAC)));; - -let m_taylor_partial_upper_array, m_taylor_partial_lower_array, m_taylor_partial_bound_array = - let gen_array bound_th = Array.init (max_dim + 1) - (fun i -> Array.init (i + 1) - (fun j -> if j < 1 then TRUTH else gen_taylor_partial_bound_th bound_th i j)) in - gen_array M_TAYLOR_PARTIAL_UPPER', gen_array M_TAYLOR_PARTIAL_LOWER', gen_array M_TAYLOR_PARTIAL_BOUND';; - - -(***************************) - - -let eval_m_taylor_partial_bounds0 mode n pp i m_taylor_th = - let bound_th = - if mode = "upper" then - m_taylor_partial_upper_array.(n).(i) - else if mode = "bound" then - m_taylor_partial_bound_array.(n).(i) - else - m_taylor_partial_lower_array.(n).(i) in - - let f_tm, domain_tm, y_tm, w_tm, f_bounds, d_bounds_list, dd_bounds_list = dest_m_taylor (concl m_taylor_th) in - let ws = dest_list (rand w_tm) and - dfs = dest_list d_bounds_list and - dds = map dest_list (dest_list dd_bounds_list) in - let ns = 1--n in - let df_lo, df_hi = dest_pair (List.nth dfs (i - 1)) and - df_vars = map (fun i -> df_vars_array.(i)) ns and - dd_vars = map (fun i -> map (fun j -> dd_vars_array.(i).(j)) (1--i)) ns and - w_vars = map (fun i -> w_vars_array.(i)) ns and - y_var = mk_var ("y", type_of y_tm) and - f_var = mk_var ("f", type_of f_tm) and - domain_var = mk_var ("domain", type_of domain_tm) in - - (* sum of second partials *) - let dd_list = map - (fun j -> if j <= i then List.nth (List.nth dds (i-1)) (j-1) else List.nth (List.nth dds (j-1)) (i-1)) ns in - - let dd_th = - let mul_dd = map2 (error_mul_f2_le_conv2 pp) ws dd_list in - end_itlist (add_ineq_hi pp) mul_dd in - - let error_tm = (rand o concl) dd_th in - - (* additional inequalities *) - let prove_ineq, bounds_inst = - if mode = "upper" then - let ineq2 = float_add_hi pp df_hi error_tm in - MY_PROVE_HYP ineq2, [(rand o concl) ineq2, hi_var_real] - else if mode = "bound" then - let ineq2 = float_add_hi pp df_hi error_tm in - let ineq3 = float_sub_lo pp df_lo error_tm in - MY_PROVE_HYP ineq2 o MY_PROVE_HYP ineq3, - [(rand o concl) ineq2, hi_var_real; (lhand o concl) ineq3, lo_var_real] - else - let ineq2 = float_sub_lo pp df_lo error_tm in - MY_PROVE_HYP ineq2, [(lhand o concl) ineq2, lo_var_real] in - - (* final step *) - let inst_list = - let inst1 = zip dfs df_vars in - let inst2 = zip (List.flatten dds) (List.flatten dd_vars) in - let inst3 = zip ws w_vars in - inst1 @ inst2 @ inst3 in - - (MY_PROVE_HYP m_taylor_th o MY_PROVE_HYP dd_th o prove_ineq o - INST ([df_hi, df_hi_var; df_lo, df_lo_var; error_tm, error_var; - y_tm, y_var; domain_tm, domain_var; f_bounds, f_bounds_var; - f_tm, f_var] @ bounds_inst @ inst_list)) bound_th;; - - -(* upper *) -let eval_m_taylor_partial_upper = eval_m_taylor_partial_bounds0 "upper";; - -(* lower *) -let eval_m_taylor_partial_lower = eval_m_taylor_partial_bounds0 "lower";; - -(* bound *) -let eval_m_taylor_partial_bound = eval_m_taylor_partial_bounds0 "bound";; - - - -end;; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal taylor intervals *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/arith/more_float.hl";; +needs "Formal_ineqs/arith/eval_interval.hl";; +needs "Formal_ineqs/list/list_conversions.hl";; +needs "Formal_ineqs/list/list_float.hl";; +needs "Formal_ineqs/list/more_list.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; + +needs "Formal_ineqs/taylor/theory/taylor_interval-compiled.hl";; +needs "Formal_ineqs/taylor/theory/multivariate_taylor-compiled.hl";; + + +module M_taylor = struct + +open Ssreflect;; +open Ssrfun;; +open Ssrbool;; +open Ssrnat;; +open Taylor_interval;; +open Multivariate_taylor;; +open Misc_functions;; +open Arith_float;; +open More_float;; +open Float_theory;; +open Eval_interval;; +open List_conversions;; +open List_float;; +open More_list;; +open Interval_arith;; +open Misc_vars;; + +let MY_RULE = UNDISCH_ALL o PURE_REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL;; +let MY_RULE_NUM = UNDISCH_ALL o Arith_nat.NUMERALS_TO_NUM o PURE_REWRITE_RULE[GSYM IMP_IMP] o SPEC_ALL;; +let MY_RULE_FLOAT = UNDISCH_ALL o Arith_nat.NUMERALS_TO_NUM o + PURE_REWRITE_RULE[FLOAT_OF_NUM; min_exp_def; GSYM IMP_IMP] o SPEC_ALL;; + + + +let max_dim = 8;; + +let inst_first_type_var ty th = + let ty_vars = type_vars_in_term (concl th) in + if ty_vars = [] then + failwith "inst_first_type: no type variables in the theorem" + else + INST_TYPE [ty, hd ty_vars] th;; + + +let float0 = mk_float 0 0 and + interval0 = mk_float_interval_small_num 0;; + + +let has_size_array = Array.init (max_dim + 1) + (fun i -> match i with + | 0 -> TRUTH + | 1 -> HAS_SIZE_1 + | _ -> HAS_SIZE_DIMINDEX_RULE(mk_finty(Num.num_of_int i)));; + +let dimindex_array = Array.init (max_dim + 1) + (fun i -> if i < 1 then TRUTH else MATCH_MP DIMINDEX_UNIQUE has_size_array.(i));; + +let n_type_array = Array.init (max_dim + 1) + (fun i -> if i < 1 then bool_ty else + let dimindex_th = dimindex_array.(i) in + (hd o snd o dest_type o snd o dest_const o rand o lhand o concl) dimindex_th);; + +let n_vector_type_array = Array.init (max_dim + 1) + (fun i -> if i < 1 then bool_ty else mk_type ("cart", [real_ty; n_type_array.(i)]));; + + +let x_var_names = Array.init (max_dim + 1) (fun i -> "x"^(string_of_int i)) and + y_var_names = Array.init (max_dim + 1) (fun i -> "y"^(string_of_int i)) and + z_var_names = Array.init (max_dim + 1) (fun i -> "z"^(string_of_int i)) and + w_var_names = Array.init (max_dim + 1) (fun i -> "w"^(string_of_int i));; + +let x_vars_array = Array.init (max_dim + 1) (fun i -> mk_var(x_var_names.(i), real_ty)) and + y_vars_array = Array.init (max_dim + 1) (fun i -> mk_var(y_var_names.(i), real_ty)) and + z_vars_array = Array.init (max_dim + 1) (fun i -> mk_var(z_var_names.(i), real_ty)) and + w_vars_array = Array.init (max_dim + 1) (fun i -> mk_var(w_var_names.(i), real_ty));; + +let df_vars_array = Array.init (max_dim + 1) (fun i -> mk_var ("df"^(string_of_int i), real_pair_ty));; +let dd_vars_array = Array.init (max_dim + 1) (fun i -> + Array.init (max_dim + 1) (fun j -> mk_var ("dd"^(string_of_int i)^(string_of_int j), real_pair_ty)));; + +let dest_vector = dest_list o rand;; + +let mk_vector list_tm = + let n = (length o dest_list) list_tm in + let ty = (hd o snd o dest_type o type_of) list_tm in + let vec = mk_const ("vector", [ty, aty; n_type_array.(n), nty]) in + mk_comb (vec, list_tm);; + +let mk_vector_list list = + mk_vector (mk_list (list, type_of (hd list)));; + +let el_thms_array = + let el_tm = `EL : num->(A)list->A` in + let gen0 n = + let e_list = mk_list (map (fun i -> mk_var ("e"^(string_of_int i), aty)) (1--n), aty) in + let el0_th = REWRITE_CONV[EL; HD] (mk_binop el_tm `0` e_list) in + Array.make n el0_th in + let array = Array.init (max_dim + 1) gen0 in + let gen_i n i = + let e_list = (rand o lhand o concl) array.(n).(i) in + let prev_thm = array.(n - 1).(i - 1) in + let i_tm = mk_small_numeral i in + let prev_i = num_CONV i_tm in + let el_th = REWRITE_CONV[prev_i; EL; HD; TL; prev_thm] (mk_binop el_tm i_tm e_list) in + array.(n).(i) <- el_th in + let _ = map (fun n -> map (fun i -> gen_i n i) (1--(n - 1))) (2--max_dim) in + array;; + + +let VECTOR_COMPONENT = prove(`!l i. i IN 1..dimindex (:N) ==> + (vector l:A^N)$i = EL (i - 1) l`, +REWRITE_TAC[IN_NUMSEG] THEN REPEAT GEN_TAC THEN DISCH_TAC THEN REWRITE_TAC[vector] THEN + MATCH_MP_TAC LAMBDA_BETA THEN ASM_REWRITE_TAC[]);; + +let gen_comp_thm n i = + let i_tm = mk_small_numeral i and + x_list = mk_list (map (fun i -> mk_var("x"^(string_of_int i), aty)) (1--n), aty) in + let th0 = (ISPECL [x_list; i_tm] o inst_first_type_var (n_type_array.(n))) VECTOR_COMPONENT in + let th1 = (CONV_RULE NUM_REDUCE_CONV o REWRITE_RULE[IN_NUMSEG; dimindex_array.(n)]) th0 in + REWRITE_RULE[el_thms_array.(n).(i - 1)] th1;; + +let comp_thms_array = Array.init (max_dim + 1) + (fun n -> Array.init (n + 1) + (fun i -> if i < 1 || n < 1 then TRUTH else gen_comp_thm n i));; + + +(************************************) +(* m_cell_domain *) + +let ALL2_ALL_ZIP = prove(`!(P:A->B->bool) l1 l2. LENGTH l1 = LENGTH l2 ==> + (ALL2 P l1 l2 <=> ALL (\p. P (FST p) (SND p)) (ZIP l1 l2))`, + GEN_TAC THEN LIST_INDUCT_TAC THENL + [ + GEN_TAC THEN REWRITE_TAC[LENGTH; EQ_SYM_EQ; LENGTH_EQ_NIL] THEN + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN + REWRITE_TAC[ZIP; ALL2; ALL]; + ALL_TAC + ] THEN + + LIST_INDUCT_TAC THEN REWRITE_TAC[LENGTH] THENL [ARITH_TAC; ALL_TAC] THEN + REWRITE_TAC[eqSS] THEN DISCH_TAC THEN + REWRITE_TAC[ALL2; ZIP; ALL] THEN + FIRST_X_ASSUM (new_rewrite [] []) THEN ASM_REWRITE_TAC[]);; + +let EL_ZIP = prove(`!(l1:(A)list) (l2:(B)list) i. LENGTH l1 = LENGTH l2 /\ i < LENGTH l1 ==> + EL i (ZIP l1 l2) = (EL i l1, EL i l2)`, + LIST_INDUCT_TAC THEN LIST_INDUCT_TAC THEN REWRITE_TAC[ZIP; LENGTH] THEN TRY ARITH_TAC THEN + case THEN REWRITE_TAC[EL; HD; TL] THEN GEN_TAC THEN + REWRITE_TAC[eqSS; ARITH_RULE `SUC n < SUC x <=> n < x`] THEN STRIP_TAC THEN + FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_REWRITE_TAC[]);; + +let LENGTH_ZIP = prove(`!l1 l2. LENGTH l1 = LENGTH l2 ==> LENGTH (ZIP l1 l2) = LENGTH l1`, + LIST_INDUCT_TAC THEN LIST_INDUCT_TAC THEN REWRITE_TAC[ZIP; LENGTH] THEN TRY ARITH_TAC THEN + REWRITE_TAC[eqSS] THEN DISCH_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_REWRITE_TAC[]);; + + + +let test_domain_xi = new_definition + `test_domain_xi xz yw <=> FST xz <= FST yw /\ FST yw <= SND xz /\ + FST yw - FST xz <= SND yw /\ SND xz - FST yw <= SND yw`;; + + +let MK_CELL_DOMAIN = prove(`!xz (yw:(real#real)list) x z y w. + LENGTH x = dimindex (:N) /\ LENGTH z = dimindex (:N) /\ + LENGTH y = dimindex (:N) /\ LENGTH w = dimindex (:N) /\ + ZIP y w = yw /\ ZIP x z = xz /\ + ALL2 test_domain_xi xz yw ==> + m_cell_domain (vector x, vector z:real^N) (vector y) (vector w)`, + REPEAT GEN_TAC THEN STRIP_TAC THEN POP_ASSUM MP_TAC THEN + SUBGOAL_THEN `LENGTH (xz:(real#real)list) = dimindex (:N) /\ LENGTH (yw:(real#real)list) = dimindex (:N)` ASSUME_TAC THENL + [ + EXPAND_TAC "yw" THEN EXPAND_TAC "xz" THEN + REPEAT (new_rewrite [] [] LENGTH_ZIP) THEN ASM_REWRITE_TAC[]; + ALL_TAC + ] THEN + rewrite [] [] ALL2_ALL_ZIP THEN ASM_REWRITE_TAC[m_cell_domain; GSYM ALL_EL] THEN DISCH_TAC THEN + REWRITE_TAC[m_cell_domain] THEN GEN_TAC THEN DISCH_TAC THEN + REPEAT (new_rewrite [] [] VECTOR_COMPONENT) THEN ASM_REWRITE_TAC[] THEN + ABBREV_TAC `j = i - 1` THEN + SUBGOAL_THEN `j < dimindex (:N)` ASSUME_TAC THENL + [ + POP_ASSUM MP_TAC THEN POP_ASSUM MP_TAC THEN REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; + ALL_TAC + ] THEN + FIRST_X_ASSUM (MP_TAC o SPEC `j:num`) THEN REWRITE_TAC[test_domain_xi] THEN + rewrite [] [] LENGTH_ZIP THEN ASM_REWRITE_TAC[] THEN + rewrite [] [] EL_ZIP THEN ASM_REWRITE_TAC[] THEN + EXPAND_TAC "xz" THEN EXPAND_TAC "yw" THEN + REPEAT (new_rewrite [] [] EL_ZIP) THEN ASM_REWRITE_TAC[] THEN + ARITH_TAC);; + + + +(* array of theorems *) +let mk_m_domain_array = + let mk_m_domain n = + let dimindex_th = dimindex_array.(n) in + let n_ty = (hd o snd o dest_type o snd o dest_const o rand o lhand o concl) dimindex_th in + let nty = `:N` in + (UNDISCH_ALL o REWRITE_RULE[float0_eq] o DISCH_ALL o RULE o + REWRITE_RULE[dimindex_th] o INST_TYPE[n_ty, nty]) MK_CELL_DOMAIN in + Array.init (max_dim + 1) (fun i -> if i < 1 then TRUTH else mk_m_domain i);; + + +let TEST_DOMAIN_XI' = (EQT_INTRO o RULE o prove)(`xz = (x,z) /\ yw = (y,w) /\ + x <= y /\ y <= z /\ y - x <= w1 /\ z - y <= w2 /\ w1 <= w /\ w2 <= w ==> test_domain_xi xz yw`, + SIMP_TAC[test_domain_xi] THEN REAL_ARITH_TAC);; + + +let eval_test_domain_xi pp test_domain_tm = + let ltm, yw = dest_comb test_domain_tm in + let xz = rand ltm in + let x, z = dest_pair xz and + y, w = dest_pair yw in + let (<=) = (fun t1 t2 -> EQT_ELIM (float_le t1 t2)) and + (-) = float_sub_hi pp in + let x_le_y = x <= y and + y_le_z = y <= z and + yx_le_w1 = y - x and + zy_le_w2 = z - y in + let w1 = (rand o concl) yx_le_w1 and + w2 = (rand o concl) zy_le_w2 in + let w1_le_w = w1 <= w and + w2_le_w = w2 <= w in + (MY_PROVE_HYP (REFL xz) o MY_PROVE_HYP (REFL yw) o + MY_PROVE_HYP x_le_y o MY_PROVE_HYP y_le_z o + MY_PROVE_HYP yx_le_w1 o MY_PROVE_HYP zy_le_w2 o + MY_PROVE_HYP w1_le_w o MY_PROVE_HYP w2_le_w o + INST[x, x_var_real; y, y_var_real; z, z_var_real; w, w_var_real; + w1, w1_var_real; w2, w2_var_real; + xz, xz_pair_var; yw, yw_pair_var]) TEST_DOMAIN_XI';; + + +(* mk_m_center_domain *) +let mk_m_center_domain n pp x_list_tm z_list_tm = + let x_list = dest_list x_list_tm and + z_list = dest_list z_list_tm in + let y_list = + let ( * ) = (fun t1 t2 -> (rand o concl) (float_mul_eq t1 t2)) and + (+) = (fun t1 t2 -> (rand o concl) (float_add_hi pp t1 t2)) in + map2 (fun x y -> if x = y then x else float_inv2 * (x + y)) x_list z_list in + + let w_list = + let (-) = (fun t1 t2 -> (rand o concl) (float_sub_hi pp t1 t2)) and + max = (fun t1 t2 -> (rand o concl) (float_max t1 t2)) in + let w1 = map2 (-) y_list x_list and + w2 = map2 (-) z_list y_list in + map2 max w1 w2 in + + let y_list_tm = mk_list (y_list, real_ty) and + w_list_tm = mk_list (w_list, real_ty) in + + let yw_zip_th = eval_zip y_list_tm w_list_tm and + xz_zip_th = eval_zip x_list_tm z_list_tm in + + let yw_list_tm = (rand o concl) yw_zip_th and + xz_list_tm = (rand o concl) xz_zip_th in + + let len_x_th = eval_length x_list_tm and + len_z_th = eval_length z_list_tm and + len_y_th = eval_length y_list_tm and + len_w_th = eval_length w_list_tm in + let th0 = (MY_PROVE_HYP len_x_th o MY_PROVE_HYP len_z_th o + MY_PROVE_HYP len_y_th o MY_PROVE_HYP len_w_th o + MY_PROVE_HYP yw_zip_th o MY_PROVE_HYP xz_zip_th o + INST[x_list_tm, x_var_real_list; z_list_tm, z_var_real_list; + y_list_tm, y_var_real_list; w_list_tm, w_var_real_list; + yw_list_tm, yw_var; xz_list_tm, xz_var]) mk_m_domain_array.(n) in + let all_th = (EQT_ELIM o all2_conv_univ (eval_test_domain_xi pp) o hd o hyp) th0 in + MY_PROVE_HYP all_th th0;; + + + +(***********************) + +let MK_M_TAYLOR_INTERVAL' = (RULE o MATCH_MP iffRL o SPEC_ALL) m_taylor_interval;; + + +let get_types_and_vars n = + let ty = n_type_array.(n) and + xty = n_vector_type_array.(n) in + let x_var = mk_var ("x", xty) and + f_var = mk_var ("f", mk_fun_ty xty real_ty) and + y_var = mk_var ("y", xty) and + w_var = mk_var ("w", xty) and + domain_var = mk_var ("domain", mk_type ("prod", [xty; xty])) in + ty, xty, x_var, f_var, y_var, w_var, domain_var;; + + + +let dest_m_cell_domain domain_tm = + let lhs, w_tm = dest_comb domain_tm in + let lhs2, y_tm = dest_comb lhs in + rand lhs2, y_tm, w_tm;; + + +(**************************************************) + +(* Given a variable of the type `:real^N`, returns the number N *) +let get_dim = Num.int_of_num o dest_finty o hd o tl o snd o dest_type o type_of;; + + +(**********************) +(* eval_m_taylor_poly *) + +let partial_pow = prove(`!i f n (y:real^N). lift o f differentiable at y ==> + partial i (\x. f x pow n) y = &n * f y pow (n - 1) * partial i f y`, + REPEAT STRIP_TAC THEN + SUBGOAL_THEN `(\x:real^N. f x pow n) = (\t. t pow n) o f` (fun th -> REWRITE_TAC[th]) THENL + [ + ONCE_REWRITE_TAC[GSYM eq_ext] THEN REWRITE_TAC[o_THM]; + ALL_TAC + ] THEN + new_rewrite [] [] partial_uni_compose THENL + [ + ASM_REWRITE_TAC[] THEN + new_rewrite [] [] REAL_DIFFERENTIABLE_POW_ATREAL THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID]; + ALL_TAC + ] THEN + new_rewrite [] [] derivative_pow THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID; derivative_x] THEN + REAL_ARITH_TAC);; + + +let nth_diff2_pow = prove(`!n y. nth_diff_strong 2 (\x. x pow n) y`, + REWRITE_TAC[nth_diff_strong2_eq] THEN REPEAT GEN_TAC THEN + EXISTS_TAC `(:real)` THEN REWRITE_TAC[REAL_OPEN_UNIV; IN_UNIV] THEN GEN_TAC THEN + new_rewrite [] [] REAL_DIFFERENTIABLE_POW_ATREAL THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID] THEN + MATCH_MP_TAC differentiable_local THEN + EXISTS_TAC `\x. &n * x pow (n - 1)` THEN EXISTS_TAC `(:real)` THEN + REWRITE_TAC[REAL_OPEN_UNIV; IN_UNIV] THEN + new_rewrite [] [] REAL_DIFFERENTIABLE_MUL_ATREAL THEN REWRITE_TAC[REAL_DIFFERENTIABLE_CONST] THENL + [ + new_rewrite [] [] REAL_DIFFERENTIABLE_POW_ATREAL THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID]; + ALL_TAC + ] THEN + GEN_TAC THEN new_rewrite [] [] derivative_pow THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID] THEN + REWRITE_TAC[derivative_x; REAL_MUL_RID]);; + + +let diff2c_pow = prove(`!f n (x:real^N). diff2c f x ==> diff2c (\x. f x pow n) x`, + REPEAT STRIP_TAC THEN + SUBGOAL_THEN `(\x:real^N. f x pow n) = (\t. t pow n) o f` (fun th -> REWRITE_TAC[th]) THENL + [ + ONCE_REWRITE_TAC[GSYM eq_ext] THEN REWRITE_TAC[o_THM]; + ALL_TAC + ] THEN + apply_tac diff2c_uni_compose THEN ASM_REWRITE_TAC[nth_diff2_pow] THEN + REWRITE_TAC[nth_derivative2] THEN + SUBGOAL_THEN `!n. derivative (\t. t pow n) = (\t. &n * t pow (n - 1))` ASSUME_TAC THENL + [ + GEN_TAC THEN REWRITE_TAC[FUN_EQ_THM] THEN GEN_TAC THEN + new_rewrite [] [] derivative_pow THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID] THEN + REWRITE_TAC[derivative_x; REAL_MUL_RID]; + ALL_TAC + ] THEN + ASM_REWRITE_TAC[] THEN + SUBGOAL_THEN `!n. derivative (\t. &n * t pow (n - 1)) = (\t. &n * derivative (\t. t pow (n - 1)) t)` ASSUME_TAC THENL + [ + GEN_TAC THEN REWRITE_TAC[FUN_EQ_THM] THEN GEN_TAC THEN + new_rewrite [] [] derivative_scale THEN REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_DIFFERENTIABLE_POW_ATREAL THEN REWRITE_TAC[REAL_DIFFERENTIABLE_ID]; + ALL_TAC + ] THEN + ASM_REWRITE_TAC[] THEN + REPEAT (MATCH_MP_TAC REAL_CONTINUOUS_LMUL) THEN + MATCH_MP_TAC REAL_CONTINUOUS_POW THEN + REWRITE_TAC[REAL_CONTINUOUS_AT_ID]);; + + +let diff2c_domain_pow = prove(`!f n domain. diff2c_domain domain f ==> + diff2c_domain domain (\x. f x pow n)`, + REWRITE_TAC[diff2c_domain] THEN REPEAT STRIP_TAC THEN ASM_SIMP_TAC[diff2c_pow]);; + + +let diff2c_domain_tm = `diff2c_domain domain`;; + +let rec gen_diff2c_domain_poly poly_tm = + let x_var, expr = dest_abs poly_tm in + let n = get_dim x_var in + let diff2c_tm = mk_icomb (diff2c_domain_tm, poly_tm) in + if frees expr = [] then + (* const *) + (SPEC_ALL o ISPEC expr o inst_first_type_var (n_type_array.(n))) diff2c_domain_const + else + let lhs, r_tm = dest_comb expr in + if lhs = neg_op_real then + (* -- *) + let r_th = gen_diff2c_domain_poly (mk_abs (x_var, r_tm)) in + prove(diff2c_tm, MATCH_MP_TAC diff2c_domain_neg THEN REWRITE_TAC[r_th]) + else + let op, l_tm = dest_comb lhs in + let name = (fst o dest_const) op in + if name = "$" then + (* x$k *) + let dim_th = dimindex_array.(n) in + prove(diff2c_tm, MATCH_MP_TAC diff2c_domain_x THEN + REWRITE_TAC[IN_NUMSEG; dim_th] THEN ARITH_TAC) + else + let l_th = gen_diff2c_domain_poly (mk_abs (x_var, l_tm)) in + if name = "real_pow" then + (* f pow n *) + prove(diff2c_tm, MATCH_MP_TAC diff2c_domain_pow THEN REWRITE_TAC[l_th]) + else + let r_th = gen_diff2c_domain_poly (mk_abs (x_var, r_tm)) in + prove(diff2c_tm, + MAP_FIRST apply_tac [diff2c_domain_add; diff2c_domain_sub; diff2c_domain_mul] THEN + REWRITE_TAC[l_th; r_th]);; + + +let gen_diff2c_poly = + let th_imp = prove(`!f. (!domain. diff2c_domain domain f) ==> !x:real^N. diff2c f x`, + REWRITE_TAC[diff2c_domain] THEN REPEAT STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN + EXISTS_TAC `(x:real^N, x:real^N)` THEN + REWRITE_TAC[INTERVAL_SING; IN_SING]) in + fun poly_tm -> + (MATCH_MP th_imp o GEN_ALL o gen_diff2c_domain_poly) poly_tm;; + + +let gen_diff_poly = + let th_imp = prove(`!f. (!domain. diff2c_domain domain f) ==> !x:real^N. lift o f differentiable at x`, + REWRITE_TAC[diff2c_domain; diff2c; diff2] THEN REPEAT STRIP_TAC THEN + FIRST_X_ASSUM (MP_TAC o SPECL [`x:real^N, x:real^N`; `x:real^N`]) THEN + REWRITE_TAC[INTERVAL_SING; IN_SING] THEN case THEN REPEAT STRIP_TAC THEN + FIRST_X_ASSUM (MP_TAC o SPEC `x:real^N`) THEN ASM_SIMP_TAC[]) in + fun poly_tm -> + (MATCH_MP th_imp o GEN_ALL o gen_diff2c_domain_poly) poly_tm;; + + +let in_tm = `IN`;; + +let add_to_hash tbl max_size key value = + let _ = if Hashtbl.length tbl >= max_size then Hashtbl.clear tbl else () in + Hashtbl.add tbl key value;; + +(* Formally computes partial derivatives of a polynomial *) +let gen_partial_poly = + let max_hash = 1000 in + let hash = Hashtbl.create max_hash in + fun i poly_tm -> + let key = (i, poly_tm) in + try Hashtbl.find hash (i, poly_tm) + with Not_found -> + let i_tm = mk_small_numeral i in + let rec gen_rec poly_tm = + let x_var, expr = dest_abs poly_tm in + let n = get_dim x_var in + if frees expr = [] then + (* const *) + (SPECL [i_tm; expr] o inst_first_type_var (n_type_array.(n))) partial_const + else + let lhs, r_tm = dest_comb expr in + if lhs = neg_op_real then + (* -- *) + let r_poly = mk_abs (x_var, r_tm) in + let r_diff = (SPEC_ALL o gen_diff_poly) r_poly and + r_partial = gen_rec r_poly in + let th0 = SPEC i_tm (MATCH_MP partial_neg r_diff) in + REWRITE_RULE[r_partial] th0 + else + let op, l_tm = dest_comb lhs in + let name = (fst o dest_const) op in + if name = "$" then + (* comp *) + let dim_th = dimindex_array.(n) in + let dim_tm = (lhand o concl) dim_th in + let i_eq_k = NUM_EQ_CONV (mk_eq (i_tm, r_tm)) in + let int_tm = mk_binop `..` `1` dim_tm in + let k_in_dim = prove(mk_comb (mk_icomb(in_tm, r_tm), int_tm), + REWRITE_TAC[IN_NUMSEG; dim_th] THEN ARITH_TAC) in + (REWRITE_RULE[i_eq_k] o MATCH_MP (SPECL [r_tm; i_tm] partial_x)) k_in_dim + else + let l_poly = mk_abs (x_var, l_tm) in + let l_partial = gen_rec l_poly in + let l_diff = (SPEC_ALL o gen_diff_poly) l_poly in + if name = "real_pow" then + (* f pow n *) + let th0 = SPECL [i_tm; r_tm] (MATCH_MP partial_pow l_diff) in + REWRITE_RULE[l_partial] th0 + else + let r_poly = mk_abs (x_var, r_tm) in + let r_partial = gen_rec r_poly in + let r_diff = (SPEC_ALL o gen_diff_poly) r_poly in + let imp_th = assoc op [add_op_real, partial_add; + sub_op_real, partial_sub; + mul_op_real, partial_mul] in + let th0 = SPEC i_tm (MATCH_MP (MATCH_MP imp_th l_diff) r_diff) in + REWRITE_RULE[l_partial; r_partial] th0 in + + let th1 = gen_rec poly_tm in + let th2 = ((NUM_REDUCE_CONV THENC REWRITE_CONV[DECIMAL] THENC REAL_POLY_CONV) o rand o concl) th1 in + let th3 = (REWRITE_RULE[ETA_AX] o ONCE_REWRITE_RULE[eq_ext] o GEN_ALL) (TRANS th1 th2) in + let _ = add_to_hash hash max_hash key th3 in + th3;; + + +let gen_partial2_poly i j poly_tm = + let partial_j = gen_partial_poly j poly_tm in + let partial_ij = gen_partial_poly i (rand (concl partial_j)) in + let pi = (rator o lhand o concl) partial_ij in + REWRITE_RULE[GSYM partial2] (TRANS (AP_TERM pi partial_j) partial_ij);; + + +(********************************************) + +let eval_diff2_poly diff2_domain_th = + fun xx zz -> + let domain_tm = mk_pair (xx, zz) in + INST[domain_tm, mk_var ("domain", type_of domain_tm)] diff2_domain_th;; + + +(*****************************) +(* m_lin_approx *) + +let CONST_INTERVAL' = RULE CONST_INTERVAL;; + +let dest_lin_approx approx_tm = + let lhs, df_bounds = dest_comb approx_tm in + let lhs2, f_bounds = dest_comb lhs in + let lhs3, x_tm = dest_comb lhs2 in + let f_tm = rand lhs3 in + f_tm, x_tm, f_bounds, df_bounds;; + +let gen_lin_approx_eq_thm n = + let ty = n_type_array.(n) in + let df_vars = map (fun i -> df_vars_array.(i)) (1--n) in + let df_bounds_list = mk_list (df_vars, real_pair_ty) in + let th0 = (SPECL[f_bounds_var; df_bounds_list] o inst_first_type_var ty) m_lin_approx in + let th1 = (CONV_RULE NUM_REDUCE_CONV o REWRITE_RULE[all_n]) th0 in + th1;; + + +let gen_lin_approx_poly_thm poly_tm diff_th partials = + let x_var, _ = dest_abs poly_tm in + let n = get_dim x_var in + let lin_eq = (REWRITE_RULE partials o SPECL [poly_tm]) (gen_lin_approx_eq_thm n) in + let x_vec = mk_vector_list (map (fun i -> x_vars_array.(i)) (1--n)) in + let th1 = (REWRITE_RULE (Array.to_list comp_thms_array.(n)) o SPEC x_vec o REWRITE_RULE[diff_th]) lin_eq in + th1;; + + +let gen_lin_approx_poly_thm0 poly_tm = + let x_var, _ = dest_abs poly_tm in + let n = get_dim x_var in + let partials = map (fun i -> gen_partial_poly i poly_tm) (1--n) in + let diff_th = gen_diff_poly poly_tm in + gen_lin_approx_poly_thm poly_tm diff_th partials;; + + +let eval_lin_approx pp0 lin_approx_th = + let poly_tm, _, _, _ = (dest_lin_approx o lhand o concl) lin_approx_th in + let x_var, _ = dest_abs poly_tm in + let n = get_dim x_var in + let th0 = lin_approx_th in + let th1 = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o MATCH_MP iffRL) th0 in + let build_eval int_hyp = + let expr, b_var = dest_binary "interval_arith" int_hyp in + (eval_constants pp0 o build_interval_fun) expr, b_var in + let int_fs = map build_eval (hyp th1) in + + let rec split_rules i_list = + match i_list with + | [] -> ([], []) + | ((i_fun, var_tm) :: es) -> + let th_list, i_list' = split_rules es in + match i_fun with + | Int_const th -> (var_tm, th) :: th_list, i_list' + | Int_var v -> (var_tm, INST[v, x_var_real] CONST_INTERVAL') :: th_list, i_list' + | _ -> th_list, (var_tm, i_fun) :: i_list' in + + let const_th_list, i_list0 = split_rules int_fs in + let th2 = itlist (fun (var_tm, th) th0 -> + let b_tm = rand (concl th) in + (MY_PROVE_HYP th o INST[b_tm, var_tm]) th0) const_th_list th1 in + let v_list, i_list' = unzip i_list0 in + let i_list = find_and_replace_all i_list' [] in + fun pp vector_tm -> + let x_vals = dest_vector vector_tm in + if length x_vals <> n then failwith (sprintf "Wrong vector size; expected size: %d" n) + else + let x_ints = map mk_const_interval x_vals in + let vars = map (fun i -> x_vars_array.(i)) (1--n) in + let th3 = INST (zip x_vals vars) th2 in + let i_vals = eval_interval_fun_list pp i_list (zip vars x_ints) in + itlist2 (fun var_tm th th0 -> + let b_tm = rand (concl th) in + (MY_PROVE_HYP th o INST[b_tm, var_tm]) th0) v_list i_vals th3;; + + +let eval_lin_approx_poly0 pp0 poly_tm = + eval_lin_approx pp0 (gen_lin_approx_poly_thm0 poly_tm);; + + + +(*************************************) + +(* 1 <= i /\ i <= n <=> i = 1 \/ i = 2 \/ ... \/ i = n *) +let i_int_array = + let i_tm = `i:num` in + let i_th0 = prove(`1 <= i /\ i <= SUC n <=> (1 <= i /\ i <= n) \/ i = SUC n`, ARITH_TAC) in + let th1 = prove(`1 <= i /\ i <= 1 <=> i = 1`, ARITH_TAC) in + let array = Array.make (max_dim + 1) th1 in + let prove_next n = + let n_tm = mk_small_numeral n in + let prev_n = num_CONV n_tm in + let tm = mk_conj (`1 <= i`, mk_binop le_op_num i_tm n_tm) in + let th = REWRITE_CONV[prev_n; i_th0; array.(n - 1)] tm in + array.(n) <- REWRITE_RULE[SYM prev_n; GSYM DISJ_ASSOC] th in + let _ = map prove_next (2--max_dim) in + array;; + + +(* (!i. 1 <= i /\ i <= n ==> P i) <=> P 1 /\ P 2 /\ ... /\ P n *) +let gen_in_interval = + let th0 = prove(`(!i:num. (i = k \/ Q i) ==> P i) <=> (P k /\ (!i. Q i ==> P i))`, MESON_TAC[]) in + let th1 = prove(`(!i:num. (i = k ==> P i)) <=> P k`, MESON_TAC[]) in + fun n -> + let n_tm = mk_small_numeral n and + i_tm = `i:num` in + let lhs1 = mk_conj (`1 <= i`, mk_binop le_op_num i_tm n_tm) in + let lhs = mk_forall (i_tm, mk_imp (lhs1, `(P:num->bool) i`)) in + REWRITE_CONV[i_int_array.(n); th0; th1] lhs;; + + +let gen_second_bounded_eq_thm n = + let ty, _, x_var, _, _, _, domain_var = get_types_and_vars n in + let dd_vars = map (fun i -> map (fun j -> dd_vars_array.(j).(i)) (1--i)) (1--n) in + let dd_bounds_list = mk_list (map (fun l -> mk_list (l, real_pair_ty)) dd_vars, real_pair_list_ty) in + let th0 = (SPECL[domain_var; dd_bounds_list] o inst_first_type_var ty) second_bounded in + let th1 = (CONV_RULE NUM_REDUCE_CONV o REWRITE_RULE[all_n]) th0 in + th1;; + + +let gen_second_bounded_poly_thm poly_tm partials2 = + let x_var, _ = dest_abs poly_tm in + let n = get_dim x_var in + let partials2' = List.flatten partials2 in + let second_th = (REWRITE_RULE partials2' o SPECL [poly_tm]) (gen_second_bounded_eq_thm n) in + second_th;; + + +let gen_second_bounded_poly_thm0 poly_tm = + let x_var, _ = dest_abs poly_tm in + let n = get_dim x_var in + let partials = map (fun i -> gen_partial_poly i poly_tm) (1--n) in + let get_partial i eq_th = + let partial_i = gen_partial_poly i (rand (concl eq_th)) in + let pi = (rator o lhand o concl) partial_i in + REWRITE_RULE[GSYM partial2] (TRANS (AP_TERM pi eq_th) partial_i) in + let partials2 = map (fun th, i -> map (fun j -> get_partial j th) (1--i)) (zip partials (1--n)) in + gen_second_bounded_poly_thm poly_tm partials2;; + +(* let eq_th = TAUT `(P ==> Q /\ R) <=> ((P ==> Q) /\ (P ==> R))` in + REWRITE_RULE[eq_th; FORALL_AND_THM; GSYM m_bounded_on_int] second_th;;*) + + +(* eval_second_bounded *) +let eval_second_bounded pp0 second_bounded_th = + let poly_tm = (lhand o rator o lhand o concl) second_bounded_th in + let th0 = second_bounded_th in + let n = (get_dim o fst o dest_abs) poly_tm in + let x_vector = mk_vector_list (map (fun i -> x_vars_array.(i)) (1--n)) and + z_vector = mk_vector_list (map (fun i -> z_vars_array.(i)) (1--n)) in + let _, _, _, _, _, _, domain_var = get_types_and_vars n in + + let th1 = INST[mk_pair (x_vector, z_vector), domain_var] th0 in + let th2 = REWRITE_RULE[IN_INTERVAL; dimindex_array.(n)] th1 in + let th3 = REWRITE_RULE[gen_in_interval n; GSYM interval_arith] th2 in + let th4 = (REWRITE_RULE[CONJ_ACI] o REWRITE_RULE (Array.to_list comp_thms_array.(n))) th3 in + let final_th0 = (UNDISCH_ALL o MATCH_MP iffRL) th4 in + + let x_var, h_tm = (dest_forall o hd o hyp) final_th0 in + let _, h2 = dest_imp h_tm in + let concl_ints = striplist dest_conj h2 in + + let i_funs = map (fun int -> + let expr, var = dest_interval_arith int in + (eval_constants pp0 o build_interval_fun) expr, var) concl_ints in + + let rec split_rules i_list = + match i_list with + | [] -> ([], []) + | ((i_fun, var_tm) :: es) -> + let th_list, i_list' = split_rules es in + match i_fun with + | Int_const th -> (var_tm, th) :: th_list, i_list' +(* | Int_var v -> (var_tm, INST[v, x_var_real] CONST_INTERVAL') :: th_list, i_list' *) + | _ -> th_list, (var_tm, i_fun) :: i_list' in + + let const_th_list, i_list0 = split_rules i_funs in + let th5 = itlist (fun (var_tm, th) th0 -> + let b_tm = rand (concl th) in + (REWRITE_RULE[th] o INST[b_tm, var_tm]) th0) const_th_list (SYM th4) in + let final_th = REWRITE_RULE[GSYM IMP_IMP] th5 in + let v_list, i_list' = unzip i_list0 in + let i_list = find_and_replace_all i_list' [] in + + fun pp x_vector_tm z_vector_tm -> + let x_vals = dest_vector x_vector_tm and + z_vals = dest_vector z_vector_tm in + if length x_vals <> n || length z_vals <> n then + failwith (sprintf "Wrong vector size; expected size: %d" n) + else + let x_vars = map (fun i -> x_vars_array.(i)) (1--n) and + z_vars = map (fun i -> z_vars_array.(i)) (1--n) in + + let inst_th = (INST (zip x_vals x_vars) o INST (zip z_vals z_vars)) final_th in + if (not o is_eq) (concl inst_th) then inst_th + else + let x_var, lhs = (dest_forall o lhand o concl) inst_th in + let hs = (butlast o striplist dest_imp) lhs in + let vars = map (rand o rator) hs in + let int_vars = zip vars (map ASSUME hs) in + + let dd_ints = eval_interval_fun_list pp i_list int_vars in + let inst_dd = map2 (fun var th -> (rand o concl) th, var) v_list dd_ints in + let inst_th2 = INST inst_dd inst_th in + + let conj_th = end_itlist CONJ dd_ints in + let lhs_th = GEN x_var (itlist DISCH hs conj_th) in + EQ_MP inst_th2 lhs_th;; + + + +let eval_second_bounded_poly0 pp0 poly_tm = + eval_second_bounded pp0 (gen_second_bounded_poly_thm0 poly_tm);; + + + +(*************************************) +(* eval_m_taylor *) + +let eval_m_taylor pp0 diff2c_th lin_th second_th = + let poly_tm = (rand o concl) diff2c_th in + let n = (get_dim o fst o dest_abs) poly_tm in + let eval_lin = eval_lin_approx pp0 lin_th and + eval_second = eval_second_bounded pp0 second_th in + + let ty, _, x_var, f_var, y_var, w_var, domain_var = get_types_and_vars n in + let th0 = (SPEC_ALL o inst_first_type_var ty) m_taylor_interval in + let th1 = INST[poly_tm, f_var] th0 in + let th2 = (UNDISCH_ALL o REWRITE_RULE[GSYM IMP_IMP] o MATCH_MP iffRL o REWRITE_RULE[diff2c_th]) th1 in + + fun p_lin p_second domain_th -> + let domain_tm, y_tm, w_tm = dest_m_cell_domain (concl domain_th) in + let x_tm, z_tm = dest_pair domain_tm in + + let lin_th = eval_lin p_lin y_tm and + second_th = eval_second p_second x_tm z_tm in + + let _, _, f_bounds, df_bounds_list = dest_lin_approx (concl lin_th) in + let dd_bounds_list = (rand o concl) second_th in + let df_var = mk_var ("d_bounds_list", type_of df_bounds_list) and + dd_var = mk_var ("dd_bounds_list", type_of dd_bounds_list) in + + (MY_PROVE_HYP domain_th o MY_PROVE_HYP lin_th o MY_PROVE_HYP second_th o + INST[domain_tm, domain_var; y_tm, y_var; w_tm, w_var; + f_bounds, f_bounds_var; df_bounds_list, df_var; dd_bounds_list, dd_var]) th2;; + + +let eval_m_taylor_poly0 pp0 poly_tm = + let diff2_th = gen_diff2c_domain_poly poly_tm in + let x_var, _ = dest_abs poly_tm in + let n = get_dim x_var in + let partials = map (fun i -> gen_partial_poly i poly_tm) (1--n) in + let get_partial i eq_th = + let partial_i = gen_partial_poly i (rand (concl eq_th)) in + let pi = (rator o lhand o concl) partial_i in + REWRITE_RULE[GSYM partial2] (TRANS (AP_TERM pi eq_th) partial_i) in + let partials2 = map2 (fun th i -> map (fun j -> get_partial j th) (1--i)) partials (1--n) in + let second_th = gen_second_bounded_poly_thm poly_tm partials2 in + let diff_th = gen_diff_poly poly_tm in + let lin_th = gen_lin_approx_poly_thm poly_tm diff_th partials in + eval_m_taylor pp0 diff2_th lin_th second_th;; + + + +(******************************************) +(* mk_eval_function *) + +let mk_eval_function_eq pp0 eq_th = + let expr_tm = (rand o concl) eq_th in + let tm0 = `!x:real^N. x IN interval [domain] ==> interval_arith (f x) f_bounds` in + let n = (get_dim o fst o dest_abs) expr_tm in + let x_vector = mk_vector_list (map (fun i -> x_vars_array.(i)) (1--n)) and + z_vector = mk_vector_list (map (fun i -> z_vars_array.(i)) (1--n)) in + let ty, _, _, _, _, _, domain_var = get_types_and_vars n and + f_var = mk_var ("f", type_of expr_tm) in + let th1 = (REWRITE_CONV[IN_INTERVAL] o subst[mk_pair(x_vector,z_vector), domain_var] o inst[ty, nty]) tm0 in + let th2 = REWRITE_RULE [dimindex_array.(n)] th1 in + let th3 = REWRITE_RULE [gen_in_interval n; GSYM interval_arith] th2 in + let th4 = (REWRITE_RULE[GSYM IMP_IMP; CONJ_ACI] o REWRITE_RULE (Array.to_list comp_thms_array.(n))) th3 in + let final_th0 = (CONV_RULE ((RAND_CONV o ONCE_DEPTH_CONV) BETA_CONV) o INST[expr_tm, f_var]) th4 in + let x_var, h_tm = (dest_forall o rand o concl) final_th0 in + let f_tm = (fst o dest_interval_arith o last o striplist dest_imp) h_tm in + let i_fun = (eval_constants pp0 o build_interval_fun) f_tm in + let i_list = find_and_replace_all [i_fun] [] in + let final_th = (PURE_REWRITE_RULE[SYM eq_th] o SYM) final_th0 in + fun pp x_vector_tm z_vector_tm -> + let x_vals = dest_vector x_vector_tm and + z_vals = dest_vector z_vector_tm in + if length x_vals <> n || length z_vals <> n then + failwith (sprintf "Wrong vector size; expected size: %d" n) + else + let x_vars = map (fun i -> x_vars_array.(i)) (1--n) and + z_vars = map (fun i -> z_vars_array.(i)) (1--n) in + + let inst_th = (INST (zip x_vals x_vars) o INST (zip z_vals z_vars)) final_th in + let x_var, lhs = (dest_forall o lhand o concl) inst_th in + let hs = (butlast o striplist dest_imp) lhs in + let vars = map (rand o rator) hs in + let int_vars = zip vars (map ASSUME hs) in + let eval_th = hd (eval_interval_fun_list pp i_list int_vars) in + let f_bounds = (rand o concl) eval_th in + let inst_th2 = INST[f_bounds, f_bounds_var] inst_th in + let lhs_th = GEN x_var (itlist DISCH hs eval_th) in + EQ_MP inst_th2 lhs_th;; + + +let mk_eval_function pp0 expr_tm = mk_eval_function_eq pp0 (REFL expr_tm);; + + + +(********************************) +(* m_taylor_error *) + +(* Sum of the list elements *) +let ITLIST2_EQ_SUM = prove(`!(f:A->B->real) l1 l2. LENGTH l1 <= LENGTH l2 ==> + ITLIST2 (\x y z. f x y + z) l1 l2 (&0) = + sum (1..(LENGTH l1)) (\i. f (EL (i - 1) l1) (EL (i - 1) l2))`, + GEN_TAC THEN + LIST_INDUCT_TAC THEN LIST_INDUCT_TAC THEN REWRITE_TAC[LENGTH; ITLIST2_DEF] THEN TRY ARITH_TAC THENL + [ + REWRITE_TAC[SUM_CLAUSES_NUMSEG; ARITH]; + REWRITE_TAC[SUM_CLAUSES_NUMSEG; ARITH]; + ALL_TAC + ] THEN + REWRITE_TAC[leqSS] THEN DISCH_TAC THEN + FIRST_X_ASSUM (MP_TAC o SPEC `t':(B)list`) THEN ASM_REWRITE_TAC[] THEN + DISCH_THEN (fun th -> REWRITE_TAC[TL; th]) THEN + REWRITE_TAC[GSYM add1n] THEN + new_rewrite [] [] SUM_ADD_SPLIT THEN REWRITE_TAC[ARITH] THEN + REWRITE_TAC[TWO; add1n; SUM_SING_NUMSEG; subnn; EL; HD] THEN + REWRITE_TAC[GSYM addn1; SUM_OFFSET; REAL_EQ_ADD_LCANCEL] THEN + MATCH_MP_TAC SUM_EQ THEN move ["i"] THEN REWRITE_TAC[IN_NUMSEG] THEN DISCH_TAC THEN + ASM_SIMP_TAC[ARITH_RULE `1 <= i ==> (i + 1) - 1 = SUC (i - 1)`; EL; TL]);; + + +let interval_arith_abs_le = prove(`!x int y. interval_arith x int ==> iabs int <= y ==> abs x <= y`, + GEN_TAC THEN case THEN REWRITE_TAC[interval_arith; IABS'] THEN REAL_ARITH_TAC);; + + +let ALL_N_ALL2 = prove(`!P (l:(A)list) i0. + (all_n i0 l P <=> if l = [] then T else ALL2 P (l_seq i0 ((i0 + LENGTH l) - 1)) l)`, + GEN_TAC THEN LIST_INDUCT_TAC THEN GEN_TAC THEN REWRITE_TAC[all_n; NOT_CONS_NIL] THEN + new_rewrite [] [] L_SEQ_CONS THEN REWRITE_TAC[LENGTH; ALL2] THEN TRY ARITH_TAC THEN + FIRST_X_ASSUM (new_rewrite [] []) THEN TRY ARITH_TAC THEN + REWRITE_TAC[addSn; addnS; addn1] THEN + SPEC_TAC (`t:(A)list`, `t:(A)list`) THEN case THEN SIMP_TAC[NOT_CONS_NIL] THEN + REWRITE_TAC[LENGTH; addn0] THEN + MP_TAC (SPECL [`SUC i0`; `SUC i0 - 1`] L_SEQ_NIL) THEN + REWRITE_TAC[ARITH_RULE `SUC i0 - 1 < SUC i0`] THEN DISCH_THEN (fun th -> REWRITE_TAC[th; ALL2]));; + + +let ALL_N_EL = prove(`!P (l:(A)list) i0. all_n i0 l P <=> (!i. i < LENGTH l ==> P (i0 + i) (EL i l))`, + REPEAT GEN_TAC THEN REWRITE_TAC[ALL_N_ALL2] THEN + SPEC_TAC (`l:(A)list`, `l:(A)list`) THEN case THEN SIMP_TAC[NOT_CONS_NIL; LENGTH; ltn0] THEN + REPEAT GEN_TAC THEN + new_rewrite [] [] ALL2_ALL_ZIP THENL + [ + REWRITE_TAC[LENGTH_L_SEQ; LENGTH] THEN ARITH_TAC; + ALL_TAC + ] THEN + REWRITE_TAC[GSYM ALL_EL] THEN + new_rewrite [] [] LENGTH_ZIP THENL + [ + REWRITE_TAC[LENGTH_L_SEQ; LENGTH] THEN ARITH_TAC; + ALL_TAC + ] THEN + REWRITE_TAC[LENGTH_L_SEQ; ARITH_RULE `((i0 + SUC a) - 1 + 1) - i0 = SUC a`] THEN + EQ_TAC THEN REPEAT STRIP_TAC THENL + [ + FIRST_X_ASSUM (MP_TAC o SPEC `i:num`) THEN ASM_REWRITE_TAC[] THEN + new_rewrite [] [] EL_ZIP THENL + [ + REWRITE_TAC[LENGTH_L_SEQ; LENGTH] THEN ASM_ARITH_TAC; + ALL_TAC + ] THEN + REWRITE_TAC[] THEN new_rewrite [] [] EL_L_SEQ THEN ASM_ARITH_TAC; + ALL_TAC + ] THEN + new_rewrite [] [] EL_ZIP THENL + [ + REWRITE_TAC[LENGTH_L_SEQ; LENGTH] THEN ASM_ARITH_TAC; + ALL_TAC + ] THEN + REWRITE_TAC[] THEN new_rewrite [] [] EL_L_SEQ THEN TRY ASM_ARITH_TAC THEN + FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_REWRITE_TAC[]);; + + +let M_TAYLOR_ERROR_ITLIST2 = prove(`!f domain y w dd_bounds_list error. + m_cell_domain domain y (vector w) ==> + diff2c_domain domain f ==> + second_bounded (f:real^N->real) domain dd_bounds_list ==> + LENGTH w = dimindex (:N) ==> + LENGTH dd_bounds_list = dimindex (:N) ==> + all_n 1 dd_bounds_list (\i list. LENGTH list = i) ==> + ITLIST2 (\list x z. x * (x * iabs (LAST list) + + &2 * ITLIST2 (\a b c. b * iabs a + c) (BUTLAST list) w (&0)) + z) + dd_bounds_list w (&0) <= error ==> + m_taylor_error f domain (vector w) error`, + REPEAT GEN_TAC THEN REWRITE_TAC[second_bounded] THEN + set_tac "s" `ITLIST2 _1 _2 _3 _4` THEN + move ["domain"; "d2f"; "second"; "lw"; "ldd"; "ldd_all"; "s_le"] THEN + ASM_SIMP_TAC[m_taylor_error_eq] THEN move ["x"; "x_in"] THEN + SUBGOAL_THEN `!i. i IN 1..dimindex (:N) ==> &0 <= EL (i - 1) w` (LABEL_TAC "w_ge0") THENL + [ + GEN_TAC THEN DISCH_TAC THEN REMOVE_THEN "domain" MP_TAC THEN new_rewrite [] [] pair_eq THEN + REWRITE_TAC[m_cell_domain] THEN + DISCH_THEN (MP_TAC o SPEC `i:num`) THEN ASM_REWRITE_TAC[] THEN + ASM_SIMP_TAC[VECTOR_COMPONENT] THEN REAL_ARITH_TAC; + ALL_TAC + ] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `s:real` THEN ASM_REWRITE_TAC[] THEN + EXPAND_TAC "s" THEN + new_rewrite [] [] ITLIST2_EQ_SUM THEN ASM_REWRITE_TAC[LE_REFL] THEN + MATCH_MP_TAC SUM_LE THEN REWRITE_TAC[FINITE_NUMSEG] THEN + move ["i"; "i_in"] THEN ASM_SIMP_TAC[VECTOR_COMPONENT] THEN + USE_THEN "i_in" (ASSUME_TAC o REWRITE_RULE[IN_NUMSEG]) THEN + MATCH_MP_TAC REAL_LE_LMUL THEN ASM_SIMP_TAC[] THEN + SUBGOAL_THEN `LENGTH (EL (i - 1) dd_bounds_list:(real#real)list) = i` (LABEL_TAC "len_i") THENL + [ + REMOVE_THEN "ldd_all" MP_TAC THEN + REWRITE_TAC[ALL_N_EL] THEN DISCH_THEN (MP_TAC o SPEC `i - 1`) THEN + ANTS_TAC THENL + [ + ASM_REWRITE_TAC[] THEN POP_ASSUM MP_TAC THEN ARITH_TAC; + ALL_TAC + ] THEN + DISCH_THEN (fun th -> REWRITE_TAC[th]) THEN POP_ASSUM MP_TAC THEN ARITH_TAC; + ALL_TAC + ] THEN + new_rewrite [] [] ITLIST2_EQ_SUM THEN ASM_REWRITE_TAC[] THENL + [ + REWRITE_TAC[LENGTH_BUTLAST] THEN POP_ASSUM MP_TAC THEN POP_ASSUM MP_TAC THEN ARITH_TAC; + ALL_TAC + ] THEN + MATCH_MP_TAC REAL_LE_ADD2 THEN CONJ_TAC THENL + [ + new_rewrite [] [] LAST_EL THENL + [ + ASM_REWRITE_TAC[GSYM LENGTH_EQ_NIL] THEN + REMOVE_THEN "i_in" MP_TAC THEN REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; + ALL_TAC + ] THEN + MATCH_MP_TAC REAL_LE_LMUL THEN ASM_SIMP_TAC[] THEN + FIRST_X_ASSUM (MP_TAC o SPEC `x:real^N`) THEN ASM_REWRITE_TAC[ALL_N_EL] THEN + DISCH_THEN (MP_TAC o SPEC `i - 1`) THEN + ANTS_TAC THENL + [ + REMOVE_THEN "i_in" MP_TAC THEN REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; + ALL_TAC + ] THEN + DISCH_THEN (MP_TAC o SPEC `i - 1`) THEN ANTS_TAC THENL + [ + UNDISCH_TAC `1 <= i /\ i <= dimindex (:N)` THEN ASM_REWRITE_TAC[] THEN ARITH_TAC; + ALL_TAC + ] THEN + SUBGOAL_THEN `1 + i - 1 = i /\ 1 + i - 1 = i` (fun th -> REWRITE_TAC[th]) THENL + [ + REPLICATE_TAC 3 (POP_ASSUM MP_TAC) THEN ARITH_TAC; + ALL_TAC + ] THEN + REWRITE_TAC[partial2] THEN + DISCH_THEN (MP_TAC o MATCH_MP interval_arith_abs_le) THEN + DISCH_THEN MATCH_MP_TAC THEN REWRITE_TAC[REAL_LE_REFL]; + ALL_TAC + ] THEN + MATCH_MP_TAC REAL_LE_LMUL THEN CONJ_TAC THENL [ REAL_ARITH_TAC; ALL_TAC ] THEN + ASM_REWRITE_TAC[LENGTH_BUTLAST] THEN + MATCH_MP_TAC SUM_LE THEN REWRITE_TAC[FINITE_NUMSEG] THEN + move ["j"; "j_in"] THEN + SUBGOAL_THEN `j IN 1..dimindex (:N)` (LABEL_TAC "j_in2") THENL + [ + REMOVE_THEN "i_in" MP_TAC THEN REMOVE_THEN "j_in" MP_TAC THEN + REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; + ALL_TAC + ] THEN + ASM_SIMP_TAC[VECTOR_COMPONENT] THEN + USE_THEN "j_in" (ASSUME_TAC o REWRITE_RULE[IN_NUMSEG]) THEN + MATCH_MP_TAC REAL_LE_LMUL THEN ASM_SIMP_TAC[] THEN + FIRST_X_ASSUM (MP_TAC o SPEC `x:real^N`) THEN ASM_REWRITE_TAC[ALL_N_EL] THEN + DISCH_THEN (MP_TAC o SPEC `i - 1`) THEN ANTS_TAC THENL + [ + REMOVE_THEN "i_in" MP_TAC THEN REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; + ALL_TAC + ] THEN + DISCH_THEN (MP_TAC o SPEC `j - 1`) THEN ANTS_TAC THENL + [ + ASM_REWRITE_TAC[] THEN POP_ASSUM MP_TAC THEN ARITH_TAC; + ALL_TAC + ] THEN + SUBGOAL_THEN `1 + j - 1 = j /\ 1 + i - 1 = i` (fun th -> REWRITE_TAC[th]) THENL + [ + REPLICATE_TAC 3 (POP_ASSUM MP_TAC) THEN ARITH_TAC; + ALL_TAC + ] THEN + REWRITE_TAC[partial2] THEN + DISCH_THEN (MP_TAC o MATCH_MP interval_arith_abs_le) THEN + DISCH_THEN MATCH_MP_TAC THEN + new_rewrite [] [] EL_BUTLAST THENL + [ + ASM_REWRITE_TAC[] THEN REMOVE_THEN "j_in" MP_TAC THEN REMOVE_THEN "i_in" MP_TAC THEN + REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; + ALL_TAC + ] THEN + REWRITE_TAC[REAL_LE_REFL]);; + + +let M_TAYLOR_ERROR_ITLIST2_ALT = prove(`!f domain y w f_lo f_hi d_bounds_list dd_bounds_list error. + m_taylor_interval f domain y (vector w:real^N) (f_lo, f_hi) d_bounds_list dd_bounds_list ==> + LENGTH w = dimindex (:N) ==> + LENGTH dd_bounds_list = dimindex (:N) ==> + all_n 1 dd_bounds_list (\i list. LENGTH list = i) ==> + ITLIST2 (\list x z. x * (x * iabs (LAST list) + + &2 * ITLIST2 (\a b c. b * iabs a + c) (BUTLAST list) w (&0)) + z) + dd_bounds_list w (&0) <= error ==> + m_taylor_error f domain (vector w) error`, + REWRITE_TAC[m_taylor_interval] THEN REPEAT STRIP_TAC THEN + MP_TAC (SPEC_ALL M_TAYLOR_ERROR_ITLIST2) THEN ASM_REWRITE_TAC[]);; + + + +(****************************) + +let M_TAYLOR_INTERVAL' = MY_RULE m_taylor_interval;; + +let dest_m_taylor m_taylor_tm = + let ltm1, dd_bounds_list = dest_comb m_taylor_tm in + let ltm2, d_bounds_list = dest_comb ltm1 in + let ltm3, f_bounds = dest_comb ltm2 in + let ltm4, w = dest_comb ltm3 in + let ltm5, y = dest_comb ltm4 in + let ltm6, domain = dest_comb ltm5 in + rand ltm6, domain, y, w, f_bounds, d_bounds_list, dd_bounds_list;; + +let dest_m_taylor_thms n = + let ty, xty, x_var, f_var, y_var, w_var, domain_var = get_types_and_vars n in + fun m_taylor_th -> + let f, domain, y, w, f_bounds, d_bounds_list, dd_bounds_list = dest_m_taylor (concl m_taylor_th) in + let th0 = (INST[f, f_var; domain, domain_var; y, y_var; w, w_var; f_bounds, f_bounds_var; + d_bounds_list, d_bounds_list_var; dd_bounds_list, dd_bounds_list_var] o + inst_first_type_var ty) M_TAYLOR_INTERVAL' in + let th1 = EQ_MP th0 m_taylor_th in + match (CONJUNCTS th1) with + | [domain_th; d2_th; lin_th; second_th] -> + domain_th, d2_th, lin_th, second_th + | _ -> failwith "dest_m_taylor_thms";; + + +(**********************) + +(* bound *) +let M_TAYLOR_BOUND' = + prove(`m_taylor_interval f domain y (vector w:real^N) (f_lo, f_hi) d_bounds_list dd_bounds_list /\ + m_taylor_error f domain (vector w) error /\ + ITLIST2 (\a b c. b * iabs a + c) d_bounds_list w (&0) <= b /\ + b + inv(&2) * error <= a /\ + lo <= f_lo - a /\ + f_hi + a <= hi /\ + LENGTH w = dimindex (:N) /\ LENGTH d_bounds_list = dimindex (:N) ==> + (!x. x IN interval [domain] ==> interval_arith (f x) (lo, hi))`, + REWRITE_TAC[GSYM m_bounded_on_int; m_taylor_interval; m_lin_approx; ALL_N_EL] THEN + set_tac "s" `ITLIST2 _1 _2 _3 _4` THEN STRIP_TAC THEN + SUBGOAL_THEN `diff2_domain domain (f:real^N->real)` ASSUME_TAC THENL + [ + UNDISCH_TAC `diff2c_domain domain (f:real^N->real)` THEN + SIMP_TAC[diff2_domain; diff2c_domain; diff2c]; + ALL_TAC + ] THEN + apply_tac m_taylor_bounds THEN + MAP_EVERY EXISTS_TAC [`y:real^N`; `vector w:real^N`; `error:real`; `f_lo:real`; `f_hi:real`; `a:real`] THEN + ASM_REWRITE_TAC[] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b + inv (&2) * error` THEN ASM_REWRITE_TAC[] THEN + REWRITE_TAC[real_div; REAL_MUL_AC] THEN + MATCH_MP_TAC REAL_LE_ADD2 THEN REWRITE_TAC[REAL_LE_REFL] THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `s:real` THEN ASM_REWRITE_TAC[] THEN + EXPAND_TAC "s" THEN new_rewrite [] [] ITLIST2_EQ_SUM THEN ASM_REWRITE_TAC[LE_REFL] THEN + MATCH_MP_TAC SUM_LE THEN REWRITE_TAC[FINITE_NUMSEG] THEN REPEAT STRIP_TAC THEN + MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_SIMP_TAC[VECTOR_COMPONENT; REAL_LE_REFL; REAL_ABS_POS] THEN + CONJ_TAC THENL + [ + UNDISCH_TAC `m_cell_domain domain (y:real^N) (vector w)` THEN + new_rewrite [] [] pair_eq THEN REWRITE_TAC[m_cell_domain] THEN + DISCH_THEN (MP_TAC o SPEC `x:num`) THEN ASM_REWRITE_TAC[] THEN + ASM_SIMP_TAC[VECTOR_COMPONENT] THEN REAL_ARITH_TAC; + ALL_TAC + ] THEN + FIRST_X_ASSUM (MP_TAC o SPEC `x - 1`) THEN + ANTS_TAC THENL + [ + POP_ASSUM MP_TAC THEN ASM_REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; + ALL_TAC + ] THEN + DISCH_THEN (MP_TAC o MATCH_MP interval_arith_abs_le) THEN + SUBGOAL_THEN `1 + x - 1 = x` (fun th -> REWRITE_TAC[th]) THENL + [ + POP_ASSUM MP_TAC THEN REWRITE_TAC[IN_NUMSEG] THEN ARITH_TAC; + ALL_TAC + ] THEN + DISCH_THEN MATCH_MP_TAC THEN REWRITE_TAC[REAL_LE_REFL]);; + + +(* upper *) +let M_TAYLOR_UPPER_BOUND' = prove(`m_taylor_interval f domain y (vector w) (f_lo, f_hi) + d_bounds_list dd_bounds_list /\ + m_taylor_error f domain (vector w:real^N) error /\ + ITLIST2 (\a b c. b * iabs a + c) d_bounds_list w (&0) <= b /\ + b + inv(&2) * error <= a /\ + f_hi + a <= hi /\ + LENGTH w = dimindex (:N) /\ LENGTH d_bounds_list = dimindex (:N) ==> + (!p. p IN interval [domain] ==> f p <= hi)`, + STRIP_TAC THEN + MP_TAC (INST[`f_lo - a:real`, `lo:real`] M_TAYLOR_BOUND') THEN + ASM_SIMP_TAC[interval_arith; REAL_LE_REFL]);; + + +(* lower *) +let M_TAYLOR_LOWER_BOUND' = prove(`m_taylor_interval f domain y (vector w:real^N) (f_lo, f_hi) + d_bounds_list dd_bounds_list /\ + m_taylor_error f domain (vector w) error /\ + ITLIST2 (\a b c. b * iabs a + c) d_bounds_list w (&0) <= b /\ + b + inv(&2) * error <= a /\ + lo <= f_lo - a /\ + LENGTH w = dimindex (:N) /\ LENGTH d_bounds_list = dimindex (:N) ==> + (!p. p IN interval [domain] ==> lo <= f p)`, + STRIP_TAC THEN + MP_TAC (INST[`f_hi + a:real`, `hi:real`] M_TAYLOR_BOUND') THEN + ASM_SIMP_TAC[interval_arith; REAL_LE_REFL]);; + + + +(* arrays *) +let gen_taylor_bound_th bound_th n = + let th0 = (DISCH_ALL o MY_RULE o REWRITE_RULE[MY_RULE M_TAYLOR_ERROR_ITLIST2_ALT]) bound_th in + let ns = 1--n in + let mk_list_hd l = mk_list (l, type_of (hd l)) in + let w_list = mk_list_hd (map (fun i -> w_vars_array.(i)) ns) in + let d_bounds_list = mk_list_hd (map (fun i -> df_vars_array.(i)) ns) in + let dd_bounds_list = mk_list_hd (map (fun i -> mk_list_hd (map (fun j -> dd_vars_array.(i).(j)) (1--i))) ns) in + let th1 = (INST[w_list, w_var_real_list; d_bounds_list, d_bounds_list_var; + dd_bounds_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) th0 in + let th2 = REWRITE_RULE[LAST; NOT_CONS_NIL; BUTLAST; all_n; ITLIST2_DEF; LENGTH; ARITH; dimindex_array.(n)] th1 in + let th3 = REWRITE_RULE[HD; TL; REAL_MUL_RZERO; REAL_ADD_RID; GSYM error_mul_f2] th2 in + (MY_RULE o REWRITE_RULE[float_inv2_th; SYM float2_eq]) th3;; + + +let m_taylor_upper_array = Array.init (max_dim + 1) + (fun i -> if i < 1 then TRUTH else gen_taylor_bound_th M_TAYLOR_UPPER_BOUND' i);; + +let m_taylor_lower_array = Array.init (max_dim + 1) + (fun i -> if i < 1 then TRUTH else gen_taylor_bound_th M_TAYLOR_LOWER_BOUND' i);; + +let m_taylor_bound_array = Array.init (max_dim + 1) + (fun i -> if i < 1 then TRUTH else gen_taylor_bound_th M_TAYLOR_BOUND' i);; + + +(***************************) +(* eval_m_taylor_bounds0 *) + +let eval_m_taylor_bounds0 mode n pp m_taylor_th = + let bound_th = + if mode = "upper" then + m_taylor_upper_array.(n) + else if mode = "bound" then + m_taylor_bound_array.(n) + else + m_taylor_lower_array.(n) in + + let f_tm, domain_tm, y_tm, w_tm, f_bounds, d_bounds_list, dd_bounds_list = dest_m_taylor (concl m_taylor_th) in + let f_lo, f_hi = dest_pair f_bounds and + ws = dest_list (rand w_tm) and + dfs = dest_list d_bounds_list and + dds = map dest_list (dest_list dd_bounds_list) in + let ns = 1--n in + let df_vars = map (fun i -> df_vars_array.(i)) ns and + dd_vars = map (fun i -> map (fun j -> dd_vars_array.(i).(j)) (1--i)) ns and + w_vars = map (fun i -> w_vars_array.(i)) ns and + y_var = mk_var ("y", type_of y_tm) and + f_var = mk_var ("f", type_of f_tm) and + domain_var = mk_var ("domain", type_of domain_tm) in + + (* sum of first partials *) + let d_th = + let mul_wd = map2 (error_mul_f2_le_conv2 pp) ws dfs in + end_itlist (add_ineq_hi pp) mul_wd in + + let b_tm = (rand o concl) d_th in + + (* sum of second partials *) + let dd_th = + let ( * ), ( + ) = mul_ineq_pos_const_hi pp, add_ineq_hi pp in + let mul_wdd = map2 (fun list i -> my_map2 (error_mul_f2_le_conv2 pp) ws list) dds ns in + let sums1 = map (end_itlist ( + ) o butlast) (tl mul_wdd) in + let sums2 = (hd o hd) mul_wdd :: map2 (fun list th1 -> last list + two_float * th1) (tl mul_wdd) sums1 in + let sums = map2 ( * ) ws sums2 in + end_itlist ( + ) sums in + + let error_tm = (rand o concl) dd_th in + + (* additional inequalities *) + let ineq1_th = + let ( * ), ( + ) = float_mul_hi pp, add_ineq_hi pp in + mk_refl_ineq b_tm + float_inv2 * error_tm in + + let a_tm = (rand o concl) ineq1_th in + + let prove_ineq2, bounds_inst = + if mode = "upper" then + let ineq2 = float_add_hi pp f_hi a_tm in + MY_PROVE_HYP ineq2, [(rand o concl) ineq2, hi_var_real] + else if mode = "bound" then + let ineq2 = float_add_hi pp f_hi a_tm in + let ineq3 = float_sub_lo pp f_lo a_tm in + MY_PROVE_HYP ineq2 o MY_PROVE_HYP ineq3, + [(rand o concl) ineq2, hi_var_real; (lhand o concl) ineq3, lo_var_real] + else + let ineq2 = float_sub_lo pp f_lo a_tm in + MY_PROVE_HYP ineq2, [(lhand o concl) ineq2, lo_var_real] in + + (* final step *) + let inst_list = + let inst1 = zip dfs df_vars in + let inst2 = zip (List.flatten dds) (List.flatten dd_vars) in + let inst3 = zip ws w_vars in + inst1 @ inst2 @ inst3 in + + (MY_PROVE_HYP m_taylor_th o MY_PROVE_HYP dd_th o MY_PROVE_HYP d_th o + MY_PROVE_HYP ineq1_th o prove_ineq2 o + INST ([f_hi, f_hi_var; f_lo, f_lo_var; error_tm, error_var; + a_tm, a_var_real; b_tm, b_var_real; + y_tm, y_var; domain_tm, domain_var; + f_tm, f_var] @ bounds_inst @ inst_list)) bound_th;; + + +(* upper *) +let eval_m_taylor_upper_bound = eval_m_taylor_bounds0 "upper";; + +(* lower *) +let eval_m_taylor_lower_bound = eval_m_taylor_bounds0 "lower";; + +(* bound *) +let eval_m_taylor_bound = eval_m_taylor_bounds0 "bound";; + + + + +(******************************) +(* taylor_upper_partial_bound *) +(* taylor_lower_partial_bound *) + + +(* bound *) +let M_TAYLOR_PARTIAL_BOUND' = + prove(`m_taylor_interval f domain (y:real^N) (vector w) f_bounds d_bounds_list dd_bounds_list /\ + i IN 1..dimindex (:N) /\ + EL (i - 1) d_bounds_list = (df_lo, df_hi) /\ + (!x. x IN interval [domain] ==> all_n 1 dd_list + (\j int. interval_arith (if j <= i then partial2 j i f x else partial2 i j f x) int)) /\ + LENGTH dd_list = dimindex (:N) /\ + LENGTH d_bounds_list = dimindex (:N) /\ + LENGTH w = dimindex (:N) /\ + ITLIST2 (\a b c. b * iabs a + c) dd_list w (&0) <= error /\ + df_hi + error <= hi ==> + lo <= df_lo - error ==> + (!x. x IN interval [domain] ==> interval_arith (partial i f x) (lo, hi))`, + REWRITE_TAC[m_taylor_interval; m_lin_approx; ALL_N_EL; GSYM m_bounded_on_int] THEN + set_tac "s" `ITLIST2 _1 _2 _3 _4` THEN REPEAT STRIP_TAC THEN + SUBGOAL_THEN `1 <= i /\ i <= dimindex (:N)` (LABEL_TAC "i_in") THENL + [ + ASM_REWRITE_TAC[GSYM IN_NUMSEG]; + ALL_TAC + ] THEN + SUBGOAL_THEN `diff2_domain domain (f:real^N->real)` ASSUME_TAC THENL + [ + UNDISCH_TAC `diff2c_domain domain (f:real^N->real)` THEN + SIMP_TAC[diff2_domain; diff2c_domain; diff2c]; + ALL_TAC + ] THEN + REWRITE_TAC[ETA_AX] THEN apply_tac m_taylor_partial_bounds THEN + MAP_EVERY EXISTS_TAC [`y:real^N`; `vector w:real^N`; `error:real`; `df_lo:real`; `df_hi:real`] THEN + ASM_REWRITE_TAC[] THEN + FIRST_X_ASSUM (MP_TAC o SPEC `i - 1`) THEN + ASM_SIMP_TAC[ARITH_RULE `1 <= i /\ i <= n ==> i - 1 < n`] THEN + ASM_SIMP_TAC[ARITH_RULE `1 <= i ==> 1 + i - 1 = i`; interval_arith] THEN + DISCH_THEN (fun th -> ALL_TAC) THEN + REWRITE_TAC[m_taylor_partial_error] THEN REPEAT STRIP_TAC THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `s:real` THEN ASM_REWRITE_TAC[] THEN + EXPAND_TAC "s" THEN new_rewrite [] [] ITLIST2_EQ_SUM THEN ASM_REWRITE_TAC[LE_REFL] THEN + MATCH_MP_TAC SUM_LE THEN REWRITE_TAC[FINITE_NUMSEG] THEN + move ["j"; "j_in"] THEN + ASM_SIMP_TAC[VECTOR_COMPONENT] THEN + MATCH_MP_TAC REAL_LE_LMUL THEN CONJ_TAC THENL + [ + UNDISCH_TAC `m_cell_domain domain (y:real^N) (vector w)` THEN + new_rewrite [] [] pair_eq THEN REWRITE_TAC[m_cell_domain] THEN + DISCH_THEN (MP_TAC o SPEC `j:num`) THEN ASM_REWRITE_TAC[] THEN + ASM_SIMP_TAC[VECTOR_COMPONENT] THEN REAL_ARITH_TAC; + ALL_TAC + ] THEN + FIRST_X_ASSUM (MP_TAC o SPEC `x:real^N`) THEN ASM_REWRITE_TAC[] THEN + DISCH_THEN (MP_TAC o SPEC `j - 1`) THEN + POP_ASSUM MP_TAC THEN REWRITE_TAC[IN_NUMSEG] THEN DISCH_TAC THEN + ASM_SIMP_TAC[ARITH_RULE `1 <= j /\ j <= n ==> j - 1 < n`] THEN + ASM_SIMP_TAC[ARITH_RULE `!i. 1 <= i ==> 1 + i - 1 = i`; GSYM partial2] THEN + DISCH_THEN (MP_TAC o MATCH_MP interval_arith_abs_le) THEN + COND_CASES_TAC THEN TRY (DISCH_THEN MATCH_MP_TAC THEN REWRITE_TAC[REAL_LE_REFL]) THEN + new_rewrite [] [] mixed_second_partials THENL + [ + UNDISCH_TAC `diff2c_domain domain (f:real^N->real)` THEN + ASM_SIMP_TAC[diff2c_domain]; + ALL_TAC + ] THEN + DISCH_THEN MATCH_MP_TAC THEN REWRITE_TAC[REAL_LE_REFL]);; + + +(* upper *) +let M_TAYLOR_PARTIAL_UPPER' = + prove(`m_taylor_interval f domain (y:real^N) (vector w) f_bounds d_bounds_list dd_bounds_list /\ + i IN 1..dimindex (:N) /\ + EL (i - 1) d_bounds_list = (df_lo, df_hi) /\ + (!x. x IN interval [domain] ==> all_n 1 dd_list + (\j int. interval_arith (if j <= i then partial2 j i f x else partial2 i j f x) int)) /\ + LENGTH dd_list = dimindex (:N) /\ + LENGTH d_bounds_list = dimindex (:N) /\ + LENGTH w = dimindex (:N) /\ + ITLIST2 (\a b c. b * iabs a + c) dd_list w (&0) <= error /\ + df_hi + error <= hi ==> + (!x. x IN interval [domain] ==> partial i f x <= hi)`, + REPEAT STRIP_TAC THEN + MP_TAC (INST[`df_lo - error`, `lo:real`] M_TAYLOR_PARTIAL_BOUND') THEN ASM_REWRITE_TAC[REAL_LE_REFL] THEN + ASM_SIMP_TAC[interval_arith]);; + + +(* lower *) +let M_TAYLOR_PARTIAL_LOWER' = + prove(`m_taylor_interval f domain (y:real^N) (vector w) f_bounds d_bounds_list dd_bounds_list /\ + i IN 1..dimindex (:N) /\ + EL (i - 1) d_bounds_list = (df_lo, df_hi) /\ + (!x. x IN interval [domain] ==> all_n 1 dd_list + (\j int. interval_arith (if j <= i then partial2 j i f x else partial2 i j f x) int)) /\ + LENGTH dd_list = dimindex (:N) /\ + LENGTH d_bounds_list = dimindex (:N) /\ + LENGTH w = dimindex (:N) /\ + ITLIST2 (\a b c. b * iabs a + c) dd_list w (&0) <= error /\ + lo <= df_lo - error ==> + (!x. x IN interval [domain] ==> lo <= partial i f x)`, + REPEAT STRIP_TAC THEN + MP_TAC (INST[`df_hi + error`, `hi:real`] M_TAYLOR_PARTIAL_BOUND') THEN ASM_REWRITE_TAC[REAL_LE_REFL] THEN + ASM_SIMP_TAC[interval_arith]);; + + + + +(* arrays *) +let gen_taylor_partial_bound_th = + let imp_and_eq = TAUT `((P ==> Q) /\ (P ==> R)) <=> (P ==> Q /\ R)` in + let mk_list_hd l = mk_list (l, type_of (hd l)) in + let dd_list_var = `dd_list : (real#real)list` in + fun bound_th n i -> + let ns = 1--n in + let i_tm = mk_small_numeral i in + let w_list = mk_list_hd (map (fun i -> w_vars_array.(i)) ns) in + let d_bounds_list = mk_list_hd (map (fun i -> df_vars_array.(i)) ns) in + let dd_bounds_list = mk_list_hd + (map (fun i -> mk_list_hd (map (fun j -> dd_vars_array.(i).(j)) (1--i))) ns) in + let dd_list = mk_list_hd + (map (fun j -> if j <= i then dd_vars_array.(i).(j) else dd_vars_array.(j).(i)) ns) in + + let th1 = (INST[w_list, w_var_real_list; d_bounds_list, d_bounds_list_var; dd_list, dd_list_var; i_tm, `i:num`; + dd_bounds_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) bound_th in + let th2 = REWRITE_RULE[REAL_ADD_RID; HD; TL; ITLIST2_DEF; LENGTH; ARITH; dimindex_array.(n)] th1 in + let th3 = REWRITE_RULE[IN_NUMSEG; ARITH; el_thms_array.(n).(i - 1)] th2 in + let th4 = (REWRITE_RULE[] o INST[`df_lo:real, df_hi:real`, df_vars_array.(i)]) th3 in + let th5 = (MY_RULE o REWRITE_RULE[GSYM imp_and_eq; GSYM AND_FORALL_THM; all_n; ARITH]) th4 in + + let m_taylor_hyp = find (can dest_m_taylor) (hyp th5) in + let t_th0 = (REWRITE_RULE[ARITH; all_n; second_bounded; m_taylor_interval] o ASSUME) m_taylor_hyp in + let t_th1 = REWRITE_RULE[GSYM imp_and_eq; GSYM AND_FORALL_THM] t_th0 in + (MY_RULE_NUM o REWRITE_RULE[GSYM error_mul_f2; t_th1] o DISCH_ALL) th5;; + + +(* The (n, i)-th element is the theorem |- i IN 1..dimindex (:n) *) +let i_in_array = Array.init (max_dim + 1) + (fun i -> Array.init (i + 1) + (fun j -> + if j < 1 then TRUTH else + let j_tm = mk_small_numeral j in + let tm0 = `j IN 1..dimindex (:N)` in + let tm1 = (subst [j_tm, `j:num`] o inst [n_type_array.(i), nty]) tm0 in + prove(tm1, REWRITE_TAC[dimindex_array.(i); IN_NUMSEG] THEN ARITH_TAC)));; + +let m_taylor_partial_upper_array, m_taylor_partial_lower_array, m_taylor_partial_bound_array = + let gen_array bound_th = Array.init (max_dim + 1) + (fun i -> Array.init (i + 1) + (fun j -> if j < 1 then TRUTH else gen_taylor_partial_bound_th bound_th i j)) in + gen_array M_TAYLOR_PARTIAL_UPPER', gen_array M_TAYLOR_PARTIAL_LOWER', gen_array M_TAYLOR_PARTIAL_BOUND';; + + +(***************************) + + +let eval_m_taylor_partial_bounds0 mode n pp i m_taylor_th = + let bound_th = + if mode = "upper" then + m_taylor_partial_upper_array.(n).(i) + else if mode = "bound" then + m_taylor_partial_bound_array.(n).(i) + else + m_taylor_partial_lower_array.(n).(i) in + + let f_tm, domain_tm, y_tm, w_tm, f_bounds, d_bounds_list, dd_bounds_list = dest_m_taylor (concl m_taylor_th) in + let ws = dest_list (rand w_tm) and + dfs = dest_list d_bounds_list and + dds = map dest_list (dest_list dd_bounds_list) in + let ns = 1--n in + let df_lo, df_hi = dest_pair (List.nth dfs (i - 1)) and + df_vars = map (fun i -> df_vars_array.(i)) ns and + dd_vars = map (fun i -> map (fun j -> dd_vars_array.(i).(j)) (1--i)) ns and + w_vars = map (fun i -> w_vars_array.(i)) ns and + y_var = mk_var ("y", type_of y_tm) and + f_var = mk_var ("f", type_of f_tm) and + domain_var = mk_var ("domain", type_of domain_tm) in + + (* sum of second partials *) + let dd_list = map + (fun j -> if j <= i then List.nth (List.nth dds (i-1)) (j-1) else List.nth (List.nth dds (j-1)) (i-1)) ns in + + let dd_th = + let mul_dd = map2 (error_mul_f2_le_conv2 pp) ws dd_list in + end_itlist (add_ineq_hi pp) mul_dd in + + let error_tm = (rand o concl) dd_th in + + (* additional inequalities *) + let prove_ineq, bounds_inst = + if mode = "upper" then + let ineq2 = float_add_hi pp df_hi error_tm in + MY_PROVE_HYP ineq2, [(rand o concl) ineq2, hi_var_real] + else if mode = "bound" then + let ineq2 = float_add_hi pp df_hi error_tm in + let ineq3 = float_sub_lo pp df_lo error_tm in + MY_PROVE_HYP ineq2 o MY_PROVE_HYP ineq3, + [(rand o concl) ineq2, hi_var_real; (lhand o concl) ineq3, lo_var_real] + else + let ineq2 = float_sub_lo pp df_lo error_tm in + MY_PROVE_HYP ineq2, [(lhand o concl) ineq2, lo_var_real] in + + (* final step *) + let inst_list = + let inst1 = zip dfs df_vars in + let inst2 = zip (List.flatten dds) (List.flatten dd_vars) in + let inst3 = zip ws w_vars in + inst1 @ inst2 @ inst3 in + + (MY_PROVE_HYP m_taylor_th o MY_PROVE_HYP dd_th o prove_ineq o + INST ([df_hi, df_hi_var; df_lo, df_lo_var; error_tm, error_var; + y_tm, y_var; domain_tm, domain_var; f_bounds, f_bounds_var; + f_tm, f_var] @ bounds_inst @ inst_list)) bound_th;; + + +(* upper *) +let eval_m_taylor_partial_upper = eval_m_taylor_partial_bounds0 "upper";; + +(* lower *) +let eval_m_taylor_partial_lower = eval_m_taylor_partial_bounds0 "lower";; + +(* bound *) +let eval_m_taylor_partial_bound = eval_m_taylor_partial_bounds0 "bound";; + + + +end;; diff --git a/Formal_ineqs/taylor/m_taylor_arith.hl b/Formal_ineqs/taylor/m_taylor_arith.hl index 16fb36ba..ad1647a1 100644 --- a/Formal_ineqs/taylor/m_taylor_arith.hl +++ b/Formal_ineqs/taylor/m_taylor_arith.hl @@ -1,11 +1,17 @@ -(* =========================================================== *) -(* Formal arithmetic of taylor intervals *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) -needs "taylor/m_taylor.hl";; -needs "misc/misc_vars.hl";; +(* -------------------------------------------------------------------------- *) +(* Formal arithmetic of taylor intervals *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/taylor/m_taylor.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; module M_taylor_arith = struct @@ -92,7 +98,7 @@ let m_lin_approx_components n m_lin_th = let ty = n_type_array.(n) in let f_var = mk_var ("f", type_of f_tm) in let x_var = mk_var ("x", type_of x_tm) in - let th0 = (INST[f_tm, f_var; x_tm, x_var; f_bounds, f_bounds_var; + let th0 = (INST[f_tm, f_var; x_tm, x_var; f_bounds, f_bounds_var; d_bounds_list, df_bounds_list_var] o inst_first_type_var ty) DEST_M_LIN_APPROX' in let th1 = EQ_MP th0 m_lin_th in match (CONJUNCTS th1) with @@ -104,9 +110,9 @@ let m_lin_approx_components n m_lin_th = (* all_n manipulations *) let ALL_N_EMPTY' = prove(`all_n n [] (s:num->A->bool)`, REWRITE_TAC[all_n]);; -let ALL_N_CONS_IMP' = (MY_RULE o prove)(`SUC n = m /\ s n (x:A) ==> +let ALL_N_CONS_IMP' = (MY_RULE o prove)(`SUC n = m /\ s n (x:A) ==> (all_n m t s <=> all_n n (CONS x t) s)`, SIMP_TAC[all_n]);; -let ALL_N_CONS_EQ' = (MY_RULE o prove)(`SUC n = m ==> +let ALL_N_CONS_EQ' = (MY_RULE o prove)(`SUC n = m ==> (all_n n (CONS x t) s <=> (s n (x:A) /\ all_n m t s))`, SIMP_TAC[all_n]);; let dest_all_n all_n_tm = @@ -217,7 +223,7 @@ let eval_all_n all_n1_th beta_flag s = let eval_all_n2 all_n1_th all_n2_th beta_flag s = let ths1', suc_ths = all_n_components all_n1_th in let ths2', _ = all_n_components all_n2_th in - let ths1, ths2 = + let ths1, ths2 = if beta_flag then map MY_BETA_RULE ths1', map MY_BETA_RULE ths2' else ths1', ths2' in let ths1, ths2, suc_ths = List.rev ths1, List.rev ths2, List.rev suc_ths in @@ -255,10 +261,10 @@ let MK_M_TAYLOR_ADD' = (MY_RULE_NUM o prove) diff2c_domain domain g ==> interval_arith (f y + g y) bounds ==> all_n 1 d_bounds_list (\i int. interval_arith (partial i f y + partial i g y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith (partial2 j i f x + partial2 j i g x) int))) ==> m_taylor_interval (\x. f x + g x) domain y w bounds d_bounds_list dd_bounds_list`, - REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded] THEN + REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded] THEN REPEAT DISCH_TAC THEN SUBGOAL_THEN `lift o f differentiable at (y:real^N) /\ lift o g differentiable at y` ASSUME_TAC THENL [ @@ -270,7 +276,7 @@ let MK_M_TAYLOR_ADD' = (MY_RULE_NUM o prove) MATCH_MP_TAC y_in_domain THEN EXISTS_TAC `w:real^N` THEN ASM_REWRITE_TAC[]; ALL_TAC ] THEN - + REPEAT STRIP_TAC THEN ASM_REWRITE_TAC[] THENL [ new_rewrite [] [] diff2c_domain_add THEN ASM_REWRITE_TAC[]; @@ -299,9 +305,9 @@ let add_second_lemma' = prove(`interval_arith (partial2 j i f (x:real^N) + parti (\j int. interval_arith (partial2 j i f x + partial2 j i g x) int) j int`, REWRITE_TAC[]);; -let add_second_lemma'' = (NUMERALS_TO_NUM o prove)(`all_n 1 list +let add_second_lemma'' = (NUMERALS_TO_NUM o prove)(`all_n 1 list (\j int. interval_arith (partial2 j i f (x:real^N) + partial2 j i g x) int) <=> - (\i list. all_n 1 list + (\i list. all_n 1 list (\j int. interval_arith (partial2 j i f x + partial2 j i g x) int)) i list`, REWRITE_TAC[]);; @@ -329,7 +335,7 @@ let eval_m_taylor_add n p_lin p_second taylor1_th taylor2_th = let bounds_th = float_interval_add p_lin bounds1_th bounds2_th in let bounds_tm = (rand o concl) bounds_th in - let add_lemma0 = (INST[f1_tm, f_var; f2_tm, g_var; y_tm, x_var] o + let add_lemma0 = (INST[f1_tm, f_var; f2_tm, g_var; y_tm, x_var] o INST_TYPE[n_type_array.(n), nty]) add_partial_lemma' in let add th1 th2 = @@ -346,10 +352,10 @@ let eval_m_taylor_add n p_lin p_second taylor1_th taylor2_th = let dd1 = second_bounded_components n second1_th in let dd2 = second_bounded_components n second2_th in - let add_second_lemma0 = (INST[f1_tm, f_var; f2_tm, g_var] o + let add_second_lemma0 = (INST[f1_tm, f_var; f2_tm, g_var] o INST_TYPE[n_type_array.(n), nty]) add_second_lemma' in - let add_second_lemma1 = (INST[f1_tm, f_var; f2_tm, g_var] o + let add_second_lemma1 = (INST[f1_tm, f_var; f2_tm, g_var] o INST_TYPE[n_type_array.(n), nty]) add_second_lemma'' in @@ -373,11 +379,11 @@ let eval_m_taylor_add n p_lin p_second taylor1_th taylor2_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP diff2_f2_th o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; f2_tm, g_var; + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP diff2_f2_th o + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; f2_tm, g_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_ADD' in let eq_th = binary_beta_gen_eq f1_tm f2_tm x_var add_op_real in @@ -395,10 +401,10 @@ let MK_M_TAYLOR_SUB' = (MY_RULE_NUM o prove) diff2c_domain domain g ==> interval_arith (f y - g y) bounds ==> all_n 1 d_bounds_list (\i int. interval_arith (partial i f y - partial i g y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith (partial2 j i f x - partial2 j i g x) int))) ==> m_taylor_interval (\x. f x - g x) domain y w bounds d_bounds_list dd_bounds_list`, - REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded] THEN + REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded] THEN REPEAT DISCH_TAC THEN SUBGOAL_THEN `lift o f differentiable at (y:real^N) /\ lift o g differentiable at y` ASSUME_TAC THENL [ @@ -410,7 +416,7 @@ let MK_M_TAYLOR_SUB' = (MY_RULE_NUM o prove) MATCH_MP_TAC y_in_domain THEN EXISTS_TAC `w:real^N` THEN ASM_REWRITE_TAC[]; ALL_TAC ] THEN - + REPEAT STRIP_TAC THEN ASM_REWRITE_TAC[] THENL [ new_rewrite [] [] diff2c_domain_sub THEN ASM_REWRITE_TAC[]; @@ -439,9 +445,9 @@ let sub_second_lemma' = prove(`interval_arith (partial2 j i f (x:real^N) - parti (\j int. interval_arith (partial2 j i f x - partial2 j i g x) int) j int`, REWRITE_TAC[]);; -let sub_second_lemma'' = (NUMERALS_TO_NUM o prove)(`all_n 1 list +let sub_second_lemma'' = (NUMERALS_TO_NUM o prove)(`all_n 1 list (\j int. interval_arith (partial2 j i f (x:real^N) - partial2 j i g x) int) <=> - (\i list. all_n 1 list + (\i list. all_n 1 list (\j int. interval_arith (partial2 j i f x - partial2 j i g x) int)) i list`, REWRITE_TAC[]);; @@ -469,7 +475,7 @@ let eval_m_taylor_sub n p_lin p_second taylor1_th taylor2_th = let bounds_th = float_interval_sub p_lin bounds1_th bounds2_th in let bounds_tm = (rand o concl) bounds_th in - let sub_lemma0 = (INST[f1_tm, f_var; f2_tm, g_var; y_tm, x_var] o + let sub_lemma0 = (INST[f1_tm, f_var; f2_tm, g_var; y_tm, x_var] o INST_TYPE[n_type_array.(n), nty]) sub_partial_lemma' in let sub th1 th2 = @@ -486,10 +492,10 @@ let eval_m_taylor_sub n p_lin p_second taylor1_th taylor2_th = let dd1 = second_bounded_components n second1_th in let dd2 = second_bounded_components n second2_th in - let sub_second_lemma0 = (INST[f1_tm, f_var; f2_tm, g_var] o + let sub_second_lemma0 = (INST[f1_tm, f_var; f2_tm, g_var] o INST_TYPE[n_type_array.(n), nty]) sub_second_lemma' in - let sub_second_lemma1 = (INST[f1_tm, f_var; f2_tm, g_var] o + let sub_second_lemma1 = (INST[f1_tm, f_var; f2_tm, g_var] o INST_TYPE[n_type_array.(n), nty]) sub_second_lemma'' in @@ -513,11 +519,11 @@ let eval_m_taylor_sub n p_lin p_second taylor1_th taylor2_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP diff2_f2_th o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; f2_tm, g_var; + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP diff2_f2_th o + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; f2_tm, g_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_SUB' in let eq_th = binary_beta_gen_eq f1_tm f2_tm x_var sub_op_real in @@ -537,11 +543,11 @@ let MK_M_TAYLOR_MUL' = (MY_RULE_NUM o prove) diff2c_domain domain g ==> interval_arith (f y * g y) bounds ==> all_n 1 d_bounds_list (\i int. interval_arith (partial i f y * g y + f y * partial i g y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith ((partial2 j i f x * g x + partial i f x * partial j g x) + partial j f x * partial i g x + f x * partial2 j i g x) int))) ==> m_taylor_interval (\x. f x * g x) domain y w bounds d_bounds_list dd_bounds_list`, - REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded] THEN + REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded] THEN REPEAT DISCH_TAC THEN SUBGOAL_THEN `lift o f differentiable at (y:real^N) /\ lift o g differentiable at y` ASSUME_TAC THENL [ @@ -553,7 +559,7 @@ let MK_M_TAYLOR_MUL' = (MY_RULE_NUM o prove) MATCH_MP_TAC y_in_domain THEN EXISTS_TAC `w:real^N` THEN ASM_REWRITE_TAC[]; ALL_TAC ] THEN - + REPEAT STRIP_TAC THEN ASM_REWRITE_TAC[] THENL [ new_rewrite [] [] diff2c_domain_mul THEN ASM_REWRITE_TAC[]; @@ -572,12 +578,12 @@ let MK_M_TAYLOR_MUL' = (MY_RULE_NUM o prove) (*************************) -let mul_partial_lemma' = +let mul_partial_lemma' = prove(`interval_arith (partial i f (y:real^N) * g y + f y * partial i g y) int <=> (\i int. interval_arith (partial i f y * g y + f y * partial i g y) int) i int`, REWRITE_TAC[]);; -let mul_second_lemma' = +let mul_second_lemma' = prove(`interval_arith ((partial2 j i f x * g (x:real^N) + partial i f x * partial j g x) + partial j f x * partial i g x + f x * partial2 j i g x) int <=> (\j int. interval_arith ((partial2 j i f x * g x + partial i f x * partial j g x) + @@ -588,7 +594,7 @@ let mul_second_lemma' = let mul_second_lemma'' = (NUMERALS_TO_NUM o prove) (`all_n 1 list (\j int. interval_arith ((partial2 j i f x * g x + partial i f x * partial j g x) + partial j f x * partial i g x + f (x:real^N) * partial2 j i g x) int) <=> - (\i list. all_n 1 list + (\i list. all_n 1 list (\j int. interval_arith ((partial2 j i f x * g x + partial i f x * partial j g x) + partial j f x * partial i g x + f x * partial2 j i g x) int)) i list`, REWRITE_TAC[]);; @@ -617,7 +623,7 @@ let eval_m_taylor_mul n p_lin p_second taylor1_th taylor2_th = let bounds_th = float_interval_mul p_lin bounds1_th bounds2_th in let bounds_tm = (rand o concl) bounds_th in - let mul_lemma0 = (INST[f1_tm, f_var; f2_tm, g_var; y_tm, y_var] o + let mul_lemma0 = (INST[f1_tm, f_var; f2_tm, g_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) mul_partial_lemma' in let mul th1 th2 = @@ -635,20 +641,20 @@ let eval_m_taylor_mul n p_lin p_second taylor1_th taylor2_th = let dd1 = second_bounded_components n second1_th in let dd2 = second_bounded_components n second2_th in - - let mul_second_lemma0 = (INST[f1_tm, f_var; f2_tm, g_var] o + + let mul_second_lemma0 = (INST[f1_tm, f_var; f2_tm, g_var] o INST_TYPE[n_type_array.(n), nty]) mul_second_lemma' in - let mul_second_lemma1 = (INST[f1_tm, f_var; f2_tm, g_var] o + let mul_second_lemma1 = (INST[f1_tm, f_var; f2_tm, g_var] o INST_TYPE[n_type_array.(n), nty]) mul_second_lemma'' in let undisch = UNDISCH o SPEC x_var in - let d1_bounds = map (fun i -> + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in - let d2_bounds = map (fun i -> + let d2_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor2_th in undisch th0) (1--n) in @@ -668,14 +674,14 @@ let eval_m_taylor_mul n p_lin p_second taylor1_th taylor2_th = let dj1 = List.nth d1_bounds (j_int - 1) and dj2 = List.nth d2_bounds (j_int - 1) in - let mul_th = + let mul_th = let ( * ), ( + ) = float_interval_mul p_second, float_interval_add p_second in (th1 * f2_bound + di1 * dj2) + (dj1 * di2 + f1_bound * th2) in let int_tm = rand (concl mul_th) in let th0 = INST[j_tm, j_var_num; int_tm, int_var] lemma in EQ_MP th0 mul_th in - + let mul_th = eval_all_n2 th1 th2 true mul_second in let list_tm = (rand o rator o concl) mul_th in let lemma1 = INST[i_tm, i_var_num; list_tm, list_var_real_pair] mul_second_lemma1 in @@ -686,11 +692,11 @@ let eval_m_taylor_mul n p_lin p_second taylor1_th taylor2_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP diff2_f2_th o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; f2_tm, g_var; + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP diff2_f2_th o + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; f2_tm, g_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_MUL' in let eq_th = binary_beta_gen_eq f1_tm f2_tm x_var mul_op_real in @@ -707,26 +713,26 @@ let eval_m_taylor_mul n p_lin p_second taylor1_th taylor2_th = -let partial_uni_compose' = +let partial_uni_compose' = REWRITE_RULE[SWAP_FORALL_THM; GSYM RIGHT_IMP_FORALL_THM] partial_uni_compose;; -let second_partial_uni_compose' = +let second_partial_uni_compose' = REWRITE_RULE[SWAP_FORALL_THM; GSYM RIGHT_IMP_FORALL_THM] second_partial_uni_compose;; -let all_n1_raw = (GEN_REWRITE_RULE (RAND_CONV o DEPTH_CONV) [SYM num1_eq] o REFL) +let all_n1_raw = (GEN_REWRITE_RULE (RAND_CONV o DEPTH_CONV) [SYM num1_eq] o REFL) `all_n 1`;; - + (* neg *) let MK_M_TAYLOR_NEG' = (MY_RULE_FLOAT o prove) (`m_cell_domain domain (y:real^N) w ==> diff2c_domain domain f ==> interval_arith (-- (f y)) bounds ==> all_n 1 d_bounds_list (\i int. interval_arith (-- partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith (-- partial2 j i f x) int))) ==> m_taylor_interval (\x. -- (f x)) domain y w bounds d_bounds_list dd_bounds_list`, - REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded] THEN + REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded] THEN REPEAT DISCH_TAC THEN SUBGOAL_THEN `lift o f differentiable at (y:real^N)` ASSUME_TAC THENL [ @@ -737,7 +743,7 @@ let MK_M_TAYLOR_NEG' = (MY_RULE_FLOAT o prove) MATCH_MP_TAC y_in_domain THEN EXISTS_TAC `w:real^N` THEN ASM_REWRITE_TAC[]; ALL_TAC ] THEN - + REPEAT STRIP_TAC THEN ASM_REWRITE_TAC[] THENL [ new_rewrite [] [] diff2c_domain_neg THEN ASM_REWRITE_TAC[]; @@ -747,7 +753,7 @@ let MK_M_TAYLOR_NEG' = (MY_RULE_FLOAT o prove) ASM_SIMP_TAC[partial_neg]; ALL_TAC ] THEN - + UNDISCH_TAC `diff2c_domain domain (f:real^N->real)` THEN REWRITE_TAC[diff2c_domain_alt; diff2_domain] THEN DISCH_THEN (MP_TAC o SPEC `x:real^N` o CONJUNCT1) THEN ASM_REWRITE_TAC[] THEN DISCH_TAC THEN @@ -865,7 +871,7 @@ let MK_M_TAYLOR_ABS_NEG' = (MY_RULE_FLOAT o prove) ASM_REWRITE_TAC[ETA_AX]; ALL_TAC ] THEN - + REPEAT STRIP_TAC THEN ASM_REWRITE_TAC[] THENL [ UNDISCH_TAC `diff2c_domain domain (f:real^N->real)` THEN @@ -922,8 +928,8 @@ let MK_M_TAYLOR_POW = prove diff2c_domain domain f ==> interval_arith ((f y) pow k) bounds ==> all_n 1 d_bounds_list (\i int. interval_arith ((&k * (f y) pow (k - 1)) * partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i - (\j int. interval_arith (((&(k * (k - 1)) * f x pow (k - 2)) * partial j f x) + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (\j int. interval_arith (((&(k * (k - 1)) * f x pow (k - 2)) * partial j f x) * partial i f x + (&k * f x pow (k - 1)) * partial2 j i f x) int))) ==> m_taylor_interval (\x. (f x) pow k) domain y w bounds d_bounds_list dd_bounds_list`, REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded; ETA_AX] THEN @@ -942,7 +948,7 @@ let MK_M_TAYLOR_POW = prove REWRITE_TAC[FUN_EQ_THM; o_THM]; ALL_TAC ] THEN - + REPEAT STRIP_TAC THEN ASM_REWRITE_TAC[] THENL [ UNDISCH_TAC `diff2c_domain domain (f:real^N->real)` THEN @@ -951,7 +957,7 @@ let MK_M_TAYLOR_POW = prove apply_tac diff_uni_compose THEN ASM_REWRITE_TAC[REAL_DIFFERENTIABLE_AT_POW]; MP_TAC (ISPECL [`y:real^N`; `f:real^N->real`] partial_uni_compose') THEN ASM_REWRITE_TAC[] THEN - DISCH_THEN (MP_TAC o SPEC `(\x:real. x pow k)`) THEN + DISCH_THEN (MP_TAC o SPEC `(\x:real. x pow k)`) THEN REWRITE_TAC[REAL_DIFFERENTIABLE_AT_POW] THEN ASM_SIMP_TAC[derivative_pow_x]; ALL_TAC @@ -971,8 +977,8 @@ let MK_M_TAYLOR_POW2' = (UNDISCH_ALL o PURE_REWRITE_RULE[all_n1_raw] o prove) diff2c_domain domain f ==> interval_arith ((f y) pow 2) bounds ==> all_n 1 d_bounds_list (\i int. interval_arith ((&2 * f y) * partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i - (\j int. interval_arith (&2 * (partial j f x * partial i f x + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (\j int. interval_arith (&2 * (partial j f x * partial i f x + f x * partial2 j i f x)) int))) ==> m_taylor_interval (\x. (f x) pow 2) domain y w bounds d_bounds_list dd_bounds_list`, REPEAT STRIP_TAC THEN MP_TAC (INST[`2`, `k:num`] MK_M_TAYLOR_POW) THEN @@ -986,7 +992,7 @@ let MK_M_TAYLOR_POW' = (UNDISCH_ALL o PURE_REWRITE_RULE[all_n1_raw] o prove) interval_arith (f y * f y pow (k - 1)) bounds ==> k >= 2 ==> all_n 1 d_bounds_list (\i int. interval_arith ((&k * f y pow (k - 1)) * partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith ((&k * f x pow (k - 2)) * ((&(k - 1) * partial j f x) * partial i f x + f x * partial2 j i f x)) int))) ==> m_taylor_interval (\x. (f x) pow k) domain y w bounds d_bounds_list dd_bounds_list`, @@ -994,7 +1000,7 @@ let MK_M_TAYLOR_POW' = (UNDISCH_ALL o PURE_REWRITE_RULE[all_n1_raw] o prove) ASM_REWRITE_TAC[GSYM REAL_OF_NUM_MUL; IMP_IMP] THEN DISCH_THEN MATCH_MP_TAC THEN CONJ_TAC THENL [ SUBGOAL_THEN `!x. x pow k = x * x pow (k - 1)` ASSUME_TAC THENL [ - GEN_TAC THEN MP_TAC (ARITH_RULE `k >= 2 ==> k = 1 + (k - 1)`) THEN ASM_REWRITE_TAC[] THEN + GEN_TAC THEN MP_TAC (ARITH_RULE `k >= 2 ==> k = 1 + (k - 1)`) THEN ASM_REWRITE_TAC[] THEN DISCH_THEN (fun th -> GEN_REWRITE_TAC (LAND_CONV o DEPTH_CONV) [th]) THEN REWRITE_TAC[REAL_POW_ADD; REAL_POW_1]; ALL_TAC @@ -1007,11 +1013,11 @@ let MK_M_TAYLOR_POW' = (UNDISCH_ALL o PURE_REWRITE_RULE[all_n1_raw] o prove) REWRITE_TAC[REAL_POW_ADD; REAL_POW_1]; ALL_TAC ] THEN - ASM_REWRITE_TAC[REAL_ARITH `(((k * k1) * f2) * j) * i + (k * f2 * f) * x = + ASM_REWRITE_TAC[REAL_ARITH `(((k * k1) * f2) * j) * i + (k * f2 * f) * x = (k * f2) * ((k1 * j) * i + f * x)`]);; (* inv *) -let MK_M_TAYLOR_INV' = (UNDISCH_ALL o PURE_REWRITE_RULE[float2_eq] o DISCH_ALL o +let MK_M_TAYLOR_INV' = (UNDISCH_ALL o PURE_REWRITE_RULE[float2_eq] o DISCH_ALL o MY_RULE_FLOAT o prove) (`m_cell_domain domain (y:real^N) w ==> (!x. x IN interval [domain] ==> interval_arith (f x) f_bounds) ==> @@ -1019,11 +1025,11 @@ let MK_M_TAYLOR_INV' = (UNDISCH_ALL o PURE_REWRITE_RULE[float2_eq] o DISCH_ALL o diff2c_domain domain f ==> interval_arith (inv (f y)) bounds ==> all_n 1 d_bounds_list (\i int. interval_arith (--inv (f y * f y) * partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith (((&2 * inv (f x * f x * f x)) * partial j f x) * partial i f x - inv (f x * f x) * partial2 j i f x) int))) ==> m_taylor_interval (\x. inv (f x)) domain y w bounds d_bounds_list dd_bounds_list`, - REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded; ETA_AX] THEN + REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded; ETA_AX] THEN REPEAT DISCH_TAC THEN SUBGOAL_THEN `lift o f differentiable at (y:real^N)` ASSUME_TAC THENL [ @@ -1094,7 +1100,7 @@ let MK_M_TAYLOR_INV' = (UNDISCH_ALL o PURE_REWRITE_RULE[float2_eq] o DISCH_ALL o (* sqrt *) -let MK_M_TAYLOR_SQRT' = (UNDISCH_ALL o PURE_REWRITE_RULE[float2_eq; float4_eq] o +let MK_M_TAYLOR_SQRT' = (UNDISCH_ALL o PURE_REWRITE_RULE[float2_eq; float4_eq] o DISCH_ALL o MY_RULE_FLOAT o prove) (`m_cell_domain domain (y:real^N) w ==> (!x. x IN interval [domain] ==> interval_arith (f x) f_bounds) ==> @@ -1102,7 +1108,7 @@ let MK_M_TAYLOR_SQRT' = (UNDISCH_ALL o PURE_REWRITE_RULE[float2_eq; float4_eq] o diff2c_domain domain f ==> interval_arith (sqrt (f y)) bounds ==> all_n 1 d_bounds_list (\i int. interval_arith (inv (&2 * sqrt (f y)) * partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith ((--inv ((&2 * sqrt (f x)) * (&2 * f x)) * partial j f x) * partial i f x + inv (&2 * sqrt (f x)) * partial2 j i f x) int))) ==> m_taylor_interval (\x. sqrt (f x)) domain y w bounds d_bounds_list dd_bounds_list`, @@ -1183,14 +1189,14 @@ let MK_M_TAYLOR_SQRT' = (UNDISCH_ALL o PURE_REWRITE_RULE[float2_eq; float4_eq] o ASM_SIMP_TAC[REAL_ARITH `&4 * a * b = (&2 * a) * (&2 * b)`]);; (* exp *) -let MK_M_TAYLOR_EXP' = (UNDISCH_ALL o PURE_REWRITE_RULE[float1_eq] o DISCH_ALL o +let MK_M_TAYLOR_EXP' = (UNDISCH_ALL o PURE_REWRITE_RULE[float1_eq] o DISCH_ALL o MY_RULE_FLOAT o prove) (`m_cell_domain domain (y:real^N) w ==> diff2c_domain domain f ==> interval_arith (exp (f y)) bounds ==> all_n 1 d_bounds_list (\i int. interval_arith (exp (f y) * partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i - (\j int. interval_arith ((exp (f x) * partial j f x) + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (\j int. interval_arith ((exp (f x) * partial j f x) * partial i f x + exp (f x) * partial2 j i f x) int))) ==> m_taylor_interval (\x. exp (f x)) domain y w bounds d_bounds_list dd_bounds_list`, REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded; ETA_AX] THEN @@ -1233,7 +1239,7 @@ let MK_M_TAYLOR_EXP' = (UNDISCH_ALL o PURE_REWRITE_RULE[float1_eq] o DISCH_ALL o ASM_SIMP_TAC[nth_derivative2; second_derivative_exp; derivative_exp]);; (* log *) -let MK_M_TAYLOR_LOG' = (UNDISCH_ALL o PURE_REWRITE_RULE[num2_eq] o +let MK_M_TAYLOR_LOG' = (UNDISCH_ALL o PURE_REWRITE_RULE[num2_eq] o DISCH_ALL o MY_RULE_FLOAT o prove) (`m_cell_domain domain (y:real^N) w ==> (!x. x IN interval [domain] ==> interval_arith (f x) f_bounds) ==> @@ -1241,7 +1247,7 @@ let MK_M_TAYLOR_LOG' = (UNDISCH_ALL o PURE_REWRITE_RULE[num2_eq] o diff2c_domain domain f ==> interval_arith (log (f y)) bounds ==> all_n 1 d_bounds_list (\i int. interval_arith (inv (f y) * partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith ((--(inv (f x) pow 2) * partial j f x) * partial i f x + inv (f x) * partial2 j i f x) int))) ==> m_taylor_interval (\x. log (f x)) domain y w bounds d_bounds_list dd_bounds_list`, @@ -1314,14 +1320,14 @@ let MK_M_TAYLOR_LOG' = (UNDISCH_ALL o PURE_REWRITE_RULE[num2_eq] o ASM_SIMP_TAC[second_derivative_log; derivative_log; REAL_INV_POW]);; (* atn *) -let MK_M_TAYLOR_ATN' = (UNDISCH_ALL o PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o DISCH_ALL o +let MK_M_TAYLOR_ATN' = (UNDISCH_ALL o PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o DISCH_ALL o MY_RULE_FLOAT o prove) (`m_cell_domain domain (y:real^N) w ==> diff2c_domain domain f ==> interval_arith (atn (f y)) bounds ==> all_n 1 d_bounds_list (\i int. interval_arith (inv (&1 + f y * f y) * partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i - (\j int. interval_arith ((((-- &2 * f x) * inv (&1 + f x * f x) pow 2) * partial j f x) + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (\j int. interval_arith ((((-- &2 * f x) * inv (&1 + f x * f x) pow 2) * partial j f x) * partial i f x + inv (&1 + f x * f x) * partial2 j i f x) int))) ==> m_taylor_interval (\x. atn (f x)) domain y w bounds d_bounds_list dd_bounds_list`, REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded; ETA_AX] THEN @@ -1366,14 +1372,14 @@ let MK_M_TAYLOR_ATN' = (UNDISCH_ALL o PURE_REWRITE_RULE[float1_eq; float2_eq; nu ASM_SIMP_TAC[]);; (* cos *) -let MK_M_TAYLOR_COS' = (UNDISCH_ALL o PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o DISCH_ALL o +let MK_M_TAYLOR_COS' = (UNDISCH_ALL o PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o DISCH_ALL o MY_RULE_FLOAT o prove) (`m_cell_domain domain (y:real^N) w ==> diff2c_domain domain f ==> interval_arith (cos (f y)) bounds ==> all_n 1 d_bounds_list (\i int. interval_arith (--sin (f y) * partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i - (\j int. interval_arith (--((cos (f x) * partial j f x) * partial i f x + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (\j int. interval_arith (--((cos (f x) * partial j f x) * partial i f x + sin (f x) * partial2 j i f x)) int))) ==> m_taylor_interval (\x. cos (f x)) domain y w bounds d_bounds_list dd_bounds_list`, REWRITE_TAC[m_taylor_interval; m_lin_approx; second_bounded; ETA_AX] THEN @@ -1415,13 +1421,13 @@ let MK_M_TAYLOR_COS' = (UNDISCH_ALL o PURE_REWRITE_RULE[float1_eq; float2_eq; nu ASM_SIMP_TAC[REAL_ARITH `(--c * p1) * p2 + --s * p3 = --((c * p1) * p2 + s * p3)`]);; (* sin *) -let MK_M_TAYLOR_SIN' = (UNDISCH_ALL o PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o DISCH_ALL o +let MK_M_TAYLOR_SIN' = (UNDISCH_ALL o PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o DISCH_ALL o MY_RULE_FLOAT o prove) (`m_cell_domain domain (y:real^N) w ==> diff2c_domain domain f ==> interval_arith (sin (f y)) bounds ==> all_n 1 d_bounds_list (\i int. interval_arith (cos (f y) * partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith (((--sin (f x)) * partial j f x) * partial i f x + cos (f x) * partial2 j i f x) int))) ==> m_taylor_interval (\x. sin (f x)) domain y w bounds d_bounds_list dd_bounds_list`, @@ -1466,16 +1472,16 @@ let MK_M_TAYLOR_SIN' = (UNDISCH_ALL o PURE_REWRITE_RULE[float1_eq; float2_eq; nu let iabs_lemma = GEN_REWRITE_RULE (RAND_CONV o RAND_CONV) [GSYM float1_eq] (REFL `iabs f_bounds < &1`);; let MK_M_TAYLOR_ACS' = (UNDISCH_ALL o PURE_ONCE_REWRITE_RULE[iabs_lemma] o - PURE_REWRITE_RULE[float1_eq; num3_eq] o DISCH_ALL o + PURE_REWRITE_RULE[float1_eq; num3_eq] o DISCH_ALL o MY_RULE_FLOAT o prove) (`m_cell_domain domain (y:real^N) w ==> (!x. x IN interval [domain] ==> interval_arith (f x) f_bounds) ==> iabs f_bounds < &1 ==> diff2c_domain domain f ==> interval_arith (acs (f y)) bounds ==> - all_n 1 d_bounds_list + all_n 1 d_bounds_list (\i int. interval_arith (--inv (sqrt (&1 - f y * f y)) * partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith ((--(f x / sqrt ((&1 - f x * f x) pow 3)) * partial j f x) * partial i f x - inv (sqrt (&1 - f x * f x)) * partial2 j i f x) int))) ==> m_taylor_interval (\x. acs (f x)) domain y w bounds d_bounds_list dd_bounds_list`, @@ -1550,16 +1556,16 @@ let MK_M_TAYLOR_ACS' = (UNDISCH_ALL o PURE_ONCE_REWRITE_RULE[iabs_lemma] o (* asn *) let MK_M_TAYLOR_ASN' = (UNDISCH_ALL o PURE_ONCE_REWRITE_RULE[iabs_lemma] o - PURE_REWRITE_RULE[float1_eq; num3_eq] o DISCH_ALL o + PURE_REWRITE_RULE[float1_eq; num3_eq] o DISCH_ALL o MY_RULE_FLOAT o prove) (`m_cell_domain domain (y:real^N) w ==> (!x. x IN interval [domain] ==> interval_arith (f x) f_bounds) ==> iabs f_bounds < &1 ==> diff2c_domain domain f ==> interval_arith (asn (f y)) bounds ==> - all_n 1 d_bounds_list + all_n 1 d_bounds_list (\i int. interval_arith (inv (sqrt (&1 - f y * f y)) * partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith (((f x / sqrt ((&1 - f x * f x) pow 3)) * partial j f x) * partial i f x + inv (sqrt (&1 - f x * f x)) * partial2 j i f x) int))) ==> m_taylor_interval (\x. asn (f x)) domain y w bounds d_bounds_list dd_bounds_list`, @@ -1636,17 +1642,17 @@ let MK_M_TAYLOR_ASN' = (UNDISCH_ALL o PURE_ONCE_REWRITE_RULE[iabs_lemma] o let float_neg1_eq = FLOAT_TO_NUM_CONV (mk_float (-1) 0);; -let MK_M_TAYLOR_MATAN' = (UNDISCH_ALL o - PURE_REWRITE_RULE[float1_eq; SYM float_neg1_eq] o DISCH_ALL o +let MK_M_TAYLOR_MATAN' = (UNDISCH_ALL o + PURE_REWRITE_RULE[float1_eq; SYM float_neg1_eq] o DISCH_ALL o MY_RULE_FLOAT o prove) (`m_cell_domain domain (y:real^N) w ==> (!x. x IN interval [domain] ==> interval_arith (f x) f_bounds) ==> interval_gt (-- &1) f_bounds ==> diff2c_domain domain f ==> interval_arith (matan (f y)) bounds ==> - all_n 1 d_bounds_list + all_n 1 d_bounds_list (\i int. interval_arith (dmatan (f y) * partial i f y) int) ==> - (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i + (!x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith ((ddmatan (f x) * partial j f x) * partial i f x + dmatan (f x) * partial2 j i f x) int))) ==> m_taylor_interval (\x. matan (f x)) domain y w bounds d_bounds_list dd_bounds_list`, @@ -1717,30 +1723,30 @@ let MK_M_TAYLOR_MATAN' = (UNDISCH_ALL o (* eval_m_taylor_pow2 *) (* ----------------------- *) -let pow2_partial_lemma' = +let pow2_partial_lemma' = prove(`interval_arith ((&2 * f y) * partial i f (y:real^N)) int <=> (\i int. interval_arith ((&2 * f y) * partial i f y) int) i int`, REWRITE_TAC[]);; -let pow2_second_lemma' = - prove(`interval_arith (&2 * (partial j f (x:real^N) * partial i f x + +let pow2_second_lemma' = + prove(`interval_arith (&2 * (partial j f (x:real^N) * partial i f x + f x * partial2 j i f x)) int <=> - (\j int. interval_arith (&2 * (partial j f x * partial i f x + + (\j int. interval_arith (&2 * (partial j f x * partial i f x + f x * partial2 j i f x)) int) j int`, REWRITE_TAC[]);; let pow2_second_lemma'' = (PURE_REWRITE_RULE[all_n1_raw] o prove) - (`all_n 1 list - (\j int. interval_arith (&2 * (partial j f (x:real^N) * partial i f x + + (`all_n 1 list + (\j int. interval_arith (&2 * (partial j f (x:real^N) * partial i f x + f x * partial2 j i f x)) int) <=> - (\i list. all_n 1 list - (\j int. interval_arith (&2 * (partial j f x * partial i f x + + (\i list. all_n 1 list + (\j int. interval_arith (&2 * (partial j f x * partial i f x + f x * partial2 j i f x)) int)) i list`, REWRITE_TAC[]);; -let eval_m_taylor_pow2 n p_lin p_second taylor1_th = +let eval_m_taylor_pow2 n p_lin p_second taylor1_th = let domain_th, diff2_f1_th, lin1_th, second1_th = dest_m_taylor_thms n taylor1_th in let f1_tm = (rand o concl) diff2_f1_th in let domain_tm, y_tm, w_tm = dest_m_cell_domain (concl domain_th) in @@ -1758,15 +1764,15 @@ let eval_m_taylor_pow2 n p_lin p_second taylor1_th = let f1_bound = undisch f1_bound0 in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_pow p_lin 2 bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* partial_lemma' *) - let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o + let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) pow2_partial_lemma' in - let u_bounds = + let u_bounds = let ( * ) = float_interval_mul p_lin in two_interval * bounds1_th in @@ -1784,15 +1790,15 @@ let eval_m_taylor_pow2 n p_lin p_second taylor1_th = let d_bounds_list = (rand o rator o concl) df_th in let dd1 = second_bounded_components n second1_th in - + (* second_lemma', second_lemma'' *) - let u_second_lemma0 = (INST[f1_tm, f_var] o + let u_second_lemma0 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) pow2_second_lemma' in let u_second_lemma1 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) pow2_second_lemma'' in - let d1_bounds = map (fun i -> + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -1808,14 +1814,14 @@ let eval_m_taylor_pow2 n p_lin p_second taylor1_th = let dj1 = List.nth d1_bounds (j_int - 1) in (* partial2 *) - let u_th = + let u_th = let ( * ), ( + ) = float_interval_mul p_second, float_interval_add p_second in two_interval * (dj1 * di1 + f1_bound * th1) in let int_tm = rand (concl u_th) in let th0 = INST[j_tm, j_var_num; int_tm, int_var] lemma in EQ_MP th0 u_th in - + let u_th = eval_all_n th1 true u_second in let list_tm = (rand o rator o concl) u_th in let lemma1 = INST[i_tm, i_var_num; list_tm, list_var_real_pair] u_second_lemma1 in @@ -1825,11 +1831,11 @@ let eval_m_taylor_pow2 n p_lin p_second taylor1_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_POW2' in let k_tm = rand (rand (rator (concl bounds_th))) in @@ -1840,26 +1846,26 @@ let eval_m_taylor_pow2 n p_lin p_second taylor1_th = (* eval_m_taylor_pow *) (* ----------------------- *) -let pow_partial_lemma' = +let pow_partial_lemma' = prove(`interval_arith ((&k * f y pow (k - 1)) * partial i f (y:real^N)) int <=> (\i int. interval_arith ((&k * f y pow (k - 1)) * partial i f y) int) i int`, REWRITE_TAC[]);; -let pow_second_lemma' = - prove(`interval_arith ((&k * f x pow (k - 2)) * ((&(k - 1) * partial j f (x:real^N)) +let pow_second_lemma' = + prove(`interval_arith ((&k * f x pow (k - 2)) * ((&(k - 1) * partial j f (x:real^N)) * partial i f x + f x * partial2 j i f x)) int <=> - (\j int. interval_arith ((&k * f x pow (k - 2)) * ((&(k - 1) * partial j f (x:real^N)) + (\j int. interval_arith ((&k * f x pow (k - 2)) * ((&(k - 1) * partial j f (x:real^N)) * partial i f x + f x * partial2 j i f x)) int) j int`, REWRITE_TAC[]);; let pow_second_lemma'' = (PURE_REWRITE_RULE[all_n1_raw] o prove) - (`all_n 1 list - (\j int. interval_arith ((&k * f x pow (k - 2)) * ((&(k - 1) * partial j f (x:real^N)) + (`all_n 1 list + (\j int. interval_arith ((&k * f x pow (k - 2)) * ((&(k - 1) * partial j f (x:real^N)) * partial i f x + f x * partial2 j i f x)) int) <=> - (\i list. all_n 1 list - (\j int. interval_arith ((&k * f x pow (k - 2)) * ((&(k - 1) * partial j f (x:real^N)) + (\i list. all_n 1 list + (\j int. interval_arith ((&k * f x pow (k - 2)) * ((&(k - 1) * partial j f (x:real^N)) * partial i f x + f x * partial2 j i f x)) int)) i list`, REWRITE_TAC[]);; @@ -1898,16 +1904,16 @@ let eval_m_taylor_pow k = let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in let bounds_pow_k1 = float_interval_pow p_lin (k - 1) bounds1_th in - let bounds_th = + let bounds_th = let ( * ) = float_interval_mul p_lin in bounds1_th * bounds_pow_k1 in let bounds_tm = (rand o concl) bounds_th in (* partial_lemma' *) - let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o + let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) pow_partial_lemma' in - let u_bounds = + let u_bounds = let ( * ) = float_interval_mul p_lin in k_interval * bounds_pow_k1 in @@ -1925,15 +1931,15 @@ let eval_m_taylor_pow k = let d_bounds_list = (rand o rator o concl) df_th in let dd1 = second_bounded_components n second1_th in - + (* second_lemma', second_lemma'' *) - let u_second_lemma0 = (INST[f1_tm, f_var] o + let u_second_lemma0 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) pow_second_lemma' in let u_second_lemma1 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) pow_second_lemma'' in - let d1_bounds = map (fun i -> + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -1955,14 +1961,14 @@ let eval_m_taylor_pow k = let dj1 = List.nth d1_bounds (j_int - 1) in (* partial2 *) - let u_th = + let u_th = let ( * ), ( + ) = float_interval_mul p_second, float_interval_add p_second in d2_th0 * ((k1_interval * dj1) * di1 + f1_bound * th1) in let int_tm = rand (concl u_th) in let th0 = INST[j_tm, j_var_num; int_tm, int_var] lemma in EQ_MP th0 u_th in - + let u_th = eval_all_n th1 true u_second in let list_tm = (rand o rator o concl) u_th in let lemma1 = INST[i_tm, i_var_num; list_tm, list_var_real_pair] u_second_lemma1 in @@ -1972,11 +1978,11 @@ let eval_m_taylor_pow k = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_POW in let eq_th = pow_beta_gen_eq f1_tm x_var k_tm in @@ -1987,12 +1993,12 @@ let eval_m_taylor_pow k = (* eval_m_taylor_inv *) (* ----------------------- *) -let inv_partial_lemma' = +let inv_partial_lemma' = prove(`interval_arith (--inv (f y * f y) * partial i f (y:real^N)) int <=> (\i int. interval_arith (--inv (f y * f y) * partial i f y) int) i int`, REWRITE_TAC[]);; -let inv_second_lemma' = +let inv_second_lemma' = prove(`interval_arith (((&2 * inv (f x * f x * f x)) * partial j f x) * partial i f (x:real^N) - inv (f x * f x) * partial2 j i f x) int <=> (\j int. interval_arith (((&2 * inv (f x * f x * f x)) * partial j f x) * partial i f x - @@ -2001,10 +2007,10 @@ let inv_second_lemma' = let inv_second_lemma'' = (PURE_REWRITE_RULE[GSYM num1_eq] o prove) - (`all_n 1 list + (`all_n 1 list (\j int. interval_arith (((&2 * inv (f x * f x * f x)) * partial j f x) * partial i f (x:real^N) - inv (f x * f x) * partial2 j i f x) int) <=> - (\i list. all_n 1 list + (\i list. all_n 1 list (\j int. interval_arith (((&2 * inv (f x * f x * f x)) * partial j f x) * partial i f x - inv (f x * f x) * partial2 j i f x) int)) i list`, REWRITE_TAC[]);; @@ -2034,15 +2040,15 @@ let eval_m_taylor_inv n p_lin p_second taylor1_th = let cond_th = check_interval_not_zero f_bounds_tm in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_inv p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* partial_lemma' *) - let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o + let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) inv_partial_lemma' in - let u_bounds = + let u_bounds = let neg, inv, ( * ) = float_interval_neg, float_interval_inv p_lin, float_interval_mul p_lin in neg (inv (bounds1_th * bounds1_th)) in @@ -2062,16 +2068,16 @@ let eval_m_taylor_inv n p_lin p_second taylor1_th = let dd1 = second_bounded_components n second1_th in - + (* second_lemma', second_lemma'' *) - let u_second_lemma0 = (INST[f1_tm, f_var] o + let u_second_lemma0 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) inv_second_lemma' in let u_second_lemma1 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) inv_second_lemma'' in - let d1_bounds = map (fun i -> + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -2079,7 +2085,7 @@ let eval_m_taylor_inv n p_lin p_second taylor1_th = let d1_th0, d2_th0 = let inv, ( * ) = float_interval_inv p_second, float_interval_mul p_second in let ff = f1_bound * f1_bound in - inv ff, + inv ff, two_interval * inv (f1_bound * ff) in @@ -2095,14 +2101,14 @@ let eval_m_taylor_inv n p_lin p_second taylor1_th = let dj1 = List.nth d1_bounds (j_int - 1) in (* partial2 *) - let u_th = + let u_th = let ( * ), ( - ) = float_interval_mul p_second, float_interval_sub p_second in (d2_th0 * dj1) * di1 - d1_th0 * th1 in let int_tm = rand (concl u_th) in let th0 = INST[j_tm, j_var_num; int_tm, int_var] lemma in EQ_MP th0 u_th in - + let u_th = eval_all_n th1 true u_second in let list_tm = (rand o rator o concl) u_th in let lemma1 = INST[i_tm, i_var_num; list_tm, list_var_real_pair] u_second_lemma1 in @@ -2112,12 +2118,12 @@ let eval_m_taylor_inv n p_lin p_second taylor1_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP cond_th o MY_PROVE_HYP f1_bound0 o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_INV' in let eq_th = unary_beta_gen_eq f1_tm x_var inv_op_real in @@ -2127,13 +2133,13 @@ let eval_m_taylor_inv n p_lin p_second taylor1_th = (* eval_m_taylor_sqrt *) (* ----------------------- *) -let sqrt_partial_lemma' = +let sqrt_partial_lemma' = prove(`interval_arith (inv (&2 * sqrt (f y)) * partial i f (y:real^N)) int <=> (\i int. interval_arith (inv (&2 * sqrt (f y)) * partial i f y) int) i int`, REWRITE_TAC[]);; -let sqrt_second_lemma' = +let sqrt_second_lemma' = prove(`interval_arith ((--inv ((&2 * sqrt (f x)) * (&2 * f x)) * partial j f x) * partial i f (x:real^N) + inv (&2 * sqrt (f x)) * partial2 j i f x) int <=> (\j int. interval_arith ((--inv ((&2 * sqrt (f x))*(&2 * f x)) * partial j f x) * partial i f x + @@ -2142,10 +2148,10 @@ let sqrt_second_lemma' = let sqrt_second_lemma'' = (PURE_REWRITE_RULE[GSYM num1_eq] o prove) - (`all_n 1 list + (`all_n 1 list (\j int. interval_arith ((--inv ((&2 * sqrt (f x)) * (&2 * f x)) * partial j f x) * partial i f x + inv (&2 * sqrt (f x)) * partial2 j i f (x:real^N)) int) <=> - (\i list. all_n 1 list + (\i list. all_n 1 list (\j int. interval_arith ((--inv ((&2 * sqrt (f x)) * (&2 * f x)) * partial j f x) * partial i f x + inv (&2 * sqrt (f x)) * partial2 j i f (x:real^N)) int)) i list`, REWRITE_TAC[]);; @@ -2173,15 +2179,15 @@ let eval_m_taylor_sqrt n p_lin p_second taylor1_th = let cond_th = check_interval_pos f_bounds_tm in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_sqrt p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* partial_lemma' *) - let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o + let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) sqrt_partial_lemma' in - let u_bounds = + let u_bounds = let inv, ( * ) = float_interval_inv p_lin, float_interval_mul p_lin in inv (two_interval * bounds_th) in @@ -2200,15 +2206,15 @@ let eval_m_taylor_sqrt n p_lin p_second taylor1_th = let dd1 = second_bounded_components n second1_th in - + (* second_lemma', second_lemma'' *) - let u_second_lemma0 = (INST[f1_tm, f_var] o + let u_second_lemma0 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) sqrt_second_lemma' in let u_second_lemma1 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) sqrt_second_lemma'' in - let d1_bounds = map (fun i -> + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -2232,14 +2238,14 @@ let eval_m_taylor_sqrt n p_lin p_second taylor1_th = let dj1 = List.nth d1_bounds (j_int - 1) in (* partial2 *) - let u_th = + let u_th = let ( * ), ( + ) = float_interval_mul p_second, float_interval_add p_second in (d2_th0 * dj1) * di1 + d1_th0 * th1 in let int_tm = rand (concl u_th) in let th0 = INST[j_tm, j_var_num; int_tm, int_var] lemma in EQ_MP th0 u_th in - + let u_th = eval_all_n th1 true u_second in let list_tm = (rand o rator o concl) u_th in let lemma1 = INST[i_tm, i_var_num; list_tm, list_var_real_pair] u_second_lemma1 in @@ -2249,12 +2255,12 @@ let eval_m_taylor_sqrt n p_lin p_second taylor1_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP cond_th o MY_PROVE_HYP f1_bound0 o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_SQRT' in let eq_th = unary_beta_gen_eq f1_tm x_var sqrt_tm in @@ -2265,32 +2271,32 @@ let eval_m_taylor_sqrt n p_lin p_second taylor1_th = (* eval_m_taylor_exp *) (* ----------------------- *) -let exp_partial_lemma' = +let exp_partial_lemma' = prove(`interval_arith (exp (f y) * partial i f (y:real^N)) int <=> (\i int. interval_arith (exp (f y) * partial i f y) int) i int`, REWRITE_TAC[]);; -let exp_second_lemma' = - prove(`interval_arith ((exp (f x) * partial j f (x:real^N)) +let exp_second_lemma' = + prove(`interval_arith ((exp (f x) * partial j f (x:real^N)) * partial i f x + exp (f x) * partial2 j i f x) int <=> - (\j int. interval_arith ((exp (f x) * partial j f x) + (\j int. interval_arith ((exp (f x) * partial j f x) * partial i f x + exp (f x) * partial2 j i f x) int) j int`, REWRITE_TAC[]);; -let exp_second_lemma'' = (PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o NUMERALS_TO_NUM o +let exp_second_lemma'' = (PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o NUMERALS_TO_NUM o PURE_REWRITE_RULE[FLOAT_OF_NUM; min_exp_def] o prove) - (`all_n 1 list - (\j int. interval_arith ((exp (f x) * partial j f (x:real^N)) + (`all_n 1 list + (\j int. interval_arith ((exp (f x) * partial j f (x:real^N)) * partial i f x + exp (f x) * partial2 j i f x) int) <=> - (\i list. all_n 1 list - (\j int. interval_arith ((exp (f x) * partial j f x) + (\i list. all_n 1 list + (\j int. interval_arith ((exp (f x) * partial j f x) * partial i f x + exp (f x) * partial2 j i f x) int)) i list`, REWRITE_TAC[]);; -let eval_m_taylor_exp n p_lin p_second taylor1_th = +let eval_m_taylor_exp n p_lin p_second taylor1_th = let domain_th, diff2_f1_th, lin1_th, second1_th = dest_m_taylor_thms n taylor1_th in let f1_tm = (rand o concl) diff2_f1_th in let domain_tm, y_tm, w_tm = dest_m_cell_domain (concl domain_th) in @@ -2308,15 +2314,15 @@ let eval_m_taylor_exp n p_lin p_second taylor1_th = let f1_bound = undisch f1_bound0 in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_exp p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* partial_lemma' *) - let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o + let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) exp_partial_lemma' in - let u_bounds = + let u_bounds = let exp = float_interval_exp p_lin in exp (bounds1_th) in @@ -2334,15 +2340,15 @@ let eval_m_taylor_exp n p_lin p_second taylor1_th = let d_bounds_list = (rand o rator o concl) df_th in let dd1 = second_bounded_components n second1_th in - + (* second_lemma', second_lemma'' *) - let u_second_lemma0 = (INST[f1_tm, f_var] o + let u_second_lemma0 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) exp_second_lemma' in let u_second_lemma1 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) exp_second_lemma'' in - let d1_bounds = map (fun i -> + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -2364,14 +2370,14 @@ let eval_m_taylor_exp n p_lin p_second taylor1_th = let dj1 = List.nth d1_bounds (j_int - 1) in (* partial2 *) - let u_th = + let u_th = let ( * ), ( + ) = float_interval_mul p_second, float_interval_add p_second in (d2_th0 * dj1) * di1 + d1_th0 * th1 in let int_tm = rand (concl u_th) in let th0 = INST[j_tm, j_var_num; int_tm, int_var] lemma in EQ_MP th0 u_th in - + let u_th = eval_all_n th1 true u_second in let list_tm = (rand o rator o concl) u_th in let lemma1 = INST[i_tm, i_var_num; list_tm, list_var_real_pair] u_second_lemma1 in @@ -2381,11 +2387,11 @@ let eval_m_taylor_exp n p_lin p_second taylor1_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_EXP' in let eq_th = unary_beta_gen_eq f1_tm x_var exp_tm in @@ -2396,13 +2402,13 @@ let eval_m_taylor_exp n p_lin p_second taylor1_th = (* eval_m_taylor_log *) (* ----------------------- *) -let log_partial_lemma' = +let log_partial_lemma' = prove(`interval_arith (inv (f y) * partial i f (y:real^N)) int <=> (\i int. interval_arith (inv (f y) * partial i f y) int) i int`, REWRITE_TAC[]);; -let log_second_lemma' = +let log_second_lemma' = prove(`interval_arith ((--(inv (f x) pow 2) * partial j f x) * partial i f (x:real^N) + inv (f x) * partial2 j i f x) int <=> (\j int. interval_arith ((--(inv (f x) pow 2) * partial j f x) * partial i f x + @@ -2411,10 +2417,10 @@ let log_second_lemma' = let log_second_lemma'' = (PURE_REWRITE_RULE[GSYM num1_eq] o prove) - (`all_n 1 list + (`all_n 1 list (\j int. interval_arith ((--(inv (f x) pow 2) * partial j f x) * partial i f x + inv (f x) * partial2 j i f (x:real^N)) int) <=> - (\i list. all_n 1 list + (\i list. all_n 1 list (\j int. interval_arith ((--(inv (f x) pow 2) * partial j f x) * partial i f x + inv (f x) * partial2 j i f (x:real^N)) int)) i list`, REWRITE_TAC[]);; @@ -2442,15 +2448,15 @@ let eval_m_taylor_log n p_lin p_second taylor1_th = let cond_th = check_interval_pos f_bounds_tm in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_log p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* partial_lemma' *) - let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o + let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) log_partial_lemma' in - let u_bounds = + let u_bounds = let inv = float_interval_inv p_lin in inv (bounds1_th) in @@ -2468,21 +2474,21 @@ let eval_m_taylor_log n p_lin p_second taylor1_th = let d_bounds_list = (rand o rator o concl) df_th in let dd1 = second_bounded_components n second1_th in - + (* second_lemma', second_lemma'' *) - let u_second_lemma0 = (INST[f1_tm, f_var] o + let u_second_lemma0 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) log_second_lemma' in let u_second_lemma1 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) log_second_lemma'' in - let d1_bounds = map (fun i -> + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in (* u'(f x), u''(f x) *) let d1_th0, d2_th0 = - let neg, pow2, inv = + let neg, pow2, inv = float_interval_neg, float_interval_pow p_second 2, float_interval_inv p_second in let inv_f = inv f1_bound in inv_f, neg (pow2 inv_f) in @@ -2499,14 +2505,14 @@ let eval_m_taylor_log n p_lin p_second taylor1_th = let dj1 = List.nth d1_bounds (j_int - 1) in (* partial2 *) - let u_th = + let u_th = let ( * ), ( + ) = float_interval_mul p_second, float_interval_add p_second in (d2_th0 * dj1) * di1 + d1_th0 * th1 in let int_tm = rand (concl u_th) in let th0 = INST[j_tm, j_var_num; int_tm, int_var] lemma in EQ_MP th0 u_th in - + let u_th = eval_all_n th1 true u_second in let list_tm = (rand o rator o concl) u_th in let lemma1 = INST[i_tm, i_var_num; list_tm, list_var_real_pair] u_second_lemma1 in @@ -2516,12 +2522,12 @@ let eval_m_taylor_log n p_lin p_second taylor1_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP cond_th o MY_PROVE_HYP f1_bound0 o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_LOG' in let eq_th = unary_beta_gen_eq f1_tm x_var log_tm in @@ -2533,33 +2539,33 @@ let eval_m_taylor_log n p_lin p_second taylor1_th = (* eval_m_taylor_atn *) (* ----------------------- *) -let atn_partial_lemma' = +let atn_partial_lemma' = prove(`interval_arith (inv (&1 + f y * f y) * partial i f (y:real^N)) int <=> (\i int. interval_arith (inv (&1 + f y * f y) * partial i f y) int) i int`, REWRITE_TAC[]);; -let atn_second_lemma' = - prove(`interval_arith ((((-- &2 * f x) * inv (&1 + f x * f x) pow 2) * partial j f (x:real^N)) +let atn_second_lemma' = + prove(`interval_arith ((((-- &2 * f x) * inv (&1 + f x * f x) pow 2) * partial j f (x:real^N)) * partial i f x + inv (&1 + f x * f x) * partial2 j i f x) int <=> - (\j int. interval_arith ((((-- &2 * f x) * inv (&1 + f x * f x) pow 2) * partial j f x) + (\j int. interval_arith ((((-- &2 * f x) * inv (&1 + f x * f x) pow 2) * partial j f x) * partial i f x + inv (&1 + f x * f x) * partial2 j i f x) int) j int`, REWRITE_TAC[]);; -let atn_second_lemma'' = (PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o NUMERALS_TO_NUM o +let atn_second_lemma'' = (PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o NUMERALS_TO_NUM o PURE_REWRITE_RULE[FLOAT_OF_NUM; min_exp_def] o prove) - (`all_n 1 list - (\j int. interval_arith ((((-- &2 * f x) * inv (&1 + f x * f x) pow 2) * partial j f (x:real^N)) + (`all_n 1 list + (\j int. interval_arith ((((-- &2 * f x) * inv (&1 + f x * f x) pow 2) * partial j f (x:real^N)) * partial i f x + inv (&1 + f x * f x) * partial2 j i f x) int) <=> - (\i list. all_n 1 list - (\j int. interval_arith ((((-- &2 * f x) * inv (&1 + f x * f x) pow 2) * partial j f x) + (\i list. all_n 1 list + (\j int. interval_arith ((((-- &2 * f x) * inv (&1 + f x * f x) pow 2) * partial j f x) * partial i f x + inv (&1 + f x * f x) * partial2 j i f x) int)) i list`, REWRITE_TAC[]);; -let eval_m_taylor_atn n p_lin p_second taylor1_th = +let eval_m_taylor_atn n p_lin p_second taylor1_th = let domain_th, diff2_f1_th, lin1_th, second1_th = dest_m_taylor_thms n taylor1_th in let f1_tm = (rand o concl) diff2_f1_th in let domain_tm, y_tm, w_tm = dest_m_cell_domain (concl domain_th) in @@ -2577,15 +2583,15 @@ let eval_m_taylor_atn n p_lin p_second taylor1_th = let f1_bound = undisch f1_bound0 in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_atn p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* partial_lemma' *) - let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o + let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) atn_partial_lemma' in - let u_bounds = + let u_bounds = let inv, ( + ), ( * ) = float_interval_inv p_lin, float_interval_add p_lin, float_interval_mul p_lin in inv (one_interval + bounds1_th * bounds1_th) in @@ -2604,15 +2610,15 @@ let eval_m_taylor_atn n p_lin p_second taylor1_th = let dd1 = second_bounded_components n second1_th in - + (* second_lemma', second_lemma'' *) - let u_second_lemma0 = (INST[f1_tm, f_var] o + let u_second_lemma0 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) atn_second_lemma' in let u_second_lemma1 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) atn_second_lemma'' in - let d1_bounds = map (fun i -> + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -2637,14 +2643,14 @@ let eval_m_taylor_atn n p_lin p_second taylor1_th = let dj1 = List.nth d1_bounds (j_int - 1) in (* partial2 *) - let u_th = + let u_th = let ( * ), ( + ) = float_interval_mul p_second, float_interval_add p_second in (d2_th0 * dj1) * di1 + d1_th0 * th1 in let int_tm = rand (concl u_th) in let th0 = INST[j_tm, j_var_num; int_tm, int_var] lemma in EQ_MP th0 u_th in - + let u_th = eval_all_n th1 true u_second in let list_tm = (rand o rator o concl) u_th in let lemma1 = INST[i_tm, i_var_num; list_tm, list_var_real_pair] u_second_lemma1 in @@ -2654,11 +2660,11 @@ let eval_m_taylor_atn n p_lin p_second taylor1_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_ATN' in let eq_th = unary_beta_gen_eq f1_tm x_var atn_tm in @@ -2669,13 +2675,13 @@ let eval_m_taylor_atn n p_lin p_second taylor1_th = (* eval_m_taylor_cos *) (* ----------------------- *) -let cos_partial_lemma' = +let cos_partial_lemma' = prove(`interval_arith (--sin (f y) * partial i f (y:real^N)) int <=> (\i int. interval_arith (--sin (f y) * partial i f y) int) i int`, REWRITE_TAC[]);; -let cos_second_lemma' = +let cos_second_lemma' = prove(`interval_arith (--((cos (f x) * partial j f (x:real^N)) * partial i f x + sin (f x) * partial2 j i f x)) int <=> (\j int. interval_arith (--((cos (f x) * partial j f x) * partial i f x @@ -2683,18 +2689,18 @@ let cos_second_lemma' = REWRITE_TAC[]);; -let cos_second_lemma'' = (PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o NUMERALS_TO_NUM o +let cos_second_lemma'' = (PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o NUMERALS_TO_NUM o PURE_REWRITE_RULE[FLOAT_OF_NUM; min_exp_def] o prove) - (`all_n 1 list + (`all_n 1 list (\j int. interval_arith (--((cos (f x) * partial j f (x:real^N)) * partial i f x + sin (f x) * partial2 j i f x)) int) <=> - (\i list. all_n 1 list + (\i list. all_n 1 list (\j int. interval_arith (--((cos (f x) * partial j f (x:real^N)) * partial i f x + sin (f x) * partial2 j i f x)) int)) i list`, REWRITE_TAC[]);; -let eval_m_taylor_cos n p_lin p_second taylor1_th = +let eval_m_taylor_cos n p_lin p_second taylor1_th = let domain_th, diff2_f1_th, lin1_th, second1_th = dest_m_taylor_thms n taylor1_th in let f1_tm = (rand o concl) diff2_f1_th in let domain_tm, y_tm, w_tm = dest_m_cell_domain (concl domain_th) in @@ -2712,15 +2718,15 @@ let eval_m_taylor_cos n p_lin p_second taylor1_th = let f1_bound = undisch f1_bound0 in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_cos p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* partial_lemma' *) - let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o + let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) cos_partial_lemma' in - let u_bounds = + let u_bounds = let sin, neg = float_interval_sin p_lin, float_interval_neg in neg (sin bounds1_th) in @@ -2738,15 +2744,15 @@ let eval_m_taylor_cos n p_lin p_second taylor1_th = let d_bounds_list = (rand o rator o concl) df_th in let dd1 = second_bounded_components n second1_th in - + (* second_lemma', second_lemma'' *) - let u_second_lemma0 = (INST[f1_tm, f_var] o + let u_second_lemma0 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) cos_second_lemma' in let u_second_lemma1 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) cos_second_lemma'' in - let d1_bounds = map (fun i -> + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -2767,7 +2773,7 @@ let eval_m_taylor_cos n p_lin p_second taylor1_th = let dj1 = List.nth d1_bounds (j_int - 1) in (* partial2 *) - let u_th = + let u_th = let ( * ), ( + ) = float_interval_mul p_second, float_interval_add p_second in let neg = float_interval_neg in neg ((d2_th0 * dj1) * di1 + d1_th0 * th1) in @@ -2775,7 +2781,7 @@ let eval_m_taylor_cos n p_lin p_second taylor1_th = let int_tm = rand (concl u_th) in let th0 = INST[j_tm, j_var_num; int_tm, int_var] lemma in EQ_MP th0 u_th in - + let u_th = eval_all_n th1 true u_second in let list_tm = (rand o rator o concl) u_th in let lemma1 = INST[i_tm, i_var_num; list_tm, list_var_real_pair] u_second_lemma1 in @@ -2785,11 +2791,11 @@ let eval_m_taylor_cos n p_lin p_second taylor1_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_COS' in let eq_th = unary_beta_gen_eq f1_tm x_var cos_tm in @@ -2800,13 +2806,13 @@ let eval_m_taylor_cos n p_lin p_second taylor1_th = (* eval_m_taylor_sin *) (* ----------------------- *) -let sin_partial_lemma' = +let sin_partial_lemma' = prove(`interval_arith (cos (f y) * partial i f (y:real^N)) int <=> (\i int. interval_arith (cos (f y) * partial i f y) int) i int`, REWRITE_TAC[]);; -let sin_second_lemma' = +let sin_second_lemma' = prove(`interval_arith ((--sin (f x) * partial j f (x:real^N)) * partial i f x + cos (f x) * partial2 j i f x) int <=> (\j int. interval_arith ((--sin (f x) * partial j f x) * partial i f x @@ -2814,18 +2820,18 @@ let sin_second_lemma' = REWRITE_TAC[]);; -let sin_second_lemma'' = (PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o NUMERALS_TO_NUM o +let sin_second_lemma'' = (PURE_REWRITE_RULE[float1_eq; float2_eq; num2_eq] o NUMERALS_TO_NUM o PURE_REWRITE_RULE[FLOAT_OF_NUM; min_exp_def] o prove) - (`all_n 1 list + (`all_n 1 list (\j int. interval_arith ((--sin (f x) * partial j f (x:real^N)) * partial i f x + cos (f x) * partial2 j i f x) int) <=> - (\i list. all_n 1 list + (\i list. all_n 1 list (\j int. interval_arith ((--sin (f x) * partial j f (x:real^N)) * partial i f x + cos (f x) * partial2 j i f x) int)) i list`, REWRITE_TAC[]);; -let eval_m_taylor_sin n p_lin p_second taylor1_th = +let eval_m_taylor_sin n p_lin p_second taylor1_th = let domain_th, diff2_f1_th, lin1_th, second1_th = dest_m_taylor_thms n taylor1_th in let f1_tm = (rand o concl) diff2_f1_th in let domain_tm, y_tm, w_tm = dest_m_cell_domain (concl domain_th) in @@ -2843,15 +2849,15 @@ let eval_m_taylor_sin n p_lin p_second taylor1_th = let f1_bound = undisch f1_bound0 in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_sin p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* partial_lemma' *) - let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o + let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) sin_partial_lemma' in - let u_bounds = + let u_bounds = let cos = float_interval_cos p_lin in cos bounds1_th in @@ -2869,15 +2875,15 @@ let eval_m_taylor_sin n p_lin p_second taylor1_th = let d_bounds_list = (rand o rator o concl) df_th in let dd1 = second_bounded_components n second1_th in - + (* second_lemma', second_lemma'' *) - let u_second_lemma0 = (INST[f1_tm, f_var] o + let u_second_lemma0 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) sin_second_lemma' in let u_second_lemma1 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) sin_second_lemma'' in - let d1_bounds = map (fun i -> + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -2899,14 +2905,14 @@ let eval_m_taylor_sin n p_lin p_second taylor1_th = let dj1 = List.nth d1_bounds (j_int - 1) in (* partial2 *) - let u_th = + let u_th = let ( * ), ( + ) = float_interval_mul p_second, float_interval_add p_second in (d2_th0 * dj1) * di1 + d1_th0 * th1 in let int_tm = rand (concl u_th) in let th0 = INST[j_tm, j_var_num; int_tm, int_var] lemma in EQ_MP th0 u_th in - + let u_th = eval_all_n th1 true u_second in let list_tm = (rand o rator o concl) u_th in let lemma1 = INST[i_tm, i_var_num; list_tm, list_var_real_pair] u_second_lemma1 in @@ -2916,11 +2922,11 @@ let eval_m_taylor_sin n p_lin p_second taylor1_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_SIN' in let eq_th = unary_beta_gen_eq f1_tm x_var sin_tm in @@ -2931,24 +2937,24 @@ let eval_m_taylor_sin n p_lin p_second taylor1_th = (* eval_m_taylor_asn *) (* ----------------------- *) -let asn_partial_lemma' = +let asn_partial_lemma' = prove(`interval_arith (inv (sqrt (&1 - f y * f y)) * partial i f (y:real^N)) int <=> (\i int. interval_arith (inv (sqrt (&1 - f y * f y)) * partial i f y) int) i int`, REWRITE_TAC[]);; -let asn_second_lemma' = +let asn_second_lemma' = prove(`interval_arith (((f x / sqrt ((&1 - f x * f x) pow 3)) * partial j f x) * partial i f (x:real^N) + inv (sqrt (&1 - f x * f x)) * partial2 j i f x) int <=> (\j int. interval_arith (((f x / sqrt ((&1 - f x * f x) pow 3)) * partial j f x) * partial i f (x:real^N) + inv (sqrt (&1 - f x * f x)) * partial2 j i f x) int) j int`, REWRITE_TAC[]);; -let asn_second_lemma'' = (PURE_REWRITE_RULE[float1_eq; float2_eq; num3_eq; num2_eq] o NUMERALS_TO_NUM o +let asn_second_lemma'' = (PURE_REWRITE_RULE[float1_eq; float2_eq; num3_eq; num2_eq] o NUMERALS_TO_NUM o PURE_REWRITE_RULE[FLOAT_OF_NUM; min_exp_def] o prove) - (`all_n 1 list + (`all_n 1 list (\j int. interval_arith (((f x / sqrt ((&1 - f x * f x) pow 3)) * partial j f x) * partial i f (x:real^N) + inv (sqrt (&1 - f x * f x)) * partial2 j i f x) int) <=> - (\i list. all_n 1 list + (\i list. all_n 1 list (\j int. interval_arith (((f x / sqrt ((&1 - f x * f x) pow 3)) * partial j f x) * partial i f (x:real^N) + inv (sqrt (&1 - f x * f x)) * partial2 j i f x) int)) i list`, REWRITE_TAC[]);; @@ -2975,15 +2981,15 @@ let eval_m_taylor_asn n p_lin p_second taylor1_th = let cond_th = EQT_ELIM (check_interval_iabs f_bounds_tm one_float) in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_asn p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* partial_lemma' *) - let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o + let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) asn_partial_lemma' in - let u_bounds = + let u_bounds = let inv, sqrt = float_interval_inv p_lin, float_interval_sqrt p_lin in let ( * ), (-) = float_interval_mul p_lin, float_interval_sub p_lin in inv (sqrt (one_interval - bounds1_th * bounds1_th)) in @@ -3003,15 +3009,15 @@ let eval_m_taylor_asn n p_lin p_second taylor1_th = let dd1 = second_bounded_components n second1_th in - + (* second_lemma', second_lemma'' *) - let u_second_lemma0 = (INST[f1_tm, f_var] o + let u_second_lemma0 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) asn_second_lemma' in let u_second_lemma1 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) asn_second_lemma'' in - let d1_bounds = map (fun i -> + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -3036,14 +3042,14 @@ let eval_m_taylor_asn n p_lin p_second taylor1_th = let dj1 = List.nth d1_bounds (j_int - 1) in (* partial2 *) - let u_th = + let u_th = let ( * ), ( + ) = float_interval_mul p_second, float_interval_add p_second in (d2_th0 * dj1) * di1 + d1_th0 * th1 in let int_tm = rand (concl u_th) in let th0 = INST[j_tm, j_var_num; int_tm, int_var] lemma in EQ_MP th0 u_th in - + let u_th = eval_all_n th1 true u_second in let list_tm = (rand o rator o concl) u_th in let lemma1 = INST[i_tm, i_var_num; list_tm, list_var_real_pair] u_second_lemma1 in @@ -3053,12 +3059,12 @@ let eval_m_taylor_asn n p_lin p_second taylor1_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP cond_th o MY_PROVE_HYP f1_bound0 o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_ASN' in let eq_th = unary_beta_gen_eq f1_tm x_var asn_tm in @@ -3069,24 +3075,24 @@ let eval_m_taylor_asn n p_lin p_second taylor1_th = (* eval_m_taylor_acs *) (* ----------------------- *) -let acs_partial_lemma' = +let acs_partial_lemma' = prove(`interval_arith (--inv (sqrt (&1 - f y * f y)) * partial i f (y:real^N)) int <=> (\i int. interval_arith (--inv (sqrt (&1 - f y * f y)) * partial i f y) int) i int`, REWRITE_TAC[]);; -let acs_second_lemma' = +let acs_second_lemma' = prove(`interval_arith ((--(f x / sqrt ((&1 - f x * f x) pow 3)) * partial j f x) * partial i f (x:real^N) - inv (sqrt (&1 - f x * f x)) * partial2 j i f x) int <=> (\j int. interval_arith ((--(f x / sqrt ((&1 - f x * f x) pow 3)) * partial j f x) * partial i f (x:real^N) - inv (sqrt (&1 - f x * f x)) * partial2 j i f x) int) j int`, REWRITE_TAC[]);; -let acs_second_lemma'' = (PURE_REWRITE_RULE[float1_eq; float2_eq; num3_eq; num2_eq] o NUMERALS_TO_NUM o +let acs_second_lemma'' = (PURE_REWRITE_RULE[float1_eq; float2_eq; num3_eq; num2_eq] o NUMERALS_TO_NUM o PURE_REWRITE_RULE[FLOAT_OF_NUM; min_exp_def] o prove) - (`all_n 1 list + (`all_n 1 list (\j int. interval_arith ((--(f x / sqrt ((&1 - f x * f x) pow 3)) * partial j f x) * partial i f (x:real^N) - inv (sqrt (&1 - f x * f x)) * partial2 j i f x) int) <=> - (\i list. all_n 1 list + (\i list. all_n 1 list (\j int. interval_arith ((--(f x / sqrt ((&1 - f x * f x) pow 3)) * partial j f x) * partial i f (x:real^N) - inv (sqrt (&1 - f x * f x)) * partial2 j i f x) int)) i list`, REWRITE_TAC[]);; @@ -3113,15 +3119,15 @@ let eval_m_taylor_acs n p_lin p_second taylor1_th = let cond_th = EQT_ELIM (check_interval_iabs f_bounds_tm one_float) in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_acs p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* partial_lemma' *) - let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o + let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) acs_partial_lemma' in - let u_bounds = + let u_bounds = let inv, sqrt, neg = float_interval_inv p_lin, float_interval_sqrt p_lin, float_interval_neg in let ( * ), (-) = float_interval_mul p_lin, float_interval_sub p_lin in neg (inv (sqrt (one_interval - bounds1_th * bounds1_th))) in @@ -3141,15 +3147,15 @@ let eval_m_taylor_acs n p_lin p_second taylor1_th = let dd1 = second_bounded_components n second1_th in - + (* second_lemma', second_lemma'' *) - let u_second_lemma0 = (INST[f1_tm, f_var] o + let u_second_lemma0 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) acs_second_lemma' in let u_second_lemma1 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) acs_second_lemma'' in - let d1_bounds = map (fun i -> + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -3174,14 +3180,14 @@ let eval_m_taylor_acs n p_lin p_second taylor1_th = let dj1 = List.nth d1_bounds (j_int - 1) in (* partial2 *) - let u_th = + let u_th = let ( * ), ( - ) = float_interval_mul p_second, float_interval_sub p_second in (d2_th0 * dj1) * di1 - d1_th0 * th1 in let int_tm = rand (concl u_th) in let th0 = INST[j_tm, j_var_num; int_tm, int_var] lemma in EQ_MP th0 u_th in - + let u_th = eval_all_n th1 true u_second in let list_tm = (rand o rator o concl) u_th in let lemma1 = INST[i_tm, i_var_num; list_tm, list_var_real_pair] u_second_lemma1 in @@ -3191,12 +3197,12 @@ let eval_m_taylor_acs n p_lin p_second taylor1_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP cond_th o MY_PROVE_HYP f1_bound0 o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_ACS' in let eq_th = unary_beta_gen_eq f1_tm x_var acs_tm in @@ -3207,27 +3213,27 @@ let eval_m_taylor_acs n p_lin p_second taylor1_th = (* eval_m_taylor_matan *) (* ----------------------- *) -let matan_partial_lemma' = +let matan_partial_lemma' = prove(`interval_arith (dmatan (f y) * partial i f (y:real^N)) int <=> (\i int. interval_arith (dmatan (f y) * partial i f y) int) i int`, REWRITE_TAC[]);; -let matan_second_lemma' = - prove(`interval_arith ((ddmatan (f x) * partial j f x) * partial i f (x:real^N) + +let matan_second_lemma' = + prove(`interval_arith ((ddmatan (f x) * partial j f x) * partial i f (x:real^N) + dmatan (f x) * partial2 j i f x) int <=> - (\j int. interval_arith ((ddmatan (f x) * partial j f x) * partial i f (x:real^N) + + (\j int. interval_arith ((ddmatan (f x) * partial j f x) * partial i f (x:real^N) + dmatan (f x) * partial2 j i f x) int) j int`, REWRITE_TAC[]);; -let matan_second_lemma'' = (NUMERALS_TO_NUM o +let matan_second_lemma'' = (NUMERALS_TO_NUM o PURE_REWRITE_RULE[FLOAT_OF_NUM; min_exp_def] o prove) - (`all_n 1 list - (\j int. interval_arith ((ddmatan (f x) * partial j f x) * partial i f (x:real^N) + (`all_n 1 list + (\j int. interval_arith ((ddmatan (f x) * partial j f x) * partial i f (x:real^N) + dmatan (f x) * partial2 j i f x) int) <=> - (\i list. all_n 1 list - (\j int. interval_arith ((ddmatan (f x) * partial j f x) * partial i f (x:real^N) + (\i list. all_n 1 list + (\j int. interval_arith ((ddmatan (f x) * partial j f x) * partial i f (x:real^N) + dmatan (f x) * partial2 j i f x) int)) i list`, REWRITE_TAC[]);; @@ -3256,15 +3262,15 @@ let eval_m_taylor_matan n p_lin p_second taylor1_th = let cond_th = check_interval_gt neg_one_float f_bounds_tm in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_matan p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* partial_lemma' *) - let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o + let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) matan_partial_lemma' in - let u_bounds = + let u_bounds = float_interval_dmatan p_lin bounds1_th in let u_lin th1 = @@ -3282,15 +3288,15 @@ let eval_m_taylor_matan n p_lin p_second taylor1_th = let dd1 = second_bounded_components n second1_th in - + (* second_lemma', second_lemma'' *) - let u_second_lemma0 = (INST[f1_tm, f_var] o + let u_second_lemma0 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) matan_second_lemma' in let u_second_lemma1 = (INST[f1_tm, f_var] o INST_TYPE[n_type_array.(n), nty]) matan_second_lemma'' in - let d1_bounds = map (fun i -> + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -3311,14 +3317,14 @@ let eval_m_taylor_matan n p_lin p_second taylor1_th = let dj1 = List.nth d1_bounds (j_int - 1) in (* partial2 *) - let u_th = + let u_th = let ( * ), ( + ) = float_interval_mul p_second, float_interval_add p_second in (d2_th0 * dj1) * di1 + d1_th0 * th1 in let int_tm = rand (concl u_th) in let th0 = INST[j_tm, j_var_num; int_tm, int_var] lemma in EQ_MP th0 u_th in - + let u_th = eval_all_n th1 true u_second in let list_tm = (rand o rator o concl) u_th in let lemma1 = INST[i_tm, i_var_num; list_tm, list_var_real_pair] u_second_lemma1 in @@ -3328,12 +3334,12 @@ let eval_m_taylor_matan n p_lin p_second taylor1_th = let dd_list = (rand o rator o concl) dd_th0 in let dd_th = GEN x_var (DISCH_ALL dd_th0) in - let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o + let th = (MY_PROVE_HYP dd_th o MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP cond_th o MY_PROVE_HYP f1_bound0 o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o - INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; + MY_PROVE_HYP bounds_th o MY_PROVE_HYP df_th o MY_PROVE_HYP domain_th o + INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; + bounds_tm, bounds_var; d_bounds_list, d_bounds_list_var; dd_list, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) MK_M_TAYLOR_MATAN' in let eq_th = unary_beta_gen_eq f1_tm x_var matan_tm in @@ -3344,8 +3350,8 @@ end;; (* -needs "informal/informal_taylor.hl";; -needs "verifier/m_verifier_build.hl";; +needs "Formal_ineqs/informal/informal_taylor.hl";; +needs "Formal_ineqs/verifier/m_verifier_build.hl";; open M_verifier_build;; #install_printer Informal_float.print_ifloat;; @@ -3442,15 +3448,15 @@ let f_bounds_tm = (rand o concl) f1_bound;; let cond_th = check_interval_gt neg_one_float f_bounds_tm;; let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th;; - + let bounds_th = float_interval_matan p_lin bounds1_th;; let bounds_tm = (rand o concl) bounds_th;; (* partial_lemma' *) -let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o +let u_lemma0 = (INST[f1_tm, f_var; y_tm, y_var] o INST_TYPE[n_type_array.(n), nty]) matan_partial_lemma';; -let u_bounds = +let u_bounds = float_interval_dmatan p_lin bounds1_th;; let u_lin th1 = @@ -3462,7 +3468,7 @@ let u_lin th1 = let i_tm = (rand o rator o rator o lhand) (concl th1) in let th0 = INST[i_tm, i_var_num; int_tm, int_var] u_lemma0 in EQ_MP th0 u_th;; - + let df_th = eval_all_n df1_th true u_lin;; let d_bounds_list = (rand o rator o concl) df_th;; *) diff --git a/Formal_ineqs/taylor/m_taylor_arith2.hl b/Formal_ineqs/taylor/m_taylor_arith2.hl index 162e50d4..ae12a1d6 100644 --- a/Formal_ineqs/taylor/m_taylor_arith2.hl +++ b/Formal_ineqs/taylor/m_taylor_arith2.hl @@ -1,10 +1,16 @@ -(* =========================================================== *) -(* Formal arithmetic of taylor intervals 2 *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) -needs "taylor/m_taylor_arith.hl";; +(* -------------------------------------------------------------------------- *) +(* Formal arithmetic of taylor intervals 2 *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/taylor/m_taylor_arith.hl";; module M_taylor_arith2 = struct @@ -63,7 +69,7 @@ let gen_taylor_arith_thm arith_th final_rule n = let dd_bounds_list = mk_list (ddfs, type_of (hd ddfs)) in let th1 = INST[d_bounds_list, d_bounds_list_var; dd_bounds_list, dd_bounds_list_var] th0 in let th2 = (CONV_RULE NUM_REDUCE_CONV o REWRITE_RULE[all_n]) th1 in - (UNDISCH_ALL o final_rule o REWRITE_RULE[GSYM CONJ_ASSOC] o + (UNDISCH_ALL o final_rule o REWRITE_RULE[GSYM CONJ_ASSOC] o NUMERALS_TO_NUM o PURE_REWRITE_RULE[FLOAT_OF_NUM; min_exp_def]) th2;; let gen_add_thm = gen_taylor_arith_thm MK_M_TAYLOR_ADD' (CONV_RULE ALL_CONV);; @@ -74,18 +80,18 @@ let gen_inv_thm = gen_taylor_arith_thm MK_M_TAYLOR_INV' (REWRITE_RULE[float2_eq] let gen_sqrt_thm = gen_taylor_arith_thm MK_M_TAYLOR_SQRT' (REWRITE_RULE[float2_eq]);; let gen_exp_thm = gen_taylor_arith_thm MK_M_TAYLOR_EXP' (CONV_RULE ALL_CONV);; -let gen_log_thm = +let gen_log_thm = let pow2_th = (SYM o REWRITE_CONV[SYM num2_eq]) `x pow 2` in gen_taylor_arith_thm MK_M_TAYLOR_LOG' (REWRITE_RULE[pow2_th]);; -let gen_atn_thm = +let gen_atn_thm = let pow2_th = (SYM o REWRITE_CONV[SYM num2_eq]) `x pow 2` in gen_taylor_arith_thm MK_M_TAYLOR_ATN' (REWRITE_RULE[float2_eq; float1_eq; pow2_th]);; -let gen_cos_thm = +let gen_cos_thm = gen_taylor_arith_thm MK_M_TAYLOR_COS' (CONV_RULE ALL_CONV);; -let gen_sin_thm = +let gen_sin_thm = gen_taylor_arith_thm MK_M_TAYLOR_SIN' (CONV_RULE ALL_CONV);; let gen_acs_thm = @@ -179,10 +185,10 @@ let eval_m_taylor_add2 n p_lin p_second taylor1_th taylor2_th = let dds = map (map (rand o concl)) dd_ths in let inst_list = union (zip dfs df_vars) (zip (List.flatten dds) (List.flatten dd_vars)) in - let th = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP diff2_f2_th o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o + let th = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP diff2_f2_th o + MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP df_th o MY_PROVE_HYP dd_th o - INST([f1_tm, f_var; f2_tm, g_var; + INST([f1_tm, f_var; f2_tm, g_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) add_ths_array.(n) in let eq_th = binary_beta_gen_eq f1_tm f2_tm x_var add_op_real in @@ -237,10 +243,10 @@ let eval_m_taylor_sub2 n p_lin p_second taylor1_th taylor2_th = let dds = map (map (rand o concl)) dd_ths in let inst_list = union (zip dfs df_vars) (zip (List.flatten dds) (List.flatten dd_vars)) in - let th = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP diff2_f2_th o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o + let th = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP diff2_f2_th o + MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP df_th o MY_PROVE_HYP dd_th o - INST([f1_tm, f_var; f2_tm, g_var; + INST([f1_tm, f_var; f2_tm, g_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) sub_ths_array.(n) in let eq_th = binary_beta_gen_eq f1_tm f2_tm x_var sub_op_real in @@ -316,10 +322,10 @@ let eval_m_taylor_mul2 n p_lin p_second taylor1_th taylor2_th = let dds = map (map (rand o concl)) dd_ths in let inst_list = union (zip dfs df_vars) (zip (List.flatten dds) (List.flatten dd_vars)) in - let th = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP diff2_f2_th o - MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o + let th = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP diff2_f2_th o + MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP df_th o MY_PROVE_HYP dd_th o - INST([f1_tm, f_var; f2_tm, g_var; + INST([f1_tm, f_var; f2_tm, g_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) mul_ths_array.(n) in let eq_th = binary_beta_gen_eq f1_tm f2_tm x_var mul_op_real in @@ -367,9 +373,9 @@ let eval_m_taylor_neg2_raw norm_flag n taylor1_th = let dds = map (map (rand o concl)) dd_ths in let inst_list = union (zip dfs df_vars) (zip (List.flatten dds) (List.flatten dd_vars)) in - let th = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP bounds_th o + let th = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP df_th o MY_PROVE_HYP dd_th o - INST([f1_tm, f_var; + INST([f1_tm, f_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) neg_ths_array.(n) in if norm_flag then @@ -385,11 +391,11 @@ let eval_m_taylor_neg2 = eval_m_taylor_neg2_raw true;; (* abs *) (* ----------------------- *) -let eval_m_taylor_abs2 n p_second taylor1_th = +let eval_m_taylor_abs2 n p_second taylor1_th = let f1_bound0 = eval_m_taylor_bound n p_second taylor1_th in let f_bounds_tm = (rand o rand o snd o dest_forall o concl) f1_bound0 in - let cond_th, taylor_th, th0, - (f1_tm, domain_tm, y_tm, w_tm, bounds_tm, d_tm, dd_tm) = + let cond_th, taylor_th, th0, + (f1_tm, domain_tm, y_tm, w_tm, bounds_tm, d_tm, dd_tm) = try let cond_th = check_interval_pos f_bounds_tm in cond_th, taylor1_th, MK_M_TAYLOR_ABS_POS', dest_m_taylor (concl taylor1_th) @@ -407,12 +413,12 @@ let eval_m_taylor_abs2 n p_second taylor1_th = w_var = mk_var ("w", ty) and f_var = mk_var ("f", type_of f1_tm) and domain_var = mk_var ("domain", type_of domain_tm) in - let th1 = (MY_PROVE_HYP f1_bound0 o + let th1 = (MY_PROVE_HYP f1_bound0 o MY_PROVE_HYP cond_th o MY_PROVE_HYP taylor_th o - INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; + INST[f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; - bounds_tm, bounds_var; d_tm, d_bounds_list_var; + bounds_tm, bounds_var; d_tm, d_bounds_list_var; dd_tm, dd_bounds_list_var] o INST_TYPE[n_type_array.(n), nty]) th0 in let eq_th = unary_beta_gen_eq f1_tm x_var abs_tm in @@ -445,19 +451,19 @@ let eval_m_taylor_inv2 n p_lin p_second taylor1_th = let cond_th = check_interval_not_zero f_bounds_tm in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_inv p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* first partials *) - let u_bounds = + let u_bounds = let neg, inv, ( * ) = float_interval_neg, float_interval_inv p_lin, float_interval_mul p_lin in neg (inv (bounds1_th * bounds1_th)) in let df1_ths' = all_n_components2 n df1_th in let df1_ths = map MY_BETA_RULE df1_ths' in - let df_ths = + let df_ths = let ( * ) = float_interval_mul p_lin in map (fun th1 -> u_bounds * th1) df1_ths in @@ -467,8 +473,8 @@ let eval_m_taylor_inv2 n p_lin p_second taylor1_th = let dd_ths = let dd1 = all_n_components2 n (second_bounded_components n second1_th) in map2 (fun i -> map MY_BETA_RULE o all_n_components2 i o MY_BETA_RULE) (1--n) dd1 in - - let d1_bounds = map (fun i -> + + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -476,7 +482,7 @@ let eval_m_taylor_inv2 n p_lin p_second taylor1_th = let d1_th0, d2_th0 = let inv, ( * ) = float_interval_inv p_second, float_interval_mul p_second in let ff = f1_bound * f1_bound in - inv ff, + inv ff, two_interval * inv (f1_bound * ff) in let dd_ths = @@ -499,7 +505,7 @@ let eval_m_taylor_inv2 n p_lin p_second taylor1_th = let th1 = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP cond_th o MY_PROVE_HYP f1_bound0 o MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP dd_th o MY_PROVE_HYP df_th o - INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; + INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) inv_ths_array.(n) in let eq_th = unary_beta_gen_eq f1_tm x_var inv_op_real in @@ -533,19 +539,19 @@ let eval_m_taylor_sqrt2 n p_lin p_second taylor1_th = let cond_th = check_interval_pos f_bounds_tm in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_sqrt p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* first partials *) - let u_bounds = + let u_bounds = let inv, ( * ) = float_interval_inv p_lin, float_interval_mul p_lin in inv (two_interval * bounds_th) in let df1_ths' = all_n_components2 n df1_th in let df1_ths = map MY_BETA_RULE df1_ths' in - let df_ths = + let df_ths = let ( * ) = float_interval_mul p_lin in map (fun th1 -> u_bounds * th1) df1_ths in @@ -555,8 +561,8 @@ let eval_m_taylor_sqrt2 n p_lin p_second taylor1_th = let dd_ths = let dd1 = all_n_components2 n (second_bounded_components n second1_th) in map2 (fun i -> map MY_BETA_RULE o all_n_components2 i o MY_BETA_RULE) (1--n) dd1 in - - let d1_bounds = map (fun i -> + + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -587,7 +593,7 @@ let eval_m_taylor_sqrt2 n p_lin p_second taylor1_th = let th1 = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP cond_th o MY_PROVE_HYP f1_bound0 o MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP dd_th o MY_PROVE_HYP df_th o - INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; + INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) sqrt_ths_array.(n) in let eq_th = unary_beta_gen_eq f1_tm x_var sqrt_tm in @@ -616,19 +622,19 @@ let eval_m_taylor_exp2 n p_lin p_second taylor1_th = let f_bounds_tm = (rand o concl) f1_bound in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_exp p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* first partials *) - let u_bounds = + let u_bounds = let exp = float_interval_exp p_lin in exp bounds1_th in let df1_ths' = all_n_components2 n df1_th in let df1_ths = map MY_BETA_RULE df1_ths' in - let df_ths = + let df_ths = let ( * ) = float_interval_mul p_lin in map (fun th1 -> u_bounds * th1) df1_ths in @@ -638,8 +644,8 @@ let eval_m_taylor_exp2 n p_lin p_second taylor1_th = let dd_ths = let dd1 = all_n_components2 n (second_bounded_components n second1_th) in map2 (fun i -> map MY_BETA_RULE o all_n_components2 i o MY_BETA_RULE) (1--n) dd1 in - - let d1_bounds = map (fun i -> + + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -669,7 +675,7 @@ let eval_m_taylor_exp2 n p_lin p_second taylor1_th = let th1 = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP f1_bound0 o MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP dd_th o MY_PROVE_HYP df_th o - INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; + INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) exp_ths_array.(n) in let eq_th = unary_beta_gen_eq f1_tm x_var exp_tm in @@ -701,19 +707,19 @@ let eval_m_taylor_log2 n p_lin p_second taylor1_th = let cond_th = check_interval_pos f_bounds_tm in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_log p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* first partials *) - let u_bounds = + let u_bounds = let inv = float_interval_inv p_lin in inv (bounds1_th) in let df1_ths' = all_n_components2 n df1_th in let df1_ths = map MY_BETA_RULE df1_ths' in - let df_ths = + let df_ths = let ( * ) = float_interval_mul p_lin in map (fun th1 -> u_bounds * th1) df1_ths in @@ -723,14 +729,14 @@ let eval_m_taylor_log2 n p_lin p_second taylor1_th = let dd_ths = let dd1 = all_n_components2 n (second_bounded_components n second1_th) in map2 (fun i -> map MY_BETA_RULE o all_n_components2 i o MY_BETA_RULE) (1--n) dd1 in - - let d1_bounds = map (fun i -> + + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in (* u'(f x), u''(f x) *) let d1_th0, d2_th0 = - let neg, pow2, inv = + let neg, pow2, inv = float_interval_neg, float_interval_pow p_second 2, float_interval_inv p_second in let inv_f = inv f1_bound in inv_f, neg (pow2 inv_f) in @@ -755,7 +761,7 @@ let eval_m_taylor_log2 n p_lin p_second taylor1_th = let th1 = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP cond_th o MY_PROVE_HYP f1_bound0 o MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP dd_th o MY_PROVE_HYP df_th o - INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; + INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) log_ths_array.(n) in let eq_th = unary_beta_gen_eq f1_tm x_var log_tm in @@ -784,20 +790,20 @@ let eval_m_taylor_atn2 n p_lin p_second taylor1_th = let f_bounds_tm = (rand o concl) f1_bound in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_atn p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* first partials *) - let u_bounds = - let inv, ( + ), ( * ) = float_interval_inv p_lin, + let u_bounds = + let inv, ( + ), ( * ) = float_interval_inv p_lin, float_interval_add p_lin, float_interval_mul p_lin in inv (one_interval + bounds1_th * bounds1_th) in let df1_ths' = all_n_components2 n df1_th in let df1_ths = map MY_BETA_RULE df1_ths' in - let df_ths = + let df_ths = let ( * ) = float_interval_mul p_lin in map (fun th1 -> u_bounds * th1) df1_ths in @@ -807,14 +813,14 @@ let eval_m_taylor_atn2 n p_lin p_second taylor1_th = let dd_ths = let dd1 = all_n_components2 n (second_bounded_components n second1_th) in map2 (fun i -> map MY_BETA_RULE o all_n_components2 i o MY_BETA_RULE) (1--n) dd1 in - - let d1_bounds = map (fun i -> + + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in (* u'(f x), u''(f x) *) let d1_th0, d2_th0 = - let neg, inv, ( + ), ( * ), pow2 = float_interval_neg, float_interval_inv p_second, + let neg, inv, ( + ), ( * ), pow2 = float_interval_neg, float_interval_inv p_second, float_interval_add p_second, float_interval_mul p_second, float_interval_pow p_second 2 in let inv_one_ff = inv (one_interval + f1_bound * f1_bound) in inv_one_ff, (neg_two_interval * f1_bound) * pow2 inv_one_ff in @@ -839,7 +845,7 @@ let eval_m_taylor_atn2 n p_lin p_second taylor1_th = let th1 = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP f1_bound0 o MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP dd_th o MY_PROVE_HYP df_th o - INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; + INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) atn_ths_array.(n) in let eq_th = unary_beta_gen_eq f1_tm x_var atn_tm in @@ -869,19 +875,19 @@ let eval_m_taylor_cos2 n p_lin p_second taylor1_th = let f_bounds_tm = (rand o concl) f1_bound in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_cos p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* first partials *) - let u_bounds = + let u_bounds = let sin, neg = float_interval_sin p_lin, float_interval_neg in neg (sin bounds1_th) in let df1_ths' = all_n_components2 n df1_th in let df1_ths = map MY_BETA_RULE df1_ths' in - let df_ths = + let df_ths = let ( * ) = float_interval_mul p_lin in map (fun th1 -> u_bounds * th1) df1_ths in @@ -891,8 +897,8 @@ let eval_m_taylor_cos2 n p_lin p_second taylor1_th = let dd_ths = let dd1 = all_n_components2 n (second_bounded_components n second1_th) in map2 (fun i -> map MY_BETA_RULE o all_n_components2 i o MY_BETA_RULE) (1--n) dd1 in - - let d1_bounds = map (fun i -> + + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -922,7 +928,7 @@ let eval_m_taylor_cos2 n p_lin p_second taylor1_th = let th1 = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP f1_bound0 o MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP dd_th o MY_PROVE_HYP df_th o - INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; + INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) cos_ths_array.(n) in let eq_th = unary_beta_gen_eq f1_tm x_var cos_tm in @@ -952,19 +958,19 @@ let eval_m_taylor_sin2 n p_lin p_second taylor1_th = let f_bounds_tm = (rand o concl) f1_bound in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_sin p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* first partials *) - let u_bounds = + let u_bounds = let cos = float_interval_cos p_lin in cos bounds1_th in let df1_ths' = all_n_components2 n df1_th in let df1_ths = map MY_BETA_RULE df1_ths' in - let df_ths = + let df_ths = let ( * ) = float_interval_mul p_lin in map (fun th1 -> u_bounds * th1) df1_ths in @@ -974,8 +980,8 @@ let eval_m_taylor_sin2 n p_lin p_second taylor1_th = let dd_ths = let dd1 = all_n_components2 n (second_bounded_components n second1_th) in map2 (fun i -> map MY_BETA_RULE o all_n_components2 i o MY_BETA_RULE) (1--n) dd1 in - - let d1_bounds = map (fun i -> + + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -1005,7 +1011,7 @@ let eval_m_taylor_sin2 n p_lin p_second taylor1_th = let th1 = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP f1_bound0 o MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP dd_th o MY_PROVE_HYP df_th o - INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; + INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) sin_ths_array.(n) in let eq_th = unary_beta_gen_eq f1_tm x_var sin_tm in @@ -1038,12 +1044,12 @@ let eval_m_taylor_asn2 n p_lin p_second taylor1_th = let cond_th = EQT_ELIM (check_interval_iabs f_bounds_tm one_float) in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_asn p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* first partials *) - let u_bounds = + let u_bounds = let inv, sqrt = float_interval_inv p_lin, float_interval_sqrt p_lin in let ( * ), ( - ) = float_interval_mul p_lin, float_interval_sub p_lin in inv (sqrt (one_interval - bounds1_th * bounds1_th)) in @@ -1051,7 +1057,7 @@ let eval_m_taylor_asn2 n p_lin p_second taylor1_th = let df1_ths' = all_n_components2 n df1_th in let df1_ths = map MY_BETA_RULE df1_ths' in - let df_ths = + let df_ths = let ( * ) = float_interval_mul p_lin in map (fun th1 -> u_bounds * th1) df1_ths in @@ -1061,8 +1067,8 @@ let eval_m_taylor_asn2 n p_lin p_second taylor1_th = let dd_ths = let dd1 = all_n_components2 n (second_bounded_components n second1_th) in map2 (fun i -> map MY_BETA_RULE o all_n_components2 i o MY_BETA_RULE) (1--n) dd1 in - - let d1_bounds = map (fun i -> + + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -1094,7 +1100,7 @@ let eval_m_taylor_asn2 n p_lin p_second taylor1_th = let th1 = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP cond_th o MY_PROVE_HYP f1_bound0 o MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP dd_th o MY_PROVE_HYP df_th o - INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; + INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) asn_ths_array.(n) in let eq_th = unary_beta_gen_eq f1_tm x_var asn_tm in @@ -1128,12 +1134,12 @@ let eval_m_taylor_acs2 n p_lin p_second taylor1_th = let cond_th = EQT_ELIM (check_interval_iabs f_bounds_tm one_float) in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_acs p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* first partials *) - let u_bounds = + let u_bounds = let inv, sqrt, neg = float_interval_inv p_lin, float_interval_sqrt p_lin, float_interval_neg in let ( * ), ( - ) = float_interval_mul p_lin, float_interval_sub p_lin in neg (inv (sqrt (one_interval - bounds1_th * bounds1_th))) in @@ -1141,7 +1147,7 @@ let eval_m_taylor_acs2 n p_lin p_second taylor1_th = let df1_ths' = all_n_components2 n df1_th in let df1_ths = map MY_BETA_RULE df1_ths' in - let df_ths = + let df_ths = let ( * ) = float_interval_mul p_lin in map (fun th1 -> u_bounds * th1) df1_ths in @@ -1151,8 +1157,8 @@ let eval_m_taylor_acs2 n p_lin p_second taylor1_th = let dd_ths = let dd1 = all_n_components2 n (second_bounded_components n second1_th) in map2 (fun i -> map MY_BETA_RULE o all_n_components2 i o MY_BETA_RULE) (1--n) dd1 in - - let d1_bounds = map (fun i -> + + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -1184,7 +1190,7 @@ let eval_m_taylor_acs2 n p_lin p_second taylor1_th = let th1 = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP cond_th o MY_PROVE_HYP f1_bound0 o MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP dd_th o MY_PROVE_HYP df_th o - INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; + INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) acs_ths_array.(n) in let eq_th = unary_beta_gen_eq f1_tm x_var acs_tm in @@ -1217,18 +1223,18 @@ let eval_m_taylor_matan2 n p_lin p_second taylor1_th = let cond_th = check_interval_gt neg_one_float f_bounds_tm in let _, bounds1_th, df1_th = m_lin_approx_components n lin1_th in - + let bounds_th = float_interval_matan p_lin bounds1_th in let bounds_tm = (rand o concl) bounds_th in (* first partials *) - let u_bounds = + let u_bounds = float_interval_dmatan p_lin bounds1_th in let df1_ths' = all_n_components2 n df1_th in let df1_ths = map MY_BETA_RULE df1_ths' in - let df_ths = + let df_ths = let ( * ) = float_interval_mul p_lin in map (fun th1 -> u_bounds * th1) df1_ths in @@ -1238,8 +1244,8 @@ let eval_m_taylor_matan2 n p_lin p_second taylor1_th = let dd_ths = let dd1 = all_n_components2 n (second_bounded_components n second1_th) in map2 (fun i -> map MY_BETA_RULE o all_n_components2 i o MY_BETA_RULE) (1--n) dd1 in - - let d1_bounds = map (fun i -> + + let d1_bounds = map (fun i -> let th0 = eval_m_taylor_partial_bound n p_second i taylor1_th in undisch th0) (1--n) in @@ -1268,7 +1274,7 @@ let eval_m_taylor_matan2 n p_lin p_second taylor1_th = let th1 = (MY_PROVE_HYP diff2_f1_th o MY_PROVE_HYP cond_th o MY_PROVE_HYP f1_bound0 o MY_PROVE_HYP bounds_th o MY_PROVE_HYP domain_th o MY_PROVE_HYP dd_th o MY_PROVE_HYP df_th o - INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; + INST([f1_tm, f_var; f_bounds_tm, f_bounds_var; domain_tm, domain_var; y_tm, y_var; w_tm, w_var; bounds_tm, bounds_var] @ inst_list)) matan_ths_array.(n) in let eq_th = unary_beta_gen_eq f1_tm x_var matan_tm in @@ -1277,8 +1283,8 @@ let eval_m_taylor_matan2 n p_lin p_second taylor1_th = end;; (* -needs "informal/informal_taylor.hl";; -needs "verifier/m_verifier_build.hl";; +needs "Formal_ineqs/informal/informal_taylor.hl";; +needs "Formal_ineqs/verifier/m_verifier_build.hl";; open M_verifier_build;; #install_printer Informal_float.print_ifloat;; diff --git a/Formal_ineqs/taylor/theory/multivariate_taylor-compiled.hl b/Formal_ineqs/taylor/theory/multivariate_taylor-compiled.hl index 409833c6..308d8597 100644 --- a/Formal_ineqs/taylor/theory/multivariate_taylor-compiled.hl +++ b/Formal_ineqs/taylor/theory/multivariate_taylor-compiled.hl @@ -1,6 +1,6 @@ -needs "lib/ssrbool.hl";; -needs "lib/ssrnat.hl";; -needs "taylor/theory/taylor_interval-compiled.hl";; +needs "Formal_ineqs/lib/ssrbool.hl";; +needs "Formal_ineqs/lib/ssrnat.hl";; +needs "Formal_ineqs/taylor/theory/taylor_interval-compiled.hl";; (* Module Multivariate_taylor*) module Multivariate_taylor = struct @@ -19,7 +19,7 @@ let IMAGE_DELETE_INJ_COMPAT = prove SET_TAC[]);; let partial = new_definition `partial i f x = derivative (f o (\t. (x:real^N) + t % basis i)) (&0)`;; let all_n = define `(all_n n [] s <=> T) /\ (all_n n (CONS h t) s <=> s n h /\ all_n (SUC n) t s)`;; -let m_lin_approx = new_definition `m_lin_approx (f:real^N->real) x f_bounds df_bounds_list <=> +let m_lin_approx = new_definition `m_lin_approx (f:real^N->real) x f_bounds df_bounds_list <=> (lift o f) differentiable at x /\ interval_arith (f x) f_bounds /\ all_n 1 df_bounds_list (\i int. interval_arith (partial i f x) int)`;; @@ -126,7 +126,7 @@ let frechet_scale = Sections.section_proof ["c"] (* Lemma frechet_add *) let frechet_add = Sections.section_proof [] -`frechet_derivative (\x. f x + g x) (at x) = +`frechet_derivative (\x. f x + g x) (at x) = (\y. frechet_derivative f (at x) y + frechet_derivative g (at x) y)` [ ((((use_arg_then2 ("EQ_SYM_EQ", [EQ_SYM_EQ]))(thm_tac (new_rewrite [] [])))) THEN ((use_arg_then2 ("FRECHET_DERIVATIVE_AT", [FRECHET_DERIVATIVE_AT])) (thm_tac apply_tac)) THEN ((use_arg_then2 ("HAS_DERIVATIVE_ADD", [HAS_DERIVATIVE_ADD])) (thm_tac apply_tac))); @@ -135,7 +135,7 @@ let frechet_add = Sections.section_proof [] (* Lemma frechet_sub *) let frechet_sub = Sections.section_proof [] -`frechet_derivative (\x. f x - g x) (at x) = +`frechet_derivative (\x. f x - g x) (at x) = (\y. frechet_derivative f (at x) y - frechet_derivative g (at x) y)` [ ((((use_arg_then2 ("EQ_SYM_EQ", [EQ_SYM_EQ]))(thm_tac (new_rewrite [] [])))) THEN ((use_arg_then2 ("FRECHET_DERIVATIVE_AT", [FRECHET_DERIVATIVE_AT])) (thm_tac apply_tac)) THEN ((use_arg_then2 ("HAS_DERIVATIVE_SUB", [HAS_DERIVATIVE_SUB])) (thm_tac apply_tac))); @@ -223,7 +223,7 @@ let has_derivative_x12 = Sections.section_proof ["y"] (* Lemma lambda_eq_vsum *) let lambda_eq_vsum = Sections.section_proof ["f"] -`(\x:A. lambda i. f i x) = +`(\x:A. lambda i. f i x) = (\x. vsum (1..dimindex (:N)) (\i. f i x % (basis i:real^N)))` [ (((((use_arg_then2 ("eq_ext", [eq_ext]))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (simp_tac)) THEN (move ["x"])); @@ -283,7 +283,7 @@ let has_derivative_mul = Sections.section_proof ["f";"g";"f'";"g'";"y"] ((fun arg_tac -> arg_tac (Arg_term (`lift o (\x. f x * g x) = (lift o (\p. p$1 * p$2)) o (\x. vector [f x; g x]:real^2)`))) (term_tac (have_gen_tac [](((conv_thm_tac DISCH_THEN)(thm_tac (new_rewrite [] []))))))); ((((((use_arg_then2 ("eq_ext", [eq_ext]))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("o_THM", [o_THM]))(thm_tac (new_rewrite [] []))))) THEN (simp_tac)) THEN (move ["x"])) THEN (repeat_tactic 1 9 (((fun arg_tac -> arg_tac (Arg_theorem (GEN_ALL VECTOR_2)))(thm_tac (new_rewrite [] []))))) THEN (done_tac)); ((fun arg_tac -> arg_tac (Arg_term (`vector [f y; g y]:real^2`))) (term_tac (set_tac "q"))); - ((fun arg_tac -> arg_tac (Arg_term (`lift o (\x. f' x * g y + f y * g' x) = + ((fun arg_tac -> arg_tac (Arg_term (`lift o (\x. f' x * g y + f y * g' x) = (lift o (\x:real^2. q$2 * x$1 + q$1 * x$2)) o (\x. vector [f' x; g' x])`))) (term_tac (have_gen_tac [](((conv_thm_tac DISCH_THEN)(thm_tac (new_rewrite [] []))))))); (((((use_arg_then2 ("eq_ext", [eq_ext]))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("o_THM", [o_THM]))(thm_tac (new_rewrite [] []))))) THEN (simp_tac) THEN (((use_arg_then2 ("q_def", []))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((fun arg_tac -> arg_tac (Arg_theorem (GEN_ALL VECTOR_2)))(thm_tac (new_rewrite [] []))))) THEN (((use_arg_then2 ("REAL_MUL_SYM", [REAL_MUL_SYM]))(thm_tac (new_rewrite [] []))))) THEN (done_tac)); ((((use_arg_then2 ("DIFF_CHAIN_AT", [DIFF_CHAIN_AT])) (thm_tac apply_tac)) THEN (simp_tac)) THEN ((((use_arg_then2 ("q_def", []))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("has_derivative_x12", [has_derivative_x12]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("has_derivative_vector2", [has_derivative_vector2]))(thm_tac (new_rewrite [] []))))) THEN (done_tac)); @@ -313,7 +313,7 @@ let frechet_mul = Sections.section_proof ["f";"g";"y"] `lift o f differentiable at y ==> lift o g differentiable at y ==> frechet_derivative (lift o (\x. f x * g x)) (at y) = - (\x. g y % frechet_derivative (lift o f) (at y) x + + (\x. g y % frechet_derivative (lift o f) (at y) x + f y % frechet_derivative (lift o g) (at y) x)` [ (((repeat_tactic 1 9 (((use_arg_then2 ("FRECHET_DERIVATIVE_WORKS", [FRECHET_DERIVATIVE_WORKS]))(thm_tac (new_rewrite [] []))))) THEN (((use_arg_then2 ("f_eq_lift_drop", [f_eq_lift_drop]))(thm_tac (new_rewrite [] [(`frechet_derivative _1 _2`)]))))) THEN (move ["df"])); @@ -373,7 +373,7 @@ Sections.begin_section "Partial";; let real_derivative_compose_frechet = Sections.section_proof ["f";"h";"t"] `(lift o f) differentiable at (h t) ==> (h o drop) differentiable at (lift t) ==> - ((f o h) has_real_derivative (drop o (frechet_derivative (lift o f) (at (h t)) o + ((f o h) has_real_derivative (drop o (frechet_derivative (lift o f) (at (h t)) o frechet_derivative (h o drop) (at (lift t))) o lift) (&1)) (atreal t)` [ (BETA_TAC THEN (move ["diff_f"]) THEN (move ["diff_h"])); @@ -397,7 +397,7 @@ let real_derivative_compose_frechet = Sections.section_proof ["f";"h";"t"] let real_derivative_compose_jacobian = Sections.section_proof ["f";"h";"t"] `(lift o f) differentiable at (h t) ==> (h o drop) differentiable at (lift t) ==> - ((f o h) has_real_derivative (jacobian (lift o f) (at (h t)) ** + ((f o h) has_real_derivative (jacobian (lift o f) (at (h t)) ** jacobian (h o drop) (at (lift t)))$1$1) (atreal t)` [ (BETA_TAC THEN (move ["df"]) THEN (move ["dh"])); @@ -409,7 +409,7 @@ let real_derivative_compose_jacobian = Sections.section_proof ["f";"h";"t"] (* Lemma diff_imp_real_diff *) let diff_imp_real_diff = Sections.section_proof ["f";"h";"t"] -`(lift o f) differentiable at (h t) ==> +`(lift o f) differentiable at (h t) ==> (h o drop) differentiable at (lift t) ==> (f o h) real_differentiable atreal t` [ @@ -442,7 +442,7 @@ let frechet_direction = Sections.section_proof ["y";"e";"t"] (* Lemma real_dir_derivative_frechet *) let real_dir_derivative_frechet = Sections.section_proof ["f";"y";"e";"t"] `(lift o f) differentiable at (y + t % e) ==> - ((f o (\t. y + t % e)) has_real_derivative + ((f o (\t. y + t % e)) has_real_derivative (drop (frechet_derivative (lift o f) (at (y + t % e)) e))) (atreal t)` [ (BETA_TAC THEN (move ["df"])); @@ -453,7 +453,7 @@ let real_dir_derivative_frechet = Sections.section_proof ["f";"y";"e";"t"] (* Lemma real_dir_derivative_jacobian *) let real_dir_derivative_jacobian = Sections.section_proof ["f";"y";"e";"t"] `(lift o f) differentiable at (y + t % e) ==> - ((f o (\t. y + t % e)) has_real_derivative + ((f o (\t. y + t % e)) has_real_derivative drop (jacobian (lift o f) (at (y + t % e)) ** e)) (atreal t)` [ (BETA_TAC THEN (move ["df"])); @@ -914,7 +914,7 @@ let real_taylor2_bound = Sections.section_proof ["f";"dd_bound"] (* Lemma real_taylor1_bound *) let real_taylor1_bound = Sections.section_proof ["f";"d_bound"] -`(!t. interval_arith t (&0, &1) ==> f real_differentiable atreal t /\ +`(!t. interval_arith t (&0, &1) ==> f real_differentiable atreal t /\ abs (derivative f t) <= d_bound) ==> abs (f (&1) - f (&0)) <= d_bound` [ @@ -1040,7 +1040,7 @@ let open_contains_open_interval = Sections.section_proof ["e";"s";"x"] (* Lemma diff2_dir *) let diff2_dir = Sections.section_proof ["f";"x";"e";"t"] -`diff2 f (x + t % e:real^N) ==> +`diff2 f (x + t % e:real^N) ==> nth_diff_strong 2 (f o (\t. x + t % e)) t` [ (((((use_arg_then2 ("diff2_eq_diff2_on_open", [diff2_eq_diff2_on_open]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("nth_diff_strong2_eq", [nth_diff_strong2_eq]))(thm_tac (new_rewrite [] []))))) THEN ALL_TAC THEN (case THEN (move ["s"])) THEN (case THEN (move ["open_s"])) THEN (case THEN (move ["xs"])) THEN (move ["df"])); @@ -1072,7 +1072,7 @@ let diff2_dir = Sections.section_proof ["f";"x";"e";"t"] let diff2_dir_derivative2 = Sections.section_proof ["f";"x";"e";"t"] `diff2 f (x + t % e:real^N) ==> nth_derivative 2 (f o (\t. x + t % e)) t = - sum (1..dimindex (:N)) (\i. sum (1..dimindex (:N)) + sum (1..dimindex (:N)) (\i. sum (1..dimindex (:N)) (\j. e$i * e$j * (partial j (partial i f) o (\t. x + t % e)) t))` [ ((((use_arg_then2 ("diff2_eq_diff2_on_open", [diff2_eq_diff2_on_open]))(thm_tac (new_rewrite [] [])))) THEN ALL_TAC THEN (case THEN (move ["s"])) THEN (case THEN (move ["open_s"])) THEN (case THEN (move ["xs"])) THEN (move ["df"])); @@ -1095,7 +1095,7 @@ let diff2_dir_derivative2 = Sections.section_proof ["f";"x";"e";"t"] (* Lemma diff2_has_derivative_partial *) let diff2_has_derivative_partial = Sections.section_proof ["f";"i";"x";"e";"t"] `diff2 f (x + t % e:real^N) ==> - (partial i f o (\t. x + t % e) has_real_derivative + (partial i f o (\t. x + t % e) has_real_derivative sum (1..dimindex (:N)) (\j. e$j * (partial j (partial i f) o (\t. x + t % e)) t)) (atreal t)` [ ((((use_arg_then2 ("diff2", [diff2]))(thm_tac (new_rewrite [] [])))) THEN ALL_TAC THEN (case THEN (move ["s"])) THEN (case THEN (move ["open_s"])) THEN (case THEN (move ["xs"])) THEN (move ["df"])); @@ -1105,7 +1105,7 @@ let diff2_has_derivative_partial = Sections.section_proof ["f";"i";"x";"e";"t"] (* Lemma diff2_derivative_partial *) let diff2_derivative_partial = Sections.section_proof ["f";"i";"x";"e";"t"] `diff2 f (x + t % e:real^N) ==> - derivative (partial i f o (\t. x + t % e)) t = + derivative (partial i f o (\t. x + t % e)) t = sum (1..dimindex (:N)) (\j. e$j * (partial j (partial i f) o (\t. x + t % e)) t)` [ ((BETA_TAC THEN (move ["df"])) THEN (((use_arg_then2 ("derivative_unique", [derivative_unique])) (disch_tac [])) THEN (clear_assumption "derivative_unique") THEN (DISCH_THEN apply_tac)) THEN (((use_arg_then2 ("diff2_has_derivative_partial", [diff2_has_derivative_partial])) (disch_tac [])) THEN (clear_assumption "diff2_has_derivative_partial") THEN (exact_tac)) THEN (done_tac)); @@ -1204,7 +1204,7 @@ let mixed_second_partials = Sections.section_proof ["f";"x";"i";"j"] ((fun arg_tac -> arg_tac (Arg_term (`!y e. (f o (\k. y + k % e)) o (\t. k + t) = f o (\t. (y + k % e) + t % e)`))) (term_tac (have_gen_tac [](move ["eq"])))); (((((use_arg_then2 ("eq_ext", [eq_ext]))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("o_THM", [o_THM]))(thm_tac (new_rewrite [] []))))) THEN (simp_tac) THEN (((fun arg_tac -> arg_tac (Arg_theorem (GEN_ALL VECTOR_ADD_RDISTRIB)))(thm_tac (new_rewrite [] [])))) THEN (((fun arg_tac -> arg_tac (Arg_theorem (GEN_ALL VECTOR_ADD_ASSOC)))(thm_tac (new_rewrite [] []))))) THEN (done_tac)); (((repeat_tactic 1 9 (((use_arg_then2 ("eq", []))(thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("partial", [partial]))(gsym_then (thm_tac (new_rewrite [] []))))))) THEN (done_tac)); - ((fun arg_tac -> arg_tac (Arg_term (`!h k. abs h <= r /\ abs k <= r ==> + ((fun arg_tac -> arg_tac (Arg_term (`!h k. abs h <= r /\ abs k <= r ==> (?t1. G h k = k * derivative (F1 h) t1 /\ abs t1 <= abs k)`))) (term_tac (have_gen_tac [](move ["Gh"])))); ((BETA_TAC THEN (move ["h"]) THEN (move ["k"]) THEN (move ["ineq"])) THEN ((((use_arg_then2 ("G_def", []))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (simp_tac))); ((fun arg_tac -> (fun arg_tac -> (fun arg_tac -> (use_arg_then2 ("real_mvt0", [real_mvt0])) (fun fst_arg -> (fun arg_tac -> arg_tac (Arg_term (`F1 h`))) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (fun arg_tac -> arg_tac (Arg_term (`derivative (F1 h)`))) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (use_arg_then2 ("k", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun arg -> thm_tac MP_TAC arg THEN ALL_TAC)); @@ -1214,7 +1214,7 @@ let mixed_second_partials = Sections.section_proof ["f";"x";"i";"j"] (BETA_TAC THEN (case THEN (move ["t1"])) THEN (case THEN (move ["t1_ineq"])) THEN (move ["eq"])); (((use_arg_then2 ("t1", [])) (term_tac exists_tac)) THEN (((use_arg_then2 ("eq", []))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("ineq", [])) (disch_tac [])) THEN (clear_assumption "ineq") THEN ((use_arg_then2 ("t1_ineq", [])) (disch_tac [])) THEN (clear_assumption "t1_ineq") THEN BETA_TAC) THEN (arith_tac) THEN (done_tac)); ((fun arg_tac -> arg_tac (Arg_term (`!h k. abs h <= r /\ abs k <= r ==> - (?t1 t2. G h k = h * k * partial i (partial j f) (x + t1 % basis j + t2 % basis i) + (?t1 t2. G h k = h * k * partial i (partial j f) (x + t1 % basis j + t2 % basis i) /\ abs t1 <= abs k /\ abs t2 <= abs h)`))) (term_tac (have_gen_tac [](move ["Ghk"])))); (BETA_TAC THEN (move ["h"]) THEN (move ["k"]) THEN (move ["ineq"])); ((((fun arg_tac -> (fun arg_tac -> (use_arg_then2 ("Gh", [])) (fun fst_arg -> (use_arg_then2 ("h", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (use_arg_then2 ("k", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (disch_tac [])) THEN BETA_TAC) THEN (((repeat_tactic 1 9 (((use_arg_then2 ("ineq", []))(thm_tac (new_rewrite [] []))))) THEN (simp_tac)) THEN ALL_TAC THEN (case THEN (move ["t1"])) THEN (case THEN ((move ["eq"]) THEN (move ["t1k"]))))); @@ -1250,7 +1250,7 @@ let mixed_second_partials = Sections.section_proof ["f";"x";"i";"j"] ((fun arg_tac -> arg_tac (Arg_term (`!y e. (f o (\h. y + h % e)) o (\t. h + t) = f o (\t. (y + h % e) + t % e)`))) (term_tac (have_gen_tac [](move ["eq"])))); (((((use_arg_then2 ("eq_ext", [eq_ext]))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("o_THM", [o_THM]))(thm_tac (new_rewrite [] []))))) THEN (simp_tac) THEN (((fun arg_tac -> arg_tac (Arg_theorem (GEN_ALL VECTOR_ADD_RDISTRIB)))(thm_tac (new_rewrite [] [])))) THEN (((fun arg_tac -> arg_tac (Arg_theorem (GEN_ALL VECTOR_ADD_ASSOC)))(thm_tac (new_rewrite [] []))))) THEN (done_tac)); (((repeat_tactic 1 9 (((use_arg_then2 ("eq", []))(thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("partial", [partial]))(gsym_then (thm_tac (new_rewrite [] []))))))) THEN (done_tac)); - ((fun arg_tac -> arg_tac (Arg_term (`!h k. abs h <= r /\ abs k <= r ==> + ((fun arg_tac -> arg_tac (Arg_term (`!h k. abs h <= r /\ abs k <= r ==> (?t3. G h k = h * derivative (F2 k) t3 /\ abs t3 <= abs h)`))) (term_tac (have_gen_tac [](move ["Gk"])))); ((BETA_TAC THEN (move ["h"]) THEN (move ["k"]) THEN (move ["ineq"])) THEN ((((use_arg_then2 ("G_eq", []))(thm_tac (new_rewrite [] [])))) THEN (simp_tac))); ((fun arg_tac -> (fun arg_tac -> (fun arg_tac -> (use_arg_then2 ("real_mvt0", [real_mvt0])) (fun fst_arg -> (fun arg_tac -> arg_tac (Arg_term (`F2 k`))) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (fun arg_tac -> arg_tac (Arg_term (`derivative (F2 k)`))) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (use_arg_then2 ("h", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun arg -> thm_tac MP_TAC arg THEN ALL_TAC)); @@ -1260,7 +1260,7 @@ let mixed_second_partials = Sections.section_proof ["f";"x";"i";"j"] (BETA_TAC THEN (case THEN (move ["t3"])) THEN (case THEN (move ["t3_ineq"])) THEN (move ["eq"])); (((use_arg_then2 ("t3", [])) (term_tac exists_tac)) THEN (((use_arg_then2 ("eq", []))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("ineq", [])) (disch_tac [])) THEN (clear_assumption "ineq") THEN ((use_arg_then2 ("t3_ineq", [])) (disch_tac [])) THEN (clear_assumption "t3_ineq") THEN BETA_TAC) THEN (arith_tac) THEN (done_tac)); ((fun arg_tac -> arg_tac (Arg_term (`!h k. abs h <= r /\ abs k <= r ==> - (?t3 t4. G h k = h * k * partial j (partial i f) (x + t4 % basis j + t3 % basis i) + (?t3 t4. G h k = h * k * partial j (partial i f) (x + t4 % basis j + t3 % basis i) /\ abs t3 <= abs h /\ abs t4 <= abs k)`))) (term_tac (have_gen_tac [](move ["Gkh"])))); (BETA_TAC THEN (move ["h"]) THEN (move ["k"]) THEN (move ["ineq"])); ((((fun arg_tac -> (fun arg_tac -> (use_arg_then2 ("Gk", [])) (fun fst_arg -> (use_arg_then2 ("h", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (use_arg_then2 ("k", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (disch_tac [])) THEN BETA_TAC) THEN (((repeat_tactic 1 9 (((use_arg_then2 ("ineq", []))(thm_tac (new_rewrite [] []))))) THEN (simp_tac)) THEN ALL_TAC THEN (case THEN (move ["t3"])) THEN (case THEN ((move ["eq"]) THEN (move ["t3h"]))))); @@ -1302,7 +1302,7 @@ let mixed_second_partials = Sections.section_proof ["f";"x";"i";"j"] (((((fun arg_tac -> (use_arg_then2 ("REAL_MUL_LINV", [REAL_MUL_LINV])) (fun fst_arg -> (fun arg_tac -> arg_tac (Arg_term (`&2`))) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg)))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 0 10 (((use_arg_then2 ("REAL_LT_LMUL", [REAL_LT_LMUL]))(thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 0 10 (((use_arg_then2 ("REAL_LT_INV", [REAL_LT_INV]))(thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 0 10 (((use_arg_then2 ("andTb", [andTb]))(thm_tac (new_rewrite [] [])))))) THEN (TRY ((arith_tac)))); ((THENL_FIRST) ((fun arg_tac -> arg_tac (Arg_term (`&2 = sqrt (&2 * &2)`))) (term_tac (have_gen_tac [](((conv_thm_tac DISCH_THEN)(thm_tac (new_rewrite [2] []))))))) (((((use_arg_then2 ("REAL_POW_2", [REAL_POW_2]))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (((use_arg_then2 ("POW_2_SQRT_ABS", [POW_2_SQRT_ABS]))(thm_tac (new_rewrite [] []))))) THEN (arith_tac) THEN (done_tac))); ((((use_arg_then2 ("SQRT_MONO_LT", [SQRT_MONO_LT]))(thm_tac (new_rewrite [] [])))) THEN (arith_tac) THEN (done_tac)); - ((fun arg_tac -> arg_tac (Arg_term (`((\y:real^2. lift (G (y$1) (y$2) / (y$1 * y$2))) --> lift (partial j (partial i f) x)) + ((fun arg_tac -> arg_tac (Arg_term (`((\y:real^2. lift (G (y$1) (y$2) / (y$1 * y$2))) --> lift (partial j (partial i f) x)) (at (vec 0) within {y | &0 < y$1 /\ &0 < y$2})`))) (term_tac (have_gen_tac [](move ["lim_ji"])))); ((((use_arg_then2 ("LIM_WITHIN", [LIM_WITHIN]))(thm_tac (new_rewrite [] [])))) THEN (move ["e"]) THEN (move ["e_gt0"])); ((((fun arg_tac -> (fun arg_tac -> (use_arg_then2 ("pc", [])) (fun fst_arg -> (use_arg_then2 ("i", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (use_arg_then2 ("j", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (disch_tac [])) THEN BETA_TAC) THEN (((((fun arg_tac -> arg_tac (Arg_theorem (GEN_ALL continuous_at)))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("dist", [dist]))(thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("o_THM", [o_THM]))(thm_tac (new_rewrite [] [])))))) THEN ((fun arg_tac -> (conv_thm_tac DISCH_THEN) (fun fst_arg -> (use_arg_then2 ("e_gt0", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (thm_tac MP_TAC)) THEN (case THEN (move ["d"])) THEN (case THEN (move ["d0"])) THEN (move ["e_ineq"]))); @@ -1322,7 +1322,7 @@ let mixed_second_partials = Sections.section_proof ["f";"x";"i";"j"] ((((use_arg_then2 ("REAL_LET_TRANS", [REAL_LET_TRANS])) (disch_tac [])) THEN (clear_assumption "REAL_LET_TRANS") THEN (DISCH_THEN apply_tac)) THEN ((fun arg_tac -> arg_tac (Arg_term (`norm (t2 % basis j:real^N) + norm (t1 % basis i:real^N)`))) (term_tac exists_tac))); ((((use_arg_then2 ("NORM_TRIANGLE", [NORM_TRIANGLE]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("andTb", [andTb]))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("NORM_MUL", [NORM_MUL]))(thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("NORM_BASIS", [NORM_BASIS]))(thm_tac (new_rewrite [] []))))) THEN ((TRY done_tac)) THEN (repeat_tactic 1 9 (((use_arg_then2 ("REAL_MUL_RID", [REAL_MUL_RID]))(thm_tac (new_rewrite [] [])))))); ((((use_arg_then2 ("yr", [])) (disch_tac [])) THEN (clear_assumption "yr") THEN ((use_arg_then2 ("t_ineq", [])) (disch_tac [])) THEN (clear_assumption "t_ineq") THEN BETA_TAC) THEN (arith_tac) THEN (done_tac)); - ((fun arg_tac -> arg_tac (Arg_term (`((\y:real^2. lift (G (y$1) (y$2) / (y$1 * y$2))) --> lift (partial i (partial j f) x)) + ((fun arg_tac -> arg_tac (Arg_term (`((\y:real^2. lift (G (y$1) (y$2) / (y$1 * y$2))) --> lift (partial i (partial j f) x)) (at (vec 0) within {y | &0 < y$1 /\ &0 < y$2})`))) (term_tac (have_gen_tac [](move ["lim_ij"])))); ((((use_arg_then2 ("LIM_WITHIN", [LIM_WITHIN]))(thm_tac (new_rewrite [] [])))) THEN (move ["e"]) THEN (move ["e_gt0"])); ((((fun arg_tac -> (fun arg_tac -> (use_arg_then2 ("pc", [])) (fun fst_arg -> (use_arg_then2 ("j", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (use_arg_then2 ("i", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (disch_tac [])) THEN BETA_TAC) THEN (((((fun arg_tac -> arg_tac (Arg_theorem (GEN_ALL continuous_at)))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("dist", [dist]))(thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("o_THM", [o_THM]))(thm_tac (new_rewrite [] [])))))) THEN ((fun arg_tac -> (conv_thm_tac DISCH_THEN) (fun fst_arg -> (use_arg_then2 ("e_gt0", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (thm_tac MP_TAC)) THEN (case THEN (move ["d"])) THEN (case THEN (move ["d0"])) THEN (move ["e_ineq"]))); @@ -1375,7 +1375,7 @@ let y_in_domain = Sections.section_proof ["domain";"y";"w"] (* Lemma domain_width *) let domain_width = Sections.section_proof ["p";"domain";"y";"w"] -`m_cell_domain domain y (w:real^N) ==> +`m_cell_domain domain y (w:real^N) ==> p IN interval [domain] ==> !i. i IN 1..dimindex (:N) ==> abs (p$i - y$i) <= w$i` [ @@ -1385,7 +1385,7 @@ let domain_width = Sections.section_proof ["p";"domain";"y";"w"] (* Lemma sum_swap1 *) let sum_swap1 = Sections.section_proof ["g";"n"] -`sum (1..n) (\i. sum (i + 1..n) (\j. g i j)) = +`sum (1..n) (\i. sum (i + 1..n) (\j. g i j)) = sum (1..n) (\i. sum (1..i - 1) (\j. g j i))` [ ((repeat_tactic 1 9 (((use_arg_then2 ("SUM_SUM_PRODUCT", [SUM_SUM_PRODUCT]))(thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 0 10 (((use_arg_then2 ("FINITE_NUMSEG", [FINITE_NUMSEG]))(thm_tac (new_rewrite [] []))))) THEN ((simp_tac THEN TRY done_tac)) THEN (repeat_tactic 1 9 (((use_arg_then2 ("IN_NUMSEG", [IN_NUMSEG]))(thm_tac (new_rewrite [] [])))))); @@ -1406,14 +1406,14 @@ let sum_swap1 = Sections.section_proof ["g";"n"] (* Lemma m_taylor_error_eq *) let m_taylor_error_eq = Sections.section_proof ["f";"domain";"w";"error"] -`diff2c_domain domain f ==> +`diff2c_domain domain f ==> (m_taylor_error f domain (w:real^N) error <=> - (!x. x IN interval [domain] ==> + (!x. x IN interval [domain] ==> sum (1..dimindex (:N)) (\i. w$i * (w$i * abs (partial2 i i f x) + &2 * sum (1..i - 1) (\j. w$j * abs (partial2 j i f x)))) <= error))` [ (((((use_arg_then2 ("diff2c_domain", [diff2c_domain]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("m_taylor_error", [m_taylor_error]))(thm_tac (new_rewrite [] []))))) THEN (move ["d2f"])); - ((fun arg_tac -> arg_tac (Arg_term (`!g1 g2. (!x. x IN interval [domain] ==> g1 x = g2 x) ==> + ((fun arg_tac -> arg_tac (Arg_term (`!g1 g2. (!x. x IN interval [domain] ==> g1 x = g2 x) ==> ((!x. x IN interval [domain] ==> g1 x <= error) <=> (!x. x IN interval [domain] ==> g2 x <= error))`))) (term_tac (have_gen_tac [](move ["eq"])))); ((BETA_TAC THEN (move ["g1"]) THEN (move ["g2"]) THEN (move ["eq"])) THEN ((split_tac) THEN (move ["cond"]) THEN (move ["x"]) THEN (move ["Px"]))); (((((use_arg_then2 ("eq", []))(gsym_then (thm_tac (new_rewrite [] []))))) THEN ((TRY done_tac)) THEN (((use_arg_then2 ("cond", []))(thm_tac (new_rewrite [] []))))) THEN (done_tac)); @@ -1447,7 +1447,7 @@ let diff2_derivative2_bound = Sections.section_proof ["domain";"y";"w";"p";"f";" p IN interval [domain] ==> diff2_domain domain f ==> m_taylor_error f domain w dd_bound ==> - (!t. interval_arith t (&0, &1) ==> + (!t. interval_arith t (&0, &1) ==> abs (nth_derivative 2 (f o (\t. y + t % (p - y))) t) <= dd_bound)` [ (((((use_arg_then2 ("diff2_domain", [diff2_domain]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("m_taylor_error", [m_taylor_error]))(thm_tac (new_rewrite [] []))))) THEN (move ["domainH"]) THEN (move ["p_in"]) THEN (move ["df"]) THEN (move ["boundedH"]) THEN (move ["t"]) THEN (move ["t_in"])); @@ -1551,7 +1551,7 @@ let diff2_derivative_partial_bound = Sections.section_proof ["domain";"y";"w";"p p IN interval [domain] ==> diff2_domain domain f ==> m_taylor_partial_error f i domain w d_bound ==> - (!t. interval_arith t (&0, &1) ==> + (!t. interval_arith t (&0, &1) ==> abs (derivative (partial i f o (\t. y + t % (p - y))) t) <= d_bound)` [ (((((use_arg_then2 ("diff2_domain", [diff2_domain]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("m_taylor_partial_error", [m_taylor_partial_error]))(thm_tac (new_rewrite [] []))))) THEN (move ["domainH"]) THEN (move ["p_in"]) THEN (move ["df"]) THEN (move ["boundedH"]) THEN (move ["t"]) THEN (move ["t_in"])); @@ -1727,7 +1727,7 @@ let has_derivative_uni_compose = Sections.section_proof ["u";"f";"u'";"f'";"x"] (* Lemma diff_uni_compose *) let diff_uni_compose = Sections.section_proof ["u";"f";"x"] -`lift o f differentiable at x ==> +`lift o f differentiable at x ==> u real_differentiable atreal (f x) ==> lift o u o f differentiable at x` [ @@ -2149,7 +2149,7 @@ let second_partial_neg = Sections.section_proof ["f"] (* Lemma second_partial_add *) let second_partial_add = Sections.section_proof ["f";"g"] -`diff2 f x ==> diff2 g x ==> +`diff2 f x ==> diff2 g x ==> partial2 i j (\x. f x + g x) x = partial2 i j f x + partial2 i j g x` [ ((repeat_tactic 1 9 (((use_arg_then2 ("diff2", [diff2]))(thm_tac (new_rewrite [] []))))) THEN ALL_TAC THEN (case THEN (move ["s"])) THEN (case THEN (move ["open_s"])) THEN (case THEN (move ["xs"])) THEN (move ["df"]) THEN (case THEN (move ["t"])) THEN (case THEN (move ["open_t"])) THEN (case THEN (move ["ys"])) THEN (move ["dg"])); @@ -2162,7 +2162,7 @@ let second_partial_add = Sections.section_proof ["f";"g"] (* Lemma second_partial_sub *) let second_partial_sub = Sections.section_proof ["f";"g"] -`diff2 f x ==> diff2 g x ==> +`diff2 f x ==> diff2 g x ==> partial2 i j (\x. f x - g x) x = partial2 i j f x - partial2 i j g x` [ ((BETA_TAC THEN (move ["d2f"]) THEN (move ["d2g"])) THEN ((((use_arg_then2 ("real_sub", [real_sub]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("second_partial_add", [second_partial_add]))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 0 10 (((use_arg_then2 ("diff2_neg", [diff2_neg]))(thm_tac (new_rewrite [] []))))) THEN ((TRY done_tac)) THEN (((use_arg_then2 ("second_partial_neg", [second_partial_neg]))(thm_tac (new_rewrite [] [])))) THEN ((TRY done_tac)) THEN (((use_arg_then2 ("real_sub", [real_sub]))(gsym_then (thm_tac (new_rewrite [] [])))))) THEN (done_tac)); @@ -2170,7 +2170,7 @@ let second_partial_sub = Sections.section_proof ["f";"g"] (* Lemma second_partial_mul *) let second_partial_mul = Sections.section_proof ["f";"g"] -`diff2 f x ==> diff2 g x ==> +`diff2 f x ==> diff2 g x ==> partial2 i j (\x. f x * g x) x = (partial2 i j f x * g x + partial j f x * partial i g x) + (partial i f x * partial j g x + f x * partial2 i j g x)` [ @@ -2675,7 +2675,7 @@ let diff2c_mul = Sections.section_proof ["f";"g"] ((use_arg_then2 ("real_cont_at_local", [real_cont_at_local])) (thm_tac apply_tac)); ((((use_arg_then2 ("d2g", [])) (disch_tac [])) THEN (clear_assumption "d2g") THEN ((use_arg_then2 ("d2f", [])) (disch_tac [])) THEN (clear_assumption "d2f") THEN BETA_TAC) THEN ((((use_arg_then2 ("diff2_eq_diff2_on_open", [diff2_eq_diff2_on_open]))(thm_tac (new_rewrite [] [])))) THEN ALL_TAC THEN (case THEN (move ["s"])) THEN (case THEN (move ["open_s"])) THEN (case THEN (move ["xs"])) THEN (move ["d2f"]))); ((((use_arg_then2 ("diff2_eq_diff2_on_open", [diff2_eq_diff2_on_open]))(thm_tac (new_rewrite [] [])))) THEN ALL_TAC THEN (case THEN (move ["t"])) THEN (case THEN (move ["open_t"])) THEN (case THEN (move ["xt"])) THEN (move ["d2g"])); - (((fun arg_tac -> arg_tac (Arg_term (`(\x. (partial2 j i f x * g x + partial i f x * partial j g x) + + (((fun arg_tac -> arg_tac (Arg_term (`(\x. (partial2 j i f x * g x + partial i f x * partial j g x) + partial j f x * partial i g x + f x * partial2 j i g x)`))) (term_tac exists_tac)) THEN ((fun arg_tac -> arg_tac (Arg_term (`s INTER t`))) (term_tac exists_tac))); ((THENL_ROT (-1)) (((repeat_tactic 1 9 (((use_arg_then2 ("IN_INTER", [IN_INTER]))(thm_tac (new_rewrite [] []))))) THEN (((use_arg_then2 ("xs", []))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("xt", []))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("OPEN_INTER", [OPEN_INTER]))(thm_tac (new_rewrite [] [])))) THEN ((simp_tac THEN TRY done_tac))) THEN (split_tac))); ((BETA_TAC THEN (move ["y"]) THEN (case THEN ((move ["ys"]) THEN (move ["yt"])))) THEN ((((use_arg_then2 ("second_partial_mul", [second_partial_mul]))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 0 10 (((use_arg_then2 ("d2f", []))(thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 0 10 (((use_arg_then2 ("d2g", []))(thm_tac (new_rewrite [] [])))))) THEN (done_tac)); @@ -3004,7 +3004,7 @@ Sections.end_section "M_LinApprox";; let second_bounded = new_definition `second_bounded f domain dd_bounds_list <=> !x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith (partial2 j i f x) int))`;; -let m_taylor_interval = +let m_taylor_interval = new_definition `m_taylor_interval f domain y w f_bounds d_bounds_list dd_bounds_list <=> m_cell_domain domain y w /\ diff2c_domain domain f /\ @@ -3410,7 +3410,7 @@ let partial_x_lemma = Sections.section_proof ["k";"i"] (* Lemma partial_x *) let partial_x = Sections.section_proof ["k";"i"] -`k IN 1..dimindex (:N) ==> +`k IN 1..dimindex (:N) ==> partial i (\x:real^N. x$k) = (\x. if i = k then &1 else &0)` [ ((BETA_TAC THEN (move ["k_ineq"])) THEN ((((use_arg_then2 ("partial_x_lemma", [partial_x_lemma]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("BASIS_COMPONENT", [BASIS_COMPONENT]))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 0 10 (((use_arg_then2 ("IN_NUMSEG", [IN_NUMSEG]))(gsym_then (thm_tac (new_rewrite [] [])))))) THEN ((TRY done_tac))) THEN (arith_tac) THEN (done_tac)); diff --git a/Formal_ineqs/taylor/theory/multivariate_taylor.vhl b/Formal_ineqs/taylor/theory/multivariate_taylor.vhl index 9ca7f9a1..42ea4349 100644 --- a/Formal_ineqs/taylor/theory/multivariate_taylor.vhl +++ b/Formal_ineqs/taylor/theory/multivariate_taylor.vhl @@ -1,14 +1,20 @@ -(* =========================================================== *) -(* Theory of multivariate taylor intervals *) -(* Requires SSReflect/HOL Light for translation *) -(* See http://code.google.com/p/flyspeck/downloads/list *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -"needs \"lib/ssrbool.hl\"". -"needs \"lib/ssrnat.hl\"". -"needs \"taylor/theory/taylor_interval-compiled.hl\"". +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Theory of multivariate taylor intervals *) +(* Requires SSReflect/HOL Light for translation *) +(* See https://github.com/monadius/hol-ssreflect *) +(* -------------------------------------------------------------------------- *) + +"needs \"Formal_ineqs/lib/ssrbool.hl\"". +"needs \"Formal_ineqs/lib/ssrnat.hl\"". +"needs \"Formal_ineqs/taylor/theory/taylor_interval-compiled.hl\"". module Multivariate_taylor. @@ -30,7 +36,7 @@ module Multivariate_taylor. "let all_n = define `(all_n n [] s <=> T) /\ (all_n n (CONS h t) s <=> s n h /\ all_n (SUC n) t s)`". -"let m_lin_approx = new_definition `m_lin_approx (f:real^N->real) x f_bounds df_bounds_list <=> +"let m_lin_approx = new_definition `m_lin_approx (f:real^N->real) x f_bounds df_bounds_list <=> (lift o f) differentiable at x /\ interval_arith (f x) f_bounds /\ all_n 1 df_bounds_list (\i int. interval_arith (partial i f x) int)`". @@ -39,7 +45,7 @@ module Multivariate_taylor. Section Misc. -Lemma f_lift_neg f : `lift o (\x. --f x) = (\x. --(lift o f) x)`. +Lemma f_lift_neg f : `lift o (\x. --f x) = (\x. --(lift o f) x)`. by rewrite -eq_ext !o_THM /= LIFT_NEG. Qed. Lemma f_lift_scale f c : `lift o (\x. c * f x) = (\x. c % (lift o f) x)`. @@ -100,13 +106,13 @@ Qed. Hypothesis dg : `g differentiable at x`. -Lemma frechet_add : `frechet_derivative (\x. f x + g x) (at x) = +Lemma frechet_add : `frechet_derivative (\x. f x + g x) (at x) = (\y. frechet_derivative f (at x) y + frechet_derivative g (at x) y)`. rewrite EQ_SYM_EQ; apply FRECHET_DERIVATIVE_AT; apply HAS_DERIVATIVE_ADD. by rewr ETA_AX; rewrite -!FRECHET_DERIVATIVE_WORKS. Qed. -Lemma frechet_sub : `frechet_derivative (\x. f x - g x) (at x) = +Lemma frechet_sub : `frechet_derivative (\x. f x - g x) (at x) = (\y. frechet_derivative f (at x) y - frechet_derivative g (at x) y)`. rewrite EQ_SYM_EQ; apply FRECHET_DERIVATIVE_AT; apply HAS_DERIVATIVE_SUB. by rewr ETA_AX; rewrite -!FRECHET_DERIVATIVE_WORKS. @@ -115,7 +121,7 @@ Qed. End MoreFrechet. -Lemma differentiable_compose_at f g x : +Lemma differentiable_compose_at f g x : `f differentiable at (g x) ==> g differentiable at x ==> (f o g) differentiable at x`. @@ -153,7 +159,7 @@ by rewrite REAL_LT_MUL2; move: ineq wnx; arith. Qed. -Lemma has_derivative_x12 y : +Lemma has_derivative_x12 y : `(lift o (\x:real^2. x$1 * x$2) has_derivative lift o (\x. y$2 * x$1 + y$1 * x$2)) (at y)`. rewrite has_derivative_at; split. rewrite linear !o_THM /= !VECTOR_ADD_COMPONENT !VECTOR_MUL_COMPONENT !LIFT_ADD !LIFT_CMUL !LIFT_ADD. @@ -178,7 +184,7 @@ by case => <-; first rewrite REAL_MUL_SYM; rewrite REAL_LT_LMUL ineq // DIMINDEX Qed. -Lemma lambda_eq_vsum f : `(\x:A. lambda i. f i x) = +Lemma lambda_eq_vsum f : `(\x:A. lambda i. f i x) = (\x. vsum (1..dimindex (:N)) (\i. f i x % (basis i:real^N)))`. rewrite -eq_ext /= => x. rewrite CART_EQ => i ineq; rewrite "GEN_ALL LAMBDA_BETA" // VSUM_COMPONENT //=. @@ -228,7 +234,7 @@ by rewrite "ARITH_RULE `~(2 = 1)`" /=; rewr ETA_AX. Qed. -Lemma has_derivative_mul f g f' g' y : +Lemma has_derivative_mul f g f' g' y : `(lift o f has_derivative lift o f') (at y) ==> (lift o g has_derivative lift o g') (at y) ==> (lift o (\x. f x * g x) has_derivative lift o (\x. f' x * g y + f y * g' x)) (at y)`. @@ -236,7 +242,7 @@ move => df dg. have ->: `lift o (\x. f x * g x) = (lift o (\p. p$1 * p$2)) o (\x. vector [f x; g x]:real^2)`. by rewrite -eq_ext !o_THM /= => x; rewrite !"GEN_ALL VECTOR_2". set q := `vector [f y; g y]:real^2`. -have ->: `lift o (\x. f' x * g y + f y * g' x) = +have ->: `lift o (\x. f' x * g y + f y * g' x) = (lift o (\x:real^2. q$2 * x$1 + q$1 * x$2)) o (\x. vector [f' x; g' x])`. by rewrite -eq_ext !o_THM /= -q_def !"GEN_ALL VECTOR_2" REAL_MUL_SYM. by apply DIFF_CHAIN_AT => /=; rewrite q_def has_derivative_x12 has_derivative_vector2. @@ -262,7 +268,7 @@ Lemma frechet_mul f g y : `lift o f differentiable at y ==> lift o g differentiable at y ==> frechet_derivative (lift o (\x. f x * g x)) (at y) = - (\x. g y % frechet_derivative (lift o f) (at y) x + + (\x. g y % frechet_derivative (lift o f) (at y) x + f y % frechet_derivative (lift o g) (at y) x)`. rewrite !FRECHET_DERIVATIVE_WORKS [`frechet_derivative _1 _2`]f_eq_lift_drop => df. rewrite [`frechet_derivative _1 _2`]f_eq_lift_drop => dg. @@ -285,7 +291,7 @@ Section Partial. Lemma real_derivative_compose_frechet f h t : `(lift o f) differentiable at (h t) ==> (h o drop) differentiable at (lift t) ==> - ((f o h) has_real_derivative (drop o (frechet_derivative (lift o f) (at (h t)) o + ((f o h) has_real_derivative (drop o (frechet_derivative (lift o f) (at (h t)) o frechet_derivative (h o drop) (at (lift t))) o lift) (&1)) (atreal t)`. move => diff_f diff_h. move: (diff_f) (diff_h); rewrite !FRECHET_DERIVATIVE_WORKS. @@ -309,7 +315,7 @@ Qed. Lemma real_derivative_compose_jacobian f h t : `(lift o f) differentiable at (h t) ==> (h o drop) differentiable at (lift t) ==> - ((f o h) has_real_derivative (jacobian (lift o f) (at (h t)) ** + ((f o h) has_real_derivative (jacobian (lift o f) (at (h t)) ** jacobian (h o drop) (at (lift t)))$1$1) (atreal t)`. move => df dh. move: (real_derivative_compose_frechet df dh). @@ -320,8 +326,8 @@ Qed. -Lemma diff_imp_real_diff f h t : - `(lift o f) differentiable at (h t) ==> +Lemma diff_imp_real_diff f h t : + `(lift o f) differentiable at (h t) ==> (h o drop) differentiable at (lift t) ==> (f o h) real_differentiable atreal t`. move => diff_f diff_h. @@ -339,7 +345,7 @@ by exists `\x. drop x % e`; apply HAS_DERIVATIVE_VMUL_DROP; rewrite HAS_DERIVATI Qed. -Lemma frechet_direction y e t : +Lemma frechet_direction y e t : `frechet_derivative ((\t. y + t % e) o drop) (at (lift t)) = (\x. drop x % e)`. rewrite f_unary_drop frechet_add. rewrite DIFFERENTIABLE_CONST andTb; apply HAS_DERIVATIVE_IMP_DIFFERENTIABLE. @@ -350,7 +356,7 @@ Qed. Lemma real_dir_derivative_frechet f y e t : `(lift o f) differentiable at (y + t % e) ==> - ((f o (\t. y + t % e)) has_real_derivative + ((f o (\t. y + t % e)) has_real_derivative (drop (frechet_derivative (lift o f) (at (y + t % e)) e))) (atreal t)`. move => df. move: (real_derivative_compose_frechet f `\t. y + t % e` t). @@ -359,7 +365,7 @@ Qed. Lemma real_dir_derivative_jacobian f y e t : `(lift o f) differentiable at (y + t % e) ==> - ((f o (\t. y + t % e)) has_real_derivative + ((f o (\t. y + t % e)) has_real_derivative drop (jacobian (lift o f) (at (y + t % e)) ** e)) (atreal t)`. move => df. move: (real_dir_derivative_frechet f e df). @@ -420,7 +426,7 @@ by rewrite derivative_composition // REAL_MUL_SYM. Qed. -Lemma projection_has_derivative i net : +Lemma projection_has_derivative i net : `i IN 1..dimindex (:N) ==> (lift o (\x:real^N. x$i) has_derivative lift o (\x. x$i)) net`. rewrite IN_NUMSEG => ineq. @@ -577,7 +583,7 @@ End Partial. Section PartialMonotone. -Lemma derivative_translation f x : +Lemma derivative_translation f x : `f real_differentiable atreal x ==> derivative f x = derivative (f o (\t. x + t)) (&0)`. move => diff_f. @@ -589,7 +595,7 @@ Qed. Implicit Type f : `:real^N->real`. -Lemma partial_increasing_left f j u x z lo : +Lemma partial_increasing_left f j u x z lo : `(!i. i IN 1..dimindex (:N) ==> ~(i = j) ==> u$i = z$i) ==> u$j = x$j ==> (!y. y IN interval [x,z] ==> (lift o f) differentiable at y) ==> @@ -636,7 +642,7 @@ have ds : `!t. t IN s ==> (g has_real_derivative (partial j f (y' + t % basis j) by rewrite -derivative_translation ?has_derivative_alt -g_def diff_imp_real_diff; rewrite diff_f //= diff_direction. have pos: `&0 <= y$j - x$j`; first by move: (y_in j_in); arith. -have := HAS_REAL_DERIVATIVE_INCREASING_IMP +have := HAS_REAL_DERIVATIVE_INCREASING_IMP g `\t. partial j f (y' + t % basis j)` s `&0` `y$j - x$j`. rewrite -{1}s_def IS_REALINTERVAL_INTERVAL /=; rewr ds /=; apply. rewrite -{2 3}s_def !IN_REAL_INTERVAL pos !REAL_LE_REFL /=. @@ -644,7 +650,7 @@ by move => t /in_s /partial_pos. Qed. -Lemma partial_decreasing_left f j u x z hi : +Lemma partial_decreasing_left f j u x z hi : `(!i. i IN 1..dimindex (:N) ==> ~(i = j) ==> u$i = z$i) ==> u$j = x$j ==> (!y. y IN interval [x,z] ==> (lift o f) differentiable at y) ==> @@ -685,7 +691,7 @@ by rewrite LINEAR_NEG ?LINEAR_FRECHET_DERIVATIVE // DROP_NEG. Qed. -Lemma partial_increasing_right f j u x z hi : +Lemma partial_increasing_right f j u x z hi : `(!i. i IN 1..dimindex (:N) ==> ~(i = j) ==> u$i = x$i) ==> u$j = z$j ==> (!y. y IN interval [x,z] ==> (lift o f) differentiable at y) ==> @@ -755,8 +761,8 @@ by rewrite -R_def /= !arithH. Qed. -Lemma real_taylor1_bound f d_bound: - `(!t. interval_arith t (&0, &1) ==> f real_differentiable atreal t /\ +Lemma real_taylor1_bound f d_bound: + `(!t. interval_arith t (&0, &1) ==> f real_differentiable atreal t /\ abs (derivative f t) <= d_bound) ==> abs (f (&1) - f (&0)) <= d_bound`. move => df. @@ -875,7 +881,7 @@ by move: t_in; rewrite IN_REAL_INTERVAL; arith. Qed. -Lemma diff2_dir f x e t : `diff2 f (x + t % e:real^N) ==> +Lemma diff2_dir f x e t : `diff2 f (x + t % e:real^N) ==> nth_diff_strong 2 (f o (\t. x + t % e)) t`. Proof. rewrite diff2_eq_diff2_on_open nth_diff_strong2_eq => [] [s] [open_s] [xs] df. @@ -907,7 +913,7 @@ Qed. Lemma diff2_dir_derivative2 f x e t : `diff2 f (x + t % e:real^N) ==> nth_derivative 2 (f o (\t. x + t % e)) t = - sum (1..dimindex (:N)) (\i. sum (1..dimindex (:N)) + sum (1..dimindex (:N)) (\i. sum (1..dimindex (:N)) (\j. e$i * e$j * (partial j (partial i f) o (\t. x + t % e)) t))`. Proof. rewrite diff2_eq_diff2_on_open => [] [s] [open_s] [xs] df. @@ -929,7 +935,7 @@ Qed. Lemma diff2_has_derivative_partial f i x e t : `diff2 f (x + t % e:real^N) ==> - (partial i f o (\t. x + t % e) has_real_derivative + (partial i f o (\t. x + t % e) has_real_derivative sum (1..dimindex (:N)) (\j. e$j * (partial j (partial i f) o (\t. x + t % e)) t)) (atreal t)`. rewrite diff2 => [] [s] [open_s] [xs] df. by rewrite real_dir_derivative_partial df. @@ -937,7 +943,7 @@ Qed. Lemma diff2_derivative_partial f i x e t : `diff2 f (x + t % e:real^N) ==> - derivative (partial i f o (\t. x + t % e)) t = + derivative (partial i f o (\t. x + t % e)) t = sum (1..dimindex (:N)) (\j. e$j * (partial j (partial i f) o (\t. x + t % e)) t)`. by move => df; apply: derivative_unique; exact: diff2_has_derivative_partial. Qed. @@ -1025,7 +1031,7 @@ have F1_der : `!h k. abs h <= r /\ abs k <= r ==> have eq: `!y e. (f o (\k. y + k % e)) o (\t. k + t) = f o (\t. (y + k % e) + t % e)`. by rewrite -eq_ext !o_THM /= "GEN_ALL VECTOR_ADD_RDISTRIB" "GEN_ALL VECTOR_ADD_ASSOC". by rewrite !eq -!partial. -have Gh: `!h k. abs h <= r /\ abs k <= r ==> +have Gh: `!h k. abs h <= r /\ abs k <= r ==> (?t1. G h k = k * derivative (F1 h) t1 /\ abs t1 <= abs k)`. move => h k ineq; rewrite -G_def /=. have := real_mvt0 `F1 h` `derivative (F1 h)` k. @@ -1035,7 +1041,7 @@ have Gh: `!h k. abs h <= r /\ abs k <= r ==> move => [t1] [t1_ineq] eq. by exists t1; rewrite eq; move: t1_ineq ineq; arith. have Ghk: `!h k. abs h <= r /\ abs k <= r ==> - (?t1 t2. G h k = h * k * partial i (partial j f) (x + t1 % basis j + t2 % basis i) + (?t1 t2. G h k = h * k * partial i (partial j f) (x + t1 % basis j + t2 % basis i) /\ abs t1 <= abs k /\ abs t2 <= abs h)`. move => h k ineq. move: (Gh h k); rewrite !ineq /= => [] [t1] [eq t1k]. @@ -1072,7 +1078,7 @@ have F2_der : `!h k. abs h <= r /\ abs k <= r ==> have eq: `!y e. (f o (\h. y + h % e)) o (\t. h + t) = f o (\t. (y + h % e) + t % e)`. by rewrite -eq_ext !o_THM /= "GEN_ALL VECTOR_ADD_RDISTRIB" "GEN_ALL VECTOR_ADD_ASSOC". by rewrite !eq -!partial. -have Gk: `!h k. abs h <= r /\ abs k <= r ==> +have Gk: `!h k. abs h <= r /\ abs k <= r ==> (?t3. G h k = h * derivative (F2 k) t3 /\ abs t3 <= abs h)`. move => h k ineq; rewrite G_eq /=. have := real_mvt0 `F2 k` `derivative (F2 k)` h. @@ -1082,7 +1088,7 @@ have Gk: `!h k. abs h <= r /\ abs k <= r ==> move => [t3] [t3_ineq] eq. by exists t3; rewrite eq; move: t3_ineq ineq; arith. have Gkh: `!h k. abs h <= r /\ abs k <= r ==> - (?t3 t4. G h k = h * k * partial j (partial i f) (x + t4 % basis j + t3 % basis i) + (?t3 t4. G h k = h * k * partial j (partial i f) (x + t4 % basis j + t3 % basis i) /\ abs t3 <= abs h /\ abs t4 <= abs k)`. move => h k ineq. move: (Gk h k); rewrite !ineq /= => [] [t3] [eq t3h]. @@ -1124,7 +1130,7 @@ have lim0: `(vec 0:real^2) limit_point_of {y | &0 < y$1 /\ &0 < y$2}`. rewrite -(REAL_MUL_LINV `&2`) ?REAL_LT_LMUL ?REAL_LT_INV ?andTb; try arith. have {2}->: `&2 = sqrt (&2 * &2)`; first by rewrite -REAL_POW_2 POW_2_SQRT_ABS; arith. by rewrite SQRT_MONO_LT; arith. -have lim_ji: `((\y:real^2. lift (G (y$1) (y$2) / (y$1 * y$2))) --> lift (partial j (partial i f) x)) +have lim_ji: `((\y:real^2. lift (G (y$1) (y$2) / (y$1 * y$2))) --> lift (partial j (partial i f) x)) (at (vec 0) within {y | &0 < y$1 /\ &0 < y$2})`. rewrite LIM_WITHIN => e e_gt0. move: (pc i j); rewrite "GEN_ALL continuous_at" !dist !o_THM => /(_ e_gt0) [d] [d0] e_ineq. @@ -1144,7 +1150,7 @@ have lim_ji: `((\y:real^2. lift (G (y$1) (y$2) / (y$1 * y$2))) --> lift (partial apply: REAL_LET_TRANS; exists `norm (t2 % basis j:real^N) + norm (t1 % basis i:real^N)`. rewrite NORM_TRIANGLE andTb !NORM_MUL !NORM_BASIS // !REAL_MUL_RID. by move: t_ineq yr; arith. -have lim_ij: `((\y:real^2. lift (G (y$1) (y$2) / (y$1 * y$2))) --> lift (partial i (partial j f) x)) +have lim_ij: `((\y:real^2. lift (G (y$1) (y$2) / (y$1 * y$2))) --> lift (partial i (partial j f) x)) (at (vec 0) within {y | &0 < y$1 /\ &0 < y$2})`. rewrite LIM_WITHIN => e e_gt0. move: (pc j i); rewrite "GEN_ALL continuous_at" !dist !o_THM => /(_ e_gt0) [d] [d0] e_ineq. @@ -1198,7 +1204,7 @@ by move/ineqs => /=. Qed. -Lemma domain_width p domain y w : `m_cell_domain domain y (w:real^N) ==> +Lemma domain_width p domain y w : `m_cell_domain domain y (w:real^N) ==> p IN interval [domain] ==> !i. i IN 1..dimindex (:N) ==> abs (p$i - y$i) <= w$i`. case: domain => x z; rewrite m_cell_domain => ineqs p_in i i_in. @@ -1206,7 +1212,7 @@ by move: p_in (ineqs i_in); rewrite "GEN_ALL IN_INTERVAL" -IN_NUMSEG => /(_ i_in Qed. -Lemma sum_swap1 g n : `sum (1..n) (\i. sum (i + 1..n) (\j. g i j)) = +Lemma sum_swap1 g n : `sum (1..n) (\i. sum (i + 1..n) (\j. g i j)) = sum (1..n) (\i. sum (1..i - 1) (\j. g j i))`. Proof. rewrite !SUM_SUM_PRODUCT ?FINITE_NUMSEG //= !IN_NUMSEG. @@ -1227,14 +1233,14 @@ Qed. (* Computation of the taylor error *) -Lemma m_taylor_error_eq f domain w error : `diff2c_domain domain f ==> +Lemma m_taylor_error_eq f domain w error : `diff2c_domain domain f ==> (m_taylor_error f domain (w:real^N) error <=> - (!x. x IN interval [domain] ==> + (!x. x IN interval [domain] ==> sum (1..dimindex (:N)) (\i. w$i * (w$i * abs (partial2 i i f x) + &2 * sum (1..i - 1) (\j. w$j * abs (partial2 j i f x)))) <= error))`. Proof. rewrite diff2c_domain m_taylor_error => d2f. -have eq: `!g1 g2. (!x. x IN interval [domain] ==> g1 x = g2 x) ==> +have eq: `!g1 g2. (!x. x IN interval [domain] ==> g1 x = g2 x) ==> ((!x. x IN interval [domain] ==> g1 x <= error) <=> (!x. x IN interval [domain] ==> g2 x <= error))`. move => g1 g2 eq; split => cond x Px. by rewrite -eq // cond. @@ -1269,7 +1275,7 @@ Lemma diff2_derivative2_bound domain y w p f dd_bound : p IN interval [domain] ==> diff2_domain domain f ==> m_taylor_error f domain w dd_bound ==> - (!t. interval_arith t (&0, &1) ==> + (!t. interval_arith t (&0, &1) ==> abs (nth_derivative 2 (f o (\t. y + t % (p - y))) t) <= dd_bound)`. Proof. rewrite diff2_domain m_taylor_error => domainH p_in df boundedH t t_in. @@ -1285,7 +1291,7 @@ rewrite -SUM_LMUL SUM_ABS_LE FINITE_NUMSEG andTb => j j_ineq /=. by rewrite !REAL_ABS_MUL !REAL_LE_MUL2 ?REAL_LE_MUL !REAL_ABS_POS // VECTOR_SUB_COMPONENT; rewrite (domain_width domainH p_in) // REAL_LE_REFL. Qed. - + Lemma m_taylor_error_lemma domain y w p f dd_bound : `m_cell_domain domain y (w:real^N) ==> @@ -1372,7 +1378,7 @@ Lemma diff2_derivative_partial_bound domain y w p f i d_bound : p IN interval [domain] ==> diff2_domain domain f ==> m_taylor_partial_error f i domain w d_bound ==> - (!t. interval_arith t (&0, &1) ==> + (!t. interval_arith t (&0, &1) ==> abs (derivative (partial i f o (\t. y + t % (p - y))) t) <= d_bound)`. Proof. rewrite diff2_domain m_taylor_partial_error => domainH p_in df boundedH t t_in. @@ -1510,7 +1516,7 @@ by rewrite DIFF_CHAIN_AT o_THM -"GEN_ALL HAS_REAL_FRECHET_DERIVATIVE_AT". Qed. -Lemma diff_uni_compose u f x : `lift o f differentiable at x ==> +Lemma diff_uni_compose u f x : `lift o f differentiable at x ==> u real_differentiable atreal (f x) ==> lift o u o f differentiable at x`. rewrite !differentiable real_differentiable => [] [f'] df [u'] du. @@ -1547,7 +1553,7 @@ Proof. by move => /diff2_abs_neg du df; rewrite diff2_uni_compose. Qed. (* pow *) Lemma diff2_pow_compose n : `diff2 f x ==> diff2 ((\x. x pow n) o f) x`. Proof. by move => df; rewrite diff2_uni_compose diff2_pow_x. Qed. - + (* inv *) Lemma diff2_inv_compose : `~(f x = &0) ==> diff2 f x ==> diff2 (inv o f) x`. Proof. by move => /diff2_inv du df; rewrite diff2_uni_compose. Qed. @@ -1763,7 +1769,7 @@ move: (open_contains_open_interval open_s xs `basis i:real^N`) => [a] [b] [ab0]; exists `real_interval (a, b)`; rewrite ab0 REAL_OPEN_REAL_INTERVAL !andTb => y y_in. by rewrite !o_THM /= eq // sub IN_IMAGE /=; exists y. Qed. - + Variables i j : `:num`. @@ -1785,7 +1791,7 @@ Proof. by move/(second_partial_scale f `--(&1)`); rewrite -!REAL_NEG_MINUS1. Qed (* Binary operations *) (* add *) -Lemma second_partial_add f g : `diff2 f x ==> diff2 g x ==> +Lemma second_partial_add f g : `diff2 f x ==> diff2 g x ==> partial2 i j (\x. f x + g x) x = partial2 i j f x + partial2 i j g x`. Proof. rewrite !diff2 => [] [s] [open_s] [xs] df [t] [open_t] [ys] dg. @@ -1798,14 +1804,14 @@ Qed. (* sub *) -Lemma second_partial_sub f g : `diff2 f x ==> diff2 g x ==> +Lemma second_partial_sub f g : `diff2 f x ==> diff2 g x ==> partial2 i j (\x. f x - g x) x = partial2 i j f x - partial2 i j g x`. -Proof. +Proof. by move => d2f d2g; rewrite real_sub second_partial_add ?diff2_neg // second_partial_neg // -real_sub. Qed. (* mul *) -Lemma second_partial_mul f g : `diff2 f x ==> diff2 g x ==> +Lemma second_partial_mul f g : `diff2 f x ==> diff2 g x ==> partial2 i j (\x. f x * g x) x = (partial2 i j f x * g x + partial j f x * partial i g x) + (partial i f x * partial j g x + f x * partial2 i j g x)`. rewrite !diff2 => [] [s] [open_s] [xs] df [t] [open_t] [ys] dg. @@ -1950,7 +1956,7 @@ rewrite open_r -r_def !IN_ELIM_THM /= o_THM LIFT_IN_IMAGE_LIFT; split; last firs move => y [z] [] [zs] fzt yz. by rewrite second_partial_uni_compose // yz d2s // d2t. split; last by exists x. -rewrite REAL_CONTINUOUS_ADD !REAL_CONTINUOUS_MUL //=; rewr ETA_AX; +rewrite REAL_CONTINUOUS_ADD !REAL_CONTINUOUS_MUL //=; rewr ETA_AX; rewrite ?(diff2_imp_partial_cont, d2s) //. rewrite andbT -[`nth_derivative 2 u _1`]o_THM; rewr ETA_AX. rewrite "GEN_ALL REAL_CONTINUOUS_CONTINUOUS1". @@ -2208,11 +2214,11 @@ move: (p2f i j) (p2g i j); rewrite -!"GEN_ALL REAL_CONTINUOUS_CONTINUOUS1" => p2 apply real_cont_at_local. move: d2f d2g; rewrite diff2_eq_diff2_on_open => [] [s] [open_s] [xs] d2f. rewrite diff2_eq_diff2_on_open => [] [t] [open_t] [xt] d2g. -exists `(\x. (partial2 j i f x * g x + partial i f x * partial j g x) + +exists `(\x. (partial2 j i f x * g x + partial i f x * partial j g x) + partial j f x * partial i g x + f x * partial2 j i g x)` `s INTER t`. rewrite !IN_INTER xs xt OPEN_INTER //=; split; last first. by move => y [ys yt]; rewrite second_partial_mul ?d2f ?d2g. -by rewrite !REAL_CONTINUOUS_ADD // !REAL_CONTINUOUS_MUL; rewr ETA_AX //; +by rewrite !REAL_CONTINUOUS_ADD // !REAL_CONTINUOUS_MUL; rewr ETA_AX //; rewrite ?p2gij ?p2fij /= ?diff2_imp_partial_cont ?diff2_imp_cont ?d2f ?d2g. Qed. @@ -2408,7 +2414,7 @@ End M_LinApprox. !x. x IN interval [domain] ==> all_n 1 dd_bounds_list (\i list_i. all_n 1 list_i (\j int. interval_arith (partial2 j i f x) int))`". -"let m_taylor_interval = +"let m_taylor_interval = new_definition `m_taylor_interval f domain y w f_bounds d_bounds_list dd_bounds_list <=> m_cell_domain domain y w /\ diff2c_domain domain f /\ @@ -2641,7 +2647,7 @@ have ->: `f y1 = g (&0)`. have y_eq : `y = y1 + (y$j - x$j) % basis j /\ y2 = y1 + (z$j - x$j) % basis j`. rewrite !CART_EQ; split => i i_in; rewrite VECTOR_ADD_COMPONENT VECTOR_MUL_COMPONENT BASIS_COMPONENT //; - case: (EXCLUDED_MIDDLE `i = j`) => /= ij; + case: (EXCLUDED_MIDDLE `i = j`) => /= ij; rewrite -?(y1_def, y2_def) !"GEN_ALL LAMBDA_BETA" //= ?REAL_MUL_RID ?REAL_SUB_ADD2 //. by rewrite ij /=; arith. by rewrite ij /=; arith. @@ -2760,7 +2766,7 @@ rewrite derivative_const /= derivative_mul ?REAL_DIFFERENTIABLE_ID ?REAL_DIFFERE by rewrite derivative_x derivative_const; arith. Qed. -Lemma partial_x k i : `k IN 1..dimindex (:N) ==> +Lemma partial_x k i : `k IN 1..dimindex (:N) ==> partial i (\x:real^N. x$k) = (\x. if i = k then &1 else &0)`. by move => k_ineq; rewrite partial_x_lemma BASIS_COMPONENT -?IN_NUMSEG //; arith. Qed. diff --git a/Formal_ineqs/taylor/theory/taylor_interval-compiled.hl b/Formal_ineqs/taylor/theory/taylor_interval-compiled.hl index 3e98bdfc..b494e6aa 100644 --- a/Formal_ineqs/taylor/theory/taylor_interval-compiled.hl +++ b/Formal_ineqs/taylor/theory/taylor_interval-compiled.hl @@ -1,7 +1,7 @@ -needs "lib/ssrbool.hl";; -needs "lib/ssrnat.hl";; -needs "arith/interval_arith.hl";; -needs "trig/matan.hl";; +needs "Formal_ineqs/lib/ssrbool.hl";; +needs "Formal_ineqs/lib/ssrnat.hl";; +needs "Formal_ineqs/arith/interval_arith.hl";; +needs "Formal_ineqs/trig/matan.hl";; needs "Multivariate/realanalysis.ml";; (* Module Taylor_interval*) @@ -17,15 +17,15 @@ prioritize_real();; let derivative = new_definition `derivative f = \y. @d. (f has_real_derivative d) (atreal y)`;; let nth_derivative = new_definition `nth_derivative n f = iter n derivative f`;; let nth_differentiable = define `(nth_differentiable 0 f x <=> f real_continuous atreal x) /\ - (nth_differentiable (SUC n) f x <=> nth_differentiable n f x /\ + (nth_differentiable (SUC n) f x <=> nth_differentiable n f x /\ nth_derivative n f real_differentiable atreal x)`;; let nth_differentiable_on = new_definition `nth_differentiable_on n s f <=> !x. x IN s ==> nth_differentiable n f x`;; let nth_differentiable_on_int = new_definition `nth_differentiable_on_int n int f <=> !x. interval_arith x int ==> nth_differentiable n f x`;; -let nth_diff_weak = new_definition `nth_diff_weak n f x <=> f real_continuous atreal x /\ +let nth_diff_weak = new_definition `nth_diff_weak n f x <=> f real_continuous atreal x /\ ?F. F 0 = f /\ !i. i < n ==> (F i has_real_derivative F (SUC i) x) (atreal x)`;; -let nth_diff_strong = new_definition `nth_diff_strong n f x <=> +let nth_diff_strong = new_definition `nth_diff_strong n f x <=> ?s. real_open s /\ x IN s /\ nth_differentiable_on n s f`;; let nth_diff_strong_int = new_definition `nth_diff_strong_int n int f <=> !x. interval_arith x int ==> nth_diff_strong n f x`;; @@ -150,7 +150,7 @@ let nth_differentiable_cond = Sections.section_proof ["n";"f";"x"] (* Lemma nth_differentiable_on_cond *) let nth_differentiable_on_cond = Sections.section_proof ["n";"s";"f"] `nth_differentiable_on n s f ==> - !x. x IN s ==> + !x. x IN s ==> !i. i < n ==> (nth_derivative i f has_real_derivative (nth_derivative (SUC i) f x)) (atreal x)` [ ((((use_arg_then2 ("nth_differentiable_on", [nth_differentiable_on]))(thm_tac (new_rewrite [] [])))) THEN (move ["cond"]) THEN (move ["x"])); @@ -735,7 +735,7 @@ let derivative_sub = Sections.section_proof [] (* Lemma derivative_div *) let derivative_div = Sections.section_proof [] -`~(g x = &0) ==> +`~(g x = &0) ==> derivative (\x. f x / g x) x = (derivative f x * g x - f x * derivative g x) / (g x * g x)` [ ((BETA_TAC THEN (move ["gn0"])) THEN (((use_arg_then2 ("derivative_unique", [derivative_unique])) (disch_tac [])) THEN (clear_assumption "derivative_unique") THEN (DISCH_THEN apply_tac)) THEN ((((use_arg_then2 ("REAL_POW_2", [REAL_POW_2]))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (((use_arg_then2 ("HAS_REAL_DERIVATIVE_DIV_ATREAL", [HAS_REAL_DERIVATIVE_DIV_ATREAL]))(thm_tac (new_rewrite [] [])))) THEN ((TRY done_tac)))); @@ -865,7 +865,7 @@ Sections.begin_section "NthDerivativeArith";; (* Lemma nth_derivative_scale_strong *) let nth_derivative_scale_strong = Sections.section_proof ["c";"i";"x"] `interval_arith x int ==> i <= n ==> - ?s. real_open s /\ x IN s /\ + ?s. real_open s /\ x IN s /\ !y. y IN s ==> nth_derivative i (\y. c * f y) y = c * nth_derivative i f y` [ ((((use_arg_then2 ("df", [])) (disch_tac [])) THEN (clear_assumption "df") THEN BETA_TAC) THEN (((((use_arg_then2 ("nth_diff_strong_int", [nth_diff_strong_int]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("nth_diff_strong", [nth_diff_strong]))(thm_tac (new_rewrite [] []))))) THEN (move ["df"]))); @@ -886,11 +886,11 @@ let nth_derivative_scale_strong = Sections.section_proof ["c";"i";"x"] (* Lemma nth_derivative_scale_strong_all *) let nth_derivative_scale_strong_all = Sections.section_proof ["c";"x"] `interval_arith x int ==> - ?s. real_open s /\ x IN s /\ + ?s. real_open s /\ x IN s /\ !i y. i <= n /\ y IN s ==> nth_derivative i (\y. c * f y) y = c * nth_derivative i f y` [ (BETA_TAC THEN (DISCH_THEN (fun snd_th -> (fun arg_tac -> (use_arg_then2 ("nth_derivative_scale_strong", [nth_derivative_scale_strong])) (fun fst_arg -> (use_arg_then2 ("c", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (thm_tac (match_mp_then snd_th MP_TAC)))) THEN (move ["h"])); - ((fun arg_tac -> arg_tac (Arg_term (`\i s. real_open s /\ x IN s /\ + ((fun arg_tac -> arg_tac (Arg_term (`\i s. real_open s /\ x IN s /\ (!y. y IN s ==> nth_derivative i (\y. c * f y) y = c * nth_derivative i f y)`))) (term_tac (set_tac "P"))); ((fun arg_tac -> arg_tac (Arg_term (`!i. i <= n:num ==> P i ((@) (P i))`))) (term_tac (have_gen_tac [](move ["sel_P"])))); ((BETA_TAC THEN (move ["i"]) THEN (move ["i_le_n"])) THEN ((use_arg_then2 ("SELECT_AX", [SELECT_AX])) (thm_tac apply_tac))); @@ -944,7 +944,7 @@ let nth_diff_scale = Sections.section_proof ["c"] (* Lemma nth_derivative_add_strong *) let nth_derivative_add_strong = Sections.section_proof ["i";"x"] `interval_arith x int ==> i <= n ==> - ?s. real_open s /\ x IN s /\ + ?s. real_open s /\ x IN s /\ !y. y IN s ==> nth_derivative i (\y. f y + g y) y = nth_derivative i f y + nth_derivative i g y` [ ((((use_arg_then2 ("dg", [])) (disch_tac [])) THEN (clear_assumption "dg") THEN ((use_arg_then2 ("df", [])) (disch_tac [])) THEN (clear_assumption "df") THEN BETA_TAC) THEN (((repeat_tactic 1 9 (((use_arg_then2 ("nth_diff_strong_int", [nth_diff_strong_int]))(thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("nth_diff_strong", [nth_diff_strong]))(thm_tac (new_rewrite [] [])))))) THEN (move ["df"]) THEN (move ["dg"]) THEN (move ["ineq"]))); @@ -968,12 +968,12 @@ let nth_derivative_add_strong = Sections.section_proof ["i";"x"] (* Lemma nth_derivative_add_strong_all *) let nth_derivative_add_strong_all = Sections.section_proof ["x"] `interval_arith x int ==> - ?s. real_open s /\ x IN s /\ - !i y. i <= n /\ y IN s ==> + ?s. real_open s /\ x IN s /\ + !i y. i <= n /\ y IN s ==> nth_derivative i (\y. f y + g y) y = nth_derivative i f y + nth_derivative i g y` [ (BETA_TAC THEN (DISCH_THEN (fun snd_th -> (use_arg_then2 ("nth_derivative_add_strong", [nth_derivative_add_strong])) (thm_tac (match_mp_then snd_th MP_TAC)))) THEN (move ["h"])); - ((fun arg_tac -> arg_tac (Arg_term (`\i s. real_open s /\ x IN s /\ + ((fun arg_tac -> arg_tac (Arg_term (`\i s. real_open s /\ x IN s /\ (!y. y IN s ==> nth_derivative i (\y. f y + g y) y = nth_derivative i f y + nth_derivative i g y)`))) (term_tac (set_tac "P"))); ((fun arg_tac -> arg_tac (Arg_term (`!i. i <= n:num ==> P i ((@) (P i))`))) (term_tac (have_gen_tac [](move ["sel_P"])))); ((BETA_TAC THEN (move ["i"]) THEN (move ["i_le_n"])) THEN ((use_arg_then2 ("SELECT_AX", [SELECT_AX])) (thm_tac apply_tac))); @@ -1029,7 +1029,7 @@ let nth_diff_add = Sections.section_proof [] (* Lemma nth_derivative_sub_strong *) let nth_derivative_sub_strong = Sections.section_proof ["i";"x"] `interval_arith x int ==> i <= n ==> - ?s. real_open s /\ x IN s /\ + ?s. real_open s /\ x IN s /\ !y. y IN s ==> nth_derivative i (\y. f y - g y) y = nth_derivative i f y - nth_derivative i g y` [ ((((use_arg_then2 ("dg", [])) (disch_tac [])) THEN (clear_assumption "dg") THEN ((use_arg_then2 ("df", [])) (disch_tac [])) THEN (clear_assumption "df") THEN BETA_TAC) THEN (((repeat_tactic 1 9 (((use_arg_then2 ("nth_diff_strong_int", [nth_diff_strong_int]))(thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("nth_diff_strong", [nth_diff_strong]))(thm_tac (new_rewrite [] [])))))) THEN (move ["df"]) THEN (move ["dg"]) THEN (move ["ineq"]))); @@ -1053,12 +1053,12 @@ let nth_derivative_sub_strong = Sections.section_proof ["i";"x"] (* Lemma nth_derivative_sub_strong_all *) let nth_derivative_sub_strong_all = Sections.section_proof ["x"] `interval_arith x int ==> - ?s. real_open s /\ x IN s /\ - !i y. i <= n /\ y IN s ==> + ?s. real_open s /\ x IN s /\ + !i y. i <= n /\ y IN s ==> nth_derivative i (\y. f y - g y) y = nth_derivative i f y - nth_derivative i g y` [ (BETA_TAC THEN (DISCH_THEN (fun snd_th -> (use_arg_then2 ("nth_derivative_sub_strong", [nth_derivative_sub_strong])) (thm_tac (match_mp_then snd_th MP_TAC)))) THEN (move ["h"])); - ((fun arg_tac -> arg_tac (Arg_term (`\i s. real_open s /\ x IN s /\ + ((fun arg_tac -> arg_tac (Arg_term (`\i s. real_open s /\ x IN s /\ (!y. y IN s ==> nth_derivative i (\y. f y - g y) y = nth_derivative i f y - nth_derivative i g y)`))) (term_tac (set_tac "P"))); ((fun arg_tac -> arg_tac (Arg_term (`!i. i <= n:num ==> P i ((@) (P i))`))) (term_tac (have_gen_tac [](move ["sel_P"])))); ((BETA_TAC THEN (move ["i"]) THEN (move ["i_le_n"])) THEN ((use_arg_then2 ("SELECT_AX", [SELECT_AX])) (thm_tac apply_tac))); @@ -1114,8 +1114,8 @@ let nth_diff_sub = Sections.section_proof [] (* Lemma nth_derivative_mul_strong *) let nth_derivative_mul_strong = Sections.section_proof ["i";"x"] `interval_arith x int ==> i <= n ==> - ?s. real_open s /\ x IN s /\ - !y. y IN s ==> nth_derivative i (\y. f y * g y) y = + ?s. real_open s /\ x IN s /\ + !y. y IN s ==> nth_derivative i (\y. f y * g y) y = sum (0..i) (\k. &(binom (i, k)) * nth_derivative k f y * nth_derivative (i - k) g y)` [ ((((use_arg_then2 ("dg", [])) (disch_tac [])) THEN (clear_assumption "dg") THEN ((use_arg_then2 ("df", [])) (disch_tac [])) THEN (clear_assumption "df") THEN BETA_TAC) THEN (((repeat_tactic 1 9 (((use_arg_then2 ("nth_diff_strong_int", [nth_diff_strong_int]))(thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("nth_diff_strong", [nth_diff_strong]))(thm_tac (new_rewrite [] [])))))) THEN (move ["df"]) THEN (move ["dg"]) THEN (move ["ineq"]))); @@ -1144,7 +1144,7 @@ let nth_derivative_mul_strong = Sections.section_proof ["i";"x"] ((((use_arg_then2 ("i_lt_n", [])) (disch_tac [])) THEN (clear_assumption "i_lt_n") THEN ((use_arg_then2 ("ineq", [])) (disch_tac [])) THEN (clear_assumption "ineq") THEN BETA_TAC) THEN (arith_tac) THEN (done_tac)); (((fun arg_tac -> arg_tac (Arg_term (`nth_derivative (SUC (i - k)) g y`))) (term_tac exists_tac)) THEN (((fun arg_tac -> (use_arg_then2 ("diff_g", [])) (fun fst_arg -> (use_arg_then2 ("ysg", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg)))(thm_tac (new_rewrite [] []))))); ((((use_arg_then2 ("i_lt_n", [])) (disch_tac [])) THEN (clear_assumption "i_lt_n") THEN ((use_arg_then2 ("ineq", [])) (disch_tac [])) THEN (clear_assumption "ineq") THEN BETA_TAC) THEN (arith_tac) THEN (done_tac)); - ((fun arg_tac -> arg_tac (Arg_term (`!k. k IN 0..i ==> + ((fun arg_tac -> arg_tac (Arg_term (`!k. k IN 0..i ==> (\y. &(binom (i, k)) * nth_derivative k f y * nth_derivative (i - k) g y) real_differentiable atreal y`))) (term_tac (have_gen_tac [](move ["diff_cond2"])))); (BETA_TAC THEN (move ["k"]) THEN (move ["k_in"]) THEN (simp_tac)); @@ -1177,14 +1177,14 @@ let nth_derivative_mul_strong = Sections.section_proof ["i";"x"] (* Lemma nth_derivative_mul_strong_all *) let nth_derivative_mul_strong_all = Sections.section_proof ["x"] `interval_arith x int ==> - ?s. real_open s /\ x IN s /\ - !i y. i <= n /\ y IN s ==> - nth_derivative i (\y. f y * g y) y = + ?s. real_open s /\ x IN s /\ + !i y. i <= n /\ y IN s ==> + nth_derivative i (\y. f y * g y) y = sum (0..i) (\k. &(binom (i, k)) * nth_derivative k f y * nth_derivative (i - k) g y)` [ (BETA_TAC THEN (DISCH_THEN (fun snd_th -> (use_arg_then2 ("nth_derivative_mul_strong", [nth_derivative_mul_strong])) (thm_tac (match_mp_then snd_th MP_TAC)))) THEN (move ["h"])); - ((fun arg_tac -> arg_tac (Arg_term (`\i s. real_open s /\ x IN s /\ - (!y. y IN s ==> nth_derivative i (\y. f y * g y) y = + ((fun arg_tac -> arg_tac (Arg_term (`\i s. real_open s /\ x IN s /\ + (!y. y IN s ==> nth_derivative i (\y. f y * g y) y = sum (0..i) (\k. &(binom (i, k)) * nth_derivative k f y * nth_derivative (i - k) g y))`))) (term_tac (set_tac "P"))); ((fun arg_tac -> arg_tac (Arg_term (`!i. i <= n:num ==> P i ((@) (P i))`))) (term_tac (have_gen_tac [](move ["sel_P"])))); ((BETA_TAC THEN (move ["i"]) THEN (move ["i_le_n"])) THEN ((use_arg_then2 ("SELECT_AX", [SELECT_AX])) (thm_tac apply_tac))); @@ -1207,7 +1207,7 @@ let nth_derivative_mul_strong_all = Sections.section_proof ["x"] (* Lemma nth_derivative_mul *) let nth_derivative_mul = Sections.section_proof ["i";"x"] `interval_arith x int ==> i <= n ==> - nth_derivative i (\x. f x * g x) x = + nth_derivative i (\x. f x * g x) x = sum (0..i) (\k. &(binom (i, k)) * nth_derivative k f x * nth_derivative (i - k) g x)` [ (BETA_TAC THEN (move ["ineq"]) THEN (move ["i_le_n"])); @@ -1353,13 +1353,13 @@ let nth_derivative_mul_strong_all = Sections.finalize_theorem nth_derivative_mul let nth_derivative_mul = Sections.finalize_theorem nth_derivative_mul;; let nth_diff_mul = Sections.finalize_theorem nth_diff_mul;; Sections.end_section "NthDerivatives";; -let lin_approx = new_definition `lin_approx f x f_bounds df_bounds <=> +let lin_approx = new_definition `lin_approx f x f_bounds df_bounds <=> interval_arith (f x) f_bounds /\ (?f'. (f has_real_derivative f') (atreal x) /\ interval_arith f' df_bounds)`;; let has_bounded_second_derivative = new_definition `has_bounded_second_derivative f int dd_bounds <=> - nth_diff_strong_int 2 int f /\ + nth_diff_strong_int 2 int f /\ bounded_on_int (nth_derivative 2 f) int dd_bounds`;; -let taylor_interval = new_definition +let taylor_interval = new_definition `taylor_interval f x y z w f_bounds df_bounds ddf_bounds <=> x <= y /\ y <= z /\ y - x <= w /\ z - y <= w /\ lin_approx f y f_bounds df_bounds /\ @@ -1496,7 +1496,7 @@ let f_continuous = Sections.section_proof [] (* Lemma taylor_error *) let taylor_error = Sections.section_proof ["t"] -`x <= t /\ t <= z ==> +`x <= t /\ t <= z ==> abs (f t - f y) <= w * iabs df_bounds + w * w * dd_bound / &2` [ (BETA_TAC THEN (move ["t_ineqs"])); @@ -1524,7 +1524,7 @@ let taylor_error = Sections.section_proof ["t"] ((((use_arg_then2 ("andTb", [andTb]))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("andbT", [andbT]))(thm_tac (new_rewrite [] []))))) THEN (((use_arg_then2 ("REAL_MUL_SYM", [REAL_MUL_SYM]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("REAL_POW_2", [REAL_POW_2]))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (((use_arg_then2 ("REAL_LE_MUL2", [REAL_LE_MUL2]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("REAL_LE_REFL", [REAL_LE_REFL]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("dd_prop", []))(thm_tac (new_rewrite [] []))))); (((((use_arg_then2 ("REAL_LE_POW_2", [REAL_LE_POW_2]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("andTb", [andTb]))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("andbT", [andbT]))(thm_tac (new_rewrite [] []))))) THEN (((use_arg_then2 ("REAL_LE_SQUARE_ABS", [REAL_LE_SQUARE_ABS]))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (((use_arg_then2 ("REAL_ABS_ABS", [REAL_ABS_ABS]))(thm_tac (new_rewrite [] []))))) THEN (((use_arg_then2 ("abs_ty", [])) (disch_tac [])) THEN (clear_assumption "abs_ty") THEN BETA_TAC) THEN (arith_tac) THEN (done_tac)); ((fun arg_tac -> arg_tac (Arg_term (`\i. if i = 0 then f else if i = 1 then f' else if i = 2 then f'' else I`))) (term_tac (set_tac "Df"))); - ((fun arg_tac -> arg_tac (Arg_theorem (ARITH_RULE `0 + 1 = 1 /\ ~(1 = 0) /\ 1 + 1 = 2 /\ + ((fun arg_tac -> arg_tac (Arg_theorem (ARITH_RULE `0 + 1 = 1 /\ ~(1 = 0) /\ 1 + 1 = 2 /\ ~(2 = 1) /\ ~(2 = 1) /\ ~(2 = 0)`))) (fun arg -> thm_tac MP_TAC arg THEN (move ["arith"]))); ((fun arg_tac -> (fun arg_tac -> (fun arg_tac -> (fun arg_tac -> (use_arg_then2 ("REAL_TAYLOR", [REAL_TAYLOR])) (fun fst_arg -> (use_arg_then2 ("Df", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (fun arg_tac -> arg_tac (Arg_term (`1`))) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (fun arg_tac -> arg_tac (Arg_term (`real_interval [x, z]`))) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (use_arg_then2 ("dd_bound", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun arg -> thm_tac MP_TAC arg THEN ALL_TAC)); (ANTS_TAC); @@ -1564,7 +1564,7 @@ let taylor_lower_bound = Sections.section_proof [] (* Lemma taylor_derivative_error *) let taylor_derivative_error = Sections.section_proof [] -`!t. x <= t /\ t <= z ==> +`!t. x <= t /\ t <= z ==> abs (derivative f t - derivative f y) <= w * dd_bound` [ (BETA_TAC THEN (move ["t"]) THEN (move ["t_ineqs"])); @@ -1581,7 +1581,7 @@ let taylor_derivative_error = Sections.section_proof [] ((((fun arg_tac -> (use_arg_then2 ("df", [])) (fun fst_arg -> (use_arg_then2 ("y", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (disch_tac [])) THEN BETA_TAC) THEN ((((fun arg_tac -> (use_arg_then2 ("pair_eq", [pair_eq])) (fun fst_arg -> (use_arg_then2 ("ddf_bounds", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg)))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("interval_arith", [interval_arith]))(thm_tac (new_rewrite [] []))))) THEN (((use_arg_then2 ("iabs", [iabs]))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("domain_ineqs", []))(thm_tac (new_rewrite [] [])))))) THEN (arith_tac) THEN (done_tac)); ((BETA_TAC THEN (move ["p"]) THEN (DISCH_THEN (fun snd_th -> (use_arg_then2 ("df", [])) (thm_tac (match_mp_then snd_th MP_TAC))))) THEN ((((use_arg_then2 ("pair_eq", [pair_eq]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("interval_arith", [interval_arith]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("iabs", [iabs]))(thm_tac (new_rewrite [] []))))) THEN (arith_tac) THEN (done_tac)); ((fun arg_tac -> arg_tac (Arg_term (`\i. if i = 0 then f' else if i = 1 then f'' else I`))) (term_tac (set_tac "Df"))); - ((fun arg_tac -> arg_tac (Arg_theorem (ARITH_RULE `0 + 1 = 1 /\ ~(1 = 0) /\ 1 + 1 = 2 /\ + ((fun arg_tac -> arg_tac (Arg_theorem (ARITH_RULE `0 + 1 = 1 /\ ~(1 = 0) /\ 1 + 1 = 2 /\ ~(2 = 1) /\ ~(2 = 1) /\ ~(2 = 0)`))) (fun arg -> thm_tac MP_TAC arg THEN (move ["arith"]))); ((fun arg_tac -> (fun arg_tac -> (fun arg_tac -> (fun arg_tac -> (use_arg_then2 ("REAL_TAYLOR", [REAL_TAYLOR])) (fun fst_arg -> (use_arg_then2 ("Df", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (fun arg_tac -> arg_tac (Arg_term (`0`))) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (fun arg_tac -> arg_tac (Arg_term (`real_interval [x, z]`))) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun fst_arg -> (use_arg_then2 ("dd_bound", [])) (fun snd_arg -> combine_args_then arg_tac fst_arg snd_arg))) (fun arg -> thm_tac MP_TAC arg THEN ALL_TAC)); (ANTS_TAC); @@ -2114,7 +2114,7 @@ let second_derivative_pow3_x = Sections.section_proof [] (* Lemma second_derivative_atn_eq *) let second_derivative_atn_eq = Sections.section_proof ["x"] -`((\x. inv (&1 + x pow 2)) has_real_derivative +`((\x. inv (&1 + x pow 2)) has_real_derivative (-- &2 * x) * inv (&1 + x pow 2) pow 2) (atreal x)` [ (((((use_arg_then2 ("REAL_POW_INV", [REAL_POW_INV]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("real_div", [real_div]))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (((use_arg_then2 ("REAL_NEG_LMUL", [REAL_NEG_LMUL]))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (((use_arg_then2 ("HAS_REAL_DERIVATIVE_INV_ATREAL", [HAS_REAL_DERIVATIVE_INV_ATREAL]))(thm_tac (new_rewrite [] []))))) THEN (split_tac)); @@ -2321,7 +2321,7 @@ let real_powS = Sections.section_proof ["x";"n"] (* Lemma second_derivative_acs *) let second_derivative_acs = Sections.section_proof ["x"] -`abs x < &1 ==> +`abs x < &1 ==> nth_derivative 2 acs x = --(x / sqrt ((&1 - x * x) pow 3))` [ ((((use_arg_then2 ("nth_derivative2", [nth_derivative2]))(thm_tac (new_rewrite [] [])))) THEN (move ["x_ineq"])); @@ -2376,7 +2376,7 @@ let diff2_acs = Sections.section_proof ["x"] (* Lemma second_derivative_asn *) let second_derivative_asn = Sections.section_proof ["x"] -`abs x < &1 ==> +`abs x < &1 ==> nth_derivative 2 asn x = x / sqrt ((&1 - x * x) pow 3)` [ ((((use_arg_then2 ("nth_derivative2", [nth_derivative2]))(thm_tac (new_rewrite [] [])))) THEN (move ["x_ineq"])); @@ -2475,7 +2475,7 @@ let REAL_CONTINUOUS_OPEN_PREIMAGE = Sections.section_proof ["f";"s";"t"] [ (BETA_TAC THEN (move ["f_cont"]) THEN (move ["open_s"]) THEN (move ["open_t"])); (((use_arg_then2 ("REAL_OPEN", [REAL_OPEN]))(thm_tac (new_rewrite [] [])))); - (((THENL_ROT 1)) ((fun arg_tac -> arg_tac (Arg_term (`IMAGE lift {x | x IN s /\ f x IN t} = + (((THENL_ROT 1)) ((fun arg_tac -> arg_tac (Arg_term (`IMAGE lift {x | x IN s /\ f x IN t} = {x | x IN (IMAGE lift s) /\ (lift o f o drop) x IN (IMAGE lift t)}`))) (term_tac (have_gen_tac [](((conv_thm_tac DISCH_THEN)(thm_tac (new_rewrite [] [])))))))); (((((use_arg_then2 ("CONTINUOUS_OPEN_PREIMAGE", [CONTINUOUS_OPEN_PREIMAGE]))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("REAL_OPEN", [REAL_OPEN]))(gsym_then (thm_tac (new_rewrite [] []))))))) THEN (((use_arg_then2 ("ETA_AX", [ETA_AX]))(fun arg -> ONCE_REWRITE_TAC[get_arg_thm arg]))) THEN (((use_arg_then2 ("REAL_CONTINUOUS_ON", [REAL_CONTINUOUS_ON]))(gsym_then (thm_tac (new_rewrite [] []))))) THEN (done_tac)); ((((((use_arg_then2 ("EXTENSION", [EXTENSION]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("IN_IMAGE_LIFT_DROP", [IN_IMAGE_LIFT_DROP]))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("IN_ELIM_THM", [IN_ELIM_THM]))(thm_tac (new_rewrite [] []))))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("IN_IMAGE_LIFT_DROP", [IN_IMAGE_LIFT_DROP]))(thm_tac (new_rewrite [] [])))))) THEN (move ["z"]) THEN (simp_tac)) THEN (split_tac)); @@ -2488,7 +2488,7 @@ let REAL_CONTINUOUS_OPEN_PREIMAGE = Sections.section_proof ["f";"s";"t"] (* Lemma second_derivative_compose *) let second_derivative_compose = Sections.section_proof ["f";"g";"x"] `nth_diff_strong 2 g x ==> nth_diff_strong 2 f (g x) ==> - nth_derivative 2 (\x. f (g x)) x = + nth_derivative 2 (\x. f (g x)) x = nth_derivative 2 f (g x) * (derivative g x) pow 2 + derivative f (g x) * nth_derivative 2 g x` [ (BETA_TAC THEN (move ["dg"]) THEN (move ["df"])); @@ -2602,7 +2602,7 @@ let second_derivative_compose_abs_neg = Sections.section_proof [] (* Lemma second_derivative_compose_pow *) let second_derivative_compose_pow = Sections.section_proof ["n"] `nth_derivative 2 (\x. f x pow n) x = - &n * (nth_derivative 2 f x * f x pow (n - 1) + &n * (nth_derivative 2 f x * f x pow (n - 1) + &(n - 1) * f x pow (n - 2) * derivative f x pow 2)` [ ((fun arg_tac -> arg_tac (Arg_term (`\x. x pow n`))) (term_tac (set_tac "g"))); @@ -2632,7 +2632,7 @@ let second_derivative_compose_pow3 = Sections.section_proof [] (* Lemma second_derivative_compose_atn *) let second_derivative_compose_atn = Sections.section_proof [] -`nth_derivative 2 (\x. atn (f x)) x = +`nth_derivative 2 (\x. atn (f x)) x = (nth_derivative 2 f x * (&1 + f x * f x) - &2 * f x * derivative f x pow 2) / (&1 + f x * f x) pow 2` [ ((((use_arg_then2 ("second_derivative_compose", [second_derivative_compose]))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 0 10 (((use_arg_then2 ("diff2_atn", [diff2_atn]))(thm_tac (new_rewrite [] []))))) THEN ((TRY done_tac)) THEN (((use_arg_then2 ("nth_derivative2", [nth_derivative2]))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("second_derivative_atn", [second_derivative_atn]))(thm_tac (new_rewrite [] [])))) THEN (simp_tac) THEN (((use_arg_then2 ("derivative_atn", [derivative_atn]))(thm_tac (new_rewrite [] [])))) THEN (simp_tac)); @@ -2649,7 +2649,7 @@ let second_derivative_compose_atn = Sections.section_proof [] (* Lemma second_derivative_compose_cos *) let second_derivative_compose_cos = Sections.section_proof [] -`nth_derivative 2 (\x. cos (f x)) x = +`nth_derivative 2 (\x. cos (f x)) x = --(nth_derivative 2 f x * sin (f x) + cos (f x) * derivative f x pow 2)` [ ((((use_arg_then2 ("second_derivative_compose", [second_derivative_compose]))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 0 10 (((use_arg_then2 ("diff2_cos", [diff2_cos]))(thm_tac (new_rewrite [] []))))) THEN ((TRY done_tac)) THEN (((use_arg_then2 ("second_derivative_cos", [second_derivative_cos]))(thm_tac (new_rewrite [] [])))) THEN (simp_tac) THEN (((use_arg_then2 ("derivative_cos", [derivative_cos]))(thm_tac (new_rewrite [] [])))) THEN (simp_tac)); @@ -2658,7 +2658,7 @@ let second_derivative_compose_cos = Sections.section_proof [] (* Lemma second_derivative_compose_sin *) let second_derivative_compose_sin = Sections.section_proof [] -`nth_derivative 2 (\x. sin (f x)) x = +`nth_derivative 2 (\x. sin (f x)) x = nth_derivative 2 f x * cos (f x) - sin (f x) * derivative f x pow 2` [ ((((use_arg_then2 ("second_derivative_compose", [second_derivative_compose]))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 0 10 (((use_arg_then2 ("diff2_sin", [diff2_sin]))(thm_tac (new_rewrite [] []))))) THEN ((TRY done_tac)) THEN (((use_arg_then2 ("second_derivative_sin", [second_derivative_sin]))(thm_tac (new_rewrite [] [])))) THEN (simp_tac) THEN (((use_arg_then2 ("derivative_sin", [derivative_sin]))(thm_tac (new_rewrite [] [])))) THEN (simp_tac)); @@ -2667,7 +2667,7 @@ let second_derivative_compose_sin = Sections.section_proof [] (* Lemma second_derivative_compose_exp *) let second_derivative_compose_exp = Sections.section_proof [] -`nth_derivative 2 (\x. exp (f x)) x = +`nth_derivative 2 (\x. exp (f x)) x = nth_derivative 2 f x * exp (f x) + exp (f x) * derivative f x pow 2` [ ((((use_arg_then2 ("second_derivative_compose", [second_derivative_compose]))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 0 10 (((use_arg_then2 ("diff2_exp", [diff2_exp]))(thm_tac (new_rewrite [] []))))) THEN ((TRY done_tac)) THEN (((use_arg_then2 ("second_derivative_exp", [second_derivative_exp]))(thm_tac (new_rewrite [] [])))) THEN (simp_tac) THEN (((use_arg_then2 ("derivative_exp", [derivative_exp]))(thm_tac (new_rewrite [] [])))) THEN (simp_tac)); @@ -2721,7 +2721,7 @@ let second_derivative_compose_sqrt = Sections.section_proof [] let second_derivative_compose_acs = Sections.section_proof [] `abs (f x) < &1 ==> nth_derivative 2 (\x. acs (f x)) x = - -- ((nth_derivative 2 f x * (&1 - f x * f x) + f x * derivative f x pow 2) / + -- ((nth_derivative 2 f x * (&1 - f x * f x) + f x * derivative f x pow 2) / sqrt ((&1 - f x * f x) pow 3))` [ (BETA_TAC THEN (move ["f_ineq"])); @@ -2742,7 +2742,7 @@ let second_derivative_compose_acs = Sections.section_proof [] let second_derivative_compose_asn = Sections.section_proof [] `abs (f x) < &1 ==> nth_derivative 2 (\x. asn (f x)) x = - ((nth_derivative 2 f x * (&1 - f x * f x) + f x * derivative f x pow 2) / + ((nth_derivative 2 f x * (&1 - f x * f x) + f x * derivative f x pow 2) / sqrt ((&1 - f x * f x) pow 3))` [ (BETA_TAC THEN (move ["f_ineq"])); @@ -2922,7 +2922,7 @@ let second_derivative_div = Sections.section_proof ["f";"g";"x"] `~(g x = &0) ==> nth_diff_strong 2 f x ==> nth_diff_strong 2 g x ==> - nth_derivative 2 (\x. f x / g x) x = + nth_derivative 2 (\x. f x / g x) x = ((nth_derivative 2 f x * g x - f x * nth_derivative 2 g x) * g x - &2 * derivative g x * (derivative f x * g x - f x * derivative g x)) / (g x pow 3)` [ @@ -2934,7 +2934,7 @@ let second_derivative_div = Sections.section_proof ["f";"g";"x"] ((((use_arg_then2 ("diff_g", [])) (disch_tac [])) THEN (clear_assumption "diff_g") THEN BETA_TAC) THEN ((((use_arg_then2 ("nth_diff_strong2_eq", [nth_diff_strong2_eq]))(thm_tac (new_rewrite [] [])))) THEN ALL_TAC THEN (case THEN (move ["s"])) THEN (case THEN (move ["open_s"])) THEN (case THEN (move ["xs"])) THEN (((conv_thm_tac DISCH_THEN)(thm_tac (new_rewrite [] [])))) THEN ((TRY done_tac))) THEN (done_tac)); ((((use_arg_then2 ("ddf_def", []))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("ddg_def", []))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("df_def", []))(thm_tac (new_rewrite [] [])))) THEN (((use_arg_then2 ("dg_def", []))(thm_tac (new_rewrite [] [])))) THEN (repeat_tactic 1 9 (((use_arg_then2 ("real_div", [real_div]))(thm_tac (new_rewrite [] [])))))); ((fun arg_tac -> arg_tac (Arg_term (`_1 + _2`))) (term_tac (set_tac "lhs"))); - (((fun arg_tac -> arg_tac (Arg_theorem (REAL_RING `!f g x. ((ddf * g x - f x * ddg) * g x - &2 * dg * (df * g x - f x * dg)) * + (((fun arg_tac -> arg_tac (Arg_theorem (REAL_RING `!f g x. ((ddf * g x - f x * ddg) * g x - &2 * dg * (df * g x - f x * dg)) * inv (g x pow 3) = f x * (&2 * dg pow 2 - ddg * g x) * inv (g x pow 3) + &2 * df * --(g x * inv (g x pow 3)) * dg + @@ -3004,7 +3004,7 @@ let second_derivative_compose_bounds = Sections.section_proof ["f";"g";"int";"g_ `nth_diff_strong_int 2 int g ==> bounded_on_int g int g_bounds ==> nth_diff_strong_int 2 g_bounds f ==> - bounded_on_int (\x. nth_derivative 2 f (g x) * derivative g x pow 2 + + bounded_on_int (\x. nth_derivative 2 f (g x) * derivative g x pow 2 + derivative f (g x) * nth_derivative 2 g x) int dd_bounds ==> has_bounded_second_derivative (\x. f (g x)) int dd_bounds` [ @@ -3165,7 +3165,7 @@ let second_derivative_cos_bounds = Sections.section_proof ["dd_bounds"] (* Lemma second_derivative_compose_cos_bounds *) let second_derivative_compose_cos_bounds = Sections.section_proof ["f";"dd_bounds"] `nth_diff_strong_int 2 int f ==> - bounded_on_int (\x. --(nth_derivative 2 f x * sin (f x) + bounded_on_int (\x. --(nth_derivative 2 f x * sin (f x) + cos (f x) * derivative f x pow 2)) int dd_bounds ==> has_bounded_second_derivative (\x. cos (f x)) int dd_bounds` [ @@ -3187,7 +3187,7 @@ let second_derivative_sin_bounds = Sections.section_proof ["dd_bounds"] (* Lemma second_derivative_compose_sin_bounds *) let second_derivative_compose_sin_bounds = Sections.section_proof ["f";"dd_bounds"] `nth_diff_strong_int 2 int f ==> - bounded_on_int (\x. nth_derivative 2 f x * cos (f x) + bounded_on_int (\x. nth_derivative 2 f x * cos (f x) - sin (f x) * derivative f x pow 2) int dd_bounds ==> has_bounded_second_derivative (\x. sin (f x)) int dd_bounds` [ @@ -3209,7 +3209,7 @@ let second_derivative_exp_bounds = Sections.section_proof ["dd_bounds"] (* Lemma second_derivative_compose_exp_bounds *) let second_derivative_compose_exp_bounds = Sections.section_proof ["f";"dd_bounds"] `nth_diff_strong_int 2 int f ==> - bounded_on_int (\x. nth_derivative 2 f x * exp (f x) + bounded_on_int (\x. nth_derivative 2 f x * exp (f x) + exp (f x) * derivative f x pow 2) int dd_bounds ==> has_bounded_second_derivative (\x. exp (f x)) int dd_bounds` [ @@ -3261,7 +3261,7 @@ let second_derivative_inv_bounds = Sections.section_proof ["dd_bounds"] let second_derivative_compose_inv_bounds = Sections.section_proof ["f";"f_bounds";"dd_bounds"] `bounded_on_int f int f_bounds ==> interval_not_zero f_bounds ==> nth_diff_strong_int 2 int f ==> - bounded_on_int (\x. (&2 * derivative f x pow 2 - nth_derivative 2 f x * f x) / f x pow 3) + bounded_on_int (\x. (&2 * derivative f x pow 2 - nth_derivative 2 f x * f x) / f x pow 3) int dd_bounds ==> has_bounded_second_derivative (\x. inv (f x)) int dd_bounds` [ @@ -3290,7 +3290,7 @@ let second_derivative_compose_sqrt_bounds = Sections.section_proof ["f";"f_bound `bounded_on_int f int f_bounds ==> interval_pos f_bounds ==> nth_diff_strong_int 2 int f ==> bounded_on_int (\x. (&2 * nth_derivative 2 f x * f x - derivative f x pow 2) / - (&4 * sqrt (f x pow 3))) + (&4 * sqrt (f x pow 3))) int dd_bounds ==> has_bounded_second_derivative (\x. sqrt (f x)) int dd_bounds` [ @@ -3578,7 +3578,7 @@ Sections.end_section "SecondDerivativeBound";; (* Section TaylorArith *) Sections.begin_section "TaylorArith";; -let cell_domain = new_definition `cell_domain x y z w <=> +let cell_domain = new_definition `cell_domain x y z w <=> x <= y /\ y <= z /\ y - x <= w /\ z - y <= w`;; (* Lemma taylor_x *) diff --git a/Formal_ineqs/taylor/theory/taylor_interval.vhl b/Formal_ineqs/taylor/theory/taylor_interval.vhl index 51fb7c2c..7438cbb1 100644 --- a/Formal_ineqs/taylor/theory/taylor_interval.vhl +++ b/Formal_ineqs/taylor/theory/taylor_interval.vhl @@ -1,15 +1,21 @@ -(* =========================================================== *) -(* Theory of univariate taylor intervals *) -(* Requires SSReflect/HOL Light for translation *) -(* See http://code.google.com/p/flyspeck/downloads/list *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -"needs \"lib/ssrbool.hl\"". -"needs \"lib/ssrnat.hl\"". -"needs \"arith/interval_arith.hl\"". -"needs \"trig/matan.hl\"". +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Theory of univariate taylor intervals *) +(* Requires SSReflect/HOL Light for translation *) +(* See https://github.com/monadius/hol-ssreflect *) +(* -------------------------------------------------------------------------- *) + +"needs \"Formal_ineqs/lib/ssrbool.hl\"". +"needs \"Formal_ineqs/lib/ssrnat.hl\"". +"needs \"Formal_ineqs/arith/interval_arith.hl\"". +"needs \"Formal_ineqs/trig/matan.hl\"". "needs \"Multivariate/realanalysis.ml\"". module Taylor_interval. @@ -27,7 +33,7 @@ module Taylor_interval. "let nth_differentiable = define `(nth_differentiable 0 f x <=> f real_continuous atreal x) /\ - (nth_differentiable (SUC n) f x <=> nth_differentiable n f x /\ + (nth_differentiable (SUC n) f x <=> nth_differentiable n f x /\ nth_derivative n f real_differentiable atreal x)`". "let nth_differentiable_on = new_definition `nth_differentiable_on n s f <=> @@ -36,11 +42,11 @@ module Taylor_interval. "let nth_differentiable_on_int = new_definition `nth_differentiable_on_int n int f <=> !x. interval_arith x int ==> nth_differentiable n f x`". -"let nth_diff_weak = new_definition `nth_diff_weak n f x <=> f real_continuous atreal x /\ +"let nth_diff_weak = new_definition `nth_diff_weak n f x <=> f real_continuous atreal x /\ ?F. F 0 = f /\ !i. i < n ==> (F i has_real_derivative F (SUC i) x) (atreal x)`". -"let nth_diff_strong = new_definition `nth_diff_strong n f x <=> +"let nth_diff_strong = new_definition `nth_diff_strong n f x <=> ?s. real_open s /\ x IN s /\ nth_differentiable_on n s f`". "let nth_diff_strong_int = new_definition `nth_diff_strong_int n int f <=> @@ -70,7 +76,7 @@ move => df; apply: (REAL_DERIVATIVE_UNIQUE_ATREAL f x). by rewrite df has_derivative_cond //; exists f'. Qed. -Lemma derivative_unique_on s f f': +Lemma derivative_unique_on s f f': `(!x. x IN s ==> (f has_real_derivative f' x) (atreal x)) ==> (!x. x IN s ==> f' x = derivative f x)`. move => df x xs. @@ -98,7 +104,7 @@ by rewrite ONE nth_derivativeS nth_derivative0. Qed. Lemma nth_derivative2 f : `nth_derivative 2 f = derivative (derivative f)`. by rewrite "ARITH_RULE `2 = SUC(SUC 0)`" nth_derivative !iterS "GEN_ALL iter". Qed. -Lemma nth_derivative_add n m f : +Lemma nth_derivative_add n m f : `nth_derivative n (nth_derivative m f) = nth_derivative (n + m) f`. by rewrite !nth_derivative iter_add. Qed. @@ -117,7 +123,7 @@ by rewrite nth_derivativeS has_derivative_alt. Qed. Lemma nth_differentiable_on_cond n s f : `nth_differentiable_on n s f ==> - !x. x IN s ==> + !x. x IN s ==> !i. i < n ==> (nth_derivative i f has_real_derivative (nth_derivative (SUC i) f x)) (atreal x)`. rewrite nth_differentiable_on => cond x. by move/cond => /nth_differentiable_cond. @@ -176,7 +182,7 @@ by rewrite !nth_derivative_add addSn cond; move: j_lt_ni; arith. Qed. -Lemma nth_diff_strong_imp_diff n f x : +Lemma nth_diff_strong_imp_diff n f x : `nth_diff_strong n f x ==> nth_differentiable n f x`. rewrite nth_diff_strong => [] [s] [_] [xs]; rewrite nth_differentiable_on => h. exact: h. @@ -310,7 +316,7 @@ Lemma derivative_composition f g x: `f real_differentiable atreal x ==> derivative (\x. g (f x)) x = derivative f x * derivative g (f x)`. Proof. move => /has_derivative_alt df /has_derivative_alt dg; apply: derivative_unique. -have := "GEN_ALL HAS_REAL_DERIVATIVE_CHAIN" +have := "GEN_ALL HAS_REAL_DERIVATIVE_CHAIN" `derivative f x` `derivative g` `\y. y = f x` f g; "ANTS_TAC"; first by move => y ->. by move => [_]; apply. @@ -511,14 +517,14 @@ Hypothesis dg : `g real_differentiable atreal x`. Lemma derivative_add : `derivative (\x. f x + g x) x = derivative f x + derivative g x`. by apply: derivative_unique; rewrite HAS_REAL_DERIVATIVE_ADD !has_derivative_alt. Qed. -Lemma derivative_mul : +Lemma derivative_mul : `derivative (\x. f x * g x) x = f x * derivative g x + derivative f x * g x`. by apply: derivative_unique; rewrite HAS_REAL_DERIVATIVE_MUL_ATREAL !has_derivative_alt. Qed. Lemma derivative_sub : `derivative (\x. f x - g x) x = derivative f x - derivative g x`. by apply: derivative_unique; rewrite HAS_REAL_DERIVATIVE_SUB !has_derivative_alt. Qed. -Lemma derivative_div : `~(g x = &0) ==> +Lemma derivative_div : `~(g x = &0) ==> derivative (\x. f x / g x) x = (derivative f x * g x - f x * derivative g x) / (g x * g x)`. move => gn0; apply: derivative_unique; rewrite -REAL_POW_2 HAS_REAL_DERIVATIVE_DIV_ATREAL //. by rewrite !has_derivative_alt. @@ -530,7 +536,7 @@ End DerivativeArith. Section MoreDerivativeArith. -Lemma differentiable_sum_numseg G n m x : +Lemma differentiable_sum_numseg G n m x : `(!i. i IN n..m ==> G i real_differentiable atreal x) ==> (\x. sum (n..m) (\i. G i x)) real_differentiable atreal x`. elim: m => [|m IHm] dG; rewrite !"GEN_ALL SUM_CLAUSES_NUMSEG". @@ -594,7 +600,7 @@ Variable n : `:num`. Hypothesis df : `nth_diff_strong_int n int f`. Lemma nth_derivative_scale_strong c i x : `interval_arith x int ==> i <= n ==> - ?s. real_open s /\ x IN s /\ + ?s. real_open s /\ x IN s /\ !y. y IN s ==> nth_derivative i (\y. c * f y) y = c * nth_derivative i f y`. move: df; rewrite nth_diff_strong_int nth_diff_strong => df. move/df => [s] [open_s] [xs]; rewrite nth_differentiable_on nth_differentiable_eq => diff. @@ -612,10 +618,10 @@ by rewrite eq. Qed. Lemma nth_derivative_scale_strong_all c x : `interval_arith x int ==> - ?s. real_open s /\ x IN s /\ + ?s. real_open s /\ x IN s /\ !i y. i <= n /\ y IN s ==> nth_derivative i (\y. c * f y) y = c * nth_derivative i f y`. move => /(nth_derivative_scale_strong c) h. -set P := `\i s. real_open s /\ x IN s /\ +set P := `\i s. real_open s /\ x IN s /\ (!y. y IN s ==> nth_derivative i (\y. c * f y) y = c * nth_derivative i f y)`. have sel_P : `!i. i <= n:num ==> P i ((@) (P i))`. move => i i_le_n; apply SELECT_AX. @@ -667,7 +673,7 @@ Hypothesis dg : `nth_diff_strong_int n int g`. (* Addition *) Lemma nth_derivative_add_strong i x : `interval_arith x int ==> i <= n ==> - ?s. real_open s /\ x IN s /\ + ?s. real_open s /\ x IN s /\ !y. y IN s ==> nth_derivative i (\y. f y + g y) y = nth_derivative i f y + nth_derivative i g y`. move: df dg; rewrite !nth_diff_strong_int !nth_diff_strong => df dg ineq. move: (df ineq) => [sf] [open_sf] [xsf]; move: df => _. @@ -688,11 +694,11 @@ by move => z z_in /=; rewrite eq. Qed. Lemma nth_derivative_add_strong_all x : `interval_arith x int ==> - ?s. real_open s /\ x IN s /\ - !i y. i <= n /\ y IN s ==> + ?s. real_open s /\ x IN s /\ + !i y. i <= n /\ y IN s ==> nth_derivative i (\y. f y + g y) y = nth_derivative i f y + nth_derivative i g y`. move => /(nth_derivative_add_strong) h. -set P := `\i s. real_open s /\ x IN s /\ +set P := `\i s. real_open s /\ x IN s /\ (!y. y IN s ==> nth_derivative i (\y. f y + g y) y = nth_derivative i f y + nth_derivative i g y)`. have sel_P : `!i. i <= n:num ==> P i ((@) (P i))`. move => i i_le_n; apply SELECT_AX. @@ -723,7 +729,7 @@ Qed. Lemma nth_diff_add : `nth_diff_strong_int n int (\x. f x + g x)`. -move: (df) (dg). +move: (df) (dg). rewrite !nth_diff_strong_int !nth_diff_strong !nth_differentiable_on !nth_differentiable_eq. move => df dg x ineq. have := nth_derivative_add_strong_all ineq. @@ -744,7 +750,7 @@ Qed. (* Subtraction *) Lemma nth_derivative_sub_strong i x : `interval_arith x int ==> i <= n ==> - ?s. real_open s /\ x IN s /\ + ?s. real_open s /\ x IN s /\ !y. y IN s ==> nth_derivative i (\y. f y - g y) y = nth_derivative i f y - nth_derivative i g y`. move: df dg; rewrite !nth_diff_strong_int !nth_diff_strong => df dg ineq. move: (df ineq) => [sf] [open_sf] [xsf]; move: df => _. @@ -765,11 +771,11 @@ by move => z z_in /=; rewrite eq. Qed. Lemma nth_derivative_sub_strong_all x : `interval_arith x int ==> - ?s. real_open s /\ x IN s /\ - !i y. i <= n /\ y IN s ==> + ?s. real_open s /\ x IN s /\ + !i y. i <= n /\ y IN s ==> nth_derivative i (\y. f y - g y) y = nth_derivative i f y - nth_derivative i g y`. move => /(nth_derivative_sub_strong) h. -set P := `\i s. real_open s /\ x IN s /\ +set P := `\i s. real_open s /\ x IN s /\ (!y. y IN s ==> nth_derivative i (\y. f y - g y) y = nth_derivative i f y - nth_derivative i g y)`. have sel_P : `!i. i <= n:num ==> P i ((@) (P i))`. move => i i_le_n; apply SELECT_AX. @@ -800,7 +806,7 @@ Qed. Lemma nth_diff_sub : `nth_diff_strong_int n int (\x. f x - g x)`. -move: (df) (dg). +move: (df) (dg). rewrite !nth_diff_strong_int !nth_diff_strong !nth_differentiable_on !nth_differentiable_eq. move => df dg x ineq. have := nth_derivative_sub_strong_all ineq. @@ -820,8 +826,8 @@ Qed. (* Multiplication *) Lemma nth_derivative_mul_strong i x : `interval_arith x int ==> i <= n ==> - ?s. real_open s /\ x IN s /\ - !y. y IN s ==> nth_derivative i (\y. f y * g y) y = + ?s. real_open s /\ x IN s /\ + !y. y IN s ==> nth_derivative i (\y. f y * g y) y = sum (0..i) (\k. &(binom (i, k)) * nth_derivative k f y * nth_derivative (i - k) g y)`. move: df dg; rewrite !nth_diff_strong_int !nth_diff_strong => df dg ineq. move: (df ineq) => [sf] [open_sf] [xsf]; move: df => _. @@ -849,7 +855,7 @@ have diff_cond : `!k. k IN 0..i ==> by move: ineq i_lt_n; arith. exists `nth_derivative (SUC (i - k)) g y`; rewrite (diff_g ysg). by move: ineq i_lt_n; arith. -have diff_cond2 : `!k. k IN 0..i ==> +have diff_cond2 : `!k. k IN 0..i ==> (\y. &(binom (i, k)) * nth_derivative k f y * nth_derivative (i - k) g y) real_differentiable atreal y`. move => k k_in /=. @@ -881,13 +887,13 @@ Qed. Lemma nth_derivative_mul_strong_all x : `interval_arith x int ==> - ?s. real_open s /\ x IN s /\ - !i y. i <= n /\ y IN s ==> - nth_derivative i (\y. f y * g y) y = + ?s. real_open s /\ x IN s /\ + !i y. i <= n /\ y IN s ==> + nth_derivative i (\y. f y * g y) y = sum (0..i) (\k. &(binom (i, k)) * nth_derivative k f y * nth_derivative (i - k) g y)`. move => /(nth_derivative_mul_strong) h. -set P := `\i s. real_open s /\ x IN s /\ - (!y. y IN s ==> nth_derivative i (\y. f y * g y) y = +set P := `\i s. real_open s /\ x IN s /\ + (!y. y IN s ==> nth_derivative i (\y. f y * g y) y = sum (0..i) (\k. &(binom (i, k)) * nth_derivative k f y * nth_derivative (i - k) g y))`. have sel_P : `!i. i <= n:num ==> P i ((@) (P i))`. move => i i_le_n; apply SELECT_AX. @@ -909,7 +915,7 @@ Qed. Lemma nth_derivative_mul i x : `interval_arith x int ==> i <= n ==> - nth_derivative i (\x. f x * g x) x = + nth_derivative i (\x. f x * g x) x = sum (0..i) (\k. &(binom (i, k)) * nth_derivative k f x * nth_derivative (i - k) g x)`. move => ineq i_le_n. have := nth_derivative_mul_strong ineq i_le_n. @@ -950,16 +956,16 @@ End NthDerivatives. (* Linear approximation and Taylor interval *) -"let lin_approx = new_definition `lin_approx f x f_bounds df_bounds <=> +"let lin_approx = new_definition `lin_approx f x f_bounds df_bounds <=> interval_arith (f x) f_bounds /\ (?f'. (f has_real_derivative f') (atreal x) /\ interval_arith f' df_bounds)`". "let has_bounded_second_derivative = new_definition `has_bounded_second_derivative f int dd_bounds <=> - nth_diff_strong_int 2 int f /\ + nth_diff_strong_int 2 int f /\ bounded_on_int (nth_derivative 2 f) int dd_bounds`". -"let taylor_interval = new_definition +"let taylor_interval = new_definition `taylor_interval f x y z w f_bounds df_bounds ddf_bounds <=> x <= y /\ y <= z /\ y - x <= w /\ z - y <= w /\ lin_approx f y f_bounds df_bounds /\ @@ -975,7 +981,7 @@ by move: (h ineq) => [s] [_] [xs]; exact. Qed. -Lemma has_bounded_second_derivative_old f int dd_bounds : +Lemma has_bounded_second_derivative_old f int dd_bounds : `has_bounded_second_derivative f int dd_bounds ==> ?f' f''. (!x. interval_arith x int ==> (f has_real_derivative f' x) (atreal x) /\ (f' has_real_derivative f'' x) (atreal x) /\ interval_arith (f'' x) dd_bounds)`. @@ -1026,7 +1032,7 @@ by move: (cond `--y`); rewrite REAL_NEG_NEG; apply; move: ineqs; arith. Qed. -Lemma continuous_leq_segment f c a b : +Lemma continuous_leq_segment f c a b : `a < b ==> f real_continuous atreal a ==> f real_continuous atreal b ==> (!x. x IN real_interval (a, b) ==> f x <= c) ==> (!x. x IN real_interval [a, b] ==> f x <= c)`. @@ -1079,7 +1085,7 @@ Qed. -Lemma taylor_error t : `x <= t /\ t <= z ==> +Lemma taylor_error t : `x <= t /\ t <= z ==> abs (f t - f y) <= w * iabs df_bounds + w * w * dd_bound / &2`. move => t_ineqs. have := tif; rewrite taylor_interval !andbA => [] [] [domain_ineqs] lin_app. @@ -1106,7 +1112,7 @@ suff: `abs (f t - (f y + f' y * (t - y) pow 1)) <= dd_bound * abs (t - y) pow (1 rewrite andTb !andbT REAL_MUL_SYM -REAL_POW_2 REAL_LE_MUL2 REAL_LE_REFL dd_prop. by rewrite REAL_LE_POW_2 andTb !andbT -REAL_LE_SQUARE_ABS REAL_ABS_ABS; move: abs_ty; arith. set Df := `\i. if i = 0 then f else if i = 1 then f' else if i = 2 then f'' else I`. -have arith := "ARITH_RULE `0 + 1 = 1 /\ ~(1 = 0) /\ 1 + 1 = 2 /\ +have arith := "ARITH_RULE `0 + 1 = 1 /\ ~(1 = 0) /\ 1 + 1 = 2 /\ ~(2 = 1) /\ ~(2 = 1) /\ ~(2 = 0)`". have := REAL_TAYLOR Df `1` `real_interval [x, z]` dd_bound. "ANTS_TAC". @@ -1140,7 +1146,7 @@ move: tif; rewrite {1}eq taylor_interval !andbA lin_approx => [] [] [_] [f_int] move: f_int; rewrite interval_arith; arith. Qed. -Lemma taylor_derivative_error : `!t. x <= t /\ t <= z ==> +Lemma taylor_derivative_error : `!t. x <= t /\ t <= z ==> abs (derivative f t - derivative f y) <= w * dd_bound`. move => t t_ineqs. have := tif; rewrite taylor_interval !andbA => [] [] [domain_ineqs] _. @@ -1156,7 +1162,7 @@ have dd_prop : `&0 <= dd_bound /\ !p. p IN real_interval [x, z] ==> abs (f'' p) by move: (df y); rewrite (pair_eq ddf_bounds) !interval_arith iabs !domain_ineqs; arith. by move => p /df; rewrite pair_eq interval_arith iabs; arith. set Df := `\i. if i = 0 then f' else if i = 1 then f'' else I`. -have arith := "ARITH_RULE `0 + 1 = 1 /\ ~(1 = 0) /\ 1 + 1 = 2 /\ +have arith := "ARITH_RULE `0 + 1 = 1 /\ ~(1 = 0) /\ 1 + 1 = 2 /\ ~(2 = 1) /\ ~(2 = 1) /\ ~(2 = 0)`". have := REAL_TAYLOR Df `0` `real_interval [x, z]` dd_bound. "ANTS_TAC". @@ -1212,7 +1218,7 @@ rewrite lin_approx real_differentiable; split => [[-> [f'] [df' int_f']] | [[f'] by rewrite (derivative_unique f f'). by rewrite andTb; exists f'; rewrite df -(derivative_unique f f' x). Qed. - + Hypothesis approx_f : `lin_approx f x f_bounds df_bounds`. @@ -1454,7 +1460,7 @@ by rewrite !h REAL_DIFFERENTIABLE_IMP_CONTINUOUS_ATREAL. Qed. -Lemma lin_approx_compose f g y g_bounds f_bounds d_bounds: +Lemma lin_approx_compose f g y g_bounds f_bounds d_bounds: `nth_diff_strong_int 2 g_bounds f ==> g real_differentiable atreal y ==> interval_arith (g y) g_bounds ==> @@ -1528,7 +1534,7 @@ rewrite -[`x pow (n - 1)`]REAL_MUL_RID HAS_REAL_DERIVATIVE_POW_ATREAL. by rewrite HAS_REAL_DERIVATIVE_ID. Qed. -Lemma second_derivative_pow_x n : +Lemma second_derivative_pow_x n : `derivative (derivative (\x. x pow n)) = (\x. &(n * (n - 1)) * x pow (n - 2))`. Proof. rewrite derivative_pow_x -eq_ext => x /=; apply: derivative_unique. @@ -1562,7 +1568,7 @@ Qed. (* atn *) -Lemma second_derivative_atn_eq x : `((\x. inv (&1 + x pow 2)) has_real_derivative +Lemma second_derivative_atn_eq x : `((\x. inv (&1 + x pow 2)) has_real_derivative (-- &2 * x) * inv (&1 + x pow 2) pow 2) (atreal x)`. Proof. rewrite REAL_POW_INV -real_div -REAL_NEG_LMUL HAS_REAL_DERIVATIVE_INV_ATREAL; split. @@ -1572,10 +1578,10 @@ rewrite REAL_POW_INV -real_div -REAL_NEG_LMUL HAS_REAL_DERIVATIVE_INV_ATREAL; sp by rewrite REAL_LT_IMP_NZ // REAL_ADD_SYM REAL_LT_ADD1 REAL_LE_POW_2. Qed. -Lemma second_derivative_atn : +Lemma second_derivative_atn : `derivative (derivative atn) = (\x. (-- &2 * x) * inv (&1 + x pow 2) pow 2)`. Proof. -rewrite derivative_atn -eq_ext => x /=; apply: derivative_unique. +rewrite derivative_atn -eq_ext => x /=; apply: derivative_unique. by rewrite -REAL_POW_2 second_derivative_atn_eq. Qed. @@ -1747,7 +1753,7 @@ Qed. Lemma real_powS x n : `x pow (SUC n) = x * x pow n`. by rewrite real_pow. Qed. -Lemma second_derivative_acs x : `abs x < &1 ==> +Lemma second_derivative_acs x : `abs x < &1 ==> nth_derivative 2 acs x = --(x / sqrt ((&1 - x * x) pow 3))`. Proof. rewrite nth_derivative2 => x_ineq. @@ -1801,7 +1807,7 @@ Qed. (* asn *) -Lemma second_derivative_asn x : `abs x < &1 ==> +Lemma second_derivative_asn x : `abs x < &1 ==> nth_derivative 2 asn x = x / sqrt ((&1 - x * x) pow 3)`. Proof. rewrite nth_derivative2 => x_ineq. @@ -1892,7 +1898,7 @@ Lemma REAL_CONTINUOUS_OPEN_PREIMAGE f s t : `f real_continuous_on s ==> real_ope real_open {x | x IN s /\ f x IN t}`. move => f_cont open_s open_t. rewrite REAL_OPEN. -suff ->: `IMAGE lift {x | x IN s /\ f x IN t} = +suff ->: `IMAGE lift {x | x IN s /\ f x IN t} = {x | x IN (IMAGE lift s) /\ (lift o f o drop) x IN (IMAGE lift t)}`. by rewrite CONTINUOUS_OPEN_PREIMAGE -!REAL_OPEN; rewr ETA_AX; rewrite -REAL_CONTINUOUS_ON. rewrite EXTENSION IN_IMAGE_LIFT_DROP !IN_ELIM_THM !IN_IMAGE_LIFT_DROP => z /=; split. @@ -1904,7 +1910,7 @@ Qed. Lemma second_derivative_compose f g x : `nth_diff_strong 2 g x ==> nth_diff_strong 2 f (g x) ==> - nth_derivative 2 (\x. f (g x)) x = + nth_derivative 2 (\x. f (g x)) x = nth_derivative 2 f (g x) * (derivative g x) pow 2 + derivative f (g x) * nth_derivative 2 g x`. move => dg df. rewrite nth_derivative2; apply: derivative_unique; apply: HAS_REAL_DERIVATIVE_LOCAL. @@ -2010,7 +2016,7 @@ Qed. Lemma second_derivative_compose_pow n : `nth_derivative 2 (\x. f x pow n) x = - &n * (nth_derivative 2 f x * f x pow (n - 1) + &n * (nth_derivative 2 f x * f x pow (n - 1) + &(n - 1) * f x pow (n - 2) * derivative f x pow 2)`. Proof. set g := `\x. x pow n`. @@ -2038,11 +2044,11 @@ Qed. (* atn *) -Lemma second_derivative_compose_atn : - `nth_derivative 2 (\x. atn (f x)) x = +Lemma second_derivative_compose_atn : + `nth_derivative 2 (\x. atn (f x)) x = (nth_derivative 2 f x * (&1 + f x * f x) - &2 * f x * derivative f x pow 2) / (&1 + f x * f x) pow 2`. Proof. -rewrite second_derivative_compose ?diff2_atn // +rewrite second_derivative_compose ?diff2_atn // nth_derivative2 second_derivative_atn /= derivative_atn /=. rewrite REAL_ADD_SYM !REAL_MUL_LNEG -real_sub. set lhs1 := `_1 * _2`; set lhs2 := `_1 * _2`. @@ -2056,8 +2062,8 @@ by move: (REAL_LE_SQUARE `f x`); arith. Qed. (* cos *) -Lemma second_derivative_compose_cos : - `nth_derivative 2 (\x. cos (f x)) x = +Lemma second_derivative_compose_cos : + `nth_derivative 2 (\x. cos (f x)) x = --(nth_derivative 2 f x * sin (f x) + cos (f x) * derivative f x pow 2)`. Proof. rewrite second_derivative_compose ?diff2_cos // second_derivative_cos /= derivative_cos /=. @@ -2065,8 +2071,8 @@ by arith. Qed. (* sin *) -Lemma second_derivative_compose_sin : - `nth_derivative 2 (\x. sin (f x)) x = +Lemma second_derivative_compose_sin : + `nth_derivative 2 (\x. sin (f x)) x = nth_derivative 2 f x * cos (f x) - sin (f x) * derivative f x pow 2`. Proof. rewrite second_derivative_compose ?diff2_sin // second_derivative_sin /= derivative_sin /=. @@ -2074,8 +2080,8 @@ by arith. Qed. (* exp *) -Lemma second_derivative_compose_exp : - `nth_derivative 2 (\x. exp (f x)) x = +Lemma second_derivative_compose_exp : + `nth_derivative 2 (\x. exp (f x)) x = nth_derivative 2 f x * exp (f x) + exp (f x) * derivative f x pow 2`. Proof. rewrite second_derivative_compose ?diff2_exp // second_derivative_exp /= derivative_exp /=. @@ -2101,7 +2107,7 @@ move => fn0. rewrite second_derivative_compose ?diff2_inv // second_derivative_inv // derivative_inv //. rewrite REAL_MUL_LNEG -real_sub real_div REAL_SUB_RDISTRIB. apply: "REAL_ARITH `!a b c d. a = c /\ b = d ==> a - b = c - d`"; split; first by arith. -rewrite REAL_INV_POW "ARITH_RULE `3 = SUC 2`" real_powS. +rewrite REAL_INV_POW "ARITH_RULE `3 = SUC 2`" real_powS. rewrite "REAL_ARITH `!a b c d. (a * b) * c * d = a * (b * c) * d`" REAL_MUL_RINV // REAL_MUL_LID. by rewrite REAL_INV_MUL -REAL_POW_2 REAL_MUL_SYM. Qed. @@ -2125,7 +2131,7 @@ Qed. (* acs *) Lemma second_derivative_compose_acs : `abs (f x) < &1 ==> nth_derivative 2 (\x. acs (f x)) x = - -- ((nth_derivative 2 f x * (&1 - f x * f x) + f x * derivative f x pow 2) / + -- ((nth_derivative 2 f x * (&1 - f x * f x) + f x * derivative f x pow 2) / sqrt ((&1 - f x * f x) pow 3))`. Proof. move => f_ineq. @@ -2145,7 +2151,7 @@ Qed. (* asn *) Lemma second_derivative_compose_asn : `abs (f x) < &1 ==> nth_derivative 2 (\x. asn (f x)) x = - ((nth_derivative 2 f x * (&1 - f x * f x) + f x * derivative f x pow 2) / + ((nth_derivative 2 f x * (&1 - f x * f x) + f x * derivative f x pow 2) / sqrt ((&1 - f x * f x) pow 3))`. Proof. move => f_ineq. @@ -2253,7 +2259,7 @@ Qed. Lemma second_derivative_div f g x : `~(g x = &0) ==> nth_diff_strong 2 f x ==> nth_diff_strong 2 g x ==> - nth_derivative 2 (\x. f x / g x) x = + nth_derivative 2 (\x. f x / g x) x = ((nth_derivative 2 f x * g x - f x * nth_derivative 2 g x) * g x - &2 * derivative g x * (derivative f x * g x - f x * derivative g x)) / (g x pow 3)`. move => gn0 diff_f diff_g. @@ -2264,7 +2270,7 @@ rewrite derivative_compose_inv. by move: diff_g; rewrite nth_diff_strong2_eq => [] [s] [open_s] [xs] -> //. rewrite ddf_def ddg_def df_def dg_def !real_div. set lhs := `_1 + _2`. -rewrite "REAL_RING `!f g x. ((ddf * g x - f x * ddg) * g x - &2 * dg * (df * g x - f x * dg)) * +rewrite "REAL_RING `!f g x. ((ddf * g x - f x * ddg) * g x - &2 * dg * (df * g x - f x * dg)) * inv (g x pow 3) = f x * (&2 * dg pow 2 - ddg * g x) * inv (g x pow 3) + &2 * df * --(g x * inv (g x pow 3)) * dg + @@ -2336,7 +2342,7 @@ Lemma second_derivative_compose_bounds f g int g_bounds dd_bounds : `nth_diff_strong_int 2 int g ==> bounded_on_int g int g_bounds ==> nth_diff_strong_int 2 g_bounds f ==> - bounded_on_int (\x. nth_derivative 2 f (g x) * derivative g x pow 2 + + bounded_on_int (\x. nth_derivative 2 f (g x) * derivative g x pow 2 + derivative f (g x) * nth_derivative 2 g x) int dd_bounds ==> has_bounded_second_derivative (\x. f (g x)) int dd_bounds`. rewrite has_bounded_second_derivative !bounded_on_int !nth_diff_strong_int /=. @@ -2347,7 +2353,7 @@ Qed. (* abs *) -Lemma second_derivative_abs_pos_bounds dd_bounds : +Lemma second_derivative_abs_pos_bounds dd_bounds : `interval_pos int ==> bounded_on_int (\x. &0) int dd_bounds ==> has_bounded_second_derivative abs int dd_bounds`. @@ -2358,7 +2364,7 @@ split => x ineq. rewrite second_derivative_abs_pos ?(interval_arith_pos x int) // (bounded ineq). Qed. -Lemma second_derivative_abs_neg_bounds dd_bounds : +Lemma second_derivative_abs_neg_bounds dd_bounds : `interval_neg int ==> bounded_on_int (\x. &0) int dd_bounds ==> has_bounded_second_derivative abs int dd_bounds`. @@ -2399,7 +2405,7 @@ Qed. (* pow *) -Lemma second_derivative_pow_bounds n dd_bounds : +Lemma second_derivative_pow_bounds n dd_bounds : `bounded_on_int (\x. &(n * (n - 1)) * x pow (n - 2)) int dd_bounds ==> has_bounded_second_derivative (\x. x pow n) int dd_bounds`. Proof. @@ -2414,13 +2420,13 @@ Lemma second_derivative_compose_pow_bounds n f dd_bounds : &(n - 1) * f x pow (n - 2) * derivative f x pow 2)) int dd_bounds ==> has_bounded_second_derivative (\x. f x pow n) int dd_bounds`. Proof. -rewrite has_bounded_second_derivative !bounded_on_int +rewrite has_bounded_second_derivative !bounded_on_int !nth_diff_strong_int /= => df bounded; split => x ineq. by rewrite diff2_compose_pow df. by rewrite second_derivative_compose_pow ?df // bounded. Qed. -Lemma second_derivative_pow2_bounds dd_bounds : +Lemma second_derivative_pow2_bounds dd_bounds : `bounded_on_int (\x. &2) int dd_bounds ==> has_bounded_second_derivative (\x. x pow 2) int dd_bounds`. Proof. @@ -2438,7 +2444,7 @@ move => df bounded; rewrite second_derivative_compose_pow_bounds df andTb. by rewrite "ARITH_RULE `2 - 1 = 1`" subnn "GEN_ALL real_pow" !REAL_MUL_LID REAL_POW_1. Qed. -Lemma second_derivative_pow3_bounds dd_bounds : +Lemma second_derivative_pow3_bounds dd_bounds : `bounded_on_int (\x. &6 * x) int dd_bounds ==> has_bounded_second_derivative (\x. x pow 3) int dd_bounds`. Proof. @@ -2458,7 +2464,7 @@ Qed. (* atn *) -Lemma second_derivative_atn_bounds dd_bounds : +Lemma second_derivative_atn_bounds dd_bounds : `bounded_on_int (\x. (-- &2 * x) * inv(&1 + x pow 2) pow 2) int dd_bounds ==> has_bounded_second_derivative atn int dd_bounds`. Proof. @@ -2480,7 +2486,7 @@ Qed. (* cos *) -Lemma second_derivative_cos_bounds dd_bounds : +Lemma second_derivative_cos_bounds dd_bounds : `bounded_on_int (\x. -- cos x) int dd_bounds ==> has_bounded_second_derivative cos int dd_bounds`. Proof. @@ -2491,7 +2497,7 @@ Qed. Lemma second_derivative_compose_cos_bounds f dd_bounds : `nth_diff_strong_int 2 int f ==> - bounded_on_int (\x. --(nth_derivative 2 f x * sin (f x) + bounded_on_int (\x. --(nth_derivative 2 f x * sin (f x) + cos (f x) * derivative f x pow 2)) int dd_bounds ==> has_bounded_second_derivative (\x. cos (f x)) int dd_bounds`. Proof. @@ -2502,7 +2508,7 @@ Qed. (* sin *) -Lemma second_derivative_sin_bounds dd_bounds : +Lemma second_derivative_sin_bounds dd_bounds : `bounded_on_int (\x. -- sin x) int dd_bounds ==> has_bounded_second_derivative sin int dd_bounds`. Proof. @@ -2513,7 +2519,7 @@ Qed. Lemma second_derivative_compose_sin_bounds f dd_bounds : `nth_diff_strong_int 2 int f ==> - bounded_on_int (\x. nth_derivative 2 f x * cos (f x) + bounded_on_int (\x. nth_derivative 2 f x * cos (f x) - sin (f x) * derivative f x pow 2) int dd_bounds ==> has_bounded_second_derivative (\x. sin (f x)) int dd_bounds`. Proof. @@ -2524,7 +2530,7 @@ Qed. (* exp *) -Lemma second_derivative_exp_bounds dd_bounds : +Lemma second_derivative_exp_bounds dd_bounds : `bounded_on_int exp int dd_bounds ==> has_bounded_second_derivative exp int dd_bounds`. Proof. @@ -2535,7 +2541,7 @@ Qed. Lemma second_derivative_compose_exp_bounds f dd_bounds : `nth_diff_strong_int 2 int f ==> - bounded_on_int (\x. nth_derivative 2 f x * exp (f x) + bounded_on_int (\x. nth_derivative 2 f x * exp (f x) + exp (f x) * derivative f x pow 2) int dd_bounds ==> has_bounded_second_derivative (\x. exp (f x)) int dd_bounds`. Proof. @@ -2547,7 +2553,7 @@ Qed. (* log *) -Lemma second_derivative_log_bounds dd_bounds : +Lemma second_derivative_log_bounds dd_bounds : `interval_pos int ==> bounded_on_int (\x. --inv (x pow 2)) int dd_bounds ==> has_bounded_second_derivative log int dd_bounds`. @@ -2576,7 +2582,7 @@ Qed. (* inv *) -Lemma second_derivative_inv_bounds dd_bounds : +Lemma second_derivative_inv_bounds dd_bounds : `interval_not_zero int ==> bounded_on_int (\x. &2 * inv (x pow 3)) int dd_bounds ==> has_bounded_second_derivative inv int dd_bounds`. @@ -2589,7 +2595,7 @@ Qed. Lemma second_derivative_compose_inv_bounds f f_bounds dd_bounds : `bounded_on_int f int f_bounds ==> interval_not_zero f_bounds ==> nth_diff_strong_int 2 int f ==> - bounded_on_int (\x. (&2 * derivative f x pow 2 - nth_derivative 2 f x * f x) / f x pow 3) + bounded_on_int (\x. (&2 * derivative f x pow 2 - nth_derivative 2 f x * f x) / f x pow 3) int dd_bounds ==> has_bounded_second_derivative (\x. inv (f x)) int dd_bounds`. Proof. @@ -2604,7 +2610,7 @@ Qed. (* sqrt *) -Lemma second_derivative_sqrt_bounds dd_bounds : +Lemma second_derivative_sqrt_bounds dd_bounds : `interval_pos int ==> bounded_on_int (\x. --inv (&4 * sqrt (x pow 3))) int dd_bounds ==> has_bounded_second_derivative sqrt int dd_bounds`. @@ -2619,7 +2625,7 @@ Lemma second_derivative_compose_sqrt_bounds f f_bounds dd_bounds : `bounded_on_int f int f_bounds ==> interval_pos f_bounds ==> nth_diff_strong_int 2 int f ==> bounded_on_int (\x. (&2 * nth_derivative 2 f x * f x - derivative f x pow 2) / - (&4 * sqrt (f x pow 3))) + (&4 * sqrt (f x pow 3))) int dd_bounds ==> has_bounded_second_derivative (\x. sqrt (f x)) int dd_bounds`. Proof. @@ -2633,7 +2639,7 @@ Qed. (* acs *) -Lemma second_derivative_acs_bounds dd_bounds : +Lemma second_derivative_acs_bounds dd_bounds : `iabs int < &1 ==> bounded_on_int (\x. --(x / sqrt ((&1 - x * x) pow 3))) int dd_bounds ==> has_bounded_second_derivative acs int dd_bounds`. @@ -2661,7 +2667,7 @@ Qed. (* asn *) -Lemma second_derivative_asn_bounds dd_bounds : +Lemma second_derivative_asn_bounds dd_bounds : `iabs int < &1 ==> bounded_on_int (\x. x / sqrt ((&1 - x * x) pow 3)) int dd_bounds ==> has_bounded_second_derivative asn int dd_bounds`. @@ -2693,7 +2699,7 @@ Qed. Lemma interval_arith_gt x int f : `interval_arith x int /\ interval_gt f int ==> f < x`. Proof. by rewrite (pair_eq int) interval_arith interval_gt; arith. Qed. -Lemma second_derivative_matan_bounds dd_bounds : +Lemma second_derivative_matan_bounds dd_bounds : `interval_gt (-- &1) int ==> bounded_on_int ddmatan int dd_bounds ==> has_bounded_second_derivative matan int dd_bounds`. @@ -2774,7 +2780,7 @@ Qed. (* div *) -Lemma second_derivative_div_bounds f g g_bounds dd_bounds : +Lemma second_derivative_div_bounds f g g_bounds dd_bounds : `bounded_on_int g int g_bounds ==> interval_not_zero g_bounds ==> nth_diff_strong_int 2 int f ==> nth_diff_strong_int 2 int g ==> @@ -2796,7 +2802,7 @@ End SecondDerivativeBound. Section TaylorArith. -"let cell_domain = new_definition `cell_domain x y z w <=> +"let cell_domain = new_definition `cell_domain x y z w <=> x <= y /\ y <= z /\ y - x <= w /\ z - y <= w`". @@ -2942,7 +2948,7 @@ by rewrite {2}REAL_ABS_NEG REAL_LE_ADD2 ?ddf1 ?ddf2 // REAL_ABS_TRIANGLE. Qed. -Lemma second_derivative_mul dd bf1 bf2 bdf1 bdf2 : +Lemma second_derivative_mul dd bf1 bf2 bdf1 bdf2 : `bounded_on f1 s bf1 ==> bounded_on f2 s bf2 ==> bounded_on (derivative f1) s bdf1 ==> bounded_on (derivative f2) s bdf2 ==> dd1 * bf2 + &2 * bdf1 * bdf2 + dd2 * bf1 <= dd ==> @@ -2967,13 +2973,13 @@ rewrite HAS_REAL_DERIVATIVE_MUL_ATREAL ?ddf1 ?ddf2 // andTb; split; last first. apply: REAL_LE_TRANS; exists `dd1 * bf2 + &2 * bdf1 * bdf2 + dd2 * bf1`. rewrite ineq andbT. apply: REAL_LE_TRANS; exists `abs (f1'' x * f2 x) + abs (&2 * f1' x * f2' x + f1 x * f2'' x)`. - rewrite REAL_ABS_TRIANGLE andTb REAL_LE_ADD2 REAL_ABS_MUL. + rewrite REAL_ABS_TRIANGLE andTb REAL_LE_ADD2 REAL_ABS_MUL. rewrite REAL_LE_MUL2 ?REAL_ABS_POS ?ddf1 ?b_f2 // andTb. apply: REAL_LE_TRANS; exists `abs (&2 * f1' x * f2' x) + abs (f1 x * f2'' x)`. rewrite REAL_ABS_TRIANGLE andTb REAL_LE_ADD2 !REAL_ABS_MUL [`dd2 * _`]REAL_MUL_SYM. rewrite !REAL_LE_MUL2 ?REAL_ABS_POS ?b_f1' ?b_f2' ?b_f1 ?ddf2 // REAL_LE_MUL ?REAL_ABS_POS //. by arith. -rewrite "REAL_ARITH `f1'' (x:real) * f2 x + &2 * f1' x * f2' x + f1 x * f2'' x = +rewrite "REAL_ARITH `f1'' (x:real) * f2 x + &2 * f1' x * f2' x + f1 x * f2'' x = (f1 x * f2'' x + f1' x * f2' x) + (f1' x * f2' x + f1'' x * f2 x)`". by rewrite HAS_REAL_DERIVATIVE_ADD !HAS_REAL_DERIVATIVE_MUL_ATREAL ?ddf1 ?ddf2. Qed. diff --git a/Formal_ineqs/tests/float_data.hl b/Formal_ineqs/tests/float_data.hl index ed0302b2..b14fb858 100644 --- a/Formal_ineqs/tests/float_data.hl +++ b/Formal_ineqs/tests/float_data.hl @@ -1,7 +1,7 @@ (* hol_log *) -needs "verifier/m_verifier_main.hl";; -needs "examples_flyspeck.hl";; +needs "Formal_ineqs/verifier/m_verifier_main.hl";; +needs "Formal_ineqs/examples_flyspeck.hl";; open M_verifier_main;; diff --git a/Formal_ineqs/tests/nat_test.hl b/Formal_ineqs/tests/nat_test.hl index f5c7f388..fdb82c9c 100644 --- a/Formal_ineqs/tests/nat_test.hl +++ b/Formal_ineqs/tests/nat_test.hl @@ -1,6 +1,6 @@ (* Run hol_test *) -needs "arith/arith_nat.hl";; +needs "Formal_ineqs/arith/arith_nat.hl";; needs "new_arith/nat_arith.hl";; #load "nums.cma";; diff --git a/Formal_ineqs/tests/results.hl b/Formal_ineqs/tests/results.hl index f8450f35..9b28fa7c 100644 --- a/Formal_ineqs/tests/results.hl +++ b/Formal_ineqs/tests/results.hl @@ -1,6 +1,6 @@ (* hol_test *) -needs "tests/test_utils.hl";; +needs "Formal_ineqs/tests/test_utils.hl";; show_result := true;; let rep = 1;; diff --git a/Formal_ineqs/tests/test_utils.hl b/Formal_ineqs/tests/test_utils.hl index 1847fae8..443ac0a4 100644 --- a/Formal_ineqs/tests/test_utils.hl +++ b/Formal_ineqs/tests/test_utils.hl @@ -1,8 +1,8 @@ (* Run hol_test *) -needs "arith/arith_nat.hl";; +needs "Formal_ineqs/arith/arith_nat.hl";; needs "new_arith/nat_arith.hl";; -needs "arith/arith_float.hl";; +needs "Formal_ineqs/arith/arith_float.hl";; #load "nums.cma";; #load "unix.cma";; diff --git a/Formal_ineqs/trig/asn_acs_eval.hl b/Formal_ineqs/trig/asn_acs_eval.hl index c78ba306..67c39af9 100644 --- a/Formal_ineqs/trig/asn_acs_eval.hl +++ b/Formal_ineqs/trig/asn_acs_eval.hl @@ -1,5 +1,17 @@ -needs "arith/float_pow.hl";; -needs "trig/atn_eval.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of functions: asn, acs *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/arith/float_pow.hl";; +needs "Formal_ineqs/trig/atn_eval.hl";; module type Asn_acs_eval_sig = sig @@ -82,7 +94,7 @@ let asn_pos_lo_th = (th_rule o prove) MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `atn t` THEN ASM_REWRITE_TAC[ATN_MONO_LE_EQ] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `x / r` THEN ASM_REWRITE_TAC[real_div] THEN MATCH_MP_TAC REAL_LE_LMUL THEN ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_INV2 THEN ASM_REWRITE_TAC[SQRT_LT_0] THEN + MATCH_MP_TAC REAL_LE_INV2 THEN ASM_REWRITE_TAC[SQRT_LT_0] THEN CONJ_TAC THENL [ REWRITE_TAC[REAL_ARITH `&0 < &1 - a <=> a < &1`; ABS_SQUARE_LT_1] THEN ASM_ARITH_TAC; @@ -107,7 +119,7 @@ let asn_neg_hi_th = (th_rule o prove) MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--x / r` THEN ASM_REWRITE_TAC[real_div] THEN REWRITE_TAC[REAL_MUL_LNEG; REAL_LE_NEG2; REAL_POW_NEG; ARITH_RULE `EVEN 2`] THEN MATCH_MP_TAC REAL_LE_LMUL THEN ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_INV2 THEN ASM_REWRITE_TAC[SQRT_LT_0] THEN + MATCH_MP_TAC REAL_LE_INV2 THEN ASM_REWRITE_TAC[SQRT_LT_0] THEN CONJ_TAC THENL [ REWRITE_TAC[REAL_ARITH `&0 < &1 - a <=> a < &1`; ABS_SQUARE_LT_1] THEN ASM_ARITH_TAC; @@ -154,7 +166,7 @@ let acs_hi_th = (th_rule o prove) c <= asn x /\ b - c <= hi ==> acs x <= hi`, - REWRITE_TAC[interval_arith] THEN STRIP_TAC THEN + REWRITE_TAC[interval_arith] THEN STRIP_TAC THEN ASM_SIMP_TAC[ACS_ASN] THEN ASM_ARITH_TAC);; let acs_lo_th = (th_rule o prove) @@ -164,11 +176,11 @@ let acs_lo_th = (th_rule o prove) asn x <= c /\ lo <= a - c ==> lo <= acs x`, - REWRITE_TAC[interval_arith] THEN STRIP_TAC THEN + REWRITE_TAC[interval_arith] THEN STRIP_TAC THEN ASM_SIMP_TAC[ACS_ASN] THEN ASM_ARITH_TAC);; let asn_interval = (th_rule o prove) - (`interval_arith x (a, b) /\ + (`interval_arith x (a, b) /\ (-- &1 <= a <=> T) /\ (b <= &1 <=> T) /\ lo <= asn a /\ @@ -182,7 +194,7 @@ let asn_interval = (th_rule o prove) ]);; let acs_interval = (th_rule o prove) - (`interval_arith x (a, b) /\ + (`interval_arith x (a, b) /\ (-- &1 <= a <=> T) /\ (b <= &1 <=> T) /\ lo <= acs b /\ @@ -196,7 +208,7 @@ let acs_interval = (th_rule o prove) ]);; let acs_interval2 = (th_rule o prove) - (`interval_arith x (a, b) /\ + (`interval_arith x (a, b) /\ (-- &1 <= a <=> T) /\ (b <= &1 <=> T) /\ interval_arith (pi / &2) (low, high) /\ diff --git a/Formal_ineqs/trig/atn.hl b/Formal_ineqs/trig/atn.hl index 0fececd2..c6c7061d 100644 --- a/Formal_ineqs/trig/atn.hl +++ b/Formal_ineqs/trig/atn.hl @@ -1,4 +1,17 @@ -needs "trig/series.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2011 Thomas C. Hales *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of functions: atn (theory) *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/trig/series.hl";; module Atn = struct @@ -44,7 +57,7 @@ let halfatn_half = prove (`!x t. abs x < t ==> abs (halfatn x) < t / &2`, REWRITE_TAC[halfatn; REAL_ABS_DIV] THEN REPEAT STRIP_TAC THEN SIMP_TAC[abs_lemma; pos_lemma2; REAL_LT_LDIV_EQ] THEN - MATCH_MP_TAC REAL_LTE_TRANS THEN EXISTS_TAC `t:real` THEN + MATCH_MP_TAC REAL_LTE_TRANS THEN EXISTS_TAC `t:real` THEN ASM_REWRITE_TAC[REAL_ARITH `t / &2 * a = t * (a / &2)`] THEN MATCH_MP_TAC (REAL_ARITH `t * (&1 + &1) <= t * a ==> t <= t * (a / &2)`) THEN MATCH_MP_TAC REAL_LE_LMUL THEN @@ -151,7 +164,7 @@ let halfatn_mono_lt = prove ASM_ARITH_TAC; ALL_TAC ] THEN - REPEAT STRIP_TAC THEN + REPEAT STRIP_TAC THEN ASM_CASES_TAC `y < &0` THENL [ ONCE_REWRITE_TAC[REAL_ARITH `a < b <=> --b < --a`] THEN REWRITE_TAC[GSYM halfatn_odd] THEN FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_ARITH_TAC; @@ -163,7 +176,7 @@ let halfatn_mono_lt = prove ALL_TAC ] THEN FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_ARITH_TAC);; - + let halfatn_mono = prove (`!x y. x <= y ==> halfatn x <= halfatn y`, REWRITE_TAC[REAL_LE_LT] THEN REPEAT STRIP_TAC THEN ASM_SIMP_TAC[halfatn_mono_lt]);; @@ -183,7 +196,7 @@ let atn_series_alt_sign = prove let f = (\i. (-- &1) pow i / &(2 * i + 1) * x pow (2 * i + 1)) in (!n. abs (f (n + 1)) <= abs (f n) /\ f (n + 1) * f n <= &0) /\ ((f ---> &0) sequentially)`, - REPEAT STRIP_TAC THEN CONV_TAC let_CONV THEN + REPEAT STRIP_TAC THEN CONV_TAC let_CONV THEN REWRITE_TAC[REAL_ABS_MUL; REAL_ABS_DIV] THEN REWRITE_TAC[abs_neg_one_pow; ARITH_RULE `2 * (n + 1) + 1 = 2 * n + 3`] THEN REWRITE_TAC[REAL_ABS_POW; REAL_ARITH `&1 / a * b = b / a`; REAL_ABS_NUM] THEN @@ -210,7 +223,7 @@ let atn_series_alt_sign = prove MAP_EVERY EXISTS_TAC [`atn x`; `0`] THEN ASM_SIMP_TAC[FROM_0; atn_series] ]);; - + let atn_poly_bound = prove (`!n x. abs x < &1 ==> abs (atn x - sum (0..n) (\i. (-- &1 pow i / &(2 * i + 1)) * x pow (2 * i + 1))) @@ -272,7 +285,7 @@ let atn_poly_pos_lower_bound = prove let atn_halfatn4_pos_upper_bound = prove (`!x n. EVEN n /\ &0 <= x - ==> atn (halfatn4 x) + ==> atn (halfatn4 x) <= sum (0..n) (\i. ((-- &1) pow i / &(2 * i + 1)) * (halfatn4 x) pow (2 * i + 1))`, REPEAT STRIP_TAC THEN MATCH_MP_TAC atn_poly_pos_upper_bound THEN ASM_SIMP_TAC[halfatn4_pos; halfatn4_bound; REAL_ARITH `abs x < inv (&8) ==> x < &1`]);; @@ -285,8 +298,8 @@ let atn_halfatn4_pos_lower_bound = prove ASM_SIMP_TAC[halfatn4_pos; halfatn4_bound; REAL_ARITH `abs x < inv (&8) ==> x < &1`]);; let real_taylor_atn_halfatn4 = prove - (`!n x. abs (atn(halfatn4 x) - - sum (0..n) (\j. (-- &1 pow j) * halfatn4 x pow (2 * j + 1) / &(2 * j+ 1))) + (`!n x. abs (atn(halfatn4 x) - + sum (0..n) (\j. (-- &1 pow j) * halfatn4 x pow (2 * j + 1) / &(2 * j+ 1))) <= inv (&8 pow (2 * n + 3) * &(2 * n + 3))`, REPEAT GEN_TAC THEN ABBREV_TAC `y = halfatn4 x` THEN SUBGOAL_THEN `abs y < inv (&8)` ASSUME_TAC THENL [ diff --git a/Formal_ineqs/trig/atn_eval.hl b/Formal_ineqs/trig/atn_eval.hl index 28590f61..1d756dd8 100644 --- a/Formal_ineqs/trig/atn_eval.hl +++ b/Formal_ineqs/trig/atn_eval.hl @@ -1,5 +1,17 @@ -needs "trig/atn.hl";; -needs "trig/poly_eval.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of functions: atn *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/trig/atn.hl";; +needs "Formal_ineqs/trig/poly_eval.hl";; module type Atn_eval_sig = sig @@ -92,7 +104,7 @@ let halfatn_pos_lo_th = (th_rule o prove) ASM_REWRITE_TAC[REAL_LE_LADD; REAL_POW_2]);; let halfatn4_eq = prove - (`halfatn(halfatn(halfatn(halfatn x))) = halfatn4 x`, + (`halfatn(halfatn(halfatn(halfatn x))) = halfatn4 x`, REWRITE_TAC[halfatn4; o_THM]);; let float16 = mk_float 16 0;; @@ -130,7 +142,7 @@ let atn_interval_th = (th_rule o prove) ASM_REWRITE_TAC[ATN_MONO_LE_EQ] ]);; -let bound_high_transformation = (PURE_REWRITE_RULE[SYM float16_eq] o +let bound_high_transformation = (PURE_REWRITE_RULE[SYM float16_eq] o GEN_REWRITE_RULE (RAND_CONV o ONCE_DEPTH_CONV) [SYM float1_eq] o prove) (`(&0 <= t /\ t < &1 ==> atn t <= t * (p1 - t pow 2 * p2)) ==> @@ -161,7 +173,7 @@ let bound_high_transformation = (PURE_REWRITE_RULE[SYM float16_eq] o MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `m * b` THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_REWRITE_TAC[]);; -let bound_low_transformation = (PURE_REWRITE_RULE[SYM float16_eq] o +let bound_low_transformation = (PURE_REWRITE_RULE[SYM float16_eq] o GEN_REWRITE_RULE (RAND_CONV o ONCE_DEPTH_CONV) [SYM float1_eq] o prove) (`(&0 <= t /\ t < &1 ==> t * (p1 - t pow 2 * p2) <= atn t) ==> @@ -215,7 +227,7 @@ let eval_halfatn_pos_hi pp x_th = let c1_th = float_add_lo pp c_tm one_float in let f_tm = lhand (concl c1_th) in let s, n1_tm, e1_tm = dest_float f_tm in - let _ = + let _ = if s <> "F" then failwith ("eval_halfatn_pos_hi: s <> F: " ^ string_of_term (concl x_th)) else () in @@ -247,7 +259,7 @@ let eval_halfatn_pos_lo pp x_th = let div_th = float_div_lo pp t_tm r_tm in let lo_tm = lhand (concl div_th) in let s, n_tm, e_tm = dest_float lo_tm in - let _ = + let _ = if s <> "F" then failwith ("eval_halfatn_pos_lo: s <> F: " ^ string_of_term (concl x_th)) else () in @@ -273,19 +285,19 @@ let eval_halfatn4_pos_lo pp x_th = let eq_th = INST[x_tm, x_var_real] halfatn4_eq in let ltm, bounds = dest_comb (concl th0) in EQ_MP (AP_TERM ltm (AP_TERM (rator bounds) eq_th)) th0;; - + (* ---------------------------------- *) (* atn upper/lower bounds *) (* ---------------------------------- *) let mk_bound_tables bound_th = - let bound = (SPEC_ALL o + let bound = (SPEC_ALL o REWRITE_RULE[poly_f_even; poly_f_odd; GSYM REAL_POW_2; REAL_POW_POW] o REWRITE_RULE[alt_sum_eq_poly_f_even; alt_sum_eq_poly_f_odd; real_div]) bound_th in (* This rule does not simplify factorials *) let reduce_rule = CONV_RULE (DEPTH_CONV (FIRST_CONV [NUM_SUC_CONV; NUM_ADD_CONV; NUM_MULT_CONV])) in - let find_poly_f = rev o find_terms (fun tm -> + let find_poly_f = rev o find_terms (fun tm -> try (rator o rator) tm = `poly_f` with Failure _ -> false) in fun pp n -> let n_tm = mk_small_numeral n in @@ -301,14 +313,14 @@ let mk_bound_tables bound_th = bound_th, zip cs_tms cs_lists;; (* Computes i such that x^(2 * i + 1) / (2 * i + 1) <= base^(-(p + 1)) and cond(i) *) -let n_of_p_atn x pp cond = +let n_of_p_atn x pp cond = let t = (float_of_int Arith_num.arith_base) ** (float_of_int (-pp - 1)) in let rec try_i i = let _ = if i > 50 then failwith "n_of_p_atn: cannot find i" else () in if cond i then let d = float_of_int (2 * i + 1) in let r = (x ** d) /. d in - if r <= t then i else try_i (i + 1) + if r <= t then i else try_i (i + 1) else try_i (i + 1) in @@ -357,13 +369,13 @@ let float_atn_pos_high = let r16_le_hi = float_mul_hi pp float16 r_tm in let hi_tm = rand (concl r16_le_hi) in let cmp_1 = float_lt t_tm one_float in - if (fst o dest_const o rand o concl) cmp_1 = "F" then + if (fst o dest_const o rand o concl) cmp_1 = "F" then failwith ("float_atn_pos_high: t >= 1: " ^ string_of_term x_tm) - else + else let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; c_tm, c_var_real; n_tm, n_var_real; m_tm, m_var_real; r_tm, r_var_real; x_tm, x_var_real; t_tm, t_var_real; hi_tm, hi_var_real] bound_th in - itlist MY_PROVE_HYP [p1_high; p2_low; t_pow2_low; n_le_mb; an_le_c; + itlist MY_PROVE_HYP [p1_high; p2_low; t_pow2_low; n_le_mb; an_le_c; tc_le_r; r16_le_hi; cmp_1; h4] th0;; (* x_th = |- interval_arith x (&0, x) *) @@ -395,13 +407,13 @@ let float_atn_pos_low = MY_PROVE_HYP h4 th0 else let cmp_1 = float_lt t_tm one_float in - if (fst o dest_const o rand o concl) cmp_1 = "F" then + if (fst o dest_const o rand o concl) cmp_1 = "F" then failwith ("float_atn_pos_low: t >= 1: " ^ string_of_term x_tm) else let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; c_tm, c_var_real; n_tm, n_var_real; m_tm, m_var_real; r_tm, r_var_real; x_tm, x_var_real; t_tm, t_var_real; lo_tm, lo_var_real] bound_th in - itlist MY_PROVE_HYP [p1_low; p2_high; t_pow2_high; mb_le_n; c_le_an; + itlist MY_PROVE_HYP [p1_low; p2_high; t_pow2_high; mb_le_n; c_le_an; r_le_tc; lo_le_r16; cmp_1; h4] th0;; (* Computes an upper bound of atn for a given floating-point number *) diff --git a/Formal_ineqs/trig/cos_bounds_eval.hl b/Formal_ineqs/trig/cos_bounds_eval.hl index 63195481..3f553e37 100644 --- a/Formal_ineqs/trig/cos_bounds_eval.hl +++ b/Formal_ineqs/trig/cos_bounds_eval.hl @@ -1,5 +1,17 @@ -needs "trig/sin_cos.hl";; -needs "trig/poly_eval.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of functions: cos *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/trig/sin_cos.hl";; +needs "Formal_ineqs/trig/poly_eval.hl";; module type Cos_bounds_eval_sig = sig @@ -39,7 +51,7 @@ let bound_high_transformation = prove(`!c x p1 p2 a b m n r. c <= p1 - x * p2 == interval_arith m (&0, x) /\ n <= m * b /\ a - n <= r - ==> c <= r`, + ==> c <= r`, REWRITE_TAC[interval_arith] THEN REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `a - n:real` THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `p1 - x * p2:real` THEN ASM_REWRITE_TAC[] THEN @@ -53,7 +65,7 @@ let bound_low_transformation = prove(`!c x p1 p2 a b m n r. p1 - x * p2 <= c ==> interval_arith x (&0, m) /\ m * b <= n /\ r <= a - n - ==> r <= c`, + ==> r <= c`, REWRITE_TAC[interval_arith] THEN REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `a - n:real` THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `p1 - x * p2:real` THEN ASM_REWRITE_TAC[] THEN @@ -63,12 +75,12 @@ let bound_low_transformation = prove(`!c x p1 p2 a b m n r. p1 - x * p2 <= c ==> let mk_cos_bound_tables cos_bound_th = - let cos_bound = (SPEC_ALL o + let cos_bound = (SPEC_ALL o REWRITE_RULE[poly_f_even; poly_f_odd; GSYM REAL_POW_2; REAL_POW_POW] o REWRITE_RULE[alt_sum_eq_poly_f_even; real_div]) cos_bound_th in (* This rule does not simplify factorials *) let reduce_rule = CONV_RULE (DEPTH_CONV (FIRST_CONV [NUM_SUC_CONV; NUM_ADD_CONV; NUM_MULT_CONV])) in - let find_poly_f = rev o find_terms (fun tm -> + let find_poly_f = rev o find_terms (fun tm -> try (rator o rator) tm = `poly_f` with Failure _ -> false) in fun pp n -> let n_tm = mk_small_numeral n in @@ -91,13 +103,13 @@ let rec x_pow_over_fact x k = (* Computes i such that x^(2(i + 1))/(2(i + 1))! <= base^(-(p + 1)) and cond(i) *) -let n_of_p_cos x pp cond = +let n_of_p_cos x pp cond = let t = (float_of_int Arith_num.arith_base) ** (float_of_int (-pp - 1)) in let rec try_i i = let _ = if i > 50 then failwith "n_of_p_cos: cannot find i" else () in if cond i then let r = x_pow_over_fact x (2 * (i + 1)) in - if r <= t then i else try_i (i + 1) + if r <= t then i else try_i (i + 1) else try_i (i + 1) in @@ -142,10 +154,10 @@ let float_cos_high_raw = let an_le_r_th = float_sub_hi pp a_tm n_tm in let r_tm = (rand o concl) an_le_r_th in let cmp_1 = float_le r_tm one_float in - if (fst o dest_const o rand o concl) cmp_1 = "F" then + if (fst o dest_const o rand o concl) cmp_1 = "F" then INST[x_tm, x_var_real] cos_high_trivial - else - let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; m_tm, m_var_real; + else + let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; m_tm, m_var_real; n_tm, n_var_real; r_tm, r_var_real; x_tm, x_var_real] bound_th in itlist MY_PROVE_HYP [p1_high_th; p2_low_th; x_pow2_low; n_le_mb_th; an_le_r_th] th0;; @@ -169,10 +181,10 @@ let float_cos_low_raw = let r_le_an = float_sub_lo pp a_tm n_tm in let r_tm = (rand o rator o concl) r_le_an in let cmp_1 = float_le neg_one_tm r_tm in - if (fst o dest_const o rand o concl) cmp_1 = "F" then + if (fst o dest_const o rand o concl) cmp_1 = "F" then INST[x_tm, x_var_real] cos_low_trivial - else - let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; m_tm, m_var_real; + else + let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; m_tm, m_var_real; n_tm, n_var_real; r_tm, r_var_real; x_tm, x_var_real] bound_th in itlist MY_PROVE_HYP [p1_low_th; p2_high_th; x_pow2_high; mb_le_n; r_le_an] th0;; diff --git a/Formal_ineqs/trig/cos_eval.hl b/Formal_ineqs/trig/cos_eval.hl index fa0ff3bf..44c3e814 100644 --- a/Formal_ineqs/trig/cos_eval.hl +++ b/Formal_ineqs/trig/cos_eval.hl @@ -1,8 +1,20 @@ -needs "misc/misc_functions.hl";; -needs "misc/misc_vars.hl";; -needs "arith/interval_arith.hl";; -needs "trig/atn_eval.hl";; -needs "trig/cos_bounds_eval.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of functions: cos *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/misc/misc_functions.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; +needs "Formal_ineqs/arith/interval_arith.hl";; +needs "Formal_ineqs/trig/atn_eval.hl";; +needs "Formal_ineqs/trig/cos_bounds_eval.hl";; module type Cos_eval_sig = sig @@ -25,11 +37,11 @@ prioritize_real();; (* Interval approximations of 2 * pi: *) (* (|- interval_arith (&2 * pi) (lo, hi), |- hi - lo <= t, (lo, hi)) *) -let two_pi_array = +let two_pi_array = let n = Array.length pi_approx_array in let pi = pi_approx_array.(n - 1) in let two_pi = float_interval_mul n two_interval pi in - Array.init (n - 1) (fun i -> + Array.init (n - 1) (fun i -> let th = float_interval_round i two_pi in let _, bounds = dest_interval_arith (concl th) in let lo_tm, hi_tm = dest_pair bounds in @@ -40,7 +52,7 @@ let pi32_array = let pi = pi_approx_array.(n - 1) in let three = mk_float_interval_small_num 3 in let pi32 = float_interval_div n (float_interval_mul n three pi) two_interval in - Array.init (n - 1) (fun i -> + Array.init (n - 1) (fun i -> let th = float_interval_round i pi32 in let _, bounds = dest_interval_arith (concl th) in let lo_tm, hi_tm = dest_pair bounds in @@ -63,9 +75,9 @@ let neg_pi_le_pi_lo_array = (* Theorems of the form: |- &2 * pi <= high, |- &0 <= low *) let two_pi_high, two_pi_low = - let interval_pos = prove(`interval_arith x (float_num F n e, f) + let interval_pos = prove(`interval_arith x (float_num F n e, f) ==> x <= f /\ &0 <= (float_num F n e)`, - REWRITE_TAC[interval_arith] THEN + REWRITE_TAC[interval_arith] THEN MP_TAC (SPECL[`n:num`; `e:num`] FLOAT_F_POS) THEN REAL_ARITH_TAC) in let two_pi_list = Array.to_list two_pi_array in @@ -73,8 +85,8 @@ let two_pi_high, two_pi_low = let list1, list2 = unzip (map pair_of_list ths1) in Array.of_list list1, Array.of_list list2;; - -(**********) + +(**********) let f_pi = 3.14159265358979323846 and f_2_pi = 6.28318530717958647693 and @@ -105,10 +117,10 @@ let cos_full_interval = prove(`!x. interval_arith (cos x) (-- &1, &1)`, let cos_case_neg1 = case_rule cos_full_interval;; - + (* 0 *) let cos_reduction_0_pi_1 = prove(`interval_arith x (a,b) /\ - (?k. integer k /\ a1 <= a + (&2 * pi) * k + (?k. integer k /\ a1 <= a + (&2 * pi) * k /\ b + (&2 * pi) * k <= b1) /\ interval_arith a1 (&0, pi) /\ b1 <= pi /\ @@ -128,7 +140,7 @@ let cos_case0 = case_rule cos_reduction_0_pi_1;; (* 1 *) let cos_reduction_0_pi_2 = prove(`interval_arith x (a,b) /\ - (?k. integer k /\ a1 <= a + (&2 * pi) * k + (?k. integer k /\ a1 <= a + (&2 * pi) * k /\ b + (&2 * pi) * k <= b1) /\ interval_arith a1 (&0, pi) /\ b1 <= &2 * pi /\ @@ -160,7 +172,7 @@ let cos_case1 = case_rule cos_reduction_0_pi_2;; (* 2 *) let cos_reduction_0_pi_2a = prove(`interval_arith x (a,b) /\ - (?k. integer k /\ a1 <= a + (&2 * pi) * k + (?k. integer k /\ a1 <= a + (&2 * pi) * k /\ b + (&2 * pi) * k <= b1) /\ interval_arith a1 (&0, pi) /\ a1 + b1 <= &2 * pi /\ @@ -193,7 +205,7 @@ let cos_case2 = case_rule cos_reduction_0_pi_2a;; (* 3 *) let cos_reduction_0_pi_2b = prove(`interval_arith x (a,b) /\ - (?k. integer k /\ a1 <= a + (&2 * pi) * k + (?k. integer k /\ a1 <= a + (&2 * pi) * k /\ b + (&2 * pi) * k <= b1) /\ interval_arith a1 (&0, pi) /\ b1 <= &2 * pi /\ &2 * pi <= a1 + b1 /\ @@ -217,10 +229,10 @@ let cos_reduction_0_pi_2b = prove(`interval_arith x (a,b) /\ REPEAT (POP_ASSUM MP_TAC) THEN REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; let cos_case3 = case_rule cos_reduction_0_pi_2b;; - + (* 4 *) let cos_reduction_neg_pi_0_1 = prove(`interval_arith x (a,b) /\ - (?k. integer k /\ a1 <= a + (&2 * pi) * k + (?k. integer k /\ a1 <= a + (&2 * pi) * k /\ b + (&2 * pi) * k <= b1) /\ interval_arith a1 (--pi, &0) /\ (b1 <= &0 <=> T) /\ @@ -241,7 +253,7 @@ let cos_case4 = case_rule cos_reduction_neg_pi_0_1;; (* 5 (not used in computations) *) let cos_reduction_neg_pi_0_2 = prove(`interval_arith x (a,b) /\ - (?k. integer k /\ a1 <= a + (&2 * pi) * k + (?k. integer k /\ a1 <= a + (&2 * pi) * k /\ b + (&2 * pi) * k <= b1) /\ interval_arith a1 (--pi, &0) /\ b1 <= pi /\ @@ -261,9 +273,9 @@ let cos_reduction_neg_pi_0_2 = prove(`interval_arith x (a,b) /\ let cos_case5 = case_rule cos_reduction_neg_pi_0_2;; -(* 6 *) +(* 6 *) let cos_reduction_neg_pi_0_2a = prove(`interval_arith x (a,b) /\ - (?k. integer k /\ a1 <= a + (&2 * pi) * k + (?k. integer k /\ a1 <= a + (&2 * pi) * k /\ b + (&2 * pi) * k <= b1) /\ interval_arith a1 (--pi, &0) /\ b2 = --b1 /\ (a1 <= b2 <=> T) /\ @@ -284,7 +296,7 @@ let cos_case6 = case_rule cos_reduction_neg_pi_0_2a;; (* 7 *) let cos_reduction_neg_pi_0_2b = prove(`interval_arith x (a,b) /\ - (?k. integer k /\ a1 <= a + (&2 * pi) * k + (?k. integer k /\ a1 <= a + (&2 * pi) * k /\ b + (&2 * pi) * k <= b1) /\ interval_arith a1 (--pi, &0) /\ b1 <= pi /\ b2 = --b1 /\ (a1 <= b2 <=> F) /\ @@ -336,7 +348,7 @@ let reduction_high_0_pi = (case_rule o prove) ==> cos x <= high`, REWRITE_TAC[interval_arith] THEN REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--cos t` THEN ASM_REWRITE_TAC[REAL_LE_NEG2] THEN - REWRITE_TAC[GSYM cos_pi_minus] THEN + REWRITE_TAC[GSYM cos_pi_minus] THEN MATCH_MP_TAC COS_MONO_LE THEN ASM_REWRITE_TAC[] THEN REPEAT (POP_ASSUM MP_TAC) THEN REAL_ARITH_TAC);; @@ -348,7 +360,7 @@ let reduction_low_0_pi = (case_rule o prove) ==> low <= cos x`, REWRITE_TAC[interval_arith] THEN REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `--cos t` THEN ASM_REWRITE_TAC[REAL_LE_NEG2] THEN - REWRITE_TAC[GSYM cos_pi_minus] THEN + REWRITE_TAC[GSYM cos_pi_minus] THEN MATCH_MP_TAC COS_MONO_LE THEN ASM_REWRITE_TAC[] THEN REPEAT (POP_ASSUM MP_TAC) THEN REAL_ARITH_TAC);; @@ -428,7 +440,7 @@ let reduction_neg_th = (th_rule o prove) z * hi <= m /\ a1 <= a - m /\ n <= z * lo /\ b - n <= b1 ==> ?k. integer k /\ a1 <= a + (&2 * pi) * k /\ b + (&2 * pi) * k <= b1`, - REWRITE_TAC[interval_arith] THEN REPEAT STRIP_TAC THEN + REWRITE_TAC[interval_arith] THEN REPEAT STRIP_TAC THEN EXISTS_TAC `-- &i` THEN SIMP_TAC[INTEGER_CLOSED] THEN CONJ_TAC THENL [ MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `a - m:real` THEN @@ -450,7 +462,7 @@ let reduction_pos_th = (th_rule o prove) m <= z * lo /\ a1 <= a + m /\ z * hi <= n /\ b + n <= b1 ==> ?k. integer k /\ a1 <= a + (&2 * pi) * k /\ b + (&2 * pi) * k <= b1`, - REWRITE_TAC[interval_arith] THEN REPEAT STRIP_TAC THEN + REWRITE_TAC[interval_arith] THEN REPEAT STRIP_TAC THEN EXISTS_TAC `&i` THEN SIMP_TAC[INTEGER_CLOSED] THEN CONJ_TAC THENL [ MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `a + m:real` THEN @@ -502,13 +514,13 @@ let ab_sum3_correction = (th_rule o prove) b1 + z <= b2 ==> b1 <= b2 /\ &2 * pi <= a1 + b2`, REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; - + let ab_sum4_correction = (th_rule o prove) (`interval_arith (&2 * pi) (lo, hi) /\ m <= a1 + b1 /\ (hi <= m <=> T) ==> &2 * pi <= a1 + b1`, REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; - + let a_b_correction = (th_rule o prove) (`a2 <= a1 /\ b1 <= b2 /\ (?k. integer k /\ a1 <= a + (&2 * pi) * k /\ b + (&2 * pi) * k <= b1) @@ -656,7 +668,7 @@ let eval_high_pi_2pi pp x_pi_2pi x_tm = else let cos_t = float_cos_high_raw pp t_tm in let r_tm = rand (concl cos_t) in - let th0 = INST[x_tm, x_var_real; + let th0 = INST[x_tm, x_var_real; lo_tm, lo_var_real; hi_tm, hi_var_real; r_tm, r_var_real; t_tm, t_var_real] reduction_high_pi_2pi_1 in itlist MY_PROVE_HYP [t_ge0; x_pi_2pi; two_pi; lo_sub_x; cos_t] th0;; @@ -697,7 +709,7 @@ let eval_low_pi_2pi pp x_pi_2pi x_tm = else let cos_t = float_cos_low_raw pp t_tm in let r_tm = rand (rator (concl cos_t)) in - let th0 = INST[x_tm, x_var_real; + let th0 = INST[x_tm, x_var_real; lo_tm, lo_var_real; hi_tm, hi_var_real; r_tm, r_var_real; t_tm, t_var_real] reduction_low_pi_2pi_1 in itlist MY_PROVE_HYP [t_le_pi; x_pi_2pi; two_pi; hi_sub_x; cos_t] th0;; @@ -714,7 +726,7 @@ let get_i f_tm = let k0 = -int_of_float (x /. f_2_pi) in let y = x +. float_of_int k0 *. f_2_pi in if y < -.f_pi then k0 + 1 - else if y > f_pi then k0 - 1 + else if y > f_pi then k0 - 1 else k0;; (* i = 0 *) @@ -727,7 +739,7 @@ let reduction_zero a_tm b_tm = (* i < 0 *) let reduction_neg pp i a_tm b_tm = - let i_eq_th, i_num_tm = float_eq_th_of_num (num (-i)) in + let i_eq_th, i_num_tm = float_eq_th_of_num (Num.num_of_int (-i)) in let z_tm = rand (concl i_eq_th) in let two_pi_th, _, (lo_tm, hi_tm) = two_pi_array.(pp) in let lo_ge0_th = two_pi_low.(pp) in @@ -744,13 +756,13 @@ let reduction_neg pp i a_tm b_tm = a_tm, a_var_real; a1_tm, a1_var_real; b_tm, b_var_real; b1_tm, b1_var_real; m_tm, m_var_real; n_tm, n_var_real] reduction_neg_th in - let th1 = itlist MY_PROVE_HYP [i_eq_th; two_pi_th; lo_ge0_th; + let th1 = itlist MY_PROVE_HYP [i_eq_th; two_pi_th; lo_ge0_th; z_hi_th; z_lo_th; a_sub_m_th; b_sub_n_th] th0 in a1_tm, b1_tm, th1;; (* i > 0 *) let reduction_pos pp i a_tm b_tm = - let i_eq_th, i_num_tm = float_eq_th_of_num (num i) in + let i_eq_th, i_num_tm = float_eq_th_of_num (Num.num_of_int i) in let z_tm = rand (concl i_eq_th) in let two_pi_th, _, (lo_tm, hi_tm) = two_pi_array.(pp) in let lo_ge0_th = two_pi_low.(pp) in @@ -767,7 +779,7 @@ let reduction_pos pp i a_tm b_tm = a_tm, a_var_real; a1_tm, a1_var_real; b_tm, b_var_real; b1_tm, b1_var_real; m_tm, m_var_real; n_tm, n_var_real] reduction_pos_th in - let th1 = itlist MY_PROVE_HYP [i_eq_th; two_pi_th; lo_ge0_th; + let th1 = itlist MY_PROVE_HYP [i_eq_th; two_pi_th; lo_ge0_th; z_hi_th; z_lo_th; a_add_m_th; b_add_n_th] th0 in a1_tm, b1_tm, th1;; @@ -831,7 +843,7 @@ let correct_ab_sum pp a1_tm b1_tm = let z_tm = rand (concl r_add_t) in let b1_add_z = float_add_hi pp b1_tm z_tm in let b2_tm = rand (concl b1_add_z) in - let th0 = INST[lo_tm, lo_var_real; hi_tm, hi_var_real; + let th0 = INST[lo_tm, lo_var_real; hi_tm, hi_var_real; n_tm, n_var_real; m_tm, m_var_real; a1_tm, a1_var_real; b1_tm, b1_var_real; t_tm, t_var_real; r_tm, r_var_real; @@ -872,8 +884,8 @@ let float_interval_cos pp x_th = try (* Reduce the interval of x *) let i = get_i a_tm in - let a1_tm, b1_tm, red_th = - if i = 0 then reduction_zero a_tm b_tm + let a1_tm, b1_tm, red_th = + if i = 0 then reduction_zero a_tm b_tm else if i < 0 then reduction_neg pp i a_tm b_tm else reduction_pos pp i a_tm b_tm in @@ -904,7 +916,7 @@ let float_interval_cos pp x_th = (* a1 in [0, pi] *) let ab_sum_case, b1_le0, b1_tm, ab_sum_th = correct_ab_sum pp a1_tm b1_tm in let red_th = correct_ab red_th a_tm b_tm a1_tm b1_tm a1_le0 b1_le0 in - let inst_tms = [x_tm, x_var_real; a_tm, a_var_real; b_tm, b_var_real; + let inst_tms = [x_tm, x_var_real; a_tm, a_var_real; b_tm, b_var_real; a1_tm, a1_var_real; b1_tm, b1_var_real] in let flag_b_pi, b_le_pi = float_prove_le_interval b1_tm pi_approx_array.(pp) in let two_pi_th, _, _ = two_pi_array.(pp) in @@ -912,7 +924,7 @@ let float_interval_cos pp x_th = if flag_b_pi then (* b1 in [0, pi] *) let b1_ge0 = float_ge0 b1_tm in - let b1_int = + let b1_int = let th0 = INST[b1_tm, b1_var_real] b1_interval_0_pi_1 in MY_PROVE_HYP b1_ge0 (MY_PROVE_HYP b_le_pi th0) in let cos_b1 = eval_low_0_pi pp b1_int b1_tm and @@ -940,7 +952,7 @@ let float_interval_cos pp x_th = INST[x_tm, x_var_real] cos_case_neg1 else (* a1 in [-pi, 0] *) - let inst_tms = [x_tm, x_var_real; a_tm, a_var_real; b_tm, b_var_real; + let inst_tms = [x_tm, x_var_real; a_tm, a_var_real; b_tm, b_var_real; a1_tm, a1_var_real; b1_tm, b1_var_real] in let b1_0 = float_le0 b1_tm in if (fst o dest_const o rand o concl) b1_0 = "T" then @@ -972,7 +984,7 @@ let float_interval_cos pp x_th = if flag_b_pi then (* b1 <= pi ==> b1 in [0, pi] *) let b1_int = - let th0 = INST[b1_tm, b1_var_real; a1_tm, a1_var_real; + let th0 = INST[b1_tm, b1_var_real; a1_tm, a1_var_real; b2_tm, b2_var_real] b1_interval_0_pi_2 in itlist MY_PROVE_HYP [b_le_pi; a1_0_th; neg_b_eq; ab_cmp] th0 in let cos_b1 = eval_low_0_pi pp b1_int b1_tm in @@ -981,7 +993,7 @@ let float_interval_cos pp x_th = itlist MY_PROVE_HYP [red_th; a1_case_th; neg_b_eq; ab_cmp; b_le_pi; x_th; cos_b1] th0 else INST[x_tm, x_var_real] cos_case_neg1 - with Correction_failed -> + with Correction_failed -> let _ = warn true (Printf.sprintf "float_interval_cos: reduction failed (%s, %s)" (string_of_term a_tm) (string_of_term b_tm)) in INST[x_tm, x_var_real] cos_case_neg1;; @@ -1051,18 +1063,18 @@ eval_low_0_pi 3 a_th x_tm;; eval_high_0_pi 4 a_th x_tm;; float_neg (mk_float 1 0);; - - + + ;; - + (**************) let i = get_i a_tm;; -let a1_tm, b1_tm, red_th = - if i = 0 then reduction_zero a_tm b_tm +let a1_tm, b1_tm, red_th = + if i = 0 then reduction_zero a_tm b_tm else if i < 0 then reduction_neg pp i a_tm b_tm else reduction_pos pp i a_tm b_tm;; (* Prove -pi <= a1; otherwise correct the value of a1 *) @@ -1125,7 +1137,7 @@ let flag_b_2pi, b_le_2pi = float_prove_le_interval b1_tm two_pi_th in 7, itlist MY_PROVE_HYP [red_th; a1_case_th; neg_b_eq; ab_cmp; b_le_pi] th0 else -1, cos_case_neg1 - with Correction_failed -> + with Correction_failed -> let _ = warn true (Printf.sprintf "cos_reduction: reduction failed (%s, %s)" (string_of_term a_tm) (string_of_term b_tm)) in -1, cos_case_neg1;; @@ -1152,7 +1164,7 @@ let flag, n_le_th = float_prove_le n_tm lo_tm;; let z_tm = rand (concl r_add_t) in let b1_add_z = float_add_hi pp b1_tm z_tm in let b2_tm = rand (concl b1_add_z) in - let th0 = INST[lo_tm, lo_var_real; hi_tm, hi_var_real; + let th0 = INST[lo_tm, lo_var_real; hi_tm, hi_var_real; n_tm, n_var_real; m_tm, m_var_real; a1_tm, a1_var_real; b1_tm, b1_var_real; t_tm, t_var_real; r_tm, r_var_real; @@ -1166,8 +1178,8 @@ let flag, n_le_th = float_prove_le n_tm lo_tm;; let reduce_cos pp a_tm b_tm = try let i = get_i a_tm in - let a1_tm, b1_tm, red_th = - if i = 0 then reduction_zero a_tm b_tm + let a1_tm, b1_tm, red_th = + if i = 0 then reduction_zero a_tm b_tm else if i < 0 then reduction_neg pp i a_tm b_tm else reduction_pos pp i a_tm b_tm in (* Prove -pi <= a1; otherwise correct the value of a1 *) @@ -1229,7 +1241,7 @@ let reduce_cos pp a_tm b_tm = 7, itlist MY_PROVE_HYP [red_th; a1_case_th; neg_b_eq; ab_cmp; b_le_pi] th0 else -1, cos_case_neg1 - with Correction_failed -> + with Correction_failed -> let _ = warn true (Printf.sprintf "cos_reduction: reduction failed (%s, %s)" (string_of_term a_tm) (string_of_term b_tm)) in -1, cos_case_neg1;; diff --git a/Formal_ineqs/trig/exp_eval.hl b/Formal_ineqs/trig/exp_eval.hl index 44830a0c..dfa1cb88 100644 --- a/Formal_ineqs/trig/exp_eval.hl +++ b/Formal_ineqs/trig/exp_eval.hl @@ -1,12 +1,18 @@ -(* =========================================================== *) -(* Formal interval evaluation of exp *) -(* Author: Alexey Solovyev *) -(* Date: 2014-12-10 *) -(* =========================================================== *) - -needs "trig/exp_log.hl";; -needs "trig/poly_eval.hl";; -needs "arith/float_pow.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of functions: exp *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/trig/exp_log.hl";; +needs "Formal_ineqs/trig/poly_eval.hl";; +needs "Formal_ineqs/arith/float_pow.hl";; module type Exp_eval_sig = sig @@ -55,7 +61,7 @@ let exp_interval = (th_rule o prove) ]);; let bound_pos_low_trans = prove - (`a <= e ==> + (`a <= e ==> interval_arith r (&0, a) ==> r <= e`, REWRITE_TAC[interval_arith] THEN REAL_ARITH_TAC);; @@ -66,7 +72,7 @@ let bound_neg_low_trans = prove(`p1 - x * p2 <= e ==> &0 <= x /\ x * b <= n /\ r <= a - n - ==> r <= e`, + ==> r <= e`, REWRITE_TAC[interval_arith] THEN REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `a - n:real` THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `p1 - x * p2:real` THEN ASM_REWRITE_TAC[] THEN @@ -80,7 +86,7 @@ let bound_neg_high_trans = prove(`e <= p1 - x * p2 ==> &0 <= x /\ n <= x * b /\ a - n <= r - ==> e <= r`, + ==> e <= r`, REWRITE_TAC[interval_arith] THEN REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `a - n:real` THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `p1 - x * p2:real` THEN ASM_REWRITE_TAC[] THEN @@ -91,17 +97,17 @@ let bound_neg_high_trans = prove(`e <= p1 - x * p2 ==> (* Argument reduction *) let exp_neg_low = (th_rule o prove) - (`r <= exp (-- float_num F n e) + (`r <= exp (-- float_num F n e) ==> r <= exp (float_num T n e)`, REWRITE_TAC[FLOAT_NEG]);; let exp_neg_high = (th_rule o prove) - (`exp (-- float_num F n e) <= r + (`exp (-- float_num F n e) <= r ==> exp (float_num T n e) <= r`, REWRITE_TAC[FLOAT_NEG]);; let exp_pos_high = (th_rule o prove) - (`&0 < t /\ t <= exp (--x) /\ inv t <= r + (`&0 < t /\ t <= exp (--x) /\ inv t <= r ==> exp x <= r`, REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `inv t` THEN ASM_REWRITE_TAC[] THEN @@ -147,13 +153,13 @@ let exp_reduce_high = (th_rule o prove) (* ----------------- *) let mk_exp_bound_tables exp_bound_th = - let exp_bound = (SPEC_ALL o + let exp_bound = (SPEC_ALL o REWRITE_RULE[poly_f_even; poly_f_odd; GSYM REAL_POW_2; REAL_POW_POW] o REWRITE_RULE[sum_eq_poly_f] o REWRITE_RULE[alt_sum_eq_poly_f; real_div]) exp_bound_th in - let reduce_rule = CONV_RULE (DEPTH_CONV + let reduce_rule = CONV_RULE (DEPTH_CONV (FIRST_CONV [NUM_SUC_CONV; NUM_ADD_CONV; NUM_MULT_CONV])) in - let find_poly_f = rev o find_terms (fun tm -> try (rator o rator) tm = `poly_f` + let find_poly_f = rev o find_terms (fun tm -> try (rator o rator) tm = `poly_f` with Failure _ -> false) in fun pp n -> let n_tm = mk_small_numeral n in @@ -175,13 +181,13 @@ let rec x_pow_over_fact x k = x /. (float_of_int k) *. x_pow_over_fact x (k - 1);; (* Computes i such that x^i / i! <= base^(-(p + 1)) and cond(i) *) -let n_of_p_exp x pp cond = +let n_of_p_exp x pp cond = let t = (float_of_int Arith_num.arith_base) ** (float_of_int (-pp - 1)) in let rec try_i i = let _ = if i > 50 then failwith "n_of_p_exp: cannot find i" else () in if cond i then let r = x_pow_over_fact x i in - if r <= t then i else try_i (i + 1) + if r <= t then i else try_i (i + 1) else try_i (i + 1) in @@ -252,7 +258,7 @@ let float_exp_neg_low_raw = let r_tm = (rand o rator o concl) r_le_an in let cmp_1 = EQT_ELIM (float_le x_tm one_float) in let cmp_0 = EQT_ELIM (float_ge0 x_tm) in - let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; n_tm, n_var_real; + let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; n_tm, n_var_real; r_tm, r_var_real; x_tm, x_var_real] bound_th in itlist MY_PROVE_HYP [p1_low_th; p2_high_th; xb_le_n; r_le_an; cmp_0; cmp_1] th0;; @@ -275,7 +281,7 @@ let float_exp_neg_high_raw = let r_tm = (rand o concl) an_le_r in let cmp_1 = EQT_ELIM (float_le x_tm one_float) in let cmp_0 = EQT_ELIM (float_ge0 x_tm) in - let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; n_tm, n_var_real; + let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; n_tm, n_var_real; r_tm, r_var_real; x_tm, x_var_real] bound_th in itlist MY_PROVE_HYP [p1_high_th; p2_low_th; n_le_xb; an_le_r; cmp_1; cmp_0] th0;; @@ -365,8 +371,8 @@ let float_exp_hi pp x_tm = let hi_tm = rand (concl exp_y) in let hi_pow = float_pow_hi pp k hi_tm in let high_tm = rand (concl hi_pow) in - let th0 = INST[t_tm, t_var_real; k_tm, k_var_num; - y_tm, y_var_real; hi_tm, hi_var_real; + let th0 = INST[t_tm, t_var_real; k_tm, k_var_num; + y_tm, y_var_real; hi_tm, hi_var_real; high_tm, high_var_real; x_tm, x_var_real] exp_reduce_high in itlist MY_PROVE_HYP [t_gt0; k_eq; div_th; exp_y; hi_pow] th0;; @@ -387,8 +393,8 @@ let float_exp_lo pp x_tm = let lo_pow = float_pow_lo pp k lo_tm in let low_tm = rand (rator (concl lo_pow)) in let lo_ge0 = EQT_ELIM (float_ge0 lo_tm) in - let th0 = INST[t_tm, t_var_real; k_tm, k_var_num; - y_tm, y_var_real; lo_tm, lo_var_real; + let th0 = INST[t_tm, t_var_real; k_tm, k_var_num; + y_tm, y_var_real; lo_tm, lo_var_real; low_tm, low_var_real; x_tm, x_var_real] exp_reduce_low in itlist MY_PROVE_HYP [t_gt0; k_eq; div_th; exp_y; lo_pow; lo_ge0] th0;; diff --git a/Formal_ineqs/trig/exp_log.hl b/Formal_ineqs/trig/exp_log.hl index c4b1f3ec..e90384f4 100644 --- a/Formal_ineqs/trig/exp_log.hl +++ b/Formal_ineqs/trig/exp_log.hl @@ -1,4 +1,16 @@ -needs "trig/series.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of functions: exp, log (theory) *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/trig/series.hl";; module Exp_log = struct @@ -11,7 +23,7 @@ prioritize_real();; (* --------------------- *) let exp_pos_poly_lower_bound = prove - (`!x n. &0 <= x ==> + (`!x n. &0 <= x ==> sum (0..n) (\k. inv (&(FACT k)) * x pow k) <= exp x`, REPEAT STRIP_TAC THEN REWRITE_TAC[GSYM exp_infsum] THEN @@ -27,7 +39,7 @@ let exp_pos_poly_lower_bound = prove EXISTS_TAC `0` THEN ASM_REWRITE_TAC[]; ALL_TAC ] THEN - REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_MUL THEN ASM_SIMP_TAC[REAL_POW_LE] THEN + REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_MUL THEN ASM_SIMP_TAC[REAL_POW_LE] THEN MATCH_MP_TAC REAL_LE_DIV THEN REWRITE_TAC[REAL_POS]);; let exp_ge1 = prove @@ -63,9 +75,9 @@ let n_le_fact = prove let exp_alt_sign_lemma = prove (`!x. &0 <= x /\ x <= &1 - ==> (!k. abs ((-- &1) pow (k + 1) / &(FACT (k + 1)) * x pow (k + 1)) + ==> (!k. abs ((-- &1) pow (k + 1) / &(FACT (k + 1)) * x pow (k + 1)) <= abs ((-- &1) pow k / &(FACT k) * x pow k) /\ - ((-- &1) pow (k + 1) / &(FACT (k + 1)) * x pow (k + 1)) * + ((-- &1) pow (k + 1) / &(FACT (k + 1)) * x pow (k + 1)) * ((-- &1) pow k / &(FACT k) * x pow k) <= &0) /\ ((\k. (-- &1) pow k / &(FACT k) * x pow k) ---> &0) sequentially`, REPEAT STRIP_TAC THENL [ @@ -79,9 +91,9 @@ let exp_alt_sign_lemma = prove MATCH_MP_TAC FACT_MONO THEN ARITH_TAC; ALL_TAC ] THEN - REWRITE_TAC[REAL_ABS_POW] THEN + REWRITE_TAC[REAL_ABS_POW] THEN MATCH_MP_TAC REAL_POW_MONO_INV THEN ASM_ARITH_TAC; - REWRITE_TAC[REAL_ARITH `(a / b * r) * (c / d * t) <= &0 + REWRITE_TAC[REAL_ARITH `(a / b * r) * (c / d * t) <= &0 <=> &0 <= ((-- &1 pow 1 * a * c) * (r * t)) / b / d`] THEN MATCH_MP_TAC REAL_LE_DIV THEN SIMP_TAC[REAL_OF_NUM_LE; FACT_LT; LT_IMP_LE] THEN MATCH_MP_TAC REAL_LE_DIV THEN SIMP_TAC[REAL_OF_NUM_LE; FACT_LT; LT_IMP_LE] THEN @@ -151,7 +163,7 @@ let exp_neg_poly_lower_bound = prove ] THEN EXPAND_TAC "f" THEN ASM_REWRITE_TAC[neg_exp_infsum] THEN REWRITE_TAC[ARITH_RULE `SUC a - 1 = a`; SUM_CLAUSES_NUMSEG; LE_0] THEN - SUBGOAL_THEN `abs (-- &1 pow (SUC (2 * m)) / &(FACT (SUC (2 * m))) * x pow SUC (2 * m)) = + SUBGOAL_THEN `abs (-- &1 pow (SUC (2 * m)) / &(FACT (SUC (2 * m))) * x pow SUC (2 * m)) = --f (SUC (2 * m))` ASSUME_TAC THENL [ EXPAND_TAC "f" THEN REWRITE_TAC[REAL_ABS_MUL] THEN REWRITE_TAC[REAL_ABS_DIV; REAL_ABS_POW; REAL_ABS_NEG; REAL_ABS_NUM] THEN @@ -161,7 +173,7 @@ let exp_neg_poly_lower_bound = prove ALL_TAC ] THEN ASM_REWRITE_TAC[] THEN REAL_ARITH_TAC);; - + let exp_reduce = prove (`!x k. 0 < k ==> exp (x / &k) pow k = exp x`, REPEAT STRIP_TAC THEN REWRITE_TAC[GSYM REAL_EXP_N; REAL_ARITH `a * b / a = b * (a / a)`] THEN @@ -197,7 +209,7 @@ let log_alt_sign_lemma = prove ] THEN REWRITE_TAC[REAL_ABS_POW] THEN MATCH_MP_TAC REAL_POW_MONO_INV THEN ASM_ARITH_TAC; - REWRITE_TAC[REAL_ARITH `(a / b * r) * (c / d * t) <= &0 + REWRITE_TAC[REAL_ARITH `(a / b * r) * (c / d * t) <= &0 <=> &0 <= ((-- &1 pow 1 * a * c) * (r * t)) / b / d`] THEN MATCH_MP_TAC REAL_LE_DIV THEN REWRITE_TAC[REAL_OF_NUM_LE; ARITH_RULE `0 <= t + 1`] THEN MATCH_MP_TAC REAL_LE_DIV THEN REWRITE_TAC[REAL_OF_NUM_LE; ARITH_RULE `0 <= t + 1`] THEN @@ -226,7 +238,7 @@ let log_poly_upper_bound = prove log (&1 + x) <= x * sum (0..n) (\k. (-- &1) pow k / &(k + 1) * x pow k)`, REWRITE_TAC[EVEN_EXISTS] THEN REPEAT STRIP_TAC THEN ABBREV_TAC `f = \k. (-- &1) pow k / &(k + 1) * x pow (k + 1)` THEN - SUBGOAL_THEN `x * sum (0..n) (\k. -- &1 pow k / &(k + 1) * x pow k) + SUBGOAL_THEN `x * sum (0..n) (\k. -- &1 pow k / &(k + 1) * x pow k) = sum (0..n) f` ASSUME_TAC THENL [ REWRITE_TAC[GSYM SUM_LMUL; REAL_ARITH `a * b / c * d = b / c * (a * d)`] THEN REWRITE_TAC[GSYM real_pow; ARITH_RULE `SUC k = k + 1`] THEN @@ -265,7 +277,7 @@ let log_poly_lower_bound = prove x * sum (0..n) (\k. (-- &1) pow k / &(k + 1) * x pow k) <= log (&1 + x)`, REWRITE_TAC[ODD_EXISTS] THEN REPEAT STRIP_TAC THEN ABBREV_TAC `f = \k. (-- &1) pow k / &(k + 1) * x pow (k + 1)` THEN - SUBGOAL_THEN `x * sum (0..n) (\k. -- &1 pow k / &(k + 1) * x pow k) + SUBGOAL_THEN `x * sum (0..n) (\k. -- &1 pow k / &(k + 1) * x pow k) = sum (0..n) f` ASSUME_TAC THENL [ REWRITE_TAC[GSYM SUM_LMUL; REAL_ARITH `a * b / c * d = b / c * (a * d)`] THEN REWRITE_TAC[GSYM real_pow; ARITH_RULE `SUC k = k + 1`] THEN diff --git a/Formal_ineqs/trig/log_eval.hl b/Formal_ineqs/trig/log_eval.hl index 89ad32a3..ecd5b8a1 100644 --- a/Formal_ineqs/trig/log_eval.hl +++ b/Formal_ineqs/trig/log_eval.hl @@ -1,12 +1,18 @@ -(* =========================================================== *) -(* Formal interval evaluation of log *) -(* Author: Alexey Solovyev *) -(* Date: 2014-12-11 *) -(* =========================================================== *) - -needs "trig/exp_log.hl";; -needs "trig/poly_eval.hl";; -needs "trig/exp_eval.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of functions: log *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/trig/exp_log.hl";; +needs "Formal_ineqs/trig/poly_eval.hl";; +needs "Formal_ineqs/trig/exp_eval.hl";; module type Log_eval_sig = sig @@ -64,7 +70,7 @@ let log_reduce = prove ASM_SIMP_TAC[LOG_DIV; REAL_EXP_POS_LT; LOG_EXP] THEN REAL_ARITH_TAC);; let log_reduce_exp_high = prove - (`t <= exp k /\ x / t <= r /\ r - &1 <= a /\ + (`t <= exp k /\ x / t <= r /\ r - &1 <= a /\ &0 < x /\ &0 < t /\ log (&1 + a) <= hi /\ hi + k <= high ==> log x <= high`, @@ -94,7 +100,7 @@ let log_reduce_exp_high = prove MATCH_MP_TAC REAL_LE_INV2 THEN ASM_REWRITE_TAC[]);; let log_reduce_exp_low = prove - (`exp k <= t /\ r <= x / t /\ a <= r - &1 /\ + (`exp k <= t /\ r <= x / t /\ a <= r - &1 /\ &0 < x /\ -- &1 < a /\ lo <= log (&1 + a) /\ low <= lo + k ==> low <= log x`, @@ -117,7 +123,7 @@ let log_reduce_exp_low = prove MATCH_MP_TAC REAL_LE_INV2 THEN ASM_REWRITE_TAC[REAL_EXP_POS_LT]);; let log_reduce_simple_high = (th_rule o REWRITE_RULE[GSYM float1_eq] o prove) - (`t <= exp b /\ x / t <= r /\ r - &1 <= a /\ + (`t <= exp b /\ x / t <= r /\ r - &1 <= a /\ &0 < x /\ &0 < t /\ a + b <= high ==> log x <= high`, @@ -132,7 +138,7 @@ let log_reduce_simple_high = (th_rule o REWRITE_RULE[GSYM float1_eq] o prove) ASM_ARITH_TAC);; let log_reduce_simple_low = (th_rule o prove) - (`inv x <= t /\ log t <= r /\ + (`inv x <= t /\ log t <= r /\ low = --r /\ &0 < x ==> low <= log x`, REPEAT STRIP_TAC THEN ONCE_REWRITE_TAC[REAL_ARITH `a <= b <=> --b <= --a`] THEN diff --git a/Formal_ineqs/trig/matan.hl b/Formal_ineqs/trig/matan.hl index 5cc37bae..1e3d180c 100644 --- a/Formal_ineqs/trig/matan.hl +++ b/Formal_ineqs/trig/matan.hl @@ -1,4 +1,16 @@ -needs "trig/series.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of functions: matan (theory) *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/trig/series.hl";; module Matan = struct @@ -8,9 +20,9 @@ prioritize_real();; (* matan *) -let matan = new_definition `matan x = +let matan = new_definition `matan x = if (x = &0) then &1 - else if (x > &0) then atn (sqrt x) / (sqrt x) + else if (x > &0) then atn (sqrt x) / (sqrt x) else (log ((&1 + sqrt( -- x))/(&1 - sqrt( -- x)))) / (&2 * sqrt (-- x))`;; (* Auxiliary definitions for matan derivatives *) @@ -126,7 +138,7 @@ let matan_power_series = prove(`!x. abs x < &1 ] THEN REWRITE_TAC[GSYM REAL_POW_POW; REAL_SQRT_POW_2] THEN ASM_SIMP_TAC[REAL_ARITH `~(x > &0) ==> abs (--x) = --x`] THEN - ONCE_REWRITE_TAC[REAL_NEG_MINUS1] THEN + ONCE_REWRITE_TAC[REAL_NEG_MINUS1] THEN REWRITE_TAC[REAL_POW_MUL; REAL_POW_ONE; REAL_MUL_RID; real_div; REAL_MUL_AC]);; @@ -181,7 +193,7 @@ let matan_has_derivative_gt0 = prove(`!x. &0 < x ==> ASM_REWRITE_TAC[matan]; ALL_TAC ] THEN - REAL_DIFF_TAC THEN + REAL_DIFF_TAC THEN SUBGOAL_THEN `~(sqrt x = &0)` ASSUME_TAC THENL [ ASM_SIMP_TAC[SQRT_EQ_0; REAL_ARITH `&0 < x ==> ~(x = &0)`]; ALL_TAC @@ -200,7 +212,7 @@ let matan_real_derivative_gt0 = prove(`!x. &0 < x ==> let matan_real_derivative_abs1 = prove(`!x. abs x < &1 ==> matan real_differentiable atreal x - /\ ((\i. ((-- &1) pow (SUC i) * &(SUC i) / &(2 * i + 3)) * x pow i) + /\ ((\i. ((-- &1) pow (SUC i) * &(SUC i) / &(2 * i + 3)) * x pow i) real_sums (real_derivative matan x)) (from 0)`, GEN_TAC THEN DISCH_TAC THEN MP_TAC (MATCH_MP power_series_has_derivative2 matan_power_series) THEN @@ -220,7 +232,7 @@ let matan_real_derivative_abs1 = prove(`!x. abs x < &1 ==> let matan_real_derivative_abs1_bound = prove(`!x m. abs x < &1 ==> - abs (real_derivative matan x + abs (real_derivative matan x - sum (0..m) (\i. ((-- &1) pow (SUC i) * &(SUC i) / &(2 * i + 3)) * x pow i)) <= abs x pow (SUC m) / (&2 * (&1 - abs x))`, REPEAT STRIP_TAC THEN @@ -246,13 +258,13 @@ let matan_real_derivative_abs1_bound = prove(`!x m. abs x < &1 ==> REWRITE_TAC[REAL_FIELD `inv (&2) * (&2 * a + b) = a + inv(&2) * b`] THEN REWRITE_TAC[ADD1; GSYM REAL_OF_NUM_ADD; REAL_LE_LADD] THEN REAL_ARITH_TAC);; - + (* matan second derivative *) let matan_has_derivative2_gt0 = prove(`!x. &0 < x ==> - ((\x. inv (&2 * x * (x + &1)) - atn(sqrt(x)) / (&2 * x * sqrt(x))) has_real_derivative - (&3 / &4) * (atn(sqrt(x)) / (x pow 2 * sqrt(x))) + ((\x. inv (&2 * x * (x + &1)) - atn(sqrt(x)) / (&2 * x * sqrt(x))) has_real_derivative + (&3 / &4) * (atn(sqrt(x)) / (x pow 2 * sqrt(x))) - (&3 / &4) * inv(x pow 2 * (x + &1)) - (&1 / &2) * inv(x * (x + &1) pow 2)) (atreal x)`, REPEAT STRIP_TAC THEN @@ -277,7 +289,7 @@ let matan_has_derivative2_gt0 = prove(`!x. &0 < x ==> let matan_second_derivative_gt0 = prove(`!x. &0 < x ==> real_derivative (real_derivative matan) x = - (&3 / &4) * (atn(sqrt(x)) / (x pow 2 * sqrt(x))) + (&3 / &4) * (atn(sqrt(x)) / (x pow 2 * sqrt(x))) - (&3 / &4) * inv(x pow 2 * (x + &1)) - (&1 / &2) * inv(x * (x + &1) pow 2)`, REPEAT STRIP_TAC THEN @@ -382,14 +394,14 @@ let matan_second_derivative_abs1_alt = prove(`!x. abs x < &1 ==> REMOVE_THEN "a" (MP_TAC o MATCH_MP REAL_INFSUM_UNIQUE) THEN REMOVE_THEN "b" (MP_TAC o MATCH_MP REAL_INFSUM_UNIQUE) THEN ASM_SIMP_TAC[] THEN REAL_ARITH_TAC);; - - + + (* matan second derivative bound *) let matan_d2_bound_ge1_high = prove(`!x. &1 <= x ==> real_derivative (real_derivative matan) x <= #0.65`, - REPEAT STRIP_TAC THEN + REPEAT STRIP_TAC THEN ASM_SIMP_TAC[REAL_ARITH `&1 <= x ==> &0 < x`; matan_second_derivative_gt0] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `(&3 / &4) * atn(sqrt x) / (x pow 2 * sqrt x)` THEN @@ -417,7 +429,7 @@ let matan_d2_bound_ge1_high = prove(`!x. &1 <= x ==> EXPAND_TAC "y" THEN MATCH_MP_TAC REAL_LE_RSQRT THEN ASM_REWRITE_TAC[REAL_POW_ONE]; ALL_TAC ] THEN - MATCH_MP_TAC REAL_LE_TRANS THEN + MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `atn (&1) / (&1 pow 5)` THEN CONJ_TAC THENL [ ONCE_REWRITE_TAC[GSYM REAL_LE_NEG2] THEN ABBREV_TAC `f = \y. --(atn y / y pow 5)` THEN @@ -435,7 +447,7 @@ let matan_d2_bound_ge1_high = prove(`!x. &1 <= x ==> EXPAND_TAC "f" THEN REAL_DIFF_TAC THEN ASM_SIMP_TAC[REAL_ARITH `&1 <= x' ==> ~(x' = &0)`; REAL_POW_NZ] THEN REWRITE_TAC[ARITH_RULE `5 - 1 = 4 /\ 5 * 2 = 10`; REAL_POW_POW] THEN - SUBGOAL_THEN `!n. ~(x' pow n = &0)` ASSUME_TAC THENL [ + SUBGOAL_THEN `!n. ~(x' pow n = &0)` ASSUME_TAC THENL [ GEN_TAC THEN MATCH_MP_TAC REAL_POW_NZ THEN UNDISCH_TAC `&1 <= x'` THEN REAL_ARITH_TAC; ALL_TAC @@ -485,7 +497,7 @@ let matan_d2_bound_ge1_high = prove(`!x. &1 <= x ==> let matan_d2_bound_ge1_low = prove(`!x. &1 <= x ==> -- #0.65 <= real_derivative (real_derivative matan) x`, - REPEAT STRIP_TAC THEN + REPEAT STRIP_TAC THEN ASM_SIMP_TAC[REAL_ARITH `&1 <= x ==> &0 < x`; matan_second_derivative_gt0] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `-- (&3 / &4) * inv (x pow 2 * (x + &1)) - (&1 / &2) * inv (x * (x + &1) pow 2)` THEN @@ -493,7 +505,7 @@ let matan_d2_bound_ge1_low = prove(`!x. &1 <= x ==> REWRITE_TAC[REAL_ARITH `-- #0.65 <= --(&3 / &4) * a - (&1 / &2) * b <=> &3 * a + &2 * b <= #2.6`] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&3 * inv(&2) + &2 * inv(&2)` THEN CONJ_TAC THENL [ ALL_TAC; REAL_ARITH_TAC ] THEN - MATCH_MP_TAC REAL_LE_ADD2 THEN + MATCH_MP_TAC REAL_LE_ADD2 THEN CONJ_TAC THEN MATCH_MP_TAC REAL_LE_LMUL THEN REWRITE_TAC[REAL_POS] THEN MATCH_MP_TAC REAL_LE_INV2 THEN REWRITE_TAC[REAL_ARITH `&0 < &2`] THEN ONCE_REWRITE_TAC[REAL_ARITH `&2 = &1 * &2`] THEN MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_REWRITE_TAC[REAL_POS] THENL [ ASM_SIMP_TAC[REAL_POW_LE_1] THEN POP_ASSUM MP_TAC THEN REAL_ARITH_TAC; ALL_TAC @@ -504,7 +516,7 @@ let matan_d2_bound_ge1_low = prove(`!x. &1 <= x ==> ] THEN REWRITE_TAC[REAL_ARITH `--(&3 / &4) * a - b <= c - (&3 / &4) * a - b <=> &0 <= c`] THEN MATCH_MP_TAC REAL_LE_MUL THEN CONJ_TAC THENL [ REAL_ARITH_TAC; ALL_TAC ] THEN - MATCH_MP_TAC REAL_LE_DIV THEN + MATCH_MP_TAC REAL_LE_DIV THEN ASM_SIMP_TAC[ATN_POS_LE; SQRT_POS_LE; REAL_ARITH `&1 <= x ==> &0 <= x`] THEN MATCH_MP_TAC REAL_LE_MUL THEN ASM_SIMP_TAC[REAL_LE_POW_2; SQRT_POS_LE; REAL_ARITH `&1 <= x ==> &0 <= x`]);; @@ -520,12 +532,12 @@ let matan_d2_bound_01 = prove(`!x. &0 <= x /\ x < &1 ==> STRIP_TAC THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `&1 / &4 + &3 / &20` THEN CONJ_TAC THEN TRY REAL_ARITH_TAC THEN - ABBREV_TAC `f = (&1 - x) / (&4 * (&1 + x) pow 2)` THEN + ABBREV_TAC `f = (&1 - x) / (&4 * (&1 + x) pow 2)` THEN ABBREV_TAC `r = real_infsum (:num) (\i. -- &1 pow i / &(2 * i + 5) * x pow i)` THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `abs f + abs (&3 / &4 * r)` THEN REWRITE_TAC[REAL_ABS_TRIANGLE] THEN MATCH_MP_TAC REAL_LE_ADD2 THEN CONJ_TAC THENL [ - EXPAND_TAC "f" THEN REWRITE_TAC[REAL_ABS_DIV; REAL_ABS_MUL; REAL_ABS_NUM] THEN + EXPAND_TAC "f" THEN REWRITE_TAC[REAL_ABS_DIV; REAL_ABS_MUL; REAL_ABS_NUM] THEN SUBGOAL_THEN `&0 < &4 /\ &0 < &4 * abs ((&1 + x) pow 2)` ASSUME_TAC THENL [ CONJ_TAC THEN TRY REAL_ARITH_TAC THEN MATCH_MP_TAC REAL_LT_MUL THEN CONJ_TAC THEN TRY REAL_ARITH_TAC THEN @@ -608,7 +620,7 @@ let matan_d2_bound_01 = prove(`!x. &0 <= x /\ x < &1 ==> ALL_TAC ] THEN REWRITE_TAC[REALLIM_1_OVER_N]);; - + let matan_d2_bound_neg_low = prove(`!x. -- &1 < x /\ x < &0 ==> &0 <= real_derivative (real_derivative matan) x`, @@ -688,9 +700,9 @@ let matan_d2_bound_neg_high = prove(`!x. -- #0.2 <= x /\ x < &0 ==> ] THEN REWRITE_TAC[ARITH_RULE `2 * 1 + 5 = 7`] THEN REAL_ARITH_TAC);; - + (* The main bound result *) - + let matan_d2_bound = prove(`!x. -- #0.2 <= x ==> abs (real_derivative (real_derivative matan) x) <= #0.65`, REPEAT STRIP_TAC THEN diff --git a/Formal_ineqs/trig/matan_eval.hl b/Formal_ineqs/trig/matan_eval.hl index cc299951..c5ab5c58 100644 --- a/Formal_ineqs/trig/matan_eval.hl +++ b/Formal_ineqs/trig/matan_eval.hl @@ -1,6 +1,18 @@ -needs "arith/more_float.hl";; -needs "trig/matan.hl";; -needs "trig/atn_eval.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of functions: matan, dmatan, ddmatan *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/arith/more_float.hl";; +needs "Formal_ineqs/trig/matan.hl";; +needs "Formal_ineqs/trig/atn_eval.hl";; module type Matan_eval_sig = sig @@ -67,12 +79,12 @@ let matan_gt0 = (th_rule o prove)( ASM_REWRITE_TAC[]);; -let matan_abs1 = - let aux_th = (REWRITE_CONV[FLOAT_OF_NUM; min_exp_def] THENC - DEPTH_CONV Arith_nat.NUMERAL_TO_NUM_CONV THENC +let matan_abs1 = + let aux_th = (REWRITE_CONV[FLOAT_OF_NUM; min_exp_def] THENC + DEPTH_CONV Arith_nat.NUMERAL_TO_NUM_CONV THENC REWRITE_CONV[Arith_num.NUM_THM]) `r1 <= &1 - u` in (th_rule o PURE_REWRITE_RULE[SYM (FLOAT_TO_NUM_CONV seven_float); aux_th] o prove)( - `!x. + `!x. interval_arith x (lo, hi) /\ interval_arith ((&1 - x / &3) + (x * x) / &5) (a, b) /\ iabs (lo, hi) = u /\ @@ -83,7 +95,7 @@ let matan_abs1 = r2 <= &7 * r1 /\ (&0 < r2 <=> T) /\ u3 / r2 <= e /\ low <= a - e /\ - b + e <= high + b + e <= high ==> interval_arith (matan x) (low, high)`, REWRITE_TAC[interval_arith; GSYM IMP_IMP; iabs] THEN GEN_TAC THEN REPEAT DISCH_TAC THEN SUBGOAL_THEN `abs x <= u` (LABEL_TAC "u") THENL [ @@ -143,14 +155,14 @@ let dmatan_gt0 = (th_rule o prove)( ASM_SIMP_TAC[matan_real_derivative_gt0; REAL_MUL_ASSOC]);; -let dmatan_abs1 = - let aux_conv = (REWRITE_CONV[FLOAT_OF_NUM; min_exp_def] THENC - DEPTH_CONV Arith_nat.NUMERAL_TO_NUM_CONV THENC +let dmatan_abs1 = + let aux_conv = (REWRITE_CONV[FLOAT_OF_NUM; min_exp_def] THENC + DEPTH_CONV Arith_nat.NUMERAL_TO_NUM_CONV THENC REWRITE_CONV[Arith_num.NUM_THM]) in let aux1 = aux_conv `r1 <= &1 - u` and aux2 = aux_conv `r2 <= &2 * r1` in (th_rule o PURE_REWRITE_RULE[aux1; aux2] o prove)( - `!x. + `!x. interval_arith x (lo, hi) /\ interval_arith (&2 * (x / &5) - &1 / &3) (a, b) /\ iabs (lo, hi) = u /\ @@ -160,7 +172,7 @@ let dmatan_abs1 = r2 <= &2 * r1 /\ (&0 < r2 <=> T) /\ u2 / r2 <= e /\ low <= a - e /\ - b + e <= high + b + e <= high ==> interval_arith (dmatan x) (low, high)`, REWRITE_TAC[interval_arith; GSYM IMP_IMP; iabs; dmatan] THEN GEN_TAC THEN REPEAT DISCH_TAC THEN SUBGOAL_THEN `abs x <= u` (LABEL_TAC "u") THENL [ @@ -204,7 +216,7 @@ let dmatan_abs1 = (* ddmatan *) let ddmatan_ge_neg02 = (th_rule o prove)( - `!x. interval_arith x (lo, hi) /\ + `!x. interval_arith x (lo, hi) /\ -- #0.2 <= lo /\ interval_arith (#0.65) (a, b) /\ b2 = --b @@ -244,7 +256,7 @@ let float_interval_matan = if not flag then failwith "float_interval_matan: the argument is out of range" else - let s_th = + let s_th = let ( * ) = float_interval_mul pp and ( / ) = float_interval_div pp and ( + ) = float_interval_add pp and @@ -286,7 +298,7 @@ let float_interval_matan = itlist MY_PROVE_HYP [x_th; lo_gt0; th1] th0;; -(* dmatan *) +(* dmatan *) let float_interval_dmatan = let e_var_real = `e:real` and @@ -307,7 +319,7 @@ let float_interval_dmatan = if not flag then failwith "float_interval_dmatan: the argument is out of range" else - let s_th = + let s_th = let ( * ) = float_interval_mul pp and ( / ) = float_interval_div pp and ( - ) = float_interval_sub pp in @@ -338,7 +350,7 @@ let float_interval_dmatan = (* inv ((2 * x) * (x + 1)) - atn (sqrt x) / ((2 * x) * (sqrt x)) *) let lo_gt0 = float_gt0 lo_tm in let r = float_interval_sqrt pp x_th in - let th1 = + let th1 = let ( * ) = float_interval_mul pp and ( / ) = float_interval_div pp and ( + ) = float_interval_add pp and diff --git a/Formal_ineqs/trig/poly.hl b/Formal_ineqs/trig/poly.hl index c6d94611..9bf9cdbf 100644 --- a/Formal_ineqs/trig/poly.hl +++ b/Formal_ineqs/trig/poly.hl @@ -1,5 +1,17 @@ -needs "list/more_list.hl";; -needs "arith/interval_arith.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of polynomials (theory) *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/list/more_list.hl";; +needs "Formal_ineqs/arith/interval_arith.hl";; module Poly = struct @@ -22,16 +34,16 @@ let poly_f_even = new_definition `poly_f_even cs x = poly_f cs (x * x)`;; let poly_f_odd = new_definition `poly_f_odd cs x = x * poly_f_even cs x`;; -let poly_f_empty = prove(`!x. poly_f [] x = &0`, +let poly_f_empty = prove(`!x. poly_f [] x = &0`, REWRITE_TAC[poly_f; ITLIST]);; -let poly_f_cons = prove(`!x h t. poly_f (CONS h t) x = h + x * poly_f t x`, +let poly_f_cons = prove(`!x h t. poly_f (CONS h t) x = h + x * poly_f t x`, REWRITE_TAC[poly_f; ITLIST]);; let poly_f_even_empty = prove(`!x. poly_f_even [] x = &0`, REWRITE_TAC[poly_f_even; poly_f_empty]);; -let poly_f_even_cons = prove(`!x h t. poly_f_even (CONS h t) x = h + (x * x) * poly_f_even t x`, +let poly_f_even_cons = prove(`!x h t. poly_f_even (CONS h t) x = h + (x * x) * poly_f_even t x`, REWRITE_TAC[poly_f_even; poly_f_cons]);; let poly_f_odd_empty = prove(`!x. poly_f_odd [] x = &0`, @@ -46,18 +58,18 @@ let poly_f_append = prove(`!x b a. poly_f (APPEND a b) x = poly_f a x + x pow (L REWRITE_TAC[real_pow; REAL_MUL_LID; REAL_ADD_LID]; ALL_TAC ] THEN - + REWRITE_TAC[APPEND; poly_f; ITLIST] THEN ASM_REWRITE_TAC[GSYM poly_f] THEN REWRITE_TAC[LENGTH; real_pow] THEN REAL_ARITH_TAC);; -let poly_f_even_append = prove(`!x b a. poly_f_even (APPEND a b) x +let poly_f_even_append = prove(`!x b a. poly_f_even (APPEND a b) x = poly_f_even a x + x pow (2 * LENGTH a) * poly_f_even b x`, - REWRITE_TAC[poly_f_even; poly_f_append] THEN + REWRITE_TAC[poly_f_even; poly_f_append] THEN REWRITE_TAC[GSYM REAL_POW_2; REAL_POW_POW]);; -let poly_f_odd_append = prove(`!x b a. poly_f_odd (APPEND a b) x +let poly_f_odd_append = prove(`!x b a. poly_f_odd (APPEND a b) x = poly_f_odd a x + x pow (2 * LENGTH a) * poly_f_odd b x`, REPEAT GEN_TAC THEN REWRITE_TAC[poly_f_odd] THEN @@ -78,14 +90,14 @@ let poly_f_odd_sing = prove(`!c x. poly_f_odd [c] x = c * x`, (* bounds *) -let poly_f_empty_high_pos = prove(`!x. interval_arith (poly_f [] x) (real_zero, &0)`, +let poly_f_empty_high_pos = prove(`!x. interval_arith (poly_f [] x) (real_zero, &0)`, REWRITE_TAC[real_zero; interval_arith; poly_f_empty; REAL_LE_REFL]);; let poly_f_sing_high_pos = prove(`!c x bounds. interval_arith c bounds ==> interval_arith (poly_f [c] x) bounds`, SIMP_TAC[poly_f_sing]);; -let poly_f_cons_high_pos_pos = prove(`!x h t a b c m r. +let poly_f_cons_high_pos_pos = prove(`!x h t a b c m r. interval_arith (poly_f t x) (real_zero, a) /\ interval_arith h (real_zero, b) /\ interval_arith x (real_zero, c) /\ @@ -102,14 +114,14 @@ let poly_f_cons_high_pos_pos = prove(`!x h t a b c m r. MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `c * a:real` THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_REWRITE_TAC[]);; -let poly_f_empty_low_pos = prove(`!x. interval_arith (&0) (real_zero, poly_f [] x)`, +let poly_f_empty_low_pos = prove(`!x. interval_arith (&0) (real_zero, poly_f [] x)`, REWRITE_TAC[real_zero; interval_arith; poly_f_empty; REAL_LE_REFL]);; let poly_f_sing_low_pos = prove(`!c b x. interval_arith b (real_zero, c) ==> interval_arith b (real_zero, poly_f [c] x)`, SIMP_TAC[poly_f_sing]);; -let poly_f_cons_low_pos_pos = prove(`!x h t a b c m r. +let poly_f_cons_low_pos_pos = prove(`!x h t a b c m r. interval_arith a (real_zero, poly_f t x) /\ interval_arith b (real_zero, h) /\ interval_arith c (real_zero, x) /\ @@ -117,7 +129,7 @@ let poly_f_cons_low_pos_pos = prove(`!x h t a b c m r. r <= b + m /\ real_zero <= r ==> interval_arith r (real_zero, poly_f (CONS h t) x)`, - REWRITE_TAC[poly_f_cons; interval_arith; real_zero] THEN + REWRITE_TAC[poly_f_cons; interval_arith; real_zero] THEN REPEAT STRIP_TAC THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `b + m:real` THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_ADD2 THEN ASM_REWRITE_TAC[] THEN @@ -135,11 +147,11 @@ let sum_eq_poly_f = prove(`!a x n. sum (0..n) (\i. a i * x pow i) = poly_f (TABL ONCE_REWRITE_TAC[TABLE_SUC] THEN ASM_REWRITE_TAC[poly_f_append; LENGTH_TABLE; poly_f_sing; REAL_MUL_AC]);; -let sum_eq_poly_f_even = prove(`!a x n. sum (0..n) (\i. a i * x pow (2 * i)) +let sum_eq_poly_f_even = prove(`!a x n. sum (0..n) (\i. a i * x pow (2 * i)) = poly_f_even (TABLE a (SUC n)) x`, REWRITE_TAC[GSYM REAL_POW_POW; sum_eq_poly_f; poly_f_even; REAL_POW_2]);; -let sum_eq_poly_f_odd = prove(`!a x n. sum (0..n) (\i. a i * x pow (2 * i + 1)) +let sum_eq_poly_f_odd = prove(`!a x n. sum (0..n) (\i. a i * x pow (2 * i + 1)) = poly_f_odd (TABLE a (SUC n)) x`, REWRITE_TAC[REAL_POW_ADD; REAL_MUL_ASSOC; SUM_RMUL; REAL_POW_1] THEN REWRITE_TAC[GSYM REAL_POW_POW; sum_eq_poly_f; poly_f_odd; poly_f_even] THEN @@ -150,7 +162,7 @@ let two_induct = prove(`!P. P 0 /\ P 1 /\ (!k. P k /\ P (k + 1) ==> P (k + 2)) = SUBGOAL_THEN `!k. P k /\ P (k + 1)` (fun th -> REWRITE_TAC[th]) THEN INDUCT_TAC THEN CONV_TAC NUM_REDUCE_CONV THEN ASM_REWRITE_TAC[ADD1] THEN ASM_SIMP_TAC[ARITH_RULE `(k + 1) + 1 = k + 2`]);; - + let alt_sum_eq_poly_f = prove(`!a x n. sum (0..n) (\i. ((-- &1) pow i * a i) * x pow i) = poly_f_even (TABLE (\i. a (2 * i)) (n DIV 2 + 1)) x diff --git a/Formal_ineqs/trig/poly_eval.hl b/Formal_ineqs/trig/poly_eval.hl index 75adde26..c29d26a2 100644 --- a/Formal_ineqs/trig/poly_eval.hl +++ b/Formal_ineqs/trig/poly_eval.hl @@ -1,7 +1,19 @@ -needs "trig/poly.hl";; -needs "misc/misc_functions.hl";; -needs "misc/misc_vars.hl";; -needs "arith/more_float.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of polynomials *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/trig/poly.hl";; +needs "Formal_ineqs/misc/misc_functions.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; +needs "Formal_ineqs/arith/more_float.hl";; module Poly_eval = struct @@ -79,15 +91,15 @@ type poly_coeff = { };; (* Creates polynomial coefficients from the given list of constant expressions *) -let mk_poly_coeffs = - let interval_pos = prove(`interval_arith x (float_num F n e, f) +let mk_poly_coeffs = + let interval_pos = prove(`interval_arith x (float_num F n e, f) ==> interval_arith x (&0, f) /\ interval_arith (float_num F n e) (&0, x)`, - REWRITE_TAC[interval_arith] THEN + REWRITE_TAC[interval_arith] THEN MP_TAC (SPECL[`n:num`; `e:num`] FLOAT_F_POS) THEN REAL_ARITH_TAC) and - interval_neg = prove(`interval_arith x (f, float_num T n e) + interval_neg = prove(`interval_arith x (f, float_num T n e) ==> interval_arith x (f, &0) /\ interval_arith (float_num T n e) (x, &0)`, - REWRITE_TAC[interval_arith] THEN + REWRITE_TAC[interval_arith] THEN MP_TAC (SPECL[`n:num`; `e:num`] FLOAT_T_NEG) THEN REAL_ARITH_TAC) in let mk_poly_coeff pp cs_tm = @@ -104,7 +116,7 @@ let mk_poly_coeffs = match (s1, s2) with | "F", _ -> 1, MATCH_MP interval_pos int_th | _, "T" -> -1, MATCH_MP interval_neg int_th - | _ -> 0, TRUTH in + | _ -> 0, TRUTH in let sign_high, sign_low = pair_of_list (CONJUNCTS sign_th) in { c_tm = c_tm; @@ -129,7 +141,7 @@ let eval_interval_poly_f = fun pp (cs_tm, cs_list) x_th -> let rec eval cs_tm cs_list x_tm x_th = match cs_list with - | [] -> + | [] -> INST[x_tm, x_var_real] poly_f_empty' | [first] -> let th0 = INST[first.c_tm, c_var_real; x_tm, x_var_real] poly_f_sing' in @@ -157,8 +169,8 @@ let eval_high_poly_f_pos_pos = poly_f_cons_high' = RULE' poly_f_cons_high_pos_pos in let check_pos c = if c.sign = 1 then () else - failwith (Printf.sprintf "eval_high_poly_f_pos_pos: non-positive coefficient: %s, %s" - (string_of_term c.c_tm) + failwith (Printf.sprintf "eval_high_poly_f_pos_pos: non-positive coefficient: %s, %s" + (string_of_term c.c_tm) (string_of_term c.bounds_tm)) in fun pp (cs_tm, cs_list) x_th -> let rec eval cs_tm cs_list x_tm x_th = @@ -167,7 +179,7 @@ let eval_high_poly_f_pos_pos = | [first] -> let _ = check_pos first in let bounds_tm = rand (concl first.sign_high_th) in - let th0 = INST[first.c_tm, c_var_real; bounds_tm, bounds_var; x_tm, x_var_real] + let th0 = INST[first.c_tm, c_var_real; bounds_tm, bounds_var; x_tm, x_var_real] poly_f_sing_high' in MY_PROVE_HYP first.sign_high_th th0 | first :: rest -> @@ -185,9 +197,9 @@ let eval_high_poly_f_pos_pos = let th0 = INST[h_tm, h_var_real; t_tm, t_var_real_list; x_tm, x_var_real; a_tm, a_var_real; b_tm, b_var_real; c_tm, c_var_real; m_tm, m_var_real; r_tm, r_var_real] poly_f_cons_high' in - MY_PROVE_HYP r_bound_th - (MY_PROVE_HYP m_bound_th - (MY_PROVE_HYP first.sign_high_th + MY_PROVE_HYP r_bound_th + (MY_PROVE_HYP m_bound_th + (MY_PROVE_HYP first.sign_high_th (MY_PROVE_HYP rest_th (MY_PROVE_HYP x_th th0)))) in let x_tm = rand (rator (concl x_th)) in @@ -197,15 +209,15 @@ let eval_high_poly_f_pos_pos = (* Evaluates a lower bound of a polynomial on a given non-negative number. * All coefficients of the polynomial must be non-negative. * The argument x should be in the form: |- interval_arith f (&0, x) *) -let eval_low_poly_f_pos_pos = +let eval_low_poly_f_pos_pos = let RULE' = (UNDISCH_ALL o REWRITE_RULE[real_zero] o DISCH_ALL o RULE) in let poly_f_empty_low' = RULE' poly_f_empty_low_pos and poly_f_sing_low' = RULE' poly_f_sing_low_pos and poly_f_cons_low' = RULE' poly_f_cons_low_pos_pos in let check_pos c = if c.sign = 1 then () else - failwith (Printf.sprintf "eval_low_poly_f_pos_pos: non-positive coefficient: %s, %s" - (string_of_term c.c_tm) + failwith (Printf.sprintf "eval_low_poly_f_pos_pos: non-positive coefficient: %s, %s" + (string_of_term c.c_tm) (string_of_term c.bounds_tm)) in fun pp (cs_tm, cs_list) x_th -> let rec eval cs_tm cs_list x_tm x_th = @@ -214,7 +226,7 @@ let eval_low_poly_f_pos_pos = | [first] -> let _ = check_pos first in let b_tm = (rand o rator o concl) first.sign_low_th in - let th0 = INST[first.c_tm, c_var_real; b_tm, b_var_real; x_tm, x_var_real] + let th0 = INST[first.c_tm, c_var_real; b_tm, b_var_real; x_tm, x_var_real] poly_f_sing_low' in MY_PROVE_HYP first.sign_low_th th0 | first :: rest -> @@ -234,9 +246,9 @@ let eval_low_poly_f_pos_pos = a_tm, a_var_real; b_tm, b_var_real; c_tm, c_var_real; m_tm, m_var_real; r_tm, r_var_real] poly_f_cons_low' in MY_PROVE_HYP r_ge0_th - (MY_PROVE_HYP r_bound_th - (MY_PROVE_HYP m_bound_th - (MY_PROVE_HYP first.sign_low_th + (MY_PROVE_HYP r_bound_th + (MY_PROVE_HYP m_bound_th + (MY_PROVE_HYP first.sign_low_th (MY_PROVE_HYP rest_th (MY_PROVE_HYP x_th th0))))) in let x_tm = rand (rand (concl x_th)) in @@ -253,7 +265,7 @@ let eval_pow2_high = MY_PROVE_HYP mul_th th0;; let eval_pow2_low = - let pow2_th = (UNDISCH_ALL o prove)(`(&0 <= a <=> T) ==> a <= x * x + let pow2_th = (UNDISCH_ALL o prove)(`(&0 <= a <=> T) ==> a <= x * x ==> interval_arith a (&0, x pow 2)`, SIMP_TAC[interval_arith; REAL_POW_2]) in fun pp f_tm -> @@ -292,7 +304,7 @@ let eval_pow4_low = REPEAT STRIP_TAC THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `a * a:real` THEN ASM_REWRITE_TAC[] THEN - MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_REWRITE_TAC[REAL_LE_POW_2; REAL_POW_2]) + MATCH_MP_TAC REAL_LE_MUL2 THEN ASM_REWRITE_TAC[REAL_LE_POW_2; REAL_POW_2]) in fun pp f_tm -> let mul1_th = float_mul_lo pp f_tm f_tm in @@ -306,7 +318,7 @@ let eval_pow4_low = let eval_pow2_pow4_high = let pow4_th = (UNDISCH_ALL o prove)(`interval_arith (x pow 2) (&0, a) ==> - a * a <= b ==> + a * a <= b ==> interval_arith (x pow 4) (&0, b)`, REWRITE_TAC[interval_arith; ARITH_RULE `4 = 2 + 2`; REAL_POW_ADD] THEN REPEAT STRIP_TAC THENL [ @@ -325,8 +337,8 @@ let eval_pow2_pow4_high = let eval_pow2_pow4_low = let pow4_th = (UNDISCH_ALL o prove)(`interval_arith a (&0, x pow 2) ==> - b <= a * a ==> - (&0 <= b <=> T) ==> + b <= a * a ==> + (&0 <= b <=> T) ==> interval_arith b (&0, x pow 4)`, REWRITE_TAC[interval_arith; ARITH_RULE `4 = 2 + 2`; REAL_POW_ADD] THEN REPEAT STRIP_TAC THENL [ diff --git a/Formal_ineqs/trig/series.hl b/Formal_ineqs/trig/series.hl index 439102d8..e0697b82 100644 --- a/Formal_ineqs/trig/series.hl +++ b/Formal_ineqs/trig/series.hl @@ -1,3 +1,15 @@ +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal power series *) +(* -------------------------------------------------------------------------- *) + needs "Multivariate/realanalysis.ml";; module Series = struct @@ -13,7 +25,7 @@ let abs_neg_one_pow = prove REWRITE_TAC[abs_neg_pow; REAL_ARITH `abs (&1) = &1`; REAL_POW_ONE]);; let real_infsum_offset = prove - (`!m k f. real_summable (from m) f /\ m <= k + (`!m k f. real_summable (from m) f /\ m <= k ==> real_infsum (from m) f = sum (m..k) f + real_infsum (from (k + 1)) f`, REPEAT STRIP_TAC THEN FIRST_ASSUM (MP_TAC o SPEC `k + 1` o MATCH_MP REAL_SUMMABLE_FROM_ELSEWHERE) THEN @@ -26,7 +38,7 @@ let real_infsum_offset = prove REAL_ARITH_TAC);; let real_infsum_offset_alt = prove - (`!m k f. real_summable (from m) f /\ m < k + (`!m k f. real_summable (from m) f /\ m < k ==> real_infsum (from m) f = sum (m..k - 1) f + real_infsum (from k) f`, REPEAT STRIP_TAC THEN SUBGOAL_THEN `?t. k = t + 1 /\ m <= t` STRIP_ASSUME_TAC THENL [ @@ -40,7 +52,7 @@ let power_series_simple_bound = prove (`!a u x m. real_summable (from m) (\i. a i * x pow i) /\ (!i. m <= i ==> abs (a i) <= u) /\ abs x < &1 - ==> abs (real_infsum (from m) (\i. a i * x pow i)) + ==> abs (real_infsum (from m) (\i. a i * x pow i)) <= u * abs x pow m / (&1 - abs x)`, REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_SERIES_BOUND THEN @@ -62,7 +74,7 @@ let power_series_simple_bound2 = prove (`!a x m. real_summable (from m) (\i. a i * x pow i) /\ (!i. m <= i ==> abs (a (SUC i)) <= abs (a i)) /\ abs x < &1 - ==> abs (real_infsum (from m) (\i. a i * x pow i)) + ==> abs (real_infsum (from m) (\i. a i * x pow i)) <= abs (a m) * abs x pow m / (&1 - abs x)`, REPEAT STRIP_TAC THEN MATCH_MP_TAC power_series_simple_bound THEN ASM_REWRITE_TAC[] THEN @@ -75,7 +87,7 @@ let real_sums_simple_bound = prove /\ n < m /\ (!i. m <= i ==> abs (a i) <= u) /\ abs x < &1 - ==> abs (t - sum (n..m - 1) (\i. a i * x pow i)) + ==> abs (t - sum (n..m - 1) (\i. a i * x pow i)) <= u * abs x pow m / (&1 - abs x)`, REPEAT STRIP_TAC THEN FIRST_ASSUM (ASSUME_TAC o MATCH_MP REAL_SUMS_SUMMABLE) THEN @@ -93,7 +105,7 @@ let real_sums_simple_bound2 = prove /\ n < m /\ (!i. m <= i ==> abs (a (SUC i)) <= abs (a i)) /\ abs x < &1 - ==> abs (t - sum (n..m - 1) (\i. a i * x pow i)) + ==> abs (t - sum (n..m - 1) (\i. a i * x pow i)) <= abs (a m) * abs x pow m / (&1 - abs x)`, REPEAT STRIP_TAC THEN FIRST_ASSUM (ASSUME_TAC o MATCH_MP REAL_SUMS_SUMMABLE) THEN @@ -123,9 +135,9 @@ let sign_mul_lemma = prove POP_ASSUM (MP_TAC o REWRITE_RULE[GSYM REAL_LT_SQUARE]) THEN ONCE_REWRITE_TAC[REAL_ARITH `(b * b) * x <= &0 <=> &0 <= (--(b * b) * x)`] THEN REWRITE_TAC[REAL_MUL_POS_LE] THEN REAL_ARITH_TAC);; - + let sum_bound_lemma = prove - (`!m n f. + (`!m n f. (!k. abs (f (k + 1)) <= abs (f k) /\ f (k + 1) * f k <= &0) ==> abs (sum (m..n) f) <= abs (f m) /\ &0 <= sum (m..n) f * f m`, GEN_TAC THEN GEN_TAC THEN @@ -191,9 +203,9 @@ let sum_bound_lemma = prove let alt_sign_converges = prove (`!f k. (!n. abs (f (n + 1)) <= abs (f n) /\ f (n + 1) * f n <= &0) /\ - ((f ---> &0) sequentially) + ((f ---> &0) sequentially) ==> ?l. (f real_sums l) (from k)`, - REWRITE_TAC[REAL_SERIES_CAUCHY; REALLIM_SEQUENTIALLY; GSYM IMP_IMP; REAL_SUB_RZERO] THEN + REWRITE_TAC[REAL_SERIES_CAUCHY; REALLIM_SEQUENTIALLY; GSYM IMP_IMP; REAL_SUB_RZERO] THEN REPEAT GEN_TAC THEN DISCH_THEN (LABEL_TAC "h1") THEN REPEAT STRIP_TAC THEN REWRITE_TAC[FROM_INTER_NUMSEG_GEN] THEN FIRST_X_ASSUM (MP_TAC o SPEC `e:real`) THEN ASM_REWRITE_TAC[] THEN STRIP_TAC THEN @@ -213,13 +225,13 @@ let alt_sign_converges = prove let alt_sign_summable = prove (`!f k. (!n. abs (f (n + 1)) <= abs (f n) /\ f (n + 1) * f n <= &0) /\ - ((f ---> &0) sequentially) + ((f ---> &0) sequentially) ==> real_summable (from k) f`, ASM_SIMP_TAC[real_summable; alt_sign_converges]);; let alt_sign_converges2 = prove (`!f k. (!n. abs (f (n + 1)) <= abs (f n) /\ f (n + 1) * f n <= &0) /\ - ((f ---> &0) sequentially) + ((f ---> &0) sequentially) ==> ?l. (f real_sums l) (from k) /\ abs l <= abs (f k) /\ &0 <= l * f k`, REPEAT STRIP_TAC THEN @@ -258,7 +270,7 @@ let alt_sign_converges2 = prove let alt_sign_bound = prove (`!f k. (!n. abs (f (n + 1)) <= abs (f n) /\ f (n + 1) * f n <= &0) /\ - ((f ---> &0) sequentially) + ((f ---> &0) sequentially) ==> abs (real_infsum (from k) f) <= abs (f k)`, REPEAT STRIP_TAC THEN MP_TAC (SPEC_ALL alt_sign_converges2) THEN ASM_REWRITE_TAC[] THEN @@ -273,7 +285,7 @@ let alt_sign_abs_bound = prove MP_TAC (SPEC_ALL alt_sign_bound) THEN ASM_SIMP_TAC[alt_sign_summable] THEN REAL_ARITH_TAC);; - + let alt_sign_upper_bound = prove (`!f m k. (!n. abs (f (n + 1)) <= abs (f n) /\ f (n + 1) * f n <= &0) /\ (f ---> &0) sequentially /\ m < k @@ -353,7 +365,7 @@ let integral_uniform_limit = prove let integral_series = prove (`!f g s r a b. - (!e. &0 < e ==> ?N. !n x. N <= n /\ x IN real_interval[a,b] + (!e. &0 < e ==> ?N. !n x. N <= n /\ x IN real_interval[a,b] ==> abs (sum (s INTER (0..n)) (\i. f i x) - g x) < e) /\ (!n. ((f n) has_real_integral (r n)) (real_interval[a,b])) ==> (?v. (g has_real_integral v) (real_interval[a,b]) /\ (r real_sums v) s)`, @@ -368,7 +380,7 @@ let integral_series = prove let power_series_uniform = prove (`!a s t y. real_summable s (\i. a i * y pow i) /\ abs t < abs y ==> (!e. &0 < e ==> ?N. !n x. N <= n /\ x IN real_interval[--abs t, abs t] - ==> abs (sum (s INTER (0..n)) (\i. a i * x pow i) + ==> abs (sum (s INTER (0..n)) (\i. a i * x pow i) - real_infsum s (\i. a i * x pow i)) < e)`, REPEAT STRIP_TAC THEN ABBREV_TAC `M = (abs y + abs t) / (abs y - abs t)` THEN @@ -473,7 +485,7 @@ let has_real_integral_pow_neg = prove let power_series_integral = prove (`!f a s r. (!x. abs x < r ==> ((\i. a i * x pow i) real_sums f x) s) ==> (!x. abs x < r ==> - let v = if x < &0 then --real_integral (real_interval[x, &0]) f + let v = if x < &0 then --real_integral (real_interval[x, &0]) f else real_integral (real_interval[&0, x]) f in ((\i. a i * x pow (i + 1) / &(i + 1)) real_sums v) s)`, REPEAT STRIP_TAC THEN LET_TAC THEN @@ -535,12 +547,12 @@ let power_series_integral = prove ALL_TAC ] THEN REWRITE_TAC[REAL_MUL_LID]);; - + let power_series_integral2 = prove - (`!f f' a s r. + (`!f f' a s r. (!x. abs x < r ==> (f has_real_derivative f' x) (atreal x)) /\ (!x. abs x < r ==> ((\i. a i * x pow i) real_sums f' x) s) - ==> (!x. abs x < r + ==> (!x. abs x < r ==> ((\i. a i * x pow (i + 1) / &(i + 1)) real_sums f x - f (&0)) s)`, REWRITE_TAC[GSYM IMP_IMP] THEN REPEAT GEN_TAC THEN DISCH_TAC THEN DISCH_THEN (MP_TAC o MATCH_MP power_series_integral) THEN REPEAT STRIP_TAC THEN @@ -578,7 +590,7 @@ let strict_increasing_num_le = prove GEN_TAC THEN DISCH_TAC THEN INDUCT_TAC THEN TRY ARITH_TAC THEN FIRST_X_ASSUM (MP_TAC o SPEC `n:num`) THEN REWRITE_TAC[ADD1] THEN POP_ASSUM MP_TAC THEN ARITH_TAC);; - + let strict_increasing_num_lt = prove (`!h:num->num. (!n. h n < h (n + 1)) ==> (!n m. n < m ==> h n < h m)`, GEN_TAC THEN DISCH_TAC THEN INDUCT_TAC THEN INDUCT_TAC THEN TRY ARITH_TAC THENL [ @@ -612,11 +624,11 @@ let strict_increasing_num_le2 = prove let real_sums_image = prove - (`!f h s l. (!n. h n < h (n + 1)) + (`!f h s l. (!n. h n < h (n + 1)) ==> ((f real_sums l) (IMAGE h s) <=> ((f o h) real_sums l) s)`, REWRITE_TAC[real_sums; REALLIM_SEQUENTIALLY] THEN REPEAT STRIP_TAC THEN EQ_TAC THEN REPEAT STRIP_TAC THENL [ - POP_ASSUM (fun th -> FIRST_X_ASSUM (STRIP_ASSUME_TAC o C MATCH_MP th)) THEN + POP_ASSUM (fun th -> FIRST_X_ASSUM (STRIP_ASSUME_TAC o C MATCH_MP th)) THEN EXISTS_TAC `N:num` THEN REPEAT STRIP_TAC THEN SUBGOAL_THEN `?m. N <= m /\ IMAGE (h:num->num) s INTER (0..m) = IMAGE h (s INTER (0..n))` STRIP_ASSUME_TAC THENL [ EXISTS_TAC `(h:num->num) n` THEN @@ -681,7 +693,7 @@ let real_sums_image = prove POP_ASSUM (MP_TAC o MATCH_MP SUM_IMAGE) THEN SIMP_TAC[]);; let real_summable_image = prove - (`!f h s. (!n. h n < h (n + 1)) + (`!f h s. (!n. h n < h (n + 1)) ==> (real_summable (IMAGE h s) f <=> real_summable s (f o h))`, REPEAT GEN_TAC THEN DISCH_THEN (MP_TAC o MATCH_MP real_sums_image) THEN REWRITE_TAC[real_summable] THEN DISCH_TAC THEN EQ_TAC THEN REPEAT STRIP_TAC THENL [ @@ -792,7 +804,7 @@ let n_le_pow_lemma = prove MATCH_MP_TAC REAL_POW_LE THEN ASM_SIMP_TAC[REAL_ARITH `&0 < e ==> &0 <= e`]);; let real_summable_lmul_eq = prove - (`!c f s. ~(c = &0) + (`!c f s. ~(c = &0) ==> (real_summable s (\i. c * f i) <=> real_summable s f)`, REPEAT STRIP_TAC THEN EQ_TAC THEN DISCH_TAC THENL [ SUBGOAL_THEN `f = (\i:num. inv c * c * f i)` (fun th -> ONCE_REWRITE_TAC[th]) THENL [ @@ -970,7 +982,7 @@ let power_series_has_derivative = prove ==> ?f'. (!x. abs x < r ==> (f has_real_derivative f' x) (atreal x) /\ ((\i. (&i * a i) * x pow (i - 1)) real_sums f' x) s)`, REPEAT STRIP_TAC THEN - SUBGOAL_THEN `(!w. &0 < w /\ abs w < r /\ w < r /\ abs w < abs r ==> + SUBGOAL_THEN `(!w. &0 < w /\ abs w < r /\ w < r /\ abs w < abs r ==> ?f'. (!x. abs x < w ==> (f has_real_derivative f' x) (atreal x) /\ ((\i. (&i * a i) * x pow (i - 1)) real_sums f' x) s)) ==> ?f'. (!x. abs x < r ==> (f has_real_derivative f' x) (atreal x) /\ @@ -983,7 +995,7 @@ let power_series_has_derivative = prove ALL_TAC ] THEN REWRITE_TAC[] THEN - DISCH_THEN MATCH_MP_TAC THEN + DISCH_THEN MATCH_MP_TAC THEN POP_ASSUM MP_TAC THEN REAL_ARITH_TAC; ALL_TAC ] THEN @@ -1096,7 +1108,7 @@ let power_series_has_derivative = prove ASM_REWRITE_TAC[]);; let power_series_has_derivative2 = prove - (`!f a r. + (`!f a r. (!x. abs x < r ==> ((\i. a i * x pow i) real_sums f x) (:num)) ==> ?f'. (!x. abs x < r ==> (f has_real_derivative f' x) (atreal x) /\ @@ -1142,7 +1154,7 @@ let real_sums_i_xi = prove let real_sums_neg_i_xi = prove (`!x. abs x < &1 - ==> ((\i. ((-- &1) pow i * &i) * x pow i) real_sums + ==> ((\i. ((-- &1) pow i * &i) * x pow i) real_sums --(x / (&1 + x) pow 2)) (:num)`, REPEAT STRIP_TAC THEN MP_TAC (SPEC `--x` real_sums_i_xi) THEN ANTS_TAC THENL [ @@ -1215,7 +1227,7 @@ let log_series = prove REWRITE_TAC[REAL_ARITH `a * x / b = a / b * x`]);; let log_infsum = prove - (`!x. abs x < &1 ==> + (`!x. abs x < &1 ==> real_infsum (from 0) (\k. ((-- &1) pow k / &(k + 1)) * x pow (k + 1)) = log (&1 + x)`, GEN_TAC THEN DISCH_THEN (MP_TAC o MATCH_MP REAL_INFSUM_UNIQUE o MATCH_MP log_series) THEN REWRITE_TAC[FROM_0]);; @@ -1224,7 +1236,7 @@ let log_infsum = prove let atn_derivative_series = prove (`!x. abs x < &1 - ==> ((\k. (if EVEN k then (-- &1) pow (k DIV 2) else &0) * x pow k) + ==> ((\k. (if EVEN k then (-- &1) pow (k DIV 2) else &0) * x pow k) real_sums inv (&1 + x pow 2)) (:num)`, REPEAT STRIP_TAC THEN MP_TAC (SPECL[`0`; `(-- &1) * x pow 2`] REAL_SUMS_GP) THEN @@ -1249,7 +1261,7 @@ let atn_series = prove REWRITE_TAC[ARITH_RULE `(2 * i) DIV 2 = i`; real_div; REAL_MUL_AC]);; let atn_infsum = prove - (`!x. abs x < &1 ==> + (`!x. abs x < &1 ==> real_infsum (from 0) (\k. ((-- &1) pow k / &(2 * k + 1)) * x pow (2 * k + 1)) = atn x`, GEN_TAC THEN DISCH_THEN (MP_TAC o MATCH_MP REAL_INFSUM_UNIQUE o MATCH_MP atn_series) THEN REWRITE_TAC[FROM_0]);; diff --git a/Formal_ineqs/trig/sin_cos.hl b/Formal_ineqs/trig/sin_cos.hl index 6bd9ede7..9cd9e5d8 100644 --- a/Formal_ineqs/trig/sin_cos.hl +++ b/Formal_ineqs/trig/sin_cos.hl @@ -1,3 +1,15 @@ +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of functions: sin, cos (theory) *) +(* -------------------------------------------------------------------------- *) + module Sin_cos = struct prioritize_real();; @@ -7,7 +19,7 @@ prioritize_real();; let cos_pi_2 = prove(`!x. cos (x + pi / &2) = --sin x`, REWRITE_TAC[COS_SIN; REAL_ARITH `a - (b + a) = --b:real`; SIN_NEG]);; -let cos_derivatives = prove(`!x n. ((\x. cos (x + &n * pi / &2)) +let cos_derivatives = prove(`!x n. ((\x. cos (x + &n * pi / &2)) has_real_derivative cos (x + &(n + 1) * pi / &2)) (atreal x)`, REPEAT GEN_TAC THEN REAL_DIFF_TAC THEN REWRITE_TAC[REAL_ARITH `(&1 + &0) * --a = --a`] THEN @@ -15,8 +27,8 @@ let cos_derivatives = prove(`!x n. ((\x. cos (x + &n * pi / &2)) REWRITE_TAC[REAL_ARITH `x + (a + &1) * t = (x + a * t) + t`] THEN REWRITE_TAC[cos_pi_2]);; -let taylor_cos_raw = prove(`!x n. - abs (cos x - sum (0..n) (\k. if (EVEN k) then ((-- &1) pow (k DIV 2) * x pow k) / &(FACT k) else &0)) +let taylor_cos_raw = prove(`!x n. + abs (cos x - sum (0..n) (\k. if (EVEN k) then ((-- &1) pow (k DIV 2) * x pow k) / &(FACT k) else &0)) <= abs x pow (n + 1) / &(FACT (n + 1))`, REPEAT GEN_TAC THEN MP_TAC (SPECL [`(\i x. cos (x + &i * pi / &2))`; `n:num`; `(:real)`; `&1`] REAL_TAYLOR) THEN @@ -62,8 +74,8 @@ let sum_pair_from_0 = prove(`!f n. sum (0..2 * n + 1) f = sum(0..n) (\i. f (2 * MP_TAC (SPECL [`f:num->real`; `0`; `n:num`] SUM_PAIR) THEN REWRITE_TAC[ARITH_RULE `2 * 0 = 0`]);; -let taylor_cos = prove(`!x n. abs (cos x - - sum (0..n) (\i. ((-- &1) pow i / &(FACT (2 * i))) * x pow (2 * i))) +let taylor_cos = prove(`!x n. abs (cos x - + sum (0..n) (\i. ((-- &1) pow i / &(FACT (2 * i))) * x pow (2 * i))) <= abs x pow (2*n + 2) / &(FACT (2*n + 2))`, REPEAT GEN_TAC THEN MP_TAC (SPECL [`x:real`; `2 * n + 1`] taylor_cos_raw) THEN @@ -137,7 +149,7 @@ let cos_poly_lower_bound = prove(`!x n. ODD n ==> let sin_pi_2 = prove(`!x. sin (x + pi / &2) = cos x`, REWRITE_TAC[SIN_COS; REAL_ARITH `a - (b + a) = --b:real`; COS_NEG]);; -let sin_derivatives = prove(`!x n. ((\x. sin (x + &n * pi / &2)) +let sin_derivatives = prove(`!x n. ((\x. sin (x + &n * pi / &2)) has_real_derivative sin (x + &(n + 1) * pi / &2)) (atreal x)`, REPEAT GEN_TAC THEN REAL_DIFF_TAC THEN REWRITE_TAC[REAL_ARITH `(&1 + &0) * a = a`] THEN @@ -145,8 +157,8 @@ let sin_derivatives = prove(`!x n. ((\x. sin (x + &n * pi / &2)) REWRITE_TAC[REAL_ARITH `x + (a + &1) * t = (x + a * t) + t`] THEN REWRITE_TAC[sin_pi_2]);; -let taylor_sin_raw = prove(`!x n. - abs (sin x - sum (0..n) (\k. if (ODD k) then ((-- &1) pow (k DIV 2) * x pow k) / &(FACT k) else &0)) +let taylor_sin_raw = prove(`!x n. + abs (sin x - sum (0..n) (\k. if (ODD k) then ((-- &1) pow (k DIV 2) * x pow k) / &(FACT k) else &0)) <= abs x pow (n + 1) / &(FACT (n + 1))`, REPEAT GEN_TAC THEN MP_TAC (SPECL [`(\i x. sin (x + &i * pi / &2))`; `n:num`; `(:real)`; `&1`] REAL_TAYLOR) THEN @@ -185,8 +197,8 @@ let taylor_sin_raw = prove(`!x n. REWRITE_TAC[REAL_MUL_LZERO]);; -let taylor_sin = prove(`!x n. abs (sin x - - sum (0..n) (\i. ((-- &1) pow i / &(FACT (2 * i + 1))) * x pow (2 * i + 1))) +let taylor_sin = prove(`!x n. abs (sin x - + sum (0..n) (\i. ((-- &1) pow i / &(FACT (2 * i + 1))) * x pow (2 * i + 1))) <= abs x pow (2 * n + 3) / &(FACT (2 * n + 3))`, REPEAT GEN_TAC THEN MP_TAC (SPECL [`x:real`; `2 * n + 2`] taylor_sin_raw) THEN diff --git a/Formal_ineqs/trig/sin_eval.hl b/Formal_ineqs/trig/sin_eval.hl index dd051c38..ce56edff 100644 --- a/Formal_ineqs/trig/sin_eval.hl +++ b/Formal_ineqs/trig/sin_eval.hl @@ -1,8 +1,20 @@ -needs "misc/misc_functions.hl";; -needs "misc/misc_vars.hl";; -needs "arith/interval_arith.hl";; -needs "trig/atn_eval.hl";; -needs "trig/cos_eval.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2014 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal interval evaluation of functions: sin *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/misc/misc_functions.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; +needs "Formal_ineqs/arith/interval_arith.hl";; +needs "Formal_ineqs/trig/atn_eval.hl";; +needs "Formal_ineqs/trig/cos_eval.hl";; module type Sin_eval_sig = sig @@ -70,12 +82,12 @@ let float_interval_sin pp x_th = let bounds = rand (concl cos_x_pi2) in let ltm, high_tm = dest_comb bounds in let low_tm = rand ltm in - let th0 = INST[x_tm, x_var_real; low_tm, low_var_real; + let th0 = INST[x_tm, x_var_real; low_tm, low_var_real; high_tm, high_var_real] sin_reduce_th in MY_PROVE_HYP cos_x_pi2 th0;; -end;; +end;; (* diff --git a/Formal_ineqs/trig/test.hl b/Formal_ineqs/trig/test.hl index 83fc3dc0..144de2a4 100644 --- a/Formal_ineqs/trig/test.hl +++ b/Formal_ineqs/trig/test.hl @@ -1,4 +1,4 @@ -needs "trig/atn_eval.hl";; +needs "Formal_ineqs/trig/atn_eval.hl";; needs "arith/float_atn.hl";; open Misc_functions;; @@ -41,7 +41,7 @@ let float_interval_halfatn pp x_th = let float_interval_halfatn4 pp x_th = let x_tm = (rand o rator o concl) x_th in - let r_th = float_interval_halfatn pp + let r_th = float_interval_halfatn pp (float_interval_halfatn pp (float_interval_halfatn pp (float_interval_halfatn pp x_th))) in let th0 = INST[x_tm, x_var_real] HALFATN4' in diff --git a/Formal_ineqs/trig/unused.hl b/Formal_ineqs/trig/unused.hl index f3881c98..9de239e9 100644 --- a/Formal_ineqs/trig/unused.hl +++ b/Formal_ineqs/trig/unused.hl @@ -17,7 +17,7 @@ let float_interval_halfatn pp x_th = let float_interval_halfatn4 pp x_th = let x_tm = (rand o rator o concl) x_th in - let r_th = float_interval_halfatn pp + let r_th = float_interval_halfatn pp (float_interval_halfatn pp (float_interval_halfatn pp (float_interval_halfatn pp x_th))) in let th0 = INST[x_tm, x_var_real] HALFATN4' in @@ -79,7 +79,7 @@ let bound_low_trans = prove x * b <= n /\ t <= a - n /\ r <= x * t - ==> r <= e`, + ==> r <= e`, REWRITE_TAC[interval_arith] THEN REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `x * (p1 - x * p2)` THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `x * (a - n)` THEN CONJ_TAC THENL [ @@ -91,7 +91,7 @@ let bound_low_trans = prove REWRITE_TAC[real_sub] THEN MATCH_MP_TAC REAL_LE_ADD2 THEN ASM_REWRITE_TAC[REAL_LE_NEG2] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `x * b:real` THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_LMUL THEN ASM_REWRITE_TAC[]);; - + let bound_high_trans = prove (`e <= x * (p1 - x * p2) ==> interval_arith p1 (&0, a) /\ @@ -100,7 +100,7 @@ let bound_high_trans = prove n <= x * b /\ a - n <= t /\ x * t <= r - ==> e <= r`, + ==> e <= r`, REWRITE_TAC[interval_arith] THEN REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `x * (p1 - x * p2)` THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `x * (a - n)` THEN CONJ_TAC THENL [ @@ -113,13 +113,13 @@ let bound_high_trans = prove ]);; let mk_log_bound_tables log_bound_th = - let log_bound = (SPEC_ALL o + let log_bound = (SPEC_ALL o REWRITE_RULE[poly_f_even; poly_f_odd; GSYM REAL_POW_2; REAL_POW_POW] o REWRITE_RULE[sum_eq_poly_f] o REWRITE_RULE[alt_sum_eq_poly_f; real_div]) log_bound_th in - let reduce_rule = CONV_RULE (DEPTH_CONV + let reduce_rule = CONV_RULE (DEPTH_CONV (FIRST_CONV [NUM_SUC_CONV; NUM_ADD_CONV; NUM_MULT_CONV])) in - let find_poly_f = rev o find_terms (fun tm -> try (rator o rator) tm = `poly_f` + let find_poly_f = rev o find_terms (fun tm -> try (rator o rator) tm = `poly_f` with Failure _ -> false) in fun pp n -> let n_tm = mk_small_numeral n in @@ -135,7 +135,7 @@ let mk_log_bound_tables log_bound_th = bound_th, zip cs_tms cs_lists;; (* Computes i such that x^i / i <= base^(-(p + 1)) and cond(i) *) -let n_of_p_log x pp cond = +let n_of_p_log x pp cond = let t = (float_of_int Arith_num.arith_base) ** (float_of_int (-pp - 1)) in let rec try_i i = let _ = if i > 50 then failwith "n_of_p_exp: cannot find i" else () in @@ -143,7 +143,7 @@ let n_of_p_log x pp cond = (* let d = float_of_int (2 * i + 1) in *) let d = float_of_int i in let r = (x ** d) /. d in - if r <= t then i else try_i (i + 1) + if r <= t then i else try_i (i + 1) else try_i (i + 1) in @@ -204,7 +204,7 @@ let log1_pos_low = let r_tm = rand (rator (concl r_le_xt)) in let cmp_1 = EQT_ELIM (float_lt x_tm one_float) in let cmp_0 = EQT_ELIM (float_ge0 x_tm) in - let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; n_tm, n_var_real; + let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; n_tm, n_var_real; t_tm, t_var_real; r_tm, r_var_real; x_tm, x_var_real] bound_th in itlist MY_PROVE_HYP [p1_low_th; p2_high_th; xb_le_n; t_le_an; r_le_xt; cmp_0; cmp_1] th0;; @@ -229,7 +229,7 @@ let log1_pos_high = let r_tm = rand (concl xt_le_r) in let cmp_1 = EQT_ELIM (float_lt x_tm one_float) in let cmp_0 = EQT_ELIM (float_ge0 x_tm) in - let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; n_tm, n_var_real; + let th0 = INST[a_tm, a_var_real; b_tm, b_var_real; n_tm, n_var_real; t_tm, t_var_real; r_tm, r_var_real; x_tm, x_var_real] bound_th in itlist MY_PROVE_HYP [p1_high_th; p2_low_th; n_le_xb; an_le_t; xt_le_r; cmp_1; cmp_0] th0;; diff --git a/Formal_ineqs/update_paths.py b/Formal_ineqs/update_paths.py new file mode 100644 index 00000000..b2a96cb0 --- /dev/null +++ b/Formal_ineqs/update_paths.py @@ -0,0 +1,47 @@ +# This script updates dependency paths for all files in this project + +import argparse +import os +import glob +import re + +def parse_args(): + parser = argparse.ArgumentParser( + description='Updates dependency paths for all files in this project') + parser.add_argument('--prefix', default='Formal_ineqs', + help='path which should be prepended to all dependency paths') + args = parser.parse_args() + args.root = os.path.dirname(os.path.realpath(__file__)) + return args + +def collect_local_files(args): + return set(glob.glob('**/*.hl', root_dir=args.root, recursive=True)) | \ + set(glob.glob('**/*.vhl', root_dir=args.root, recursive=True)) + +def update_file(path, local_files, args): + print(f'Updating: {path}') + with open(path, 'r') as f: + text = f.read() + def sub(m): + if m[2] in local_files: + return f'{m[1]}{args.prefix}/{m[2]}{m[3]}' + else: + return m[0] + text2 = re.sub(r'\b((?:needs|loads|loadt)\s*\\?")([^\\"]+)(\\?")', sub, text) + if text == text2: + print('No changes') + else: + with open(path, 'w') as f: + f.write(text2) + +def main(): + args = parse_args() + if not args.prefix: + print('Empty prefix') + return + local_files = collect_local_files(args) + for fname in sorted(local_files): + update_file(os.path.join(args.root, fname), local_files, args) + +if __name__ == "__main__": + main() \ No newline at end of file diff --git a/Formal_ineqs/verifier/certificate.hl b/Formal_ineqs/verifier/certificate.hl index 2c4e3e6b..f8fa22c1 100644 --- a/Formal_ineqs/verifier/certificate.hl +++ b/Formal_ineqs/verifier/certificate.hl @@ -1,4 +1,16 @@ -needs "misc/report.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal verification certificate *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/misc/report.hl";; module Certificate = struct @@ -18,7 +30,7 @@ type mono_status = { type result_tree = | Result_false of (float list * float list) (* function number, raw flag *) - | Result_pass of (int * bool) + | Result_pass of (int * bool) | Result_pass_mono of mono_status | Result_pass_ref of int | Result_mono of mono_status list * result_tree @@ -28,7 +40,7 @@ type result_tree = type p_status = { pp : int; };; - + type p_result_tree = | P_result_pass of p_status * int * bool | P_result_mono of p_status * mono_status list * p_result_tree @@ -67,11 +79,11 @@ type certificate_stats = { };; let dummy_stats = { - pass = 0; - pass_raw = 0; + pass = 0; + pass_raw = 0; pass_mono = 0; - mono = 0; - glue = 0; + mono = 0; + glue = 0; glue_convex = 0; };; @@ -84,7 +96,7 @@ let result_stats result = pass_mono = ref 0 and pass_raw = ref 0 and glue_convex = ref 0 in - + let rec count r = match r with | Result_false _ -> failwith "False result" @@ -126,7 +138,7 @@ let result_p_stats glue_flag p_result = count r1; count r2 in let _ = count p_result in - let s = Hashtbl.fold + let s = Hashtbl.fold (fun p c s -> (sprintf "p = %d: %d\n" p c) ^ s) p_table "" in report s;; @@ -198,7 +210,7 @@ let transform_result x z r = (* get_domain *) (* Subdivides the given domain (x,z) according to the given path *) let domain_hash = Hashtbl.create 1000 in - let find_hash, mem_hash, add_hash = + let find_hash, mem_hash, add_hash = Hashtbl.find domain_hash, Hashtbl.mem domain_hash, Hashtbl.add domain_hash in let get_domain path = @@ -218,7 +230,7 @@ let transform_result x z r = let ( ++ ), ( / ) = (+.), (/.) in let yj = (List.nth x j ++ List.nth z j) / 2.0 in let delta b v = table (fun i -> if i = j && b then yj else List.nth v i) in - if s = "l" then + if s = "l" then delta false x, delta true z else delta true x, delta false z @@ -259,11 +271,11 @@ let transform_result x z r = let t1 = rec_transform p1 r1 in let t2 = rec_transform p2 r2 in Result_glue (j, convex_flag, t1, t2) - | Result_pass_mono m -> + | Result_pass_mono m -> let path' = rev (get_m m :: path) in let x', z' = get_domain path' in let _, i = find_domain x' z' in - (* let _ = report (sprintf "p = %s, d = %s, found: %d" + (* let _ = report (sprintf "p = %s, d = %s, found: %d" (domain_str x' z') (path_str path') i) in *) if i >= 0 then Result_mono ([m], Result_pass_ref (-i)) else r | _ -> r in diff --git a/Formal_ineqs/verifier/m_verifier.hl b/Formal_ineqs/verifier/m_verifier.hl index c88de178..b020782b 100644 --- a/Formal_ineqs/verifier/m_verifier.hl +++ b/Formal_ineqs/verifier/m_verifier.hl @@ -1,14 +1,20 @@ -(* =========================================================== *) -(* Formal verification functions *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -needs "taylor/m_taylor.hl";; -needs "verifier/certificate.hl";; -needs "misc/misc_vars.hl";; -needs "verifier_options.hl";; -needs "verifier/m_verifier_build.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Formal verification functions *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/taylor/m_taylor.hl";; +needs "Formal_ineqs/verifier/certificate.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; +needs "Formal_ineqs/verifier_options.hl";; +needs "Formal_ineqs/verifier/m_verifier_build.hl";; module M_verifier = struct @@ -81,7 +87,7 @@ let eval_subset_trans = s_var = mk_var ("s", type_of s_tm) and t_var = mk_var ("t", type_of t_tm) and u_var = mk_var ("u", type_of u_tm) in - (MY_PROVE_HYP st_th o MY_PROVE_HYP tu_th o + (MY_PROVE_HYP st_th o MY_PROVE_HYP tu_th o INST[s_tm, s_var; t_tm, t_var; u_tm, u_var] o inst_first_type_var ty) SUBSET_TRANS';; let eval_subset_refl = @@ -90,7 +96,7 @@ let eval_subset_refl = let ty = (hd o snd o dest_type o type_of) s_tm and s_var = mk_var ("s", type_of s_tm) in (INST[s_tm, s_var] o inst_first_type_var ty) SUBSET_REFL';; - + let SUBSET_INTERVAL_IMP = prove(`!a b c d. (!i. i IN 1..dimindex (:N) ==> a$i <= c$i /\ d$i <= b$i) ==> @@ -115,7 +121,7 @@ let gen_subset_interval_lemma n = MY_RULE th2;; -let subset_interval_thms_array = Array.init (max_dim + 1) +let subset_interval_thms_array = Array.init (max_dim + 1) (fun n -> if n < 1 then TRUTH else gen_subset_interval_lemma n);; @@ -131,7 +137,7 @@ let m_subset_interval n a_tm b_tm c_tm d_tm = c_s = dest_vector c_tm and d_s = dest_vector d_tm in - let th0 = (INST (zip a_s a_vars) o INST (zip b_s b_vars) o + let th0 = (INST (zip a_s a_vars) o INST (zip b_s b_vars) o INST (zip c_s c_vars) o INST (zip d_s d_vars)) subset_interval_thms_array.(n) in let prove_le tm = let ltm, rtm = dest_binop le_op_real tm in @@ -143,9 +149,9 @@ let m_subset_interval n a_tm b_tm c_tm d_tm = (*************************************) let M_RESTRICT_RIGHT_LEMMA = prove(`!j x z y w u y' w'. m_cell_domain (x:real^N,z) y w /\ - (!i. i IN 1..dimindex (:N) ==> ~(i = j) ==> + (!i. i IN 1..dimindex (:N) ==> ~(i = j) ==> u$i = x$i /\ y'$i = y$i /\ w'$i = w$i) /\ - u$j = z$j /\ y'$j = z$j /\ w'$j = &0 ==> + u$j = z$j /\ y'$j = z$j /\ w'$j = &0 ==> m_cell_domain (u,z) y' w' /\ interval [u,z] SUBSET interval [x,z]`, REWRITE_TAC[m_cell_domain; SUBSET_INTERVAL; GSYM IN_NUMSEG] THEN REPEAT GEN_TAC THEN STRIP_TAC THEN CONJ_TAC THENL @@ -173,9 +179,9 @@ let M_RESTRICT_RIGHT_LEMMA = prove(`!j x z y w u y' w'. m_cell_domain (x:real^N, let M_RESTRICT_LEFT_LEMMA = prove(`!j x z y w u y' w'. m_cell_domain (x:real^N,z) y w /\ - (!i. i IN 1..dimindex (:N) ==> ~(i = j) ==> + (!i. i IN 1..dimindex (:N) ==> ~(i = j) ==> u$i = z$i /\ y'$i = y$i /\ w'$i = w$i) /\ - u$j = x$j /\ y'$j = x$j /\ w'$j = &0 ==> + u$j = x$j /\ y'$j = x$j /\ w'$j = &0 ==> m_cell_domain (x,u) y' w' /\ interval [x,u] SUBSET interval [x,z]`, REWRITE_TAC[m_cell_domain; SUBSET_INTERVAL; GSYM IN_NUMSEG] THEN REPEAT GEN_TAC THEN STRIP_TAC THEN CONJ_TAC THENL @@ -216,19 +222,19 @@ let gen_restrict_lemma n j left_flag = y'_tm = mk_vector_list (map (fun i -> List.nth (if i = j then b else ys) (i - 1)) (1--n)) and w'_tm = mk_vector_list (map (fun i -> if i = j then `&0` else List.nth ws (i - 1)) (1--n)) in - let th0 = (SPEC_ALL o ISPECL [j_tm; x_tm; z_tm; y_tm; w_tm; u_tm; y'_tm; w'_tm]) + let th0 = (SPEC_ALL o ISPECL [j_tm; x_tm; z_tm; y_tm; w_tm; u_tm; y'_tm; w'_tm]) (if left_flag then M_RESTRICT_LEFT_LEMMA else M_RESTRICT_RIGHT_LEMMA) in let th1 = REWRITE_RULE[dimindex_array.(n); IN_NUMSEG; gen_in_interval n; ARITH] th0 in let th2 = REWRITE_RULE (Array.to_list comp_thms_array.(n)) th1 in MY_RULE_FLOAT th2;; -let left_restrict_thms_array = Array.init (max_dim + 1) +let left_restrict_thms_array = Array.init (max_dim + 1) (fun n -> Array.init (n + 1) (fun j -> if j < 1 then TRUTH else gen_restrict_lemma n j true));; -let right_restrict_thms_array = Array.init (max_dim + 1) +let right_restrict_thms_array = Array.init (max_dim + 1) (fun n -> Array.init (n + 1) (fun j -> if j < 1 then TRUTH else gen_restrict_lemma n j false));; @@ -303,7 +309,7 @@ let m_taylor_cell_pass n pp m_taylor_th = failwith "m_taylor_cell_pass: hi < &0 <=> F" else (MY_PROVE_HYP upper_th o MY_PROVE_HYP hi_lt0_th o - INST[f_tm, f_var; domain_tm, domain_var; hi_tm, hi_var_real] o + INST[f_tm, f_var; domain_tm, domain_var; hi_tm, hi_var_real] o inst_first_type_var n_type_array.(n)) M_CELL_PASS_LEMMA';; let m_taylor_cell_list_pass n pp m_taylor_th = @@ -321,7 +327,7 @@ let m_taylor_cell_list_pass n pp m_taylor_th = failwith "m_taylor_cell_list_pass: hi < &0 <=> F" else (MY_PROVE_HYP upper_th o MY_PROVE_HYP hi_lt0_th o - INST[f_tm, f_var; domain_tm, domain_var; hi_tm, hi_var_real] o + INST[f_tm, f_var; domain_tm, domain_var; hi_tm, hi_var_real] o inst_first_type_var n_type_array.(n)) M_CELL_LIST_PASS1_LEMMA';; @@ -398,8 +404,8 @@ let m_cell_list_pass_subdomain domain2_tm pass_th = (******************************) -let GLUE_LEMMA = prove(`!j x z v u P Q. - (!i. 1 <= i /\ i <= dimindex (:N) ==> ~(i = j) ==> +let GLUE_LEMMA = prove(`!j x z v u P Q. + (!i. 1 <= i /\ i <= dimindex (:N) ==> ~(i = j) ==> u$i = x$i /\ v$i = z$i) ==> v$j = u$j ==> (!p. p IN interval [x,v] ==> P p) ==> @@ -433,11 +439,11 @@ let GLUE_LEMMA = prove(`!j x z v u P Q. USE_THEN "eq1" (new_rewrite [] []) THEN ASM_REWRITE_TAC[] THEN USE_THEN "ineq" (new_rewrite [] []) THEN ASM_REWRITE_TAC[]);; -let M_CELL_PASS_GLUE_LEMMA = prove(`!j x z v u f. - (!i. 1 <= i /\ i <= dimindex (:N) ==> ~(i = j) ==> +let M_CELL_PASS_GLUE_LEMMA = prove(`!j x z v u f. + (!i. 1 <= i /\ i <= dimindex (:N) ==> ~(i = j) ==> u$i = x$i /\ v$i = z$i) ==> v$j = u$j ==> - m_cell_pass f (x,v) ==> + m_cell_pass f (x,v) ==> m_cell_pass f (u,z) ==> m_cell_pass f (x,z:real^N)`, REPEAT GEN_TAC THEN REWRITE_TAC[m_cell_pass] THEN @@ -450,11 +456,11 @@ let ITLIST_DISJ_APPEND = prove(`!P l1 l2. ITLIST (\a r. P a \/ r) (APPEND l1 l2) GEN_TAC THEN LIST_INDUCT_TAC THEN LIST_INDUCT_TAC THEN REWRITE_TAC[APPEND; ITLIST; APPEND_NIL] THEN ASM_REWRITE_TAC[ITLIST; DISJ_ACI]);; -let M_CELL_LIST_PASS_GLUE_LEMMA = prove(`!j x z v u fs1 fs2. - (!i. 1 <= i /\ i <= dimindex (:N) ==> ~(i = j) ==> +let M_CELL_LIST_PASS_GLUE_LEMMA = prove(`!j x z v u fs1 fs2. + (!i. 1 <= i /\ i <= dimindex (:N) ==> ~(i = j) ==> u$i = x$i /\ v$i = z$i) ==> v$j = u$j ==> - m_cell_list_pass fs1 (x,v) ==> + m_cell_list_pass fs1 (x,v) ==> m_cell_list_pass fs2 (u,z) ==> m_cell_list_pass (APPEND fs1 fs2) (x,z:real^N)`, REPEAT GEN_TAC THEN REWRITE_TAC[m_cell_list_pass; ITLIST_DISJ_APPEND] THEN @@ -479,7 +485,7 @@ let gen_glue_lemma n j = gen_th M_CELL_PASS_GLUE_LEMMA, gen_th M_CELL_LIST_PASS_GLUE_LEMMA;; -let glue_thms_array = Array.init (max_dim + 1) +let glue_thms_array = Array.init (max_dim + 1) (fun n -> Array.init (n + 1) (fun j -> if j < 1 then TRUTH, TRUTH else gen_glue_lemma n j));; @@ -492,15 +498,15 @@ let CELL_LIST_PASS_ACC_INTRO = prove(`m_cell_list_pass fs1 domain <=> m_cell_lis let CELL_LIST_PASS_ACC_ELIM = SYM CELL_LIST_PASS_ACC_INTRO;; -let CELL_LIST_PASS_ACC_REV = prove(`m_cell_list_pass (APPEND acc (CONS h fs2)) domain +let CELL_LIST_PASS_ACC_REV = prove(`m_cell_list_pass (APPEND acc (CONS h fs2)) domain <=> m_cell_list_pass (APPEND (CONS h acc) fs2) domain`, REWRITE_TAC[m_cell_list_pass; ITLIST_DISJ_APPEND; ITLIST; DISJ_ACI]);; -let CELL_LIST_PASS_NIL1 = prove(`m_cell_list_pass (APPEND (APPEND [] fs2) acc) domain +let CELL_LIST_PASS_NIL1 = prove(`m_cell_list_pass (APPEND (APPEND [] fs2) acc) domain <=> m_cell_list_pass (APPEND fs2 acc) domain`, REWRITE_TAC[APPEND]);; -let CELL_LIST_PASS_NIL2 = prove(`m_cell_list_pass (APPEND (APPEND fs1 []) acc) domain +let CELL_LIST_PASS_NIL2 = prove(`m_cell_list_pass (APPEND (APPEND fs1 []) acc) domain <=> m_cell_list_pass (APPEND fs1 acc) domain`, REWRITE_TAC[APPEND_NIL]);; @@ -602,7 +608,7 @@ let DIFF2_DOMAIN_IMP_CONTINUOUS_ON = prove(`!(f:real^N->real) domain. diff2_doma -let M_CELL_INCREASING_PASS_LEMMA = prove(`!j x z u domain lo f. +let M_CELL_INCREASING_PASS_LEMMA = prove(`!j x z u domain lo f. interval [x,z] SUBSET interval [domain] ==> diff2c_domain domain f ==> (!i. i IN 1..dimindex (:N) ==> ~(i = j) ==> u$i = x$i) ==> @@ -669,7 +675,7 @@ let M_CELL_INCREASING_PASS_LEMMA = prove(`!j x z u domain lo f. -let M_CELL_DECREASING_PASS_LEMMA = prove(`!j x z u domain hi f. +let M_CELL_DECREASING_PASS_LEMMA = prove(`!j x z u domain hi f. interval [x,z] SUBSET interval [domain] ==> diff2c_domain domain f ==> (!i. i IN 1..dimindex (:N) ==> ~(i = j) ==> u$i = z$i) ==> @@ -735,7 +741,7 @@ let M_CELL_DECREASING_PASS_LEMMA = prove(`!j x z u domain hi f. -let M_CELL_CONVEX_PASS_LEMMA = prove(`!j x z u v lo f. +let M_CELL_CONVEX_PASS_LEMMA = prove(`!j x z u v lo f. diff2c_domain (x,z) f ==> (!i. i IN 1..dimindex (:N) ==> ~(i = j) ==> u$i = z$i /\ v$i = x$i) ==> u$j = x$j ==> v$j = z$j ==> @@ -837,7 +843,7 @@ let ITLIST_DISJ_EXISTS = prove(`!P l. ITLIST (\x r. P x \/ r) l F <=> ?x:A. MEM DISJ2_TAC THEN EXISTS_TAC `x:A` THEN ASM_REWRITE_TAC[] ]);; -let M_CELL_LIST_PASS_ALT = prove(`!fs domain. m_cell_list_pass fs domain <=> +let M_CELL_LIST_PASS_ALT = prove(`!fs domain. m_cell_list_pass fs domain <=> (!y:real^N. y IN interval [domain] ==> ?f. MEM f fs /\ f y < &0)`, REWRITE_TAC[m_cell_list_pass; ITLIST_DISJ_EXISTS]);; @@ -855,18 +861,18 @@ let ONE_POINT_INTERVAL = prove(`!x:real^N. interval [x,x] = {x}`, ARITH_TAC);; -let M_CELL_LIST_INCREASING_PASS_LEMMA = prove(`!j x z u domain fs. +let M_CELL_LIST_INCREASING_PASS_LEMMA = prove(`!j x z u domain fs. interval [x,z] SUBSET interval [domain] ==> (!i. i IN 1..dimindex (:N) ==> ~(i = j) ==> u$i = x$i) ==> u$j = z$j ==> - ALL (\f. diff2c_domain domain f /\ + ALL (\f. diff2c_domain domain f /\ (!y. y IN interval [domain] ==> &0 <= partial j f y)) fs ==> m_cell_list_pass fs (u,z) ==> m_cell_list_pass fs (x,z:real^N)`, - REWRITE_TAC[SUBSET; M_CELL_LIST_PASS_ALT] THEN REPEAT STRIP_TAC THEN + REWRITE_TAC[SUBSET; M_CELL_LIST_PASS_ALT] THEN REPEAT STRIP_TAC THEN ABBREV_TAC `(xx:real^N) = lambda i. if i = j then (x:real^N)$i else (y:real^N)$i` THEN ABBREV_TAC `(zz:real^N) = lambda i. if i = j then (z:real^N)$i else (y:real^N)$i` THEN - + SUBGOAL_THEN `interval [xx:real^N,zz] SUBSET interval [domain]` ASSUME_TAC THENL [ MATCH_MP_TAC SUBSET_TRANS THEN EXISTS_TAC `interval [x:real^N,z]` THEN ASM_REWRITE_TAC[SUBSET] THEN REWRITE_TAC[IN_INTERVAL] THEN X_GEN_TAC `t:real^N` THEN DISCH_THEN (LABEL_TAC "in1") THEN @@ -879,7 +885,7 @@ let M_CELL_LIST_INCREASING_PASS_LEMMA = prove(`!j x z u domain fs. ARITH_TAC; ALL_TAC ] THEN - + SUBGOAL_THEN `y IN interval [xx:real^N,zz]` ASSUME_TAC THENL [ REWRITE_TAC[IN_INTERVAL] THEN GEN_TAC THEN DISCH_TAC THEN EXPAND_TAC "xx" THEN EXPAND_TAC "zz" THEN ASM_SIMP_TAC[LAMBDA_BETA] THEN @@ -888,7 +894,7 @@ let M_CELL_LIST_INCREASING_PASS_LEMMA = prove(`!j x z u domain fs. DISCH_THEN (MP_TAC o SPEC `i:num`) THEN ASM_REWRITE_TAC[]; ALL_TAC ] THEN - + FIRST_X_ASSUM (MP_TAC o SPEC `zz:real^N`) THEN ANTS_TAC THENL [ REWRITE_TAC[IN_INTERVAL] THEN GEN_TAC THEN DISCH_TAC THEN EXPAND_TAC "zz" THEN ASM_SIMP_TAC[LAMBDA_BETA] THEN @@ -914,15 +920,15 @@ let M_CELL_LIST_INCREASING_PASS_LEMMA = prove(`!j x z u domain fs. -let M_CELL_LIST_DECREASING_PASS_LEMMA = prove(`!j x z u domain fs. +let M_CELL_LIST_DECREASING_PASS_LEMMA = prove(`!j x z u domain fs. interval [x,z] SUBSET interval [domain] ==> (!i. i IN 1..dimindex (:N) ==> ~(i = j) ==> u$i = z$i) ==> u$j = x$j ==> - ALL (\f. diff2c_domain domain f /\ + ALL (\f. diff2c_domain domain f /\ (!y. y IN interval [domain] ==> partial j f y <= &0)) fs ==> m_cell_list_pass fs (x,u) ==> m_cell_list_pass fs (x,z:real^N)`, - REWRITE_TAC[SUBSET; M_CELL_LIST_PASS_ALT] THEN REPEAT STRIP_TAC THEN + REWRITE_TAC[SUBSET; M_CELL_LIST_PASS_ALT] THEN REPEAT STRIP_TAC THEN ABBREV_TAC `(xx:real^N) = lambda i. if i = j then (x:real^N)$i else (y:real^N)$i` THEN ABBREV_TAC `(zz:real^N) = lambda i. if i = j then (z:real^N)$i else (y:real^N)$i` THEN SUBGOAL_THEN `interval [xx:real^N,zz] SUBSET interval [domain]` ASSUME_TAC THENL [ @@ -967,7 +973,7 @@ let M_CELL_LIST_DECREASING_PASS_LEMMA = prove(`!j x z u domain fs. ] THEN ASM_SIMP_TAC[ONE_POINT_INTERVAL; IN_SING]);; - + (*********************) @@ -994,11 +1000,11 @@ let gen_mono_lemma0 th = let domain_var = mk_var ("domain", type_of domain_tm) in (UNDISCH_ALL o REWRITE_RULE[SUBSET_REFL] o DISCH_ALL o INST[domain_tm, domain_var]) th;; -let incr_gen_thms_array = Array.init (max_dim + 1) +let incr_gen_thms_array = Array.init (max_dim + 1) (fun n -> Array.init (n + 1) (fun j -> if j < 1 then TRUTH else gen_increasing_lemma n j));; -let incr_thms_array = Array.init (max_dim + 1) +let incr_thms_array = Array.init (max_dim + 1) (fun n -> Array.init (n + 1) (fun j -> if j < 1 then TRUTH else gen_mono_lemma0 incr_gen_thms_array.(n).(j)));; @@ -1019,11 +1025,11 @@ let gen_decreasing_lemma n j = (UNDISCH_ALL o ONCE_REWRITE_RULE[GSYM ZERO_EQ_ZERO_CONST] o DISCH (last (hyp th3))) th3;; -let decr_gen_thms_array = Array.init (max_dim + 1) +let decr_gen_thms_array = Array.init (max_dim + 1) (fun n -> Array.init (n + 1) (fun j -> if j < 1 then TRUTH else gen_decreasing_lemma n j));; -let decr_thms_array = Array.init (max_dim + 1) +let decr_thms_array = Array.init (max_dim + 1) (fun n -> Array.init (n + 1) (fun j -> if j < 1 then TRUTH else gen_mono_lemma0 decr_gen_thms_array.(n).(j)));; @@ -1045,7 +1051,7 @@ let gen_convex_max_lemma n j = (UNDISCH_ALL o ONCE_REWRITE_RULE[GSYM ZERO_EQ_ZERO_CONST] o DISCH (last (hyp th3))) th3;; -let convex_thms_array = Array.init (max_dim + 1) +let convex_thms_array = Array.init (max_dim + 1) (fun n -> Array.init (n + 1) (fun j -> if j < 1 then TRUTH else gen_convex_max_lemma n j));; @@ -1067,8 +1073,8 @@ let m_glue_cells n j pass_th1 pass_th2 = z_vars = map (fun i -> z_vars_array.(i)) (1--n) and f_var = mk_var ("f", type_of f_tm) and t_tm = List.nth x2s (j - 1) in - - let th0 = (INST[t_tm, t_var_real; f_tm, f_var] o + + let th0 = (INST[t_tm, t_var_real; f_tm, f_var] o INST (zip z2s z_vars) o INST (zip x1s x_vars)) (fst glue_thms_array.(n).(j)) in (MY_PROVE_HYP pass_th1 o MY_PROVE_HYP pass_th2) th0;; @@ -1088,8 +1094,8 @@ let m_glue_cells_list n j pass_th1 pass_th2 = fs1_var = mk_var ("fs1", type_of fs1_tm) and fs2_var = mk_var ("fs2", type_of fs2_tm) and t_tm = List.nth x2s (j - 1) in - - let th0 = (INST[t_tm, t_var_real; fs1_tm, fs1_var; fs2_tm, fs2_var] o + + let th0 = (INST[t_tm, t_var_real; fs1_tm, fs1_var; fs2_tm, fs2_var] o INST (zip z2s z_vars) o INST (zip x1s x_vars)) (snd glue_thms_array.(n).(j)) in (MY_PROVE_HYP pass_th1 o MY_PROVE_HYP pass_th2) th0;; @@ -1112,13 +1118,13 @@ let m_mono_pass_gen n j decr_flag diff2_th partial_mono_th sub_th pass_th = bound_var = mk_var ((if decr_flag then "hi" else "lo"), real_ty) in let le_th0 = (if decr_flag then float_le0 else float_ge0) bound_tm in - let le_th = try EQT_ELIM le_th0 with Failure _ -> + let le_th = try EQT_ELIM le_th0 with Failure _ -> failwith (sprintf "m_mono_pass_gen: j = %d, th = %s" j (string_of_thm le_th0)) in let th0 = (INST[f_tm, f_var; bound_tm, bound_var; domain, domain_var] o - INST (zip xs x_vars) o INST (zip zs z_vars)) + INST (zip xs x_vars) o INST (zip zs z_vars)) (if decr_flag then decr_gen_thms_array.(n).(j) else incr_gen_thms_array.(n).(j)) in - (MY_PROVE_HYP le_th o MY_PROVE_HYP pass_th o MY_PROVE_HYP diff2_th o + (MY_PROVE_HYP le_th o MY_PROVE_HYP pass_th o MY_PROVE_HYP diff2_th o MY_PROVE_HYP sub_th o MY_PROVE_HYP partial_mono_th) th0;; @@ -1140,9 +1146,9 @@ let m_incr_pass n pp j m_taylor_th pass_th0 = let x_vars = map (fun i -> x_vars_array.(i)) (1--n) and z_vars = map (fun i -> z_vars_array.(i)) (1--n) and f_var = mk_var ("f", type_of f_tm) in - let th0 = (INST[f_tm, f_var; lo_tm, lo_var_real] o + let th0 = (INST[f_tm, f_var; lo_tm, lo_var_real] o INST (zip zs z_vars) o INST (zip xs x_vars)) incr_thms_array.(n).(j) in - (MY_PROVE_HYP lo_ge0_th o MY_PROVE_HYP pass_th0 o + (MY_PROVE_HYP lo_ge0_th o MY_PROVE_HYP pass_th0 o MY_PROVE_HYP diff2_th o MY_PROVE_HYP partial_bound) th0;; @@ -1163,9 +1169,9 @@ let m_decr_pass n pp j m_taylor_th pass_th0 = let x_vars = map (fun i -> x_vars_array.(i)) (1--n) and z_vars = map (fun i -> z_vars_array.(i)) (1--n) and f_var = mk_var ("f", type_of f_tm) in - let th0 = (INST[f_tm, f_var; hi_tm, hi_var_real] o + let th0 = (INST[f_tm, f_var; hi_tm, hi_var_real] o INST (zip zs z_vars) o INST (zip xs x_vars)) decr_thms_array.(n).(j) in - (MY_PROVE_HYP hi_le0_th o MY_PROVE_HYP pass_th0 o + (MY_PROVE_HYP hi_le0_th o MY_PROVE_HYP pass_th0 o MY_PROVE_HYP diff2_th o MY_PROVE_HYP partial_bound) th0;; (*************************) @@ -1186,7 +1192,7 @@ let m_convex_pass n j diff2_th partial2_bound_th pass1_th pass2_th = f_var = mk_var ("f", type_of f_tm) in let le_th0 = float_ge0 bound_tm in - let le_th = + let le_th = try EQT_ELIM le_th0 with Failure _ -> failwith ("m_convex_pass: "^string_of_thm le_th0) in let th0 = (INST[f_tm, f_var; bound_tm, lo_var_real] o @@ -1201,7 +1207,7 @@ let m_convex_pass n j diff2_th partial2_bound_th pass1_th pass2_th = (* split_domain *) -let split_domain n pp j domain_th = +let split_domain n pp j domain_th = let domain_tm, y_tm, _ = dest_m_cell_domain (concl domain_th) in let x_tm, z_tm = dest_pair domain_tm in let xs = dest_vector x_tm and @@ -1231,7 +1237,7 @@ let restrict_domain n j left_flag domain_th = (if left_flag then left_restrict_thms_array.(n).(j) else right_restrict_thms_array.(n).(j)) in let ths = CONJUNCTS (MY_PROVE_HYP domain_th th0) in hd ths, hd (tl ths);; - + @@ -1248,7 +1254,7 @@ let m_verify_raw (report_start, total_size) n pp fs certificate domain_th0 th_li let rec apply_trans sub_ths th0 acc = match sub_ths with | [] -> rev acc - | th :: ths -> + | th :: ths -> let th' = eval_subset_trans th th0 in apply_trans ths th' (th' :: acc) in @@ -1265,7 +1271,7 @@ let m_verify_raw (report_start, total_size) n pp fs certificate domain_th0 th_li let xx, zz = dest_pair domain in let df0_flags = itlist (fun m b -> m.df0_flag && b) mono true in let _ = !info_print_level < 2 || (report (sprintf "df0_flags = %b" df0_flags); true) in - let taylor_th, diff2_th = + let taylor_th, diff2_th = if df0_flags then TRUTH, fs.diff2_f xx zz else @@ -1306,13 +1312,13 @@ let m_verify_raw (report_start, total_size) n pp fs certificate domain_th0 th_li match certificate with | Result_mono (mono, r1) -> let _ = !info_print_level < 2 || - (let mono_strs = - map (fun m -> sprintf "%s%d (%b)" (if m.decr_flag then "-" else "") + (let mono_strs = + map (fun m -> sprintf "%s%d (%b)" (if m.decr_flag then "-" else "") m.variable m.df0_flag) mono in report (sprintf "Mono: [%s]" (String.concat ";" mono_strs)); true) in verify_mono mono domain_th r1 - | Result_pass (_, f0_flag) -> + | Result_pass (_, f0_flag) -> let _ = k := !k + 1 in let _ = !info_print_level < 2 || (report (sprintf "Verifying: %d/%d (f0_flag = %b)" !k r_size f0_flag); true) in @@ -1325,8 +1331,8 @@ let m_verify_raw (report_start, total_size) n pp fs certificate domain_th0 th_li m_taylor_cell_pass0 n (fs.f pp xx zz) else let taylor_th = fs.taylor pp pp domain_th in - m_taylor_cell_pass n pp taylor_th - + m_taylor_cell_pass n pp taylor_th + | Result_glue (i, convex_flag, r1, r2) -> let domain1_th, domain2_th = if convex_flag then @@ -1350,15 +1356,15 @@ let m_verify_raw (report_start, total_size) n pp fs certificate domain_th0 th_li | Result_pass_ref i -> let _ = !info_print_level < 2 || (report (sprintf "Ref: %d" i); true) in - if i > 0 then + if i > 0 then List.nth th_list (i - 1) else let domain, _, _ = (dest_m_cell_domain o concl) domain_th in let pass_th = List.nth th_list (-i - 1) in m_cell_pass_subdomain domain pass_th - + | _ -> failwith "False result" in - + rec_verify domain_th0 certificate;; @@ -1368,11 +1374,11 @@ let m_verify_raw (report_start, total_size) n pp fs certificate domain_th0 th_li let m_verify_raw0 n pp fs certificate xx zz = m_verify_raw (0, 0) n pp fs certificate (mk_m_center_domain n pp xx zz) [];; - + let m_verify_list n pp fs certificate_list xx zz = let domain_hash = Hashtbl.create (length certificate_list * 10) in - let mem, find, add = Hashtbl.mem domain_hash, + let mem, find, add = Hashtbl.mem domain_hash, Hashtbl.find domain_hash, Hashtbl.add domain_hash in let get_m_cell_domain n pp domain0 path = @@ -1381,7 +1387,7 @@ let m_verify_list n pp fs certificate_list xx zz = | [] -> domain_th | (s, j) :: ps -> let hash' = hash^s^(string_of_int j) in - if mem hash' then + if mem hash' then get_rec (find hash') ps hash' else if s = "l" || s = "r" then @@ -1433,7 +1439,7 @@ let m_p_verify_raw (report_start, total_size) n p_split fs certificate domain_th let rec apply_trans sub_ths th0 acc = match sub_ths with | [] -> rev acc - | th :: ths -> + | th :: ths -> let th' = eval_subset_trans th th0 in apply_trans ths th' (th' :: acc) in @@ -1450,7 +1456,7 @@ let m_p_verify_raw (report_start, total_size) n p_split fs certificate domain_th let xx, zz = dest_pair domain in let df0_flags = itlist (fun m b -> m.df0_flag && b) mono true in let _ = !info_print_level < 2 || (report (sprintf "df0_flags = %b" df0_flags); true) in - let taylor_th, diff2_th = + let taylor_th, diff2_th = if df0_flags then TRUTH, fs.diff2_f xx zz else @@ -1486,15 +1492,15 @@ let m_p_verify_raw (report_start, total_size) n p_split fs certificate domain_th match certificate with | P_result_mono (p_stat, mono, r1) -> let _ = !info_print_level < 2 || - (let mono_strs = - map (fun m -> sprintf "%s%d (%b)" (if m.decr_flag then "-" else "") + (let mono_strs = + map (fun m -> sprintf "%s%d (%b)" (if m.decr_flag then "-" else "") m.variable m.df0_flag) mono in report (sprintf "Mono: [%s]" (String.concat ";" mono_strs)); true) in verify_mono p_stat mono domain_th r1 | P_result_pass (p_stat, _, f0_flag) -> let _ = k := !k + 1; kk := !kk + 1 in - let _ = !info_print_level < 2 || + let _ = !info_print_level < 2 || (report (sprintf "Verifying: %d/%d (f0_flag = %b)" !k r_size f0_flag); true) in let _ = !info_print_level <> 1 || (let r = int_of_float (float_of_int !kk /. r_size2 *. 100.0) in @@ -1505,8 +1511,8 @@ let m_p_verify_raw (report_start, total_size) n p_split fs certificate domain_th m_taylor_cell_pass0 n (fs.f p_stat.pp xx zz) else let taylor_th = fs.taylor p_stat.pp p_stat.pp domain_th in - m_taylor_cell_pass n p_stat.pp taylor_th - + m_taylor_cell_pass n p_stat.pp taylor_th + | P_result_glue (p_stat, i, convex_flag, r1, r2) -> let domain1_th, domain2_th = if convex_flag then @@ -1530,13 +1536,13 @@ let m_p_verify_raw (report_start, total_size) n p_split fs certificate domain_th | P_result_ref i -> let _ = !info_print_level < 2 || (report (sprintf "Ref: %d" i); true) in - if i > 0 then + if i > 0 then List.nth th_list (i - 1) else let domain, _, _ = (dest_m_cell_domain o concl) domain_th in let pass_th = List.nth th_list (-i - 1) in m_cell_pass_subdomain domain pass_th in - + rec_verify domain_th0 certificate;; (*****************) @@ -1544,11 +1550,11 @@ let m_p_verify_raw (report_start, total_size) n p_split fs certificate domain_th let m_p_verify_raw0 n p_split fs certificate xx zz = m_p_verify_raw (0, 0) n p_split fs certificate (mk_m_center_domain n p_split xx zz) [];; - + let m_p_verify_list n p_split fs certificate_list xx zz = let domain_hash = Hashtbl.create (length certificate_list * 10) in - let mem, find, add = Hashtbl.mem domain_hash, + let mem, find, add = Hashtbl.mem domain_hash, Hashtbl.find domain_hash, Hashtbl.add domain_hash in let get_m_cell_domain n pp domain0 path = @@ -1557,7 +1563,7 @@ let m_p_verify_list n p_split fs certificate_list xx zz = | [] -> domain_th | (s, j) :: ps -> let hash' = hash^s^(string_of_int j) in - if mem hash' then + if mem hash' then get_rec (find hash') ps hash' else if s = "l" || s = "r" then @@ -1581,7 +1587,7 @@ let m_p_verify_list n p_split fs certificate_list xx zz = let k = ref 0 in let kk = ref 0 in let total_size = end_itlist (+) (map (p_result_size o snd) certificate_list) in - + let rec rec_verify certificate_list th_list = match certificate_list with | [] -> last th_list @@ -1610,7 +1616,7 @@ let m_verify_disj_raw (report_start, total_size) n pp fs_list certificate domain let rec apply_trans sub_ths th0 acc = match sub_ths with | [] -> rev acc - | th :: ths -> + | th :: ths -> let th' = eval_subset_trans th th0 in apply_trans ths th' (th' :: acc) in @@ -1627,7 +1633,7 @@ let m_verify_disj_raw (report_start, total_size) n pp fs_list certificate domain let xx, zz = dest_pair domain in let df0_flags = itlist (fun m b -> m.df0_flag && b) mono true in let _ = !info_print_level < 2 || (report (sprintf "df0_flags = %b" df0_flags); true) in - let taylor_th, diff2_th = + let taylor_th, diff2_th = if df0_flags then TRUTH, fs.diff2_f xx zz else @@ -1670,14 +1676,14 @@ let m_verify_disj_raw (report_start, total_size) n pp fs_list certificate domain failwith "Mono: not implemented" (* let _ = !info_print_level < 2 || - (let mono_strs = - map (fun m -> sprintf "%s%d (%b)" (if m.decr_flag then "-" else "") + (let mono_strs = + map (fun m -> sprintf "%s%d (%b)" (if m.decr_flag then "-" else "") m.variable m.df0_flag) mono in report (sprintf "Mono: [%s]" (String.concat ";" mono_strs)); true) in verify_mono mono domain_th r1 *) - | Result_pass (j, f0_flag) -> + | Result_pass (j, f0_flag) -> let _ = k := !k + 1 in let _ = !info_print_level < 2 || (report (sprintf "Verifying: %d/%d (f0_flag = %b)" !k r_size f0_flag); true) in @@ -1692,8 +1698,8 @@ let m_verify_disj_raw (report_start, total_size) n pp fs_list certificate domain m_taylor_cell_list_pass0 n (fs.f pp xx zz) else let taylor_th = fs.taylor pp pp domain_th in - m_taylor_cell_list_pass n pp taylor_th - + m_taylor_cell_list_pass n pp taylor_th + | Result_glue (i, convex_flag, r1, r2) -> let domain1_th, domain2_th = if convex_flag then @@ -1721,15 +1727,15 @@ let m_verify_disj_raw (report_start, total_size) n pp fs_list certificate domain | Result_pass_ref i -> let _ = !info_print_level < 2 || (report (sprintf "Ref: %d" i); true) in - if i > 0 then + if i > 0 then List.nth th_list (i - 1) else let domain, _, _ = (dest_m_cell_domain o concl) domain_th in let pass_th = List.nth th_list (-i - 1) in m_cell_list_pass_subdomain domain pass_th - + | _ -> failwith "False result" in - + rec_verify domain_th0 certificate;; @@ -1755,9 +1761,9 @@ let m_p_verify_disj_raw (report_start, total_size) n p_split fs_list certificate | P_result_mono (p_stat, mono, r1) -> failwith "Mono: not implemented" - | P_result_pass (p_stat, j, f0_flag) -> + | P_result_pass (p_stat, j, f0_flag) -> let _ = k := !k + 1; kk := !kk + 1 in - let _ = !info_print_level < 2 || + let _ = !info_print_level < 2 || (report (sprintf "Verifying: %d/%d (f0_flag = %b)" !k r_size f0_flag); true) in let _ = !info_print_level <> 1 || (let r = int_of_float (float_of_int !kk /. r_size2 *. 100.0) in @@ -1769,8 +1775,8 @@ let m_p_verify_disj_raw (report_start, total_size) n p_split fs_list certificate m_taylor_cell_list_pass0 n (fs.f p_stat.pp xx zz) else let taylor_th = fs.taylor p_stat.pp p_stat.pp domain_th in - m_taylor_cell_list_pass n p_stat.pp taylor_th - + m_taylor_cell_list_pass n p_stat.pp taylor_th + | P_result_glue (p_stat, i, convex_flag, r1, r2) -> let domain1_th, domain2_th = if convex_flag then @@ -1789,14 +1795,14 @@ let m_p_verify_disj_raw (report_start, total_size) n p_split fs_list certificate | P_result_ref i -> let _ = !info_print_level < 2 || (report (sprintf "Ref: %d" i); true) in - if i > 0 then + if i > 0 then List.nth th_list (i - 1) else let domain, _, _ = (dest_m_cell_domain o concl) domain_th in let pass_th = List.nth th_list (-i - 1) in m_cell_list_pass_subdomain domain pass_th - - | _ -> failwith "False result" + + | _ -> failwith "False result" in rec_verify domain_th0 certificate;; diff --git a/Formal_ineqs/verifier/m_verifier_build.hl b/Formal_ineqs/verifier/m_verifier_build.hl index 8f73c2d5..e554e069 100644 --- a/Formal_ineqs/verifier/m_verifier_build.hl +++ b/Formal_ineqs/verifier/m_verifier_build.hl @@ -1,15 +1,21 @@ -(* =========================================================== *) -(* Auxiliary formal verification functions *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -needs "arith/eval_interval.hl";; -needs "arith/more_float.hl";; -needs "taylor/m_taylor.hl";; -needs "informal/informal_verifier.hl";; -needs "verifier_options.hl";; -needs "misc/misc_vars.hl";; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Auxiliary formal verification functions *) +(* -------------------------------------------------------------------------- *) + +needs "Formal_ineqs/arith/eval_interval.hl";; +needs "Formal_ineqs/arith/more_float.hl";; +needs "Formal_ineqs/taylor/m_taylor.hl";; +needs "Formal_ineqs/informal/informal_verifier.hl";; +needs "Formal_ineqs/verifier_options.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; module M_verifier_build = struct @@ -40,7 +46,7 @@ type verification_funs = };; -(**********************************) +(**********************************) (* mk_verification_functions *) let mk_verification_functions_poly pp0 poly_tm = @@ -49,16 +55,16 @@ let mk_verification_functions_poly pp0 poly_tm = let n = get_dim x_tm in let _ = !info_print_level <= 1 || (report0 (sprintf "Computing partial derivatives (%d)..." n); true) in - let partials = map (fun i -> + let partials = map (fun i -> let _ = !info_print_level <= 1 || (report0 (sprintf " %d" i); true) in gen_partial_poly i new_f) (1--n) in - let get_partial i eq_th = + let get_partial i eq_th = let partial_i = gen_partial_poly i (rand (concl eq_th)) in let pi = (rator o lhand o concl) partial_i in REWRITE_RULE[GSYM partial2] (TRANS (AP_TERM pi eq_th) partial_i) in - let partials2 = map (fun j -> + let partials2 = map (fun j -> let th = List.nth partials (j - 1) in - map (fun i -> + map (fun i -> let _ = !info_print_level <= 1 || (report0 (sprintf " %d,%d" j i); true) in get_partial i th) (1--j)) (1--n) in @@ -70,12 +76,12 @@ let mk_verification_functions_poly pp0 poly_tm = let second_th = gen_second_bounded_poly_thm new_f partials2 in let replace_numeral i th = - let num_eq = (REWRITE_RULE[Arith_num.NUM_THM] o Arith_nat.NUMERAL_TO_NUM_CONV) + let num_eq = (REWRITE_RULE[Arith_num.NUM_THM] o Arith_nat.NUMERAL_TO_NUM_CONV) (mk_small_numeral i) in GEN_REWRITE_RULE (LAND_CONV o RATOR_CONV o DEPTH_CONV) [num_eq] th in let eval0 = mk_eval_function pp0 new_f in - let eval1 = map (fun i -> + let eval1 = map (fun i -> let d_th = List.nth partials (i - 1) in let eq_th = replace_numeral i d_th in mk_eval_function_eq pp0 eq_th) (1--n) in @@ -98,5 +104,5 @@ let mk_verification_functions_poly pp0 poly_tm = - + end;; diff --git a/Formal_ineqs/verifier/m_verifier_main.hl b/Formal_ineqs/verifier/m_verifier_main.hl index bbd71bc6..c3c12b42 100644 --- a/Formal_ineqs/verifier/m_verifier_main.hl +++ b/Formal_ineqs/verifier/m_verifier_main.hl @@ -1,17 +1,23 @@ -(* =========================================================== *) -(* Main formal verification functions *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) -needs "verifier/m_verifier.hl";; -needs "verifier/m_verifier_build.hl";; -needs "taylor/m_taylor_arith2.hl";; -needs "misc/misc_vars.hl";; +(* -------------------------------------------------------------------------- *) +(* Main formal verification functions *) +(* -------------------------------------------------------------------------- *) -needs "verifier/certificate.hl";; -needs "informal/informal_search.hl";; -needs "informal/informal_verifier.hl";; +needs "Formal_ineqs/verifier/m_verifier.hl";; +needs "Formal_ineqs/verifier/m_verifier_build.hl";; +needs "Formal_ineqs/taylor/m_taylor_arith2.hl";; +needs "Formal_ineqs/misc/misc_vars.hl";; + +needs "Formal_ineqs/verifier/certificate.hl";; +needs "Formal_ineqs/informal/informal_search.hl";; +needs "Formal_ineqs/informal/informal_verifier.hl";; module M_verifier_main = struct @@ -115,7 +121,7 @@ let test_expression bin_ops unary_ops = let lhs, rhs = dest_comb tm in let c, _ = dest_const lhs in if mem c unary_ops then test rhs else false - with Failure _ -> false + with Failure _ -> false in fun tm -> frees tm = [] || @@ -139,7 +145,7 @@ let is_poly = let rec mk_funs = (* add *) let mk_add n (f1, ti1) (f2, ti2) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x and b = f2 p1 p2 x in eval_m_taylor_add2 n p1 p2 a b), @@ -149,7 +155,7 @@ let rec mk_funs = Informal_taylor.eval_m_taylor_add p1 p2 a b) in (* sub *) let mk_sub n (f1, ti1) (f2, ti2) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x and b = f2 p1 p2 x in eval_m_taylor_sub2 n p1 p2 a b), @@ -159,7 +165,7 @@ let rec mk_funs = Informal_taylor.eval_m_taylor_sub p1 p2 a b) in (* mul *) let mk_mul n (f1, ti1) (f2, ti2) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x and b = f2 p1 p2 x in eval_m_taylor_mul2 n p1 p2 a b), @@ -179,7 +185,7 @@ let rec mk_funs = ieval_pow p1 p2 a) in (* neg *) let mk_neg n (f1, ti1) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x in eval_m_taylor_neg2 n a), (fun p1 p2 x -> @@ -187,7 +193,7 @@ let rec mk_funs = Informal_taylor.eval_m_taylor_neg a) in (* abs *) let mk_real_abs n (f1, ti1) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x in eval_m_taylor_abs2 n p2 a), (fun p1 p2 x -> @@ -195,7 +201,7 @@ let rec mk_funs = Informal_taylor.eval_m_taylor_abs p2 a) in (* sqrt *) let mk_sqrt n (f1, ti1) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x in eval_m_taylor_sqrt2 n p1 p2 a), (fun p1 p2 x -> @@ -203,7 +209,7 @@ let rec mk_funs = Informal_taylor.eval_m_taylor_sqrt p1 p2 a) in (* inv *) let mk_inv n (f1, ti1) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x in eval_m_taylor_inv2 n p1 p2 a), (fun p1 p2 x -> @@ -211,7 +217,7 @@ let rec mk_funs = Informal_taylor.eval_m_taylor_inv p1 p2 a) in (* exp *) let mk_exp n (f1, ti1) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x in eval_m_taylor_exp2 n p1 p2 a), (fun p1 p2 x -> @@ -219,7 +225,7 @@ let rec mk_funs = Informal_taylor.eval_m_taylor_exp p1 p2 a) in (* log *) let mk_log n (f1, ti1) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x in eval_m_taylor_log2 n p1 p2 a), (fun p1 p2 x -> @@ -227,7 +233,7 @@ let rec mk_funs = Informal_taylor.eval_m_taylor_log p1 p2 a) in (* atn *) let mk_atn n (f1, ti1) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x in eval_m_taylor_atn2 n p1 p2 a), (fun p1 p2 x -> @@ -235,7 +241,7 @@ let rec mk_funs = Informal_taylor.eval_m_taylor_atn p1 p2 a) in (* cos *) let mk_cos n (f1, ti1) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x in eval_m_taylor_cos2 n p1 p2 a), (fun p1 p2 x -> @@ -243,7 +249,7 @@ let rec mk_funs = Informal_taylor.eval_m_taylor_cos p1 p2 a) in (* sin *) let mk_sin n (f1, ti1) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x in eval_m_taylor_sin2 n p1 p2 a), (fun p1 p2 x -> @@ -251,7 +257,7 @@ let rec mk_funs = Informal_taylor.eval_m_taylor_sin p1 p2 a) in (* acs *) let mk_acs n (f1, ti1) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x in eval_m_taylor_acs2 n p1 p2 a), (fun p1 p2 x -> @@ -259,7 +265,7 @@ let rec mk_funs = Informal_taylor.eval_m_taylor_acs p1 p2 a) in (* asn *) let mk_asn n (f1, ti1) = - (fun p1 p2 x -> + (fun p1 p2 x -> let a = f1 p1 p2 x in eval_m_taylor_asn2 n p1 p2 a), (fun p1 p2 x -> @@ -274,8 +280,8 @@ let rec mk_funs = let a = ti1 p1 p2 x in Informal_taylor.eval_m_taylor_matan p1 p2 a) in - (* binary operations *) - let bin_ops = + (* binary operations *) + let bin_ops = ["real_add", mk_add; "real_sub", mk_sub; "real_mul", mk_mul] in @@ -323,7 +329,7 @@ let rec mk_funs = try mk_bin n pp x_var body_tm with Failure _ -> mk_unary n pp x_var body_tm;; - + (* Prepares verification functions *) (* fun_tm must be in the form `\x. f x` *) let mk_verification_functions = @@ -341,15 +347,15 @@ let mk_verification_functions = let eval0 = mk_eval_function pp fun_tm in let eval0_informal = Informal_taylor.mk_eval_function pp fun_tm in let _ = params := { - !params with - raw_intervals1 = false; + !params with + raw_intervals1 = false; convex_flag = false - } in + } in { - taylor = eval_taylor; - f = eval0; - df = dummy_df; - ddf = dummy_ddf; + taylor = eval_taylor; + f = eval0; + df = dummy_df; + ddf = dummy_ddf; diff2_f = dummy_diff2 }, { @@ -359,7 +365,7 @@ let mk_verification_functions = Informal_verifier.ddf = dummy_ddf };; - + (********************************) let convert_to_float_list pp lo_flag list_tm = @@ -375,7 +381,7 @@ let convert_to_float_list pp lo_flag list_tm = let mk_float_domain pp (xx_tm, zz_tm) = let xx_list = dest_list xx_tm and zz_list = dest_list zz_tm in - let n = length xx_list in + let n = length xx_list in let get_intervals tms = let i_funs = map build_interval_fun tms in map (fun f -> eval_interval_fun pp f [] []) i_funs in @@ -390,9 +396,9 @@ let mk_float_domain pp (xx_tm, zz_tm) = c_vars = mk_real_vars n "c" and d_vars = mk_real_vars n "d" in let th0 = (INST (zip xx_list c_vars) o INST (zip zz_list d_vars) o - INST (zip a_vals a_vars) o INST (zip b_vals b_vars)) + INST (zip a_vals a_vars) o INST (zip b_vals b_vars)) subset_interval_thms_array.(n) in - itlist MY_PROVE_HYP (xx_ineqs @ zz_ineqs) th0, + itlist MY_PROVE_HYP (xx_ineqs @ zz_ineqs) th0, (mk_list (a_vals, real_ty), mk_list (b_vals, real_ty));; @@ -405,18 +411,18 @@ let mk_standard_ineq = lemma_le = REAL_ARITH `a <= b <=> a - b <= &0` and lemma_imp = TAUT `(P <=> Q) ==> (Q ==> P)` in fun thms tm -> - let th0 = (REWRITE_CONV([real_gt; real_ge; real_div] @ thms) + let th0 = (REWRITE_CONV([real_gt; real_ge; real_div] @ thms) THENC DEPTH_CONV let_CONV) tm in let rhs = rand (concl th0) in - let th1 = (ONCE_REWRITE_CONV[lemma_le] THENC - ONCE_REWRITE_CONV[lemma_lt] THENC + let th1 = (ONCE_REWRITE_CONV[lemma_le] THENC + ONCE_REWRITE_CONV[lemma_lt] THENC PURE_REWRITE_CONV[REAL_NEG_0; REAL_SUB_RZERO; REAL_SUB_LZERO]) rhs in let th2 = TRANS th0 th1 in let ineqs = striplist dest_disj (rand (concl th2)) in let le_flags = map (can (dest_binop le_op_real)) ineqs in - if exists I le_flags then - let q_le_tms, q_lt_tms = unzip - (map2 (fun i b -> + if exists I le_flags then + let q_le_tms, q_lt_tms = unzip + (map2 (fun i b -> let var = mk_var ("Q" ^ string_of_int i, real_ty) in mk_binop (if b then le_op_real else lt_op_real) var `&0`, mk_binop lt_op_real var `&0`) @@ -425,7 +431,7 @@ let mk_standard_ineq = qs_lt = end_itlist (curry mk_disj) q_lt_tms and p_tm = `P:bool` in (* |- (P <=> Q1 <= &0 \/ ...) ==> (Q1 < &0 \/ ... ==> P) *) - let q_th = MESON[REAL_LT_IMP_LE] + let q_th = MESON[REAL_LT_IMP_LE] (mk_imp (mk_eq (p_tm, qs_le), mk_imp (qs_lt, p_tm))) in (* |- Q1 < &0 \/ ... ==> P *) MATCH_MP q_th th2 @@ -441,10 +447,10 @@ let expr_to_vector_fun = let x_var = mk_var ("x", n_vector_type_array.(if n = 0 then 1 else n)) in let x_tm = mk_icomb (comp_op, x_var) in let vars2 = map (fun i -> mk_comb (x_tm, mk_small_numeral i)) (1--n) in - mk_abs (x_var, subst (zip vars2 vars) expr_tm), + mk_abs (x_var, subst (zip vars2 vars) expr_tm), (if n = 0 then mk_vector_list [x_var] else mk_vector_list vars);; -(* Converts a list of terms in the form [`x + y`; `z`] +(* Converts a list of terms in the form [`x + y`; `z`] into a list of terms [`\x:real^3. x$1 + x$2`; `\x:real^3. x$3` *) let exprs_to_vector_fun = let comp_op = `$` in @@ -484,19 +490,19 @@ let dest_ineq ineq_tm = let lhs, rhs = dest_binop le_op_real tm in let lo_flag = (frees lhs = []) in let name = (fst o dest_var) (if lo_flag then rhs else lhs) in - let val_ref = + let val_ref = try assoc name !ineqs - with Failure _ -> - let val_ref = ref (x_var_real, x_var_real) in + with Failure _ -> + let val_ref = ref (x_var_real, x_var_real) in ineqs := ((name, val_ref) :: !ineqs); val_ref in val_ref := if lo_flag then (lhs, snd !val_ref) else (fst !val_ref, rhs) in - let _ = map (fun tm -> + let _ = map (fun tm -> (try decode_ineq tm with Failure _ -> failwith ("Bad variable bound inequality: "^string_of_term tm))) conds in let names, bounds0 = unzip !ineqs in let lo, hi = unzip (map (fun r -> !r) bounds0) in let test_bounds bounds bound_name = - let _ = map2 (fun tm name -> if frees tm <> [] then + let _ = map2 (fun tm name -> if frees tm <> [] then failwith (bound_name^" bound is not defined for "^name) else ()) bounds names in () in let _ = test_bounds hi "Upper"; test_bounds lo "Lower" in @@ -514,7 +520,7 @@ let normalize_result norm_flag v1 imp_th1 domain_sub_th pass_thm = let th3 = MP imp_th1 th2 in let dom_th = (UNDISCH_ALL o SPEC v1 o REWRITE_RULE[SUBSET]) domain_sub_th in let th4 = (DISCH_ALL o MY_PROVE_HYP dom_th) th3 in - let th5 = REWRITE_RULE[IN_INTERVAL; dimindex_array.(n); + let th5 = REWRITE_RULE[IN_INTERVAL; dimindex_array.(n); gen_in_interval n; comp_thms] th4 in if norm_flag then GEN_ALL th5 else th4;; @@ -569,7 +575,7 @@ let verify_ineq0 params0 norm_flag pp ineq_tm var_names (lo_tm, hi_tm) rewrite_t let _ = !info_print_level < 1 || (Certificate.report_stats stats; true) in let c1 = Certificate.transform_result xx_float zz_float certificate in - let start, finish, result = + let start, finish, result = if !params.adaptive_precision then let _ = !info_print_level < 1 || (report0 "Informal verification... "; true) in let c1p = Informal_verifier.m_verify_list pp 1 pp [ti] c1 xx2 zz2 in @@ -593,7 +599,7 @@ let verify_ineq0 params0 norm_flag pp ineq_tm var_names (lo_tm, hi_tm) rewrite_t (* -(* A simple verification function which accepts +(* A simple verification function which accepts a list of rewrite theorems which are applied to the inequality before verification *) let verify_ineq_and_rewrite rewrite_thms params pp ineq_tm = @@ -656,7 +662,7 @@ let verify_disj_ineq0 params0 norm_flag pp ineq_tm var_names (lo_tm, hi_tm) rewr let _ = params0.convex_flag = false || (report "WARNING: covex_flag should be false"; true) in let _ = params0.allow_derivatives = false || (report "WARNING: allow_derivatives should be false"; true) in - let params = ref {params0 with + let params = ref {params0 with mono_pass_flag = false; convex_flag = false; allow_derivatives = false} in @@ -703,7 +709,7 @@ let verify_disj_ineq0 params0 norm_flag pp ineq_tm var_names (lo_tm, hi_tm) rewr -(* A simple verification function which accepts +(* A simple verification function which accepts a list of rewrite theorems which are applied to the inequality before verification *) let verify_ineq_and_rewrite rewrite_thms params pp ineq_tm = @@ -720,7 +726,7 @@ end;; let h0 = new_definition `h0 = #1.26` and lmfun = new_definition`lmfun h = if (h<=h0) then (h0 - h)/(h0 - &1) else &0` and hplus = new_definition `hplus = #1.3254` and - marchal_quartic = new_definition `marchal_quartic h = + marchal_quartic = new_definition `marchal_quartic h = (sqrt(&2)-h)*(h- hplus )*(&9*(h pow 2) - &17*h + &3)/ ((sqrt(&2) - &1)* &5 *(hplus - &1))`;; @@ -734,7 +740,7 @@ let lmfun_cond_ge = (UNDISCH_ALL o prove)(`h0 <= x ==> lmfun x = &0`, REAL_ARITH_TAC);; let eq_tm = `marchal_quartic x - lmfun x`;; -let eq_th1, eq_th2 = +let eq_th1, eq_th2 = let ths = [marchal_quartic; h0; hplus] in REWRITE_CONV (lmfun_cond_le :: ths) eq_tm, REWRITE_CONV (REAL_SUB_RZERO :: lmfun_cond_ge :: ths) eq_tm;; diff --git a/Formal_ineqs/verifier_options.hl b/Formal_ineqs/verifier_options.hl index ad99b454..e5c924c3 100644 --- a/Formal_ineqs/verifier_options.hl +++ b/Formal_ineqs/verifier_options.hl @@ -1,18 +1,24 @@ -(* =========================================================== *) -(* Options of the verification library *) -(* Author: Alexey Solovyev *) -(* Date: 2012-10-27 *) -(* =========================================================== *) - -module Verifier_options = struct - -let report0 s = - Format.print_string s; Format.print_flush();; - -(* Debug/info printing level: - 0 - no debug/info printing - 1 - print important messages only - 2 - print all information *) -let info_print_level = ref 1;; - -end;; +(* ========================================================================== *) +(* Formal verification of nonlinear inequalities in HOL Light *) +(* *) +(* Copyright (c) 2012 Alexey Solovyev *) +(* *) +(* This file is distributed under the terms of the MIT licence *) +(* ========================================================================== *) + +(* -------------------------------------------------------------------------- *) +(* Options of the verification library *) +(* -------------------------------------------------------------------------- *) + +module Verifier_options = struct + +let report0 s = + Format.print_string s; Format.print_flush();; + +(* Debug/info printing level: + 0 - no debug/info printing + 1 - print important messages only + 2 - print all information *) +let info_print_level = ref 1;; + +end;; diff --git a/bignum_zarith.ml b/bignum_zarith.ml index 1746ede7..d01a96b5 100644 --- a/bignum_zarith.ml +++ b/bignum_zarith.ml @@ -102,6 +102,14 @@ module Num = struct try Q.of_bigint (Z.of_string s) with _ -> failwith "num_of_string" + let big_int_of_num = Q.to_bigint + + let num_of_big_int = Q.of_bigint + +end;; + +module Big_int = struct + include Big_int_Z end;; let (=/) x y = Num.eq_num x y;;