diff -r dd58f13a8eb4 -r eb85850d3eb7 src/HOL/Hyperreal/Integration.thy --- a/src/HOL/Hyperreal/Integration.thy Fri Nov 17 02:19:55 2006 +0100 +++ b/src/HOL/Hyperreal/Integration.thy Fri Nov 17 02:20:03 2006 +0100 @@ -13,38 +13,43 @@ text{*We follow John Harrison in formalizing the Gauge integral.*} definition - --{*Partitions and tagged partitions etc.*} - partition :: "[(real*real),nat => real] => bool" + partition :: "[(real*real),nat => real] => bool" where "partition = (%(a,b) D. D 0 = a & (\N. (\n < N. D(n) < D(Suc n)) & (\n \ N. D(n) = b)))" - psize :: "(nat => real) => nat" +definition + psize :: "(nat => real) => nat" where "psize D = (SOME N. (\n < N. D(n) < D(Suc n)) & (\n \ N. D(n) = D(N)))" - tpart :: "[(real*real),((nat => real)*(nat =>real))] => bool" +definition + tpart :: "[(real*real),((nat => real)*(nat =>real))] => bool" where "tpart = (%(a,b) (D,p). partition(a,b) D & (\n. D(n) \ p(n) & p(n) \ D(Suc n)))" --{*Gauges and gauge-fine divisions*} - gauge :: "[real => bool, real => real] => bool" +definition + gauge :: "[real => bool, real => real] => bool" where "gauge E g = (\x. E x --> 0 < g(x))" - fine :: "[real => real, ((nat => real)*(nat => real))] => bool" +definition + fine :: "[real => real, ((nat => real)*(nat => real))] => bool" where "fine = (%g (D,p). \n. n < (psize D) --> D(Suc n) - D(n) < g(p n))" --{*Riemann sum*} - rsum :: "[((nat=>real)*(nat=>real)),real=>real] => real" +definition + rsum :: "[((nat=>real)*(nat=>real)),real=>real] => real" where "rsum = (%(D,p) f. \n=0..real,real] => bool" +definition + Integral :: "[(real*real),real=>real,real] => bool" where "Integral = (%(a,b) f k. \e > 0. (\g. gauge(%x. a \ x & x \ b) g & (\D p. tpart(a,b) (D,p) & fine(g)(D,p) -->