src/HOL/Hyperreal/hypreal_arith.ML
author paulson
Sun, 15 Feb 2004 10:46:37 +0100
changeset 14387 e96d5c42c4b0
parent 14370 b0064703967b
child 15184 d2c19aea17bc
permissions -rw-r--r--
Polymorphic treatment of binary arithmetic using axclasses
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14309
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
     1
(*  Title:      HOL/Hyperreal/hypreal_arith.ML
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     2
    ID:         $Id$
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     3
    Author:     Tobias Nipkow, TU Muenchen
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     4
    Copyright   1999 TU Muenchen
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     5
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
     6
Simprocs for common factor cancellation & Rational coefficient handling
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
     7
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
     8
Instantiation of the generic linear arithmetic package for type hypreal.
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
     9
*)
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    10
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    11
(*FIXME DELETE*)
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 14369
diff changeset
    12
val hypreal_mult_less_mono2 = thm"hypreal_mult_less_mono2";
b0064703967b simplifications in the hyperreals
paulson
parents: 14369
diff changeset
    13
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    14
val hypreal_mult_left_mono =
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    15
    read_instantiate_sg(sign_of (the_context())) [("a","?a::hypreal")] mult_left_mono;
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    16
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    17
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    18
(****Instantiation of the generic linear arithmetic package****)
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    19
14309
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
    20
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    21
local
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    22
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    23
fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    24
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    25
val hypreal_mult_mono_thms =
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    26
 [(rotate_prems 1 hypreal_mult_less_mono2,
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    27
   cvar(hypreal_mult_less_mono2, hd(prems_of hypreal_mult_less_mono2))),
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    28
  (hypreal_mult_left_mono,
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    29
   cvar(hypreal_mult_left_mono, hd(tl(prems_of hypreal_mult_left_mono))))]
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    30
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    31
val real_inj_thms = [hypreal_of_real_le_iff RS iffD2, 
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    32
                     hypreal_of_real_less_iff RS iffD2,
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    33
                     hypreal_of_real_eq_iff RS iffD2];
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    34
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    35
in
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    36
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14370
diff changeset
    37
val hyprealT = Type("Rational.hypreal", []);
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14370
diff changeset
    38
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    39
val fast_hypreal_arith_simproc =
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    40
    Simplifier.simproc (Theory.sign_of (the_context ()))
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    41
      "fast_hypreal_arith" 
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    42
      ["(m::hypreal) < n", "(m::hypreal) <= n", "(m::hypreal) = n"]
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    43
    Fast_Arith.lin_arith_prover;
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    44
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    45
val hypreal_arith_setup =
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    46
 [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    47
   {add_mono_thms = add_mono_thms,
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    48
    mult_mono_thms = mult_mono_thms @ hypreal_mult_mono_thms,
14369
c50188fe6366 tidying up arithmetic for the hyperreals
paulson
parents: 14352
diff changeset
    49
    inj_thms = inj_thms @ real_inj_thms, 
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    50
    lessD = lessD,  (*Can't change LA_Data_Ref.lessD: the hypreals are dense!*)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14370
diff changeset
    51
    simpset = simpset}),
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14370
diff changeset
    52
  arith_inj_const ("HyperDef.hypreal_of_real", HOLogic.realT --> hyprealT),
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    53
  arith_discrete ("HyperDef.hypreal",false),
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    54
  Simplifier.change_simpset_of (op addsimprocs) [fast_hypreal_arith_simproc]];
10751
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    55
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    56
end;
a81ea5d3dd41 separation of HOL-Hyperreal from HOL-Real
paulson
parents:
diff changeset
    57
14352
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    58
a8b1a44d8264 Modified real arithmetic simplification
paulson
parents: 14334
diff changeset
    59