src/HOL/Hyperreal/Integration.ML
author paulson
Wed, 28 Jul 2004 10:49:29 +0200
changeset 15079 2ef899e4526d
parent 14477 cc61fd03e589
child 15082 6c3276a2735b
permissions -rw-r--r--
conversion of Hyperreal/MacLaurin_lemmas to Isar script
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
     1
(*  Title       : Integration.ML
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
     3
    Copyright   : 2000  University of Edinburgh
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
     4
    Description : Theory of integration (c.f. Harison's HOL development)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
     5
*)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
     6
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14365
diff changeset
     7
val mult_2 = thm"mult_2";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14365
diff changeset
     8
val mult_2_right = thm"mult_2_right";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14365
diff changeset
     9
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14477
diff changeset
    10
fun ARITH_PROVE str = prove_goal thy str
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14477
diff changeset
    11
                      (fn prems => [cut_facts_tac prems 1,arith_tac 1]);
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14477
diff changeset
    12
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    13
Goalw [psize_def] "a = b ==> psize (%n. if n = 0 then a else b) = 0";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    14
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    15
qed "partition_zero";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    16
Addsimps [partition_zero];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    17
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    18
Goalw [psize_def] "a < b ==> psize (%n. if n = 0 then a else b) = 1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    19
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    20
by (rtac some_equality 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    21
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    22
by (dres_inst_tac [("x","1")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    23
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    24
qed "partition_one";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    25
Addsimps [partition_one];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    26
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    27
Goalw [partition_def] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    28
   "a <= b ==> partition(a,b)(%n. if n = 0 then a else b)";
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14336
diff changeset
    29
by (auto_tac (claset(),simpset() addsimps [order_le_less]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    30
qed "partition_single";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    31
Addsimps [partition_single];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    32
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    33
Goalw [partition_def] "partition(a,b) D ==> (D(0) = a)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    34
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    35
qed "partition_lhs";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    36
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    37
Goalw [partition_def]
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    38
       "(partition(a,b) D) = \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    39
\       ((D 0 = a) & \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    40
\        (ALL n. n < (psize D) --> D n < D(Suc n)) & \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    41
\        (ALL n. (psize D) <= n --> (D n = b)))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    42
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    43
by (ALLGOALS(subgoal_tac "psize D = N"));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    44
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    45
by (ALLGOALS(simp_tac (simpset() addsimps [psize_def])));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    46
by (ALLGOALS(rtac some_equality));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    47
by (Blast_tac 1 THEN Blast_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    48
by (ALLGOALS(rtac ccontr));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    49
by (ALLGOALS(dtac (ARITH_PROVE "n ~= (N::nat) ==> n < N | N < n")));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    50
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    51
by (dres_inst_tac [("x","Na")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    52
by (rotate_tac 3 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    53
by (dres_inst_tac [("x","Suc Na")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    54
by (Asm_full_simp_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    55
by (rotate_tac 2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    56
by (dres_inst_tac [("x","N")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    57
by (Asm_full_simp_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    58
by (dres_inst_tac [("x","Na")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    59
by (rotate_tac 3 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    60
by (dres_inst_tac [("x","Suc Na")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    61
by (Asm_full_simp_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    62
by (rotate_tac 2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    63
by (dres_inst_tac [("x","N")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    64
by (Asm_full_simp_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    65
qed "partition";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    66
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    67
Goal "partition(a,b) D ==> (D(psize D) = b)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    68
by (dtac (partition RS subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    69
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    70
by (REPEAT(dres_inst_tac [("x","psize D")] spec 1));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    71
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    72
qed "partition_rhs";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    73
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    74
Goal "[|partition(a,b) D; psize D <= n |] ==> (D n = b)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    75
by (dtac (partition RS subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    76
by (Blast_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    77
qed "partition_rhs2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    78
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    79
Goal 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    80
 "partition(a,b) D & m + Suc d <= n & n <= (psize D) --> D(m) < D(m + Suc d)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    81
by (induct_tac "d" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    82
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    83
by (ALLGOALS(dtac (partition RS subst)));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    84
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    85
by (ALLGOALS(dtac (ARITH_PROVE "Suc m <= n ==> m < n")));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    86
by (ALLGOALS(dtac less_le_trans));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    87
by (assume_tac 1 THEN assume_tac 2);
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14336
diff changeset
    88
by (ALLGOALS(blast_tac (claset() addIs [order_less_trans])));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    89
qed_spec_mp "lemma_partition_lt_gen";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    90
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    91
Goal "m < n ==> EX d. n = m + Suc d";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    92
by (auto_tac (claset(),simpset() addsimps [less_iff_Suc_add]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    93
qed "less_eq_add_Suc";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    94
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    95
Goal "[|partition(a,b) D; m < n; n <= (psize D)|] ==> D(m) < D(n)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    96
by (auto_tac (claset() addDs [less_eq_add_Suc] addIs [lemma_partition_lt_gen],
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    97
    simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    98
qed "partition_lt_gen";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    99
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   100
Goal "partition(a,b) D ==> (n < (psize D) --> D(0) < D(Suc n))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   101
by (dtac (partition RS subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   102
by (induct_tac "n" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   103
by (Blast_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   104
by (blast_tac (claset() addSDs [leI] addDs 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   105
    [(ARITH_PROVE "m <= n ==> m <= Suc n"),
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14336
diff changeset
   106
    le_less_trans,order_less_trans]) 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   107
qed_spec_mp "partition_lt";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   108
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   109
Goal "partition(a,b) D ==> a <= b";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   110
by (ftac (partition RS subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   111
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   112
by (rotate_tac 2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   113
by (dres_inst_tac [("x","psize D")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   114
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   115
by (rtac ccontr 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   116
by (case_tac "psize D = 0" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   117
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   118
by (dres_inst_tac [("n","psize D - 1")] partition_lt 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   119
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   120
qed "partition_le";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   121
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   122
Goal "[|partition(a,b) D; n < (psize D)|] ==> D(n) < D(psize D)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   123
by (auto_tac (claset() addIs [partition_lt_gen],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   124
qed "partition_gt";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   125
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   126
Goal "partition(a,b) D ==> ((a = b) = (psize D = 0))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   127
by (ftac (partition RS subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   128
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   129
by (rotate_tac 2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   130
by (dres_inst_tac [("x","psize D")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   131
by (rtac ccontr 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   132
by (dres_inst_tac [("n","psize D - 1")] partition_lt 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   133
by (Blast_tac 3 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   134
qed "partition_eq";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   135
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   136
Goal "partition(a,b) D ==> a <= D(r)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   137
by (ftac (partition RS subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   138
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   139
by (induct_tac "r" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   140
by (cut_inst_tac [("m","Suc n"),("n","psize D")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   141
    (ARITH_PROVE "m < n | n <= (m::nat)") 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   142
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   143
by (eres_inst_tac [("j","D n")] real_le_trans 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   144
by (dres_inst_tac [("x","n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   145
by (blast_tac (claset() addIs [less_trans,order_less_imp_le]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   146
by (res_inst_tac [("j","b")] real_le_trans 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   147
by (etac partition_le 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   148
by (Blast_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   149
qed "partition_lb";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   150
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   151
Goal "[| partition(a,b) D; psize D ~= 0 |] ==> a < D(Suc n)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   152
by (res_inst_tac [("t","a")] (partition_lhs RS subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   153
by (assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   154
by (cut_inst_tac [("m","psize D"),("n","Suc n")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   155
    (ARITH_PROVE "m < n | n <= (m::nat)") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   156
by (ftac (partition RS subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   157
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   158
by (rotate_tac 3 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   159
by (dres_inst_tac [("x","Suc n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   160
by (etac impE 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   161
by (etac less_imp_le 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   162
by (ftac partition_rhs 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   163
by (dtac partition_gt 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   164
by (Asm_simp_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   165
by (blast_tac (claset() addIs [partition_lt,less_le_trans]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   166
qed "partition_lb_lt";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   167
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   168
Goal "partition(a,b) D ==> D(r) <= b";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   169
by (ftac (partition RS subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   170
by (cut_inst_tac [("m","r"),("n","psize D")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   171
    (ARITH_PROVE "m < n | n <= (m::nat)") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   172
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   173
by (Blast_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   174
by (subgoal_tac "ALL x. D((psize D) - x) <= b" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   175
by (rotate_tac 4 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   176
by (dres_inst_tac [("x","psize D - r")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   177
by (dtac (ARITH_PROVE "p < m ==> m - (m - p) = (p::nat)") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   178
by (thin_tac "ALL n. psize D <= n --> D n = b" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   179
by (Asm_full_simp_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   180
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   181
by (induct_tac "x" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   182
by (Simp_tac 1 THEN Blast_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   183
by (case_tac "psize D - Suc n = 0" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   184
by (thin_tac "ALL n. psize D <= n --> D n = b" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   185
by (asm_simp_tac (simpset() addsimps [partition_le]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   186
by (rtac real_le_trans 1 THEN assume_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   187
by (rtac (ARITH_PROVE "0 < m - Suc n ==> (m - n) = Suc(m - Suc n)" 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   188
          RS ssubst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   189
by (dres_inst_tac [("x","psize D - Suc n")] spec 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   190
by (thin_tac "ALL n. psize D <= n --> D n = b" 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   191
by (Asm_full_simp_tac 2);
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14305
diff changeset
   192
by (arith_tac 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   193
qed "partition_ub";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   194
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   195
Goal "[| partition(a,b) D; n < psize D |] ==> D(n) < b"; 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   196
by (blast_tac (claset() addIs [partition_rhs RS subst] addIs
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   197
    [partition_gt]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   198
qed "partition_ub_lt";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   199
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   200
Goal "[| partition (a, b) D1; partition (b, c) D2 |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   201
\      ==> (ALL n. \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   202
\            n < psize D1 + psize D2 --> \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   203
\            (if n < psize D1 then D1 n else D2 (n - psize D1)) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   204
\            < (if Suc n < psize D1 then D1 (Suc n) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   205
\               else D2 (Suc n - psize D1))) & \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   206
\        (ALL n. \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   207
\            psize D1 + psize D2 <= n --> \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   208
\            (if n < psize D1 then D1 n else D2 (n - psize D1)) = \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   209
\            (if psize D1 + psize D2 < psize D1 then D1 (psize D1 + psize D2) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   210
\             else D2 (psize D1 + psize D2 - psize D1)))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   211
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   212
by (auto_tac (claset() addIs [partition_lt_gen],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   213
by (dres_inst_tac [("m","psize D1")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   214
    (leI RSN(2,ARITH_PROVE "[| n < m; m <= Suc n |] ==> m = Suc n")) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   215
by (assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   216
by (auto_tac (claset() addSIs [partition_lt_gen],
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   217
    simpset() addsimps [partition_lhs,partition_ub_lt]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   218
by (blast_tac (claset() addIs [ARITH_PROVE "~n < m ==> n - m < Suc n - m"]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   219
by (auto_tac (claset(),simpset() addsimps [
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   220
    ARITH_PROVE "~ n < m ==> (Suc n - m <= p) = (Suc n <= m + p)"]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   221
by (dtac (ARITH_PROVE "p + q <= (n::nat) ==> q <= n - p") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   222
by (auto_tac (claset() addSIs [partition_rhs2], simpset() addsimps 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   223
    [partition_rhs]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   224
qed "lemma_partition_append1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   225
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   226
Goal "[| partition (a, b) D1; partition (b, c) D2; N < psize D1 |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   227
\     ==> D1(N) < D2 (psize D2)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   228
by (res_inst_tac [("y","D1(psize D1)")] order_less_le_trans 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   229
by (etac partition_gt 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   230
by (auto_tac (claset(),simpset() addsimps [partition_rhs,partition_le]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   231
qed "lemma_psize1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   232
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   233
Goal "[| partition (a, b) D1; partition (b, c) D2 |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   234
\     ==> psize (%n. if n < psize D1 then D1 n else D2 (n - psize D1)) = \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   235
\         psize D1 + psize D2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   236
by (res_inst_tac 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   237
    [("D2","%n. if n < psize D1 then D1 n else D2 (n - psize D1)")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   238
    ((psize_def RS meta_eq_to_obj_eq) RS ssubst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   239
by (rtac some1_equality 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   240
by (blast_tac (claset() addSIs [lemma_partition_append1]) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   241
by (rtac ex1I 1 THEN rtac lemma_partition_append1 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   242
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   243
by (rtac ccontr 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   244
by (dtac (ARITH_PROVE "m ~= n ==> m < n | n < (m::nat)") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   245
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   246
by (rotate_tac 3 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   247
by (dres_inst_tac [("x","psize D1 + psize D2")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   248
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   249
by (case_tac "N < psize D1" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   250
by (auto_tac (claset() addDs [lemma_psize1],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   251
by (dtac leI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   252
by (dtac (ARITH_PROVE "[|p <= n; n < p + m|] ==> n - p < (m::nat)") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   253
by (assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   254
by (dres_inst_tac [("a","b"),("b","c")] partition_gt 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   255
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   256
by (dres_inst_tac [("x","psize D1 + psize D2")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   257
by (auto_tac (claset(),simpset() addsimps [partition_rhs2]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   258
qed "lemma_partition_append2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   259
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   260
Goalw [tpart_def]
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   261
"[|psize D = 0; tpart(a,b) (D,p)|] ==> a = b";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   262
by (auto_tac (claset(),simpset() addsimps [partition_eq]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   263
qed "tpart_eq_lhs_rhs";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   264
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   265
Goalw [tpart_def] "tpart(a,b) (D,p) ==> partition(a,b) D";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   266
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   267
qed "tpart_partition";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   268
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   269
Goal "[| tpart(a,b) (D1,p1); fine(g) (D1,p1); \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   270
\        tpart(b,c) (D2,p2); fine(g) (D2,p2) |] \ 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   271
\      ==> EX D p. tpart(a,c) (D,p) & fine(g) (D,p)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   272
by (res_inst_tac 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   273
    [("x","%n. if n < psize D1 then D1 n else D2 (n - psize D1)")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   274
by (res_inst_tac 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   275
    [("x","%n. if n < psize D1 then p1 n else p2 (n - psize D1)")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   276
by (cut_inst_tac [("m","psize D1")] (ARITH_PROVE "0 < (m::nat) | m = 0") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   277
by (auto_tac (claset() addDs [tpart_eq_lhs_rhs],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   278
by (asm_full_simp_tac (simpset() addsimps [fine_def,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   279
    [tpart_partition,tpart_partition] MRS lemma_partition_append2]) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   280
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   281
by (dres_inst_tac [("m","psize D1"),("n","n")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   282
    (leI RSN(2,ARITH_PROVE "[| n < m; m <= Suc n |] ==> m = Suc n")) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   283
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   284
by (dtac (tpart_partition RS partition_rhs) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   285
by (dtac (tpart_partition RS partition_lhs) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   286
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   287
by (rotate_tac 3 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   288
by (dres_inst_tac [("x","n - psize D1")] spec 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   289
by (auto_tac (claset(),simpset() addsimps 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   290
    [ARITH_PROVE "[| (0::nat) < p; p <= n |] ==> (n - p < m) = (n < p + m)",
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   291
     ARITH_PROVE "~n < p ==> Suc n - p = Suc(n - p)"]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   292
by (auto_tac (claset(),simpset() addsimps [tpart_def,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   293
    ARITH_PROVE "~n < p ==> Suc n - p = Suc(n - p)"]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   294
by (dres_inst_tac [("m","psize D1"),("n","n")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   295
    (leI RSN(2,ARITH_PROVE "[| n < m; m <= Suc n |] ==> m = Suc n")) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   296
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   297
by (dtac partition_rhs 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   298
by (dtac partition_lhs 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   299
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   300
by (rtac (partition RS ssubst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   301
by (rtac (lemma_partition_append2 RS ssubst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   302
by (rtac conjI 3);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   303
by (dtac lemma_partition_append1 4);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   304
by (auto_tac (claset(),simpset() addsimps [partition_lhs,partition_rhs]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   305
qed "partition_append";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   306
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   307
(* ------------------------------------------------------------------------ *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   308
(* We can always find a division which is fine wrt any gauge                *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   309
(* ------------------------------------------------------------------------ *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   310
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   311
Goal "[| a <= b; gauge(%x. a <= x & x <= b) g |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   312
\     ==> EX D p. tpart(a,b) (D,p) & fine g (D,p)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   313
by (cut_inst_tac 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   314
 [("P","%(u,v). a <= u & v <= b --> (EX D p. tpart(u,v) (D,p) & fine(g) (D,p))")]
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   315
 lemma_BOLZANO2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   316
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   317
by (REPEAT(blast_tac (claset() addIs [real_le_trans]) 1));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   318
by (auto_tac (claset() addIs [partition_append],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   319
by (case_tac "a <= x & x <= b" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   320
by (res_inst_tac [("x","1")] exI 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   321
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   322
by (res_inst_tac [("x","g x")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   323
by (auto_tac (claset(),simpset() addsimps [gauge_def]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   324
by (res_inst_tac [("x","%n. if n = 0 then aa else ba")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   325
by (res_inst_tac [("x","%n. if n = 0 then x else ba")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   326
by (auto_tac (claset(),simpset() addsimps [tpart_def,fine_def]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   327
qed "partition_exists";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   328
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   329
(* ------------------------------------------------------------------------ *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   330
(* Lemmas about combining gauges                                            *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   331
(* ------------------------------------------------------------------------ *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   332
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   333
Goalw [gauge_def] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   334
     "[| gauge(E) g1; gauge(E) g2 |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   335
\     ==> gauge(E) (%x. if g1(x) < g2(x) then g1(x) else g2(x))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   336
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   337
qed "gauge_min";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   338
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   339
Goalw [fine_def]
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   340
      "fine (%x. if g1(x) < g2(x) then g1(x) else g2(x)) (D,p) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   341
\      ==> fine(g1) (D,p) & fine(g2) (D,p)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   342
by (auto_tac (claset(),simpset() addsimps [split_if]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   343
qed "fine_min";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   344
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   345
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   346
(* ------------------------------------------------------------------------ *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   347
(* The integral is unique if it exists                                      *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   348
(* ------------------------------------------------------------------------ *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   349
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   350
Goalw [Integral_def] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   351
    "[| a <= b; Integral(a,b) f k1; Integral(a,b) f k2 |] ==> k1 = k2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   352
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   353
by (REPEAT(dres_inst_tac [("x","abs(k1 - k2)/2")] spec 1));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   354
by (Auto_tac THEN TRYALL(arith_tac));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   355
by (dtac gauge_min 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   356
by (dres_inst_tac [("g","%x. if g x < ga x then g x else ga x")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   357
    partition_exists 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   358
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   359
by (dtac fine_min 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   360
by (REPEAT(dtac spec 1) THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   361
by (subgoal_tac 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   362
    "abs((rsum(D,p) f - k2) - (rsum(D,p) f - k1)) < abs(k1 - k2)" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   363
by (arith_tac 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   364
by (dtac add_strict_mono 1 THEN assume_tac 1);
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14294
diff changeset
   365
by (auto_tac (claset(),
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   366
    HOL_ss addsimps [left_distrib RS sym,
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14365
diff changeset
   367
                     mult_2_right RS sym, mult_less_cancel_right]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   368
by (ALLGOALS(arith_tac));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   369
qed "Integral_unique";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   370
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   371
Goalw [Integral_def] "Integral(a,a) f 0";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   372
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   373
by (res_inst_tac [("x","%x. 1")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   374
by (auto_tac (claset() addDs [partition_eq],simpset() addsimps [gauge_def,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   375
     tpart_def,rsum_def]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   376
qed "Integral_zero"; 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   377
Addsimps [Integral_zero];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   378
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   379
Goal "sumr 0 m (%n. D (Suc n) - D n) = D(m) - D 0";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   380
by (induct_tac "m" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   381
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   382
qed "sumr_partition_eq_diff_bounds";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   383
Addsimps [sumr_partition_eq_diff_bounds];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   384
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   385
Goal "a <= b ==> Integral(a,b) (%x. 1) (b - a)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   386
by (dtac real_le_imp_less_or_eq 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   387
by (auto_tac (claset(),simpset() addsimps [rsum_def,Integral_def]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   388
by (res_inst_tac [("x","%x. b - a")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   389
by (auto_tac (claset(),simpset() addsimps [gauge_def,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   390
    abs_interval_iff,tpart_def,partition]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   391
qed "Integral_eq_diff_bounds";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   392
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   393
Goal "a <= b ==> Integral(a,b) (%x. c)  (c*(b - a))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   394
by (dtac real_le_imp_less_or_eq 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   395
by (auto_tac (claset(),simpset() addsimps [rsum_def,Integral_def]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   396
by (res_inst_tac [("x","%x. b - a")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   397
by (auto_tac (claset(),simpset() addsimps 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   398
    [sumr_mult RS sym,gauge_def,abs_interval_iff,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   399
     right_diff_distrib RS sym,partition,tpart_def]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   400
qed "Integral_mult_const";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   401
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   402
Goal "[| a <= b; Integral(a,b) f k |] ==> Integral(a,b) (%x. c * f x) (c * k)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   403
by (dtac real_le_imp_less_or_eq 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   404
by (auto_tac (claset() addDs [[real_le_refl,Integral_zero] MRS Integral_unique],
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   405
    simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   406
by (auto_tac (claset(),simpset() addsimps [rsum_def,Integral_def,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   407
    sumr_mult RS sym,real_mult_assoc]));
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14277
diff changeset
   408
by (res_inst_tac [("a2","c")] ((abs_ge_zero RS real_le_imp_less_or_eq)
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   409
    RS disjE) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   410
by (dtac sym 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   411
by (Asm_full_simp_tac 2 THEN Blast_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   412
by (dres_inst_tac [("x","e/abs c")] spec 1 THEN Auto_tac);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   413
by (asm_full_simp_tac (simpset() addsimps [zero_less_mult_iff,
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   414
    real_divide_def]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   415
by (rtac exI 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   416
by (REPEAT(dtac spec 1) THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   417
by (res_inst_tac [("z1","inverse(abs c)")] (real_mult_less_iff1 RS iffD1) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   418
by (fold_tac [real_divide_def]);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   419
by (auto_tac (claset(),
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   420
      simpset() addsimps [right_diff_distrib RS sym,
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   421
                     abs_mult, real_mult_assoc RS sym,
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   422
    ARITH_PROVE "c ~= 0 ==> abs (c::real) ~= 0",
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   423
    positive_imp_inverse_positive]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   424
qed "Integral_mult";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   425
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   426
(* ------------------------------------------------------------------------ *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   427
(* Fundamental theorem of calculus (Part I)                                 *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   428
(* ------------------------------------------------------------------------ *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   429
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   430
(* "Straddle Lemma" : Swartz & Thompson: AMM 95(7) 1988 *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   431
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   432
Goal "ALL x. P(x) --> (EX y. Q x y) ==> EX f. (ALL x. P(x) --> Q x (f x))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   433
by (cut_inst_tac [("S","Collect (%x. P x)"),("Q","Q")] (CLAIM_SIMP 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   434
    "(ALL x:S. (EX y. Q x y)) --> (EX f. (ALL x:S. Q x (f x)))" [bchoice]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   435
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   436
qed "choiceP";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   437
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   438
Goal "ALL x. P(x) --> (EX y. R(y) & (EX z. Q x y z)) ==> \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   439
\     EX f fa. (ALL x. P(x) --> R(f x) & Q x (f x) (fa x))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   440
by (dtac (CLAIM "(ALL x. P(x) --> (EX y. R(y) & (EX z. Q x y z))) = \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   441
\                (ALL x. P(x) --> (EX y z. R(y) & Q x y z))" RS iffD1) 1);;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   442
by (dtac choiceP 1 THEN Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   443
by (dtac choiceP 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   444
qed "choiceP2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   445
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   446
Goal "ALL x. (EX y. R(y) & (EX z. Q x y z)) ==> \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   447
\     EX f fa. (ALL x. R(f x) & Q x (f x) (fa x))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   448
by (dtac (CLAIM "(ALL x. (EX y. R(y) & (EX z. Q x y z))) = \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   449
\                (ALL x. (EX y z. R(y) & Q x y z))" RS iffD1) 1);;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   450
by (dtac choice 1 THEN Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   451
by (dtac choice 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   452
qed "choice2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   453
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   454
(* new simplifications e.g. (y < x/n) = (y * n < x) are a real nuisance 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   455
   they break the original proofs and make new proofs longer!                 *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   456
Goalw [gauge_def]
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   457
     "[| ALL x. a <= x & x <= b --> DERIV f x :> f'(x); 0 < e |]\
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   458
\     ==> EX g. gauge(%x. a <= x & x <= b) g & \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   459
\               (ALL x u v. a <= u & u <= x & x <= v & v <= b & (v - u) < g(x) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   460
\                 --> abs((f(v) - f(u)) - (f'(x) * (v - u))) <= e * (v - u))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   461
by (subgoal_tac "ALL x. a <= x & x <= b --> \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   462
\                  (EX d. 0 < d & \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   463
\                     (ALL u v. u <= x & x <= v & (v - u) < d --> \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   464
\                     abs((f(v) - f(u)) - (f'(x) * (v - u))) <= e * (v - u)))" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   465
by (dtac choiceP 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   466
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   467
by (dtac spec 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   468
by (auto_tac (claset(),simpset() addsimps [DERIV_iff2,LIM_def]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   469
by (dres_inst_tac [("x","e/2")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   470
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   471
by (subgoal_tac "ALL z. abs(z - x) < s --> \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   472
\        abs((f(z) - f(x)) - (f'(x) * (z - x))) <= (e / 2) * abs(z - x)" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   473
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   474
by (case_tac "0 < abs(z - x)" 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   475
by (dtac (ARITH_PROVE "~ 0 < abs z ==> z = (0::real)") 3);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   476
by (asm_full_simp_tac (simpset() addsimps 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   477
    [ARITH_PROVE "(z - x = 0) = (z = (x::real))"]) 3);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   478
by (dres_inst_tac [("x","z")] spec 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   479
by (res_inst_tac [("z1","abs(inverse(z - x))")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   480
         (real_mult_le_cancel_iff2 RS iffD1) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   481
by (Asm_full_simp_tac 2);
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14277
diff changeset
   482
by (asm_full_simp_tac (simpset() 
14336
8f731d3cd65b Deleting more redundant theorems
paulson
parents: 14334
diff changeset
   483
     delsimps [abs_inverse, abs_mult]
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14277
diff changeset
   484
     addsimps [abs_mult RS sym, real_mult_assoc RS sym]) 2);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   485
by (subgoal_tac "inverse (z - x) * (f z - f x - f' x * (z - x)) = \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   486
\                (f z - f x)/(z - x) - f' x" 2);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   487
by (asm_full_simp_tac (simpset() addsimps [abs_mult RS sym] @ mult_ac) 2);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   488
by (asm_full_simp_tac (simpset() addsimps [real_diff_def]) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   489
by (rtac (real_mult_commute RS subst) 2);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   490
by (asm_full_simp_tac (simpset() addsimps [left_distrib,
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   491
    real_diff_def]) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   492
by (dtac (CLAIM "z ~= x ==> z + -x ~= (0::real)") 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   493
by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   494
    real_divide_def]) 2);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   495
by (simp_tac (simpset() addsimps [left_distrib]) 2);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   496
by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   497
by (res_inst_tac [("u1","u"),("v1","v")] (ARITH_PROVE "u < v | v <= (u::real)"
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   498
    RS disjE) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   499
by (dres_inst_tac [("x","v")] real_le_imp_less_or_eq 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   500
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   501
(*29*)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   502
by (res_inst_tac [("j","abs((f(v) - f(x)) - (f'(x) * (v - x))) + \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   503
\                  abs((f(x) - f(u)) - (f'(x) * (x - u)))")] real_le_trans 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   504
by (rtac (abs_triangle_ineq RSN (2,real_le_trans)) 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   505
by (asm_full_simp_tac (simpset() addsimps [right_diff_distrib]) 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   506
by (arith_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   507
by (res_inst_tac [("t","e*(v - u)")] (real_sum_of_halves RS subst) 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   508
by (rtac add_mono 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   509
by (res_inst_tac [("j","(e / 2) * abs(v - x)")] real_le_trans 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   510
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   511
by (Asm_full_simp_tac 2 THEN arith_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   512
by (ALLGOALS (thin_tac "ALL xa. \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   513
\             xa ~= x & abs (xa + - x) < s --> \
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14435
diff changeset
   514
\             abs ((f xa - f x) / (xa - x) + - f' x) * 2 < e"));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   515
by (dres_inst_tac [("x","v")] spec 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   516
by (arith_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   517
by (dres_inst_tac [("x","u")] spec 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   518
by (arith_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   519
by (subgoal_tac "abs (f u - f x - f' x * (u - x)) = \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   520
\                abs (f x - f u - f' x * (x - u))" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   521
by (Asm_full_simp_tac 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   522
by (asm_full_simp_tac (simpset() addsimps [right_distrib,
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   523
    real_diff_def]) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   524
by (arith_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   525
by(rtac real_le_trans 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   526
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   527
by (arith_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   528
qed "lemma_straddle";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   529
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   530
Goal "((number_of c - 1) * x <= 0) =((number_of c ::real) * x <= x)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   531
by (rtac ((ARITH_PROVE "(x <= y) = (x - y <= (0::real))") RS ssubst) 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   532
by (simp_tac (simpset() addsimps [left_diff_distrib,
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   533
    CLAIM_SIMP "c * x - x = (c - 1) * (x::real)" [left_diff_distrib]]) 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   534
qed "lemma_number_of_mult_le";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   535
 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   536
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   537
Goal "[|a <= b; ALL x. a <= x & x <= b --> DERIV f x :> f'(x) |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   538
\     ==> Integral(a,b) f' (f(b) - f(a))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   539
by (dtac real_le_imp_less_or_eq 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   540
by (auto_tac (claset(),simpset() addsimps [Integral_def]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   541
by (rtac ccontr 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   542
by (subgoal_tac "ALL e. 0 < e --> \ 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   543
\                   (EX g. gauge (%x. a <= x & x <= b) g & \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   544
\                   (ALL D p. \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   545
\                       tpart (a, b) (D, p) & fine g (D, p) --> \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   546
\                       abs (rsum (D, p) f' - (f b - f a)) <= e))" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   547
by (rotate_tac 3 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   548
by (dres_inst_tac [("x","e/2")] spec 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   549
by (dtac spec 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   550
by (REPEAT(dtac spec 1) THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   551
by (dres_inst_tac [("e","ea/(b - a)")] lemma_straddle 1);
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 13958
diff changeset
   552
by (auto_tac (claset(),simpset() addsimps [thm"Ring_and_Field.zero_less_divide_iff"]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   553
by (rtac exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   554
by (auto_tac (claset(),simpset() addsimps [tpart_def,rsum_def]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   555
by (subgoal_tac "sumr 0 (psize D) (%n. f(D(Suc n)) - f(D n)) = \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   556
\   f b - f a" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   557
by (cut_inst_tac [("D","%n. f(D n)"),("m","psize D")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   558
    sumr_partition_eq_diff_bounds 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   559
by (asm_full_simp_tac (simpset() addsimps [partition_lhs,partition_rhs]) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   560
by (dtac sym 1 THEN Asm_full_simp_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   561
by (simp_tac (simpset() addsimps [sumr_diff]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   562
by (rtac (sumr_rabs RS real_le_trans) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   563
by (subgoal_tac 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   564
    "ea = sumr 0 (psize D) (%n. (ea / (b - a)) * (D(Suc n) - (D n)))" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   565
by (asm_full_simp_tac (simpset() addsimps 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   566
    [ARITH_PROVE "abs(y - (x::real)) = abs(x - y)"]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   567
by (res_inst_tac [("t","ea")] ssubst 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   568
by (rtac sumr_le2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   569
by (rtac (sumr_mult RS subst) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   570
by (auto_tac (claset(),simpset() addsimps [partition_rhs,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   571
    partition_lhs,partition_lb,partition_ub,fine_def]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   572
qed "FTC1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   573
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   574
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   575
Goal "[| Integral(a,b) f k1; k2 = k1 |] ==> Integral(a,b) f k2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   576
by (Asm_simp_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   577
qed "Integral_subst";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   578
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   579
Goal "[| a <= b; b <= c; Integral(a,b) f' k1; Integral(b,c) f' k2; \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   580
\        ALL x. a <= x & x <= c --> DERIV f x :> f' x |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   581
\    ==> Integral(a,c) f' (k1 + k2)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   582
by (rtac (FTC1 RS Integral_subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   583
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   584
by (ftac FTC1 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   585
by (forw_inst_tac [("a","b")] FTC1 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   586
by (dres_inst_tac [("x","x")] spec 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   587
by (dres_inst_tac [("k2.0","f b - f a")] Integral_unique 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   588
by (dres_inst_tac [("k2.0","f c - f b")] Integral_unique 3);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   589
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   590
qed "Integral_add";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   591
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   592
Goal "partition(a,b) D ==> psize D = (LEAST n. D(n) = b)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   593
by (auto_tac (claset() addSIs [Least_equality RS sym,partition_rhs],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   594
by (rtac ccontr 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   595
by (dtac partition_ub_lt 1);
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14336
diff changeset
   596
by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   597
qed "partition_psize_Least";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   598
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   599
Goal "partition (a, c) D ==> ~ (EX n. c < D(n))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   600
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   601
by (dres_inst_tac [("r","n")] partition_ub 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   602
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   603
qed "lemma_partition_bounded";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   604
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   605
Goal "partition (a, c) D ==> D = (%n. if D n < c then D n else c)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   606
by (rtac ext 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   607
by (auto_tac (claset() addSDs [lemma_partition_bounded],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   608
by (dres_inst_tac [("x","n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   609
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   610
qed "lemma_partition_eq";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   611
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   612
Goal "partition (a, c) D ==> D = (%n. if D n <= c then D n else c)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   613
by (rtac ext 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   614
by (auto_tac (claset() addSDs [lemma_partition_bounded],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   615
by (dres_inst_tac [("x","n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   616
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   617
qed "lemma_partition_eq2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   618
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   619
Goal "[| partition(a,b) D; n < psize D |] ==> D n < D (Suc n)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   620
by (auto_tac (claset(),simpset() addsimps [partition]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   621
qed "partition_lt_Suc";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   622
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   623
Goal "tpart(a,c) (D,p) ==> p = (%n. if D n < c then p n else c)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   624
by (rtac ext 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   625
by (auto_tac (claset(),simpset() addsimps [tpart_def]));
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   626
by (dtac (linorder_not_less RS iffD1) 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   627
by (dres_inst_tac [("r","Suc n")] partition_ub 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   628
by (dres_inst_tac [("x","n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   629
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   630
qed "tpart_tag_eq";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   631
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   632
(*----------------------------------------------------------------------------*)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   633
(* Lemmas for Additivity Theorem of gauge integral                            *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   634
(*----------------------------------------------------------------------------*)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   635
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   636
Goal "[| a <= D n; D n < b; partition(a,b) D |] ==> n < psize D";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   637
by (dtac (partition RS iffD1) 1 THEN Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   638
by (rtac ccontr 1 THEN dtac leI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   639
by Auto_tac;
14435
9e22eeccf129 Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
paulson
parents: 14387
diff changeset
   640
qed "lemma_additivity1";
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   641
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   642
Goal "[| a <= D n; partition(a,D n) D |] ==> psize D <= n";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   643
by (rtac ccontr 1 THEN dtac not_leE 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   644
by (ftac (partition RS iffD1) 1 THEN Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   645
by (forw_inst_tac [("r","Suc n")] partition_ub 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   646
by (auto_tac (claset() addSDs [spec],simpset()));
14435
9e22eeccf129 Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
paulson
parents: 14387
diff changeset
   647
qed "lemma_additivity2";
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   648
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   649
Goal "[| partition(a,b) D; psize D < m |] ==> D(m) = D(psize D)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   650
by (auto_tac (claset(),simpset() addsimps [partition]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   651
qed "partition_eq_bound";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   652
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   653
Goal "[| partition(a,b) D; psize D < m |] ==> D(r) <= D(m)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   654
by (ftac (partition RS iffD1) 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   655
by (etac partition_ub 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   656
qed "partition_ub2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   657
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   658
Goalw [tpart_def]
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   659
    "[| tpart(a,b) (D,p); psize D <= m |] ==> p(m) = D(m)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   660
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   661
by (dres_inst_tac [("x","m")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   662
by (auto_tac (claset(),simpset() addsimps [partition_rhs2]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   663
qed "tag_point_eq_partition_point";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   664
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   665
Goal "[| partition(a,b) D; D m < D n |] ==> m < n";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   666
by (cut_inst_tac [("m","n"),("n","psize D")] less_linear 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   667
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   668
by (rtac ccontr 1 THEN dtac leI 1 THEN dtac le_imp_less_or_eq 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   669
by (cut_inst_tac [("m","m"),("n","psize D")] less_linear 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   670
by (auto_tac (claset() addDs [partition_gt],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   671
by (dres_inst_tac [("n","m")] partition_lt_gen 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   672
by (ftac partition_eq_bound 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   673
by (dtac partition_gt 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   674
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   675
by (rtac ccontr 1 THEN dtac leI 1 THEN dtac le_imp_less_or_eq 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   676
by (auto_tac (claset() addDs [partition_eq_bound],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   677
by (rtac ccontr 1 THEN dtac leI 1 THEN dtac le_imp_less_or_eq 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   678
by (ftac partition_eq_bound 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   679
by (dres_inst_tac [("m","m")] partition_eq_bound 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   680
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   681
qed "partition_lt_cancel";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   682
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   683
Goalw [psize_def] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   684
     "[| a <= D n; D n < b; partition (a, b) D |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   685
\     ==> psize (%x. if D x < D n then D(x) else D n) = n";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   686
by (ftac lemma_additivity1 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   687
by (assume_tac 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   688
by (rtac some_equality 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   689
by (auto_tac (claset() addIs [partition_lt_Suc],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   690
by (dres_inst_tac [("n","n")] partition_lt_gen 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   691
by (assume_tac 1 THEN arith_tac 1 THEN arith_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   692
by (cut_inst_tac [("m","na"),("n","psize D")] less_linear 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   693
by (auto_tac (claset() addDs [partition_lt_cancel],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   694
by (rtac ccontr 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   695
by (dtac (ARITH_PROVE "m ~= (n::nat) ==> m < n | n < m") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   696
by (etac disjE 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   697
by (rotate_tac 5 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   698
by (dres_inst_tac [("x","n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   699
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   700
by (dres_inst_tac [("n","n")] partition_lt_gen 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   701
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   702
by (arith_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   703
by (dres_inst_tac [("x","n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   704
by (asm_full_simp_tac (simpset() addsimps [split_if]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   705
qed "lemma_additivity4_psize_eq";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   706
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   707
Goal "partition (a, b) D  \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   708
\     ==> psize (%x. if D x < D n then D(x) else D n) <= psize D";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   709
by (forw_inst_tac [("r","n")] partition_ub 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   710
by (dres_inst_tac [("x","D n")] real_le_imp_less_or_eq 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   711
by (auto_tac (claset(),simpset() addsimps 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   712
    [lemma_partition_eq RS sym]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   713
by (forw_inst_tac [("r","n")] partition_lb 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   714
by (dtac lemma_additivity4_psize_eq 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   715
by (rtac ccontr 3 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   716
by (ftac (not_leE RSN (2,partition_eq_bound)) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   717
by (auto_tac (claset(),simpset() addsimps [partition_rhs]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   718
qed "lemma_psize_left_less_psize";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   719
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   720
Goal "[| partition(a,b) D; na < psize (%x. if D x < D n then D(x) else D n) |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   721
\     ==> na < psize D";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   722
by (etac (lemma_psize_left_less_psize RSN (2,less_le_trans)) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   723
by (assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   724
qed "lemma_psize_left_less_psize2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   725
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   726
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   727
Goal "[| partition(a,b) D; D na < D n; D n < D (Suc na); \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   728
\        n < psize D |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   729
\     ==> False";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   730
by (cut_inst_tac [("m","n"),("n","Suc na")] less_linear 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   731
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   732
by (dres_inst_tac [("n","n")] partition_lt_gen 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   733
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   734
by (cut_inst_tac [("m","psize D"),("n","na")] less_linear 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   735
by (auto_tac (claset(),simpset() addsimps [partition_rhs2]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   736
by (dtac (ARITH_PROVE "n < Suc na ==> n < na | n = na") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   737
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   738
by (dres_inst_tac [("n","na")] partition_lt_gen 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   739
by Auto_tac;
14435
9e22eeccf129 Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
paulson
parents: 14387
diff changeset
   740
qed "lemma_additivity3";
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   741
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   742
Goalw [psize_def] "psize (%x. k) = 0";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   743
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   744
qed "psize_const";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   745
Addsimps [psize_const];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   746
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   747
Goal "[| partition(a,b) D; D na < D n; D n < D (Suc na); \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   748
\        na < psize D |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   749
\     ==> False";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   750
by (forw_inst_tac [("m","n")] partition_lt_cancel 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   751
by (auto_tac (claset() addIs [lemma_additivity3],simpset()));
14435
9e22eeccf129 Conversion of Poly to Isar script, and other tidying of HOL/Hyperreal
paulson
parents: 14387
diff changeset
   752
qed "lemma_additivity3a";
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   753
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   754
Goal "[| partition(a,b) D; D n < b |] ==> psize (%x. D (x + n)) <= psize D - n";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   755
by (res_inst_tac [("D2","(%x. D (x + n))")] (psize_def RS 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   756
    meta_eq_to_obj_eq RS ssubst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   757
by (res_inst_tac [("a","psize D - n")] someI2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   758
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   759
by (dtac (ARITH_PROVE "(m::nat) < n - p  ==> m + p < n") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   760
by (asm_full_simp_tac (simpset() addsimps [partition]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   761
by (dtac (ARITH_PROVE "(m::nat) - n <= p  ==> m <= p + n") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   762
by (case_tac "psize D <= n" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   763
by (dtac not_leE 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   764
by (asm_simp_tac (simpset() addsimps 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   765
    [ARITH_PROVE "p <= n ==> p - n = (0::nat)"]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   766
by (blast_tac (claset() addDs [partition_rhs2]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   767
by (asm_full_simp_tac (simpset() addsimps [partition]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   768
by (rtac ccontr 1 THEN dtac not_leE 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   769
by (dres_inst_tac [("x","psize D - n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   770
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   771
by (ftac partition_rhs 1 THEN Step_tac 1); 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   772
by (ftac partition_lt_cancel 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   773
by (dtac (partition  RS iffD1) 1 THEN Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   774
by (subgoal_tac "~  D (psize D - n + n) < D (Suc (psize D - n + n))" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   775
by (Blast_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   776
by (rotate_tac 6 1 THEN dres_inst_tac [("x","Suc (psize D)")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   777
by (Asm_simp_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   778
qed "better_lemma_psize_right_eq1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   779
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   780
Goal "partition (a, D n) D ==> psize D <= n";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   781
by (rtac ccontr 1 THEN dtac not_leE 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   782
by (ftac partition_lt_Suc 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   783
by (forw_inst_tac [("r","Suc n")] partition_ub 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   784
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   785
qed "psize_le_n";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   786
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   787
Goal "partition(a,D n) D ==> psize (%x. D (x + n)) <= psize D - n";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   788
by (res_inst_tac [("D2","(%x. D (x + n))")] (psize_def RS 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   789
    meta_eq_to_obj_eq RS ssubst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   790
by (res_inst_tac [("a","psize D - n")] someI2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   791
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   792
by (dtac (ARITH_PROVE "(m::nat) < n - p  ==> m + p < n") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   793
by (asm_full_simp_tac (simpset() addsimps [partition]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   794
by (dtac (ARITH_PROVE "(m::nat) - n <= p  ==> m <= p + n") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   795
by (case_tac "psize D <= n" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   796
by (dtac not_leE 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   797
by (asm_simp_tac (simpset() addsimps 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   798
    [ARITH_PROVE "p <= n ==> p - n = (0::nat)"]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   799
by (blast_tac (claset() addDs [partition_rhs2]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   800
by (asm_full_simp_tac (simpset() addsimps [partition]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   801
by (rtac ccontr 1 THEN dtac not_leE 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   802
by (ftac psize_le_n 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   803
by (dres_inst_tac [("x","psize D - n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   804
by (asm_full_simp_tac (simpset() addsimps 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   805
    [ARITH_PROVE "p <= n ==> p - n = (0::nat)"]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   806
by (dtac (partition  RS iffD1) 1 THEN Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   807
by (rotate_tac 5 1 THEN dres_inst_tac [("x","Suc n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   808
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   809
qed "better_lemma_psize_right_eq1a";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   810
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   811
Goal "partition(a,b) D ==> psize (%x. D (x + n)) <= psize D - n";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   812
by (forw_inst_tac [("r1","n")] (partition_ub RS real_le_imp_less_or_eq) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   813
by (blast_tac (claset() addIs [better_lemma_psize_right_eq1a,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   814
    better_lemma_psize_right_eq1]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   815
qed "better_lemma_psize_right_eq";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   816
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   817
Goal "[| partition(a,b) D; D n < b |] ==> psize (%x. D (x + n)) <= psize D";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   818
by (res_inst_tac [("D2","(%x. D (x + n))")] (psize_def RS 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   819
    meta_eq_to_obj_eq RS ssubst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   820
by (res_inst_tac [("a","psize D - n")] someI2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   821
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   822
by (dtac (ARITH_PROVE "(m::nat) < n - p  ==> m + p < n") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   823
by (asm_full_simp_tac (simpset() addsimps [partition]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   824
by (subgoal_tac "n <= psize D" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   825
by (dtac (ARITH_PROVE "(m::nat) - n <= p  ==> m <= p + n") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   826
by (asm_full_simp_tac (simpset() addsimps [partition]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   827
by (rtac ccontr 1 THEN dtac not_leE 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   828
by (dtac (less_imp_le RSN (2,partition_rhs2)) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   829
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   830
by (rtac ccontr 1 THEN dtac not_leE 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   831
by (dres_inst_tac [("x","psize D")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   832
by (asm_full_simp_tac (simpset() addsimps [partition]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   833
qed "lemma_psize_right_eq1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   834
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   835
(* should be combined with previous theorem; also proof has redundancy *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   836
Goal "partition(a,D n) D ==> psize (%x. D (x + n)) <= psize D";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   837
by (res_inst_tac [("D2","(%x. D (x + n))")] (psize_def RS 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   838
    meta_eq_to_obj_eq RS ssubst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   839
by (res_inst_tac [("a","psize D - n")] someI2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   840
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   841
by (dtac (ARITH_PROVE "(m::nat) < n - p  ==> m + p < n") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   842
by (asm_full_simp_tac (simpset() addsimps [partition]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   843
by (dtac (ARITH_PROVE "(m::nat) - n <= p  ==> m <= p + n") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   844
by (case_tac "psize D <= n" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   845
by (dtac not_leE 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   846
by (asm_simp_tac (simpset() addsimps 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   847
    [ARITH_PROVE "p <= n ==> p - n = (0::nat)"]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   848
by (blast_tac (claset() addDs [partition_rhs2]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   849
by (asm_full_simp_tac (simpset() addsimps [partition]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   850
by (rtac ccontr 1 THEN dtac not_leE 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   851
by (dres_inst_tac [("x","psize D")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   852
by (asm_full_simp_tac (simpset() addsimps [partition]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   853
qed "lemma_psize_right_eq1a";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   854
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   855
Goal "[| partition(a,b) D |] ==> psize (%x. D (x + n)) <= psize D";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   856
by (forw_inst_tac [("r1","n")] (partition_ub RS real_le_imp_less_or_eq) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   857
by (blast_tac (claset() addIs [lemma_psize_right_eq1a,lemma_psize_right_eq1]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   858
qed "lemma_psize_right_eq";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   859
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   860
Goal "[| a <= D n; tpart (a, b) (D, p) |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   861
\     ==> tpart(a, D n) (%x. if D x < D n then D(x) else D n, \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   862
\         %x. if D x < D n then p(x) else D n)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   863
by (forw_inst_tac [("r","n")] (tpart_partition RS partition_ub) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   864
by (dres_inst_tac [("x","D n")] real_le_imp_less_or_eq 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   865
by (auto_tac (claset(),simpset() addsimps 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   866
    [tpart_partition RS lemma_partition_eq RS sym,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   867
     tpart_tag_eq RS sym]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   868
by (ftac (tpart_partition RSN (3,lemma_additivity1)) 1); 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   869
by (auto_tac (claset(),simpset() addsimps [tpart_def]));
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   870
by (dtac ((linorder_not_less RS iffD1) RS real_le_imp_less_or_eq) 2);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   871
by (Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   872
by (blast_tac (claset() addDs [lemma_additivity3]) 2);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   873
by (dtac (linorder_not_less RS iffD1) 2 THEN dres_inst_tac [("x","na")] spec 2);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   874
by (arith_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   875
by (ftac lemma_additivity4_psize_eq 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   876
by (REPEAT(assume_tac 1));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   877
by (rtac (partition RS iffD2) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   878
by (ftac (partition RS iffD1) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   879
by (auto_tac (claset() addIs [partition_lt_gen],simpset() addsimps []));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   880
by (dres_inst_tac [("n","n")] partition_lt_gen 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   881
by (assume_tac 1 THEN arith_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   882
by (Blast_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   883
by (dtac partition_lt_cancel 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   884
qed "tpart_left1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   885
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   886
Goal "[| a <= D n; tpart (a, b) (D, p); gauge (%x. a <= x & x <= D n) g;\
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   887
\        fine (%x. if x < D n then min (g x) ((D n - x)/ 2) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   888
\                else if x = D n then min (g (D n)) (ga (D n)) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   889
\                     else min (ga x) ((x - D n)/ 2)) (D, p) |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   890
\     ==> fine g \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   891
\          (%x. if D x < D n then D(x) else D n, \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   892
\           %x. if D x < D n then p(x) else D n)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   893
by (auto_tac (claset(),simpset() addsimps [fine_def,tpart_def,gauge_def]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   894
by (ALLGOALS (ftac lemma_psize_left_less_psize2));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   895
by (TRYALL(assume_tac));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   896
by (ALLGOALS (dres_inst_tac [("x","na")] spec) THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   897
by (ALLGOALS(dres_inst_tac [("x","na")] spec));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   898
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   899
by (asm_full_simp_tac (simpset() addsimps [split_if]) 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   900
by (dtac ((linorder_not_less RS iffD1) RS real_le_imp_less_or_eq) 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   901
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   902
by (blast_tac (claset() addDs [lemma_additivity3a]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   903
by (dtac sym 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   904
qed "fine_left1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   905
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   906
Goal "[| a <= D n; tpart (a, b) (D, p) |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   907
\     ==> tpart(D n, b) (%x. D(x + n),%x. p(x + n))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   908
by (asm_full_simp_tac (simpset() addsimps [tpart_def,partition_def]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   909
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   910
by (res_inst_tac [("x","N - n")] exI 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   911
by (rotate_tac 2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   912
by (dres_inst_tac [("x","na + n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   913
by (rotate_tac 3 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   914
by (dres_inst_tac [("x","na + n")] spec 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   915
by (ALLGOALS(arith_tac));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   916
qed "tpart_right1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   917
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   918
Goal "[| a <= D n; tpart (a, b) (D, p); gauge (%x. D n <= x & x <= b) ga;\
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   919
\        fine (%x. if x < D n then min (g x) ((D n - x)/ 2) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   920
\                else if x = D n then min (g (D n)) (ga (D n)) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   921
\                     else min (ga x) ((x - D n)/ 2)) (D, p) |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   922
\     ==> fine ga (%x. D(x + n),%x. p(x + n))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   923
by (auto_tac (claset(),simpset() addsimps [fine_def,gauge_def]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   924
by (dres_inst_tac [("x","na + n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   925
by (forw_inst_tac [("n","n")] (tpart_partition RS better_lemma_psize_right_eq) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   926
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   927
by (arith_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   928
by (asm_full_simp_tac (simpset() addsimps [tpart_def]) 1 THEN Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   929
by (subgoal_tac "D n <= p (na + n)" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   930
by (dres_inst_tac [("y","p (na + n)")] real_le_imp_less_or_eq 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   931
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   932
by (asm_full_simp_tac (simpset() addsimps [split_if]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   933
by (Asm_full_simp_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   934
by (dtac less_le_trans 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   935
by (rotate_tac 5 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   936
by (dres_inst_tac [("x","na + n")] spec 1 THEN Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   937
by (rtac real_le_trans 1 THEN assume_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   938
by (case_tac "na = 0" 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   939
by (rtac (partition_lt_gen RS order_less_imp_le) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   940
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   941
by (arith_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   942
qed "fine_right1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   943
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   944
Goalw [rsum_def] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   945
   "rsum (D, p) (%x. f x + g x) =  rsum (D, p) f + rsum(D, p) g";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   946
by (auto_tac (claset(),simpset() addsimps [sumr_add,left_distrib]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   947
qed "rsum_add";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   948
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   949
(* Bartle/Sherbert: Theorem 10.1.5 p. 278 *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   950
Goalw [Integral_def] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   951
    "[| a <= b; Integral(a,b) f k1; Integral(a,b) g k2 |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   952
\    ==> Integral(a,b) (%x. f x + g x) (k1 + k2)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   953
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   954
by (REPEAT(dres_inst_tac [("x","e/2")] spec 1));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   955
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   956
by (dtac gauge_min 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   957
by (res_inst_tac [("x","(%x. if ga x < gaa x then ga x else gaa x)")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   958
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   959
by (dtac fine_min 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   960
by (REPEAT(dtac spec 1) THEN Auto_tac);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   961
by (dres_inst_tac  [("a","abs (rsum (D, p) f - k1)* 2"),
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   962
    ("c","abs (rsum (D, p) g - k2) * 2")] 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   963
    add_strict_mono 1 THEN assume_tac 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   964
by (auto_tac (claset(),HOL_ss addsimps [rsum_add,left_distrib RS sym,
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14365
diff changeset
   965
    mult_2_right RS sym,real_mult_less_iff1,CLAIM "(0::real) < 2"]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   966
by (arith_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   967
qed "Integral_add_fun";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   968
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   969
Goal "[| partition(a,b) D; r < psize D |] ==> 0 < D (Suc r) - D r";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   970
by (auto_tac (claset(),simpset() addsimps [partition]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   971
qed "partition_lt_gen2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   972
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   973
Goalw [tpart_def] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   974
     "[| ALL x. a <= x & x <= b --> f x <= g x; \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   975
\        tpart(a,b) (D,p) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   976
\     |] ==> ALL n. n <= psize D --> f (p n) <= g (p n)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   977
by (Auto_tac THEN ftac (partition RS iffD1) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   978
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   979
by (dres_inst_tac [("x","p n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   980
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   981
by (case_tac "n = 0" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   982
by (rtac (partition_lt_gen RS order_less_le_trans RS order_less_imp_le) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   983
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   984
by (dtac le_imp_less_or_eq 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   985
by (dres_inst_tac [("x","psize D")] spec 2 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   986
by (dres_inst_tac [("r","Suc n")] partition_ub 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   987
by (dres_inst_tac [("x","n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   988
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   989
qed "lemma_Integral_le";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   990
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   991
Goalw [rsum_def] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   992
     "[| ALL x. a <= x & x <= b --> f x <= g x; \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   993
\        tpart(a,b) (D,p) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   994
\     |] ==> rsum(D,p) f <= rsum(D,p) g";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   995
by (auto_tac (claset() addSIs [sumr_le2] addDs 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   996
    [tpart_partition RS partition_lt_gen2] addSDs 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   997
    [lemma_Integral_le],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   998
qed "lemma_Integral_rsum_le";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
   999
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1000
Goalw [Integral_def] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1001
    "[| a <= b; \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1002
\       ALL x. a <= x & x <= b --> f(x) <= g(x); \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1003
\       Integral(a,b) f k1; Integral(a,b) g k2 \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1004
\    |] ==> k1 <= k2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1005
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1006
by (rotate_tac 2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1007
by (dres_inst_tac [("x","abs(k1 - k2)/2")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1008
by (dres_inst_tac [("x","abs(k1 - k2)/2")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1009
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1010
by (dtac gauge_min 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1011
by (dres_inst_tac [("g","%x. if ga x < gaa x then ga x else gaa x")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1012
    partition_exists 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1013
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1014
by (dtac fine_min 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1015
by (dres_inst_tac [("x","D")] spec 1 THEN dres_inst_tac [("x","D")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1016
by (dres_inst_tac [("x","p")] spec 1 THEN dres_inst_tac [("x","p")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1017
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1018
by (ftac lemma_Integral_rsum_le 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1019
by (subgoal_tac 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1020
    "abs((rsum(D,p) f - k1) - (rsum(D,p) g - k2)) < abs(k1 - k2)" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1021
by (arith_tac 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1022
by (dtac add_strict_mono 1 THEN assume_tac 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1023
by (auto_tac (claset(),HOL_ss addsimps [left_distrib RS sym,
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14365
diff changeset
  1024
    mult_2_right RS sym,real_mult_less_iff1,CLAIM "(0::real) < 2"]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1025
by (arith_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1026
qed "Integral_le";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1027
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1028
Goalw [Integral_def] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1029
     "(EX k. Integral(a,b) f k) ==> \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1030
\     (ALL e. 0 < e --> \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1031
\              (EX g. gauge (%x. a <= x & x <= b) g & \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1032
\                      (ALL D1 D2 p1 p2. \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1033
\                           tpart(a,b) (D1, p1) & fine g (D1,p1) & \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1034
\                           tpart(a,b) (D2, p2) & fine g (D2,p2) --> \ 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1035
\                           abs(rsum(D1,p1) f - rsum(D2,p2) f) < e)))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1036
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1037
by (dres_inst_tac [("x","e/2")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1038
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1039
by (rtac exI 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1040
by (forw_inst_tac [("x","D1")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1041
by (forw_inst_tac [("x","D2")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1042
by (REPEAT(dtac spec 1) THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1043
by (thin_tac "0 < e" 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1044
by (dtac add_strict_mono 1 THEN assume_tac 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
  1045
by (auto_tac (claset(),HOL_ss addsimps [left_distrib RS sym,
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14365
diff changeset
  1046
    mult_2_right RS sym,real_mult_less_iff1,CLAIM "(0::real) < 2"]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1047
by (arith_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1048
qed "Integral_imp_Cauchy";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1049
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1050
Goalw [Cauchy_def] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1051
     "Cauchy X = \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1052
\     (ALL j. (EX M. ALL m n. M <= m & M <= n --> \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1053
\              abs (X m + - X n) < inverse(real (Suc j))))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1054
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1055
by (dtac reals_Archimedean 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1056
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1057
by (dres_inst_tac [("x","n")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1058
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1059
by (res_inst_tac [("x","M")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1060
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1061
by (dres_inst_tac [("x","m")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1062
by (dres_inst_tac [("x","na")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1063
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1064
qed "Cauchy_iff2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1065
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1066
Goal "[| a <= b; ALL n. gauge (%x. a <= x & x <= b) (fa n) |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1067
\     ==> ALL n. EX D p. tpart (a, b) (D, p) & fine (fa n) (D, p)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1068
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1069
by (rtac partition_exists 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1070
qed "partition_exists2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1071
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1072
Goal "[| a <= b; ALL c. a <= c & c <= b --> f' c <= g' c; \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1073
\        ALL x. DERIV f x :> f' x; ALL x. DERIV g x :> g' x \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1074
\     |] ==> f b - f a <= g b - g a";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1075
by (rtac Integral_le 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1076
by (assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1077
by (rtac FTC1 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1078
by (rtac FTC1 4);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1079
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
  1080
qed "monotonic_anti_derivative";