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