| 12196 |      1 | (*  Title       : NthRoot.ML
 | 
|  |      2 |     Author      : Jacques D. Fleuriot
 | 
|  |      3 |     Copyright   : 1998  University of Cambridge
 | 
|  |      4 |     Description : Existence of nth root. Adapted from
 | 
|  |      5 |                    http://www.math.unl.edu/~webnotes
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
|  |      8 | (*----------------------------------------------------------------------
 | 
|  |      9 |                          Existence of nth root
 | 
|  |     10 |    Various lemmas needed for this result. We follow the proof
 | 
|  |     11 |    given by John Lindsay Orr (jorr@math.unl.edu) in his Analysis
 | 
|  |     12 |    Webnotes available on the www at http://www.math.unl.edu/~webnotes
 | 
|  |     13 |    Lemmas about sequences of reals are used to reach the result.
 | 
|  |     14 |  ---------------------------------------------------------------------*)
 | 
|  |     15 | 
 | 
|  |     16 | Goal "[| (0::real) < a; 0 < n |] \
 | 
|  |     17 | \     ==> EX s. s : {x. x ^ n <= a & 0 < x}";
 | 
|  |     18 | by (case_tac "1 <= a" 1);
 | 
|  |     19 | by (res_inst_tac [("x","1")] exI 1);
 | 
|  |     20 | by (dtac not_real_leE 2);
 | 
|  |     21 | by (dtac (less_not_refl2 RS  not0_implies_Suc) 2);
 | 
|  |     22 | by (auto_tac (claset() addSIs [realpow_Suc_le_self],
 | 
|  |     23 |     simpset() addsimps [real_zero_less_one]));
 | 
|  |     24 | qed "lemma_nth_realpow_non_empty";
 | 
|  |     25 | 
 | 
|  |     26 | Goal "[| (0::real) < a; 0 < n |] \
 | 
|  |     27 | \     ==> EX u. isUb (UNIV::real set) {x. x ^ n <= a & 0 < x} u";
 | 
|  |     28 | by (case_tac "1 <= a" 1);
 | 
|  |     29 | by (res_inst_tac [("x","a")] exI 1);
 | 
|  |     30 | by (dtac not_real_leE 2);
 | 
|  |     31 | by (res_inst_tac [("x","1")] exI 2);
 | 
|  |     32 | by (ALLGOALS(rtac (setleI RS isUbI)));
 | 
|  |     33 | by (Auto_tac);
 | 
|  |     34 | by (ALLGOALS(rtac ccontr));
 | 
|  |     35 | by (ALLGOALS(dtac not_real_leE));
 | 
|  |     36 | by (dtac realpow_ge_self2 1 THEN assume_tac 1);
 | 
|  |     37 | by (dres_inst_tac [("n","n")] (conjI 
 | 
|  |     38 |     RSN (2,conjI RS realpow_less)) 1);
 | 
|  |     39 | by (REPEAT(assume_tac 1));
 | 
|  |     40 | by (dtac real_le_trans 1 THEN assume_tac 1);
 | 
|  |     41 | by (dres_inst_tac [("y","y ^ n")] order_less_le_trans 1);
 | 
|  |     42 | by (assume_tac 1 THEN etac real_less_irrefl 1);
 | 
|  |     43 | by (dres_inst_tac [("n","n")] ((real_zero_less_one) RS (conjI 
 | 
|  |     44 |     RSN (2,conjI RS realpow_less))) 1);
 | 
|  |     45 | by (Auto_tac);
 | 
|  |     46 | qed "lemma_nth_realpow_isUb_ex";
 | 
|  |     47 | 
 | 
|  |     48 | Goal "[| (0::real) < a; 0 < n |] \
 | 
|  |     49 | \    ==> EX u. isLub (UNIV::real set) {x. x ^ n <= a & 0 < x} u";
 | 
|  |     50 | by (blast_tac (claset() addIs [lemma_nth_realpow_isUb_ex,
 | 
|  |     51 |     lemma_nth_realpow_non_empty,reals_complete]) 1);
 | 
|  |     52 | qed "nth_realpow_isLub_ex";
 | 
|  |     53 |  
 | 
|  |     54 | (*---------------------------------------------
 | 
|  |     55 |           First Half -- lemmas first
 | 
|  |     56 |  --------------------------------------------*)
 | 
|  |     57 | Goal "isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u \
 | 
|  |     58 | \          ==> u + inverse(real_of_posnat k) ~: {x. x ^ n <= a & 0 < x}";
 | 
|  |     59 | by (Step_tac 1 THEN dtac isLubD2 1 THEN Blast_tac 1);
 | 
|  |     60 | by (asm_full_simp_tac (simpset() addsimps [real_le_def]) 1);
 | 
|  |     61 | val lemma_nth_realpow_seq = result();
 | 
|  |     62 | 
 | 
|  |     63 | Goal "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u; \
 | 
|  |     64 | \        0 < a; 0 < n |] ==> 0 < u";
 | 
|  |     65 | by (dtac lemma_nth_realpow_non_empty 1 THEN Auto_tac);
 | 
|  |     66 | by (dres_inst_tac [("y","s")] (isLub_isUb RS isUbD) 1);
 | 
|  |     67 | by (auto_tac (claset() addIs [order_less_le_trans],simpset()));
 | 
|  |     68 | val lemma_nth_realpow_isLub_gt_zero = result();
 | 
|  |     69 | 
 | 
|  |     70 | Goal "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u; \
 | 
|  |     71 | \        0 < a; 0 < n |] ==> ALL k. a <= (u + inverse(real_of_posnat k)) ^ n";
 | 
|  |     72 | by (Step_tac 1);
 | 
| 12486 |     73 | by (ftac lemma_nth_realpow_seq 1 THEN Step_tac 1);
 | 
| 12196 |     74 | by (auto_tac (claset() addEs [real_less_asym],
 | 
|  |     75 |     simpset() addsimps [real_le_def]));
 | 
|  |     76 | by (fold_tac [real_le_def]);
 | 
|  |     77 | by (rtac real_less_trans 1);
 | 
|  |     78 | by (auto_tac (claset() addIs [real_inv_real_of_posnat_gt_zero,
 | 
|  |     79 |     lemma_nth_realpow_isLub_gt_zero],simpset()));
 | 
|  |     80 | val lemma_nth_realpow_isLub_ge = result();
 | 
|  |     81 | 
 | 
|  |     82 | (*-----------------------
 | 
|  |     83 |    First result we want
 | 
|  |     84 |  ----------------------*)
 | 
|  |     85 | Goal "[| (0::real) < a; 0 < n; \
 | 
|  |     86 | \    isLub (UNIV::real set) \
 | 
|  |     87 | \    {x. x ^ n <= a & 0 < x} u |] ==> a <= u ^ n";
 | 
| 12486 |     88 | by (ftac lemma_nth_realpow_isLub_ge 1 THEN Step_tac 1);
 | 
| 12196 |     89 | by (rtac (LIMSEQ_inverse_real_of_posnat_add RS LIMSEQ_pow 
 | 
|  |     90 |     RS LIMSEQ_le_const) 1);
 | 
|  |     91 | by (auto_tac (claset(),simpset() addsimps [real_of_nat_def,
 | 
|  |     92 |     real_of_posnat_Suc]));
 | 
|  |     93 | qed "realpow_nth_ge";
 | 
|  |     94 | 
 | 
|  |     95 | (*---------------------------------------------
 | 
|  |     96 |           Second Half
 | 
|  |     97 |  --------------------------------------------*)
 | 
|  |     98 | 
 | 
|  |     99 | Goal "[| isLub (UNIV::real set) S u; x < u |] \
 | 
|  |    100 | \          ==> ~ isUb (UNIV::real set) S x";
 | 
|  |    101 | by (Step_tac 1);
 | 
|  |    102 | by (dtac isLub_le_isUb 1);
 | 
|  |    103 | by (assume_tac 1);
 | 
|  |    104 | by (dtac order_less_le_trans 1);
 | 
|  |    105 | by (auto_tac (claset(),simpset() 
 | 
|  |    106 |    addsimps [real_less_not_refl]));
 | 
|  |    107 | qed "less_isLub_not_isUb";
 | 
|  |    108 | 
 | 
|  |    109 | Goal "~ isUb (UNIV::real set) S u \
 | 
|  |    110 | \              ==> EX x: S. u < x";
 | 
|  |    111 | by (rtac ccontr 1 THEN etac swap 1);
 | 
|  |    112 | by (rtac (setleI RS isUbI) 1);
 | 
|  |    113 | by (auto_tac (claset(),simpset() addsimps [real_le_def]));
 | 
|  |    114 | qed "not_isUb_less_ex";
 | 
|  |    115 | 
 | 
|  |    116 | Goal "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u; \
 | 
|  |    117 | \      0 < a; 0 < n |] ==> ALL k. (u*(1 + -inverse(real_of_posnat k))) ^ n <= a";
 | 
|  |    118 | by (Step_tac 1);
 | 
|  |    119 | by (forward_tac [less_isLub_not_isUb RS not_isUb_less_ex] 1);
 | 
|  |    120 | by (res_inst_tac [("n","k")] real_mult_less_self 1);
 | 
|  |    121 | by (blast_tac (claset() addIs [lemma_nth_realpow_isLub_gt_zero]) 1);
 | 
|  |    122 | by (Step_tac 1);
 | 
|  |    123 | by (dres_inst_tac [("n","k")] (lemma_nth_realpow_isLub_gt_zero RS 
 | 
|  |    124 |           real_mult_add_one_minus_ge_zero) 1);
 | 
|  |    125 | by (dtac (conjI RS realpow_le2) 3);
 | 
|  |    126 | by (rtac (CLAIM "x < y ==> (x::real) <= y") 3);
 | 
|  |    127 | by (auto_tac (claset() addIs [real_le_trans],simpset()));
 | 
|  |    128 | val lemma_nth_realpow_isLub_le = result();
 | 
|  |    129 | 
 | 
|  |    130 | (*-----------------------
 | 
|  |    131 |    Second result we want
 | 
|  |    132 |  ----------------------*)
 | 
|  |    133 | Goal "[| (0::real) < a; 0 < n; \
 | 
|  |    134 | \    isLub (UNIV::real set) \
 | 
|  |    135 | \    {x. x ^ n <= a & 0 < x} u |] ==> u ^ n <= a";
 | 
| 12486 |    136 | by (ftac lemma_nth_realpow_isLub_le 1 THEN Step_tac 1);
 | 
| 12196 |    137 | by (rtac (LIMSEQ_inverse_real_of_posnat_add_minus_mult RS LIMSEQ_pow 
 | 
|  |    138 |     RS LIMSEQ_le_const2) 1);
 | 
|  |    139 | by (auto_tac (claset(),simpset() addsimps [real_of_nat_def,real_of_posnat_Suc]));
 | 
|  |    140 | qed "realpow_nth_le";
 | 
|  |    141 | 
 | 
|  |    142 | (*----------- The theorem at last! -----------*)
 | 
|  |    143 | Goal "[| (0::real) < a; 0 < n |] ==> EX r. r ^ n = a";
 | 
| 12486 |    144 | by (ftac nth_realpow_isLub_ex 1 THEN Auto_tac);
 | 
| 12196 |    145 | by (auto_tac (claset() addIs [realpow_nth_le,
 | 
|  |    146 |     realpow_nth_ge,real_le_anti_sym],simpset()));
 | 
|  |    147 | qed "realpow_nth";
 | 
|  |    148 | 
 | 
|  |    149 | (* positive only *)
 | 
|  |    150 | Goal "[| (0::real) < a; 0 < n |] ==> EX r. 0 < r & r ^ n = a";
 | 
| 12486 |    151 | by (ftac nth_realpow_isLub_ex 1 THEN Auto_tac);
 | 
| 12196 |    152 | by (auto_tac (claset() addIs [realpow_nth_le,
 | 
|  |    153 |     realpow_nth_ge,real_le_anti_sym,
 | 
|  |    154 |     lemma_nth_realpow_isLub_gt_zero],simpset()));
 | 
|  |    155 | qed "realpow_pos_nth";
 | 
|  |    156 | 
 | 
|  |    157 | Goal "(0::real) < a  ==> EX r. 0 < r & r ^ Suc n = a";
 | 
|  |    158 | by (blast_tac (claset() addIs [realpow_pos_nth]) 1);
 | 
|  |    159 | qed "realpow_pos_nth2";
 | 
|  |    160 | 
 | 
|  |    161 | (* uniqueness of nth positive root *)
 | 
|  |    162 | Goal "[| (0::real) < a; 0 < n |] ==> EX! r. 0 < r & r ^ n = a";
 | 
|  |    163 | by (auto_tac (claset() addSIs [realpow_pos_nth],simpset()));
 | 
|  |    164 | by (cut_inst_tac [("R1.0","r"),("R2.0","y")] real_linear 1);
 | 
|  |    165 | by (Auto_tac);
 | 
|  |    166 | by (dres_inst_tac [("x","r")] (conjI RS realpow_less) 1);
 | 
|  |    167 | by (dres_inst_tac [("x","y")] (conjI RS realpow_less) 3);
 | 
|  |    168 | by (auto_tac (claset(),simpset() addsimps [real_less_not_refl]));
 | 
|  |    169 | qed "realpow_pos_nth_unique";
 | 
|  |    170 | 
 |