src/HOL/Real/RComplete.ML
author paulson
Tue, 16 Jan 2001 12:20:52 +0100
changeset 10919 144ede948e58
parent 10784 27e4d90b35b5
child 11701 3d51fbf81c17
permissions -rw-r--r--
renamings: real_of_nat, real_of_int -> (overloaded) real inf_close -> approx SReal -> Reals SNat -> Nats
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9428
c8eb573114de rename_numerals: use implicit theory context;
wenzelm
parents: 9065
diff changeset
     1
(*  Title       : HOL/Real/RComplete.ML
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 7127
diff changeset
     2
    ID          : $Id$
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     3
    Author      : Jacques D. Fleuriot
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     4
    Copyright   : 1998  University of Cambridge
9428
c8eb573114de rename_numerals: use implicit theory context;
wenzelm
parents: 9065
diff changeset
     5
c8eb573114de rename_numerals: use implicit theory context;
wenzelm
parents: 9065
diff changeset
     6
Completeness theorems for positive reals and reals.
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     7
*) 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     8
5521
7970832271cc added wrapper for bspec
oheimb
parents: 5143
diff changeset
     9
claset_ref() := claset() delWrapper "bspec";
7970832271cc added wrapper for bspec
oheimb
parents: 5143
diff changeset
    10
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10677
diff changeset
    11
Goal "x/#2 + x/#2 = (x::real)";
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10677
diff changeset
    12
by (Simp_tac 1);
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10677
diff changeset
    13
qed "real_sum_of_halves";
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10677
diff changeset
    14
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    15
(*---------------------------------------------------------
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    16
       Completeness of reals: use supremum property of 
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    17
       preal and theorems about real_preal. Theorems 
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    18
       previously in Real.ML. 
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    19
 ---------------------------------------------------------*)
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    20
 (*a few lemmas*)
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    21
Goal "ALL x:P. #0 < x ==> \ 
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    22
\       ((EX x:P. y < x) = (EX X. real_of_preal X : P & \
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    23
\                          y < real_of_preal X))";
9065
15f82c9aa331 full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
    24
by (blast_tac (claset() addSDs [bspec, 
9428
c8eb573114de rename_numerals: use implicit theory context;
wenzelm
parents: 9065
diff changeset
    25
		    rename_numerals real_gt_zero_preal_Ex RS iffD1]) 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    26
qed "real_sup_lemma1";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    27
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    28
Goal "[| ALL x:P. #0 < x; EX x. x: P; EX y. ALL x: P. x < y |] \
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    29
\         ==> (EX X. X: {w. real_of_preal w : P}) & \
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    30
\             (EX Y. ALL X: {w. real_of_preal w : P}. X < Y)";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    31
by (rtac conjI 1);
9065
15f82c9aa331 full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
    32
by (blast_tac (claset() addDs [bspec, 
9428
c8eb573114de rename_numerals: use implicit theory context;
wenzelm
parents: 9065
diff changeset
    33
                rename_numerals real_gt_zero_preal_Ex RS iffD1]) 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    34
by Auto_tac;
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    35
by (dtac bspec 1 THEN assume_tac 1);
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7219
diff changeset
    36
by (ftac bspec 1  THEN assume_tac 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10677
diff changeset
    37
by (dtac order_less_trans 1 THEN assume_tac 1);
9428
c8eb573114de rename_numerals: use implicit theory context;
wenzelm
parents: 9065
diff changeset
    38
by (dtac ((rename_numerals real_gt_zero_preal_Ex) RS iffD1) 1
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
    39
    THEN etac exE 1);    
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    40
by (res_inst_tac [("x","ya")] exI 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    41
by Auto_tac;
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    42
by (dres_inst_tac [("x","real_of_preal X")] bspec 1 THEN assume_tac 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    43
by (etac real_of_preal_lessD 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    44
qed "real_sup_lemma2";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    45
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    46
(*-------------------------------------------------------------
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    47
            Completeness of Positive Reals
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    48
 -------------------------------------------------------------*)
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    49
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
    50
(**
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
    51
 Supremum property for the set of positive reals
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
    52
 FIXME: long proof - should be improved - need 
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
    53
 only have one case split 
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
    54
**)
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
    55
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    56
Goal "[| ALL x:P. (#0::real) < x; EX x. x: P; EX y. ALL x: P. x < y |] \
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    57
\         ==> (EX S. ALL y. (EX x: P. y < x) = (y < S))";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    58
by (res_inst_tac 
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    59
   [("x","real_of_preal (psup({w. real_of_preal w : P}))")] exI 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    60
by Auto_tac;
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7219
diff changeset
    61
by (ftac real_sup_lemma2 1 THEN Auto_tac);
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
    62
by (case_tac "#0 < ya" 1);
9428
c8eb573114de rename_numerals: use implicit theory context;
wenzelm
parents: 9065
diff changeset
    63
by (dtac ((rename_numerals real_gt_zero_preal_Ex) RS iffD1) 1);
c8eb573114de rename_numerals: use implicit theory context;
wenzelm
parents: 9065
diff changeset
    64
by (dtac (rename_numerals real_less_all_real2) 2);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    65
by Auto_tac;
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    66
by (rtac (preal_complete RS spec RS iffD1) 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    67
by Auto_tac;
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7219
diff changeset
    68
by (ftac real_gt_preal_preal_Ex 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    69
by Auto_tac;
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    70
(* second part *)
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    71
by (rtac (real_sup_lemma1 RS iffD2) 1 THEN assume_tac 1);
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
    72
by (case_tac "#0 < ya" 1);
9428
c8eb573114de rename_numerals: use implicit theory context;
wenzelm
parents: 9065
diff changeset
    73
by (auto_tac (claset() addSDs (map rename_numerals
9065
15f82c9aa331 full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
    74
			       [real_less_all_real2,
15f82c9aa331 full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
    75
				real_gt_zero_preal_Ex RS iffD1]),
15f82c9aa331 full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
    76
	      simpset()));
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7219
diff changeset
    77
by (ftac real_sup_lemma2 2 THEN Auto_tac);
23e090051cb8 isatool expandshort;
wenzelm
parents: 7219
diff changeset
    78
by (ftac real_sup_lemma2 1 THEN Auto_tac);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    79
by (rtac (preal_complete RS spec RS iffD2 RS bexE) 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    80
by (Blast_tac 3);
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
    81
by (ALLGOALS(Blast_tac));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    82
qed "posreal_complete";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    83
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    84
(*--------------------------------------------------------
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    85
   Completeness properties using isUb, isLub etc.
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    86
 -------------------------------------------------------*)
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
    87
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 7077
diff changeset
    88
Goal "[| isLub R S x; isLub R S y |] ==> x = (y::real)";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7219
diff changeset
    89
by (ftac isLub_isUb 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    90
by (forw_inst_tac [("x","y")] isLub_isUb 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    91
by (blast_tac (claset() addSIs [real_le_anti_sym]
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 7077
diff changeset
    92
                        addSDs [isLub_le_isUb]) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    93
qed "real_isLub_unique";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    94
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 7077
diff changeset
    95
Goalw [setle_def,setge_def] "[| (x::real) <=* S'; S <= S' |] ==> x <=* S";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    96
by (Blast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    97
qed "real_order_restrict";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    98
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    99
(*----------------------------------------------------------------
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   100
           Completeness theorem for the positive reals(again)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   101
 ----------------------------------------------------------------*)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   102
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   103
Goal "[| ALL x: S. #0 < x; \
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 7077
diff changeset
   104
\        EX x. x: S; \
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 7077
diff changeset
   105
\        EX u. isUb (UNIV::real set) S u \
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 7077
diff changeset
   106
\     |] ==> EX t. isLub (UNIV::real set) S t";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   107
by (res_inst_tac 
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   108
    [("x","real_of_preal(psup({w. real_of_preal w : S}))")] exI 1);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   109
by (auto_tac (claset(), simpset() addsimps 
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   110
    [isLub_def,leastP_def,isUb_def]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   111
by (auto_tac (claset() addSIs [setleI,setgeI] 
9428
c8eb573114de rename_numerals: use implicit theory context;
wenzelm
parents: 9065
diff changeset
   112
	         addSDs [(rename_numerals real_gt_zero_preal_Ex) RS iffD1],
9065
15f82c9aa331 full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
   113
	      simpset()));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   114
by (forw_inst_tac [("x","y")] bspec 1 THEN assume_tac 1);
9428
c8eb573114de rename_numerals: use implicit theory context;
wenzelm
parents: 9065
diff changeset
   115
by (dtac ((rename_numerals real_gt_zero_preal_Ex) RS iffD1) 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
   116
by (auto_tac (claset(), simpset() addsimps [real_of_preal_le_iff]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   117
by (rtac preal_psup_leI2a 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
   118
by (forw_inst_tac [("y","real_of_preal ya")] setleD 1 THEN assume_tac 1);
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7219
diff changeset
   119
by (ftac real_ge_preal_preal_Ex 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   120
by (Step_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   121
by (res_inst_tac [("x","y")] exI 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
   122
by (blast_tac (claset() addSDs [setleD] addIs [real_of_preal_le_iff RS iffD1]) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   123
by (forw_inst_tac [("x","x")] bspec 1 THEN assume_tac 1);
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7219
diff changeset
   124
by (ftac isUbD2 1);
9428
c8eb573114de rename_numerals: use implicit theory context;
wenzelm
parents: 9065
diff changeset
   125
by (dtac ((rename_numerals real_gt_zero_preal_Ex) RS iffD1) 1);
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   126
by (auto_tac (claset() addSDs [isUbD, real_ge_preal_preal_Ex],
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
   127
	      simpset() addsimps [real_of_preal_le_iff]));
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   128
by (blast_tac (claset() addSDs [setleD] addSIs [psup_le_ub1] 
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
   129
	                addIs [real_of_preal_le_iff RS iffD1]) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   130
qed "posreals_complete";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   131
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   132
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   133
(*-------------------------------
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   134
    Lemmas
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   135
 -------------------------------*)
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   136
Goal "ALL y : {z. EX x: P. z = x + (-xa) + #1} Int {x. #0 < x}. #0 < y";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   137
by Auto_tac;
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   138
qed "real_sup_lemma3";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   139
 
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 7077
diff changeset
   140
Goal "(xa <= S + X + (-Z)) = (xa + (-X) + Z <= (S::real))";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   141
by (Auto_tac);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   142
qed "lemma_le_swap2";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   143
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   144
Goal "[| (x::real) + (-X) + #1 <= S; xa <= x |] ==> xa <= S + X + (-#1)";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   145
by (Auto_tac);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   146
qed "lemma_real_complete2b";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   147
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   148
(*----------------------------------------------------------
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   149
      reals Completeness (again!)
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   150
 ----------------------------------------------------------*)
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 7077
diff changeset
   151
Goal "[| EX X. X: S;  EX Y. isUb (UNIV::real set) S Y |]  \
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 7077
diff changeset
   152
\     ==> EX t. isLub (UNIV :: real set) S t";
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   153
by (Step_tac 1);
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   154
by (subgoal_tac "EX u. u: {z. EX x: S. z = x + (-X) + #1} \
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   155
\                Int {x. #0 < x}" 1);
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   156
by (subgoal_tac "isUb (UNIV::real set) ({z. EX x: S. z = x + (-X) + #1} \
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   157
\                Int {x. #0 < x})  (Y + (-X) + #1)" 1); 
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   158
by (cut_inst_tac [("P","S"),("xa","X")] real_sup_lemma3 1);
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 7077
diff changeset
   159
by (EVERY1[forward_tac [exI RSN (3,posreals_complete)], Blast_tac, Blast_tac, 
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 7077
diff changeset
   160
	   Step_tac]);
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   161
by (res_inst_tac [("x","t + X + (-#1)")] exI 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   162
by (rtac isLubI2 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   163
by (rtac setgeI 2 THEN Step_tac 2);
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   164
by (subgoal_tac "isUb (UNIV:: real set) ({z. EX x: S. z = x + (-X) + #1} \
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   165
\                Int {x. #0 < x})  (y + (-X) + #1)" 2); 
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   166
by (dres_inst_tac [("y","(y + (- X) + #1)")] isLub_le_isUb 2 
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   167
      THEN assume_tac 2);
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   168
by (full_simp_tac
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   169
    (simpset() addsimps [real_diff_def, real_diff_le_eq RS sym] @
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   170
                        real_add_ac) 2);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   171
by (rtac (setleI RS isUbI) 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   172
by (Step_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   173
by (res_inst_tac [("R1.0","x"),("R2.0","y")] real_linear_less2 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   174
by (stac lemma_le_swap2 1);
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7219
diff changeset
   175
by (ftac isLubD2 1 THEN assume_tac 2);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   176
by (Step_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   177
by (Blast_tac 1);
10677
36625483213f further round of tidying
paulson
parents: 10606
diff changeset
   178
by (arith_tac 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   179
by (stac lemma_le_swap2 1);
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7219
diff changeset
   180
by (ftac isLubD2 1 THEN assume_tac 2);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   181
by (Blast_tac 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   182
by (rtac lemma_real_complete2b 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10677
diff changeset
   183
by (etac order_less_imp_le 2);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   184
by (blast_tac (claset() addSIs [isLubD2]) 1 THEN Step_tac 1);
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   185
by (full_simp_tac (simpset() addsimps [real_add_assoc]) 1);
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   186
by (blast_tac (claset() addDs [isUbD] addSIs [setleI RS isUbI]
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   187
                        addIs [real_add_le_mono1]) 1);
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   188
by (blast_tac (claset() addDs [isUbD] addSIs [setleI RS isUbI]
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   189
                        addIs [real_add_le_mono1]) 1);
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   190
by (Auto_tac);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   191
qed "reals_complete";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   192
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   193
(*----------------------------------------------------------------
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
   194
        Related: Archimedean property of reals
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   195
 ----------------------------------------------------------------*)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   196
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   197
Goal "#0 < real (Suc n)";
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   198
by (res_inst_tac [("y","real n")] order_le_less_trans 1); 
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   199
by (rtac (rename_numerals real_of_nat_ge_zero) 1);
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   200
by (simp_tac (simpset() addsimps [real_of_nat_Suc]) 1); 
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   201
qed "real_of_nat_Suc_gt_zero";
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   202
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   203
Goal "#0 < x ==> EX n. inverse (real(Suc n)) < x";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   204
by (rtac ccontr 1);
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   205
by (subgoal_tac "ALL n. x * real (Suc n) <= #1" 1);
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   206
by (asm_full_simp_tac
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   207
    (simpset() addsimps [linorder_not_less, real_inverse_eq_divide]) 2); 
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   208
by (Clarify_tac 2);
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   209
by (dres_inst_tac [("x","n")] spec 2); 
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   210
by (dres_inst_tac [("k","real (Suc n)")] (real_mult_le_mono1) 2); 
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   211
by (rtac real_of_nat_ge_zero 2);
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   212
by (asm_full_simp_tac (simpset()  
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   213
	 addsimps [real_of_nat_Suc_gt_zero RS real_not_refl2 RS not_sym, 
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   214
                   real_mult_commute]) 2); 
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
   215
by (subgoal_tac "isUb (UNIV::real set) \
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   216
\                     {z. EX n. z = x*(real (Suc n))} #1" 1);
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   217
by (subgoal_tac "EX X. X : {z. EX n. z = x*(real (Suc n))}" 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   218
by (dtac reals_complete 1);
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   219
by (auto_tac (claset() addIs [isUbI,setleI],simpset()));
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   220
by (subgoal_tac "ALL m. x*(real(Suc m)) <= t" 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   221
by (asm_full_simp_tac (simpset() addsimps 
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   222
                       [real_of_nat_Suc, real_add_mult_distrib2]) 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   223
by (blast_tac (claset() addIs [isLubD2]) 2);
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   224
by (asm_full_simp_tac
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   225
    (simpset() addsimps [real_le_diff_eq RS sym, real_diff_def]) 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5588
diff changeset
   226
by (subgoal_tac "isUb (UNIV::real set) \
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   227
\                  {z. EX n. z = x*(real (Suc n))} (t + (-x))" 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   228
by (blast_tac (claset() addSIs [isUbI,setleI]) 2);
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 7077
diff changeset
   229
by (dres_inst_tac [("y","t+(-x)")] isLub_le_isUb 1);
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   230
by (auto_tac (claset(), 
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   231
	      simpset() addsimps [real_of_nat_Suc,real_add_mult_distrib2]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   232
qed "reals_Archimedean";
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   233
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   234
(*There must be other proofs, e.g. Suc of the largest integer in the
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   235
  cut representing x*)
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   236
Goal "EX n. (x::real) < real (n::nat)";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 7583
diff changeset
   237
by (res_inst_tac [("R1.0","x"),("R2.0","#0")] real_linear_less2 1);
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   238
by (res_inst_tac [("x","0")] exI 1);
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   239
by (res_inst_tac [("x","1")] exI 2);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10677
diff changeset
   240
by (auto_tac (claset() addEs [order_less_trans],
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   241
	      simpset() addsimps [real_of_nat_one]));
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9825
diff changeset
   242
by (ftac ((rename_numerals real_inverse_gt_zero) RS reals_Archimedean) 1);
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   243
by (Step_tac 1 THEN res_inst_tac [("x","Suc n")] exI 1);
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9825
diff changeset
   244
by (forw_inst_tac [("y","inverse x")]
9428
c8eb573114de rename_numerals: use implicit theory context;
wenzelm
parents: 9065
diff changeset
   245
    (rename_numerals real_mult_less_mono1) 1);
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   246
by Auto_tac;  
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   247
by (dres_inst_tac [("y","#1"),("z","real (Suc n)")]
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   248
    (rotate_prems 1 real_mult_less_mono2) 1);
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   249
by (auto_tac (claset(),
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   250
	      simpset() addsimps [real_of_nat_Suc_gt_zero,
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   251
				  real_not_refl2 RS not_sym,
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents: 5521
diff changeset
   252
				  real_mult_assoc RS sym]));
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   253
qed "reals_Archimedean2";