src/HOL/Real/real_arith0.ML
changeset 14352 a8b1a44d8264
parent 14351 cd3ef10d02be
child 14353 79f9fbef9106
equal deleted inserted replaced
14351:cd3ef10d02be 14352:a8b1a44d8264
     1 (*  Title:      HOL/Real/real_arith0.ML
       
     2     ID:         $Id$
       
     3     Author:     Tobias Nipkow, TU Muenchen
       
     4     Copyright   1999 TU Muenchen
       
     5 
       
     6 Instantiation of the generic linear arithmetic package for type real.
       
     7 *)
       
     8 
       
     9 val add_zero_left = thm"Ring_and_Field.add_0";
       
    10 val add_zero_right = thm"Ring_and_Field.add_0_right";
       
    11 
       
    12 val real_mult_left_mono =
       
    13     read_instantiate_sg(sign_of RealBin.thy) [("a","?a::real")] mult_left_mono;
       
    14 
       
    15 
       
    16 local
       
    17 
       
    18 (* reduce contradictory <= to False *)
       
    19 val add_rules = 
       
    20     [order_less_irrefl, real_numeral_0_eq_0, real_numeral_1_eq_1,
       
    21      real_minus_1_eq_m1, 
       
    22      add_real_number_of, minus_real_number_of, diff_real_number_of,
       
    23      mult_real_number_of, eq_real_number_of, less_real_number_of,
       
    24      le_real_number_of_eq_not_less, real_diff_def,
       
    25      minus_add_distrib, minus_minus, mult_assoc, minus_zero,
       
    26      add_zero_left, add_zero_right, left_minus, right_minus,
       
    27      mult_left_zero, mult_right_zero, mult_1, mult_1_right,
       
    28      minus_mult_left RS sym, minus_mult_right RS sym];
       
    29 
       
    30 val simprocs = [Real_Times_Assoc.conv, Real_Numeral_Simprocs.combine_numerals]@
       
    31                Real_Numeral_Simprocs.cancel_numerals @
       
    32                Real_Numeral_Simprocs.eval_numerals;
       
    33 
       
    34 val mono_ss = simpset() addsimps
       
    35                 [add_mono,add_strict_mono,
       
    36                  real_add_less_le_mono,real_add_le_less_mono];
       
    37 
       
    38 val add_mono_thms_real =
       
    39   map (fn s => prove_goal (the_context ()) s
       
    40                  (fn prems => [cut_facts_tac prems 1, asm_simp_tac mono_ss 1]))
       
    41     ["(i <= j) & (k <= l) ==> i + k <= j + (l::real)",
       
    42      "(i  = j) & (k <= l) ==> i + k <= j + (l::real)",
       
    43      "(i <= j) & (k  = l) ==> i + k <= j + (l::real)",
       
    44      "(i  = j) & (k  = l) ==> i + k  = j + (l::real)",
       
    45      "(i < j) & (k = l)   ==> i + k < j + (l::real)",
       
    46      "(i = j) & (k < l)   ==> i + k < j + (l::real)",
       
    47      "(i < j) & (k <= l)  ==> i + k < j + (l::real)",
       
    48      "(i <= j) & (k < l)  ==> i + k < j + (l::real)",
       
    49      "(i < j) & (k < l)   ==> i + k < j + (l::real)"];
       
    50 
       
    51 fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
       
    52 
       
    53 val real_mult_mono_thms =
       
    54  [(rotate_prems 1 real_mult_less_mono2,
       
    55    cvar(real_mult_less_mono2, hd(prems_of real_mult_less_mono2))),
       
    56   (real_mult_left_mono,
       
    57    cvar(real_mult_left_mono, hd(tl(prems_of real_mult_left_mono))))]
       
    58 
       
    59 in
       
    60 
       
    61 val fast_real_arith_simproc = Simplifier.simproc (Theory.sign_of (the_context ()))
       
    62   "fast_real_arith" ["(m::real) < n","(m::real) <= n", "(m::real) = n"]
       
    63   Fast_Arith.lin_arith_prover;
       
    64 
       
    65 val real_arith_setup =
       
    66  [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
       
    67    {add_mono_thms = add_mono_thms @ add_mono_thms_real,
       
    68     mult_mono_thms = mult_mono_thms @ real_mult_mono_thms,
       
    69     inj_thms = inj_thms, (*FIXME: add real*)
       
    70     lessD = lessD,  (*We don't change LA_Data_Ref.lessD because the real ordering is dense!*)
       
    71     simpset = simpset addsimps add_rules
       
    72                       addsimprocs simprocs}),
       
    73   arith_discrete ("RealDef.real",false),
       
    74   Simplifier.change_simpset_of (op addsimprocs) [fast_real_arith_simproc]];
       
    75 
       
    76 (* some thms for injection nat => real:
       
    77 real_of_nat_zero
       
    78 ?zero_eq_numeral_0
       
    79 real_of_nat_add
       
    80 *)
       
    81 
       
    82 end;
       
    83 
       
    84 
       
    85 (* Some test data [omitting examples that assume the ordering to be discrete!]
       
    86 Goal "!!a::real. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d";
       
    87 by (fast_arith_tac 1);
       
    88 qed "";
       
    89 
       
    90 Goal "!!a::real. [| a <= b; b+b <= c |] ==> a+a <= c";
       
    91 by (fast_arith_tac 1);
       
    92 qed "";
       
    93 
       
    94 Goal "!!a::real. [| a+b <= i+j; a<=b; i<=j |] ==> a+a <= j+j";
       
    95 by (fast_arith_tac 1);
       
    96 qed "";
       
    97 
       
    98 Goal "!!a::real. a+b+c <= i+j+k & a<=b & b<=c & i<=j & j<=k --> a+a+a <= k+k+k";
       
    99 by (arith_tac 1);
       
   100 qed "";
       
   101 
       
   102 Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
       
   103 \     ==> a <= l";
       
   104 by (fast_arith_tac 1);
       
   105 qed "";
       
   106 
       
   107 Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
       
   108 \     ==> a+a+a+a <= l+l+l+l";
       
   109 by (fast_arith_tac 1);
       
   110 qed "";
       
   111 
       
   112 Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
       
   113 \     ==> a+a+a+a+a <= l+l+l+l+i";
       
   114 by (fast_arith_tac 1);
       
   115 qed "";
       
   116 
       
   117 Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
       
   118 \     ==> a+a+a+a+a+a <= l+l+l+l+i+l";
       
   119 by (fast_arith_tac 1);
       
   120 qed "";
       
   121 
       
   122 Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
       
   123 \     ==> 6*a <= 5*l+i";
       
   124 by (fast_arith_tac 1);
       
   125 qed "";
       
   126 
       
   127 Goal "a<=b ==> a < b+(1::real)";
       
   128 by (fast_arith_tac 1);
       
   129 qed "";
       
   130 
       
   131 Goal "a<=b ==> a-(3::real) < b";
       
   132 by (fast_arith_tac 1);
       
   133 qed "";
       
   134 
       
   135 Goal "a<=b ==> a-(1::real) < b";
       
   136 by (fast_arith_tac 1);
       
   137 qed "";
       
   138 
       
   139 *)