src/HOL/Complex/NSComplexBin.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:      NSComplexBin.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
    Descrition: Binary arithmetic for the nonstandard complex numbers
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     5
*)
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
(** hcomplex_of_complex (coercion from complex to nonstandard complex) **)
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
Goal "hcomplex_of_complex (number_of w) = number_of w";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    10
by (simp_tac (simpset() addsimps [hcomplex_number_of_def]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    11
qed "hcomplex_number_of";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    12
Addsimps [hcomplex_number_of];
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 [hypreal_of_real_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    15
     "hcomplex_of_hypreal (hypreal_of_real x) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    16
\     hcomplex_of_complex(complex_of_real x)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    17
by (simp_tac (simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    18
    hcomplex_of_complex_def,complex_of_real_def]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    19
qed "hcomplex_of_hypreal_eq_hcomplex_of_complex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    20
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    21
Goalw [complex_number_of_def,hypreal_number_of_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    22
  "hcomplex_of_complex (number_of w) = hcomplex_of_hypreal(number_of w)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    23
by (rtac (hcomplex_of_hypreal_eq_hcomplex_of_complex RS sym) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    24
qed "hcomplex_hypreal_number_of";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    25
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    26
Goalw [hcomplex_number_of_def] "Numeral0 = (0::hcomplex)";
14320
fb7a114826be tidying up hcomplex arithmetic
paulson
parents: 14318
diff changeset
    27
by(Simp_tac 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    28
qed "hcomplex_numeral_0_eq_0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    29
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    30
Goalw [hcomplex_number_of_def] "Numeral1 = (1::hcomplex)";
14320
fb7a114826be tidying up hcomplex arithmetic
paulson
parents: 14318
diff changeset
    31
by(Simp_tac 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    32
qed "hcomplex_numeral_1_eq_1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    33
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    34
(*
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    35
Goal "z + hcnj z = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    36
\     hcomplex_of_hypreal (2 * hRe(z))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    37
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    38
by (auto_tac (claset(),HOL_ss addsimps [hRe,hcnj,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    39
    hypreal_mult,hcomplex_of_hypreal,complex_add_cnj]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    40
qed "hcomplex_add_hcnj";
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
Goal "z - hcnj z = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    43
\     hcomplex_of_hypreal (hypreal_of_real #2 * hIm(z)) * iii";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    44
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    45
by (auto_tac (claset(),simpset() addsimps [hIm,hcnj,hcomplex_diff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    46
    hypreal_of_real_def,hypreal_mult,hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    47
    complex_diff_cnj,iii_def,hcomplex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    48
qed "hcomplex_diff_hcnj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    49
*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    50
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    51
(** Addition **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    52
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    53
Goal "(number_of v :: hcomplex) + number_of v' = number_of (bin_add v v')";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    54
by (simp_tac
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    55
    (HOL_ss addsimps [hcomplex_number_of_def, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    56
                      hcomplex_of_complex_add RS sym, add_complex_number_of]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    57
qed "add_hcomplex_number_of";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    58
Addsimps [add_hcomplex_number_of];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    59
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    60
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    61
(** Subtraction **)
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 [hcomplex_number_of_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    64
     "- (number_of w :: hcomplex) = number_of (bin_minus w)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    65
by (simp_tac
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    66
    (HOL_ss addsimps [minus_complex_number_of, hcomplex_of_complex_minus RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    67
qed "minus_hcomplex_number_of";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    68
Addsimps [minus_hcomplex_number_of];
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 [hcomplex_number_of_def, hcomplex_diff_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    71
     "(number_of v :: hcomplex) - number_of w = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    72
\     number_of (bin_add v (bin_minus w))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    73
by (Simp_tac 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    74
qed "diff_hcomplex_number_of";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    75
Addsimps [diff_hcomplex_number_of];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    76
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
(** Multiplication **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    79
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    80
Goal "(number_of v :: hcomplex) * number_of v' = number_of (bin_mult v v')";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    81
by (simp_tac
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    82
    (HOL_ss addsimps [hcomplex_number_of_def, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    83
	              hcomplex_of_complex_mult RS sym, mult_complex_number_of]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    84
qed "mult_hcomplex_number_of";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    85
Addsimps [mult_hcomplex_number_of];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    86
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    87
Goal "(2::hcomplex) = 1 + 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    88
by (simp_tac (simpset() addsimps [hcomplex_numeral_1_eq_1 RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    89
val lemma = result();
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    90
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    91
(*For specialist use: NOT as default simprules*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    92
Goal "2 * z = (z+z::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    93
by (simp_tac (simpset() addsimps [lemma, hcomplex_add_mult_distrib]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    94
qed "hcomplex_mult_2";
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 "z * 2 = (z+z::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    97
by (stac hcomplex_mult_commute 1 THEN rtac hcomplex_mult_2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    98
qed "hcomplex_mult_2_right";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    99
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   100
(** Equals (=) **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   101
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   102
Goal "((number_of v :: hcomplex) = number_of v') = \
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14377
diff changeset
   103
\     iszero (number_of (bin_add v (bin_minus v')) :: int)";
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   104
by (simp_tac
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   105
    (HOL_ss addsimps [hcomplex_number_of_def, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   106
	              hcomplex_of_complex_eq_iff, eq_complex_number_of]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   107
qed "eq_hcomplex_number_of";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   108
Addsimps [eq_hcomplex_number_of];
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
(*** New versions of existing theorems involving 0, 1hc ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   111
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   112
Goal "- 1 = (-1::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   113
by (simp_tac (simpset() addsimps [hcomplex_numeral_1_eq_1 RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   114
qed "hcomplex_minus_1_eq_m1";
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
Goal "-1 * z = -(z::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   117
by (simp_tac (simpset() addsimps [hcomplex_minus_1_eq_m1 RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   118
qed "hcomplex_mult_minus1";
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
Goal "z * -1 = -(z::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   121
by (stac hcomplex_mult_commute 1 THEN rtac hcomplex_mult_minus1 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   122
qed "hcomplex_mult_minus1_right";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   123
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   124
Addsimps [hcomplex_mult_minus1,hcomplex_mult_minus1_right];
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
(*Maps 0 to Numeral0 and 1 to Numeral1 and -Numeral1 to -1*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   127
val hcomplex_numeral_ss = 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   128
    complex_numeral_ss addsimps [hcomplex_numeral_0_eq_0 RS sym, hcomplex_numeral_1_eq_1 RS sym, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   129
		                 hcomplex_minus_1_eq_m1];
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
fun rename_numerals th = 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   132
    asm_full_simplify hcomplex_numeral_ss (Thm.transfer (the_context ()) th);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   133
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
(*Now insert some identities previously stated for 0 and 1hc*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   136
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   137
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   138
Addsimps [hcomplex_numeral_0_eq_0,hcomplex_numeral_1_eq_1];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   139
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   140
Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   141
by (auto_tac (claset(),simpset() addsimps [hcomplex_add_assoc RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   142
qed "hcomplex_add_number_of_left";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   143
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   144
Goal "number_of v *(number_of w * z) = (number_of(bin_mult v w) * z::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   145
by (simp_tac (simpset() addsimps [hcomplex_mult_assoc RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   146
qed "hcomplex_mult_number_of_left";
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
Goalw [hcomplex_diff_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   149
    "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   150
by (rtac hcomplex_add_number_of_left 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   151
qed "hcomplex_add_number_of_diff1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   152
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   153
Goal "number_of v + (c - number_of w) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   154
\     number_of (bin_add v (bin_minus w)) + (c::hcomplex)";
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14320
diff changeset
   155
by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_def]@ add_ac));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   156
qed "hcomplex_add_number_of_diff2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   157
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   158
Addsimps [hcomplex_add_number_of_left, hcomplex_mult_number_of_left,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   159
	  hcomplex_add_number_of_diff1, hcomplex_add_number_of_diff2]; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   160
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
(**** Simprocs for numeric literals ****)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   163
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   164
structure HComplex_Numeral_Simprocs =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   165
struct
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
(*Utilities*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   168
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   169
val hcomplexT = Type("NSComplex.hcomplex",[]);
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
fun mk_numeral n = HOLogic.number_of_const hcomplexT $ HOLogic.mk_bin n;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   172
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   173
val dest_numeral = Complex_Numeral_Simprocs.dest_numeral;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   174
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   175
val find_first_numeral = Complex_Numeral_Simprocs.find_first_numeral;
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
val zero = mk_numeral 0;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   178
val mk_plus = HOLogic.mk_binop "op +";
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
val uminus_const = Const ("uminus", hcomplexT --> hcomplexT);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   181
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   182
(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   183
fun mk_sum []        = zero
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   184
  | mk_sum [t,u]     = mk_plus (t, u)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   185
  | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   186
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   187
(*this version ALWAYS includes a trailing zero*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   188
fun long_mk_sum []        = zero
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   189
  | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   190
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   191
val dest_plus = HOLogic.dest_bin "op +" hcomplexT;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   192
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   193
(*decompose additions AND subtractions as a sum*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   194
fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   195
        dest_summing (pos, t, dest_summing (pos, u, ts))
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   196
  | dest_summing (pos, Const ("op -", _) $ t $ u, ts) =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   197
        dest_summing (pos, t, dest_summing (not pos, u, ts))
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   198
  | dest_summing (pos, t, ts) =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   199
	if pos then t::ts else uminus_const$t :: ts;
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
fun dest_sum t = dest_summing (true, t, []);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   202
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   203
val mk_diff = HOLogic.mk_binop "op -";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   204
val dest_diff = HOLogic.dest_bin "op -" hcomplexT;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   205
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   206
val one = mk_numeral 1;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   207
val mk_times = HOLogic.mk_binop "op *";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   208
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   209
fun mk_prod [] = one
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   210
  | mk_prod [t] = t
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   211
  | mk_prod (t :: ts) = if t = one then mk_prod ts
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   212
                        else mk_times (t, mk_prod ts);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   213
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   214
val dest_times = HOLogic.dest_bin "op *" hcomplexT;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   215
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   216
fun dest_prod t =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   217
      let val (t,u) = dest_times t 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   218
      in  dest_prod t @ dest_prod u  end
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   219
      handle TERM _ => [t];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   220
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   221
(*DON'T do the obvious simplifications; that would create special cases*) 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   222
fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   223
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   224
(*Express t as a product of (possibly) a numeral with other sorted terms*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   225
fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   226
  | dest_coeff sign t =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   227
    let val ts = sort Term.term_ord (dest_prod t)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   228
	val (n, ts') = find_first_numeral [] ts
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   229
                          handle TERM _ => (1, ts)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   230
    in (sign*n, mk_prod ts') end;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   231
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   232
(*Find first coefficient-term THAT MATCHES u*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   233
fun find_first_coeff past u [] = raise TERM("find_first_coeff", []) 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   234
  | find_first_coeff past u (t::terms) =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   235
	let val (n,u') = dest_coeff 1 t
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   236
	in  if u aconv u' then (n, rev past @ terms)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   237
			  else find_first_coeff (t::past) u terms
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   238
	end
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   239
	handle TERM _ => find_first_coeff (t::past) u terms;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   240
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   241
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   242
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   243
(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   244
val add_0s = map rename_numerals [hcomplex_add_zero_left, hcomplex_add_zero_right];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   245
val mult_plus_1s = map rename_numerals [hcomplex_mult_one_left, hcomplex_mult_one_right];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   246
val mult_minus_1s = map rename_numerals [hcomplex_mult_minus1, hcomplex_mult_minus1_right];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   247
val mult_1s = mult_plus_1s @ mult_minus_1s;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   248
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   249
(*To perform binary arithmetic*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   250
val bin_simps =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   251
    [hcomplex_numeral_0_eq_0 RS sym, hcomplex_numeral_1_eq_1 RS sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   252
     add_hcomplex_number_of, hcomplex_add_number_of_left, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   253
     minus_hcomplex_number_of, diff_hcomplex_number_of, mult_hcomplex_number_of, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   254
     hcomplex_mult_number_of_left] @ bin_arith_simps @ bin_rel_simps;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   255
14123
b300efca4ae0 fixed simprocs
paulson
parents: 13957
diff changeset
   256
(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
b300efca4ae0 fixed simprocs
paulson
parents: 13957
diff changeset
   257
  during re-arrangement*)
b300efca4ae0 fixed simprocs
paulson
parents: 13957
diff changeset
   258
val non_add_bin_simps = 
b300efca4ae0 fixed simprocs
paulson
parents: 13957
diff changeset
   259
    bin_simps \\ [hcomplex_add_number_of_left, add_hcomplex_number_of];
b300efca4ae0 fixed simprocs
paulson
parents: 13957
diff changeset
   260
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   261
(*To evaluate binary negations of coefficients*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   262
val hcomplex_minus_simps = NCons_simps @
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   263
                   [hcomplex_minus_1_eq_m1,minus_hcomplex_number_of, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   264
		    bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   265
		    bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   266
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   267
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   268
(*To let us treat subtraction as addition*)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14320
diff changeset
   269
val diff_simps = [hcomplex_diff_def, minus_add_distrib, minus_minus];
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   270
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   271
(*push the unary minus down: - x * y = x * - y *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   272
val hcomplex_minus_mult_eq_1_to_2 = 
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14320
diff changeset
   273
    [minus_mult_left RS sym, minus_mult_right] MRS trans 
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   274
    |> standard;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   275
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   276
(*to extract again any uncancelled minuses*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   277
val hcomplex_minus_from_mult_simps = 
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14320
diff changeset
   278
    [minus_minus, minus_mult_left RS sym, minus_mult_right RS sym];
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   279
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   280
(*combine unary minus with numeric literals, however nested within a product*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   281
val hcomplex_mult_minus_simps =
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14320
diff changeset
   282
    [hcomplex_mult_assoc, minus_mult_left, hcomplex_minus_mult_eq_1_to_2];
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   283
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   284
(*Final simplification: cancel + and *  *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   285
val simplify_meta_eq = 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   286
    Int_Numeral_Simprocs.simplify_meta_eq
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14320
diff changeset
   287
         [add_zero_left, add_zero_right,
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14335
diff changeset
   288
 	  mult_zero_left, mult_zero_right, mult_1, mult_1_right];
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   289
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   290
val prep_simproc = Complex_Numeral_Simprocs.prep_simproc;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   291
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
structure CancelNumeralsCommon =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   294
  struct
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   295
  val mk_sum    	= mk_sum
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   296
  val dest_sum		= dest_sum
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   297
  val mk_coeff		= mk_coeff
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   298
  val dest_coeff	= dest_coeff 1
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   299
  val find_first_coeff	= find_first_coeff []
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   300
  val trans_tac         = Real_Numeral_Simprocs.trans_tac
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   301
  val norm_tac = 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   302
     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14320
diff changeset
   303
                                         hcomplex_minus_simps@add_ac))
14123
b300efca4ae0 fixed simprocs
paulson
parents: 13957
diff changeset
   304
     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@hcomplex_mult_minus_simps))
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   305
     THEN ALLGOALS
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   306
              (simp_tac (HOL_ss addsimps hcomplex_minus_from_mult_simps@
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14320
diff changeset
   307
                                         add_ac@mult_ac))
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   308
  val numeral_simp_tac	= ALLGOALS (simp_tac (HOL_ss addsimps add_0s@bin_simps))
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   309
  val simplify_meta_eq  = simplify_meta_eq
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   310
  end;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   311
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
structure EqCancelNumerals = CancelNumeralsFun
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   314
 (open CancelNumeralsCommon
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   315
  val prove_conv = Bin_Simprocs.prove_conv
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   316
  val mk_bal   = HOLogic.mk_eq
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   317
  val dest_bal = HOLogic.dest_bin "op =" hcomplexT
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14353
diff changeset
   318
  val bal_add1 = eq_add_iff1 RS trans
67a628beb981 tidying of the complex numbers
paulson
parents: 14353
diff changeset
   319
  val bal_add2 = eq_add_iff2 RS trans
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   320
);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   321
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
val cancel_numerals = 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   324
  map prep_simproc
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   325
   [("hcomplexeq_cancel_numerals",
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   326
      ["(l::hcomplex) + m = n", "(l::hcomplex) = m + n", 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   327
		"(l::hcomplex) - m = n", "(l::hcomplex) = m - n", 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   328
		"(l::hcomplex) * m = n", "(l::hcomplex) = m * n"], 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   329
     EqCancelNumerals.proc)];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   330
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   331
structure CombineNumeralsData =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   332
  struct
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   333
  val add		= op + : int*int -> int 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   334
  val mk_sum    	= long_mk_sum    (*to work for e.g. #2*x + #3*x *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   335
  val dest_sum		= dest_sum
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   336
  val mk_coeff		= mk_coeff
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   337
  val dest_coeff	= dest_coeff 1
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14353
diff changeset
   338
  val left_distrib	= combine_common_factor RS trans
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   339
  val prove_conv	= Bin_Simprocs.prove_conv_nohyps
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   340
  val trans_tac         = Real_Numeral_Simprocs.trans_tac
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   341
  val norm_tac = 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   342
     ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14320
diff changeset
   343
                                         hcomplex_minus_simps@add_ac))
14123
b300efca4ae0 fixed simprocs
paulson
parents: 13957
diff changeset
   344
     THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@hcomplex_mult_minus_simps))
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   345
     THEN ALLGOALS (simp_tac (HOL_ss addsimps hcomplex_minus_from_mult_simps@
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14320
diff changeset
   346
                                              add_ac@mult_ac))
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   347
  val numeral_simp_tac	= ALLGOALS 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   348
                    (simp_tac (HOL_ss addsimps add_0s@bin_simps))
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   349
  val simplify_meta_eq  = simplify_meta_eq
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   350
  end;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   351
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   352
structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   353
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   354
val combine_numerals = 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   355
    prep_simproc ("hcomplex_combine_numerals",
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   356
		  ["(i::hcomplex) + j", "(i::hcomplex) - j"],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   357
		  CombineNumerals.proc);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   358
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   359
(** Declarations for ExtractCommonTerm **)
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
(*this version ALWAYS includes a trailing one*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   362
fun long_mk_prod []        = one
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   363
  | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   364
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   365
(*Find first term that matches u*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   366
fun find_first past u []         = raise TERM("find_first", []) 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   367
  | find_first past u (t::terms) =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   368
	if u aconv t then (rev past @ terms)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   369
        else find_first (t::past) u terms
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   370
	handle TERM _ => find_first (t::past) u terms;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   371
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   372
(*Final simplification: cancel + and *  *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   373
fun cancel_simplify_meta_eq cancel_th th = 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   374
    Int_Numeral_Simprocs.simplify_meta_eq 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   375
        [hcomplex_mult_one_left, hcomplex_mult_one_right] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   376
        (([th, cancel_th]) MRS trans);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   377
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   378
(*** Making constant folding work for 0 and 1 too ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   379
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   380
structure HComplexAbstractNumeralsData =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   381
  struct
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   382
  val dest_eq         = HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   383
  val is_numeral      = Bin_Simprocs.is_numeral
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   384
  val numeral_0_eq_0  = hcomplex_numeral_0_eq_0
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   385
  val numeral_1_eq_1  = hcomplex_numeral_1_eq_1
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   386
  val prove_conv      = Bin_Simprocs.prove_conv_nohyps_novars
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   387
  fun norm_tac simps  = ALLGOALS (simp_tac (HOL_ss addsimps simps))
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   388
  val simplify_meta_eq = Bin_Simprocs.simplify_meta_eq
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   389
  end
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
structure HComplexAbstractNumerals = AbstractNumeralsFun (HComplexAbstractNumeralsData)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   392
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   393
(*For addition, we already have rules for the operand 0.
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   394
  Multiplication is omitted because there are already special rules for
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   395
  both 0 and 1 as operands.  Unary minus is trivial, just have - 1 = -1.
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   396
  For the others, having three patterns is a compromise between just having
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   397
  one (many spurious calls) and having nine (just too many!) *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   398
val eval_numerals =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   399
  map prep_simproc
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   400
   [("hcomplex_add_eval_numerals",
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   401
     ["(m::hcomplex) + 1", "(m::hcomplex) + number_of v"],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   402
     HComplexAbstractNumerals.proc add_hcomplex_number_of),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   403
    ("hcomplex_diff_eval_numerals",
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   404
     ["(m::hcomplex) - 1", "(m::hcomplex) - number_of v"],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   405
     HComplexAbstractNumerals.proc diff_hcomplex_number_of),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   406
    ("hcomplex_eq_eval_numerals",
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   407
     ["(m::hcomplex) = 0", "(m::hcomplex) = 1", "(m::hcomplex) = number_of v"],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   408
     HComplexAbstractNumerals.proc eq_hcomplex_number_of)]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   409
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   410
end;
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
Addsimprocs HComplex_Numeral_Simprocs.eval_numerals;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   413
Addsimprocs HComplex_Numeral_Simprocs.cancel_numerals;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   414
Addsimprocs [HComplex_Numeral_Simprocs.combine_numerals];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   415
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   416
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   417
(*examples:
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   418
print_depth 22;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   419
set timing;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   420
set trace_simp;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   421
fun test s = (Goal s, by (Simp_tac 1)); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   422
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   423
test "l +  2 +  2 +  2 + (l +  2) + (oo +  2) = (uu::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   424
test " 2*u = (u::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   425
test "(i + j + 12 + (k::hcomplex)) - 15 = y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   426
test "(i + j + 12 + (k::hcomplex)) -  5 = y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   427
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   428
test "( 2*x - (u*v) + y) - v* 3*u = (w::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   429
test "( 2*x*u*v + (u*v)* 4 + y) - v*u* 4 = (w::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   430
test "( 2*x*u*v + (u*v)* 4 + y) - v*u = (w::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   431
test "u*v - (x*u*v + (u*v)* 4 + y) = (w::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   432
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   433
test "(i + j + 12 + (k::hcomplex)) = u + 15 + y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   434
test "(i + j* 2 + 12 + (k::hcomplex)) = j +  5 + y";
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
test " 2*y +  3*z +  6*w +  2*y +  3*z +  2*u =  2*y' +  3*z' +  6*w' +  2*y' +  3*z' + u + (vv::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   437
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   438
test "a + -(b+c) + b = (d::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   439
test "a + -(b+c) - b = (d::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   440
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   441
(*negative numerals*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   442
test "(i + j +  -2 + (k::hcomplex)) - (u +  5 + y) = zz";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   443
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   444
test "(i + j +  -12 + (k::hcomplex)) - 15 = y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   445
test "(i + j + 12 + (k::hcomplex)) -  -15 = y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   446
test "(i + j +  -12 + (k::hcomplex)) - -15 = y";
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   449
(** Constant folding for hcomplex plus and times **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   450
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   451
structure HComplex_Times_Assoc_Data : ASSOC_FOLD_DATA =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   452
struct
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   453
  val ss		= HOL_ss
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   454
  val eq_reflection	= eq_reflection
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   455
  val sg_ref    = Sign.self_ref (Theory.sign_of (the_context ()))
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   456
  val T	     = HComplex_Numeral_Simprocs.hcomplexT
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   457
  val plus   = Const ("op *", [T,T] ---> T)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14320
diff changeset
   458
  val add_ac = mult_ac
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   459
end;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   460
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   461
structure HComplex_Times_Assoc = Assoc_Fold (HComplex_Times_Assoc_Data);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   462
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   463
Addsimprocs [HComplex_Times_Assoc.conv];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   464
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   465
Addsimps [hcomplex_of_complex_zero_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   466
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   467
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   468
(** extra thms **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   469
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   470
Goal "(hcnj z = 0) = (z = 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   471
by (auto_tac (claset(),simpset() addsimps [hcomplex_hcnj_zero_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   472
qed "hcomplex_hcnj_num_zero_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   473
Addsimps [hcomplex_hcnj_num_zero_iff];
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 "0 = Abs_hcomplex (hcomplexrel `` {%n. 0})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   476
by (simp_tac (simpset() addsimps [hcomplex_zero_def RS meta_eq_to_obj_eq RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   477
qed "hcomplex_zero_num";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   478
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   479
Goal "1 =  Abs_hcomplex (hcomplexrel `` {%n. 1})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   480
by (simp_tac (simpset() addsimps [hcomplex_one_def RS meta_eq_to_obj_eq RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   481
qed "hcomplex_one_num";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   482
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   483
(*** Real and imaginary stuff ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   484
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14373
diff changeset
   485
(*Convert???
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   486
Goalw [hcomplex_number_of_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   487
  "((number_of xa :: hcomplex) + iii * number_of ya = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   488
\       number_of xb + iii * number_of yb) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   489
\  (((number_of xa :: hcomplex) = number_of xb) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   490
\   ((number_of ya :: hcomplex) = number_of yb))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   491
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   492
     hcomplex_hypreal_number_of]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   493
qed "hcomplex_number_of_eq_cancel_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   494
Addsimps [hcomplex_number_of_eq_cancel_iff];
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
Goalw [hcomplex_number_of_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   497
  "((number_of xa :: hcomplex) + number_of ya * iii = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   498
\       number_of xb + number_of yb * iii) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   499
\  (((number_of xa :: hcomplex) = number_of xb) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   500
\   ((number_of ya :: hcomplex) = number_of yb))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   501
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffA,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   502
    hcomplex_hypreal_number_of]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   503
qed "hcomplex_number_of_eq_cancel_iffA";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   504
Addsimps [hcomplex_number_of_eq_cancel_iffA];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   505
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   506
Goalw [hcomplex_number_of_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   507
  "((number_of xa :: hcomplex) + number_of ya * iii = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   508
\       number_of xb + iii * number_of yb) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   509
\  (((number_of xa :: hcomplex) = number_of xb) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   510
\   ((number_of ya :: hcomplex) = number_of yb))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   511
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffB,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   512
    hcomplex_hypreal_number_of]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   513
qed "hcomplex_number_of_eq_cancel_iffB";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   514
Addsimps [hcomplex_number_of_eq_cancel_iffB];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   515
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   516
Goalw [hcomplex_number_of_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   517
  "((number_of xa :: hcomplex) + iii * number_of ya = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   518
\       number_of xb + number_of yb * iii) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   519
\  (((number_of xa :: hcomplex) = number_of xb) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   520
\   ((number_of ya :: hcomplex) = number_of yb))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   521
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffC,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   522
     hcomplex_hypreal_number_of]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   523
qed "hcomplex_number_of_eq_cancel_iffC";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   524
Addsimps [hcomplex_number_of_eq_cancel_iffC];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   525
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   526
Goalw [hcomplex_number_of_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   527
  "((number_of xa :: hcomplex) + iii * number_of ya = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   528
\       number_of xb) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   529
\  (((number_of xa :: hcomplex) = number_of xb) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   530
\   ((number_of ya :: hcomplex) = 0))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   531
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff2,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   532
    hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   533
qed "hcomplex_number_of_eq_cancel_iff2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   534
Addsimps [hcomplex_number_of_eq_cancel_iff2];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   535
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   536
Goalw [hcomplex_number_of_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   537
  "((number_of xa :: hcomplex) + number_of ya * iii = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   538
\       number_of xb) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   539
\  (((number_of xa :: hcomplex) = number_of xb) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   540
\   ((number_of ya :: hcomplex) = 0))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   541
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff2a,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   542
    hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   543
qed "hcomplex_number_of_eq_cancel_iff2a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   544
Addsimps [hcomplex_number_of_eq_cancel_iff2a];
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
Goalw [hcomplex_number_of_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   547
  "((number_of xa :: hcomplex) + iii * number_of ya = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   548
\    iii * number_of yb) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   549
\  (((number_of xa :: hcomplex) = 0) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   550
\   ((number_of ya :: hcomplex) = number_of yb))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   551
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff3,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   552
    hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   553
qed "hcomplex_number_of_eq_cancel_iff3";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   554
Addsimps [hcomplex_number_of_eq_cancel_iff3];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   555
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   556
Goalw [hcomplex_number_of_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   557
  "((number_of xa :: hcomplex) + number_of ya * iii= \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   558
\    iii * number_of yb) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   559
\  (((number_of xa :: hcomplex) = 0) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   560
\   ((number_of ya :: hcomplex) = number_of yb))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   561
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff3a,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   562
    hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   563
qed "hcomplex_number_of_eq_cancel_iff3a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   564
Addsimps [hcomplex_number_of_eq_cancel_iff3a];
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14373
diff changeset
   565
*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   566
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   567
Goalw [hcomplex_number_of_def] "hcnj (number_of v :: hcomplex) = number_of v";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   568
by (rtac (hcomplex_hypreal_number_of RS ssubst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   569
by (rtac hcomplex_hcnj_hcomplex_of_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   570
qed "hcomplex_number_of_hcnj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   571
Addsimps [hcomplex_number_of_hcnj];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   572
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   573
Goalw [hcomplex_number_of_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   574
      "hcmod(number_of v :: hcomplex) = abs (number_of v :: hypreal)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   575
by (rtac (hcomplex_hypreal_number_of RS ssubst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   576
by (auto_tac (claset(), HOL_ss addsimps [hcmod_hcomplex_of_hypreal]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   577
qed "hcomplex_number_of_hcmod";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   578
Addsimps [hcomplex_number_of_hcmod];
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
Goalw [hcomplex_number_of_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   581
      "hRe(number_of v :: hcomplex) = number_of v";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   582
by (rtac (hcomplex_hypreal_number_of RS ssubst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   583
by (auto_tac (claset(), HOL_ss addsimps [hRe_hcomplex_of_hypreal]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   584
qed "hcomplex_number_of_hRe";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   585
Addsimps [hcomplex_number_of_hRe];
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
Goalw [hcomplex_number_of_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   588
      "hIm(number_of v :: hcomplex) = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   589
by (rtac (hcomplex_hypreal_number_of RS ssubst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   590
by (auto_tac (claset(), HOL_ss addsimps [hIm_hcomplex_of_hypreal]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   591
qed "hcomplex_number_of_hIm";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   592
Addsimps [hcomplex_number_of_hIm];
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   595