src/HOL/Complex/CStar.ML
author paulson
Tue, 10 Feb 2004 12:02:11 +0100
changeset 14378 69c4d5997669
parent 14377 f454b3004f8f
permissions -rw-r--r--
generic of_nat and of_int functions, and generalization of iszero and neg
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     1
(*  Title       : CStar.ML
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     3
    Copyright   : 2001 University of Edinburgh
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     4
    Description : defining *-transforms in NSA which extends sets of complex numbers, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     5
                  and complex functions
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     6
*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     7
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     8
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     9
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    10
(*-----------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    11
(*    Properties of the *-transform applied to sets of reals                         *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    12
(* ----------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    13
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    14
Goalw [starsetC_def] "*sc*(UNIV::complex set) = (UNIV::hcomplex set)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    15
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    16
qed "STARC_complex_set";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    17
Addsimps [STARC_complex_set];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    18
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    19
Goalw [starsetC_def] "*sc* {} = {}";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    20
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    21
qed "STARC_empty_set";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    22
Addsimps [STARC_empty_set];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    23
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    24
Goalw [starsetC_def] "*sc* (A Un B) = *sc* A Un *sc* B";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    25
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    26
by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    27
by (dtac bspec 1 THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    28
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    29
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    30
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    31
qed "STARC_Un";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    32
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    33
Goalw [starsetC_n_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    34
      "*scn* (%n. (A n) Un (B n)) = *scn* A Un *scn* B";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    35
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    36
by (dres_inst_tac [("x","Xa")] bspec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    37
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    38
by (auto_tac (claset() addSDs [bspec], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    39
by (TRYALL(Ultra_tac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    40
qed "starsetC_n_Un";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    41
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    42
Goalw [InternalCSets_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    43
     "[| X : InternalCSets; Y : InternalCSets |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    44
\     ==> (X Un Y) : InternalCSets";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    45
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    46
         simpset() addsimps [starsetC_n_Un RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    47
qed "InternalCSets_Un";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    48
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    49
Goalw [starsetC_def] "*sc* (A Int B) = *sc* A Int *sc* B";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    50
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    51
by (blast_tac (claset() addIs [FreeUltrafilterNat_Int,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    52
               FreeUltrafilterNat_subset]) 3);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    53
by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    54
qed "STARC_Int";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    55
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    56
Goalw [starsetC_n_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    57
      "*scn* (%n. (A n) Int (B n)) = *scn* A Int *scn* B";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    58
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    59
by (auto_tac (claset() addSDs [bspec],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    60
by (TRYALL(Ultra_tac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    61
qed "starsetC_n_Int";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    62
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    63
Goalw [InternalCSets_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    64
     "[| X : InternalCSets; Y : InternalCSets |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    65
\     ==> (X Int Y) : InternalCSets";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    66
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    67
         simpset() addsimps [starsetC_n_Int RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    68
qed "InternalCSets_Int";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    69
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    70
Goalw [starsetC_def] "*sc* -A = -( *sc* A)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    71
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    72
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    73
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    74
by (REPEAT(Step_tac 1) THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    75
by (ALLGOALS(Ultra_tac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    76
qed "STARC_Compl";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    77
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    78
Goalw [starsetC_n_def] "*scn* ((%n. - A n)) = -( *scn* A)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    79
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    80
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    81
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    82
by (REPEAT(Step_tac 1) THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    83
by (TRYALL(Ultra_tac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    84
qed "starsetC_n_Compl";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    85
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    86
Goalw [InternalCSets_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    87
     "X :InternalCSets ==> -X : InternalCSets";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    88
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    89
         simpset() addsimps [starsetC_n_Compl RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    90
qed "InternalCSets_Compl";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    91
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    92
Goal "x ~: *sc* F ==> x : *sc* (- F)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    93
by (auto_tac (claset(),simpset() addsimps [STARC_Compl]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    94
qed "STARC_mem_Compl";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    95
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    96
Goal "*sc* (A - B) = *sc* A - *sc* B";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    97
by (auto_tac (claset(),simpset() addsimps 
14370
b0064703967b simplifications in the hyperreals
paulson
parents: 13957
diff changeset
    98
         [Diff_eq,STARC_Int,STARC_Compl]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    99
qed "STARC_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   100
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   101
Goalw [starsetC_n_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   102
      "*scn* (%n. (A n) - (B n)) = *scn* A - *scn* B";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   103
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   104
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   105
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 3);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   106
by (auto_tac (claset() addSDs [bspec], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   107
by (TRYALL(Ultra_tac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   108
qed "starsetC_n_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   109
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   110
Goalw [InternalCSets_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   111
     "[| X : InternalCSets; Y : InternalCSets |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   112
\     ==> (X - Y) : InternalCSets";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   113
by (auto_tac (claset(), simpset() addsimps [starsetC_n_diff RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   114
qed "InternalCSets_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   115
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   116
Goalw [starsetC_def] "A <= B ==> *sc* A <= *sc* B";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   117
by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   118
qed "STARC_subset";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   119
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   120
Goalw [starsetC_def,hcomplex_of_complex_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   121
          "a : A ==> hcomplex_of_complex a : *sc* A";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   122
by (auto_tac (claset() addIs [FreeUltrafilterNat_subset],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   123
         simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   124
qed "STARC_mem";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   125
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   126
Goalw [starsetC_def] "hcomplex_of_complex ` A <= *sc* A";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   127
by (auto_tac (claset(), simpset() addsimps [hcomplex_of_complex_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   128
by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   129
qed "STARC_hcomplex_of_complex_image_subset";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   130
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   131
Goal "SComplex <= *sc* (UNIV:: complex set)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   132
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   133
qed "STARC_SComplex_subset";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   134
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   135
Goalw [starsetC_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   136
     "*sc* X Int SComplex = hcomplex_of_complex ` X";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   137
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   138
         simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   139
           [hcomplex_of_complex_def,SComplex_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   140
by (fold_tac [hcomplex_of_complex_def]);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   141
by (rtac imageI 1 THEN rtac ccontr 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   142
by (dtac bspec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   143
by (rtac lemma_hcomplexrel_refl 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   144
by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   145
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   146
qed "STARC_hcomplex_of_complex_Int";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   147
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   148
Goal "x ~: hcomplex_of_complex ` A ==> ALL y: A. x ~= hcomplex_of_complex y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   149
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   150
qed "lemma_not_hcomplexA";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   151
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   152
Goalw [starsetC_n_def,starsetC_def] "*sc* X = *scn* (%n. X)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   153
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   154
qed "starsetC_starsetC_n_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   155
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   156
Goalw [InternalCSets_def] "( *sc* X) : InternalCSets";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   157
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   158
         simpset() addsimps [starsetC_starsetC_n_eq]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   159
qed "InternalCSets_starsetC_n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   160
Addsimps [InternalCSets_starsetC_n];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   161
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   162
Goal "X : InternalCSets ==> UNIV - X : InternalCSets";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   163
by (auto_tac (claset() addIs [InternalCSets_Compl], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   164
qed "InternalCSets_UNIV_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   165
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   166
(*-----------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   167
(* Nonstandard extension of a set (defined using a constant sequence) as a special   *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   168
(* case of an internal set                                                           *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   169
(*-----------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   170
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   171
Goalw [starsetC_n_def,starsetC_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   172
     "ALL n. (As n = A) ==> *scn* As = *sc* A";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   173
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   174
qed "starsetC_n_starsetC";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   175
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   176
(*-----------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   177
(* Theorems about nonstandard extensions of functions                                *)   
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   178
(*-----------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   179
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   180
Goalw [starfunC_n_def,starfunC_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   181
     "ALL n. (F n = f) ==> *fcn* F = *fc* f";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   182
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   183
qed "starfunC_n_starfunC";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   184
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   185
Goalw [starfunRC_n_def,starfunRC_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   186
     "ALL n. (F n = f) ==> *fRcn* F = *fRc* f";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   187
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   188
qed "starfunRC_n_starfunRC";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   189
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   190
Goalw [starfunCR_n_def,starfunCR_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   191
     "ALL n. (F n = f) ==> *fcRn* F = *fcR* f";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   192
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   193
qed "starfunCR_n_starfunCR";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   194
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   195
Goalw [congruent_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   196
      "congruent hcomplexrel (%X. hcomplexrel``{%n. f (X n)})";
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   197
by (auto_tac (clasimpset() addIffs [hcomplexrel_iff]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   198
by (ALLGOALS(Ultra_tac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   199
qed "starfunC_congruent";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   200
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   201
(* f::complex => complex *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   202
Goalw [starfunC_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   203
      "( *fc* f) (Abs_hcomplex(hcomplexrel``{%n. X n})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   204
\      Abs_hcomplex(hcomplexrel `` {%n. f (X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   205
by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1);
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   206
by (auto_tac (clasimpset() addIffs [hcomplexrel_iff]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   207
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   208
qed "starfunC";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   209
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   210
Goalw [starfunRC_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   211
      "( *fRc* f) (Abs_hypreal(hyprel``{%n. X n})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   212
\      Abs_hcomplex(hcomplexrel `` {%n. f (X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   213
by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   214
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   215
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   216
qed "starfunRC";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   217
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   218
Goalw [starfunCR_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   219
      "( *fcR* f) (Abs_hcomplex(hcomplexrel``{%n. X n})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   220
\      Abs_hypreal(hyprel `` {%n. f (X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   221
by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   222
by (auto_tac (clasimpset() addIffs [hcomplexrel_iff]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   223
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   224
qed "starfunCR";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   225
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   226
(**  multiplication: ( *f ) x ( *g ) = *(f x g) **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   227
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   228
Goal "( *fc* f) z * ( *fc* g) z = ( *fc* (%x. f x * g x)) z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   229
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   230
by (auto_tac (claset(), simpset() addsimps [starfunC,hcomplex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   231
qed "starfunC_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   232
Addsimps [starfunC_mult RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   233
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   234
Goal "( *fRc* f) z * ( *fRc* g) z = ( *fRc* (%x. f x * g x)) z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   235
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   236
by (auto_tac (claset(), simpset() addsimps [starfunRC,hcomplex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   237
qed "starfunRC_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   238
Addsimps [starfunRC_mult RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   239
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   240
Goal "( *fcR* f) z * ( *fcR* g) z = ( *fcR* (%x. f x * g x)) z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   241
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   242
by (auto_tac (claset(), simpset() addsimps [starfunCR,hypreal_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   243
qed "starfunCR_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   244
Addsimps [starfunCR_mult RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   245
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   246
(**  addition: ( *f ) + ( *g ) = *(f + g)  **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   247
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   248
Goal "( *fc* f) z + ( *fc* g) z = ( *fc* (%x. f x + g x)) z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   249
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   250
by (auto_tac (claset(), simpset() addsimps [starfunC,hcomplex_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   251
qed "starfunC_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   252
Addsimps [starfunC_add RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   253
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   254
Goal "( *fRc* f) z + ( *fRc* g) z = ( *fRc* (%x. f x + g x)) z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   255
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   256
by (auto_tac (claset(), simpset() addsimps [starfunRC,hcomplex_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   257
qed "starfunRC_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   258
Addsimps [starfunRC_add RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   259
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   260
Goal "( *fcR* f) z + ( *fcR* g) z = ( *fcR* (%x. f x + g x)) z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   261
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   262
by (auto_tac (claset(), simpset() addsimps [starfunCR,hypreal_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   263
qed "starfunCR_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   264
Addsimps [starfunCR_add RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   265
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   266
(**  uminus **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   267
Goal "( *fc* (%x. - f x)) x = - ( *fc* f) x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   268
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   269
by (auto_tac (claset(),simpset() addsimps [starfunC, hcomplex_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   270
qed "starfunC_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   271
Addsimps [starfunC_minus];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   272
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   273
Goal "( *fRc* (%x. - f x)) x = - ( *fRc* f) x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   274
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   275
by (auto_tac (claset(),simpset() addsimps [starfunRC, hcomplex_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   276
qed "starfunRC_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   277
Addsimps [starfunRC_minus];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   278
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   279
Goal "( *fcR* (%x. - f x)) x = - ( *fcR* f) x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   280
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   281
by (auto_tac (claset(),simpset() addsimps [starfunCR, hypreal_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   282
qed "starfunCR_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   283
Addsimps [starfunCR_minus];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   284
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   285
(**  addition: ( *f ) - ( *g ) = *(f - g)  **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   286
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   287
Goalw [hcomplex_diff_def,complex_diff_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   288
  "( *fc* f) xa  - ( *fc* g) xa = ( *fc* (%x. f x - g x)) xa";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   289
by (auto_tac (claset(),simpset() addsimps [starfunC_minus,starfunC_add RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   290
qed "starfunC_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   291
Addsimps [starfunC_diff RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   292
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   293
Goalw [hcomplex_diff_def,complex_diff_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   294
  "( *fRc* f) xa  - ( *fRc* g) xa = ( *fRc* (%x. f x - g x)) xa";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   295
by (auto_tac (claset(),simpset() addsimps [starfunRC_minus,starfunRC_add RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   296
qed "starfunRC_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   297
Addsimps [starfunRC_diff RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   298
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   299
Goalw [hypreal_diff_def,real_diff_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   300
  "( *fcR* f) xa  - ( *fcR* g) xa = ( *fcR* (%x. f x - g x)) xa";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   301
by (auto_tac (claset(),simpset() addsimps [starfunCR_minus,starfunCR_add RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   302
qed "starfunCR_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   303
Addsimps [starfunCR_diff RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   304
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   305
(**  composition: ( *f ) o ( *g ) = *(f o g) **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   306
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   307
Goal "(%x. ( *fc* f) (( *fc* g) x)) = *fc* (%x. f (g x))"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   308
by (rtac ext 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   309
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   310
by (auto_tac (claset(),simpset() addsimps [starfunC]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   311
qed "starfunC_o2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   312
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   313
Goalw [o_def] "( *fc* f) o ( *fc* g) = ( *fc* (f o g))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   314
by (simp_tac (simpset() addsimps [starfunC_o2]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   315
qed "starfunC_o";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   316
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   317
Goal "(%x. ( *fc* f) (( *fRc* g) x)) = *fRc* (%x. f (g x))"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   318
by (rtac ext 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   319
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   320
by (auto_tac (claset(),simpset() addsimps [starfunRC,starfunC]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   321
qed "starfunC_starfunRC_o2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   322
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   323
Goal "(%x. ( *f* f) (( *fcR* g) x)) = *fcR* (%x. f (g x))"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   324
by (rtac ext 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   325
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   326
by (auto_tac (claset(),simpset() addsimps [starfunCR,starfun]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   327
qed "starfun_starfunCR_o2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   328
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   329
Goalw [o_def] "( *fc* f) o ( *fRc* g) = ( *fRc* (f o g))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   330
by (simp_tac (simpset() addsimps [starfunC_starfunRC_o2]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   331
qed "starfunC_starfunRC_o";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   332
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   333
Goalw [o_def] "( *f* f) o ( *fcR* g) = ( *fcR* (f o g))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   334
by (simp_tac (simpset() addsimps [starfun_starfunCR_o2]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   335
qed "starfun_starfunCR_o";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   336
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   337
Goal "( *fc* (%x. k)) z = hcomplex_of_complex k";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   338
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   339
by (auto_tac (claset(), simpset() addsimps [starfunC, hcomplex_of_complex_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   340
qed "starfunC_const_fun";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   341
Addsimps [starfunC_const_fun];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   342
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   343
Goal "( *fRc* (%x. k)) z = hcomplex_of_complex k";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   344
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   345
by (auto_tac (claset(), simpset() addsimps [starfunRC, hcomplex_of_complex_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   346
qed "starfunRC_const_fun";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   347
Addsimps [starfunRC_const_fun];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   348
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   349
Goal "( *fcR* (%x. k)) z = hypreal_of_real k";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   350
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   351
by (auto_tac (claset(), simpset() addsimps [starfunCR, hypreal_of_real_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   352
qed "starfunCR_const_fun";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   353
Addsimps [starfunCR_const_fun];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   354
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   355
Goal "inverse (( *fc* f) x) = ( *fc* (%x. inverse (f x))) x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   356
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   357
by (auto_tac (claset(), simpset() addsimps [starfunC, hcomplex_inverse]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   358
qed "starfunC_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   359
Addsimps [starfunC_inverse RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   360
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   361
Goal "inverse (( *fRc* f) x) = ( *fRc* (%x. inverse (f x))) x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   362
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   363
by (auto_tac (claset(), simpset() addsimps [starfunRC, hcomplex_inverse]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   364
qed "starfunRC_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   365
Addsimps [starfunRC_inverse RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   366
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   367
Goal "inverse (( *fcR* f) x) = ( *fcR* (%x. inverse (f x))) x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   368
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   369
by (auto_tac (claset(), simpset() addsimps [starfunCR, hypreal_inverse]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   370
qed "starfunCR_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   371
Addsimps [starfunCR_inverse RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   372
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   373
Goal "( *fc* f) (hcomplex_of_complex a) = hcomplex_of_complex (f a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   374
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   375
      simpset() addsimps [starfunC,hcomplex_of_complex_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   376
qed "starfunC_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   377
Addsimps [starfunC_eq];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   378
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   379
Goal "( *fRc* f) (hypreal_of_real a) = hcomplex_of_complex (f a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   380
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   381
      simpset() addsimps [starfunRC,hcomplex_of_complex_def,hypreal_of_real_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   382
qed "starfunRC_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   383
Addsimps [starfunRC_eq];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   384
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   385
Goal "( *fcR* f) (hcomplex_of_complex a) = hypreal_of_real (f a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   386
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   387
      simpset() addsimps [starfunCR,hcomplex_of_complex_def,hypreal_of_real_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   388
qed "starfunCR_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   389
Addsimps [starfunCR_eq];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   390
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   391
Goal "( *fc* f) (hcomplex_of_complex a) @c= hcomplex_of_complex (f a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   392
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   393
qed "starfunC_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   394
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   395
Goal "( *fRc* f) (hypreal_of_real a) @c= hcomplex_of_complex (f a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   396
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   397
qed "starfunRC_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   398
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   399
Goal "( *fcR* f) (hcomplex_of_complex a) @= hypreal_of_real (f a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   400
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   401
qed "starfunCR_approx";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   402
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   403
(*
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   404
Goal "( *fcNat* (%n. z ^ n)) N = (hcomplex_of_complex z) hcpow N";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   405
*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   406
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14377
diff changeset
   407
Goal "( *fc* (%z. z ^ n)) Z = Z hcpow hypnat_of_nat n";
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   408
by (res_inst_tac [("z","Z")] eq_Abs_hcomplex 1);
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14377
diff changeset
   409
by (auto_tac (claset(), simpset() addsimps [hcpow,starfunC,hypnat_of_nat_eq]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   410
qed "starfunC_hcpow";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   411
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   412
Goal "( *fc* (%h. f (x + h))) xa  = ( *fc* f) (hcomplex_of_complex  x + xa)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   413
by (res_inst_tac [("z","xa")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   414
by (auto_tac (claset(),simpset() addsimps [starfunC,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   415
    hcomplex_of_complex_def,hcomplex_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   416
qed "starfunC_lambda_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   417
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   418
Goal "( *fcR* (%h. f (x + h))) xa  = ( *fcR* f) (hcomplex_of_complex  x + xa)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   419
by (res_inst_tac [("z","xa")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   420
by (auto_tac (claset(),simpset() addsimps [starfunCR,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   421
    hcomplex_of_complex_def,hcomplex_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   422
qed "starfunCR_lambda_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   423
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   424
Goal "( *fRc* (%h. f (x + h))) xa  = ( *fRc* f) (hypreal_of_real x + xa)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   425
by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   426
by (auto_tac (claset(),simpset() addsimps [starfunRC,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   427
    hypreal_of_real_def,hypreal_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   428
qed "starfunRC_lambda_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   429
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   430
Goal "( *fc* (%h. f(g(x + h)))) xa = ( *fc* (f o g)) (hcomplex_of_complex x + xa)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   431
by (res_inst_tac [("z","xa")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   432
by (auto_tac (claset(),simpset() addsimps [starfunC,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   433
    hcomplex_of_complex_def,hcomplex_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   434
qed "starfunC_lambda_cancel2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   435
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   436
Goal "( *fcR* (%h. f(g(x + h)))) xa = ( *fcR* (f o g)) (hcomplex_of_complex x + xa)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   437
by (res_inst_tac [("z","xa")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   438
by (auto_tac (claset(),simpset() addsimps [starfunCR,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   439
    hcomplex_of_complex_def,hcomplex_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   440
qed "starfunCR_lambda_cancel2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   441
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   442
Goal "( *fRc* (%h. f(g(x + h)))) xa = ( *fRc* (f o g)) (hypreal_of_real x + xa)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   443
by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   444
by (auto_tac (claset(),simpset() addsimps [starfunRC,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   445
    hypreal_of_real_def,hypreal_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   446
qed "starfunRC_lambda_cancel2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   447
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   448
Goal "[| ( *fc* f) xa @c= l; ( *fc* g) xa @c= m; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   449
\                 l: CFinite; m: CFinite  \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   450
\              |] ==>  ( *fc* (%x. f x * g x)) xa @c= l * m";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   451
by (dtac capprox_mult_CFinite 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   452
by (REPEAT(assume_tac 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   453
by (auto_tac (claset() addIs [capprox_sym RSN (2,capprox_CFinite)],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   454
              simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   455
qed "starfunC_mult_CFinite_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   456
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   457
Goal "[| ( *fcR* f) xa @= l; ( *fcR* g) xa @= m; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   458
\                 l: HFinite; m: HFinite  \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   459
\              |] ==>  ( *fcR* (%x. f x * g x)) xa @= l * m";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   460
by (dtac approx_mult_HFinite 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   461
by (REPEAT(assume_tac 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   462
by (auto_tac (claset() addIs [approx_sym RSN (2,approx_HFinite)],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   463
              simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   464
qed "starfunCR_mult_HFinite_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   465
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   466
Goal "[| ( *fRc* f) xa @c= l; ( *fRc* g) xa @c= m; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   467
\                 l: CFinite; m: CFinite  \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   468
\              |] ==>  ( *fRc* (%x. f x * g x)) xa @c= l * m";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   469
by (dtac capprox_mult_CFinite 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   470
by (REPEAT(assume_tac 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   471
by (auto_tac (claset() addIs [capprox_sym RSN (2,capprox_CFinite)],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   472
              simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   473
qed "starfunRC_mult_CFinite_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   474
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   475
Goal "[| ( *fc* f) xa @c= l; ( *fc* g) xa @c= m \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   476
\              |] ==>  ( *fc* (%x. f x + g x)) xa @c= l + m";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   477
by (auto_tac (claset() addIs [capprox_add], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   478
qed "starfunC_add_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   479
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   480
Goal "[| ( *fRc* f) xa @c= l; ( *fRc* g) xa @c= m \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   481
\              |] ==>  ( *fRc* (%x. f x + g x)) xa @c= l + m";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   482
by (auto_tac (claset() addIs [capprox_add], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   483
qed "starfunRC_add_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   484
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   485
Goal "[| ( *fcR* f) xa @= l; ( *fcR* g) xa @= m \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   486
\              |] ==>  ( *fcR* (%x. f x + g x)) xa @= l + m";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   487
by (auto_tac (claset() addIs [approx_add], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   488
qed "starfunCR_add_approx";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   489
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   490
Goal "*fcR* cmod = hcmod";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   491
by (rtac ext 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   492
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   493
by (auto_tac (claset(),simpset() addsimps [starfunCR,hcmod]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   494
qed "starfunCR_cmod";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   495
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   496
Goal "( *fc* inverse) x = inverse(x)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   497
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   498
by (auto_tac (claset(),simpset() addsimps [starfunC,hcomplex_inverse]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   499
qed "starfunC_inverse_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   500
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   501
Goalw [hcomplex_divide_def,complex_divide_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   502
  "( *fc* f) xa  / ( *fc* g) xa = ( *fc* (%x. f x / g x)) xa";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   503
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   504
qed "starfunC_divide";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   505
Addsimps [starfunC_divide RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   506
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   507
Goalw [hypreal_divide_def,real_divide_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   508
  "( *fcR* f) xa  / ( *fcR* g) xa = ( *fcR* (%x. f x / g x)) xa";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   509
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   510
qed "starfunCR_divide";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   511
Addsimps [starfunCR_divide RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   512
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   513
Goalw [hcomplex_divide_def,complex_divide_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   514
  "( *fRc* f) xa  / ( *fRc* g) xa = ( *fRc* (%x. f x / g x)) xa";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   515
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   516
qed "starfunRC_divide";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   517
Addsimps [starfunRC_divide RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   518
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   519
(*-----------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   520
(* Internal functions - some redundancy with *fc* now                                *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   521
(*-----------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   522
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   523
Goalw [congruent_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   524
      "congruent hcomplexrel (%X. hcomplexrel``{%n. f n (X n)})";
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   525
by (auto_tac (clasimpset() addIffs [hcomplexrel_iff]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   526
by (ALLGOALS(Fuf_tac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   527
qed "starfunC_n_congruent";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   528
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   529
Goalw [starfunC_n_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   530
     "( *fcn* f) (Abs_hcomplex(hcomplexrel``{%n. X n})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   531
\     Abs_hcomplex(hcomplexrel `` {%n. f n (X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   532
by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1);
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   533
by (auto_tac (clasimpset() addIffs [hcomplexrel_iff]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   534
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   535
qed "starfunC_n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   536
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   537
(**  multiplication: ( *fn ) x ( *gn ) = *(fn x gn) **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   538
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   539
Goal "( *fcn* f) z * ( *fcn* g) z = ( *fcn* (% i x. f i x * g i x)) z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   540
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   541
by (auto_tac (claset(), simpset() addsimps [starfunC_n,hcomplex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   542
qed "starfunC_n_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   543
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   544
(**  addition: ( *fn ) + ( *gn ) = *(fn + gn) **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   545
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   546
Goal "( *fcn* f) z + ( *fcn* g) z = ( *fcn* (%i x. f i x + g i x)) z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   547
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   548
by (auto_tac (claset(), simpset() addsimps [starfunC_n,hcomplex_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   549
qed "starfunC_n_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   550
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   551
(** uminus **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   552
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   553
Goal "- ( *fcn* g) z = ( *fcn* (%i x. - g i x)) z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   554
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   555
by (auto_tac (claset(), simpset() addsimps [starfunC_n,hcomplex_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   556
qed "starfunC_n_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   557
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   558
(** subtraction: ( *fn ) - ( *gn ) = *(fn - gn) **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   559
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   560
Goalw [hcomplex_diff_def,complex_diff_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   561
   "( *fcn* f) z - ( *fcn* g) z = ( *fcn* (%i x. f i x - g i x)) z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   562
by (auto_tac (claset(), 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   563
          simpset() addsimps [starfunC_n_add,starfunC_n_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   564
qed "starfunNat_n_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   565
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   566
(** composition: ( *fn ) o ( *gn ) = *(fn o gn) **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   567
 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   568
Goal "( *fcn* (%i x. k)) z = hcomplex_of_complex  k";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   569
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   570
by (auto_tac (claset(), 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   571
       simpset() addsimps [starfunC_n, hcomplex_of_complex_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   572
qed "starfunC_n_const_fun";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   573
Addsimps [starfunC_n_const_fun];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   574
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   575
Goal "( *fcn* f) (hcomplex_of_complex n) = Abs_hcomplex(hcomplexrel `` {%i. f i n})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   576
by (auto_tac (claset(), simpset() addsimps [starfunC_n,hcomplex_of_complex_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   577
qed "starfunC_n_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   578
Addsimps [starfunC_n_eq];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   579
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   580
Goal "(( *fc* f) = ( *fc* g)) = (f = g)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   581
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   582
by (rtac ext 1 THEN rtac ccontr 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   583
by (dres_inst_tac [("x","hcomplex_of_complex(x)")] fun_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   584
by (auto_tac (claset(), simpset() addsimps [starfunC,hcomplex_of_complex_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   585
qed "starfunC_eq_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   586
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   587
Goal "(( *fRc* f) = ( *fRc* g)) = (f = g)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   588
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   589
by (rtac ext 1 THEN rtac ccontr 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   590
by (dres_inst_tac [("x","hypreal_of_real(x)")] fun_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   591
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   592
qed "starfunRC_eq_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   593
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   594
Goal "(( *fcR* f) = ( *fcR* g)) = (f = g)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   595
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   596
by (rtac ext 1 THEN rtac ccontr 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   597
by (dres_inst_tac [("x","hcomplex_of_complex(x)")] fun_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   598
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   599
qed "starfunCR_eq_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   600
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   601
(*** more theorems ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   602
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   603
Goal "(( *fc* f) x = z) = ((( *fcR* (%x. Re(f x))) x = hRe (z)) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   604
\                         (( *fcR* (%x. Im(f x))) x = hIm (z)))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   605
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   606
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   607
by (auto_tac (claset(),simpset() addsimps [starfunCR,starfunC,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   608
    hIm,hRe,complex_Re_Im_cancel_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   609
by (ALLGOALS(Ultra_tac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   610
qed "starfunC_eq_Re_Im_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   611
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   612
Goal "(( *fc* f) x @c= z) = ((( *fcR* (%x. Re(f x))) x @= hRe (z)) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   613
\                           (( *fcR* (%x. Im(f x))) x @= hIm (z)))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   614
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   615
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   616
by (auto_tac (claset(),simpset() addsimps [starfunCR,starfunC,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   617
    hIm,hRe,capprox_approx_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   618
qed "starfunC_approx_Re_Im_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   619
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   620
Goal "x @c= hcomplex_of_complex a ==> ( *fc* (%x. x)) x @c= hcomplex_of_complex  a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   621
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   622
by (auto_tac (claset(),simpset() addsimps [starfunC]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   623
qed "starfunC_Idfun_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   624
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   625
Goal "( *fc* (%x. x)) x = x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   626
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   627
by (auto_tac (claset(),simpset() addsimps [starfunC]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   628
qed "starfunC_Id";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   629
Addsimps [starfunC_Id];