| author | urbanc | 
| Thu, 22 Mar 2007 11:17:32 +0100 | |
| changeset 22495 | c54748fd1f43 | 
| parent 21243 | afffe1f72143 | 
| child 22548 | 6ce4bddf3bcb | 
| permissions | -rw-r--r-- | 
| 20254 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 1 | (* Title: HOL/Real/rat_arith.ML | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 4 | Copyright 2004 University of Cambridge | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 5 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 6 | Simprocs for common factor cancellation & Rational coefficient handling | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 7 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 8 | Instantiation of the generic linear arithmetic package for type rat. | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 9 | *) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 10 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 11 | local | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 12 | |
| 14390 | 13 | val simprocs = field_cancel_numeral_factors | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 14 | |
| 14390 | 15 | val simps = [order_less_irrefl, neg_less_iff_less, True_implies_equals, | 
| 14369 | 16 | inst "a" "(number_of ?v)" right_distrib, | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 17 | divide_1, divide_zero_left, | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 18 | times_divide_eq_right, times_divide_eq_left, | 
| 14390 | 19 | minus_divide_left RS sym, minus_divide_right RS sym, | 
| 20254 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 20 | of_int_0, of_int_1, of_int_add, of_int_minus, of_int_diff, | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 21 | of_int_mult, of_int_of_nat_eq] | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 22 | |
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 23 | val nat_inj_thms = [of_nat_le_iff RS iffD2, | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 24 | of_nat_eq_iff RS iffD2] | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 25 | (* not needed because x < (y::nat) can be rewritten as Suc x <= y: | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 26 | of_nat_less_iff RS iffD2 *) | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 27 | |
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 28 | val int_inj_thms = [of_int_le_iff RS iffD2, | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 29 | of_int_eq_iff RS iffD2] | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 30 | (* not needed because x < (y::int) can be rewritten as x + 1 <= y: | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 31 | of_int_less_iff RS iffD2 *) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 32 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 33 | in | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 34 | |
| 20254 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 35 | val fast_rat_arith_simproc = | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 36 | Simplifier.simproc (the_context ()) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 37 | "fast_rat_arith" ["(m::rat) < n","(m::rat) <= n", "(m::rat) = n"] | 
| 20254 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 38 | Fast_Arith.lin_arith_prover | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 39 | |
| 20254 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 40 | val ratT = Type ("Rational.rat", [])
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 41 | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 42 | val rat_arith_setup = | 
| 18708 | 43 |   Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
 | 
| 15121 
1198032bad25
Initial changes to extend arithmetic from individual types to type classes.
 nipkow parents: 
14390diff
changeset | 44 |    {add_mono_thms = add_mono_thms,
 | 
| 15184 | 45 | mult_mono_thms = mult_mono_thms, | 
| 20254 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 46 | inj_thms = int_inj_thms @ nat_inj_thms @ inj_thms, | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 47 | lessD = lessD, (*Can't change LA_Data_Ref.lessD: the rats are dense!*) | 
| 15923 | 48 | neqE = neqE, | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 49 | simpset = simpset addsimps simps | 
| 18708 | 50 | addsimprocs simprocs}) #> | 
| 21243 | 51 |   arith_inj_const ("Nat.of_nat", HOLogic.natT --> ratT) #>
 | 
| 20254 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 52 |   arith_inj_const ("IntDef.of_int", HOLogic.intT --> ratT) #>
 | 
| 
58b71535ed00
lin_arith_prover splits certain operators (e.g. min, max, abs)
 webertj parents: 
18708diff
changeset | 53 | (fn thy => (change_simpset_of thy (fn ss => ss addsimprocs [fast_rat_arith_simproc]); thy)) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 54 | |
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: diff
changeset | 55 | end; |