| author | wenzelm | 
| Mon, 18 Sep 2006 19:12:46 +0200 | |
| changeset 20574 | a10885a269cb | 
| parent 20563 | 44eda2314aab | 
| child 20792 | add17d26151b | 
| permissions | -rw-r--r-- | 
| 15944 | 1 | (* ID : $Id$ | 
| 13958 | 2 | Author : Jacques D. Fleuriot | 
| 3 | Copyright : 2000 University of Edinburgh | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 4 | Conversion to Isar and new proofs by Lawrence C Paulson, 2004 | 
| 13958 | 5 | *) | 
| 6 | ||
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 7 | header{*Theory of Integration*}
 | 
| 13958 | 8 | |
| 15131 | 9 | theory Integration | 
| 15140 | 10 | imports MacLaurin | 
| 15131 | 11 | begin | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 12 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 13 | text{*We follow John Harrison in formalizing the Gauge integral.*}
 | 
| 13958 | 14 | |
| 19765 | 15 | definition | 
| 13958 | 16 | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 17 |   --{*Partitions and tagged partitions etc.*}
 | 
| 13958 | 18 | |
| 19 | partition :: "[(real*real),nat => real] => bool" | |
| 19765 | 20 | "partition = (%(a,b) D. D 0 = a & | 
| 15360 | 21 | (\<exists>N. (\<forall>n < N. D(n) < D(Suc n)) & | 
| 19765 | 22 | (\<forall>n \<ge> N. D(n) = b)))" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 23 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 24 | psize :: "(nat => real) => nat" | 
| 19765 | 25 | "psize D = (SOME N. (\<forall>n < N. D(n) < D(Suc n)) & | 
| 26 | (\<forall>n \<ge> N. D(n) = D(N)))" | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 27 | |
| 13958 | 28 | tpart :: "[(real*real),((nat => real)*(nat =>real))] => bool" | 
| 19765 | 29 | "tpart = (%(a,b) (D,p). partition(a,b) D & | 
| 30 | (\<forall>n. D(n) \<le> p(n) & p(n) \<le> D(Suc n)))" | |
| 13958 | 31 | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 32 |   --{*Gauges and gauge-fine divisions*}
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 33 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 34 | gauge :: "[real => bool, real => real] => bool" | 
| 19765 | 35 | "gauge E g = (\<forall>x. E x --> 0 < g(x))" | 
| 13958 | 36 | |
| 37 | fine :: "[real => real, ((nat => real)*(nat => real))] => bool" | |
| 19765 | 38 | "fine = (%g (D,p). \<forall>n. n < (psize D) --> D(Suc n) - D(n) < g(p n))" | 
| 13958 | 39 | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 40 |   --{*Riemann sum*}
 | 
| 13958 | 41 | |
| 42 | rsum :: "[((nat=>real)*(nat=>real)),real=>real] => real" | |
| 19765 | 43 | "rsum = (%(D,p) f. \<Sum>n=0..<psize(D). f(p n) * (D(Suc n) - D(n)))" | 
| 13958 | 44 | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 45 |   --{*Gauge integrability (definite)*}
 | 
| 13958 | 46 | |
| 19765 | 47 | Integral :: "[(real*real),real=>real,real] => bool" | 
| 48 | "Integral = (%(a,b) f k. \<forall>e > 0. | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 49 | (\<exists>g. gauge(%x. a \<le> x & x \<le> b) g & | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 50 | (\<forall>D p. tpart(a,b) (D,p) & fine(g)(D,p) --> | 
| 19765 | 51 | \<bar>rsum(D,p) f - k\<bar> < e)))" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 52 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 53 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 54 | lemma partition_zero [simp]: "a = b ==> psize (%n. if n = 0 then a else b) = 0" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 55 | by (auto simp add: psize_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 56 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 57 | lemma partition_one [simp]: "a < b ==> psize (%n. if n = 0 then a else b) = 1" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 58 | apply (simp add: psize_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 59 | apply (rule some_equality, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 60 | apply (drule_tac x = 1 in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 61 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 62 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 63 | lemma partition_single [simp]: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 64 | "a \<le> b ==> partition(a,b)(%n. if n = 0 then a else b)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 65 | by (auto simp add: partition_def order_le_less) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 66 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 67 | lemma partition_lhs: "partition(a,b) D ==> (D(0) = a)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 68 | by (simp add: partition_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 69 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 70 | lemma partition: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 71 | "(partition(a,b) D) = | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 72 | ((D 0 = a) & | 
| 15360 | 73 | (\<forall>n < psize D. D n < D(Suc n)) & | 
| 74 | (\<forall>n \<ge> psize D. D n = b))" | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 75 | apply (simp add: partition_def, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 76 | apply (subgoal_tac [!] "psize D = N", auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 77 | apply (simp_all (no_asm) add: psize_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 78 | apply (rule_tac [!] some_equality, blast) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 79 | prefer 2 apply blast | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 80 | apply (rule_tac [!] ccontr) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 81 | apply (simp_all add: linorder_neq_iff, safe) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 82 | apply (drule_tac x = Na in spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 83 | apply (rotate_tac 3) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 84 | apply (drule_tac x = "Suc Na" in spec, simp) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 85 | apply (rotate_tac 2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 86 | apply (drule_tac x = N in spec, simp) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 87 | apply (drule_tac x = Na in spec) | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 88 | apply (drule_tac x = "Suc Na" and P = "%n. Na\<le>n \<longrightarrow> D n = D Na" in spec, auto) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 89 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 90 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 91 | lemma partition_rhs: "partition(a,b) D ==> (D(psize D) = b)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 92 | by (simp add: partition) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 93 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 94 | lemma partition_rhs2: "[|partition(a,b) D; psize D \<le> n |] ==> (D n = b)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 95 | by (simp add: partition) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 96 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 97 | lemma lemma_partition_lt_gen [rule_format]: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 98 | "partition(a,b) D & m + Suc d \<le> n & n \<le> (psize D) --> D(m) < D(m + Suc d)" | 
| 15251 | 99 | apply (induct "d", auto simp add: partition) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 100 | apply (blast dest: Suc_le_lessD intro: less_le_trans order_less_trans) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 101 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 102 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 103 | lemma less_eq_add_Suc: "m < n ==> \<exists>d. n = m + Suc d" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 104 | by (auto simp add: less_iff_Suc_add) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 105 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 106 | lemma partition_lt_gen: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 107 | "[|partition(a,b) D; m < n; n \<le> (psize D)|] ==> D(m) < D(n)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 108 | by (auto dest: less_eq_add_Suc intro: lemma_partition_lt_gen) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 109 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 110 | lemma partition_lt: "partition(a,b) D ==> n < (psize D) ==> D(0) < D(Suc n)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 111 | apply (induct "n") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 112 | apply (auto simp add: partition) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 113 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 114 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 115 | lemma partition_le: "partition(a,b) D ==> a \<le> b" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 116 | apply (frule partition [THEN iffD1], safe) | 
| 15219 | 117 | apply (drule_tac x = "psize D" and P="%n. psize D \<le> n --> ?P n" in spec, safe) | 
| 118 | apply (case_tac "psize D = 0") | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 119 | apply (drule_tac [2] n = "psize D - 1" in partition_lt, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 120 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 121 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 122 | lemma partition_gt: "[|partition(a,b) D; n < (psize D)|] ==> D(n) < D(psize D)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 123 | by (auto intro: partition_lt_gen) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 124 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 125 | lemma partition_eq: "partition(a,b) D ==> ((a = b) = (psize D = 0))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 126 | apply (frule partition [THEN iffD1], safe) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 127 | apply (rotate_tac 2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 128 | apply (drule_tac x = "psize D" in spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 129 | apply (rule ccontr) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 130 | apply (drule_tac n = "psize D - 1" in partition_lt) | 
| 15219 | 131 | apply auto | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 132 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 133 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 134 | lemma partition_lb: "partition(a,b) D ==> a \<le> D(r)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 135 | apply (frule partition [THEN iffD1], safe) | 
| 15251 | 136 | apply (induct "r") | 
| 137 | apply (cut_tac [2] y = "Suc r" and x = "psize D" in linorder_le_less_linear) | |
| 138 | apply (auto intro: partition_le) | |
| 139 | apply (drule_tac x = r in spec) | |
| 140 | apply arith; | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 141 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 142 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 143 | lemma partition_lb_lt: "[| partition(a,b) D; psize D ~= 0 |] ==> a < D(Suc n)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 144 | apply (rule_tac t = a in partition_lhs [THEN subst], assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 145 | apply (cut_tac x = "Suc n" and y = "psize D" in linorder_le_less_linear) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 146 | apply (frule partition [THEN iffD1], safe) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 147 | apply (blast intro: partition_lt less_le_trans) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 148 | apply (rotate_tac 3) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 149 | apply (drule_tac x = "Suc n" in spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 150 | apply (erule impE) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 151 | apply (erule less_imp_le) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 152 | apply (frule partition_rhs) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 153 | apply (drule partition_gt, assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 154 | apply (simp (no_asm_simp)) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 155 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 156 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 157 | lemma partition_ub: "partition(a,b) D ==> D(r) \<le> b" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 158 | apply (frule partition [THEN iffD1]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 159 | apply (cut_tac x = "psize D" and y = r in linorder_le_less_linear, safe, blast) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 160 | apply (subgoal_tac "\<forall>x. D ((psize D) - x) \<le> b") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 161 | apply (rotate_tac 4) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 162 | apply (drule_tac x = "psize D - r" in spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 163 | apply (subgoal_tac "psize D - (psize D - r) = r") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 164 | apply simp | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 165 | apply arith | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 166 | apply safe | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 167 | apply (induct_tac "x") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 168 | apply (simp (no_asm), blast) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 169 | apply (case_tac "psize D - Suc n = 0") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 170 | apply (erule_tac V = "\<forall>n. psize D \<le> n --> D n = b" in thin_rl) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 171 | apply (simp (no_asm_simp) add: partition_le) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 172 | apply (rule order_trans) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 173 | prefer 2 apply assumption | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 174 | apply (subgoal_tac "psize D - n = Suc (psize D - Suc n)") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 175 | prefer 2 apply arith | 
| 20432 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20256diff
changeset | 176 | apply (drule_tac x = "psize D - Suc n" in spec, simp) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 177 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 178 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 179 | lemma partition_ub_lt: "[| partition(a,b) D; n < psize D |] ==> D(n) < b" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 180 | by (blast intro: partition_rhs [THEN subst] partition_gt) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 181 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 182 | lemma lemma_partition_append1: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 183 | "[| partition (a, b) D1; partition (b, c) D2 |] | 
| 15360 | 184 | ==> (\<forall>n < psize D1 + psize D2. | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 185 | (if n < psize D1 then D1 n else D2 (n - psize D1)) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 186 | < (if Suc n < psize D1 then D1 (Suc n) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 187 | else D2 (Suc n - psize D1))) & | 
| 15360 | 188 | (\<forall>n \<ge> psize D1 + psize D2. | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 189 | (if n < psize D1 then D1 n else D2 (n - psize D1)) = | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 190 | (if psize D1 + psize D2 < psize D1 then D1 (psize D1 + psize D2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 191 | else D2 (psize D1 + psize D2 - psize D1)))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 192 | apply (auto intro: partition_lt_gen) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 193 | apply (subgoal_tac "psize D1 = Suc n") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 194 | apply (auto intro!: partition_lt_gen simp add: partition_lhs partition_ub_lt) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 195 | apply (auto intro!: partition_rhs2 simp add: partition_rhs | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 196 | split: nat_diff_split) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 197 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 198 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 199 | lemma lemma_psize1: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 200 | "[| partition (a, b) D1; partition (b, c) D2; N < psize D1 |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 201 | ==> D1(N) < D2 (psize D2)" | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 202 | apply (rule_tac y = "D1 (psize D1)" in order_less_le_trans) | 
| 15219 | 203 | apply (erule partition_gt) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 204 | apply (auto simp add: partition_rhs partition_le) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 205 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 206 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 207 | lemma lemma_partition_append2: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 208 | "[| partition (a, b) D1; partition (b, c) D2 |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 209 | ==> psize (%n. if n < psize D1 then D1 n else D2 (n - psize D1)) = | 
| 15219 | 210 | psize D1 + psize D2" | 
| 211 | apply (unfold psize_def | |
| 212 | [of "%n. if n < psize D1 then D1 n else D2 (n - psize D1)"]) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 213 | apply (rule some1_equality) | 
| 15219 | 214 | prefer 2 apply (blast intro!: lemma_partition_append1) | 
| 215 | apply (rule ex1I, rule lemma_partition_append1) | |
| 216 | apply (simp_all split: split_if_asm) | |
| 217 |  txt{*The case @{term "N < psize D1"}*} 
 | |
| 218 | apply (drule_tac x = "psize D1 + psize D2" and P="%n. ?P n & ?Q n" in spec) | |
| 219 | apply (force dest: lemma_psize1) | |
| 220 | apply (rule order_antisym); | |
| 221 |  txt{*The case @{term "psize D1 \<le> N"}: 
 | |
| 222 |        proving @{term "N \<le> psize D1 + psize D2"}*}
 | |
| 223 | apply (drule_tac x = "psize D1 + psize D2" in spec) | |
| 224 | apply (simp add: partition_rhs2) | |
| 225 | apply (case_tac "N - psize D1 < psize D2") | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 226 | prefer 2 apply arith | 
| 15219 | 227 |  txt{*Proving @{term "psize D1 + psize D2 \<le> N"}*}
 | 
| 228 | apply (drule_tac x = "psize D1 + psize D2" and P="%n. N\<le>n --> ?P n" in spec, simp) | |
| 229 | apply (drule_tac a = b and b = c in partition_gt, assumption, simp) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 230 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 231 | |
| 15219 | 232 | lemma tpart_eq_lhs_rhs: "[|psize D = 0; tpart(a,b) (D,p)|] ==> a = b" | 
| 233 | by (auto simp add: tpart_def partition_eq) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 234 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 235 | lemma tpart_partition: "tpart(a,b) (D,p) ==> partition(a,b) D" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 236 | by (simp add: tpart_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 237 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 238 | lemma partition_append: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 239 | "[| tpart(a,b) (D1,p1); fine(g) (D1,p1); | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 240 | tpart(b,c) (D2,p2); fine(g) (D2,p2) |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 241 | ==> \<exists>D p. tpart(a,c) (D,p) & fine(g) (D,p)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 242 | apply (rule_tac x = "%n. if n < psize D1 then D1 n else D2 (n - psize D1)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 243 | in exI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 244 | apply (rule_tac x = "%n. if n < psize D1 then p1 n else p2 (n - psize D1)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 245 | in exI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 246 | apply (case_tac "psize D1 = 0") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 247 | apply (auto dest: tpart_eq_lhs_rhs) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 248 | prefer 2 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 249 | apply (simp add: fine_def | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 250 | lemma_partition_append2 [OF tpart_partition tpart_partition]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 251 |   --{*But must not expand @{term fine} in other subgoals*}
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 252 | apply auto | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 253 | apply (subgoal_tac "psize D1 = Suc n") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 254 | prefer 2 apply arith | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 255 | apply (drule tpart_partition [THEN partition_rhs]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 256 | apply (drule tpart_partition [THEN partition_lhs]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 257 | apply (auto split: nat_diff_split) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 258 | apply (auto simp add: tpart_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 259 | defer 1 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 260 | apply (subgoal_tac "psize D1 = Suc n") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 261 | prefer 2 apply arith | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 262 | apply (drule partition_rhs) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 263 | apply (drule partition_lhs, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 264 | apply (simp split: nat_diff_split) | 
| 15944 | 265 | apply (subst partition) | 
| 266 | apply (subst (1 2) lemma_partition_append2, assumption+) | |
| 267 | apply (rule conjI) | |
| 268 | apply (simp add: partition_lhs) | |
| 269 | apply (drule lemma_partition_append1) | |
| 270 | apply assumption; | |
| 271 | apply (simp add: partition_rhs) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 272 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 273 | |
| 15481 | 274 | |
| 15219 | 275 | text{*We can always find a division that is fine wrt any gauge*}
 | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 276 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 277 | lemma partition_exists: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 278 | "[| a \<le> b; gauge(%x. a \<le> x & x \<le> b) g |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 279 | ==> \<exists>D p. tpart(a,b) (D,p) & fine g (D,p)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 280 | apply (cut_tac P = "%(u,v). a \<le> u & v \<le> b --> | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 281 | (\<exists>D p. tpart (u,v) (D,p) & fine (g) (D,p))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 282 | in lemma_BOLZANO2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 283 | apply safe | 
| 15219 | 284 | apply (blast intro: order_trans)+ | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 285 | apply (auto intro: partition_append) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 286 | apply (case_tac "a \<le> x & x \<le> b") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 287 | apply (rule_tac [2] x = 1 in exI, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 288 | apply (rule_tac x = "g x" in exI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 289 | apply (auto simp add: gauge_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 290 | apply (rule_tac x = "%n. if n = 0 then aa else ba" in exI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 291 | apply (rule_tac x = "%n. if n = 0 then x else ba" in exI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 292 | apply (auto simp add: tpart_def fine_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 293 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 294 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 295 | text{*Lemmas about combining gauges*}
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 296 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 297 | lemma gauge_min: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 298 | "[| gauge(E) g1; gauge(E) g2 |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 299 | ==> gauge(E) (%x. if g1(x) < g2(x) then g1(x) else g2(x))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 300 | by (simp add: gauge_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 301 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 302 | lemma fine_min: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 303 | "fine (%x. if g1(x) < g2(x) then g1(x) else g2(x)) (D,p) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 304 | ==> fine(g1) (D,p) & fine(g2) (D,p)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 305 | by (auto simp add: fine_def split: split_if_asm) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 306 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 307 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 308 | text{*The integral is unique if it exists*}
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 309 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 310 | lemma Integral_unique: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 311 | "[| a \<le> b; Integral(a,b) f k1; Integral(a,b) f k2 |] ==> k1 = k2" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 312 | apply (simp add: Integral_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 313 | apply (drule_tac x = "\<bar>k1 - k2\<bar> /2" in spec)+ | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 314 | apply auto | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 315 | apply (drule gauge_min, assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 316 | apply (drule_tac g = "%x. if g x < ga x then g x else ga x" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 317 | in partition_exists, assumption, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 318 | apply (drule fine_min) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 319 | apply (drule spec)+ | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 320 | apply auto | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 321 | apply (subgoal_tac "\<bar>(rsum (D,p) f - k2) - (rsum (D,p) f - k1)\<bar> < \<bar>k1 - k2\<bar>") | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 322 | apply arith | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 323 | apply (drule add_strict_mono, assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 324 | apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
16924diff
changeset | 325 | mult_less_cancel_right) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 326 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 327 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 328 | lemma Integral_zero [simp]: "Integral(a,a) f 0" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 329 | apply (auto simp add: Integral_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 330 | apply (rule_tac x = "%x. 1" in exI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 331 | apply (auto dest: partition_eq simp add: gauge_def tpart_def rsum_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 332 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 333 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 334 | lemma sumr_partition_eq_diff_bounds [simp]: | 
| 15539 | 335 | "(\<Sum>n=0..<m. D (Suc n) - D n::real) = D(m) - D 0" | 
| 15251 | 336 | by (induct "m", auto) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 337 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 338 | lemma Integral_eq_diff_bounds: "a \<le> b ==> Integral(a,b) (%x. 1) (b - a)" | 
| 15219 | 339 | apply (auto simp add: order_le_less rsum_def Integral_def) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 340 | apply (rule_tac x = "%x. b - a" in exI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 341 | apply (auto simp add: gauge_def abs_interval_iff tpart_def partition) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 342 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 343 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 344 | lemma Integral_mult_const: "a \<le> b ==> Integral(a,b) (%x. c) (c*(b - a))" | 
| 15219 | 345 | apply (auto simp add: order_le_less rsum_def Integral_def) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 346 | apply (rule_tac x = "%x. b - a" in exI) | 
| 19279 | 347 | apply (auto simp add: setsum_right_distrib [symmetric] gauge_def abs_interval_iff | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 348 | right_diff_distrib [symmetric] partition tpart_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 349 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 350 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 351 | lemma Integral_mult: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 352 | "[| a \<le> b; Integral(a,b) f k |] ==> Integral(a,b) (%x. c * f x) (c * k)" | 
| 15221 | 353 | apply (auto simp add: order_le_less | 
| 354 | dest: Integral_unique [OF order_refl Integral_zero]) | |
| 19279 | 355 | apply (auto simp add: rsum_def Integral_def setsum_right_distrib[symmetric] mult_assoc) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 356 | apply (rule_tac a2 = c in abs_ge_zero [THEN real_le_imp_less_or_eq, THEN disjE]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 357 | prefer 2 apply force | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 358 | apply (drule_tac x = "e/abs c" in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 359 | apply (simp add: zero_less_mult_iff divide_inverse) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 360 | apply (rule exI, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 361 | apply (drule spec)+ | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 362 | apply auto | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 363 | apply (rule_tac z1 = "inverse (abs c)" in real_mult_less_iff1 [THEN iffD1]) | 
| 16924 | 364 | apply (auto simp add: abs_mult divide_inverse [symmetric] right_diff_distrib [symmetric]) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 365 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 366 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 367 | text{*Fundamental theorem of calculus (Part I)*}
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 368 | |
| 15105 | 369 | text{*"Straddle Lemma" : Swartz and Thompson: AMM 95(7) 1988 *}
 | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 370 | |
| 16588 | 371 | lemma choiceP: "\<forall>x. P(x) --> (\<exists>y. Q x y) ==> \<exists>f. (\<forall>x. P(x) --> Q x (f x))" | 
| 372 | by (insert bchoice [of "Collect P" Q], simp) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 373 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 374 | (*UNUSED | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 375 | lemma choice2: "\<forall>x. (\<exists>y. R(y) & (\<exists>z. Q x y z)) ==> | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 376 | \<exists>f fa. (\<forall>x. R(f x) & Q x (f x) (fa x))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 377 | *) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 378 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 379 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 380 | (* new simplifications e.g. (y < x/n) = (y * n < x) are a real nuisance | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 381 | they break the original proofs and make new proofs longer!*) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 382 | lemma strad1: | 
| 20563 | 383 | "\<lbrakk>\<forall>xa::real. xa \<noteq> x \<and> \<bar>xa - x\<bar> < s \<longrightarrow> | 
| 384 | \<bar>(f xa - f x) / (xa - x) - f' x\<bar> * 2 < e; | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 385 | 0 < e; a \<le> x; x \<le> b; 0 < s\<rbrakk> | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 386 | \<Longrightarrow> \<forall>z. \<bar>z - x\<bar> < s -->\<bar>f z - f x - f' x * (z - x)\<bar> * 2 \<le> e * \<bar>z - x\<bar>" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 387 | apply auto | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 388 | apply (case_tac "0 < \<bar>z - x\<bar>") | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 389 | prefer 2 apply (simp add: zero_less_abs_iff) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 390 | apply (drule_tac x = z in spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 391 | apply (rule_tac z1 = "\<bar>inverse (z - x)\<bar>" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 392 | in real_mult_le_cancel_iff2 [THEN iffD1]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 393 | apply simp | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 394 | apply (simp del: abs_inverse abs_mult add: abs_mult [symmetric] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 395 | mult_assoc [symmetric]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 396 | apply (subgoal_tac "inverse (z - x) * (f z - f x - f' x * (z - x)) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 397 | = (f z - f x) / (z - x) - f' x") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 398 | apply (simp add: abs_mult [symmetric] mult_ac diff_minus) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 399 | apply (subst mult_commute) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 400 | apply (simp add: left_distrib diff_minus) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 401 | apply (simp add: mult_assoc divide_inverse) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 402 | apply (simp add: left_distrib) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 403 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 404 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 405 | lemma lemma_straddle: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 406 | "[| \<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x); 0 < e |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 407 | ==> \<exists>g. gauge(%x. a \<le> x & x \<le> b) g & | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 408 | (\<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: 
15093diff
changeset | 409 | --> \<bar>(f(v) - f(u)) - (f'(x) * (v - u))\<bar> \<le> e * (v - u))" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 410 | apply (simp add: gauge_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 411 | apply (subgoal_tac "\<forall>x. a \<le> x & x \<le> b --> | 
| 15360 | 412 | (\<exists>d > 0. \<forall>u v. u \<le> x & x \<le> v & (v - u) < d --> | 
| 413 | \<bar>(f (v) - f (u)) - (f' (x) * (v - u))\<bar> \<le> e * (v - u))") | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 414 | apply (drule choiceP, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 415 | apply (drule spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 416 | apply (auto simp add: DERIV_iff2 LIM_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 417 | apply (drule_tac x = "e/2" in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 418 | apply (frule strad1, assumption+) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 419 | apply (rule_tac x = s in exI, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 420 | apply (rule_tac x = u and y = v in linorder_cases, auto) | 
| 15219 | 421 | apply (rule_tac y = "\<bar>(f (v) - f (x)) - (f' (x) * (v - x))\<bar> + | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 422 | \<bar>(f (x) - f (u)) - (f' (x) * (x - u))\<bar>" | 
| 15219 | 423 | in order_trans) | 
| 424 | apply (rule abs_triangle_ineq [THEN [2] order_trans]) | |
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19765diff
changeset | 425 | apply (simp add: right_diff_distrib) | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 426 | apply (rule_tac t = "e* (v - u)" in real_sum_of_halves [THEN subst]) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 427 | apply (rule add_mono) | 
| 15219 | 428 | apply (rule_tac y = "(e/2) * \<bar>v - x\<bar>" in order_trans) | 
| 15229 | 429 | prefer 2 apply simp | 
| 20563 | 430 | apply (erule_tac [!] V= "\<forall>x'. x' ~= x & \<bar>x' - x\<bar> < s --> ?P x'" in thin_rl) | 
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 431 | apply (drule_tac x = v in spec, simp add: times_divide_eq) | 
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16588diff
changeset | 432 | apply (drule_tac x = u in spec, auto) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 433 | apply (subgoal_tac "\<bar>f u - f x - f' x * (u - x)\<bar> = \<bar>f x - f u - f' x * (x - u)\<bar>") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 434 | apply (rule order_trans) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 435 | apply (auto simp add: abs_le_interval_iff) | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19765diff
changeset | 436 | apply (simp add: right_diff_distrib) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 437 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 438 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 439 | lemma FTC1: "[|a \<le> b; \<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x) |] | 
| 15219 | 440 | ==> Integral(a,b) f' (f(b) - f(a))" | 
| 441 | apply (drule order_le_imp_less_or_eq, auto) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 442 | apply (auto simp add: Integral_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 443 | apply (rule ccontr) | 
| 15360 | 444 | apply (subgoal_tac "\<forall>e > 0. \<exists>g. gauge (%x. a \<le> x & x \<le> b) g & (\<forall>D p. tpart (a, b) (D, p) & fine g (D, p) --> \<bar>rsum (D, p) f' - (f b - f a)\<bar> \<le> e)") | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 445 | apply (rotate_tac 3) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 446 | apply (drule_tac x = "e/2" in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 447 | apply (drule spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 448 | apply ((drule spec)+, auto) | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 449 | apply (drule_tac e = "ea/ (b - a)" in lemma_straddle) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 450 | apply (auto simp add: zero_less_divide_iff) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 451 | apply (rule exI) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 452 | apply (auto simp add: tpart_def rsum_def) | 
| 15539 | 453 | apply (subgoal_tac "(\<Sum>n=0..<psize D. f(D(Suc n)) - f(D n)) = f b - f a") | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 454 | prefer 2 | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 455 | apply (cut_tac D = "%n. f (D n)" and m = "psize D" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 456 | in sumr_partition_eq_diff_bounds) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 457 | apply (simp add: partition_lhs partition_rhs) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 458 | apply (drule sym, simp) | 
| 15536 | 459 | apply (simp (no_asm) add: setsum_subtractf[symmetric]) | 
| 460 | apply (rule setsum_abs [THEN order_trans]) | |
| 15539 | 461 | apply (subgoal_tac "ea = (\<Sum>n=0..<psize D. (ea / (b - a)) * (D (Suc n) - (D n)))") | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 462 | apply (simp add: abs_minus_commute) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 463 | apply (rule_tac t = ea in ssubst, assumption) | 
| 15539 | 464 | apply (rule setsum_mono) | 
| 19279 | 465 | apply (rule_tac [2] setsum_right_distrib [THEN subst]) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 466 | apply (auto simp add: partition_rhs partition_lhs partition_lb partition_ub | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 467 | fine_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 468 | done | 
| 13958 | 469 | |
| 470 | ||
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 471 | lemma Integral_subst: "[| Integral(a,b) f k1; k2=k1 |] ==> Integral(a,b) f k2" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 472 | by simp | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 473 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 474 | lemma Integral_add: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 475 | "[| a \<le> b; b \<le> c; Integral(a,b) f' k1; Integral(b,c) f' k2; | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 476 | \<forall>x. a \<le> x & x \<le> c --> DERIV f x :> f' x |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 477 | ==> Integral(a,c) f' (k1 + k2)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 478 | apply (rule FTC1 [THEN Integral_subst], auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 479 | apply (frule FTC1, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 480 | apply (frule_tac a = b in FTC1, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 481 | apply (drule_tac x = x in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 482 | apply (drule_tac ?k2.0 = "f b - f a" in Integral_unique) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 483 | apply (drule_tac [3] ?k2.0 = "f c - f b" in Integral_unique, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 484 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 485 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 486 | lemma partition_psize_Least: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 487 | "partition(a,b) D ==> psize D = (LEAST n. D(n) = b)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 488 | apply (auto intro!: Least_equality [symmetric] partition_rhs) | 
| 15219 | 489 | apply (auto dest: partition_ub_lt simp add: linorder_not_less [symmetric]) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 490 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 491 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 492 | lemma lemma_partition_bounded: "partition (a, c) D ==> ~ (\<exists>n. c < D(n))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 493 | apply safe | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 494 | apply (drule_tac r = n in partition_ub, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 495 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 496 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 497 | lemma lemma_partition_eq: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 498 | "partition (a, c) D ==> D = (%n. if D n < c then D n else c)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 499 | apply (rule ext, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 500 | apply (auto dest!: lemma_partition_bounded) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 501 | apply (drule_tac x = n in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 502 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 503 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 504 | lemma lemma_partition_eq2: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 505 | "partition (a, c) D ==> D = (%n. if D n \<le> c then D n else c)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 506 | apply (rule ext, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 507 | apply (auto dest!: lemma_partition_bounded) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 508 | apply (drule_tac x = n in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 509 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 510 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 511 | lemma partition_lt_Suc: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 512 | "[| partition(a,b) D; n < psize D |] ==> D n < D (Suc n)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 513 | by (auto simp add: partition) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 514 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 515 | lemma tpart_tag_eq: "tpart(a,c) (D,p) ==> p = (%n. if D n < c then p n else c)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 516 | apply (rule ext) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 517 | apply (auto simp add: tpart_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 518 | apply (drule linorder_not_less [THEN iffD1]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 519 | apply (drule_tac r = "Suc n" in partition_ub) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 520 | apply (drule_tac x = n in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 521 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 522 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 523 | subsection{*Lemmas for Additivity Theorem of Gauge Integral*}
 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 524 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 525 | lemma lemma_additivity1: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 526 | "[| a \<le> D n; D n < b; partition(a,b) D |] ==> n < psize D" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 527 | by (auto simp add: partition linorder_not_less [symmetric]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 528 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 529 | lemma lemma_additivity2: "[| a \<le> D n; partition(a,D n) D |] ==> psize D \<le> n" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 530 | apply (rule ccontr, drule not_leE) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 531 | apply (frule partition [THEN iffD1], safe) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 532 | apply (frule_tac r = "Suc n" in partition_ub) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 533 | apply (auto dest!: spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 534 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 535 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 536 | lemma partition_eq_bound: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 537 | "[| partition(a,b) D; psize D < m |] ==> D(m) = D(psize D)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 538 | by (auto simp add: partition) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 539 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 540 | lemma partition_ub2: "[| partition(a,b) D; psize D < m |] ==> D(r) \<le> D(m)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 541 | by (simp add: partition partition_ub) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 542 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 543 | lemma tag_point_eq_partition_point: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 544 | "[| tpart(a,b) (D,p); psize D \<le> m |] ==> p(m) = D(m)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 545 | apply (simp add: tpart_def, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 546 | apply (drule_tac x = m in spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 547 | apply (auto simp add: partition_rhs2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 548 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 549 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 550 | lemma partition_lt_cancel: "[| partition(a,b) D; D m < D n |] ==> m < n" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 551 | apply (cut_tac m = n and n = "psize D" in less_linear, auto) | 
| 15219 | 552 | apply (cut_tac m = m and n = n in less_linear) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 553 | apply (cut_tac m = m and n = "psize D" in less_linear) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 554 | apply (auto dest: partition_gt) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 555 | apply (drule_tac n = m in partition_lt_gen, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 556 | apply (frule partition_eq_bound) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 557 | apply (drule_tac [2] partition_gt, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 558 | apply (rule ccontr, drule leI, drule le_imp_less_or_eq) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 559 | apply (auto dest: partition_eq_bound) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 560 | apply (rule ccontr, drule leI, drule le_imp_less_or_eq) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 561 | apply (frule partition_eq_bound, assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 562 | apply (drule_tac m = m in partition_eq_bound, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 563 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 564 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 565 | lemma lemma_additivity4_psize_eq: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 566 | "[| a \<le> D n; D n < b; partition (a, b) D |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 567 | ==> psize (%x. if D x < D n then D(x) else D n) = n" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 568 | apply (unfold psize_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 569 | apply (frule lemma_additivity1) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 570 | apply (assumption, assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 571 | apply (rule some_equality) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 572 | apply (auto intro: partition_lt_Suc) | 
| 15219 | 573 | apply (drule_tac n = n in partition_lt_gen, assumption) | 
| 574 | apply (arith, arith) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 575 | apply (cut_tac m = na and n = "psize D" in less_linear) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 576 | apply (auto dest: partition_lt_cancel) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 577 | apply (rule_tac x=N and y=n in linorder_cases) | 
| 15219 | 578 | apply (drule_tac x = n and P="%m. N \<le> m --> ?f m = ?g m" in spec, simp) | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19765diff
changeset | 579 | apply (drule_tac n = n in partition_lt_gen, auto) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 580 | apply (drule_tac x = n in spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 581 | apply (simp split: split_if_asm) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 582 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 583 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 584 | lemma lemma_psize_left_less_psize: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 585 | "partition (a, b) D | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 586 | ==> psize (%x. if D x < D n then D(x) else D n) \<le> psize D" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 587 | apply (frule_tac r = n in partition_ub) | 
| 15219 | 588 | apply (drule_tac x = "D n" in order_le_imp_less_or_eq) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 589 | apply (auto simp add: lemma_partition_eq [symmetric]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 590 | apply (frule_tac r = n in partition_lb) | 
| 15219 | 591 | apply (drule (2) lemma_additivity4_psize_eq) | 
| 592 | apply (rule ccontr, auto) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 593 | apply (frule_tac not_leE [THEN [2] partition_eq_bound]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 594 | apply (auto simp add: partition_rhs) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 595 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 596 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 597 | lemma lemma_psize_left_less_psize2: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 598 | "[| partition(a,b) D; na < psize (%x. if D x < D n then D(x) else D n) |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 599 | ==> na < psize D" | 
| 15219 | 600 | by (erule lemma_psize_left_less_psize [THEN [2] less_le_trans]) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 601 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 602 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 603 | lemma lemma_additivity3: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 604 | "[| partition(a,b) D; D na < D n; D n < D (Suc na); | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 605 | n < psize D |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 606 | ==> False" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 607 | apply (cut_tac m = n and n = "Suc na" in less_linear, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 608 | apply (drule_tac [2] n = n in partition_lt_gen, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 609 | apply (cut_tac m = "psize D" and n = na in less_linear) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 610 | apply (auto simp add: partition_rhs2 less_Suc_eq) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 611 | apply (drule_tac n = na in partition_lt_gen, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 612 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 613 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 614 | lemma psize_const [simp]: "psize (%x. k) = 0" | 
| 15219 | 615 | by (auto simp add: psize_def) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 616 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 617 | lemma lemma_additivity3a: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 618 | "[| partition(a,b) D; D na < D n; D n < D (Suc na); | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 619 | na < psize D |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 620 | ==> False" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 621 | apply (frule_tac m = n in partition_lt_cancel) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 622 | apply (auto intro: lemma_additivity3) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 623 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 624 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 625 | lemma better_lemma_psize_right_eq1: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 626 | "[| partition(a,b) D; D n < b |] ==> psize (%x. D (x + n)) \<le> psize D - n" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 627 | apply (simp add: psize_def [of "(%x. D (x + n))"]); | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 628 | apply (rule_tac a = "psize D - n" in someI2, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 629 | apply (simp add: partition less_diff_conv) | 
| 15219 | 630 | apply (simp add: le_diff_conv partition_rhs2 split: nat_diff_split) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 631 | apply (drule_tac x = "psize D - n" in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 632 | apply (frule partition_rhs, safe) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 633 | apply (frule partition_lt_cancel, assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 634 | apply (drule partition [THEN iffD1], safe) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 635 | apply (subgoal_tac "~ D (psize D - n + n) < D (Suc (psize D - n + n))") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 636 | apply blast | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 637 | apply (drule_tac x = "Suc (psize D)" and P="%n. ?P n \<longrightarrow> D n = D (psize D)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 638 | in spec) | 
| 15219 | 639 | apply simp | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 640 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 641 | |
| 15219 | 642 | lemma psize_le_n: "partition (a, D n) D ==> psize D \<le> n" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 643 | apply (rule ccontr, drule not_leE) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 644 | apply (frule partition_lt_Suc, assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 645 | apply (frule_tac r = "Suc n" in partition_ub, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 646 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 647 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 648 | lemma better_lemma_psize_right_eq1a: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 649 | "partition(a,D n) D ==> psize (%x. D (x + n)) \<le> psize D - n" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 650 | apply (simp add: psize_def [of "(%x. D (x + n))"]); | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 651 | apply (rule_tac a = "psize D - n" in someI2, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 652 | apply (simp add: partition less_diff_conv) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 653 | apply (simp add: le_diff_conv) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 654 | apply (case_tac "psize D \<le> n") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 655 | apply (force intro: partition_rhs2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 656 | apply (simp add: partition linorder_not_le) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 657 | apply (rule ccontr, drule not_leE) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 658 | apply (frule psize_le_n) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 659 | apply (drule_tac x = "psize D - n" in spec, simp) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 660 | apply (drule partition [THEN iffD1], safe) | 
| 15219 | 661 | apply (drule_tac x = "Suc n" and P="%na. ?s \<le> na \<longrightarrow> D na = D n" in spec, auto) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 662 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 663 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 664 | lemma better_lemma_psize_right_eq: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 665 | "partition(a,b) D ==> psize (%x. D (x + n)) \<le> psize D - n" | 
| 15219 | 666 | apply (frule_tac r1 = n in partition_ub [THEN order_le_imp_less_or_eq]) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 667 | apply (blast intro: better_lemma_psize_right_eq1a better_lemma_psize_right_eq1) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 668 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 669 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 670 | lemma lemma_psize_right_eq1: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 671 | "[| partition(a,b) D; D n < b |] ==> psize (%x. D (x + n)) \<le> psize D" | 
| 15219 | 672 | apply (simp add: psize_def [of "(%x. D (x + n))"]) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 673 | apply (rule_tac a = "psize D - n" in someI2, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 674 | apply (simp add: partition less_diff_conv) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 675 | apply (subgoal_tac "n \<le> psize D") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 676 | apply (simp add: partition le_diff_conv) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 677 | apply (rule ccontr, drule not_leE) | 
| 15219 | 678 | apply (drule_tac less_imp_le [THEN [2] partition_rhs2], assumption, simp) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 679 | apply (drule_tac x = "psize D" in spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 680 | apply (simp add: partition) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 681 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 682 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 683 | (* should be combined with previous theorem; also proof has redundancy *) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 684 | lemma lemma_psize_right_eq1a: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 685 | "partition(a,D n) D ==> psize (%x. D (x + n)) \<le> psize D" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 686 | apply (simp add: psize_def [of "(%x. D (x + n))"]); | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 687 | apply (rule_tac a = "psize D - n" in someI2, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 688 | apply (simp add: partition less_diff_conv) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 689 | apply (case_tac "psize D \<le> n") | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 690 | apply (force intro: partition_rhs2 simp add: le_diff_conv) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 691 | apply (simp add: partition le_diff_conv) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 692 | apply (rule ccontr, drule not_leE) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 693 | apply (drule_tac x = "psize D" in spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 694 | apply (simp add: partition) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 695 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 696 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 697 | lemma lemma_psize_right_eq: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 698 | "[| partition(a,b) D |] ==> psize (%x. D (x + n)) \<le> psize D" | 
| 15219 | 699 | apply (frule_tac r1 = n in partition_ub [THEN order_le_imp_less_or_eq]) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 700 | apply (blast intro: lemma_psize_right_eq1a lemma_psize_right_eq1) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 701 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 702 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 703 | lemma tpart_left1: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 704 | "[| a \<le> D n; tpart (a, b) (D, p) |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 705 | ==> tpart(a, D n) (%x. if D x < D n then D(x) else D n, | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 706 | %x. if D x < D n then p(x) else D n)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 707 | apply (frule_tac r = n in tpart_partition [THEN partition_ub]) | 
| 15219 | 708 | apply (drule_tac x = "D n" in order_le_imp_less_or_eq) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 709 | apply (auto simp add: tpart_partition [THEN lemma_partition_eq, symmetric] tpart_tag_eq [symmetric]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 710 | apply (frule_tac tpart_partition [THEN [3] lemma_additivity1]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 711 | apply (auto simp add: tpart_def) | 
| 15219 | 712 | apply (drule_tac [2] linorder_not_less [THEN iffD1, THEN order_le_imp_less_or_eq], auto) | 
| 713 | prefer 3 apply (drule_tac x=na in spec, arith) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 714 | prefer 2 apply (blast dest: lemma_additivity3) | 
| 15219 | 715 | apply (frule (2) lemma_additivity4_psize_eq) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 716 | apply (rule partition [THEN iffD2]) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 717 | apply (frule partition [THEN iffD1]) | 
| 15219 | 718 | apply safe | 
| 719 | apply (auto simp add: partition_lt_gen) | |
| 15197 | 720 | apply (drule (1) partition_lt_cancel, arith) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 721 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 722 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 723 | lemma fine_left1: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 724 | "[| a \<le> D n; tpart (a, b) (D, p); gauge (%x. a \<le> x & x \<le> D n) g; | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 725 | fine (%x. if x < D n then min (g x) ((D n - x)/ 2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 726 | else if x = D n then min (g (D n)) (ga (D n)) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 727 | else min (ga x) ((x - D n)/ 2)) (D, p) |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 728 | ==> fine g | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 729 | (%x. if D x < D n then D(x) else D n, | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 730 | %x. if D x < D n then p(x) else D n)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 731 | apply (auto simp add: fine_def tpart_def gauge_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 732 | apply (frule_tac [!] na=na in lemma_psize_left_less_psize2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 733 | apply (drule_tac [!] x = na in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 734 | apply (drule_tac [!] x = na in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 735 | apply (auto dest: lemma_additivity3a simp add: split_if_asm) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 736 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 737 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 738 | lemma tpart_right1: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 739 | "[| a \<le> D n; tpart (a, b) (D, p) |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 740 | ==> tpart(D n, b) (%x. D(x + n),%x. p(x + n))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 741 | apply (simp add: tpart_def partition_def, safe) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 742 | apply (rule_tac x = "N - n" in exI, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 743 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 744 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 745 | lemma fine_right1: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 746 | "[| a \<le> D n; tpart (a, b) (D, p); gauge (%x. D n \<le> x & x \<le> b) ga; | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 747 | fine (%x. if x < D n then min (g x) ((D n - x)/ 2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 748 | else if x = D n then min (g (D n)) (ga (D n)) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 749 | else min (ga x) ((x - D n)/ 2)) (D, p) |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 750 | ==> fine ga (%x. D(x + n),%x. p(x + n))" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 751 | apply (auto simp add: fine_def gauge_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 752 | apply (drule_tac x = "na + n" in spec) | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19765diff
changeset | 753 | apply (frule_tac n = n in tpart_partition [THEN better_lemma_psize_right_eq], auto) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 754 | apply (simp add: tpart_def, safe) | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 755 | apply (subgoal_tac "D n \<le> p (na + n)") | 
| 15219 | 756 | apply (drule_tac y = "p (na + n)" in order_le_imp_less_or_eq) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 757 | apply safe | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 758 | apply (simp split: split_if_asm, simp) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 759 | apply (drule less_le_trans, assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 760 | apply (rotate_tac 5) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 761 | apply (drule_tac x = "na + n" in spec, safe) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 762 | apply (rule_tac y="D (na + n)" in order_trans) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 763 | apply (case_tac "na = 0", auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 764 | apply (erule partition_lt_gen [THEN order_less_imp_le], arith+) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 765 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 766 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 767 | lemma rsum_add: "rsum (D, p) (%x. f x + g x) = rsum (D, p) f + rsum(D, p) g" | 
| 15536 | 768 | by (simp add: rsum_def setsum_addf left_distrib) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 769 | |
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 770 | text{* Bartle/Sherbert: Theorem 10.1.5 p. 278 *}
 | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 771 | lemma Integral_add_fun: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 772 | "[| a \<le> b; Integral(a,b) f k1; Integral(a,b) g k2 |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 773 | ==> Integral(a,b) (%x. f x + g x) (k1 + k2)" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 774 | apply (simp add: Integral_def, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 775 | apply ((drule_tac x = "e/2" in spec)+) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 776 | apply auto | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 777 | apply (drule gauge_min, assumption) | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 778 | apply (rule_tac x = " (%x. if ga x < gaa x then ga x else gaa x)" in exI) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 779 | apply auto | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 780 | apply (drule fine_min) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 781 | apply ((drule spec)+, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 782 | apply (drule_tac a = "\<bar>rsum (D, p) f - k1\<bar> * 2" and c = "\<bar>rsum (D, p) g - k2\<bar> * 2" in add_strict_mono, assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 783 | apply (auto simp only: rsum_add left_distrib [symmetric] | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19765diff
changeset | 784 | mult_2_right [symmetric] real_mult_less_iff1) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 785 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 786 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 787 | lemma partition_lt_gen2: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 788 | "[| partition(a,b) D; r < psize D |] ==> 0 < D (Suc r) - D r" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 789 | by (auto simp add: partition) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 790 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 791 | lemma lemma_Integral_le: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 792 | "[| \<forall>x. a \<le> x & x \<le> b --> f x \<le> g x; | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 793 | tpart(a,b) (D,p) | 
| 15360 | 794 | |] ==> \<forall>n \<le> psize D. f (p n) \<le> g (p n)" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 795 | apply (simp add: tpart_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 796 | apply (auto, frule partition [THEN iffD1], auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 797 | apply (drule_tac x = "p n" in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 798 | apply (case_tac "n = 0", simp) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 799 | apply (rule partition_lt_gen [THEN order_less_le_trans, THEN order_less_imp_le], auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 800 | apply (drule le_imp_less_or_eq, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 801 | apply (drule_tac [2] x = "psize D" in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 802 | apply (drule_tac r = "Suc n" in partition_ub) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 803 | apply (drule_tac x = n in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 804 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 805 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 806 | lemma lemma_Integral_rsum_le: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 807 | "[| \<forall>x. a \<le> x & x \<le> b --> f x \<le> g x; | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 808 | tpart(a,b) (D,p) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 809 | |] ==> rsum(D,p) f \<le> rsum(D,p) g" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 810 | apply (simp add: rsum_def) | 
| 15539 | 811 | apply (auto intro!: setsum_mono dest: tpart_partition [THEN partition_lt_gen2] | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 812 | dest!: lemma_Integral_le) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 813 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 814 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 815 | lemma Integral_le: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 816 | "[| a \<le> b; | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 817 | \<forall>x. a \<le> x & x \<le> b --> f(x) \<le> g(x); | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 818 | Integral(a,b) f k1; Integral(a,b) g k2 | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 819 | |] ==> k1 \<le> k2" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 820 | apply (simp add: Integral_def) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 821 | apply (rotate_tac 2) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 822 | apply (drule_tac x = "\<bar>k1 - k2\<bar> /2" in spec) | 
| 15221 | 823 | apply (drule_tac x = "\<bar>k1 - k2\<bar> /2" in spec, auto) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 824 | apply (drule gauge_min, assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 825 | apply (drule_tac g = "%x. if ga x < gaa x then ga x else gaa x" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 826 | in partition_exists, assumption, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 827 | apply (drule fine_min) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 828 | apply (drule_tac x = D in spec, drule_tac x = D in spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 829 | apply (drule_tac x = p in spec, drule_tac x = p in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 830 | apply (frule lemma_Integral_rsum_le, assumption) | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 831 | apply (subgoal_tac "\<bar>(rsum (D,p) f - k1) - (rsum (D,p) g - k2)\<bar> < \<bar>k1 - k2\<bar>") | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 832 | apply arith | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 833 | apply (drule add_strict_mono, assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 834 | apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19765diff
changeset | 835 | real_mult_less_iff1) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 836 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 837 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 838 | lemma Integral_imp_Cauchy: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 839 | "(\<exists>k. Integral(a,b) f k) ==> | 
| 15360 | 840 | (\<forall>e > 0. \<exists>g. gauge (%x. a \<le> x & x \<le> b) g & | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 841 | (\<forall>D1 D2 p1 p2. | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 842 | tpart(a,b) (D1, p1) & fine g (D1,p1) & | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 843 | tpart(a,b) (D2, p2) & fine g (D2,p2) --> | 
| 15360 | 844 | \<bar>rsum(D1,p1) f - rsum(D2,p2) f\<bar> < e))" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 845 | apply (simp add: Integral_def, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 846 | apply (drule_tac x = "e/2" in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 847 | apply (rule exI, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 848 | apply (frule_tac x = D1 in spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 849 | apply (frule_tac x = D2 in spec) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 850 | apply ((drule spec)+, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 851 | apply (erule_tac V = "0 < e" in thin_rl) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 852 | apply (drule add_strict_mono, assumption) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 853 | apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19765diff
changeset | 854 | real_mult_less_iff1) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 855 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 856 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 857 | lemma Cauchy_iff2: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 858 | "Cauchy X = | 
| 20563 | 859 | (\<forall>j. (\<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. \<bar>X m - X n\<bar> < inverse(real (Suc j))))" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 860 | apply (simp add: Cauchy_def, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 861 | apply (drule reals_Archimedean, safe) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 862 | apply (drule_tac x = n in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 863 | apply (rule_tac x = M in exI, auto) | 
| 15360 | 864 | apply (drule_tac x = m in spec, simp) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 865 | apply (drule_tac x = na in spec, auto) | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 866 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 867 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 868 | lemma partition_exists2: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 869 | "[| a \<le> b; \<forall>n. gauge (%x. a \<le> x & x \<le> b) (fa n) |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 870 | ==> \<forall>n. \<exists>D p. tpart (a, b) (D, p) & fine (fa n) (D, p)" | 
| 15219 | 871 | by (blast dest: partition_exists) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 872 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 873 | lemma monotonic_anti_derivative: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 874 | "[| a \<le> b; \<forall>c. a \<le> c & c \<le> b --> f' c \<le> g' c; | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 875 | \<forall>x. DERIV f x :> f' x; \<forall>x. DERIV g x :> g' x |] | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 876 | ==> f b - f a \<le> g b - g a" | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 877 | apply (rule Integral_le, assumption) | 
| 15219 | 878 | apply (auto intro: FTC1) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 879 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 880 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 881 | end |