src/HOL/Real/Hyperreal/HyperDef.ML
author paulson
Tue, 19 Dec 2000 15:16:21 +0100
changeset 10702 9e6898befad4
parent 10690 cd80241125b0
child 10712 351ba950d4d9
permissions -rw-r--r--
new simprule zero_less_abs_iff
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
     1
(*  Title       : HOL/Real/Hyperreal/Hyper.ML
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
     2
    ID          : $Id$
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
     3
    Author      : Jacques D. Fleuriot
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
     4
    Copyright   : 1998  University of Cambridge
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
     5
    Description : Ultrapower construction of hyperreals
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
     6
*) 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
     7
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
     8
(*------------------------------------------------------------------------
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
     9
             Proof that the set of naturals is not finite
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    10
 ------------------------------------------------------------------------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    11
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    12
(*** based on James' proof that the set of naturals is not finite ***)
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
    13
Goal "finite (A::nat set) --> (EX n. ALL m. Suc (n + m) ~: A)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    14
by (rtac impI 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    15
by (eres_inst_tac [("F","A")] finite_induct 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    16
by (Blast_tac 1 THEN etac exE 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    17
by (res_inst_tac [("x","n + x")] exI 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    18
by (rtac allI 1 THEN eres_inst_tac [("x","x + m")] allE 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    19
by (auto_tac (claset(), simpset() addsimps add_ac));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    20
by (auto_tac (claset(),
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    21
	      simpset() addsimps [add_assoc RS sym,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    22
				  less_add_Suc2 RS less_not_refl2]));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    23
qed_spec_mp "finite_exhausts";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    24
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
    25
Goal "finite (A :: nat set) --> (EX n. n ~:A)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    26
by (rtac impI 1 THEN dtac finite_exhausts 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    27
by (Blast_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    28
qed_spec_mp "finite_not_covers";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    29
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    30
Goal "~ finite(UNIV:: nat set)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    31
by (fast_tac (claset() addSDs [finite_exhausts]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    32
qed "not_finite_nat";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    33
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    34
(*------------------------------------------------------------------------
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    35
   Existence of free ultrafilter over the naturals and proof of various 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    36
   properties of the FreeUltrafilterNat- an arbitrary free ultrafilter
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    37
 ------------------------------------------------------------------------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    38
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    39
Goal "EX U. U: FreeUltrafilter (UNIV::nat set)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    40
by (rtac (not_finite_nat RS FreeUltrafilter_Ex) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    41
qed "FreeUltrafilterNat_Ex";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    42
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    43
Goalw [FreeUltrafilterNat_def] 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    44
     "FreeUltrafilterNat: FreeUltrafilter(UNIV:: nat set)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    45
by (rtac (FreeUltrafilterNat_Ex RS exE) 1);
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9432
diff changeset
    46
by (rtac someI2 1 THEN ALLGOALS(assume_tac));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    47
qed "FreeUltrafilterNat_mem";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    48
Addsimps [FreeUltrafilterNat_mem];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    49
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    50
Goalw [FreeUltrafilterNat_def] "finite x ==> x ~: FreeUltrafilterNat";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    51
by (rtac (FreeUltrafilterNat_Ex RS exE) 1);
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9432
diff changeset
    52
by (rtac someI2 1 THEN assume_tac 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    53
by (blast_tac (claset() addDs [mem_FreeUltrafiltersetD1]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    54
qed "FreeUltrafilterNat_finite";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    55
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    56
Goal "x: FreeUltrafilterNat ==> ~ finite x";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    57
by (blast_tac (claset() addDs [FreeUltrafilterNat_finite]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    58
qed "FreeUltrafilterNat_not_finite";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    59
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    60
Goalw [FreeUltrafilterNat_def] "{} ~: FreeUltrafilterNat";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    61
by (rtac (FreeUltrafilterNat_Ex RS exE) 1);
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9432
diff changeset
    62
by (rtac someI2 1 THEN assume_tac 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    63
by (blast_tac (claset() addDs [FreeUltrafilter_Ultrafilter,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    64
			       Ultrafilter_Filter,Filter_empty_not_mem]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    65
qed "FreeUltrafilterNat_empty";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    66
Addsimps [FreeUltrafilterNat_empty];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    67
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    68
Goal "[| X: FreeUltrafilterNat;  Y: FreeUltrafilterNat |]  \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    69
\     ==> X Int Y : FreeUltrafilterNat";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    70
by (cut_facts_tac [FreeUltrafilterNat_mem] 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    71
by (blast_tac (claset() addDs [FreeUltrafilter_Ultrafilter,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    72
			       Ultrafilter_Filter,mem_FiltersetD1]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    73
qed "FreeUltrafilterNat_Int";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    74
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    75
Goal "[| X: FreeUltrafilterNat;  X <= Y |] \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    76
\     ==> Y : FreeUltrafilterNat";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    77
by (cut_facts_tac [FreeUltrafilterNat_mem] 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    78
by (blast_tac (claset() addDs [FreeUltrafilter_Ultrafilter,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    79
			       Ultrafilter_Filter,mem_FiltersetD2]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    80
qed "FreeUltrafilterNat_subset";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    81
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    82
Goal "X: FreeUltrafilterNat ==> -X ~: FreeUltrafilterNat";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    83
by (Step_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    84
by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    85
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    86
qed "FreeUltrafilterNat_Compl";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    87
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    88
Goal "X~: FreeUltrafilterNat ==> -X : FreeUltrafilterNat";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    89
by (cut_facts_tac [FreeUltrafilterNat_mem RS (FreeUltrafilter_iff RS iffD1)] 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    90
by (Step_tac 1 THEN dres_inst_tac [("x","X")] bspec 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
    91
by (auto_tac (claset(), simpset() addsimps [UNIV_diff_Compl]));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    92
qed "FreeUltrafilterNat_Compl_mem";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    93
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    94
Goal "(X ~: FreeUltrafilterNat) = (-X: FreeUltrafilterNat)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    95
by (blast_tac (claset() addDs [FreeUltrafilterNat_Compl,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    96
			       FreeUltrafilterNat_Compl_mem]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    97
qed "FreeUltrafilterNat_Compl_iff1";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    98
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
    99
Goal "(X: FreeUltrafilterNat) = (-X ~: FreeUltrafilterNat)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   100
by (auto_tac (claset(),
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   101
	      simpset() addsimps [FreeUltrafilterNat_Compl_iff1 RS sym]));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   102
qed "FreeUltrafilterNat_Compl_iff2";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   103
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   104
Goal "(UNIV::nat set) : FreeUltrafilterNat";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   105
by (rtac (FreeUltrafilterNat_mem RS FreeUltrafilter_Ultrafilter RS 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   106
          Ultrafilter_Filter RS mem_FiltersetD4) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   107
qed "FreeUltrafilterNat_UNIV";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   108
Addsimps [FreeUltrafilterNat_UNIV];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   109
9391
a6ab3a442da6 changed / to // for quotienting; general tidying
paulson
parents: 9387
diff changeset
   110
Goal "UNIV : FreeUltrafilterNat";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   111
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   112
qed "FreeUltrafilterNat_Nat_set";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   113
Addsimps [FreeUltrafilterNat_Nat_set];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   114
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   115
Goal "{n. P(n) = P(n)} : FreeUltrafilterNat";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   116
by (Simp_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   117
qed "FreeUltrafilterNat_Nat_set_refl";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   118
AddIs [FreeUltrafilterNat_Nat_set_refl];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   119
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   120
Goal "{n::nat. P} : FreeUltrafilterNat ==> P";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   121
by (rtac ccontr 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   122
by (rotate_tac 1 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   123
by (Asm_full_simp_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   124
qed "FreeUltrafilterNat_P";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   125
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   126
Goal "{n. P(n)} : FreeUltrafilterNat ==> EX n. P(n)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   127
by (rtac ccontr 1 THEN rotate_tac 1 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   128
by (Asm_full_simp_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   129
qed "FreeUltrafilterNat_Ex_P";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   130
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   131
Goal "ALL n. P(n) ==> {n. P(n)} : FreeUltrafilterNat";
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   132
by (auto_tac (claset() addIs [FreeUltrafilterNat_Nat_set], simpset()));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   133
qed "FreeUltrafilterNat_all";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   134
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   135
(*-------------------------------------------------------
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   136
     Define and use Ultrafilter tactics
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   137
 -------------------------------------------------------*)
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   138
use "fuf.ML";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   139
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   140
(*-------------------------------------------------------
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   141
  Now prove one further property of our free ultrafilter
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   142
 -------------------------------------------------------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   143
Goal "X Un Y: FreeUltrafilterNat \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   144
\     ==> X: FreeUltrafilterNat | Y: FreeUltrafilterNat";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   145
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   146
by (Ultra_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   147
qed "FreeUltrafilterNat_Un";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   148
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   149
(*-------------------------------------------------------
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   150
   Properties of hyprel
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   151
 -------------------------------------------------------*)
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   152
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   153
(** Proving that hyprel is an equivalence relation **)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   154
(** Natural deduction for hyprel **)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   155
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   156
Goalw [hyprel_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   157
   "((X,Y): hyprel) = ({n. X n = Y n}: FreeUltrafilterNat)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   158
by (Fast_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   159
qed "hyprel_iff";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   160
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   161
Goalw [hyprel_def] 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   162
     "{n. X n = Y n}: FreeUltrafilterNat  ==> (X,Y): hyprel";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   163
by (Fast_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   164
qed "hyprelI";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   165
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   166
Goalw [hyprel_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   167
  "p: hyprel --> (EX X Y. \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   168
\                 p = (X,Y) & {n. X n = Y n} : FreeUltrafilterNat)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   169
by (Fast_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   170
qed "hyprelE_lemma";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   171
9432
8b7aad2abcc9 avoid referencing thy value;
wenzelm
parents: 9391
diff changeset
   172
val [major,minor] = goal (the_context ())
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   173
  "[| p: hyprel;  \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   174
\     !!X Y. [| p = (X,Y); {n. X n = Y n}: FreeUltrafilterNat\
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   175
\                    |] ==> Q |] ==> Q";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   176
by (cut_facts_tac [major RS (hyprelE_lemma RS mp)] 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   177
by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   178
qed "hyprelE";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   179
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   180
AddSIs [hyprelI];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   181
AddSEs [hyprelE];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   182
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   183
Goalw [hyprel_def] "(x,x): hyprel";
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   184
by (auto_tac (claset(),
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   185
              simpset() addsimps [FreeUltrafilterNat_Nat_set]));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   186
qed "hyprel_refl";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   187
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   188
Goal "{n. X n = Y n} = {n. Y n = X n}";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   189
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   190
qed "lemma_perm";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   191
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   192
Goalw [hyprel_def] "(x,y): hyprel --> (y,x):hyprel";
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   193
by (auto_tac (claset() addIs [lemma_perm RS subst], simpset()));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   194
qed_spec_mp "hyprel_sym";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   195
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   196
Goalw [hyprel_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   197
      "(x,y): hyprel --> (y,z):hyprel --> (x,z):hyprel";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   198
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   199
by (Ultra_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   200
qed_spec_mp "hyprel_trans";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   201
9391
a6ab3a442da6 changed / to // for quotienting; general tidying
paulson
parents: 9387
diff changeset
   202
Goalw [equiv_def, refl_def, sym_def, trans_def] "equiv UNIV hyprel";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   203
by (auto_tac (claset() addSIs [hyprel_refl] 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   204
                       addSEs [hyprel_sym,hyprel_trans] 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   205
                       delrules [hyprelI,hyprelE],
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   206
	      simpset() addsimps [FreeUltrafilterNat_Nat_set]));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   207
qed "equiv_hyprel";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   208
9391
a6ab3a442da6 changed / to // for quotienting; general tidying
paulson
parents: 9387
diff changeset
   209
(* (hyprel ^^ {x} = hyprel ^^ {y}) = ((x,y) : hyprel) *)
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9071
diff changeset
   210
bind_thm ("equiv_hyprel_iff",
9391
a6ab3a442da6 changed / to // for quotienting; general tidying
paulson
parents: 9387
diff changeset
   211
    	  [equiv_hyprel, UNIV_I, UNIV_I] MRS eq_equiv_class_iff);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   212
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   213
Goalw  [hypreal_def,hyprel_def,quotient_def] "hyprel^^{x}:hypreal";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   214
by (Blast_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   215
qed "hyprel_in_hypreal";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   216
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   217
Goal "inj_on Abs_hypreal hypreal";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   218
by (rtac inj_on_inverseI 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   219
by (etac Abs_hypreal_inverse 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   220
qed "inj_on_Abs_hypreal";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   221
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   222
Addsimps [equiv_hyprel_iff,inj_on_Abs_hypreal RS inj_on_iff,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   223
          hyprel_iff, hyprel_in_hypreal, Abs_hypreal_inverse];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   224
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   225
Addsimps [equiv_hyprel RS eq_equiv_class_iff];
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9071
diff changeset
   226
bind_thm ("eq_hyprelD", equiv_hyprel RSN (2,eq_equiv_class));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   227
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   228
Goal "inj(Rep_hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   229
by (rtac inj_inverseI 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   230
by (rtac Rep_hypreal_inverse 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   231
qed "inj_Rep_hypreal";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   232
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   233
Goalw [hyprel_def] "x: hyprel ^^ {x}";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   234
by (Step_tac 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   235
by (auto_tac (claset() addSIs [FreeUltrafilterNat_Nat_set], simpset()));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   236
qed "lemma_hyprel_refl";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   237
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   238
Addsimps [lemma_hyprel_refl];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   239
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   240
Goalw [hypreal_def] "{} ~: hypreal";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   241
by (auto_tac (claset() addSEs [quotientE], simpset()));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   242
qed "hypreal_empty_not_mem";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   243
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   244
Addsimps [hypreal_empty_not_mem];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   245
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   246
Goal "Rep_hypreal x ~= {}";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   247
by (cut_inst_tac [("x","x")] Rep_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   248
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   249
qed "Rep_hypreal_nonempty";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   250
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   251
Addsimps [Rep_hypreal_nonempty];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   252
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   253
(*------------------------------------------------------------------------
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   254
   hypreal_of_real: the injection from real to hypreal
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   255
 ------------------------------------------------------------------------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   256
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   257
Goal "inj(hypreal_of_real)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   258
by (rtac injI 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   259
by (rewtac hypreal_of_real_def);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   260
by (dtac (inj_on_Abs_hypreal RS inj_onD) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   261
by (REPEAT (rtac hyprel_in_hypreal 1));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   262
by (dtac eq_equiv_class 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   263
by (rtac equiv_hyprel 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   264
by (Fast_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   265
by (rtac ccontr 1 THEN rotate_tac 1 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   266
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   267
qed "inj_hypreal_of_real";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   268
9432
8b7aad2abcc9 avoid referencing thy value;
wenzelm
parents: 9391
diff changeset
   269
val [prem] = goal (the_context ())
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   270
    "(!!x y. z = Abs_hypreal(hyprel^^{x}) ==> P) ==> P";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   271
by (res_inst_tac [("x1","z")] 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   272
    (rewrite_rule [hypreal_def] Rep_hypreal RS quotientE) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   273
by (dres_inst_tac [("f","Abs_hypreal")] arg_cong 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   274
by (res_inst_tac [("x","x")] prem 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   275
by (asm_full_simp_tac (simpset() addsimps [Rep_hypreal_inverse]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   276
qed "eq_Abs_hypreal";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   277
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   278
(**** hypreal_minus: additive inverse on hypreal ****)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   279
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   280
Goalw [congruent_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   281
  "congruent hyprel (%X. hyprel^^{%n. - (X n)})";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   282
by Safe_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   283
by (ALLGOALS Ultra_tac);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   284
qed "hypreal_minus_congruent";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   285
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   286
Goalw [hypreal_minus_def]
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   287
   "- (Abs_hypreal(hyprel^^{%n. X n})) = Abs_hypreal(hyprel ^^ {%n. -(X n)})";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   288
by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   289
by (simp_tac (simpset() addsimps 
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   290
      [hyprel_in_hypreal RS Abs_hypreal_inverse,
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   291
       [equiv_hyprel, hypreal_minus_congruent] MRS UN_equiv_class]) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   292
qed "hypreal_minus";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   293
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   294
Goal "- (- z) = (z::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   295
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   296
by (asm_simp_tac (simpset() addsimps [hypreal_minus]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   297
qed "hypreal_minus_minus";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   298
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   299
Addsimps [hypreal_minus_minus];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   300
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   301
Goal "inj(%r::hypreal. -r)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   302
by (rtac injI 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   303
by (dres_inst_tac [("f","uminus")] arg_cong 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   304
by (asm_full_simp_tac (simpset() addsimps [hypreal_minus_minus]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   305
qed "inj_hypreal_minus";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   306
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   307
Goalw [hypreal_zero_def] "-0 = (0::hypreal)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   308
by (simp_tac (simpset() addsimps [hypreal_minus]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   309
qed "hypreal_minus_zero";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   310
Addsimps [hypreal_minus_zero];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   311
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   312
Goal "(-x = 0) = (x = (0::hypreal))"; 
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   313
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   314
by (auto_tac (claset(),
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   315
       simpset() addsimps [hypreal_zero_def, hypreal_minus, eq_commute] @ 
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   316
                          real_add_ac));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   317
qed "hypreal_minus_zero_iff";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   318
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   319
Addsimps [hypreal_minus_zero_iff];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   320
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   321
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   322
(**** hyperreal addition: hypreal_add  ****)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   323
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   324
Goalw [congruent2_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   325
    "congruent2 hyprel (%X Y. hyprel^^{%n. X n + Y n})";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   326
by Safe_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   327
by (ALLGOALS(Ultra_tac));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   328
qed "hypreal_add_congruent2";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   329
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   330
Goalw [hypreal_add_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   331
  "Abs_hypreal(hyprel^^{%n. X n}) + Abs_hypreal(hyprel^^{%n. Y n}) = \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   332
\  Abs_hypreal(hyprel^^{%n. X n + Y n})";
9391
a6ab3a442da6 changed / to // for quotienting; general tidying
paulson
parents: 9387
diff changeset
   333
by (simp_tac (simpset() addsimps 
a6ab3a442da6 changed / to // for quotienting; general tidying
paulson
parents: 9387
diff changeset
   334
         [[equiv_hyprel, hypreal_add_congruent2] MRS UN_equiv_class2]) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   335
qed "hypreal_add";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   336
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   337
Goal "(z::hypreal) + w = w + z";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   338
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   339
by (res_inst_tac [("z","w")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   340
by (asm_simp_tac (simpset() addsimps (real_add_ac @ [hypreal_add])) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   341
qed "hypreal_add_commute";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   342
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   343
Goal "((z1::hypreal) + z2) + z3 = z1 + (z2 + z3)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   344
by (res_inst_tac [("z","z1")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   345
by (res_inst_tac [("z","z2")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   346
by (res_inst_tac [("z","z3")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   347
by (asm_simp_tac (simpset() addsimps [hypreal_add, real_add_assoc]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   348
qed "hypreal_add_assoc";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   349
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   350
(*For AC rewriting*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   351
Goal "(x::hypreal)+(y+z)=y+(x+z)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   352
by (rtac (hypreal_add_commute RS trans) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   353
by (rtac (hypreal_add_assoc RS trans) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   354
by (rtac (hypreal_add_commute RS arg_cong) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   355
qed "hypreal_add_left_commute";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   356
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   357
(* hypreal addition is an AC operator *)
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9071
diff changeset
   358
bind_thms ("hypreal_add_ac", [hypreal_add_assoc,hypreal_add_commute,
9fff97d29837 bind_thm(s);
wenzelm
parents: 9071
diff changeset
   359
                      hypreal_add_left_commute]);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   360
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   361
Goalw [hypreal_zero_def] "(0::hypreal) + z = z";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   362
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   363
by (asm_full_simp_tac (simpset() addsimps 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   364
    [hypreal_add]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   365
qed "hypreal_add_zero_left";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   366
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   367
Goal "z + (0::hypreal) = z";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   368
by (simp_tac (simpset() addsimps 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   369
    [hypreal_add_zero_left,hypreal_add_commute]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   370
qed "hypreal_add_zero_right";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   371
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   372
Goalw [hypreal_zero_def] "z + -z = (0::hypreal)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   373
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
10648
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10607
diff changeset
   374
by (asm_full_simp_tac (simpset() addsimps [hypreal_minus, hypreal_add]) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   375
qed "hypreal_add_minus";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   376
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   377
Goal "-z + z = (0::hypreal)";
10648
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10607
diff changeset
   378
by (simp_tac (simpset() addsimps [hypreal_add_commute, hypreal_add_minus]) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   379
qed "hypreal_add_minus_left";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   380
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   381
Addsimps [hypreal_add_minus,hypreal_add_minus_left,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   382
          hypreal_add_zero_left,hypreal_add_zero_right];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   383
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   384
Goal "EX y. (x::hypreal) + y = 0";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   385
by (fast_tac (claset() addIs [hypreal_add_minus]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   386
qed "hypreal_minus_ex";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   387
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   388
Goal "EX! y. (x::hypreal) + y = 0";
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   389
by (auto_tac (claset() addIs [hypreal_add_minus], simpset()));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   390
by (dres_inst_tac [("f","%x. ya+x")] arg_cong 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   391
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   392
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   393
qed "hypreal_minus_ex1";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   394
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   395
Goal "EX! y. y + (x::hypreal) = 0";
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   396
by (auto_tac (claset() addIs [hypreal_add_minus_left], simpset()));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   397
by (dres_inst_tac [("f","%x. x+ya")] arg_cong 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   398
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   399
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   400
qed "hypreal_minus_left_ex1";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   401
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   402
Goal "x + y = (0::hypreal) ==> x = -y";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   403
by (cut_inst_tac [("z","y")] hypreal_add_minus_left 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   404
by (res_inst_tac [("x1","y")] (hypreal_minus_left_ex1 RS ex1E) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   405
by (Blast_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   406
qed "hypreal_add_minus_eq_minus";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   407
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   408
Goal "EX y::hypreal. x = -y";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   409
by (cut_inst_tac [("x","x")] hypreal_minus_ex 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   410
by (etac exE 1 THEN dtac hypreal_add_minus_eq_minus 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   411
by (Fast_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   412
qed "hypreal_as_add_inverse_ex";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   413
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   414
Goal "-(x + (y::hypreal)) = -x + -y";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   415
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   416
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10648
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10607
diff changeset
   417
by (auto_tac (claset(),
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10607
diff changeset
   418
              simpset() addsimps [hypreal_minus, hypreal_add,
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10607
diff changeset
   419
                                  real_minus_add_distrib]));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   420
qed "hypreal_minus_add_distrib";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   421
Addsimps [hypreal_minus_add_distrib];
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   422
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   423
Goal "-(y + -(x::hypreal)) = x + -y";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   424
by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   425
qed "hypreal_minus_distrib1";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   426
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   427
Goal "(x + - (y::hypreal)) + (y + - z) = x + -z";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   428
by (res_inst_tac [("w1","y")] (hypreal_add_commute RS subst) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   429
by (simp_tac (simpset() addsimps [hypreal_add_left_commute,
10648
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10607
diff changeset
   430
                                  hypreal_add_assoc]) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   431
by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   432
qed "hypreal_add_minus_cancel1";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   433
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   434
Goal "((x::hypreal) + y = x + z) = (y = z)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   435
by (Step_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   436
by (dres_inst_tac [("f","%t.-x + t")] arg_cong 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   437
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   438
qed "hypreal_add_left_cancel";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   439
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   440
Goal "z + (x + (y + -z)) = x + (y::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   441
by (simp_tac (simpset() addsimps hypreal_add_ac) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   442
qed "hypreal_add_minus_cancel2";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   443
Addsimps [hypreal_add_minus_cancel2];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   444
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   445
Goal "y + -(x + y) = -(x::hypreal)";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   446
by (Full_simp_tac 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   447
by (rtac (hypreal_add_left_commute RS subst) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   448
by (Full_simp_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   449
qed "hypreal_add_minus_cancel";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   450
Addsimps [hypreal_add_minus_cancel];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   451
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   452
Goal "y + -(y + x) = -(x::hypreal)";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   453
by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   454
qed "hypreal_add_minus_cancelc";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   455
Addsimps [hypreal_add_minus_cancelc];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   456
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   457
Goal "(z + -x) + (y + -z) = (y + -(x::hypreal))";
10648
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10607
diff changeset
   458
by (full_simp_tac
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10607
diff changeset
   459
    (simpset() addsimps [hypreal_minus_add_distrib RS sym, 
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10607
diff changeset
   460
                         hypreal_add_left_cancel] @ hypreal_add_ac 
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10607
diff changeset
   461
               delsimps [hypreal_minus_add_distrib]) 1); 
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   462
qed "hypreal_add_minus_cancel3";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   463
Addsimps [hypreal_add_minus_cancel3];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   464
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   465
Goal "(y + (x::hypreal)= z + x) = (y = z)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   466
by (simp_tac (simpset() addsimps [hypreal_add_commute,
10648
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10607
diff changeset
   467
                                  hypreal_add_left_cancel]) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   468
qed "hypreal_add_right_cancel";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   469
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   470
Goal "z + (y + -z) = (y::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   471
by (simp_tac (simpset() addsimps hypreal_add_ac) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   472
qed "hypreal_add_minus_cancel4";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   473
Addsimps [hypreal_add_minus_cancel4];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   474
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   475
Goal "z + (w + (x + (-z + y))) = w + x + (y::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   476
by (simp_tac (simpset() addsimps hypreal_add_ac) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   477
qed "hypreal_add_minus_cancel5";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   478
Addsimps [hypreal_add_minus_cancel5];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   479
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   480
Goal "z + ((- z) + w) = (w::hypreal)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   481
by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   482
qed "hypreal_add_minus_cancelA";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   483
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   484
Goal "(-z) + (z + w) = (w::hypreal)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   485
by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   486
qed "hypreal_minus_add_cancelA";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   487
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   488
Addsimps [hypreal_add_minus_cancelA, hypreal_minus_add_cancelA];
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   489
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   490
(**** hyperreal multiplication: hypreal_mult  ****)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   491
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   492
Goalw [congruent2_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   493
    "congruent2 hyprel (%X Y. hyprel^^{%n. X n * Y n})";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   494
by Safe_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   495
by (ALLGOALS(Ultra_tac));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   496
qed "hypreal_mult_congruent2";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   497
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   498
Goalw [hypreal_mult_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   499
  "Abs_hypreal(hyprel^^{%n. X n}) * Abs_hypreal(hyprel^^{%n. Y n}) = \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   500
\  Abs_hypreal(hyprel^^{%n. X n * Y n})";
9391
a6ab3a442da6 changed / to // for quotienting; general tidying
paulson
parents: 9387
diff changeset
   501
by (simp_tac (simpset() addsimps 
a6ab3a442da6 changed / to // for quotienting; general tidying
paulson
parents: 9387
diff changeset
   502
      [[equiv_hyprel, hypreal_mult_congruent2] MRS UN_equiv_class2]) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   503
qed "hypreal_mult";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   504
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   505
Goal "(z::hypreal) * w = w * z";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   506
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   507
by (res_inst_tac [("z","w")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   508
by (asm_simp_tac (simpset() addsimps ([hypreal_mult] @ real_mult_ac)) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   509
qed "hypreal_mult_commute";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   510
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   511
Goal "((z1::hypreal) * z2) * z3 = z1 * (z2 * z3)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   512
by (res_inst_tac [("z","z1")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   513
by (res_inst_tac [("z","z2")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   514
by (res_inst_tac [("z","z3")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   515
by (asm_simp_tac (simpset() addsimps [hypreal_mult,real_mult_assoc]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   516
qed "hypreal_mult_assoc";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   517
9432
8b7aad2abcc9 avoid referencing thy value;
wenzelm
parents: 9391
diff changeset
   518
qed_goal "hypreal_mult_left_commute" (the_context ())
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   519
    "(z1::hypreal) * (z2 * z3) = z2 * (z1 * z3)"
10648
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10607
diff changeset
   520
 (fn _ => [rtac (hypreal_mult_commute RS trans) 1, 
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10607
diff changeset
   521
           rtac (hypreal_mult_assoc RS trans) 1,
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   522
           rtac (hypreal_mult_commute RS arg_cong) 1]);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   523
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   524
(* hypreal multiplication is an AC operator *)
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9071
diff changeset
   525
bind_thms ("hypreal_mult_ac", [hypreal_mult_assoc, hypreal_mult_commute, 
9fff97d29837 bind_thm(s);
wenzelm
parents: 9071
diff changeset
   526
                       hypreal_mult_left_commute]);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   527
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   528
Goalw [hypreal_one_def] "1hr * z = z";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   529
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   530
by (asm_full_simp_tac (simpset() addsimps [hypreal_mult]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   531
qed "hypreal_mult_1";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   532
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   533
Goal "z * 1hr = z";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   534
by (simp_tac (simpset() addsimps [hypreal_mult_commute,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   535
    hypreal_mult_1]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   536
qed "hypreal_mult_1_right";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   537
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   538
Goalw [hypreal_zero_def] "0 * z = (0::hypreal)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   539
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   540
by (asm_full_simp_tac (simpset() addsimps [hypreal_mult,real_mult_0]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   541
qed "hypreal_mult_0";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   542
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   543
Goal "z * 0 = (0::hypreal)";
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   544
by (simp_tac (simpset() addsimps [hypreal_mult_commute, hypreal_mult_0]) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   545
qed "hypreal_mult_0_right";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   546
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   547
Addsimps [hypreal_mult_0,hypreal_mult_0_right];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   548
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   549
Goal "-(x * y) = -x * (y::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   550
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   551
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   552
by (auto_tac (claset(),
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   553
	      simpset() addsimps [hypreal_minus, hypreal_mult] 
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   554
                                 @ real_mult_ac @ real_add_ac));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   555
qed "hypreal_minus_mult_eq1";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   556
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   557
Goal "-(x * y) = (x::hypreal) * -y";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   558
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   559
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   560
by (auto_tac (claset(), simpset() addsimps [hypreal_minus, hypreal_mult] 
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   561
                                           @ real_mult_ac @ real_add_ac));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   562
qed "hypreal_minus_mult_eq2";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   563
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   564
(*Pull negations out*)
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   565
Addsimps [hypreal_minus_mult_eq2 RS sym, hypreal_minus_mult_eq1 RS sym];
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   566
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   567
Goal "-x*y = (x::hypreal)*-y";
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   568
by Auto_tac;
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   569
qed "hypreal_minus_mult_commute";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   570
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   571
(*-----------------------------------------------------------------------------
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   572
    A few more theorems
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   573
 ----------------------------------------------------------------------------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   574
Goal "(z::hypreal) + v = z' + v' ==> z + (v + w) = z' + (v' + w)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   575
by (asm_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   576
qed "hypreal_add_assoc_cong";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   577
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   578
Goal "(z::hypreal) + (v + w) = v + (z + w)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   579
by (REPEAT (ares_tac [hypreal_add_commute RS hypreal_add_assoc_cong] 1));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   580
qed "hypreal_add_assoc_swap";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   581
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   582
Goal "((z1::hypreal) + z2) * w = (z1 * w) + (z2 * w)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   583
by (res_inst_tac [("z","z1")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   584
by (res_inst_tac [("z","z2")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   585
by (res_inst_tac [("z","w")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   586
by (asm_simp_tac (simpset() addsimps [hypreal_mult,hypreal_add,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   587
     real_add_mult_distrib]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   588
qed "hypreal_add_mult_distrib";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   589
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   590
val hypreal_mult_commute'= read_instantiate [("z","w")] hypreal_mult_commute;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   591
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   592
Goal "(w::hypreal) * (z1 + z2) = (w * z1) + (w * z2)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   593
by (simp_tac (simpset() addsimps [hypreal_mult_commute',hypreal_add_mult_distrib]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   594
qed "hypreal_add_mult_distrib2";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   595
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9071
diff changeset
   596
bind_thms ("hypreal_mult_simps", [hypreal_mult_1, hypreal_mult_1_right]);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   597
Addsimps hypreal_mult_simps;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   598
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   599
(* 07/00 *)
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   600
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   601
Goalw [hypreal_diff_def] "((z1::hypreal) - z2) * w = (z1 * w) - (z2 * w)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   602
by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   603
qed "hypreal_diff_mult_distrib";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   604
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   605
Goal "(w::hypreal) * (z1 - z2) = (w * z1) - (w * z2)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   606
by (simp_tac (simpset() addsimps [hypreal_mult_commute', 
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   607
				  hypreal_diff_mult_distrib]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   608
qed "hypreal_diff_mult_distrib2";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   609
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   610
(*** one and zero are distinct ***)
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   611
Goalw [hypreal_zero_def,hypreal_one_def] "0 ~= 1hr";
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   612
by (auto_tac (claset(), simpset() addsimps [real_zero_not_eq_one]));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   613
qed "hypreal_zero_not_eq_one";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   614
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   615
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   616
(**** multiplicative inverse on hypreal ****)
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   617
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   618
Goalw [congruent_def]
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   619
  "congruent hyprel (%X. hyprel^^{%n. if X n = #0 then #0 else inverse(X n)})";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   620
by (Auto_tac THEN Ultra_tac 1);
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   621
qed "hypreal_inverse_congruent";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   622
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   623
Goalw [hypreal_inverse_def]
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   624
      "inverse (Abs_hypreal(hyprel^^{%n. X n})) = \
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   625
\      Abs_hypreal(hyprel ^^ {%n. if X n = #0 then #0 else inverse(X n)})";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   626
by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   627
by (simp_tac (simpset() addsimps 
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   628
   [hyprel_in_hypreal RS Abs_hypreal_inverse,
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   629
    [equiv_hyprel, hypreal_inverse_congruent] MRS UN_equiv_class]) 1);
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   630
qed "hypreal_inverse";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   631
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   632
Goal "inverse 0 = (0::hypreal)";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   633
by (simp_tac (simpset() addsimps [hypreal_inverse, hypreal_zero_def]) 1);
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   634
qed "HYPREAL_INVERSE_ZERO";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   635
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   636
Goal "a / (0::hypreal) = 0";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   637
by (simp_tac
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   638
    (simpset() addsimps [hypreal_divide_def, HYPREAL_INVERSE_ZERO]) 1);
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   639
qed "HYPREAL_DIVISION_BY_ZERO";  (*NOT for adding to default simpset*)
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   640
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   641
fun hypreal_div_undefined_case_tac s i =
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   642
  case_tac s i THEN 
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   643
  asm_simp_tac 
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   644
       (simpset() addsimps [HYPREAL_DIVISION_BY_ZERO, HYPREAL_INVERSE_ZERO]) i;
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   645
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   646
Goal "inverse (inverse (z::hypreal)) = z";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   647
by (hypreal_div_undefined_case_tac "z=0" 1);
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   648
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   649
by (asm_full_simp_tac (simpset() addsimps 
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   650
                       [hypreal_inverse, hypreal_zero_def]) 1);
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   651
qed "hypreal_inverse_inverse";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   652
Addsimps [hypreal_inverse_inverse];
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   653
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   654
Goalw [hypreal_one_def] "inverse(1hr) = 1hr";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   655
by (full_simp_tac (simpset() addsimps [hypreal_inverse,
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   656
                                       real_zero_not_eq_one RS not_sym]) 1);
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   657
qed "hypreal_inverse_1";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   658
Addsimps [hypreal_inverse_1];
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   659
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   660
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   661
(*** existence of inverse ***)
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   662
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   663
Goalw [hypreal_one_def,hypreal_zero_def] 
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   664
     "x ~= 0 ==> x*inverse(x) = 1hr";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   665
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   666
by (rotate_tac 1 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   667
by (asm_full_simp_tac (simpset() addsimps [hypreal_inverse, hypreal_mult]) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   668
by (dtac FreeUltrafilterNat_Compl_mem 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   669
by (blast_tac (claset() addSIs [real_mult_inv_right,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   670
    FreeUltrafilterNat_subset]) 1);
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   671
qed "hypreal_mult_inverse";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   672
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   673
Goal "x ~= 0 ==> inverse(x)*x = 1hr";
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   674
by (asm_simp_tac (simpset() addsimps [hypreal_mult_inverse,
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   675
				      hypreal_mult_commute]) 1);
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   676
qed "hypreal_mult_inverse_left";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   677
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   678
Goal "(c::hypreal) ~= 0 ==> (c*a=c*b) = (a=b)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   679
by Auto_tac;
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   680
by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1);
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   681
by (asm_full_simp_tac (simpset() addsimps [hypreal_mult_inverse] @ hypreal_mult_ac)  1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   682
qed "hypreal_mult_left_cancel";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   683
    
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   684
Goal "(c::hypreal) ~= 0 ==> (a*c=b*c) = (a=b)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   685
by (Step_tac 1);
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   686
by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1);
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   687
by (asm_full_simp_tac (simpset() addsimps [hypreal_mult_inverse] @ hypreal_mult_ac)  1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   688
qed "hypreal_mult_right_cancel";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   689
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   690
Goalw [hypreal_zero_def] "x ~= 0 ==> inverse (x::hypreal) ~= 0";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   691
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   692
by (asm_full_simp_tac (simpset() addsimps [hypreal_inverse, hypreal_mult]) 1);
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   693
qed "hypreal_inverse_not_zero";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   694
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   695
Addsimps [hypreal_mult_inverse,hypreal_mult_inverse_left];
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   696
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   697
Goal "[| x ~= 0; y ~= 0 |] ==> x * y ~= (0::hypreal)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   698
by (Step_tac 1);
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   699
by (dres_inst_tac [("f","%z. inverse x*z")] arg_cong 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   700
by (asm_full_simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   701
qed "hypreal_mult_not_0";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   702
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   703
bind_thm ("hypreal_mult_not_0E",hypreal_mult_not_0 RS notE);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   704
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   705
Goal "x*y = (0::hypreal) ==> x = 0 | y = 0";
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   706
by (auto_tac (claset() addIs [ccontr] addDs [hypreal_mult_not_0],
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   707
              simpset()));
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   708
qed "hypreal_mult_zero_disj";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   709
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   710
Goal "x ~= 0 ==> x * x ~= (0::hypreal)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   711
by (blast_tac (claset() addDs [hypreal_mult_not_0]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   712
qed "hypreal_mult_self_not_zero";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   713
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   714
Goal "[| x ~= 0; y ~= 0 |] ==> inverse(x*y) = inverse(x)*inverse(y::hypreal)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   715
by (res_inst_tac [("c1","x")] (hypreal_mult_left_cancel RS iffD1) 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   716
by (auto_tac (claset(),
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   717
        simpset() addsimps [hypreal_mult_assoc RS sym, hypreal_mult_not_0]));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   718
by (res_inst_tac [("c1","y")] (hypreal_mult_right_cancel RS iffD1) 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   719
by (auto_tac (claset(),
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   720
        simpset() addsimps [hypreal_mult_not_0] @ hypreal_mult_ac));
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   721
by (auto_tac (claset(),
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   722
        simpset() addsimps [hypreal_mult_assoc RS sym, hypreal_mult_not_0]));
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   723
qed "inverse_mult_eq";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   724
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   725
Goal "inverse(-x) = -inverse(x::hypreal)";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   726
by (hypreal_div_undefined_case_tac "x=0" 1);
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   727
by (rtac (hypreal_mult_right_cancel RS iffD1) 1);
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   728
by (stac hypreal_mult_inverse_left 2);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   729
by Auto_tac;
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   730
qed "hypreal_minus_inverse";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   731
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   732
Goal "inverse(x*y) = inverse(x)*inverse(y::hypreal)";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   733
by (hypreal_div_undefined_case_tac "x=0" 1);
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   734
by (hypreal_div_undefined_case_tac "y=0" 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   735
by (forw_inst_tac [("y","y")] hypreal_mult_not_0 1 THEN assume_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   736
by (res_inst_tac [("c1","x")] (hypreal_mult_left_cancel RS iffD1) 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   737
by (auto_tac (claset(), simpset() addsimps [hypreal_mult_assoc RS sym]));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   738
by (res_inst_tac [("c1","y")] (hypreal_mult_left_cancel RS iffD1) 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   739
by (auto_tac (claset(), simpset() addsimps [hypreal_mult_left_commute]));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   740
by (asm_simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym]) 1);
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   741
qed "hypreal_inverse_distrib";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   742
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   743
(*------------------------------------------------------------------
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   744
                   Theorems for ordering 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   745
 ------------------------------------------------------------------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   746
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   747
(* prove introduction and elimination rules for hypreal_less *)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   748
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   749
Goalw [hypreal_less_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   750
 "P < (Q::hypreal) = (EX X Y. X : Rep_hypreal(P) & \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   751
\                             Y : Rep_hypreal(Q) & \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   752
\                             {n. X n < Y n} : FreeUltrafilterNat)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   753
by (Fast_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   754
qed "hypreal_less_iff";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   755
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   756
Goalw [hypreal_less_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   757
 "[| {n. X n < Y n} : FreeUltrafilterNat; \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   758
\         X : Rep_hypreal(P); \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   759
\         Y : Rep_hypreal(Q) |] ==> P < (Q::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   760
by (Fast_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   761
qed "hypreal_lessI";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   762
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   763
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   764
Goalw [hypreal_less_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   765
     "!! R1. [| R1 < (R2::hypreal); \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   766
\         !!X Y. {n. X n < Y n} : FreeUltrafilterNat ==> P; \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   767
\         !!X. X : Rep_hypreal(R1) ==> P; \ 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   768
\         !!Y. Y : Rep_hypreal(R2) ==> P |] \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   769
\     ==> P";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   770
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   771
qed "hypreal_lessE";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   772
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   773
Goalw [hypreal_less_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   774
 "R1 < (R2::hypreal) ==> (EX X Y. {n. X n < Y n} : FreeUltrafilterNat & \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   775
\                                  X : Rep_hypreal(R1) & \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   776
\                                  Y : Rep_hypreal(R2))";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   777
by (Fast_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   778
qed "hypreal_lessD";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   779
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   780
Goal "~ (R::hypreal) < R";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   781
by (res_inst_tac [("z","R")] eq_Abs_hypreal 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   782
by (auto_tac (claset(), simpset() addsimps [hypreal_less_def]));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   783
by (Ultra_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   784
qed "hypreal_less_not_refl";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   785
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   786
(*** y < y ==> P ***)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   787
bind_thm("hypreal_less_irrefl",hypreal_less_not_refl RS notE);
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   788
AddSEs [hypreal_less_irrefl];
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   789
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   790
Goal "!!(x::hypreal). x < y ==> x ~= y";
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   791
by (auto_tac (claset(), simpset() addsimps [hypreal_less_not_refl]));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   792
qed "hypreal_not_refl2";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   793
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   794
Goal "!!(R1::hypreal). [| R1 < R2; R2 < R3 |] ==> R1 < R3";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   795
by (res_inst_tac [("z","R1")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   796
by (res_inst_tac [("z","R2")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   797
by (res_inst_tac [("z","R3")] eq_Abs_hypreal 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   798
by (auto_tac (claset() addSIs [exI], simpset() addsimps [hypreal_less_def]));
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   799
by (ultra_tac (claset() addIs [real_less_trans], simpset()) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   800
qed "hypreal_less_trans";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   801
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   802
Goal "!! (R1::hypreal). [| R1 < R2; R2 < R1 |] ==> P";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   803
by (dtac hypreal_less_trans 1 THEN assume_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   804
by (asm_full_simp_tac (simpset() addsimps 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   805
    [hypreal_less_not_refl]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   806
qed "hypreal_less_asym";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   807
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   808
(*-------------------------------------------------------
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   809
  TODO: The following theorem should have been proved 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   810
  first and then used througout the proofs as it probably 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   811
  makes many of them more straightforward. 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   812
 -------------------------------------------------------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   813
Goalw [hypreal_less_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   814
      "(Abs_hypreal(hyprel^^{%n. X n}) < \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   815
\           Abs_hypreal(hyprel^^{%n. Y n})) = \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   816
\      ({n. X n < Y n} : FreeUltrafilterNat)";
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   817
by (auto_tac (claset() addSIs [lemma_hyprel_refl], simpset()));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   818
by (Ultra_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   819
qed "hypreal_less";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   820
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   821
(*---------------------------------------------------------------------------------
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   822
             Hyperreals as a linearly ordered field
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   823
 ---------------------------------------------------------------------------------*)
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   824
(*** sum order 
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   825
Goalw [hypreal_zero_def] 
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   826
      "[| 0 < x; 0 < y |] ==> (0::hypreal) < x + y";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   827
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   828
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   829
by (auto_tac (claset(),simpset() addsimps
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   830
    [hypreal_less_def,hypreal_add]));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   831
by (auto_tac (claset() addSIs [exI],simpset() addsimps
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   832
    [hypreal_less_def,hypreal_add]));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   833
by (ultra_tac (claset() addIs [real_add_order],simpset()) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   834
qed "hypreal_add_order";
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   835
***)
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   836
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
   837
(*** mult order 
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   838
Goalw [hypreal_zero_def] 
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   839
          "[| 0 < x; 0 < y |] ==> (0::hypreal) < x * y";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   840
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   841
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   842
by (auto_tac (claset() addSIs [exI],simpset() addsimps
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   843
    [hypreal_less_def,hypreal_mult]));
9432
8b7aad2abcc9 avoid referencing thy value;
wenzelm
parents: 9391
diff changeset
   844
by (ultra_tac (claset() addIs [rename_numerals real_mult_order],
9071
paulson
parents: 9055
diff changeset
   845
	       simpset()) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   846
qed "hypreal_mult_order";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   847
****)
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   848
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   849
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   850
(*---------------------------------------------------------------------------------
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   851
                         Trichotomy of the hyperreals
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   852
  --------------------------------------------------------------------------------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   853
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   854
Goalw [hyprel_def] "EX x. x: hyprel ^^ {%n. #0}";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8856
diff changeset
   855
by (res_inst_tac [("x","%n. #0")] exI 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   856
by (Step_tac 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   857
by (auto_tac (claset() addSIs [FreeUltrafilterNat_Nat_set], simpset()));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   858
qed "lemma_hyprel_0r_mem";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   859
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   860
Goalw [hypreal_zero_def]"0 <  x | x = 0 | x < (0::hypreal)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   861
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   862
by (auto_tac (claset(),simpset() addsimps [hypreal_less_def]));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   863
by (cut_facts_tac [lemma_hyprel_0r_mem] 1 THEN etac exE 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   864
by (dres_inst_tac [("x","xa")] spec 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   865
by (dres_inst_tac [("x","x")] spec 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   866
by (cut_inst_tac [("x","x")] lemma_hyprel_refl 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   867
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   868
by (dres_inst_tac [("x","x")] spec 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   869
by (dres_inst_tac [("x","xa")] spec 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   870
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   871
by (Ultra_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   872
by (auto_tac (claset() addIs [real_linear_less2],simpset()));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   873
qed "hypreal_trichotomy";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   874
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   875
val prems = Goal "[| (0::hypreal) < x ==> P; \
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   876
\                 x = 0 ==> P; \
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   877
\                 x < 0 ==> P |] ==> P";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   878
by (cut_inst_tac [("x","x")] hypreal_trichotomy 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   879
by (REPEAT (eresolve_tac (disjE::prems) 1));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   880
qed "hypreal_trichotomyE";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   881
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   882
(*----------------------------------------------------------------------------
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   883
            More properties of <
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   884
 ----------------------------------------------------------------------------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   885
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   886
Goal "((x::hypreal) < y) = (0 < y + -x)";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   887
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   888
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   889
by (auto_tac (claset(),simpset() addsimps [hypreal_add,
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   890
    hypreal_zero_def,hypreal_minus,hypreal_less]));
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   891
by (ALLGOALS(Ultra_tac));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   892
qed "hypreal_less_minus_iff"; 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   893
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   894
Goal "((x::hypreal) < y) = (x + -y < 0)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   895
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   896
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   897
by (auto_tac (claset(),simpset() addsimps [hypreal_add,
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   898
    hypreal_zero_def,hypreal_minus,hypreal_less]));
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   899
by (ALLGOALS(Ultra_tac));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   900
qed "hypreal_less_minus_iff2";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   901
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   902
Goal "((x::hypreal) = y) = (0 = x + - y)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   903
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   904
by (res_inst_tac [("x1","-y")] (hypreal_add_right_cancel RS iffD1) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   905
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   906
qed "hypreal_eq_minus_iff"; 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   907
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   908
Goal "((x::hypreal) = y) = (0 = y + - x)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   909
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   910
by (res_inst_tac [("x1","-x")] (hypreal_add_right_cancel RS iffD1) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   911
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   912
qed "hypreal_eq_minus_iff2"; 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   913
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   914
(* 07/00 *)
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   915
Goal "(0::hypreal) - x = -x";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   916
by (simp_tac (simpset() addsimps [hypreal_diff_def]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   917
qed "hypreal_diff_zero";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   918
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   919
Goal "x - (0::hypreal) = x";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   920
by (simp_tac (simpset() addsimps [hypreal_diff_def]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   921
qed "hypreal_diff_zero_right";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   922
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   923
Goal "x - x = (0::hypreal)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   924
by (simp_tac (simpset() addsimps [hypreal_diff_def]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   925
qed "hypreal_diff_self";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   926
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   927
Addsimps [hypreal_diff_zero, hypreal_diff_zero_right, hypreal_diff_self];
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   928
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   929
Goal "(x = y + z) = (x + -z = (y::hypreal))";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   930
by (auto_tac (claset(),simpset() addsimps [hypreal_add_assoc]));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   931
qed "hypreal_eq_minus_iff3";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   932
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   933
Goal "(x = z + y) = (x + -z = (y::hypreal))";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   934
by (auto_tac (claset(),simpset() addsimps hypreal_add_ac));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   935
qed "hypreal_eq_minus_iff4";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   936
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
   937
Goal "(x ~= a) = (x + -a ~= (0::hypreal))";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   938
by (auto_tac (claset() addDs [sym RS 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   939
    (hypreal_eq_minus_iff RS iffD2)],simpset())); 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   940
qed "hypreal_not_eq_minus_iff";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   941
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   942
(*** linearity ***)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   943
Goal "(x::hypreal) < y | x = y | y < x";
7322
d16d7ddcc842 isatool expandshort;
wenzelm
parents: 7218
diff changeset
   944
by (stac hypreal_eq_minus_iff2 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   945
by (res_inst_tac [("x1","x")] (hypreal_less_minus_iff RS ssubst) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   946
by (res_inst_tac [("x1","y")] (hypreal_less_minus_iff2 RS ssubst) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   947
by (rtac hypreal_trichotomyE 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   948
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   949
qed "hypreal_linear";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   950
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   951
Goal "((w::hypreal) ~= z) = (w<z | z<w)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   952
by (cut_facts_tac [hypreal_linear] 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   953
by (Blast_tac 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   954
qed "hypreal_neq_iff";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
   955
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   956
Goal "!!(x::hypreal). [| x < y ==> P;  x = y ==> P; \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   957
\          y < x ==> P |] ==> P";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   958
by (cut_inst_tac [("x","x"),("y","y")] hypreal_linear 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   959
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   960
qed "hypreal_linear_less2";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   961
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   962
(*------------------------------------------------------------------------------
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   963
                            Properties of <=
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   964
 ------------------------------------------------------------------------------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   965
(*------ hypreal le iff reals le a.e ------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   966
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   967
Goalw [hypreal_le_def,real_le_def]
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   968
      "(Abs_hypreal(hyprel^^{%n. X n}) <= \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   969
\           Abs_hypreal(hyprel^^{%n. Y n})) = \
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   970
\      ({n. X n <= Y n} : FreeUltrafilterNat)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   971
by (auto_tac (claset(),simpset() addsimps [hypreal_less]));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   972
by (ALLGOALS(Ultra_tac));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   973
qed "hypreal_le";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   974
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   975
(*---------------------------------------------------------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   976
(*---------------------------------------------------------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   977
Goalw [hypreal_le_def] 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   978
     "~(w < z) ==> z <= (w::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   979
by (assume_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   980
qed "hypreal_leI";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   981
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   982
Goalw [hypreal_le_def] 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   983
      "z<=w ==> ~(w<(z::hypreal))";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   984
by (assume_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   985
qed "hypreal_leD";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   986
9108
9fff97d29837 bind_thm(s);
wenzelm
parents: 9071
diff changeset
   987
bind_thm ("hypreal_leE", make_elim hypreal_leD);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   988
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   989
Goal "(~(w < z)) = (z <= (w::hypreal))";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   990
by (fast_tac (claset() addSIs [hypreal_leI,hypreal_leD]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   991
qed "hypreal_less_le_iff";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   992
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   993
Goalw [hypreal_le_def] "~ z <= w ==> w<(z::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   994
by (Fast_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   995
qed "not_hypreal_leE";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   996
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   997
Goalw [hypreal_le_def] "z < w ==> z <= (w::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   998
by (fast_tac (claset() addEs [hypreal_less_asym]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
   999
qed "hypreal_less_imp_le";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1000
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1001
Goalw [hypreal_le_def] "!!(x::hypreal). x <= y ==> x < y | x = y";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1002
by (cut_facts_tac [hypreal_linear] 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1003
by (fast_tac (claset() addEs [hypreal_less_irrefl,hypreal_less_asym]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1004
qed "hypreal_le_imp_less_or_eq";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1005
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1006
Goalw [hypreal_le_def] "z<w | z=w ==> z <=(w::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1007
by (cut_facts_tac [hypreal_linear] 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1008
by (fast_tac (claset() addEs [hypreal_less_irrefl,hypreal_less_asym]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1009
qed "hypreal_less_or_eq_imp_le";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1010
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1011
Goal "(x <= (y::hypreal)) = (x < y | x=y)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1012
by (REPEAT(ares_tac [iffI, hypreal_less_or_eq_imp_le, hypreal_le_imp_less_or_eq] 1));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1013
qed "hypreal_le_eq_less_or_eq";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1014
val hypreal_le_less = hypreal_le_eq_less_or_eq;
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1015
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1016
Goal "w <= (w::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1017
by (simp_tac (simpset() addsimps [hypreal_le_eq_less_or_eq]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1018
qed "hypreal_le_refl";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1019
Addsimps [hypreal_le_refl];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1020
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1021
(* Axiom 'linorder_linear' of class 'linorder': *)
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1022
Goal "(z::hypreal) <= w | w <= z";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1023
by (simp_tac (simpset() addsimps [hypreal_le_less]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1024
by (cut_facts_tac [hypreal_linear] 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1025
by (Blast_tac 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1026
qed "hypreal_le_linear";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1027
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1028
Goal "[| i <= j; j < k |] ==> i < (k::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1029
by (dtac hypreal_le_imp_less_or_eq 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1030
by (fast_tac (claset() addIs [hypreal_less_trans]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1031
qed "hypreal_le_less_trans";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1032
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1033
Goal "!! (i::hypreal). [| i < j; j <= k |] ==> i < k";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1034
by (dtac hypreal_le_imp_less_or_eq 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1035
by (fast_tac (claset() addIs [hypreal_less_trans]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1036
qed "hypreal_less_le_trans";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1037
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1038
Goal "[| i <= j; j <= k |] ==> i <= (k::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1039
by (EVERY1 [dtac hypreal_le_imp_less_or_eq, dtac hypreal_le_imp_less_or_eq,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1040
            rtac hypreal_less_or_eq_imp_le, fast_tac (claset() addIs [hypreal_less_trans])]);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1041
qed "hypreal_le_trans";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1042
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1043
Goal "[| z <= w; w <= z |] ==> z = (w::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1044
by (EVERY1 [dtac hypreal_le_imp_less_or_eq, dtac hypreal_le_imp_less_or_eq,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1045
            fast_tac (claset() addEs [hypreal_less_irrefl,hypreal_less_asym])]);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1046
qed "hypreal_le_anti_sym";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1047
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1048
Goal "[| ~ y < x; y ~= x |] ==> x < (y::hypreal)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1049
by (rtac not_hypreal_leE 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1050
by (fast_tac (claset() addDs [hypreal_le_imp_less_or_eq]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1051
qed "not_less_not_eq_hypreal_less";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1052
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1053
(* Axiom 'order_less_le' of class 'order': *)
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1054
Goal "(w::hypreal) < z = (w <= z & w ~= z)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1055
by (simp_tac (simpset() addsimps [hypreal_le_def, hypreal_neq_iff]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1056
by (blast_tac (claset() addIs [hypreal_less_asym]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1057
qed "hypreal_less_le";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1058
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
  1059
Goal "(0 < -R) = (R < (0::hypreal))";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1060
by (res_inst_tac [("z","R")] eq_Abs_hypreal 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1061
by (auto_tac (claset(),simpset() addsimps [hypreal_zero_def,
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1062
    hypreal_less,hypreal_minus]));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1063
qed "hypreal_minus_zero_less_iff";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1064
Addsimps [hypreal_minus_zero_less_iff];
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1065
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
  1066
Goal "(-R < 0) = ((0::hypreal) < R)";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1067
by (res_inst_tac [("z","R")] eq_Abs_hypreal 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1068
by (auto_tac (claset(),simpset() addsimps [hypreal_zero_def,
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1069
    hypreal_less,hypreal_minus]));
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1070
by (ALLGOALS(Ultra_tac));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1071
qed "hypreal_minus_zero_less_iff2";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1072
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1073
Goalw [hypreal_le_def] "((0::hypreal) <= -r) = (r <= (0::hypreal))";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1074
by (simp_tac (simpset() addsimps 
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1075
    [hypreal_minus_zero_less_iff2]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1076
qed "hypreal_minus_zero_le_iff";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1077
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1078
(*----------------------------------------------------------
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1079
  hypreal_of_real preserves field and order properties
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1080
 -----------------------------------------------------------*)
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1081
Goalw [hypreal_of_real_def] 
9071
paulson
parents: 9055
diff changeset
  1082
     "hypreal_of_real (z1 + z2) = \
paulson
parents: 9055
diff changeset
  1083
\     hypreal_of_real z1 + hypreal_of_real z2";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1084
by (asm_simp_tac (simpset() addsimps [hypreal_add,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1085
       hypreal_add_mult_distrib]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1086
qed "hypreal_of_real_add";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1087
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1088
Goalw [hypreal_of_real_def] 
9071
paulson
parents: 9055
diff changeset
  1089
     "hypreal_of_real (z1 * z2) = hypreal_of_real z1 * hypreal_of_real z2";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1090
by (full_simp_tac (simpset() addsimps [hypreal_mult,
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1091
        hypreal_add_mult_distrib2]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1092
qed "hypreal_of_real_mult";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1093
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1094
Goalw [hypreal_less_def,hypreal_of_real_def] 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1095
            "(z1 < z2) = (hypreal_of_real z1 <  hypreal_of_real z2)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1096
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1097
by (res_inst_tac [("x","%n. z1")] exI 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1098
by (Step_tac 1); 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1099
by (res_inst_tac [("x","%n. z2")] exI 2);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1100
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1101
by (rtac FreeUltrafilterNat_P 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1102
by (Ultra_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1103
qed "hypreal_of_real_less_iff";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1104
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1105
Addsimps [hypreal_of_real_less_iff RS sym];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1106
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1107
Goalw [hypreal_le_def,real_le_def] 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1108
            "(z1 <= z2) = (hypreal_of_real z1 <=  hypreal_of_real z2)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1109
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1110
qed "hypreal_of_real_le_iff";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1111
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1112
Goalw [hypreal_of_real_def] "hypreal_of_real (-r) = - hypreal_of_real  r";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1113
by (auto_tac (claset(),simpset() addsimps [hypreal_minus]));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1114
qed "hypreal_of_real_minus";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1115
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1116
(*DON'T insert this or the next one as default simprules.
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1117
  They are used in both orientations and anyway aren't the ones we finally
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1118
  need, which would use binary literals.*)
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8856
diff changeset
  1119
Goalw [hypreal_of_real_def,hypreal_one_def] "hypreal_of_real  #1 = 1hr";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1120
by (Step_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1121
qed "hypreal_of_real_one";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1122
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1123
Goalw [hypreal_of_real_def,hypreal_zero_def] "hypreal_of_real #0 = 0";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1124
by (Step_tac 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1125
qed "hypreal_of_real_zero";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1126
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1127
Goal "(hypreal_of_real r = 0) = (r = #0)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1128
by (auto_tac (claset() addIs [FreeUltrafilterNat_P],
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1129
    simpset() addsimps [hypreal_of_real_def,
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1130
                        hypreal_zero_def,FreeUltrafilterNat_Nat_set]));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1131
qed "hypreal_of_real_zero_iff";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1132
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1133
(*FIXME: delete*)
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1134
Goal "(hypreal_of_real r ~= 0) = (r ~= #0)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1135
by (full_simp_tac (simpset() addsimps [hypreal_of_real_zero_iff]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1136
qed "hypreal_of_real_not_zero_iff";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1137
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1138
Goal "inverse (hypreal_of_real r) = hypreal_of_real (inverse r)";
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1139
by (case_tac "r=#0" 1);
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1140
by (asm_simp_tac (simpset() addsimps [REAL_DIVIDE_ZERO, INVERSE_ZERO, 
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1141
                              HYPREAL_INVERSE_ZERO, hypreal_of_real_zero]) 1);
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1142
by (res_inst_tac [("c1","hypreal_of_real r")] 
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1143
    (hypreal_mult_left_cancel RS iffD1) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1144
by (etac (hypreal_of_real_not_zero_iff RS iffD2) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1145
by (forward_tac [hypreal_of_real_not_zero_iff RS iffD2] 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1146
by (auto_tac (claset(),
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1147
      simpset() addsimps [hypreal_of_real_one, hypreal_of_real_mult RS sym]));
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
  1148
qed "hypreal_of_real_inverse";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1149
10690
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1150
Goal "hypreal_of_real (z1 / z2) = hypreal_of_real z1 / hypreal_of_real z2";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1151
by (simp_tac (simpset() addsimps [hypreal_divide_def, real_divide_def, 
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1152
                          hypreal_of_real_mult, hypreal_of_real_inverse]) 1);
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1153
qed "hypreal_of_real_divide"; 
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1154
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1155
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1156
(*** Division lemmas ***)
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1157
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1158
Goal "(0::hypreal)/x = 0";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1159
by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); 
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1160
qed "hypreal_zero_divide";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1161
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1162
Goal "x/1hr = x";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1163
by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); 
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1164
qed "hypreal_divide_one";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1165
Addsimps [hypreal_zero_divide, hypreal_divide_one];
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1166
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1167
Goal "(x::hypreal) * (y/z) = (x*y)/z";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1168
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 1); 
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1169
qed "hypreal_times_divide1_eq";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1170
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1171
Goal "(y/z) * (x::hypreal) = (y*x)/z";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1172
by (simp_tac (simpset() addsimps [hypreal_divide_def]@hypreal_mult_ac) 1); 
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1173
qed "hypreal_times_divide2_eq";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1174
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1175
Addsimps [hypreal_times_divide1_eq, hypreal_times_divide2_eq];
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1176
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1177
Goal "(x::hypreal) / (y/z) = (x*z)/y";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1178
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_inverse_distrib]@
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1179
                                 hypreal_mult_ac) 1); 
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1180
qed "hypreal_divide_divide1_eq";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1181
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1182
Goal "((x::hypreal) / y) / z = x/(y*z)";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1183
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_inverse_distrib, 
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1184
                                  hypreal_mult_assoc]) 1); 
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1185
qed "hypreal_divide_divide2_eq";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1186
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1187
Addsimps [hypreal_divide_divide1_eq, hypreal_divide_divide2_eq];
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1188
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1189
(** As with multiplication, pull minus signs OUT of the / operator **)
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1190
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1191
Goal "(-x) / (y::hypreal) = - (x/y)";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1192
by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); 
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1193
qed "hypreal_minus_divide_eq";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1194
Addsimps [hypreal_minus_divide_eq];
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1195
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1196
Goal "(x / -(y::hypreal)) = - (x/y)";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1197
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_minus_inverse]) 1); 
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1198
qed "hypreal_divide_minus_eq";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1199
Addsimps [hypreal_divide_minus_eq];
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1200
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1201
Goal "(x+y)/(z::hypreal) = x/z + y/z";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1202
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_add_mult_distrib]) 1); 
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1203
qed "hypreal_add_divide_distrib";
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1204
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
  1205
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1206
Goal "x+x=x*(1hr+1hr)";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1207
by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib2]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1208
qed "hypreal_add_self";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1209
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
  1210
(*FIXME: DELETE (used in Lim.ML) *)
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
  1211
Goal "(z::hypreal) ~= 0 ==> x*y = (x*inverse(z))*(z*y)";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1212
by (asm_simp_tac (simpset() addsimps hypreal_mult_ac)  1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1213
qed "lemma_chain";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1214
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
  1215
Goal "[|(x::hypreal) ~= 0; y ~= 0 |] ==> \
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
  1216
\                   inverse(x) + inverse(y) = (x + y)*inverse(x*y)";
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
  1217
by (asm_full_simp_tac (simpset() addsimps [hypreal_inverse_distrib,
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1218
             hypreal_add_mult_distrib,hypreal_mult_assoc RS sym]) 1);
7322
d16d7ddcc842 isatool expandshort;
wenzelm
parents: 7218
diff changeset
  1219
by (stac hypreal_mult_assoc 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1220
by (rtac (hypreal_mult_left_commute RS subst) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1221
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
10607
352f6f209775 converted rinv and hrinv to inverse;
bauerg
parents: 10043
diff changeset
  1222
qed "hypreal_inverse_add";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1223
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
  1224
Goal "x = -x ==> x = (0::hypreal)";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1225
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1226
by (auto_tac (claset(),simpset() addsimps [hypreal_minus,
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1227
    hypreal_zero_def]));
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1228
by (Ultra_tac 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1229
qed "hypreal_self_eq_minus_self_zero";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1230
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
  1231
Goal "(x + x = 0) = (x = (0::hypreal))";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1232
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1233
by (auto_tac (claset(),simpset() addsimps [hypreal_add,
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1234
    hypreal_zero_def]));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1235
qed "hypreal_add_self_zero_cancel";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1236
Addsimps [hypreal_add_self_zero_cancel];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1237
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
  1238
Goal "(x + x + y = y) = (x = (0::hypreal))";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1239
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1240
by (dtac (hypreal_eq_minus_iff RS iffD1) 1 THEN dtac sym 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1241
by (auto_tac (claset(),simpset() addsimps [hypreal_add_assoc]));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1242
qed "hypreal_add_self_zero_cancel2";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1243
Addsimps [hypreal_add_self_zero_cancel2];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1244
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
  1245
Goal "(x + (x + y) = y) = (x = (0::hypreal))";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1246
by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1247
qed "hypreal_add_self_zero_cancel2a";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1248
Addsimps [hypreal_add_self_zero_cancel2a];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1249
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1250
Goal "(b = -a) = (-b = (a::hypreal))";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1251
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1252
qed "hypreal_minus_eq_swap";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1253
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1254
Goal "(-b = -a) = (b = (a::hypreal))";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1255
by (asm_full_simp_tac (simpset() addsimps 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1256
    [hypreal_minus_eq_swap]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1257
qed "hypreal_minus_eq_cancel";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1258
Addsimps [hypreal_minus_eq_cancel];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1259
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1260
Goal "x < x + 1hr";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1261
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1262
by (auto_tac (claset(),simpset() addsimps [hypreal_add,
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1263
    hypreal_one_def,hypreal_less]));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1264
qed "hypreal_less_self_add_one";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1265
Addsimps [hypreal_less_self_add_one];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1266
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1267
Goal "((x::hypreal) + x = y + y) = (x = y)";
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1268
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1269
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1270
by (auto_tac (claset(),simpset() addsimps [hypreal_add]));
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1271
by (ALLGOALS(Ultra_tac));
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1272
qed "hypreal_add_self_cancel";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1273
Addsimps [hypreal_add_self_cancel];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1274
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1275
Goal "(y = x + - y + x) = (y = (x::hypreal))";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1276
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1277
by (dres_inst_tac [("x1","y")] 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1278
    (hypreal_add_right_cancel RS iffD2) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1279
by (auto_tac (claset(),simpset() addsimps hypreal_add_ac));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1280
qed "hypreal_add_self_minus_cancel";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1281
Addsimps [hypreal_add_self_minus_cancel];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1282
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1283
Goal "(y = x + (- y + x)) = (y = (x::hypreal))";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1284
by (asm_full_simp_tac (simpset() addsimps 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1285
         [hypreal_add_assoc RS sym])1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1286
qed "hypreal_add_self_minus_cancel2";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1287
Addsimps [hypreal_add_self_minus_cancel2];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1288
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1289
(* of course, can prove this by "transfer" as well *)
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1290
Goal "z + -x = y + (y + (-x + -z)) = (y = (z::hypreal))";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1291
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1292
by (dres_inst_tac [("x1","z")] 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1293
    (hypreal_add_right_cancel RS iffD2) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1294
by (asm_full_simp_tac (simpset() addsimps 
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1295
    [hypreal_minus_add_distrib RS sym] @ hypreal_add_ac
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1296
    delsimps [hypreal_minus_add_distrib]) 1);
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1297
by (asm_full_simp_tac (simpset() addsimps 
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1298
     [hypreal_add_assoc RS sym,hypreal_add_right_cancel]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1299
qed "hypreal_add_self_minus_cancel3";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1300
Addsimps [hypreal_add_self_minus_cancel3];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1301
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
  1302
Goal "(x * x = 0) = (x = (0::hypreal))";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1303
by Auto_tac;
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1304
by (blast_tac (claset() addIs [hypreal_mult_not_0E]) 1);
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1305
qed "hypreal_mult_self_eq_zero_iff";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1306
Addsimps [hypreal_mult_self_eq_zero_iff];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1307
9055
f020e00c6304 replacing 0hr by (0::hypreal)
paulson
parents: 9043
diff changeset
  1308
Goal "(0 = x * x) = (x = (0::hypreal))";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1309
by (auto_tac (claset() addDs [sym],simpset()));
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1310
qed "hypreal_mult_self_eq_zero_iff2";
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1311
Addsimps [hypreal_mult_self_eq_zero_iff2];
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1312
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1313
Goalw [hypreal_diff_def] "(x<y) = (x-y < (0::hypreal))";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1314
by (rtac hypreal_less_minus_iff2 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1315
qed "hypreal_less_eq_diff";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1316
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1317
(*** Subtraction laws ***)
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1318
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1319
Goal "x + (y - z) = (x + y) - (z::hypreal)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1320
by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1321
qed "hypreal_add_diff_eq";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1322
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1323
Goal "(x - y) + z = (x + z) - (y::hypreal)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1324
by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1325
qed "hypreal_diff_add_eq";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1326
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1327
Goal "(x - y) - z = x - (y + (z::hypreal))";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1328
by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1329
qed "hypreal_diff_diff_eq";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1330
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1331
Goal "x - (y - z) = (x + z) - (y::hypreal)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1332
by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1333
qed "hypreal_diff_diff_eq2";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1334
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1335
Goal "(x-y < z) = (x < z + (y::hypreal))";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1336
by (stac hypreal_less_eq_diff 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1337
by (res_inst_tac [("y1", "z")] (hypreal_less_eq_diff RS ssubst) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1338
by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1339
qed "hypreal_diff_less_eq";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1340
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1341
Goal "(x < z-y) = (x + (y::hypreal) < z)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1342
by (stac hypreal_less_eq_diff 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1343
by (res_inst_tac [("y1", "z-y")] (hypreal_less_eq_diff RS ssubst) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1344
by (simp_tac (simpset() addsimps hypreal_diff_def::hypreal_add_ac) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1345
qed "hypreal_less_diff_eq";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1346
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1347
Goalw [hypreal_le_def] "(x-y <= z) = (x <= z + (y::hypreal))";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1348
by (simp_tac (simpset() addsimps [hypreal_less_diff_eq]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1349
qed "hypreal_diff_le_eq";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1350
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1351
Goalw [hypreal_le_def] "(x <= z-y) = (x + (y::hypreal) <= z)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1352
by (simp_tac (simpset() addsimps [hypreal_diff_less_eq]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1353
qed "hypreal_le_diff_eq";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1354
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1355
Goalw [hypreal_diff_def] "(x-y = z) = (x = z + (y::hypreal))";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1356
by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc]));
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1357
qed "hypreal_diff_eq_eq";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1358
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1359
Goalw [hypreal_diff_def] "(x = z-y) = (x + (y::hypreal) = z)";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1360
by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc]));
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1361
qed "hypreal_eq_diff_eq";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1362
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1363
(*This list of rewrites simplifies (in)equalities by bringing subtractions
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1364
  to the top and then moving negative terms to the other side.  
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1365
  Use with hypreal_add_ac*)
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1366
val hypreal_compare_rls = 
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1367
  [symmetric hypreal_diff_def,
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1368
   hypreal_add_diff_eq, hypreal_diff_add_eq, hypreal_diff_diff_eq, hypreal_diff_diff_eq2, 
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1369
   hypreal_diff_less_eq, hypreal_less_diff_eq, hypreal_diff_le_eq, hypreal_le_diff_eq, 
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1370
   hypreal_diff_eq_eq, hypreal_eq_diff_eq];
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1371
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1372
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1373
(** For the cancellation simproc.
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1374
    The idea is to cancel like terms on opposite sides by subtraction **)
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1375
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1376
Goal "(x::hypreal) - y = x' - y' ==> (x<y) = (x'<y')";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1377
by (stac hypreal_less_eq_diff 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1378
by (res_inst_tac [("y1", "y")] (hypreal_less_eq_diff RS ssubst) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1379
by (Asm_simp_tac 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1380
qed "hypreal_less_eqI";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1381
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1382
Goal "(x::hypreal) - y = x' - y' ==> (y<=x) = (y'<=x')";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1383
by (dtac hypreal_less_eqI 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1384
by (asm_simp_tac (simpset() addsimps [hypreal_le_def]) 1);
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1385
qed "hypreal_le_eqI";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1386
10043
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1387
Goal "(x::hypreal) - y = x' - y' ==> (x=y) = (x'=y')";
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1388
by Safe_tac;
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1389
by (ALLGOALS
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1390
    (asm_full_simp_tac
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1391
     (simpset() addsimps [hypreal_eq_diff_eq, hypreal_diff_eq_eq])));
a0364652e115 Updated Files with new theorems
fleuriot
parents: 9969
diff changeset
  1392
qed "hypreal_eq_eqI";
7218
bfa767b4dc51 new theory Real/Hyperreal/HyperDef and file fuf.ML
paulson
parents:
diff changeset
  1393