src/HOL/Hyperreal/hypreal_arith0.ML
author paulson
Thu, 01 Jan 2004 10:06:32 +0100
changeset 14334 6137d24eef79
parent 14331 8dbbb7cf3637
permissions -rw-r--r--
tweaking of lemmas in RealDef, RealOrd

(*  Title:      HOL/Hyperreal/hypreal_arith.ML
    ID:         $Id$
    Author:     Tobias Nipkow, TU Muenchen
    Copyright   1999 TU Muenchen

Instantiation of the generic linear arithmetic package for type hypreal.
*)

local

(* reduce contradictory <= to False *)
val add_rules =
    [order_less_irrefl, hypreal_numeral_0_eq_0, hypreal_numeral_1_eq_1,
     add_hypreal_number_of, minus_hypreal_number_of, diff_hypreal_number_of,
     mult_hypreal_number_of, eq_hypreal_number_of, less_hypreal_number_of,
     le_hypreal_number_of_eq_not_less, hypreal_diff_def,
     minus_add_distrib, minus_minus, mult_assoc, minus_zero,
     hypreal_add_zero_left, hypreal_add_zero_right,
     hypreal_add_minus, hypreal_add_minus_left,
     mult_left_zero, mult_right_zero,
     mult_1, mult_1_right,
     hypreal_mult_minus_1, hypreal_mult_minus_1_right];

val simprocs = [Hyperreal_Times_Assoc.conv, 
                Hyperreal_Numeral_Simprocs.combine_numerals]@
               Hyperreal_Numeral_Simprocs.cancel_numerals @
               Hyperreal_Numeral_Simprocs.eval_numerals;

val mono_ss = simpset() addsimps
                [add_mono, add_strict_mono,
                 hypreal_add_less_le_mono,hypreal_add_le_less_mono];

val add_mono_thms_hypreal =
  map (fn s => prove_goal (the_context ()) s
                 (fn prems => [cut_facts_tac prems 1, asm_simp_tac mono_ss 1]))
    ["(i <= j) & (k <= l) ==> i + k <= j + (l::hypreal)",
     "(i  = j) & (k <= l) ==> i + k <= j + (l::hypreal)",
     "(i <= j) & (k  = l) ==> i + k <= j + (l::hypreal)",
     "(i  = j) & (k  = l) ==> i + k  = j + (l::hypreal)",
     "(i < j) & (k = l)   ==> i + k < j + (l::hypreal)",
     "(i = j) & (k < l)   ==> i + k < j + (l::hypreal)",
     "(i < j) & (k <= l)  ==> i + k < j + (l::hypreal)",
     "(i <= j) & (k < l)  ==> i + k < j + (l::hypreal)",
     "(i < j) & (k < l)   ==> i + k < j + (l::hypreal)"];

fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;

val hypreal_mult_mono_thms =
 [(rotate_prems 1 hypreal_mult_less_mono2,
   cvar(hypreal_mult_less_mono2, hd(prems_of hypreal_mult_less_mono2))),
  (mult_left_mono,
   cvar(mult_left_mono, hd(tl(prems_of mult_left_mono))))]

in

val fast_hypreal_arith_simproc =
  Simplifier.simproc (Theory.sign_of (the_context ()))
    "fast_hypreal_arith" ["(m::hypreal) < n", "(m::hypreal) <= n", "(m::hypreal) = n"]
    Fast_Arith.lin_arith_prover;

val hypreal_arith_setup =
 [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
   {add_mono_thms = add_mono_thms @ add_mono_thms_hypreal,
    mult_mono_thms = mult_mono_thms @ hypreal_mult_mono_thms,
    inj_thms = inj_thms, (*FIXME: add hypreal*)
    lessD = lessD,  (*We don't change LA_Data_Ref.lessD because the hypreal ordering is dense!*)
    simpset = simpset addsimps add_rules addsimprocs simprocs}),
  arith_discrete ("HyperDef.hypreal",false),
  Simplifier.change_simpset_of (op addsimprocs) [fast_hypreal_arith_simproc]];

end;


(* Some test data [omitting examples that assume the ordering to be discrete!]
Goal "!!a::hypreal. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d";
by (fast_arith_tac 1);
qed "";

Goal "!!a::hypreal. [| a <= b; b+b <= c |] ==> a+a <= c";
by (fast_arith_tac 1);
qed "";

Goal "!!a::hypreal. [| a+b <= i+j; a<=b; i<=j |] ==> a+a <= j+j";
by (fast_arith_tac 1);
qed "";

Goal "!!a::hypreal. a+b+c <= i+j+k & a<=b & b<=c & i<=j & j<=k --> a+a+a <= k+k+k";
by (arith_tac 1);
qed "";

Goal "!!a::hypreal. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
\     ==> a <= l";
by (fast_arith_tac 1);
qed "";

Goal "!!a::hypreal. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
\     ==> a+a+a+a <= l+l+l+l";
by (fast_arith_tac 1);
qed "";

Goal "!!a::hypreal. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
\     ==> a+a+a+a+a <= l+l+l+l+i";
by (fast_arith_tac 1);
qed "";

Goal "!!a::hypreal. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
\     ==> a+a+a+a+a+a <= l+l+l+l+i+l";
by (fast_arith_tac 1);
qed "";

Goal "!!a::hypreal. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
\     ==> 6*a <= 5*l+i";
by (fast_arith_tac 1);
qed "";
*)