src/HOL/Real/real_arith.ML
changeset 10722 55c8367bab05
parent 10693 9e4a0e84d0d6
child 10758 9d766f21cf41
equal deleted inserted replaced
10721:12b166418455 10722:55c8367bab05
     1 (*  Title:      HOL/Real/real_arith.ML
     1 (*  Title:      HOL/Real/real_arith.ML
     2     ID:         $Id$
     2     ID:         $Id$
     3     Author:     Tobias Nipkow, TU Muenchen
     3     Author:     Tobias Nipkow, TU Muenchen
     4     Copyright   1999 TU Muenchen
     4     Copyright   1999 TU Muenchen
     5 
     5 
     6 Instantiation of the generic linear arithmetic package for type real.
     6 Augmentation of real linear arithmetic with rational coefficient handling
     7 *)
     7 *)
     8 
     8 
     9 local
     9 local
    10 
    10 
    11 (* reduce contradictory <= to False *)
    11 (* reduce contradictory <= to False *)
    12 val simps = [order_less_irrefl, zero_eq_numeral_0, one_eq_numeral_1,
    12 val simps = [True_implies_equals,inst "w" "number_of ?v" real_add_mult_distrib2,
    13              add_real_number_of, minus_real_number_of, diff_real_number_of,
    13              real_divide_1,real_times_divide1_eq,real_times_divide2_eq];
    14              mult_real_number_of, eq_real_number_of, less_real_number_of,
       
    15              le_real_number_of_eq_not_less, real_diff_def,
       
    16              real_minus_add_distrib, real_minus_minus];
       
    17 
    14 
    18 val add_rules =
    15 val simprocs = [hd(rev real_cancel_numeral_factors)];
    19     map rename_numerals
       
    20         [real_add_zero_left, real_add_zero_right,
       
    21          real_add_minus, real_add_minus_left,
       
    22          real_mult_0, real_mult_0_right,
       
    23          real_mult_1, real_mult_1_right,
       
    24          real_mult_minus_1, real_mult_minus_1_right];
       
    25 
       
    26 val simprocs = [Real_Times_Assoc.conv, Real_Numeral_Simprocs.combine_numerals]@
       
    27                Real_Numeral_Simprocs.cancel_numerals(* @ real_cancel_numeral_factors*);
       
    28 
       
    29 val mono_ss = simpset() addsimps
       
    30                 [real_add_le_mono,real_add_less_mono,
       
    31                  real_add_less_le_mono,real_add_le_less_mono];
       
    32 
       
    33 val add_mono_thms_real =
       
    34   map (fn s => prove_goal (the_context ()) s
       
    35                  (fn prems => [cut_facts_tac prems 1, asm_simp_tac mono_ss 1]))
       
    36     ["(i <= j) & (k <= l) ==> i + k <= j + (l::real)",
       
    37      "(i  = j) & (k <= l) ==> i + k <= j + (l::real)",
       
    38      "(i <= j) & (k  = l) ==> i + k <= j + (l::real)",
       
    39      "(i  = j) & (k  = l) ==> i + k  = j + (l::real)",
       
    40      "(i < j) & (k = l)   ==> i + k < j + (l::real)",
       
    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 
       
    46 val real_arith_simproc_pats =
       
    47   map (fn s => Thm.read_cterm (Theory.sign_of (the_context ())) (s, HOLogic.boolT))
       
    48       ["(m::real) < n","(m::real) <= n", "(m::real) = n"];
       
    49 
    16 
    50 fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
    17 fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
    51 
    18 
    52 val real_mult_mono_thms =
    19 val real_mult_mono_thms =
    53  [(rotate_prems 1 real_mult_less_mono2,
    20  [(rotate_prems 1 real_mult_less_mono2,
    55   (real_mult_le_mono2,
    22   (real_mult_le_mono2,
    56    cvar(real_mult_le_mono2, hd(tl(prems_of real_mult_le_mono2))))]
    23    cvar(real_mult_le_mono2, hd(tl(prems_of real_mult_le_mono2))))]
    57 
    24 
    58 in
    25 in
    59 
    26 
    60 val fast_real_arith_simproc = mk_simproc
       
    61   "fast_real_arith" real_arith_simproc_pats Fast_Arith.lin_arith_prover;
       
    62 
       
    63 val real_arith_setup =
    27 val real_arith_setup =
    64  [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
    28  [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
    65    {add_mono_thms = add_mono_thms @ add_mono_thms_real,
    29    {add_mono_thms = add_mono_thms,
    66     mult_mono_thms = mult_mono_thms @ real_mult_mono_thms,
    30     mult_mono_thms = mult_mono_thms @ real_mult_mono_thms,
    67     inj_thms = inj_thms, (*FIXME: add real*)
    31     inj_thms = inj_thms,
    68     lessD = lessD,  (*We don't change LA_Data_Ref.lessD because the real ordering is dense!*)
    32     lessD = lessD,
    69     simpset = simpset addsimps (True_implies_equals :: add_rules @ simps)
    33     simpset = simpset addsimps simps addsimprocs simprocs})];
    70                       addsimprocs simprocs
       
    71                       addcongs [if_weak_cong]}),
       
    72   arith_discrete ("RealDef.real",false),
       
    73   Simplifier.change_simpset_of (op addsimprocs) [fast_real_arith_simproc]];
       
    74 
    34 
    75 end;
    35 end;
    76 
    36 
    77 
    37 (*
    78 (* Some test data [omitting examples that assume the ordering to be discrete!]
    38 Procedure "assoc_fold" needed?
    79 Goal "!!a::real. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d";
       
    80 by (fast_arith_tac 1);
       
    81 qed "";
       
    82 
       
    83 Goal "!!a::real. [| a <= b; b+b <= c |] ==> a+a <= c";
       
    84 by (fast_arith_tac 1);
       
    85 qed "";
       
    86 
       
    87 Goal "!!a::real. [| a+b <= i+j; a<=b; i<=j |] ==> a+a <= j+j";
       
    88 by (fast_arith_tac 1);
       
    89 qed "";
       
    90 
       
    91 Goal "!!a::real. a+b+c <= i+j+k & a<=b & b<=c & i<=j & j<=k --> a+a+a <= k+k+k";
       
    92 by (arith_tac 1);
       
    93 qed "";
       
    94 
       
    95 Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
       
    96 \     ==> a <= l";
       
    97 by (fast_arith_tac 1);
       
    98 qed "";
       
    99 
       
   100 Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
       
   101 \     ==> a+a+a+a <= l+l+l+l";
       
   102 by (fast_arith_tac 1);
       
   103 qed "";
       
   104 
       
   105 Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
       
   106 \     ==> a+a+a+a+a <= l+l+l+l+i";
       
   107 by (fast_arith_tac 1);
       
   108 qed "";
       
   109 
       
   110 Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
       
   111 \     ==> a+a+a+a+a+a <= l+l+l+l+i+l";
       
   112 by (fast_arith_tac 1);
       
   113 qed "";
       
   114 
       
   115 Goal "!!a::real. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
       
   116 \     ==> #6*a <= #5*l+i";
       
   117 by (fast_arith_tac 1);
       
   118 qed "";
       
   119 *)
    39 *)