| author | wenzelm | 
| Sat, 21 Sep 2013 19:48:46 +0200 | |
| changeset 53777 | 06a6216f733e | 
| parent 53755 | b8e62805566b | 
| child 54230 | b1d955791529 | 
| permissions | -rw-r--r-- | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31366 
diff
changeset
 | 
1  | 
(* Author: Jacques D. Fleuriot, University of Edinburgh  | 
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
2  | 
Conversion to Isar and new proofs by Lawrence C Paulson, 2004  | 
| 35328 | 3  | 
|
4  | 
Replaced by ~~/src/HOL/Multivariate_Analysis/Real_Integral.thy .  | 
|
| 13958 | 5  | 
*)  | 
6  | 
||
| 35328 | 7  | 
header{*Theory of Integration on real intervals*}
 | 
8  | 
||
9  | 
theory Gauge_Integration  | 
|
10  | 
imports Complex_Main  | 
|
11  | 
begin  | 
|
12  | 
||
13  | 
text {*
 | 
|
| 13958 | 14  | 
|
| 35328 | 15  | 
\textbf{Attention}: This theory defines the Integration on real
 | 
16  | 
intervals. This is just a example theory for historical / expository interests.  | 
|
17  | 
A better replacement is found in the Multivariate Analysis library. This defines  | 
|
18  | 
the gauge integral on real vector spaces and in the Real Integral theory  | 
|
19  | 
is a specialization to the integral on arbitrary real intervals. The  | 
|
20  | 
Multivariate Analysis package also provides a better support for analysis on  | 
|
21  | 
integrals.  | 
|
22  | 
||
23  | 
*}  | 
|
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
24  | 
|
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
25  | 
text{*We follow John Harrison in formalizing the Gauge integral.*}
 | 
| 13958 | 26  | 
|
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
27  | 
subsection {* Gauges *}
 | 
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
28  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20792 
diff
changeset
 | 
29  | 
definition  | 
| 31253 | 30  | 
gauge :: "[real set, real => real] => bool" where  | 
| 37765 | 31  | 
"gauge E g = (\<forall>x\<in>E. 0 < g(x))"  | 
| 13958 | 32  | 
|
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
33  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
34  | 
subsection {* Gauge-fine divisions *}
 | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
35  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
36  | 
inductive  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
37  | 
fine :: "[real \<Rightarrow> real, real \<times> real, (real \<times> real \<times> real) list] \<Rightarrow> bool"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
38  | 
for  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
39  | 
\<delta> :: "real \<Rightarrow> real"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
40  | 
where  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
41  | 
fine_Nil:  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
42  | 
"fine \<delta> (a, a) []"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
43  | 
| fine_Cons:  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
44  | 
"\<lbrakk>fine \<delta> (b, c) D; a < b; a \<le> x; x \<le> b; b - a < \<delta> x\<rbrakk>  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
45  | 
\<Longrightarrow> fine \<delta> (a, c) ((a, x, b) # D)"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
46  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
47  | 
lemmas fine_induct [induct set: fine] =  | 
| 45605 | 48  | 
fine.induct [of "\<delta>" "(a,b)" "D" "split P", unfolded split_conv] for \<delta> a b D P  | 
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
49  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
50  | 
lemma fine_single:  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
51  | 
"\<lbrakk>a < b; a \<le> x; x \<le> b; b - a < \<delta> x\<rbrakk> \<Longrightarrow> fine \<delta> (a, b) [(a, x, b)]"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
52  | 
by (rule fine_Cons [OF fine_Nil])  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
53  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
54  | 
lemma fine_append:  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
55  | 
"\<lbrakk>fine \<delta> (a, b) D; fine \<delta> (b, c) D'\<rbrakk> \<Longrightarrow> fine \<delta> (a, c) (D @ D')"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
56  | 
by (induct set: fine, simp, simp add: fine_Cons)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
57  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
58  | 
lemma fine_imp_le: "fine \<delta> (a, b) D \<Longrightarrow> a \<le> b"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
59  | 
by (induct set: fine, simp_all)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
60  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
61  | 
lemma nonempty_fine_imp_less: "\<lbrakk>fine \<delta> (a, b) D; D \<noteq> []\<rbrakk> \<Longrightarrow> a < b"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
62  | 
apply (induct set: fine, simp)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
63  | 
apply (drule fine_imp_le, simp)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
64  | 
done  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
65  | 
|
| 35441 | 66  | 
lemma fine_Nil_iff: "fine \<delta> (a, b) [] \<longleftrightarrow> a = b"  | 
67  | 
by (auto elim: fine.cases intro: fine.intros)  | 
|
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
68  | 
|
| 35441 | 69  | 
lemma fine_same_iff: "fine \<delta> (a, a) D \<longleftrightarrow> D = []"  | 
70  | 
proof  | 
|
71  | 
assume "fine \<delta> (a, a) D" thus "D = []"  | 
|
72  | 
by (metis nonempty_fine_imp_less less_irrefl)  | 
|
73  | 
next  | 
|
74  | 
assume "D = []" thus "fine \<delta> (a, a) D"  | 
|
75  | 
by (simp add: fine_Nil)  | 
|
76  | 
qed  | 
|
77  | 
||
78  | 
lemma empty_fine_imp_eq: "\<lbrakk>fine \<delta> (a, b) D; D = []\<rbrakk> \<Longrightarrow> a = b"  | 
|
79  | 
by (simp add: fine_Nil_iff)  | 
|
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
80  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
81  | 
lemma mem_fine:  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
82  | 
"\<lbrakk>fine \<delta> (a, b) D; (u, x, v) \<in> set D\<rbrakk> \<Longrightarrow> u < v \<and> u \<le> x \<and> x \<le> v"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
83  | 
by (induct set: fine, simp, force)  | 
| 13958 | 84  | 
|
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
85  | 
lemma mem_fine2: "\<lbrakk>fine \<delta> (a, b) D; (u, z, v) \<in> set D\<rbrakk> \<Longrightarrow> a \<le> u \<and> v \<le> b"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
86  | 
apply (induct arbitrary: z u v set: fine, auto)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
87  | 
apply (simp add: fine_imp_le)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
88  | 
apply (erule order_trans [OF less_imp_le], simp)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
89  | 
done  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
90  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
91  | 
lemma mem_fine3: "\<lbrakk>fine \<delta> (a, b) D; (u, z, v) \<in> set D\<rbrakk> \<Longrightarrow> v - u < \<delta> z"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
92  | 
by (induct arbitrary: z u v set: fine) auto  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
93  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
94  | 
lemma BOLZANO:  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
95  | 
fixes P :: "real \<Rightarrow> real \<Rightarrow> bool"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
96  | 
assumes 1: "a \<le> b"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
97  | 
assumes 2: "\<And>a b c. \<lbrakk>P a b; P b c; a \<le> b; b \<le> c\<rbrakk> \<Longrightarrow> P a c"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
98  | 
assumes 3: "\<And>x. \<exists>d>0. \<forall>a b. a \<le> x & x \<le> b & (b-a) < d \<longrightarrow> P a b"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
99  | 
shows "P a b"  | 
| 
51480
 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 
hoelzl 
parents: 
49962 
diff
changeset
 | 
100  | 
using 1 2 3 by (rule Bolzano)  | 
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
101  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
102  | 
text{*We can always find a division that is fine wrt any gauge*}
 | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
103  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
104  | 
lemma fine_exists:  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
105  | 
  assumes "a \<le> b" and "gauge {a..b} \<delta>" shows "\<exists>D. fine \<delta> (a, b) D"
 | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
106  | 
proof -  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
107  | 
  {
 | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
108  | 
fix u v :: real assume "u \<le> v"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
109  | 
have "a \<le> u \<Longrightarrow> v \<le> b \<Longrightarrow> \<exists>D. fine \<delta> (u, v) D"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
110  | 
apply (induct u v rule: BOLZANO, rule `u \<le> v`)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
111  | 
apply (simp, fast intro: fine_append)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
112  | 
apply (case_tac "a \<le> x \<and> x \<le> b")  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
113  | 
apply (rule_tac x="\<delta> x" in exI)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
114  | 
apply (rule conjI)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
115  | 
        apply (simp add: `gauge {a..b} \<delta>` [unfolded gauge_def])
 | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
116  | 
apply (clarify, rename_tac u v)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
117  | 
apply (case_tac "u = v")  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
118  | 
apply (fast intro: fine_Nil)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
119  | 
apply (subgoal_tac "u < v", fast intro: fine_single, simp)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
120  | 
apply (rule_tac x="1" in exI, clarsimp)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
121  | 
done  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
122  | 
}  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
123  | 
with `a \<le> b` show ?thesis by auto  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
124  | 
qed  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
125  | 
|
| 31364 | 126  | 
lemma fine_covers_all:  | 
127  | 
assumes "fine \<delta> (a, c) D" and "a < x" and "x \<le> c"  | 
|
128  | 
shows "\<exists> N < length D. \<forall> d t e. D ! N = (d,t,e) \<longrightarrow> d < x \<and> x \<le> e"  | 
|
129  | 
using assms  | 
|
130  | 
proof (induct set: fine)  | 
|
131  | 
case (2 b c D a t)  | 
|
132  | 
thus ?case  | 
|
133  | 
proof (cases "b < x")  | 
|
134  | 
case True  | 
|
135  | 
with 2 obtain N where *: "N < length D"  | 
|
136  | 
and **: "\<And> d t e. D ! N = (d,t,e) \<Longrightarrow> d < x \<and> x \<le> e" by auto  | 
|
137  | 
hence "Suc N < length ((a,t,b)#D) \<and>  | 
|
138  | 
(\<forall> d t' e. ((a,t,b)#D) ! Suc N = (d,t',e) \<longrightarrow> d < x \<and> x \<le> e)" by auto  | 
|
139  | 
thus ?thesis by auto  | 
|
140  | 
next  | 
|
141  | 
case False with 2  | 
|
142  | 
have "0 < length ((a,t,b)#D) \<and>  | 
|
143  | 
(\<forall> d t' e. ((a,t,b)#D) ! 0 = (d,t',e) \<longrightarrow> d < x \<and> x \<le> e)" by auto  | 
|
144  | 
thus ?thesis by auto  | 
|
145  | 
qed  | 
|
146  | 
qed auto  | 
|
147  | 
||
148  | 
lemma fine_append_split:  | 
|
149  | 
assumes "fine \<delta> (a,b) D" and "D2 \<noteq> []" and "D = D1 @ D2"  | 
|
150  | 
shows "fine \<delta> (a,fst (hd D2)) D1" (is "?fine1")  | 
|
151  | 
and "fine \<delta> (fst (hd D2), b) D2" (is "?fine2")  | 
|
152  | 
proof -  | 
|
153  | 
from assms  | 
|
154  | 
have "?fine1 \<and> ?fine2"  | 
|
155  | 
proof (induct arbitrary: D1 D2)  | 
|
156  | 
case (2 b c D a' x D1 D2)  | 
|
157  | 
note induct = this  | 
|
158  | 
||
159  | 
thus ?case  | 
|
160  | 
proof (cases D1)  | 
|
161  | 
case Nil  | 
|
162  | 
hence "fst (hd D2) = a'" using 2 by auto  | 
|
163  | 
with fine_Cons[OF `fine \<delta> (b,c) D` induct(3,4,5)] Nil induct  | 
|
164  | 
show ?thesis by (auto intro: fine_Nil)  | 
|
165  | 
next  | 
|
166  | 
case (Cons d1 D1')  | 
|
167  | 
with induct(2)[OF `D2 \<noteq> []`, of D1'] induct(8)  | 
|
168  | 
have "fine \<delta> (b, fst (hd D2)) D1'" and "fine \<delta> (fst (hd D2), c) D2" and  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31366 
diff
changeset
 | 
169  | 
"d1 = (a', x, b)" by auto  | 
| 31364 | 170  | 
with fine_Cons[OF this(1) induct(3,4,5), OF induct(6)] Cons  | 
171  | 
show ?thesis by auto  | 
|
172  | 
qed  | 
|
173  | 
qed auto  | 
|
174  | 
thus ?fine1 and ?fine2 by auto  | 
|
175  | 
qed  | 
|
176  | 
||
177  | 
lemma fine_\<delta>_expand:  | 
|
178  | 
assumes "fine \<delta> (a,b) D"  | 
|
| 35441 | 179  | 
and "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<delta> x \<le> \<delta>' x"  | 
| 31364 | 180  | 
shows "fine \<delta>' (a,b) D"  | 
181  | 
using assms proof induct  | 
|
182  | 
case 1 show ?case by (rule fine_Nil)  | 
|
183  | 
next  | 
|
184  | 
case (2 b c D a x)  | 
|
185  | 
show ?case  | 
|
186  | 
proof (rule fine_Cons)  | 
|
187  | 
show "fine \<delta>' (b,c) D" using 2 by auto  | 
|
188  | 
from fine_imp_le[OF 2(1)] 2(6) `x \<le> b`  | 
|
189  | 
show "b - a < \<delta>' x"  | 
|
190  | 
using 2(7)[OF `a \<le> x`] by auto  | 
|
191  | 
qed (auto simp add: 2)  | 
|
192  | 
qed  | 
|
193  | 
||
194  | 
lemma fine_single_boundaries:  | 
|
195  | 
assumes "fine \<delta> (a,b) D" and "D = [(d, t, e)]"  | 
|
196  | 
shows "a = d \<and> b = e"  | 
|
197  | 
using assms proof induct  | 
|
198  | 
case (2 b c D a x)  | 
|
199  | 
hence "D = []" and "a = d" and "b = e" by auto  | 
|
200  | 
moreover  | 
|
201  | 
from `fine \<delta> (b,c) D` `D = []` have "b = c"  | 
|
202  | 
by (rule empty_fine_imp_eq)  | 
|
203  | 
ultimately show ?case by simp  | 
|
204  | 
qed auto  | 
|
205  | 
||
| 35328 | 206  | 
lemma fine_listsum_eq_diff:  | 
207  | 
fixes f :: "real \<Rightarrow> real"  | 
|
208  | 
shows "fine \<delta> (a, b) D \<Longrightarrow> (\<Sum>(u, x, v)\<leftarrow>D. f v - f u) = f b - f a"  | 
|
209  | 
by (induct set: fine) simp_all  | 
|
210  | 
||
211  | 
text{*Lemmas about combining gauges*}
 | 
|
212  | 
||
213  | 
lemma gauge_min:  | 
|
214  | 
"[| gauge(E) g1; gauge(E) g2 |]  | 
|
215  | 
==> gauge(E) (%x. min (g1(x)) (g2(x)))"  | 
|
216  | 
by (simp add: gauge_def)  | 
|
217  | 
||
218  | 
lemma fine_min:  | 
|
219  | 
"fine (%x. min (g1(x)) (g2(x))) (a,b) D  | 
|
220  | 
==> fine(g1) (a,b) D & fine(g2) (a,b) D"  | 
|
221  | 
apply (erule fine.induct)  | 
|
222  | 
apply (simp add: fine_Nil)  | 
|
223  | 
apply (simp add: fine_Cons)  | 
|
224  | 
done  | 
|
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
225  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
226  | 
subsection {* Riemann sum *}
 | 
| 13958 | 227  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20792 
diff
changeset
 | 
228  | 
definition  | 
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
229  | 
rsum :: "[(real \<times> real \<times> real) list, real \<Rightarrow> real] \<Rightarrow> real" where  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
230  | 
"rsum D f = (\<Sum>(u, x, v)\<leftarrow>D. f x * (v - u))"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
231  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
232  | 
lemma rsum_Nil [simp]: "rsum [] f = 0"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
233  | 
unfolding rsum_def by simp  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
234  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
235  | 
lemma rsum_Cons [simp]: "rsum ((u, x, v) # D) f = f x * (v - u) + rsum D f"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
236  | 
unfolding rsum_def by simp  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
237  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
238  | 
lemma rsum_zero [simp]: "rsum D (\<lambda>x. 0) = 0"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
239  | 
by (induct D, auto)  | 
| 13958 | 240  | 
|
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
241  | 
lemma rsum_left_distrib: "rsum D f * c = rsum D (\<lambda>x. f x * c)"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
242  | 
by (induct D, auto simp add: algebra_simps)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
243  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
244  | 
lemma rsum_right_distrib: "c * rsum D f = rsum D (\<lambda>x. c * f x)"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
245  | 
by (induct D, auto simp add: algebra_simps)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
246  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
247  | 
lemma rsum_add: "rsum D (\<lambda>x. f x + g x) = rsum D f + rsum D g"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
248  | 
by (induct D, auto simp add: algebra_simps)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
249  | 
|
| 31364 | 250  | 
lemma rsum_append: "rsum (D1 @ D2) f = rsum D1 f + rsum D2 f"  | 
251  | 
unfolding rsum_def map_append listsum_append ..  | 
|
252  | 
||
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
253  | 
|
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
254  | 
subsection {* Gauge integrability (definite) *}
 | 
| 13958 | 255  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20792 
diff
changeset
 | 
256  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20792 
diff
changeset
 | 
257  | 
Integral :: "[(real*real),real=>real,real] => bool" where  | 
| 37765 | 258  | 
"Integral = (%(a,b) f k. \<forall>e > 0.  | 
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
259  | 
                               (\<exists>\<delta>. gauge {a .. b} \<delta> &
 | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
260  | 
(\<forall>D. fine \<delta> (a,b) D -->  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
261  | 
\<bar>rsum D f - k\<bar> < e)))"  | 
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
262  | 
|
| 35441 | 263  | 
lemma Integral_eq:  | 
264  | 
"Integral (a, b) f k \<longleftrightarrow>  | 
|
265  | 
    (\<forall>e>0. \<exists>\<delta>. gauge {a..b} \<delta> \<and> (\<forall>D. fine \<delta> (a,b) D \<longrightarrow> \<bar>rsum D f - k\<bar> < e))"
 | 
|
266  | 
unfolding Integral_def by simp  | 
|
267  | 
||
268  | 
lemma IntegralI:  | 
|
269  | 
assumes "\<And>e. 0 < e \<Longrightarrow>  | 
|
270  | 
    \<exists>\<delta>. gauge {a..b} \<delta> \<and> (\<forall>D. fine \<delta> (a, b) D \<longrightarrow> \<bar>rsum D f - k\<bar> < e)"
 | 
|
271  | 
shows "Integral (a, b) f k"  | 
|
272  | 
using assms unfolding Integral_def by auto  | 
|
273  | 
||
274  | 
lemma IntegralE:  | 
|
275  | 
assumes "Integral (a, b) f k" and "0 < e"  | 
|
276  | 
  obtains \<delta> where "gauge {a..b} \<delta>" and "\<forall>D. fine \<delta> (a, b) D \<longrightarrow> \<bar>rsum D f - k\<bar> < e"
 | 
|
277  | 
using assms unfolding Integral_def by auto  | 
|
278  | 
||
| 31252 | 279  | 
lemma Integral_def2:  | 
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
280  | 
  "Integral = (%(a,b) f k. \<forall>e>0. (\<exists>\<delta>. gauge {a..b} \<delta> &
 | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
281  | 
(\<forall>D. fine \<delta> (a,b) D -->  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
282  | 
\<bar>rsum D f - k\<bar> \<le> e)))"  | 
| 31252 | 283  | 
unfolding Integral_def  | 
284  | 
apply (safe intro!: ext)  | 
|
285  | 
apply (fast intro: less_imp_le)  | 
|
286  | 
apply (drule_tac x="e/2" in spec)  | 
|
287  | 
apply force  | 
|
288  | 
done  | 
|
289  | 
||
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
290  | 
text{*The integral is unique if it exists*}
 | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
291  | 
|
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
292  | 
lemma Integral_unique:  | 
| 35441 | 293  | 
assumes le: "a \<le> b"  | 
294  | 
assumes 1: "Integral (a, b) f k1"  | 
|
295  | 
assumes 2: "Integral (a, b) f k2"  | 
|
296  | 
shows "k1 = k2"  | 
|
297  | 
proof (rule ccontr)  | 
|
298  | 
assume "k1 \<noteq> k2"  | 
|
299  | 
hence e: "0 < \<bar>k1 - k2\<bar> / 2" by simp  | 
|
300  | 
  obtain d1 where "gauge {a..b} d1" and
 | 
|
301  | 
d1: "\<forall>D. fine d1 (a, b) D \<longrightarrow> \<bar>rsum D f - k1\<bar> < \<bar>k1 - k2\<bar> / 2"  | 
|
302  | 
using 1 e by (rule IntegralE)  | 
|
303  | 
  obtain d2 where "gauge {a..b} d2" and
 | 
|
304  | 
d2: "\<forall>D. fine d2 (a, b) D \<longrightarrow> \<bar>rsum D f - k2\<bar> < \<bar>k1 - k2\<bar> / 2"  | 
|
305  | 
using 2 e by (rule IntegralE)  | 
|
306  | 
  have "gauge {a..b} (\<lambda>x. min (d1 x) (d2 x))"
 | 
|
307  | 
    using `gauge {a..b} d1` and `gauge {a..b} d2`
 | 
|
308  | 
by (rule gauge_min)  | 
|
309  | 
then obtain D where "fine (\<lambda>x. min (d1 x) (d2 x)) (a, b) D"  | 
|
310  | 
using fine_exists [OF le] by fast  | 
|
311  | 
hence "fine d1 (a, b) D" and "fine d2 (a, b) D"  | 
|
312  | 
by (auto dest: fine_min)  | 
|
313  | 
hence "\<bar>rsum D f - k1\<bar> < \<bar>k1 - k2\<bar> / 2" and "\<bar>rsum D f - k2\<bar> < \<bar>k1 - k2\<bar> / 2"  | 
|
314  | 
using d1 d2 by simp_all  | 
|
315  | 
hence "\<bar>rsum D f - k1\<bar> + \<bar>rsum D f - k2\<bar> < \<bar>k1 - k2\<bar> / 2 + \<bar>k1 - k2\<bar> / 2"  | 
|
316  | 
by (rule add_strict_mono)  | 
|
317  | 
thus False by auto  | 
|
318  | 
qed  | 
|
319  | 
||
320  | 
lemma Integral_zero: "Integral(a,a) f 0"  | 
|
321  | 
apply (rule IntegralI)  | 
|
322  | 
apply (rule_tac x = "\<lambda>x. 1" in exI)  | 
|
323  | 
apply (simp add: fine_same_iff gauge_def)  | 
|
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
324  | 
done  | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
325  | 
|
| 35441 | 326  | 
lemma Integral_same_iff [simp]: "Integral (a, a) f k \<longleftrightarrow> k = 0"  | 
327  | 
by (auto intro: Integral_zero Integral_unique)  | 
|
328  | 
||
329  | 
lemma Integral_zero_fun: "Integral (a,b) (\<lambda>x. 0) 0"  | 
|
330  | 
apply (rule IntegralI)  | 
|
331  | 
apply (rule_tac x="\<lambda>x. 1" in exI, simp add: gauge_def)  | 
|
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
332  | 
done  | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
333  | 
|
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
334  | 
lemma fine_rsum_const: "fine \<delta> (a,b) D \<Longrightarrow> rsum D (\<lambda>x. c) = (c * (b - a))"  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
335  | 
unfolding rsum_def  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
336  | 
by (induct set: fine, auto simp add: algebra_simps)  | 
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
337  | 
|
| 35441 | 338  | 
lemma Integral_mult_const: "a \<le> b \<Longrightarrow> Integral(a,b) (\<lambda>x. c) (c * (b - a))"  | 
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
339  | 
apply (cases "a = b", simp)  | 
| 35441 | 340  | 
apply (rule IntegralI)  | 
341  | 
apply (rule_tac x = "\<lambda>x. b - a" in exI)  | 
|
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
342  | 
apply (rule conjI, simp add: gauge_def)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
343  | 
apply (clarify)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
344  | 
apply (subst fine_rsum_const, assumption, simp)  | 
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
345  | 
done  | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
346  | 
|
| 35441 | 347  | 
lemma Integral_eq_diff_bounds: "a \<le> b \<Longrightarrow> Integral(a,b) (\<lambda>x. 1) (b - a)"  | 
348  | 
using Integral_mult_const [of a b 1] by simp  | 
|
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
349  | 
|
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
350  | 
lemma Integral_mult:  | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
351  | 
"[| a \<le> b; Integral(a,b) f k |] ==> Integral(a,b) (%x. c * f x) (c * k)"  | 
| 35441 | 352  | 
apply (auto simp add: order_le_less)  | 
353  | 
apply (cases "c = 0", simp add: Integral_zero_fun)  | 
|
354  | 
apply (rule IntegralI)  | 
|
355  | 
apply (erule_tac e="e / \<bar>c\<bar>" in IntegralE, simp add: divide_pos_pos)  | 
|
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
356  | 
apply (rule_tac x="\<delta>" in exI, clarify)  | 
| 
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
357  | 
apply (drule_tac x="D" in spec, clarify)  | 
| 31257 | 358  | 
apply (simp add: pos_less_divide_eq abs_mult [symmetric]  | 
359  | 
algebra_simps rsum_right_distrib)  | 
|
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
360  | 
done  | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
361  | 
|
| 31364 | 362  | 
lemma Integral_add:  | 
363  | 
assumes "Integral (a, b) f x1"  | 
|
364  | 
assumes "Integral (b, c) f x2"  | 
|
365  | 
assumes "a \<le> b" and "b \<le> c"  | 
|
366  | 
shows "Integral (a, c) f (x1 + x2)"  | 
|
| 35441 | 367  | 
proof (cases "a < b \<and> b < c", rule IntegralI)  | 
| 31364 | 368  | 
fix \<epsilon> :: real assume "0 < \<epsilon>"  | 
369  | 
hence "0 < \<epsilon> / 2" by auto  | 
|
370  | 
||
371  | 
assume "a < b \<and> b < c"  | 
|
372  | 
hence "a < b" and "b < c" by auto  | 
|
373  | 
||
374  | 
  obtain \<delta>1 where \<delta>1_gauge: "gauge {a..b} \<delta>1"
 | 
|
| 35441 | 375  | 
and I1: "\<And> D. fine \<delta>1 (a,b) D \<Longrightarrow> \<bar> rsum D f - x1 \<bar> < (\<epsilon> / 2)"  | 
376  | 
using IntegralE [OF `Integral (a, b) f x1` `0 < \<epsilon>/2`] by auto  | 
|
| 31364 | 377  | 
|
378  | 
  obtain \<delta>2 where \<delta>2_gauge: "gauge {b..c} \<delta>2"
 | 
|
| 35441 | 379  | 
and I2: "\<And> D. fine \<delta>2 (b,c) D \<Longrightarrow> \<bar> rsum D f - x2 \<bar> < (\<epsilon> / 2)"  | 
380  | 
using IntegralE [OF `Integral (b, c) f x2` `0 < \<epsilon>/2`] by auto  | 
|
| 31364 | 381  | 
|
382  | 
def \<delta> \<equiv> "\<lambda> x. if x < b then min (\<delta>1 x) (b - x)  | 
|
383  | 
else if x = b then min (\<delta>1 b) (\<delta>2 b)  | 
|
384  | 
else min (\<delta>2 x) (x - b)"  | 
|
385  | 
||
386  | 
  have "gauge {a..c} \<delta>"
 | 
|
387  | 
using \<delta>1_gauge \<delta>2_gauge unfolding \<delta>_def gauge_def by auto  | 
|
| 35441 | 388  | 
|
| 31364 | 389  | 
  moreover {
 | 
390  | 
fix D :: "(real \<times> real \<times> real) list"  | 
|
391  | 
assume fine: "fine \<delta> (a,c) D"  | 
|
392  | 
from fine_covers_all[OF this `a < b` `b \<le> c`]  | 
|
393  | 
obtain N where "N < length D"  | 
|
394  | 
and *: "\<forall> d t e. D ! N = (d, t, e) \<longrightarrow> d < b \<and> b \<le> e"  | 
|
395  | 
by auto  | 
|
396  | 
obtain d t e where D_eq: "D ! N = (d, t, e)" by (cases "D!N", auto)  | 
|
397  | 
with * have "d < b" and "b \<le> e" by auto  | 
|
398  | 
have in_D: "(d, t, e) \<in> set D"  | 
|
399  | 
using D_eq[symmetric] using `N < length D` by auto  | 
|
400  | 
||
401  | 
from mem_fine[OF fine in_D]  | 
|
402  | 
have "d < e" and "d \<le> t" and "t \<le> e" by auto  | 
|
403  | 
||
404  | 
have "t = b"  | 
|
405  | 
proof (rule ccontr)  | 
|
406  | 
assume "t \<noteq> b"  | 
|
407  | 
with mem_fine3[OF fine in_D] `b \<le> e` `d \<le> t` `t \<le> e` `d < b` \<delta>_def  | 
|
408  | 
show False by (cases "t < b") auto  | 
|
409  | 
qed  | 
|
410  | 
||
411  | 
let ?D1 = "take N D"  | 
|
412  | 
let ?D2 = "drop N D"  | 
|
413  | 
def D1 \<equiv> "take N D @ [(d, t, b)]"  | 
|
414  | 
def D2 \<equiv> "(if b = e then [] else [(b, t, e)]) @ drop (Suc N) D"  | 
|
415  | 
||
| 46501 | 416  | 
from hd_drop_conv_nth[OF `N < length D`]  | 
| 31364 | 417  | 
have "fst (hd ?D2) = d" using `D ! N = (d, t, e)` by auto  | 
418  | 
with fine_append_split[OF _ _ append_take_drop_id[symmetric]]  | 
|
419  | 
have fine1: "fine \<delta> (a,d) ?D1" and fine2: "fine \<delta> (d,c) ?D2"  | 
|
420  | 
using `N < length D` fine by auto  | 
|
421  | 
||
422  | 
have "fine \<delta>1 (a,b) D1" unfolding D1_def  | 
|
423  | 
proof (rule fine_append)  | 
|
424  | 
show "fine \<delta>1 (a, d) ?D1"  | 
|
425  | 
proof (rule fine1[THEN fine_\<delta>_expand])  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31366 
diff
changeset
 | 
426  | 
fix x assume "a \<le> x" "x \<le> d"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31366 
diff
changeset
 | 
427  | 
hence "x \<le> b" using `d < b` `x \<le> d` by auto  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31366 
diff
changeset
 | 
428  | 
thus "\<delta> x \<le> \<delta>1 x" unfolding \<delta>_def by auto  | 
| 31364 | 429  | 
qed  | 
430  | 
||
431  | 
have "b - d < \<delta>1 t"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31366 
diff
changeset
 | 
432  | 
using mem_fine3[OF fine in_D] \<delta>_def `b \<le> e` `t = b` by auto  | 
| 31364 | 433  | 
from `d < b` `d \<le> t` `t = b` this  | 
434  | 
show "fine \<delta>1 (d, b) [(d, t, b)]" using fine_single by auto  | 
|
435  | 
qed  | 
|
436  | 
note rsum1 = I1[OF this]  | 
|
437  | 
||
438  | 
have drop_split: "drop N D = [D ! N] @ drop (Suc N) D"  | 
|
439  | 
using nth_drop'[OF `N < length D`] by simp  | 
|
440  | 
||
441  | 
have fine2: "fine \<delta>2 (e,c) (drop (Suc N) D)"  | 
|
442  | 
proof (cases "drop (Suc N) D = []")  | 
|
443  | 
case True  | 
|
444  | 
note * = fine2[simplified drop_split True D_eq append_Nil2]  | 
|
445  | 
have "e = c" using fine_single_boundaries[OF * refl] by auto  | 
|
446  | 
thus ?thesis unfolding True using fine_Nil by auto  | 
|
447  | 
next  | 
|
448  | 
case False  | 
|
449  | 
note * = fine_append_split[OF fine2 False drop_split]  | 
|
450  | 
from fine_single_boundaries[OF *(1)]  | 
|
451  | 
have "fst (hd (drop (Suc N) D)) = e" using D_eq by auto  | 
|
452  | 
with *(2) have "fine \<delta> (e,c) (drop (Suc N) D)" by auto  | 
|
453  | 
thus ?thesis  | 
|
454  | 
proof (rule fine_\<delta>_expand)  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31366 
diff
changeset
 | 
455  | 
fix x assume "e \<le> x" and "x \<le> c"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31366 
diff
changeset
 | 
456  | 
thus "\<delta> x \<le> \<delta>2 x" using `b \<le> e` unfolding \<delta>_def by auto  | 
| 31364 | 457  | 
qed  | 
458  | 
qed  | 
|
459  | 
||
460  | 
have "fine \<delta>2 (b, c) D2"  | 
|
461  | 
proof (cases "e = b")  | 
|
462  | 
case True thus ?thesis using fine2 by (simp add: D1_def D2_def)  | 
|
463  | 
next  | 
|
464  | 
case False  | 
|
465  | 
have "e - b < \<delta>2 b"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31366 
diff
changeset
 | 
466  | 
using mem_fine3[OF fine in_D] \<delta>_def `d < b` `t = b` by auto  | 
| 31364 | 467  | 
with False `t = b` `b \<le> e`  | 
468  | 
show ?thesis using D2_def  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31366 
diff
changeset
 | 
469  | 
by (auto intro!: fine_append[OF _ fine2] fine_single  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31366 
diff
changeset
 | 
470  | 
simp del: append_Cons)  | 
| 31364 | 471  | 
qed  | 
472  | 
note rsum2 = I2[OF this]  | 
|
473  | 
||
474  | 
have "rsum D f = rsum (take N D) f + rsum [D ! N] f + rsum (drop (Suc N) D) f"  | 
|
475  | 
using rsum_append[symmetric] nth_drop'[OF `N < length D`] by auto  | 
|
476  | 
also have "\<dots> = rsum D1 f + rsum D2 f"  | 
|
| 31366 | 477  | 
by (cases "b = e", auto simp add: D1_def D2_def D_eq rsum_append algebra_simps)  | 
| 31364 | 478  | 
finally have "\<bar>rsum D f - (x1 + x2)\<bar> < \<epsilon>"  | 
479  | 
using add_strict_mono[OF rsum1 rsum2] by simp  | 
|
480  | 
}  | 
|
481  | 
  ultimately show "\<exists> \<delta>. gauge {a .. c} \<delta> \<and>
 | 
|
482  | 
(\<forall>D. fine \<delta> (a,c) D \<longrightarrow> \<bar>rsum D f - (x1 + x2)\<bar> < \<epsilon>)"  | 
|
483  | 
by blast  | 
|
484  | 
next  | 
|
485  | 
case False  | 
|
486  | 
hence "a = b \<or> b = c" using `a \<le> b` and `b \<le> c` by auto  | 
|
487  | 
thus ?thesis  | 
|
488  | 
proof (rule disjE)  | 
|
489  | 
assume "a = b" hence "x1 = 0"  | 
|
| 35441 | 490  | 
using `Integral (a, b) f x1` by simp  | 
491  | 
thus ?thesis using `a = b` `Integral (b, c) f x2` by simp  | 
|
| 31364 | 492  | 
next  | 
493  | 
assume "b = c" hence "x2 = 0"  | 
|
| 35441 | 494  | 
using `Integral (b, c) f x2` by simp  | 
495  | 
thus ?thesis using `b = c` `Integral (a, b) f x1` by simp  | 
|
| 31364 | 496  | 
qed  | 
497  | 
qed  | 
|
| 
31259
 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 
huffman 
parents: 
31257 
diff
changeset
 | 
498  | 
|
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
499  | 
text{*Fundamental theorem of calculus (Part I)*}
 | 
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
500  | 
|
| 15105 | 501  | 
text{*"Straddle Lemma" : Swartz and Thompson: AMM 95(7) 1988 *}
 | 
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
502  | 
|
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
503  | 
lemma strad1:  | 
| 53755 | 504  | 
fixes z x s e :: real  | 
505  | 
assumes P: "(\<And>z. z \<noteq> x \<Longrightarrow> \<bar>z - x\<bar> < s \<Longrightarrow> \<bar>(f z - f x) / (z - x) - f' x\<bar> < e / 2)"  | 
|
506  | 
assumes "\<bar>z - x\<bar> < s"  | 
|
507  | 
shows "\<bar>f z - f x - f' x * (z - x)\<bar> \<le> e / 2 * \<bar>z - x\<bar>"  | 
|
508  | 
proof (cases "z = x")  | 
|
509  | 
case True then show ?thesis by simp  | 
|
510  | 
next  | 
|
511  | 
case False  | 
|
512  | 
then have "inverse (z - x) * (f z - f x - f' x * (z - x)) = (f z - f x) / (z - x) - f' x"  | 
|
513  | 
apply (subst mult_commute)  | 
|
514  | 
apply (simp add: distrib_right diff_minus)  | 
|
515  | 
apply (simp add: mult_assoc divide_inverse)  | 
|
516  | 
apply (simp add: distrib_right)  | 
|
517  | 
done  | 
|
518  | 
moreover from False `\<bar>z - x\<bar> < s` have "\<bar>(f z - f x) / (z - x) - f' x\<bar> < e / 2"  | 
|
519  | 
by (rule P)  | 
|
520  | 
ultimately have "\<bar>inverse (z - x)\<bar> * (\<bar>f z - f x - f' x * (z - x)\<bar> * 2)  | 
|
521  | 
\<le> \<bar>inverse (z - x)\<bar> * (e * \<bar>z - x\<bar>)"  | 
|
522  | 
using False by (simp del: abs_inverse add: abs_mult [symmetric] ac_simps)  | 
|
523  | 
with False have "\<bar>f z - f x - f' x * (z - x)\<bar> * 2 \<le> e * \<bar>z - x\<bar>"  | 
|
524  | 
by simp  | 
|
525  | 
then show ?thesis by simp  | 
|
526  | 
qed  | 
|
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
527  | 
|
| 
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
528  | 
lemma lemma_straddle:  | 
| 31252 | 529  | 
assumes f': "\<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x)" and "0 < e"  | 
| 31253 | 530  | 
  shows "\<exists>g. gauge {a..b} g &
 | 
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
531  | 
(\<forall>x u v. a \<le> u & u \<le> x & x \<le> v & v \<le> b & (v - u) < g(x)  | 
| 
15094
 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 
paulson 
parents: 
15093 
diff
changeset
 | 
532  | 
--> \<bar>(f(v) - f(u)) - (f'(x) * (v - u))\<bar> \<le> e * (v - u))"  | 
| 31252 | 533  | 
proof -  | 
| 31253 | 534  | 
  have "\<forall>x\<in>{a..b}.
 | 
| 15360 | 535  | 
(\<exists>d > 0. \<forall>u v. u \<le> x & x \<le> v & (v - u) < d -->  | 
| 31252 | 536  | 
\<bar>(f(v) - f(u)) - (f'(x) * (v - u))\<bar> \<le> e * (v - u))"  | 
| 31253 | 537  | 
proof (clarsimp)  | 
| 31252 | 538  | 
fix x :: real assume "a \<le> x" and "x \<le> b"  | 
539  | 
with f' have "DERIV f x :> f'(x)" by simp  | 
|
540  | 
then have "\<forall>r>0. \<exists>s>0. \<forall>z. z \<noteq> x \<and> \<bar>z - x\<bar> < s \<longrightarrow> \<bar>(f z - f x) / (z - x) - f' x\<bar> < r"  | 
|
| 
31338
 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 
huffman 
parents: 
31336 
diff
changeset
 | 
541  | 
by (simp add: DERIV_iff2 LIM_eq)  | 
| 31252 | 542  | 
with `0 < e` obtain s  | 
| 53755 | 543  | 
where "\<And>z. z \<noteq> x \<Longrightarrow> \<bar>z - x\<bar> < s \<Longrightarrow> \<bar>(f z - f x) / (z - x) - f' x\<bar> < e/2" and "0 < s"  | 
| 31252 | 544  | 
by (drule_tac x="e/2" in spec, auto)  | 
| 53755 | 545  | 
with strad1 [of x s f f' e] have strad:  | 
546  | 
"\<And>z. \<bar>z - x\<bar> < s \<Longrightarrow> \<bar>f z - f x - f' x * (z - x)\<bar> \<le> e/2 * \<bar>z - x\<bar>"  | 
|
547  | 
by auto  | 
|
| 31252 | 548  | 
show "\<exists>d>0. \<forall>u v. u \<le> x \<and> x \<le> v \<and> v - u < d \<longrightarrow> \<bar>f v - f u - f' x * (v - u)\<bar> \<le> e * (v - u)"  | 
549  | 
proof (safe intro!: exI)  | 
|
550  | 
show "0 < s" by fact  | 
|
551  | 
next  | 
|
552  | 
fix u v :: real assume "u \<le> x" and "x \<le> v" and "v - u < s"  | 
|
553  | 
have "\<bar>f v - f u - f' x * (v - u)\<bar> =  | 
|
554  | 
\<bar>(f v - f x - f' x * (v - x)) + (f x - f u - f' x * (x - u))\<bar>"  | 
|
555  | 
by (simp add: right_diff_distrib)  | 
|
556  | 
also have "\<dots> \<le> \<bar>f v - f x - f' x * (v - x)\<bar> + \<bar>f x - f u - f' x * (x - u)\<bar>"  | 
|
557  | 
by (rule abs_triangle_ineq)  | 
|
558  | 
also have "\<dots> = \<bar>f v - f x - f' x * (v - x)\<bar> + \<bar>f u - f x - f' x * (u - x)\<bar>"  | 
|
559  | 
by (simp add: right_diff_distrib)  | 
|
560  | 
also have "\<dots> \<le> (e/2) * \<bar>v - x\<bar> + (e/2) * \<bar>u - x\<bar>"  | 
|
561  | 
using `u \<le> x` `x \<le> v` `v - u < s` by (intro add_mono strad, simp_all)  | 
|
562  | 
also have "\<dots> \<le> e * (v - u) / 2 + e * (v - u) / 2"  | 
|
563  | 
using `u \<le> x` `x \<le> v` `0 < e` by (intro add_mono, simp_all)  | 
|
564  | 
also have "\<dots> = e * (v - u)"  | 
|
565  | 
by simp  | 
|
566  | 
finally show "\<bar>f v - f u - f' x * (v - u)\<bar> \<le> e * (v - u)" .  | 
|
567  | 
qed  | 
|
568  | 
qed  | 
|
569  | 
thus ?thesis  | 
|
| 31253 | 570  | 
by (simp add: gauge_def) (drule bchoice, auto)  | 
| 31252 | 571  | 
qed  | 
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
572  | 
|
| 35328 | 573  | 
lemma fundamental_theorem_of_calculus:  | 
| 35441 | 574  | 
assumes "a \<le> b"  | 
575  | 
assumes f': "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> DERIV f x :> f'(x)"  | 
|
576  | 
shows "Integral (a, b) f' (f(b) - f(a))"  | 
|
577  | 
proof (cases "a = b")  | 
|
578  | 
assume "a = b" thus ?thesis by simp  | 
|
579  | 
next  | 
|
580  | 
assume "a \<noteq> b" with `a \<le> b` have "a < b" by simp  | 
|
581  | 
show ?thesis  | 
|
582  | 
proof (simp add: Integral_def2, clarify)  | 
|
583  | 
fix e :: real assume "0 < e"  | 
|
584  | 
with `a < b` have "0 < e / (b - a)" by (simp add: divide_pos_pos)  | 
|
585  | 
||
586  | 
from lemma_straddle [OF f' this]  | 
|
587  | 
    obtain \<delta> where "gauge {a..b} \<delta>"
 | 
|
588  | 
and \<delta>: "\<And>x u v. \<lbrakk>a \<le> u; u \<le> x; x \<le> v; v \<le> b; v - u < \<delta> x\<rbrakk> \<Longrightarrow>  | 
|
589  | 
\<bar>f v - f u - f' x * (v - u)\<bar> \<le> e * (v - u) / (b - a)" by auto  | 
|
590  | 
||
591  | 
have "\<forall>D. fine \<delta> (a, b) D \<longrightarrow> \<bar>rsum D f' - (f b - f a)\<bar> \<le> e"  | 
|
592  | 
proof (clarify)  | 
|
593  | 
fix D assume D: "fine \<delta> (a, b) D"  | 
|
594  | 
hence "(\<Sum>(u, x, v)\<leftarrow>D. f v - f u) = f b - f a"  | 
|
595  | 
by (rule fine_listsum_eq_diff)  | 
|
596  | 
hence "\<bar>rsum D f' - (f b - f a)\<bar> = \<bar>rsum D f' - (\<Sum>(u, x, v)\<leftarrow>D. f v - f u)\<bar>"  | 
|
597  | 
by simp  | 
|
598  | 
also have "\<dots> = \<bar>(\<Sum>(u, x, v)\<leftarrow>D. f v - f u) - rsum D f'\<bar>"  | 
|
599  | 
by (rule abs_minus_commute)  | 
|
600  | 
also have "\<dots> = \<bar>\<Sum>(u, x, v)\<leftarrow>D. (f v - f u) - f' x * (v - u)\<bar>"  | 
|
601  | 
by (simp only: rsum_def listsum_subtractf split_def)  | 
|
602  | 
also have "\<dots> \<le> (\<Sum>(u, x, v)\<leftarrow>D. \<bar>(f v - f u) - f' x * (v - u)\<bar>)"  | 
|
603  | 
by (rule ord_le_eq_trans [OF listsum_abs], simp add: o_def split_def)  | 
|
604  | 
also have "\<dots> \<le> (\<Sum>(u, x, v)\<leftarrow>D. (e / (b - a)) * (v - u))"  | 
|
605  | 
apply (rule listsum_mono, clarify, rename_tac u x v)  | 
|
606  | 
using D apply (simp add: \<delta> mem_fine mem_fine2 mem_fine3)  | 
|
607  | 
done  | 
|
608  | 
also have "\<dots> = e"  | 
|
609  | 
using fine_listsum_eq_diff [OF D, where f="\<lambda>x. x"]  | 
|
610  | 
unfolding split_def listsum_const_mult  | 
|
611  | 
using `a < b` by simp  | 
|
612  | 
finally show "\<bar>rsum D f' - (f b - f a)\<bar> \<le> e" .  | 
|
613  | 
qed  | 
|
614  | 
||
615  | 
    with `gauge {a..b} \<delta>`
 | 
|
616  | 
    show "\<exists>\<delta>. gauge {a..b} \<delta> \<and> (\<forall>D. fine \<delta> (a, b) D \<longrightarrow> \<bar>rsum D f' - (f b - f a)\<bar> \<le> e)"
 | 
|
617  | 
by auto  | 
|
618  | 
qed  | 
|
619  | 
qed  | 
|
| 13958 | 620  | 
|
| 
15093
 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 
paulson 
parents: 
13958 
diff
changeset
 | 
621  | 
end  |