# HG changeset patch # User paulson # Date 1052151751 -7200 # Node ID 10dbf16be15ffddfd76490e0b7535f9788959643 # Parent 8fe7e12290e1637008f826500ab5db681a13b330 new session Complex for the complex numbers diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/CLim.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/CLim.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,1191 @@ +(* Title : CLim.ML + Author : Jacques D. Fleuriot + Copyright : 2001 University of Edinburgh + Description : A first theory of limits, continuity and + differentiation for complex functions +*) + +(*------------------------------------------------------------------------------------*) +(* Limit of complex to complex function *) +(*------------------------------------------------------------------------------------*) + +Goalw [NSCLIM_def,NSCRLIM_def] + "f -- a --NSC> L ==> (%x. Re(f x)) -- a --NSCR> Re(L)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunC_approx_Re_Im_iff, + hRe_hcomplex_of_complex])); +qed "NSCLIM_NSCRLIM_Re"; + +Goalw [NSCLIM_def,NSCRLIM_def] + "f -- a --NSC> L ==> (%x. Im(f x)) -- a --NSCR> Im(L)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunC_approx_Re_Im_iff, + hIm_hcomplex_of_complex])); +qed "NSCLIM_NSCRLIM_Im"; + +Goalw [CLIM_def,NSCLIM_def,capprox_def] + "f -- x --C> L ==> f -- x --NSC> L"; +by Auto_tac; +by (res_inst_tac [("z","xa")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_complex_def, + starfunC,hcomplex_diff,CInfinitesimal_hcmod_iff,hcmod, + Infinitesimal_FreeUltrafilterNat_iff])); +by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2); +by (Step_tac 1); +by (dres_inst_tac [("x","u")] spec 1 THEN Auto_tac); +by (dres_inst_tac [("x","s")] spec 1 THEN Auto_tac); +by (Ultra_tac 1); +by (dtac sym 1 THEN Auto_tac); +qed "CLIM_NSCLIM"; + +Goal "(ALL t. P t) = (ALL X. P (Abs_hcomplex(hcomplexrel `` {X})))"; +by Auto_tac; +by (res_inst_tac [("z","t")] eq_Abs_hcomplex 1); +by Auto_tac; +qed "eq_Abs_hcomplex_ALL"; + +Goal "ALL s. 0 < s --> (EX xa. xa ~= x & \ +\ cmod (xa - x) < s & r <= cmod (f xa - L)) \ +\ ==> ALL (n::nat). EX xa. xa ~= x & \ +\ cmod(xa - x) < inverse(real(Suc n)) & r <= cmod(f xa - L)"; +by (Clarify_tac 1); +by (cut_inst_tac [("n1","n")] + (real_of_nat_Suc_gt_zero RS real_inverse_gt_0) 1); +by Auto_tac; +val lemma_CLIM = result(); + +(* not needed? *) +Goal "ALL x z. EX y. Q x z y ==> EX f. ALL x z. Q x z (f x z)"; +by (rtac choice 1 THEN Step_tac 1); +by (blast_tac (claset() addIs [choice]) 1); +qed "choice2"; + +Goal "ALL s. 0 < s --> (EX xa. xa ~= x & \ +\ cmod (xa - x) < s & r <= cmod (f xa - L)) \ +\ ==> EX X. ALL (n::nat). X n ~= x & \ +\ cmod(X n - x) < inverse(real(Suc n)) & r <= cmod(f (X n) - L)"; +by (dtac lemma_CLIM 1); +by (dtac choice 1); +by (Blast_tac 1); +val lemma_skolemize_CLIM2 = result(); + +Goal "ALL n. X n ~= x & \ +\ cmod (X n - x) < inverse (real(Suc n)) & \ +\ r <= cmod (f (X n) - L) ==> \ +\ ALL n. cmod (X n - x) < inverse (real(Suc n))"; +by (Auto_tac ); +val lemma_csimp = result(); + +Goalw [CLIM_def,NSCLIM_def] + "f -- x --NSC> L ==> f -- x --C> L"; +by (auto_tac (claset(),simpset() addsimps [eq_Abs_hcomplex_ALL, + starfunC,CInfinitesimal_capprox_minus RS sym,hcomplex_diff, + CInfinitesimal_hcmod_iff,hcomplex_of_complex_def, + Infinitesimal_FreeUltrafilterNat_iff,hcmod])); +by (EVERY1[rtac ccontr, Asm_full_simp_tac]); +by (fold_tac [real_le_def]); +by (dtac lemma_skolemize_CLIM2 1); +by (Step_tac 1); +by (dres_inst_tac [("x","X")] spec 1); +by Auto_tac; +by (dtac (lemma_csimp RS complex_seq_to_hcomplex_CInfinitesimal) 1); +by (asm_full_simp_tac (simpset() addsimps [CInfinitesimal_hcmod_iff, + hcomplex_of_complex_def,Infinitesimal_FreeUltrafilterNat_iff, + hcomplex_diff,hcmod]) 1); +by (Blast_tac 1); +by (dres_inst_tac [("x","r")] spec 1); +by (Clarify_tac 1); +by (dtac FreeUltrafilterNat_all 1); +by (Ultra_tac 1); +by (arith_tac 1); +qed "NSCLIM_CLIM"; + +(**** First key result ****) + +Goal "(f -- x --C> L) = (f -- x --NSC> L)"; +by (blast_tac (claset() addIs [CLIM_NSCLIM,NSCLIM_CLIM]) 1); +qed "CLIM_NSCLIM_iff"; + +(*------------------------------------------------------------------------------------*) +(* Limit of complex to real function *) +(*------------------------------------------------------------------------------------*) + +Goalw [CRLIM_def,NSCRLIM_def,capprox_def] + "f -- x --CR> L ==> f -- x --NSCR> L"; +by Auto_tac; +by (res_inst_tac [("z","xa")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_complex_def, + starfunCR,hcomplex_diff,CInfinitesimal_hcmod_iff,hcmod,hypreal_diff, + Infinitesimal_FreeUltrafilterNat_iff,Infinitesimal_approx_minus RS sym, + hypreal_of_real_def])); +by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2); +by (Step_tac 1); +by (dres_inst_tac [("x","u")] spec 1 THEN Auto_tac); +by (dres_inst_tac [("x","s")] spec 1 THEN Auto_tac); +by (Ultra_tac 1); +by (dtac sym 1 THEN Auto_tac); +qed "CRLIM_NSCRLIM"; + +Goal "ALL s. 0 < s --> (EX xa. xa ~= x & \ +\ cmod (xa - x) < s & r <= abs (f xa - L)) \ +\ ==> ALL (n::nat). EX xa. xa ~= x & \ +\ cmod(xa - x) < inverse(real(Suc n)) & r <= abs (f xa - L)"; +by (Clarify_tac 1); +by (cut_inst_tac [("n1","n")] + (real_of_nat_Suc_gt_zero RS real_inverse_gt_0) 1); +by Auto_tac; +val lemma_CRLIM = result(); + +Goal "ALL s. 0 < s --> (EX xa. xa ~= x & \ +\ cmod (xa - x) < s & r <= abs (f xa - L)) \ +\ ==> EX X. ALL (n::nat). X n ~= x & \ +\ cmod(X n - x) < inverse(real(Suc n)) & r <= abs (f (X n) - L)"; +by (dtac lemma_CRLIM 1); +by (dtac choice 1); +by (Blast_tac 1); +val lemma_skolemize_CRLIM2 = result(); + +Goal "ALL n. X n ~= x & \ +\ cmod (X n - x) < inverse (real(Suc n)) & \ +\ r <= abs (f (X n) - L) ==> \ +\ ALL n. cmod (X n - x) < inverse (real(Suc n))"; +by (Auto_tac ); +val lemma_crsimp = result(); + +Goalw [CRLIM_def,NSCRLIM_def,capprox_def] + "f -- x --NSCR> L ==> f -- x --CR> L"; +by (auto_tac (claset(),simpset() addsimps [eq_Abs_hcomplex_ALL, + starfunCR,hcomplex_diff,hcomplex_of_complex_def,hypreal_diff, + CInfinitesimal_hcmod_iff,hcmod,Infinitesimal_approx_minus RS sym, + Infinitesimal_FreeUltrafilterNat_iff])); +by (EVERY1[rtac ccontr, Asm_full_simp_tac]); +by (fold_tac [real_le_def]); +by (dtac lemma_skolemize_CRLIM2 1); +by (Step_tac 1); +by (dres_inst_tac [("x","X")] spec 1); +by Auto_tac; +by (dtac (lemma_crsimp RS complex_seq_to_hcomplex_CInfinitesimal) 1); +by (asm_full_simp_tac (simpset() addsimps [CInfinitesimal_hcmod_iff, + hcomplex_of_complex_def,Infinitesimal_FreeUltrafilterNat_iff, + hcomplex_diff,hcmod]) 1); +by (Blast_tac 1); +by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_def, + hypreal_diff])); +by (dres_inst_tac [("x","r")] spec 1); +by (Clarify_tac 1); +by (dtac FreeUltrafilterNat_all 1); +by (Ultra_tac 1); +qed "NSCRLIM_CRLIM"; + +(** second key result **) +Goal "(f -- x --CR> L) = (f -- x --NSCR> L)"; +by (blast_tac (claset() addIs [CRLIM_NSCRLIM,NSCRLIM_CRLIM]) 1); +qed "CRLIM_NSCRLIM_iff"; + +(** get this result easily now **) +Goal "f -- a --C> L ==> (%x. Re(f x)) -- a --CR> Re(L)"; +by (auto_tac (claset() addDs [NSCLIM_NSCRLIM_Re],simpset() + addsimps [CLIM_NSCLIM_iff,CRLIM_NSCRLIM_iff RS sym])); +qed "CLIM_CRLIM_Re"; + +Goal "f -- a --C> L ==> (%x. Im(f x)) -- a --CR> Im(L)"; +by (auto_tac (claset() addDs [NSCLIM_NSCRLIM_Im],simpset() + addsimps [CLIM_NSCLIM_iff,CRLIM_NSCRLIM_iff RS sym])); +qed "CLIM_CRLIM_Im"; + +Goal "f -- a --C> L ==> (%x. cnj (f x)) -- a --C> cnj L"; +by (auto_tac (claset(),simpset() addsimps [CLIM_def, + complex_cnj_diff RS sym])); +qed "CLIM_cnj"; + +Goal "((%x. cnj (f x)) -- a --C> cnj L) = (f -- a --C> L)"; +by (auto_tac (claset(),simpset() addsimps [CLIM_def, + complex_cnj_diff RS sym])); +qed "CLIM_cnj_iff"; + +(*** NSLIM_add hence CLIM_add *) + +Goalw [NSCLIM_def] + "[| f -- x --NSC> l; g -- x --NSC> m |] \ +\ ==> (%x. f(x) + g(x)) -- x --NSC> (l + m)"; +by (auto_tac (claset() addSIs [capprox_add], simpset())); +qed "NSCLIM_add"; + +Goal "[| f -- x --C> l; g -- x --C> m |] \ +\ ==> (%x. f(x) + g(x)) -- x --C> (l + m)"; +by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_add]) 1); +qed "CLIM_add"; + +(*** NSLIM_mult hence CLIM_mult *) + +Goalw [NSCLIM_def] + "[| f -- x --NSC> l; g -- x --NSC> m |] \ +\ ==> (%x. f(x) * g(x)) -- x --NSC> (l * m)"; +by (auto_tac (claset() addSIs [capprox_mult_CFinite], simpset())); +qed "NSCLIM_mult"; + +Goal "[| f -- x --C> l; g -- x --C> m |] \ +\ ==> (%x. f(x) * g(x)) -- x --C> (l * m)"; +by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_mult]) 1); +qed "CLIM_mult"; + +(*** NSCLIM_const and CLIM_const ***) + +Goalw [NSCLIM_def] "(%x. k) -- x --NSC> k"; +by Auto_tac; +qed "NSCLIM_const"; +Addsimps [NSCLIM_const]; + +Goalw [CLIM_def] "(%x. k) -- x --C> k"; +by Auto_tac; +qed "CLIM_const"; +Addsimps [CLIM_const]; + +(*** NSCLIM_minus and CLIM_minus ***) + +Goalw [NSCLIM_def] + "f -- a --NSC> L ==> (%x. -f(x)) -- a --NSC> -L"; +by Auto_tac; +qed "NSCLIM_minus"; + +Goal "f -- a --C> L ==> (%x. -f(x)) -- a --C> -L"; +by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_minus]) 1); +qed "CLIM_minus"; + +(*** NSCLIM_diff hence CLIM_diff ***) + +Goalw [complex_diff_def] + "[| f -- x --NSC> l; g -- x --NSC> m |] \ +\ ==> (%x. f(x) - g(x)) -- x --NSC> (l - m)"; +by (auto_tac (claset(), simpset() addsimps [NSCLIM_add,NSCLIM_minus])); +qed "NSCLIM_diff"; + +Goal "[| f -- x --C> l; g -- x --C> m |] \ +\ ==> (%x. f(x) - g(x)) -- x --C> (l - m)"; +by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_diff]) 1); +qed "CLIM_diff"; + +(*** NSCLIM_inverse and hence CLIM_inverse *) + +Goalw [NSCLIM_def] + "[| f -- a --NSC> L; L ~= 0 |] \ +\ ==> (%x. inverse(f(x))) -- a --NSC> (inverse L)"; +by (Clarify_tac 1); +by (dtac spec 1); +by (auto_tac (claset(), + simpset() addsimps [hcomplex_of_complex_capprox_inverse])); +qed "NSCLIM_inverse"; + +Goal "[| f -- a --C> L; L ~= 0 |] \ +\ ==> (%x. inverse(f(x))) -- a --C> (inverse L)"; +by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_inverse]) 1); +qed "CLIM_inverse"; + +(*** NSCLIM_zero, CLIM_zero, etc. ***) + +Goal "f -- a --NSC> l ==> (%x. f(x) - l) -- a --NSC> 0"; +by (res_inst_tac [("z1","l")] (complex_add_minus_right_zero RS subst) 1); +by (rewtac complex_diff_def); +by (rtac NSCLIM_add 1 THEN Auto_tac); +qed "NSCLIM_zero"; + +Goal "f -- a --C> l ==> (%x. f(x) - l) -- a --C> 0"; +by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_zero]) 1); +qed "CLIM_zero"; + +Goal "(%x. f(x) - l) -- x --NSC> 0 ==> f -- x --NSC> l"; +by (dres_inst_tac [("g","%x. l"),("m","l")] NSCLIM_add 1); +by Auto_tac; +qed "NSCLIM_zero_cancel"; + +Goal "(%x. f(x) - l) -- x --C> 0 ==> f -- x --C> l"; +by (dres_inst_tac [("g","%x. l"),("m","l")] CLIM_add 1); +by Auto_tac; +qed "CLIM_zero_cancel"; + +(*** NSCLIM_not zero and hence CLIM_not_zero ***) + +(*not in simpset?*) +Addsimps [hypreal_epsilon_not_zero]; + +Goalw [NSCLIM_def] "k ~= 0 ==> ~ ((%x. k) -- x --NSC> 0)"; +by (auto_tac (claset(),simpset() delsimps [hcomplex_of_complex_zero])); +by (res_inst_tac [("x","hcomplex_of_complex x + hcomplex_of_hypreal epsilon")] exI 1); +by (auto_tac (claset() addIs [CInfinitesimal_add_capprox_self RS capprox_sym],simpset() + delsimps [hcomplex_of_complex_zero])); +qed "NSCLIM_not_zero"; + +(* [| k ~= 0; (%x. k) -- x --NSC> 0 |] ==> R *) +bind_thm("NSCLIM_not_zeroE", NSCLIM_not_zero RS notE); + +Goal "k ~= 0 ==> ~ ((%x. k) -- x --C> 0)"; +by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_not_zero]) 1); +qed "CLIM_not_zero"; + +(*** NSCLIM_const hence CLIM_const ***) + +Goal "(%x. k) -- x --NSC> L ==> k = L"; +by (rtac ccontr 1); +by (dtac NSCLIM_zero 1); +by (rtac NSCLIM_not_zeroE 1 THEN assume_tac 2); +by Auto_tac; +qed "NSCLIM_const_eq"; + +Goal "(%x. k) -- x --C> L ==> k = L"; +by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff,NSCLIM_const_eq]) 1); +qed "CLIM_const_eq"; + +(*** NSCLIM and hence CLIM are unique ***) + +Goal "[| f -- x --NSC> L; f -- x --NSC> M |] ==> L = M"; +by (dtac NSCLIM_minus 1); +by (dtac NSCLIM_add 1 THEN assume_tac 1); +by (auto_tac (claset() addSDs [NSCLIM_const_eq RS sym], simpset())); +qed "NSCLIM_unique"; + +Goal "[| f -- x --C> L; f -- x --C> M |] ==> L = M"; +by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_unique]) 1); +qed "CLIM_unique"; + +(*** NSCLIM_mult_zero and CLIM_mult_zero ***) + +Goal "[| f -- x --NSC> 0; g -- x --NSC> 0 |] \ +\ ==> (%x. f(x)*g(x)) -- x --NSC> 0"; +by (dtac NSCLIM_mult 1 THEN Auto_tac); +qed "NSCLIM_mult_zero"; + +Goal "[| f -- x --C> 0; g -- x --C> 0 |] \ +\ ==> (%x. f(x)*g(x)) -- x --C> 0"; +by (dtac CLIM_mult 1 THEN Auto_tac); +qed "CLIM_mult_zero"; + +(*** NSCLIM_self hence CLIM_self ***) + +Goalw [NSCLIM_def] "(%x. x) -- a --NSC> a"; +by (auto_tac (claset() addIs [starfunC_Idfun_capprox],simpset())); +qed "NSCLIM_self"; + +Goal "(%x. x) -- a --C> a"; +by (simp_tac (simpset() addsimps [CLIM_NSCLIM_iff,NSCLIM_self]) 1); +qed "CLIM_self"; + +(** another equivalence result **) +Goalw [NSCLIM_def,NSCRLIM_def] + "(f -- x --NSC> L) = ((%y. cmod(f y - L)) -- x --NSCR> 0)"; +by (auto_tac (claset(),simpset() addsimps [CInfinitesimal_capprox_minus + RS sym,CInfinitesimal_hcmod_iff])); +by (ALLGOALS(dtac spec) THEN Auto_tac); +by (ALLGOALS(res_inst_tac [("z","xa")] eq_Abs_hcomplex)); +by (auto_tac (claset(),simpset() addsimps [hcomplex_diff, + starfunC,starfunCR,hcomplex_of_complex_def,hcmod,mem_infmal_iff])); +qed "NSCLIM_NSCRLIM_iff"; + +(** much, much easier standard proof **) +Goalw [CLIM_def,CRLIM_def] + "(f -- x --C> L) = ((%y. cmod(f y - L)) -- x --CR> 0)"; +by Auto_tac; +qed "CLIM_CRLIM_iff"; + +(* so this is nicer nonstandard proof *) +Goal "(f -- x --NSC> L) = ((%y. cmod(f y - L)) -- x --NSCR> 0)"; +by (auto_tac (claset(),simpset() addsimps [CRLIM_NSCRLIM_iff RS sym, + CLIM_CRLIM_iff,CLIM_NSCLIM_iff RS sym])); +qed "NSCLIM_NSCRLIM_iff2"; + +Goal "(f -- a --NSC> L) = ((%x. Re(f x)) -- a --NSCR> Re(L) & \ +\ (%x. Im(f x)) -- a --NSCR> Im(L))"; +by (auto_tac (claset() addIs [NSCLIM_NSCRLIM_Re,NSCLIM_NSCRLIM_Im],simpset())); +by (auto_tac (claset(),simpset() addsimps [NSCLIM_def,NSCRLIM_def])); +by (REPEAT(dtac spec 1) THEN Auto_tac); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [capprox_approx_iff,starfunC, + hcomplex_of_complex_def,starfunCR,hypreal_of_real_def])); +qed "NSCLIM_NSCRLIM_Re_Im_iff"; + +Goal "(f -- a --C> L) = ((%x. Re(f x)) -- a --CR> Re(L) & \ +\ (%x. Im(f x)) -- a --CR> Im(L))"; +by (auto_tac (claset(),simpset() addsimps [CLIM_NSCLIM_iff,CRLIM_NSCRLIM_iff, + NSCLIM_NSCRLIM_Re_Im_iff])); +qed "CLIM_CRLIM_Re_Im_iff"; + + +(*------------------------------------------------------------------------------------*) +(* Continuity *) +(*------------------------------------------------------------------------------------*) + +Goalw [isNSContc_def] + "[| isNSContc f a; y @c= hcomplex_of_complex a |] \ +\ ==> ( *fc* f) y @c= hcomplex_of_complex (f a)"; +by (Blast_tac 1); +qed "isNSContcD"; + +Goalw [isNSContc_def,NSCLIM_def] + "isNSContc f a ==> f -- a --NSC> (f a) "; +by (Blast_tac 1); +qed "isNSContc_NSCLIM"; + +Goalw [isNSContc_def,NSCLIM_def] + "f -- a --NSC> (f a) ==> isNSContc f a"; +by Auto_tac; +by (res_inst_tac [("Q","y = hcomplex_of_complex a")] + (excluded_middle RS disjE) 1); +by Auto_tac; +qed "NSCLIM_isNSContc"; + +(*--------------------------------------------------*) +(* NS continuity can be defined using NS Limit in *) +(* similar fashion to standard def of continuity *) +(* -------------------------------------------------*) + +Goal "(isNSContc f a) = (f -- a --NSC> (f a))"; +by (blast_tac (claset() addIs [isNSContc_NSCLIM,NSCLIM_isNSContc]) 1); +qed "isNSContc_NSCLIM_iff"; + +Goal "(isNSContc f a) = (f -- a --C> (f a))"; +by (asm_full_simp_tac (simpset() addsimps + [CLIM_NSCLIM_iff,isNSContc_NSCLIM_iff]) 1); +qed "isNSContc_CLIM_iff"; + +(*** key result for continuity ***) +Goalw [isContc_def] "(isNSContc f a) = (isContc f a)"; +by (rtac isNSContc_CLIM_iff 1); +qed "isNSContc_isContc_iff"; + +Goal "isContc f a ==> isNSContc f a"; +by (etac (isNSContc_isContc_iff RS iffD2) 1); +qed "isContc_isNSContc"; + +Goal "isNSContc f a ==> isContc f a"; +by (etac (isNSContc_isContc_iff RS iffD1) 1); +qed "isNSContc_isContc"; + +(*--------------------------------------------------*) +(* Alternative definition of continuity *) +(* -------------------------------------------------*) + +Goalw [NSCLIM_def] + "(f -- a --NSC> L) = ((%h. f(a + h)) -- 0 --NSC> L)"; +by Auto_tac; +by (dres_inst_tac [("x","hcomplex_of_complex a + x")] spec 1); +by (dres_inst_tac [("x","- hcomplex_of_complex a + x")] spec 2); +by (Step_tac 1); +by (Asm_full_simp_tac 1); +by (rtac ((mem_cinfmal_iff RS iffD2) RS + (CInfinitesimal_add_capprox_self RS capprox_sym)) 1); +by (rtac (capprox_minus_iff2 RS iffD1) 4); +by (asm_full_simp_tac (simpset() addsimps [hcomplex_add_commute]) 3); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 2); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 4); +by (auto_tac (claset(), + simpset() addsimps [starfunC, hcomplex_of_complex_def, + hcomplex_minus, hcomplex_add])); +qed "NSCLIM_h_iff"; + +Goal "(f -- a --NSC> f a) = ((%h. f(a + h)) -- 0 --NSC> f a)"; +by (rtac NSCLIM_h_iff 1); +qed "NSCLIM_isContc_iff"; + +Goal "(f -- a --C> f a) = ((%h. f(a + h)) -- 0 --C> f(a))"; +by (simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_isContc_iff]) 1); +qed "CLIM_isContc_iff"; + +Goalw [isContc_def] "(isContc f x) = ((%h. f(x + h)) -- 0 --C> f(x))"; +by (simp_tac (simpset() addsimps [CLIM_isContc_iff]) 1); +qed "isContc_iff"; + +Goal "[| isContc f a; isContc g a |] ==> isContc (%x. f(x) + g(x)) a"; +by (auto_tac (claset() addIs [capprox_add], + simpset() addsimps [isNSContc_isContc_iff RS sym, isNSContc_def])); +qed "isContc_add"; + +Goal "[| isContc f a; isContc g a |] ==> isContc (%x. f(x) * g(x)) a"; +by (auto_tac (claset() addSIs [starfunC_mult_CFinite_capprox], + simpset() delsimps [starfunC_mult RS sym] + addsimps [isNSContc_isContc_iff RS sym, isNSContc_def])); +qed "isContc_mult"; + +(*** more theorems: note simple proofs ***) + +Goal "[| isContc f a; isContc g (f a) |] \ +\ ==> isContc (g o f) a"; +by (auto_tac (claset(),simpset() addsimps [isNSContc_isContc_iff RS sym, + isNSContc_def,starfunC_o RS sym])); +qed "isContc_o"; + +Goal "[| isContc f a; isContc g (f a) |] \ +\ ==> isContc (%x. g (f x)) a"; +by (auto_tac (claset() addDs [isContc_o],simpset() addsimps [o_def])); +qed "isContc_o2"; + +Goalw [isNSContc_def] "isNSContc f a ==> isNSContc (%x. - f x) a"; +by Auto_tac; +qed "isNSContc_minus"; + +Goal "isContc f a ==> isContc (%x. - f x) a"; +by (auto_tac (claset(),simpset() addsimps [isNSContc_isContc_iff RS sym, + isNSContc_minus])); +qed "isContc_minus"; + +Goalw [isContc_def] + "[| isContc f x; f x ~= 0 |] ==> isContc (%x. inverse (f x)) x"; +by (blast_tac (claset() addIs [CLIM_inverse]) 1); +qed "isContc_inverse"; + +Goal "[| isNSContc f x; f x ~= 0 |] ==> isNSContc (%x. inverse (f x)) x"; +by (auto_tac (claset() addIs [isContc_inverse],simpset() addsimps + [isNSContc_isContc_iff])); +qed "isNSContc_inverse"; + +Goalw [complex_diff_def] + "[| isContc f a; isContc g a |] ==> isContc (%x. f(x) - g(x)) a"; +by (auto_tac (claset() addIs [isContc_add,isContc_minus],simpset())); +qed "isContc_diff"; + +Goalw [isContc_def] "isContc (%x. k) a"; +by (Simp_tac 1); +qed "isContc_const"; +Addsimps [isContc_const]; + +Goalw [isNSContc_def] "isNSContc (%x. k) a"; +by (Simp_tac 1); +qed "isNSContc_const"; +Addsimps [isNSContc_const]; + + +(*------------------------------------------------------------------------------------*) +(* functions from complex to reals *) +(* -----------------------------------------------------------------------------------*) + +Goalw [isNSContCR_def] + "[| isNSContCR f a; y @c= hcomplex_of_complex a |] \ +\ ==> ( *fcR* f) y @= hypreal_of_real (f a)"; +by (Blast_tac 1); +qed "isNSContCRD"; + +Goalw [isNSContCR_def,NSCRLIM_def] + "isNSContCR f a ==> f -- a --NSCR> (f a) "; +by (Blast_tac 1); +qed "isNSContCR_NSCRLIM"; + +Goalw [isNSContCR_def,NSCRLIM_def] + "f -- a --NSCR> (f a) ==> isNSContCR f a"; +by Auto_tac; +by (res_inst_tac [("Q","y = hcomplex_of_complex a")] + (excluded_middle RS disjE) 1); +by Auto_tac; +qed "NSCRLIM_isNSContCR"; + +Goal "(isNSContCR f a) = (f -- a --NSCR> (f a))"; +by (blast_tac (claset() addIs [isNSContCR_NSCRLIM,NSCRLIM_isNSContCR]) 1); +qed "isNSContCR_NSCRLIM_iff"; + +Goal "(isNSContCR f a) = (f -- a --CR> (f a))"; +by (asm_full_simp_tac (simpset() addsimps + [CRLIM_NSCRLIM_iff,isNSContCR_NSCRLIM_iff]) 1); +qed "isNSContCR_CRLIM_iff"; + +(*** another key result for continuity ***) +Goalw [isContCR_def] "(isNSContCR f a) = (isContCR f a)"; +by (rtac isNSContCR_CRLIM_iff 1); +qed "isNSContCR_isContCR_iff"; + +Goal "isContCR f a ==> isNSContCR f a"; +by (etac (isNSContCR_isContCR_iff RS iffD2) 1); +qed "isContCR_isNSContCR"; + +Goal "isNSContCR f a ==> isContCR f a"; +by (etac (isNSContCR_isContCR_iff RS iffD1) 1); +qed "isNSContCR_isContCR"; + +Goalw [isNSContCR_def] "isNSContCR cmod (a)"; +by (auto_tac (claset() addIs [capprox_hcmod_approx], + simpset() addsimps [starfunCR_cmod,hcmod_hcomplex_of_complex + RS sym])); +qed "isNSContCR_cmod"; +Addsimps [isNSContCR_cmod]; + +Goal "isContCR cmod (a)"; +by (auto_tac (claset(),simpset() addsimps [isNSContCR_isContCR_iff RS sym])); +qed "isContCR_cmod"; +Addsimps [isContCR_cmod]; + +Goalw [isContc_def,isContCR_def] + "isContc f a ==> isContCR (%x. Re (f x)) a"; +by (etac CLIM_CRLIM_Re 1); +qed "isContc_isContCR_Re"; + +Goalw [isContc_def,isContCR_def] + "isContc f a ==> isContCR (%x. Im (f x)) a"; +by (etac CLIM_CRLIM_Im 1); +qed "isContc_isContCR_Im"; + +(*------------------------------------------------------------------------------------*) +(* Derivatives *) +(*------------------------------------------------------------------------------------*) + +Goalw [cderiv_def] + "(CDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --C> D)"; +by (Blast_tac 1); +qed "CDERIV_iff"; + +Goalw [cderiv_def] + "(CDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --NSC> D)"; +by (simp_tac (simpset() addsimps [CLIM_NSCLIM_iff]) 1); +qed "CDERIV_NSC_iff"; + +Goalw [cderiv_def] + "CDERIV f x :> D \ +\ ==> (%h. (f(x + h) - f(x))/h) -- 0 --C> D"; +by (Blast_tac 1); +qed "CDERIVD"; + +Goalw [cderiv_def] + "CDERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) -- 0 --NSC> D"; +by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff]) 1); +qed "NSC_DERIVD"; + +(*** Uniqueness ***) + +Goalw [cderiv_def] + "[| CDERIV f x :> D; CDERIV f x :> E |] ==> D = E"; +by (blast_tac (claset() addIs [CLIM_unique]) 1); +qed "CDERIV_unique"; + +(*** uniqueness: a nonstandard proof ***) +Goalw [nscderiv_def] + "[| NSCDERIV f x :> D; NSCDERIV f x :> E |] ==> D = E"; +by (auto_tac (claset() addSDs [inst "x" "hcomplex_of_hypreal epsilon" bspec] + addSIs [inj_hcomplex_of_complex RS injD] + addDs [capprox_trans3], + simpset())); +qed "NSCDeriv_unique"; + + +(*------------------------------------------------------------------------------------*) +(* Differentiability *) +(*------------------------------------------------------------------------------------*) + +Goalw [cdifferentiable_def] + "f cdifferentiable x ==> EX D. CDERIV f x :> D"; +by (assume_tac 1); +qed "cdifferentiableD"; + +Goalw [cdifferentiable_def] + "CDERIV f x :> D ==> f cdifferentiable x"; +by (Blast_tac 1); +qed "cdifferentiableI"; + +Goalw [NSCdifferentiable_def] + "f NSCdifferentiable x ==> EX D. NSCDERIV f x :> D"; +by (assume_tac 1); +qed "NSCdifferentiableD"; + +Goalw [NSCdifferentiable_def] + "NSCDERIV f x :> D ==> f NSCdifferentiable x"; +by (Blast_tac 1); +qed "NSCdifferentiableI"; + + +(*------------------------------------------------------------------------------------*) +(* Alternative definition for differentiability *) +(*------------------------------------------------------------------------------------*) + +Goalw [CLIM_def] + "((%h. (f(a + h) - f(a))/h) -- 0 --C> D) = \ +\ ((%x. (f(x) - f(a)) / (x - a)) -- a --C> D)"; +by (Step_tac 1); +by (ALLGOALS(dtac spec)); +by (Step_tac 1); +by (Blast_tac 1 THEN Blast_tac 2); +by (ALLGOALS(res_inst_tac [("x","s")] exI)); +by (Step_tac 1); +by (dres_inst_tac [("x","x - a")] spec 1); +by (dres_inst_tac [("x","x + a")] spec 2); +by (auto_tac (claset(), simpset() addsimps complex_add_ac)); +qed "CDERIV_CLIM_iff"; + +Goalw [cderiv_def] "(CDERIV f x :> D) = \ +\ ((%z. (f(z) - f(x)) / (z - x)) -- x --C> D)"; +by (simp_tac (simpset() addsimps [CDERIV_CLIM_iff]) 1); +qed "CDERIV_iff2"; + + +(*------------------------------------------------------------------------------------*) +(* Equivalence of NS and standard defs of differentiation *) +(*------------------------------------------------------------------------------------*) + +(*** first equivalence ***) +Goalw [nscderiv_def,NSCLIM_def] + "(NSCDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --NSC> D)"; +by Auto_tac; +by (dres_inst_tac [("x","xa")] bspec 1); +by (rtac ccontr 3); +by (dres_inst_tac [("x","h")] spec 3); +by (auto_tac (claset(), + simpset() addsimps [mem_cinfmal_iff, starfunC_lambda_cancel])); +qed "NSCDERIV_NSCLIM_iff"; + +(*** 2nd equivalence ***) +Goal "(NSCDERIV f x :> D) = \ +\ ((%z. (f(z) - f(x)) / (z - x)) -- x --NSC> D)"; +by (full_simp_tac (simpset() addsimps + [NSCDERIV_NSCLIM_iff, CDERIV_CLIM_iff, CLIM_NSCLIM_iff RS sym]) 1); +qed "NSCDERIV_NSCLIM_iff2"; + +Goal "(NSCDERIV f x :> D) = \ +\ (ALL xa. xa ~= hcomplex_of_complex x & xa @c= hcomplex_of_complex x --> \ +\ ( *fc* (%z. (f z - f x) / (z - x))) xa @c= hcomplex_of_complex D)"; +by (auto_tac (claset(), simpset() addsimps [NSCDERIV_NSCLIM_iff2, NSCLIM_def])); +qed "NSCDERIV_iff2"; + +Goalw [cderiv_def] "(NSCDERIV f x :> D) = (CDERIV f x :> D)"; +by (simp_tac (simpset() addsimps [NSCDERIV_NSCLIM_iff,CLIM_NSCLIM_iff]) 1); +qed "NSCDERIV_CDERIV_iff"; + +Goalw [nscderiv_def] + "NSCDERIV f x :> D ==> isNSContc f x"; +by (auto_tac (claset(),simpset() addsimps [isNSContc_NSCLIM_iff, + NSCLIM_def,hcomplex_diff_def])); +by (dtac (capprox_minus_iff RS iffD1) 1); +by (dtac (CLAIM "x ~= a ==> x + - a ~= (0::hcomplex)") 1); +by (dres_inst_tac [("x","- hcomplex_of_complex x + xa")] bspec 1); +by (asm_full_simp_tac (simpset() addsimps [hcomplex_add_assoc RS sym]) 2); +by (auto_tac (claset(),simpset() addsimps + [mem_cinfmal_iff RS sym,hcomplex_add_commute])); +by (dres_inst_tac [("c","xa + - hcomplex_of_complex x")] capprox_mult1 1); +by (auto_tac (claset() addIs [CInfinitesimal_subset_CFinite + RS subsetD],simpset() addsimps [hcomplex_mult_assoc])); +by (dres_inst_tac [("x3","D")] (CFinite_hcomplex_of_complex RSN + (2,CInfinitesimal_CFinite_mult) RS (mem_cinfmal_iff RS iffD1)) 1); +by (blast_tac (claset() addIs [capprox_trans,hcomplex_mult_commute RS subst, + (capprox_minus_iff RS iffD2)]) 1); +qed "NSCDERIV_isNSContc"; + +Goal "CDERIV f x :> D ==> isContc f x"; +by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff RS sym, + isNSContc_isContc_iff RS sym,NSCDERIV_isNSContc]) 1); +qed "CDERIV_isContc"; + +(*------------------------------------------------------------------------------------*) +(* Differentiation rules for combinations of functions follow from clear, *) +(* straightforard, algebraic manipulations *) +(*------------------------------------------------------------------------------------*) + +(* use simple constant nslimit theorem *) +Goal "(NSCDERIV (%x. k) x :> 0)"; +by (simp_tac (simpset() addsimps [NSCDERIV_NSCLIM_iff]) 1); +qed "NSCDERIV_const"; +Addsimps [NSCDERIV_const]; + +Goal "(CDERIV (%x. k) x :> 0)"; +by (simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff RS sym]) 1); +qed "CDERIV_const"; +Addsimps [CDERIV_const]; + +Goal "[| NSCDERIV f x :> Da; NSCDERIV g x :> Db |] \ +\ ==> NSCDERIV (%x. f x + g x) x :> Da + Db"; +by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_NSCLIM_iff, + NSCLIM_def]) 1 THEN REPEAT(Step_tac 1)); +by (auto_tac (claset(), + simpset() addsimps [hcomplex_add_divide_distrib,hcomplex_diff_def])); +by (dres_inst_tac [("b","hcomplex_of_complex Da"), + ("d","hcomplex_of_complex Db")] capprox_add 1); +by (auto_tac (claset(), simpset() addsimps hcomplex_add_ac)); +qed "NSCDERIV_add"; + +Goal "[| CDERIV f x :> Da; CDERIV g x :> Db |] \ +\ ==> CDERIV (%x. f x + g x) x :> Da + Db"; +by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_add, + NSCDERIV_CDERIV_iff RS sym]) 1); +qed "CDERIV_add"; + +(*** lemmas for multiplication ***) + +Goal "((a::hcomplex)*b) - (c*d) = (b*(a - c)) + (c*(b - d))"; +by (simp_tac (simpset() addsimps [hcomplex_diff_mult_distrib2]) 1); +val lemma_nscderiv1 = result(); + +Goal "[| (x + y) / z = hcomplex_of_complex D + yb; z ~= 0; \ +\ z : CInfinitesimal; yb : CInfinitesimal |] \ +\ ==> x + y @c= 0"; +by (forw_inst_tac [("c1","z")] (hcomplex_mult_right_cancel RS iffD2) 1 + THEN assume_tac 1); +by (thin_tac "(x + y) / z = hcomplex_of_complex D + yb" 1); +by (auto_tac (claset() addSIs [CInfinitesimal_CFinite_mult2, CFinite_add], + simpset() addsimps [mem_cinfmal_iff RS sym])); +by (etac (CInfinitesimal_subset_CFinite RS subsetD) 1); +val lemma_nscderiv2 = result(); + +Goal "[| NSCDERIV f x :> Da; NSCDERIV g x :> Db |] \ +\ ==> NSCDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))"; +by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_NSCLIM_iff, NSCLIM_def]) 1 + THEN REPEAT(Step_tac 1)); +by (auto_tac (claset(), + simpset() addsimps [starfunC_lambda_cancel, lemma_nscderiv1, + hcomplex_of_complex_zero])); +by (simp_tac (simpset() addsimps [hcomplex_add_divide_distrib]) 1); +by (REPEAT(dtac (bex_CInfinitesimal_iff2 RS iffD2) 1)); +by (auto_tac (claset(), + simpset() delsimps [hcomplex_times_divide1_eq] + addsimps [hcomplex_times_divide1_eq RS sym])); +by (rewtac hcomplex_diff_def); +by (dres_inst_tac [("D","Db")] lemma_nscderiv2 1); +by (dtac (capprox_minus_iff RS iffD2 RS (bex_CInfinitesimal_iff2 RS iffD2)) 4); +by (auto_tac (claset() addSIs [capprox_add_mono1], + simpset() addsimps [hcomplex_add_mult_distrib, hcomplex_add_mult_distrib2, + hcomplex_mult_commute, hcomplex_add_assoc])); +by (res_inst_tac [("w1","hcomplex_of_complex Db * hcomplex_of_complex (f x)")] + (hcomplex_add_commute RS subst) 1); +by (auto_tac (claset() addSIs [CInfinitesimal_add_capprox_self2 RS capprox_sym, + CInfinitesimal_add, CInfinitesimal_mult, + CInfinitesimal_hcomplex_of_complex_mult, + CInfinitesimal_hcomplex_of_complex_mult2], + simpset() addsimps [hcomplex_add_assoc RS sym])); +qed "NSCDERIV_mult"; + +Goal "[| CDERIV f x :> Da; CDERIV g x :> Db |] \ +\ ==> CDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))"; +by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_mult, + NSCDERIV_CDERIV_iff RS sym]) 1); +qed "CDERIV_mult"; + +Goal "NSCDERIV f x :> D ==> NSCDERIV (%x. c * f x) x :> c*D"; +by (asm_full_simp_tac + (simpset() addsimps [complex_times_divide1_eq RS sym, NSCDERIV_NSCLIM_iff, + complex_minus_mult_eq2, complex_add_mult_distrib2 RS sym, + complex_diff_def] + delsimps [complex_times_divide1_eq, complex_minus_mult_eq2 RS sym]) 1); +by (etac (NSCLIM_const RS NSCLIM_mult) 1); +qed "NSCDERIV_cmult"; + +Goal "CDERIV f x :> D ==> CDERIV (%x. c * f x) x :> c*D"; +by (auto_tac (claset(),simpset() addsimps [NSCDERIV_cmult,NSCDERIV_CDERIV_iff + RS sym])); +qed "CDERIV_cmult"; + +Goal "NSCDERIV f x :> D ==> NSCDERIV (%x. -(f x)) x :> -D"; +by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_NSCLIM_iff,complex_diff_def]) 1); +by (res_inst_tac [("t","f x")] (complex_minus_minus RS subst) 1); +by (asm_simp_tac (simpset() addsimps [complex_minus_add_distrib RS sym] + delsimps [complex_minus_add_distrib, complex_minus_minus]) 1); +by (etac NSCLIM_minus 1); +qed "NSCDERIV_minus"; + +Goal "CDERIV f x :> D ==> CDERIV (%x. -(f x)) x :> -D"; +by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_minus,NSCDERIV_CDERIV_iff RS sym]) 1); +qed "CDERIV_minus"; + +Goal "[| NSCDERIV f x :> Da; NSCDERIV g x :> Db |] \ +\ ==> NSCDERIV (%x. f x + -g x) x :> Da + -Db"; +by (blast_tac (claset() addDs [NSCDERIV_add,NSCDERIV_minus]) 1); +qed "NSCDERIV_add_minus"; + +Goal "[| CDERIV f x :> Da; CDERIV g x :> Db |] \ +\ ==> CDERIV (%x. f x + -g x) x :> Da + -Db"; +by (blast_tac (claset() addDs [CDERIV_add,CDERIV_minus]) 1); +qed "CDERIV_add_minus"; + +Goalw [complex_diff_def] + "[| NSCDERIV f x :> Da; NSCDERIV g x :> Db |] \ +\ ==> NSCDERIV (%x. f x - g x) x :> Da - Db"; +by (blast_tac (claset() addIs [NSCDERIV_add_minus]) 1); +qed "NSCDERIV_diff"; + +Goalw [complex_diff_def] + "[| CDERIV f x :> Da; CDERIV g x :> Db |] \ +\ ==> CDERIV (%x. f x - g x) x :> Da - Db"; +by (blast_tac (claset() addIs [CDERIV_add_minus]) 1); +qed "CDERIV_diff"; + + +(*--------------------------------------------------*) +(* Chain rule *) +(*--------------------------------------------------*) + +(* lemmas *) +Goalw [nscderiv_def] + "[| NSCDERIV g x :> D; \ +\ ( *fc* g) (hcomplex_of_complex(x) + xa) = hcomplex_of_complex(g x);\ +\ xa : CInfinitesimal; xa ~= 0 \ +\ |] ==> D = 0"; +by (dtac bspec 1); +by Auto_tac; +qed "NSCDERIV_zero"; + +Goalw [nscderiv_def] + "[| NSCDERIV f x :> D; h: CInfinitesimal; h ~= 0 |] \ +\ ==> ( *fc* f) (hcomplex_of_complex(x) + h) - hcomplex_of_complex(f x) @c= 0"; +by (asm_full_simp_tac (simpset() addsimps [mem_cinfmal_iff RS sym]) 1); +by (rtac CInfinitesimal_ratio 1); +by (rtac capprox_hcomplex_of_complex_CFinite 3); +by Auto_tac; +qed "NSCDERIV_capprox"; + + +(*--------------------------------------------------*) +(* from one version of differentiability *) +(* *) +(* f(x) - f(a) *) +(* --------------- @= Db *) +(* x - a *) +(* -------------------------------------------------*) + +Goal "[| NSCDERIV f (g x) :> Da; \ +\ ( *fc* g) (hcomplex_of_complex(x) + xa) ~= hcomplex_of_complex (g x); \ +\ ( *fc* g) (hcomplex_of_complex(x) + xa) @c= hcomplex_of_complex (g x) \ +\ |] ==> (( *fc* f) (( *fc* g) (hcomplex_of_complex(x) + xa)) \ +\ - hcomplex_of_complex (f (g x))) \ +\ / (( *fc* g) (hcomplex_of_complex(x) + xa) - hcomplex_of_complex (g x)) \ +\ @c= hcomplex_of_complex (Da)"; +by (auto_tac (claset(),simpset() addsimps [NSCDERIV_NSCLIM_iff2, NSCLIM_def])); +qed "NSCDERIVD1"; + +(*--------------------------------------------------*) +(* from other version of differentiability *) +(* *) +(* f(x + h) - f(x) *) +(* ----------------- @= Db *) +(* h *) +(*--------------------------------------------------*) + +Goal "[| NSCDERIV g x :> Db; xa: CInfinitesimal; xa ~= 0 |] \ +\ ==> (( *fc* g) (hcomplex_of_complex(x) + xa) - hcomplex_of_complex(g x)) / xa \ +\ @c= hcomplex_of_complex (Db)"; +by (auto_tac (claset(), + simpset() addsimps [NSCDERIV_NSCLIM_iff, NSCLIM_def, + mem_cinfmal_iff, starfunC_lambda_cancel])); +qed "NSCDERIVD2"; + +Goal "(z::hcomplex) ~= 0 ==> x*y = (x*inverse(z))*(z*y)"; +by Auto_tac; +qed "lemma_complex_chain"; + +(*** chain rule ***) + +Goal "[| NSCDERIV f (g x) :> Da; NSCDERIV g x :> Db |] \ +\ ==> NSCDERIV (f o g) x :> Da * Db"; +by (asm_simp_tac (simpset() addsimps [NSCDERIV_NSCLIM_iff, + NSCLIM_def,mem_cinfmal_iff RS sym]) 1 THEN Step_tac 1); +by (forw_inst_tac [("f","g")] NSCDERIV_capprox 1); +by (auto_tac (claset(), + simpset() addsimps [starfunC_lambda_cancel2, starfunC_o RS sym])); +by (case_tac "( *fc* g) (hcomplex_of_complex(x) + xa) = hcomplex_of_complex (g x)" 1); +by (dres_inst_tac [("g","g")] NSCDERIV_zero 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_divide_def])); +by (res_inst_tac [("z1","( *fc* g) (hcomplex_of_complex(x) + xa) - hcomplex_of_complex (g x)"), + ("y1","inverse xa")] (lemma_complex_chain RS ssubst) 1); +by (Asm_simp_tac 1); +by (rtac capprox_mult_hcomplex_of_complex 1); +by (fold_tac [hcomplex_divide_def]); +by (blast_tac (claset() addIs [NSCDERIVD2]) 2); +by (auto_tac (claset() addSIs [NSCDERIVD1] addIs [capprox_minus_iff RS iffD2], + simpset() addsimps [symmetric hcomplex_diff_def])); +qed "NSCDERIV_chain"; + +(* standard version *) +Goal "[| CDERIV f (g x) :> Da; CDERIV g x :> Db |] \ +\ ==> CDERIV (f o g) x :> Da * Db"; +by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff RS sym, + NSCDERIV_chain]) 1); +qed "CDERIV_chain"; + +Goal "[| CDERIV f (g x) :> Da; CDERIV g x :> Db |] \ +\ ==> CDERIV (%x. f (g x)) x :> Da * Db"; +by (auto_tac (claset() addDs [CDERIV_chain], simpset() addsimps [o_def])); +qed "CDERIV_chain2"; + +(*------------------------------------------------------------------------------------*) +(* Differentiation of natural number powers *) +(*------------------------------------------------------------------------------------*) + +Goal "NSCDERIV (%x. x) x :> 1"; +by (auto_tac (claset(), + simpset() addsimps [NSCDERIV_NSCLIM_iff,NSCLIM_def])); +qed "NSCDERIV_Id"; +Addsimps [NSCDERIV_Id]; + +Goal "CDERIV (%x. x) x :> 1"; +by (simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff RS sym]) 1); +qed "CDERIV_Id"; +Addsimps [CDERIV_Id]; + +bind_thm ("isContc_Id", CDERIV_Id RS CDERIV_isContc); + +(*derivative of linear multiplication*) +Goal "CDERIV (op * c) x :> c"; +by (cut_inst_tac [("c","c"),("x","x")] (CDERIV_Id RS CDERIV_cmult) 1); +by (Asm_full_simp_tac 1); +qed "CDERIV_cmult_Id"; +Addsimps [CDERIV_cmult_Id]; + +Goal "NSCDERIV (op * c) x :> c"; +by (simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff]) 1); +qed "NSCDERIV_cmult_Id"; +Addsimps [NSCDERIV_cmult_Id]; + +Goal "CDERIV (%x. x ^ n) x :> (complex_of_real (real n)) * (x ^ (n - 1))"; +by (induct_tac "n" 1); +by (dtac (CDERIV_Id RS CDERIV_mult) 2); +by (auto_tac (claset(), + simpset() addsimps [complex_of_real_add RS sym, + complex_add_mult_distrib,real_of_nat_Suc] delsimps [complex_of_real_add])); +by (case_tac "n" 1); +by (auto_tac (claset(), + simpset() addsimps [complex_mult_assoc, complex_add_commute])); +by (auto_tac (claset(),simpset() addsimps [complex_mult_commute])); +qed "CDERIV_pow"; +Addsimps [CDERIV_pow,simplify (simpset()) CDERIV_pow]; + +(* NS version *) +Goal "NSCDERIV (%x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - 1))"; +by (simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff]) 1); +qed "NSCDERIV_pow"; + +Goal "\\ CDERIV f x :> D; D = E \\ \\ CDERIV f x :> E"; +by Auto_tac; +qed "lemma_CDERIV_subst"; + +(*used once, in NSCDERIV_inverse*) +Goal "[| h: CInfinitesimal; x ~= 0 |] ==> hcomplex_of_complex x + h ~= 0"; +by Auto_tac; +qed "CInfinitesimal_add_not_zero"; + +(*** +Goal "[|(x::hcomplex) ~= 0; y ~= 0 |] \ +\ ==> inverse(x) + inverse(y) = (x + y)*inverse(x*y)"; +by (asm_full_simp_tac (simpset() addsimps [hcomplex_inverse_distrib, + hcomplex_add_mult_distrib,hcomplex_mult_assoc RS sym]) 1); +qed "hcomplex_inverse_add"; +***) + +(*Can't get rid of x ~= 0 because it isn't continuous at zero*) + +Goalw [nscderiv_def] + "x ~= 0 ==> NSCDERIV (%x. inverse(x)) x :> (- (inverse x ^ 2))"; +by (rtac ballI 1 THEN Asm_full_simp_tac 1 THEN Step_tac 1); +by (forward_tac [CInfinitesimal_add_not_zero] 1); +by (asm_full_simp_tac (simpset() addsimps [hcomplex_add_commute,two_eq_Suc_Suc]) 2); +by (auto_tac (claset(), + simpset() addsimps [starfunC_inverse_inverse,hcomplex_diff_def] + delsimps [hcomplex_minus_mult_eq1 RS sym, + hcomplex_minus_mult_eq2 RS sym])); +by (asm_simp_tac + (simpset() addsimps [hcomplex_inverse_add, + hcomplex_inverse_distrib RS sym, hcomplex_minus_inverse RS sym] + @ hcomplex_add_ac @ hcomplex_mult_ac + delsimps [hcomplex_minus_mult_eq1 RS sym, + hcomplex_minus_mult_eq2 RS sym] ) 1); +by (asm_simp_tac (simpset() addsimps [hcomplex_mult_assoc RS sym, + hcomplex_add_mult_distrib2] + delsimps [hcomplex_minus_mult_eq1 RS sym, + hcomplex_minus_mult_eq2 RS sym]) 1); +by (res_inst_tac [("y"," inverse(- hcomplex_of_complex x * hcomplex_of_complex x)")] + capprox_trans 1); +by (rtac inverse_add_CInfinitesimal_capprox2 1); +by (auto_tac (claset() addSDs [hcomplex_of_complex_CFinite_diff_CInfinitesimal] addIs [CFinite_mult], + simpset() addsimps [hcomplex_minus_inverse RS sym])); +by (rtac CInfinitesimal_CFinite_mult2 1); +by Auto_tac; +qed "NSCDERIV_inverse"; + +Goal "x ~= 0 ==> CDERIV (%x. inverse(x)) x :> (-(inverse x ^ 2))"; +by (asm_simp_tac (simpset() addsimps [NSCDERIV_inverse, + NSCDERIV_CDERIV_iff RS sym] delsimps [complexpow_Suc]) 1); +qed "CDERIV_inverse"; + + +(*------------------------------------------------------------------------------------*) +(* Derivative of inverse *) +(*------------------------------------------------------------------------------------*) + +Goal "[| CDERIV f x :> d; f(x) ~= 0 |] \ +\ ==> CDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))"; +by (rtac (complex_mult_commute RS subst) 1); +by (asm_simp_tac (simpset() addsimps [complex_minus_mult_eq1, + complexpow_inverse] delsimps [complexpow_Suc, + complex_minus_mult_eq1 RS sym]) 1); +by (fold_goals_tac [o_def]); +by (blast_tac (claset() addSIs [CDERIV_chain,CDERIV_inverse]) 1); +qed "CDERIV_inverse_fun"; + +Goal "[| NSCDERIV f x :> d; f(x) ~= 0 |] \ +\ ==> NSCDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))"; +by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff, + CDERIV_inverse_fun] delsimps [complexpow_Suc]) 1); +qed "NSCDERIV_inverse_fun"; + +(*------------------------------------------------------------------------------------*) +(* Derivative of quotient *) +(*------------------------------------------------------------------------------------*) + + +Goal "x ~= (0::complex) \\ (x * inverse(x) ^ 2) = inverse x"; +by (auto_tac (claset(),simpset() addsimps [two_eq_Suc_Suc])); +qed "lemma_complex_mult_inverse_squared"; +Addsimps [lemma_complex_mult_inverse_squared]; + +Goalw [complex_diff_def] + "[| CDERIV f x :> d; CDERIV g x :> e; g(x) ~= 0 |] \ +\ ==> CDERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ 2)"; +by (dres_inst_tac [("f","g")] CDERIV_inverse_fun 1); +by (dtac CDERIV_mult 2); +by (REPEAT(assume_tac 1)); +by (asm_full_simp_tac + (simpset() addsimps [complex_divide_def, complex_add_mult_distrib2, + complexpow_inverse,complex_minus_mult_eq1] @ complex_mult_ac + delsimps [complexpow_Suc, complex_minus_mult_eq1 RS sym, + complex_minus_mult_eq2 RS sym]) 1); +qed "CDERIV_quotient"; + +Goal "[| NSCDERIV f x :> d; NSCDERIV g x :> e; g(x) ~= 0 |] \ +\ ==> NSCDERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ 2)"; +by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff, + CDERIV_quotient] delsimps [complexpow_Suc]) 1); +qed "NSCDERIV_quotient"; + + +(*------------------------------------------------------------------------------------*) +(* Caratheodory formulation of derivative at a point: standard proof *) +(*------------------------------------------------------------------------------------*) + + +Goalw [CLIM_def] + "[| ALL x. x ~= a --> (f x = g x) |] \ +\ ==> (f -- a --C> l) = (g -- a --C> l)"; +by (auto_tac (claset(), simpset() addsimps [complex_add_minus_iff])); +qed "CLIM_equal"; + +Goal "[| (%x. f(x) + -g(x)) -- a --C> 0; \ +\ g -- a --C> l |] \ +\ ==> f -- a --C> l"; +by (dtac CLIM_add 1 THEN assume_tac 1); +by (auto_tac (claset(), simpset() addsimps [complex_add_assoc])); +qed "CLIM_trans"; + +Goal "(CDERIV f x :> l) = \ +\ (EX g. (ALL z. f z - f x = g z * (z - x)) & isContc g x & g x = l)"; +by (Step_tac 1); +by (res_inst_tac + [("x","%z. if z = x then l else (f(z) - f(x)) / (z - x)")] exI 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult_assoc, + CLAIM "z ~= x ==> z - x ~= (0::complex)"])); +by (auto_tac (claset(),simpset() addsimps [isContc_iff,CDERIV_iff])); +by (ALLGOALS(rtac (CLIM_equal RS iffD1))); +by Auto_tac; +qed "CARAT_CDERIV"; + +Goal "NSCDERIV f x :> l ==> \ +\ EX g. (ALL z. f z - f x = g z * (z - x)) & isNSContc g x & g x = l"; +by (auto_tac (claset(),simpset() addsimps [NSCDERIV_CDERIV_iff, + isNSContc_isContc_iff,CARAT_CDERIV])); +qed "CARAT_NSCDERIV"; + +(* How about a NS proof? *) +Goal "(ALL z. f z - f x = g z * (z - x)) & isNSContc g x & g x = l \ +\ ==> NSCDERIV f x :> l"; +by (auto_tac (claset(), + simpset() delsimprocs complex_cancel_factor + addsimps [NSCDERIV_iff2])); +by (asm_full_simp_tac (simpset() addsimps [isNSContc_def]) 1); +qed "CARAT_CDERIVD"; + diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/CLim.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/CLim.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,82 @@ +(* Title : CLim.thy + Author : Jacques D. Fleuriot + Copyright : 2001 University of Edinburgh + Description : A first theory of limits, continuity and + differentiation for complex functions +*) + +CLim = CSeries + + +constdefs + + CLIM :: [complex=>complex,complex,complex] => bool + ("((_)/ -- (_)/ --C> (_))" [60, 0, 60] 60) + "f -- a --C> L == + ALL r. 0 < r --> + (EX s. 0 < s & (ALL x. (x ~= a & (cmod(x - a) < s) + --> cmod(f x - L) < r)))" + + NSCLIM :: [complex=>complex,complex,complex] => bool + ("((_)/ -- (_)/ --NSC> (_))" [60, 0, 60] 60) + "f -- a --NSC> L == (ALL x. (x ~= hcomplex_of_complex a & + x @c= hcomplex_of_complex a + --> ( *fc* f) x @c= hcomplex_of_complex L))" + + (* f: C --> R *) + CRLIM :: [complex=>real,complex,real] => bool + ("((_)/ -- (_)/ --CR> (_))" [60, 0, 60] 60) + "f -- a --CR> L == + ALL r. 0 < r --> + (EX s. 0 < s & (ALL x. (x ~= a & (cmod(x - a) < s) + --> abs(f x - L) < r)))" + + NSCRLIM :: [complex=>real,complex,real] => bool + ("((_)/ -- (_)/ --NSCR> (_))" [60, 0, 60] 60) + "f -- a --NSCR> L == (ALL x. (x ~= hcomplex_of_complex a & + x @c= hcomplex_of_complex a + --> ( *fcR* f) x @= hypreal_of_real L))" + + + isContc :: [complex=>complex,complex] => bool + "isContc f a == (f -- a --C> (f a))" + + (* NS definition dispenses with limit notions *) + isNSContc :: [complex=>complex,complex] => bool + "isNSContc f a == (ALL y. y @c= hcomplex_of_complex a --> + ( *fc* f) y @c= hcomplex_of_complex (f a))" + + isContCR :: [complex=>real,complex] => bool + "isContCR f a == (f -- a --CR> (f a))" + + (* NS definition dispenses with limit notions *) + isNSContCR :: [complex=>real,complex] => bool + "isNSContCR f a == (ALL y. y @c= hcomplex_of_complex a --> + ( *fcR* f) y @= hypreal_of_real (f a))" + + (* differentiation: D is derivative of function f at x *) + cderiv:: [complex=>complex,complex,complex] => bool + ("(CDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60) + "CDERIV f x :> D == ((%h. (f(x + h) - f(x))/h) -- 0 --C> D)" + + nscderiv :: [complex=>complex,complex,complex] => bool + ("(NSCDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60) + "NSCDERIV f x :> D == (ALL h: CInfinitesimal - {0}. + (( *fc* f)(hcomplex_of_complex x + h) + - hcomplex_of_complex (f x))/h @c= hcomplex_of_complex D)" + + cdifferentiable :: [complex=>complex,complex] => bool (infixl 60) + "f cdifferentiable x == (EX D. CDERIV f x :> D)" + + NSCdifferentiable :: [complex=>complex,complex] => bool (infixl 60) + "f NSCdifferentiable x == (EX D. NSCDERIV f x :> D)" + + + isUContc :: (complex=>complex) => bool + "isUContc f == (ALL r. 0 < r --> + (EX s. 0 < s & (ALL x y. cmod(x - y) < s + --> cmod(f x - f y) < r)))" + + isNSUContc :: (complex=>complex) => bool + "isNSUContc f == (ALL x y. x @c= y --> ( *fc* f) x @c= ( *fc* f) y)" + +end diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/CSeries.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/CSeries.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,210 @@ +(* Title : CSeries.ML + Author : Jacques D. Fleuriot + Copyright : 2002 University of Edinburgh + Description : Finite summation and infinite series for complex numbers +*) + + +Goal "sumc (Suc n) n f = 0"; +by (induct_tac "n" 1); +by (Auto_tac); +qed "sumc_Suc_zero"; +Addsimps [sumc_Suc_zero]; + +Goal "sumc m m f = 0"; +by (induct_tac "m" 1); +by (Auto_tac); +qed "sumc_eq_bounds"; +Addsimps [sumc_eq_bounds]; + +Goal "sumc m (Suc m) f = f(m)"; +by (Auto_tac); +qed "sumc_Suc_eq"; +Addsimps [sumc_Suc_eq]; + +Goal "sumc (m+k) k f = 0"; +by (induct_tac "k" 1); +by (Auto_tac); +qed "sumc_add_lbound_zero"; +Addsimps [sumc_add_lbound_zero]; + +Goal "sumc m n f + sumc m n g = sumc m n (%n. f n + g n)"; +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps real_add_ac)); +qed "sumc_add"; + +Goal "r * sumc m n f = sumc m n (%n. r * f n)"; +by (induct_tac "n" 1); +by (Auto_tac); +by (auto_tac (claset(),simpset() addsimps + [complex_add_mult_distrib2])); +qed "sumc_mult"; + +Goal "n < p --> sumc 0 n f + sumc n p f = sumc 0 p f"; +by (induct_tac "p" 1); +by (auto_tac (claset() addSDs [CLAIM "n < Suc na ==> n <= na", + leI] addDs [le_anti_sym], simpset())); +qed_spec_mp "sumc_split_add"; + +Goal "n < p ==> sumc 0 p f + \ +\ - sumc 0 n f = sumc n p f"; +by (dres_inst_tac [("f1","f")] (sumc_split_add RS sym) 1); +by (asm_simp_tac (simpset() addsimps complex_add_ac) 1); +qed "sumc_split_add_minus"; + +Goal "cmod(sumc m n f) <= sumr m n (%i. cmod(f i))"; +by (induct_tac "n" 1); +by (auto_tac (claset() addIs [complex_mod_triangle_ineq RS order_trans], + simpset())); +qed "sumc_cmod"; + +Goal "!!f g. (ALL r. m <= r & r < n --> f r = g r) \ +\ --> sumc m n f = sumc m n g"; +by (induct_tac "n" 1); +by (Auto_tac); +qed_spec_mp "sumc_fun_eq"; + +Goal "sumc 0 n (%i. r) = complex_of_real (real n) * r"; +by (induct_tac "n" 1); +by (auto_tac (claset(), + simpset() addsimps [complex_add_mult_distrib, + complex_of_real_add RS sym, + real_of_nat_Suc])); +qed "sumc_const"; +Addsimps [sumc_const]; + +Goal "sumc 0 n f + -(complex_of_real(real n) * r) = sumc 0 n (%i. f i + -r)"; +by (full_simp_tac (simpset() addsimps [sumc_add RS sym]) 1); +qed "sumc_add_mult_const"; + +Goalw [complex_diff_def] + "sumc 0 n f - (complex_of_real(real n)*r) = sumc 0 n (%i. f i - r)"; +by (full_simp_tac (simpset() addsimps [sumc_add_mult_const]) 1); +qed "sumc_diff_mult_const"; + +Goal "n < m --> sumc m n f = 0"; +by (induct_tac "n" 1); +by Auto_tac; +qed_spec_mp "sumc_less_bounds_zero"; +Addsimps [sumc_less_bounds_zero]; + +Goal "sumc m n (%i. - f i) = - sumc m n f"; +by (induct_tac "n" 1); +by Auto_tac; +qed "sumc_minus"; + +Goal "sumc (m+k) (n+k) f = sumc m n (%i. f(i + k))"; +by (induct_tac "n" 1); +by (Auto_tac); +qed "sumc_shift_bounds"; + +Goal "sumc 0 (2*n) (%i. (-1) ^ Suc i) = 0"; +by (induct_tac "n" 1); +by (Auto_tac); +qed "sumc_minus_one_complexpow_zero"; +Addsimps [sumc_minus_one_complexpow_zero]; + +Goal "(ALL n. m <= Suc n --> f n = r) & m <= na \ +\ --> sumc m na f = (complex_of_real(real (na - m)) * r)"; +by (induct_tac "na" 1); +by (auto_tac (claset(),simpset() addsimps [real_add_mult_distrib, Suc_diff_n, + real_of_nat_Suc,complex_of_real_add RS sym, + complex_add_mult_distrib])); +qed_spec_mp "sumc_interval_const"; + +Goal "(ALL n. m <= n --> f n = r) & m <= na \ +\ --> sumc m na f = (complex_of_real(real (na - m)) * r)"; +by (induct_tac "na" 1); +by (auto_tac (claset(),simpset() addsimps [real_add_mult_distrib, Suc_diff_n, + real_of_nat_Suc,complex_of_real_add RS sym, + complex_add_mult_distrib])); +qed_spec_mp "sumc_interval_const2"; + +(*** +Goal "(ALL n. m <= n --> 0 <= cmod(f n)) & m < k --> cmod(sumc 0 m f) <= cmod(sumc 0 k f)"; +by (induct_tac "k" 1); +by (Step_tac 1); +by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [less_Suc_eq_le]))); +by (ALLGOALS(dres_inst_tac [("x","n")] spec)); +by (Step_tac 1); +by (dtac le_imp_less_or_eq 1 THEN Step_tac 1); +by (dtac real_add_le_mono 2); +by (dres_inst_tac [("i","sumr 0 m f")] (order_refl RS real_add_le_mono) 1); +by (Auto_tac); +qed_spec_mp "sumc_le"; + +Goal "!!f g. (ALL r. m <= r & r < n --> f r <= g r) \ +\ --> sumc m n f <= sumc m n g"; +by (induct_tac "n" 1); +by (auto_tac (claset() addIs [real_add_le_mono], + simpset() addsimps [le_def])); +qed_spec_mp "sumc_le2"; + +Goal "(ALL n. 0 <= f n) --> 0 <= sumc m n f"; +by (induct_tac "n" 1); +by Auto_tac; +by (dres_inst_tac [("x","n")] spec 1); +by (arith_tac 1); +qed_spec_mp "sumc_ge_zero"; + +Goal "(ALL n. m <= n --> 0 <= f n) --> 0 <= sumc m n f"; +by (induct_tac "n" 1); +by Auto_tac; +by (dres_inst_tac [("x","n")] spec 1); +by (arith_tac 1); +qed_spec_mp "sumc_ge_zero2"; +***) + +Goal "0 <= sumr m n (%n. cmod (f n))"; +by (induct_tac "n" 1); +by Auto_tac; +by (res_inst_tac [("j","0")] real_le_trans 1); +by Auto_tac; +qed "sumr_cmod_ge_zero"; +Addsimps [sumr_cmod_ge_zero]; +AddSIs [sumr_cmod_ge_zero]; + +Goal "abs (sumr m n (%n. cmod (f n))) = (sumr m n (%n. cmod (f n)))"; +by (rtac (abs_eqI1 RS ssubst) 1 THEN Auto_tac); +qed "rabs_sumc_cmod_cancel"; +Addsimps [rabs_sumc_cmod_cancel]; + +Goal "ALL n. N <= n --> f n = 0 \ +\ ==> ALL m n. N <= m --> sumc m n f = 0"; +by (Step_tac 1); +by (induct_tac "n" 1); +by (Auto_tac); +qed "sumc_zero"; + +Goal "ALL n. N <= n --> f (Suc n) = 0 \ +\ ==> ALL m n. Suc N <= m --> sumc m n f = 0"; +by (rtac sumc_zero 1 THEN Step_tac 1); +by (dres_inst_tac [("x","n - 1")] spec 1); +by Auto_tac; +by (arith_tac 1); +qed "fun_zero_sumc_zero"; + +Goal "sumc 1 n (%n. f(n) * 0 ^ n) = 0"; +by (induct_tac "n" 1); +by (case_tac "n" 2); +by Auto_tac; +qed "sumc_one_lb_complexpow_zero"; +Addsimps [sumc_one_lb_complexpow_zero]; + +Goalw [complex_diff_def] "sumc m n f - sumc m n g = sumc m n (%n. f n - g n)"; +by (simp_tac (simpset() addsimps [sumc_add RS sym,sumc_minus]) 1); +qed "sumc_diff"; + +Goal "(ALL p. (m <= p & p < m + n --> (f p = g p))) --> sumc m n f = sumc m n g"; +by (induct_tac "n" 1); +by (Auto_tac); +qed_spec_mp "sumc_subst"; + +Goal "sumc 0 n (%m. sumc (m * k) (m*k + k) f) = sumc 0 (n * k) f"; +by (subgoal_tac "k = 0 | 0 < k" 1); +by (Auto_tac); +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps [sumc_split_add,add_commute])); +qed "sumc_group"; +Addsimps [sumc_group]; + diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/CSeries.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/CSeries.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,30 @@ +(* Title : CSeries.thy + Author : Jacques D. Fleuriot + Copyright : 2002 University of Edinburgh + Description : Finite summation and infinite series for complex numbers +*) + +CSeries = CStar + + +consts sumc :: "[nat,nat,(nat=>complex)] => complex" +primrec + sumc_0 "sumc m 0 f = 0" + sumc_Suc "sumc m (Suc n) f = (if n < m then 0 else sumc m n f + f(n))" + +(* +constdefs + + needs convergence of complex sequences + + csums :: [nat=>complex,complex] => bool (infixr 80) + "f sums s == (%n. sumr 0 n f) ----C> s" + + csummable :: (nat=>complex) => bool + "csummable f == (EX s. f csums s)" + + csuminf :: (nat=>complex) => complex + "csuminf f == (@s. f csums s)" +*) + +end + diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/CStar.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/CStar.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,630 @@ +(* Title : CStar.ML + Author : Jacques D. Fleuriot + Copyright : 2001 University of Edinburgh + Description : defining *-transforms in NSA which extends sets of complex numbers, + and complex functions +*) + + + +(*-----------------------------------------------------------------------------------*) +(* Properties of the *-transform applied to sets of reals *) +(* ----------------------------------------------------------------------------------*) + +Goalw [starsetC_def] "*sc*(UNIV::complex set) = (UNIV::hcomplex set)"; +by (Auto_tac); +qed "STARC_complex_set"; +Addsimps [STARC_complex_set]; + +Goalw [starsetC_def] "*sc* {} = {}"; +by (Auto_tac); +qed "STARC_empty_set"; +Addsimps [STARC_empty_set]; + +Goalw [starsetC_def] "*sc* (A Un B) = *sc* A Un *sc* B"; +by (Auto_tac); +by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2)); +by (dtac bspec 1 THEN assume_tac 1); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (Auto_tac); +by (Ultra_tac 1); +qed "STARC_Un"; + +Goalw [starsetC_n_def] + "*scn* (%n. (A n) Un (B n)) = *scn* A Un *scn* B"; +by Auto_tac; +by (dres_inst_tac [("x","Xa")] bspec 1); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 2); +by (auto_tac (claset() addSDs [bspec], simpset())); +by (TRYALL(Ultra_tac)); +qed "starsetC_n_Un"; + +Goalw [InternalCSets_def] + "[| X : InternalCSets; Y : InternalCSets |] \ +\ ==> (X Un Y) : InternalCSets"; +by (auto_tac (claset(), + simpset() addsimps [starsetC_n_Un RS sym])); +qed "InternalCSets_Un"; + +Goalw [starsetC_def] "*sc* (A Int B) = *sc* A Int *sc* B"; +by (Auto_tac); +by (blast_tac (claset() addIs [FreeUltrafilterNat_Int, + FreeUltrafilterNat_subset]) 3); +by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1)); +qed "STARC_Int"; + +Goalw [starsetC_n_def] + "*scn* (%n. (A n) Int (B n)) = *scn* A Int *scn* B"; +by (Auto_tac); +by (auto_tac (claset() addSDs [bspec],simpset())); +by (TRYALL(Ultra_tac)); +qed "starsetC_n_Int"; + +Goalw [InternalCSets_def] + "[| X : InternalCSets; Y : InternalCSets |] \ +\ ==> (X Int Y) : InternalCSets"; +by (auto_tac (claset(), + simpset() addsimps [starsetC_n_Int RS sym])); +qed "InternalCSets_Int"; + +Goalw [starsetC_def] "*sc* -A = -( *sc* A)"; +by (Auto_tac); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 2); +by (REPEAT(Step_tac 1) THEN Auto_tac); +by (ALLGOALS(Ultra_tac)); +qed "STARC_Compl"; + +Goalw [starsetC_n_def] "*scn* ((%n. - A n)) = -( *scn* A)"; +by (Auto_tac); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 2); +by (REPEAT(Step_tac 1) THEN Auto_tac); +by (TRYALL(Ultra_tac)); +qed "starsetC_n_Compl"; + +Goalw [InternalCSets_def] + "X :InternalCSets ==> -X : InternalCSets"; +by (auto_tac (claset(), + simpset() addsimps [starsetC_n_Compl RS sym])); +qed "InternalCSets_Compl"; + +Goal "x ~: *sc* F ==> x : *sc* (- F)"; +by (auto_tac (claset(),simpset() addsimps [STARC_Compl])); +qed "STARC_mem_Compl"; + +Goal "*sc* (A - B) = *sc* A - *sc* B"; +by (auto_tac (claset(),simpset() addsimps + [set_diff_iff2,STARC_Int,STARC_Compl])); +qed "STARC_diff"; + +Goalw [starsetC_n_def] + "*scn* (%n. (A n) - (B n)) = *scn* A - *scn* B"; +by (Auto_tac); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 2); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 3); +by (auto_tac (claset() addSDs [bspec], simpset())); +by (TRYALL(Ultra_tac)); +qed "starsetC_n_diff"; + +Goalw [InternalCSets_def] + "[| X : InternalCSets; Y : InternalCSets |] \ +\ ==> (X - Y) : InternalCSets"; +by (auto_tac (claset(), simpset() addsimps [starsetC_n_diff RS sym])); +qed "InternalCSets_diff"; + +Goalw [starsetC_def] "A <= B ==> *sc* A <= *sc* B"; +by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1)); +qed "STARC_subset"; + +Goalw [starsetC_def,hcomplex_of_complex_def] + "a : A ==> hcomplex_of_complex a : *sc* A"; +by (auto_tac (claset() addIs [FreeUltrafilterNat_subset], + simpset())); +qed "STARC_mem"; + +Goalw [starsetC_def] "hcomplex_of_complex ` A <= *sc* A"; +by (auto_tac (claset(), simpset() addsimps [hcomplex_of_complex_def])); +by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1); +qed "STARC_hcomplex_of_complex_image_subset"; + +Goal "SComplex <= *sc* (UNIV:: complex set)"; +by Auto_tac; +qed "STARC_SComplex_subset"; + +Goalw [starsetC_def] + "*sc* X Int SComplex = hcomplex_of_complex ` X"; +by (auto_tac (claset(), + simpset() addsimps + [hcomplex_of_complex_def,SComplex_def])); +by (fold_tac [hcomplex_of_complex_def]); +by (rtac imageI 1 THEN rtac ccontr 1); +by (dtac bspec 1); +by (rtac lemma_hcomplexrel_refl 1); +by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2); +by (Auto_tac); +qed "STARC_hcomplex_of_complex_Int"; + +Goal "x ~: hcomplex_of_complex ` A ==> ALL y: A. x ~= hcomplex_of_complex y"; +by (Auto_tac); +qed "lemma_not_hcomplexA"; + +Goalw [starsetC_n_def,starsetC_def] "*sc* X = *scn* (%n. X)"; +by Auto_tac; +qed "starsetC_starsetC_n_eq"; + +Goalw [InternalCSets_def] "( *sc* X) : InternalCSets"; +by (auto_tac (claset(), + simpset() addsimps [starsetC_starsetC_n_eq])); +qed "InternalCSets_starsetC_n"; +Addsimps [InternalCSets_starsetC_n]; + +Goal "X : InternalCSets ==> UNIV - X : InternalCSets"; +by (auto_tac (claset() addIs [InternalCSets_Compl], simpset())); +qed "InternalCSets_UNIV_diff"; + +(*-----------------------------------------------------------------------------------*) +(* Nonstandard extension of a set (defined using a constant sequence) as a special *) +(* case of an internal set *) +(*-----------------------------------------------------------------------------------*) + +Goalw [starsetC_n_def,starsetC_def] + "ALL n. (As n = A) ==> *scn* As = *sc* A"; +by (Auto_tac); +qed "starsetC_n_starsetC"; + +(*-----------------------------------------------------------------------------------*) +(* Theorems about nonstandard extensions of functions *) +(*-----------------------------------------------------------------------------------*) + +Goalw [starfunC_n_def,starfunC_def] + "ALL n. (F n = f) ==> *fcn* F = *fc* f"; +by (Auto_tac); +qed "starfunC_n_starfunC"; + +Goalw [starfunRC_n_def,starfunRC_def] + "ALL n. (F n = f) ==> *fRcn* F = *fRc* f"; +by (Auto_tac); +qed "starfunRC_n_starfunRC"; + +Goalw [starfunCR_n_def,starfunCR_def] + "ALL n. (F n = f) ==> *fcRn* F = *fcR* f"; +by (Auto_tac); +qed "starfunCR_n_starfunCR"; + +Goalw [congruent_def] + "congruent hcomplexrel (%X. hcomplexrel``{%n. f (X n)})"; +by (safe_tac (claset())); +by (ALLGOALS(Ultra_tac)); +qed "starfunC_congruent"; + +(* f::complex => complex *) +Goalw [starfunC_def] + "( *fc* f) (Abs_hcomplex(hcomplexrel``{%n. X n})) = \ +\ Abs_hcomplex(hcomplexrel `` {%n. f (X n)})"; +by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1); +by Auto_tac; +by (Ultra_tac 1); +qed "starfunC"; + +Goalw [starfunRC_def] + "( *fRc* f) (Abs_hypreal(hyprel``{%n. X n})) = \ +\ Abs_hcomplex(hcomplexrel `` {%n. f (X n)})"; +by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1); +by Auto_tac; +by (Ultra_tac 1); +qed "starfunRC"; + +Goalw [starfunCR_def] + "( *fcR* f) (Abs_hcomplex(hcomplexrel``{%n. X n})) = \ +\ Abs_hypreal(hyprel `` {%n. f (X n)})"; +by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1); +by Auto_tac; +by (Ultra_tac 1); +qed "starfunCR"; + +(** multiplication: ( *f ) x ( *g ) = *(f x g) **) + +Goal "( *fc* f) z * ( *fc* g) z = ( *fc* (%x. f x * g x)) z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), simpset() addsimps [starfunC,hcomplex_mult])); +qed "starfunC_mult"; +Addsimps [starfunC_mult RS sym]; + +Goal "( *fRc* f) z * ( *fRc* g) z = ( *fRc* (%x. f x * g x)) z"; +by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); +by (auto_tac (claset(), simpset() addsimps [starfunRC,hcomplex_mult])); +qed "starfunRC_mult"; +Addsimps [starfunRC_mult RS sym]; + +Goal "( *fcR* f) z * ( *fcR* g) z = ( *fcR* (%x. f x * g x)) z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), simpset() addsimps [starfunCR,hypreal_mult])); +qed "starfunCR_mult"; +Addsimps [starfunCR_mult RS sym]; + +(** addition: ( *f ) + ( *g ) = *(f + g) **) + +Goal "( *fc* f) z + ( *fc* g) z = ( *fc* (%x. f x + g x)) z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), simpset() addsimps [starfunC,hcomplex_add])); +qed "starfunC_add"; +Addsimps [starfunC_add RS sym]; + +Goal "( *fRc* f) z + ( *fRc* g) z = ( *fRc* (%x. f x + g x)) z"; +by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); +by (auto_tac (claset(), simpset() addsimps [starfunRC,hcomplex_add])); +qed "starfunRC_add"; +Addsimps [starfunRC_add RS sym]; + +Goal "( *fcR* f) z + ( *fcR* g) z = ( *fcR* (%x. f x + g x)) z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), simpset() addsimps [starfunCR,hypreal_add])); +qed "starfunCR_add"; +Addsimps [starfunCR_add RS sym]; + +(** uminus **) +Goal "( *fc* (%x. - f x)) x = - ( *fc* f) x"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunC, hcomplex_minus])); +qed "starfunC_minus"; +Addsimps [starfunC_minus]; + +Goal "( *fRc* (%x. - f x)) x = - ( *fRc* f) x"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [starfunRC, hcomplex_minus])); +qed "starfunRC_minus"; +Addsimps [starfunRC_minus]; + +Goal "( *fcR* (%x. - f x)) x = - ( *fcR* f) x"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunCR, hypreal_minus])); +qed "starfunCR_minus"; +Addsimps [starfunCR_minus]; + +(** addition: ( *f ) - ( *g ) = *(f - g) **) + +Goalw [hcomplex_diff_def,complex_diff_def] + "( *fc* f) xa - ( *fc* g) xa = ( *fc* (%x. f x - g x)) xa"; +by (auto_tac (claset(),simpset() addsimps [starfunC_minus,starfunC_add RS sym])); +qed "starfunC_diff"; +Addsimps [starfunC_diff RS sym]; + +Goalw [hcomplex_diff_def,complex_diff_def] + "( *fRc* f) xa - ( *fRc* g) xa = ( *fRc* (%x. f x - g x)) xa"; +by (auto_tac (claset(),simpset() addsimps [starfunRC_minus,starfunRC_add RS sym])); +qed "starfunRC_diff"; +Addsimps [starfunRC_diff RS sym]; + +Goalw [hypreal_diff_def,real_diff_def] + "( *fcR* f) xa - ( *fcR* g) xa = ( *fcR* (%x. f x - g x)) xa"; +by (auto_tac (claset(),simpset() addsimps [starfunCR_minus,starfunCR_add RS sym])); +qed "starfunCR_diff"; +Addsimps [starfunCR_diff RS sym]; + +(** composition: ( *f ) o ( *g ) = *(f o g) **) + +Goal "(%x. ( *fc* f) (( *fc* g) x)) = *fc* (%x. f (g x))"; +by (rtac ext 1); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunC])); +qed "starfunC_o2"; + +Goalw [o_def] "( *fc* f) o ( *fc* g) = ( *fc* (f o g))"; +by (simp_tac (simpset() addsimps [starfunC_o2]) 1); +qed "starfunC_o"; + +Goal "(%x. ( *fc* f) (( *fRc* g) x)) = *fRc* (%x. f (g x))"; +by (rtac ext 1); +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [starfunRC,starfunC])); +qed "starfunC_starfunRC_o2"; + +Goal "(%x. ( *f* f) (( *fcR* g) x)) = *fcR* (%x. f (g x))"; +by (rtac ext 1); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunCR,starfun])); +qed "starfun_starfunCR_o2"; + +Goalw [o_def] "( *fc* f) o ( *fRc* g) = ( *fRc* (f o g))"; +by (simp_tac (simpset() addsimps [starfunC_starfunRC_o2]) 1); +qed "starfunC_starfunRC_o"; + +Goalw [o_def] "( *f* f) o ( *fcR* g) = ( *fcR* (f o g))"; +by (simp_tac (simpset() addsimps [starfun_starfunCR_o2]) 1); +qed "starfun_starfunCR_o"; + +Goal "( *fc* (%x. k)) z = hcomplex_of_complex k"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), simpset() addsimps [starfunC, hcomplex_of_complex_def])); +qed "starfunC_const_fun"; +Addsimps [starfunC_const_fun]; + +Goal "( *fRc* (%x. k)) z = hcomplex_of_complex k"; +by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); +by (auto_tac (claset(), simpset() addsimps [starfunRC, hcomplex_of_complex_def])); +qed "starfunRC_const_fun"; +Addsimps [starfunRC_const_fun]; + +Goal "( *fcR* (%x. k)) z = hypreal_of_real k"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), simpset() addsimps [starfunCR, hypreal_of_real_def])); +qed "starfunCR_const_fun"; +Addsimps [starfunCR_const_fun]; + +Goal "inverse (( *fc* f) x) = ( *fc* (%x. inverse (f x))) x"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), simpset() addsimps [starfunC, hcomplex_inverse])); +qed "starfunC_inverse"; +Addsimps [starfunC_inverse RS sym]; + +Goal "inverse (( *fRc* f) x) = ( *fRc* (%x. inverse (f x))) x"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (auto_tac (claset(), simpset() addsimps [starfunRC, hcomplex_inverse])); +qed "starfunRC_inverse"; +Addsimps [starfunRC_inverse RS sym]; + +Goal "inverse (( *fcR* f) x) = ( *fcR* (%x. inverse (f x))) x"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), simpset() addsimps [starfunCR, hypreal_inverse])); +qed "starfunCR_inverse"; +Addsimps [starfunCR_inverse RS sym]; + +Goal "( *fc* f) (hcomplex_of_complex a) = hcomplex_of_complex (f a)"; +by (auto_tac (claset(), + simpset() addsimps [starfunC,hcomplex_of_complex_def])); +qed "starfunC_eq"; +Addsimps [starfunC_eq]; + +Goal "( *fRc* f) (hypreal_of_real a) = hcomplex_of_complex (f a)"; +by (auto_tac (claset(), + simpset() addsimps [starfunRC,hcomplex_of_complex_def,hypreal_of_real_def])); +qed "starfunRC_eq"; +Addsimps [starfunRC_eq]; + +Goal "( *fcR* f) (hcomplex_of_complex a) = hypreal_of_real (f a)"; +by (auto_tac (claset(), + simpset() addsimps [starfunCR,hcomplex_of_complex_def,hypreal_of_real_def])); +qed "starfunCR_eq"; +Addsimps [starfunCR_eq]; + +Goal "( *fc* f) (hcomplex_of_complex a) @c= hcomplex_of_complex (f a)"; +by (Auto_tac); +qed "starfunC_capprox"; + +Goal "( *fRc* f) (hypreal_of_real a) @c= hcomplex_of_complex (f a)"; +by (Auto_tac); +qed "starfunRC_capprox"; + +Goal "( *fcR* f) (hcomplex_of_complex a) @= hypreal_of_real (f a)"; +by (Auto_tac); +qed "starfunCR_approx"; + +(* +Goal "( *fcNat* (%n. z ^ n)) N = (hcomplex_of_complex z) hcpow N"; +*) + +Goalw [hypnat_of_nat_def] + "( *fc* (%z. z ^ n)) Z = Z hcpow hypnat_of_nat n"; +by (res_inst_tac [("z","Z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), simpset() addsimps [hcpow,starfunC])); +qed "starfunC_hcpow"; + +Goal "( *fc* (%h. f (x + h))) xa = ( *fc* f) (hcomplex_of_complex x + xa)"; +by (res_inst_tac [("z","xa")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunC, + hcomplex_of_complex_def,hcomplex_add])); +qed "starfunC_lambda_cancel"; + +Goal "( *fcR* (%h. f (x + h))) xa = ( *fcR* f) (hcomplex_of_complex x + xa)"; +by (res_inst_tac [("z","xa")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunCR, + hcomplex_of_complex_def,hcomplex_add])); +qed "starfunCR_lambda_cancel"; + +Goal "( *fRc* (%h. f (x + h))) xa = ( *fRc* f) (hypreal_of_real x + xa)"; +by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [starfunRC, + hypreal_of_real_def,hypreal_add])); +qed "starfunRC_lambda_cancel"; + +Goal "( *fc* (%h. f(g(x + h)))) xa = ( *fc* (f o g)) (hcomplex_of_complex x + xa)"; +by (res_inst_tac [("z","xa")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunC, + hcomplex_of_complex_def,hcomplex_add])); +qed "starfunC_lambda_cancel2"; + +Goal "( *fcR* (%h. f(g(x + h)))) xa = ( *fcR* (f o g)) (hcomplex_of_complex x + xa)"; +by (res_inst_tac [("z","xa")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunCR, + hcomplex_of_complex_def,hcomplex_add])); +qed "starfunCR_lambda_cancel2"; + +Goal "( *fRc* (%h. f(g(x + h)))) xa = ( *fRc* (f o g)) (hypreal_of_real x + xa)"; +by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [starfunRC, + hypreal_of_real_def,hypreal_add])); +qed "starfunRC_lambda_cancel2"; + +Goal "[| ( *fc* f) xa @c= l; ( *fc* g) xa @c= m; \ +\ l: CFinite; m: CFinite \ +\ |] ==> ( *fc* (%x. f x * g x)) xa @c= l * m"; +by (dtac capprox_mult_CFinite 1); +by (REPEAT(assume_tac 1)); +by (auto_tac (claset() addIs [capprox_sym RSN (2,capprox_CFinite)], + simpset())); +qed "starfunC_mult_CFinite_capprox"; + +Goal "[| ( *fcR* f) xa @= l; ( *fcR* g) xa @= m; \ +\ l: HFinite; m: HFinite \ +\ |] ==> ( *fcR* (%x. f x * g x)) xa @= l * m"; +by (dtac approx_mult_HFinite 1); +by (REPEAT(assume_tac 1)); +by (auto_tac (claset() addIs [approx_sym RSN (2,approx_HFinite)], + simpset())); +qed "starfunCR_mult_HFinite_capprox"; + +Goal "[| ( *fRc* f) xa @c= l; ( *fRc* g) xa @c= m; \ +\ l: CFinite; m: CFinite \ +\ |] ==> ( *fRc* (%x. f x * g x)) xa @c= l * m"; +by (dtac capprox_mult_CFinite 1); +by (REPEAT(assume_tac 1)); +by (auto_tac (claset() addIs [capprox_sym RSN (2,capprox_CFinite)], + simpset())); +qed "starfunRC_mult_CFinite_capprox"; + +Goal "[| ( *fc* f) xa @c= l; ( *fc* g) xa @c= m \ +\ |] ==> ( *fc* (%x. f x + g x)) xa @c= l + m"; +by (auto_tac (claset() addIs [capprox_add], simpset())); +qed "starfunC_add_capprox"; + +Goal "[| ( *fRc* f) xa @c= l; ( *fRc* g) xa @c= m \ +\ |] ==> ( *fRc* (%x. f x + g x)) xa @c= l + m"; +by (auto_tac (claset() addIs [capprox_add], simpset())); +qed "starfunRC_add_capprox"; + +Goal "[| ( *fcR* f) xa @= l; ( *fcR* g) xa @= m \ +\ |] ==> ( *fcR* (%x. f x + g x)) xa @= l + m"; +by (auto_tac (claset() addIs [approx_add], simpset())); +qed "starfunCR_add_approx"; + +Goal "*fcR* cmod = hcmod"; +by (rtac ext 1); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunCR,hcmod])); +qed "starfunCR_cmod"; + +Goal "( *fc* inverse) x = inverse(x)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunC,hcomplex_inverse])); +qed "starfunC_inverse_inverse"; + +Goalw [hcomplex_divide_def,complex_divide_def] + "( *fc* f) xa / ( *fc* g) xa = ( *fc* (%x. f x / g x)) xa"; +by Auto_tac; +qed "starfunC_divide"; +Addsimps [starfunC_divide RS sym]; + +Goalw [hypreal_divide_def,real_divide_def] + "( *fcR* f) xa / ( *fcR* g) xa = ( *fcR* (%x. f x / g x)) xa"; +by Auto_tac; +qed "starfunCR_divide"; +Addsimps [starfunCR_divide RS sym]; + +Goalw [hcomplex_divide_def,complex_divide_def] + "( *fRc* f) xa / ( *fRc* g) xa = ( *fRc* (%x. f x / g x)) xa"; +by Auto_tac; +qed "starfunRC_divide"; +Addsimps [starfunRC_divide RS sym]; + +(*-----------------------------------------------------------------------------------*) +(* Internal functions - some redundancy with *fc* now *) +(*-----------------------------------------------------------------------------------*) + +Goalw [congruent_def] + "congruent hcomplexrel (%X. hcomplexrel``{%n. f n (X n)})"; +by (safe_tac (claset())); +by (ALLGOALS(Fuf_tac)); +qed "starfunC_n_congruent"; + +Goalw [starfunC_n_def] + "( *fcn* f) (Abs_hcomplex(hcomplexrel``{%n. X n})) = \ +\ Abs_hcomplex(hcomplexrel `` {%n. f n (X n)})"; +by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1); +by Auto_tac; +by (Ultra_tac 1); +qed "starfunC_n"; + +(** multiplication: ( *fn ) x ( *gn ) = *(fn x gn) **) + +Goal "( *fcn* f) z * ( *fcn* g) z = ( *fcn* (% i x. f i x * g i x)) z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), simpset() addsimps [starfunC_n,hcomplex_mult])); +qed "starfunC_n_mult"; + +(** addition: ( *fn ) + ( *gn ) = *(fn + gn) **) + +Goal "( *fcn* f) z + ( *fcn* g) z = ( *fcn* (%i x. f i x + g i x)) z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), simpset() addsimps [starfunC_n,hcomplex_add])); +qed "starfunC_n_add"; + +(** uminus **) + +Goal "- ( *fcn* g) z = ( *fcn* (%i x. - g i x)) z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), simpset() addsimps [starfunC_n,hcomplex_minus])); +qed "starfunC_n_minus"; + +(** subtraction: ( *fn ) - ( *gn ) = *(fn - gn) **) + +Goalw [hcomplex_diff_def,complex_diff_def] + "( *fcn* f) z - ( *fcn* g) z = ( *fcn* (%i x. f i x - g i x)) z"; +by (auto_tac (claset(), + simpset() addsimps [starfunC_n_add,starfunC_n_minus])); +qed "starfunNat_n_diff"; + +(** composition: ( *fn ) o ( *gn ) = *(fn o gn) **) + +Goal "( *fcn* (%i x. k)) z = hcomplex_of_complex k"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), + simpset() addsimps [starfunC_n, hcomplex_of_complex_def])); +qed "starfunC_n_const_fun"; +Addsimps [starfunC_n_const_fun]; + +Goal "( *fcn* f) (hcomplex_of_complex n) = Abs_hcomplex(hcomplexrel `` {%i. f i n})"; +by (auto_tac (claset(), simpset() addsimps [starfunC_n,hcomplex_of_complex_def])); +qed "starfunC_n_eq"; +Addsimps [starfunC_n_eq]; + +Goal "(( *fc* f) = ( *fc* g)) = (f = g)"; +by Auto_tac; +by (rtac ext 1 THEN rtac ccontr 1); +by (dres_inst_tac [("x","hcomplex_of_complex(x)")] fun_cong 1); +by (auto_tac (claset(), simpset() addsimps [starfunC,hcomplex_of_complex_def])); +qed "starfunC_eq_iff"; + +Goal "(( *fRc* f) = ( *fRc* g)) = (f = g)"; +by Auto_tac; +by (rtac ext 1 THEN rtac ccontr 1); +by (dres_inst_tac [("x","hypreal_of_real(x)")] fun_cong 1); +by Auto_tac; +qed "starfunRC_eq_iff"; + +Goal "(( *fcR* f) = ( *fcR* g)) = (f = g)"; +by Auto_tac; +by (rtac ext 1 THEN rtac ccontr 1); +by (dres_inst_tac [("x","hcomplex_of_complex(x)")] fun_cong 1); +by Auto_tac; +qed "starfunCR_eq_iff"; + +(*** more theorems ***) + +Goal "(( *fc* f) x = z) = ((( *fcR* (%x. Re(f x))) x = hRe (z)) & \ +\ (( *fcR* (%x. Im(f x))) x = hIm (z)))"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunCR,starfunC, + hIm,hRe,complex_Re_Im_cancel_iff])); +by (ALLGOALS(Ultra_tac)); +qed "starfunC_eq_Re_Im_iff"; + +Goal "(( *fc* f) x @c= z) = ((( *fcR* (%x. Re(f x))) x @= hRe (z)) & \ +\ (( *fcR* (%x. Im(f x))) x @= hIm (z)))"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunCR,starfunC, + hIm,hRe,capprox_approx_iff])); +qed "starfunC_approx_Re_Im_iff"; + +Goal "x @c= hcomplex_of_complex a ==> ( *fc* (%x. x)) x @c= hcomplex_of_complex a"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunC])); +qed "starfunC_Idfun_capprox"; + +Goal "( *fc* (%x. x)) x = x"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [starfunC])); +qed "starfunC_Id"; +Addsimps [starfunC_Id]; diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/CStar.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/CStar.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,57 @@ +(* Title : CStar.thy + Author : Jacques D. Fleuriot + Copyright : 2001 University of Edinburgh + Description : defining *-transforms in NSA which extends sets of complex numbers, + and complex functions +*) + +CStar = NSCA + + +constdefs + + (* nonstandard extension of sets *) + starsetC :: complex set => hcomplex set ("*sc* _" [80] 80) + "*sc* A == {x. ALL X: Rep_hcomplex(x). {n::nat. X n : A}: FreeUltrafilterNat}" + + (* internal sets *) + starsetC_n :: (nat => complex set) => hcomplex set ("*scn* _" [80] 80) + "*scn* As == {x. ALL X: Rep_hcomplex(x). {n::nat. X n : (As n)}: FreeUltrafilterNat}" + + InternalCSets :: "hcomplex set set" + "InternalCSets == {X. EX As. X = *scn* As}" + + (* star transform of functions f: Complex --> Complex *) + + starfunC :: (complex => complex) => hcomplex => hcomplex ("*fc* _" [80] 80) + "*fc* f == (%x. Abs_hcomplex(UN X: Rep_hcomplex(x). hcomplexrel``{%n. f (X n)}))" + + starfunC_n :: (nat => (complex => complex)) => hcomplex => hcomplex ("*fcn* _" [80] 80) + "*fcn* F == (%x. Abs_hcomplex(UN X: Rep_hcomplex(x). hcomplexrel``{%n. (F n)(X n)}))" + + InternalCFuns :: (hcomplex => hcomplex) set + "InternalCFuns == {X. EX F. X = *fcn* F}" + + + (* star transform of functions f: Real --> Complex *) + + starfunRC :: (real => complex) => hypreal => hcomplex ("*fRc* _" [80] 80) + "*fRc* f == (%x. Abs_hcomplex(UN X: Rep_hypreal(x). hcomplexrel``{%n. f (X n)}))" + + starfunRC_n :: (nat => (real => complex)) => hypreal => hcomplex ("*fRcn* _" [80] 80) + "*fRcn* F == (%x. Abs_hcomplex(UN X: Rep_hypreal(x). hcomplexrel``{%n. (F n)(X n)}))" + + InternalRCFuns :: (hypreal => hcomplex) set + "InternalRCFuns == {X. EX F. X = *fRcn* F}" + + (* star transform of functions f: Complex --> Real; needed for Re and Im parts *) + + starfunCR :: (complex => real) => hcomplex => hypreal ("*fcR* _" [80] 80) + "*fcR* f == (%x. Abs_hypreal(UN X: Rep_hcomplex(x). hyprel``{%n. f (X n)}))" + + starfunCR_n :: (nat => (complex => real)) => hcomplex => hypreal ("*fcRn* _" [80] 80) + "*fcRn* F == (%x. Abs_hypreal(UN X: Rep_hcomplex(x). hyprel``{%n. (F n)(X n)}))" + + InternalCRFuns :: (hcomplex => hypreal) set + "InternalCRFuns == {X. EX F. X = *fcRn* F}" + +end \ No newline at end of file diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/Complex.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/Complex.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,1588 @@ +(* Title: Complex.ML + Author: Jacques D. Fleuriot + Copyright: 2001 University of Edinburgh + Description: Complex numbers +*) + +Goal "inj Rep_complex"; +by (rtac inj_inverseI 1); +by (rtac Rep_complex_inverse 1); +qed "inj_Rep_complex"; + +Goal "inj Abs_complex"; +by (rtac inj_inverseI 1); +by (rtac Abs_complex_inverse 1); +by (simp_tac (simpset() addsimps [complex_def]) 1); +qed "inj_Abs_complex"; +Addsimps [inj_Abs_complex RS injD]; + +Goal "(Abs_complex x = Abs_complex y) = (x = y)"; +by (auto_tac (claset() addDs [inj_Abs_complex RS injD],simpset())); +qed "Abs_complex_cancel_iff"; +Addsimps [Abs_complex_cancel_iff]; + +Goalw [complex_def] "(x,y) : complex"; +by (Auto_tac); +qed "pair_mem_complex"; +Addsimps [pair_mem_complex]; + +Goal "Rep_complex (Abs_complex (x,y)) = (x,y)"; +by (simp_tac (simpset() addsimps [Abs_complex_inverse]) 1); +qed "Abs_complex_inverse2"; +Addsimps [Abs_complex_inverse2]; + +val [prem] = goal Complex.thy + "(!!x y. z = Abs_complex(x,y) ==> P) ==> P"; +by (res_inst_tac [("p","Rep_complex z")] PairE 1); +by (dres_inst_tac [("f","Abs_complex")] arg_cong 1); +by (res_inst_tac [("x","x"),("y","y")] prem 1); +by (asm_full_simp_tac (simpset() addsimps [Rep_complex_inverse]) 1); +qed "eq_Abs_complex"; + +Goalw [Re_def] "Re(Abs_complex(x,y)) = x"; +by (Simp_tac 1); +qed "Re"; +Addsimps [Re]; + +Goalw [Im_def] "Im(Abs_complex(x,y)) = y"; +by (Simp_tac 1); +qed "Im"; +Addsimps [Im]; + +Goal "Abs_complex(Re(z),Im(z)) = z"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (Asm_simp_tac 1); +qed "Abs_complex_cancel"; +Addsimps [Abs_complex_cancel]; + +Goal "(w=z) = (Re(w) = Re(z) & Im(w) = Im(z))"; +by (res_inst_tac [("z","w")] eq_Abs_complex 1); +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset() addDs [inj_Abs_complex RS injD],simpset())); +qed "complex_Re_Im_cancel_iff"; + +Goalw [complex_zero_def] "Re 0 = 0"; +by (Simp_tac 1); +qed "complex_Re_zero"; + +Goalw [complex_zero_def] "Im 0 = 0"; +by (Simp_tac 1); +qed "complex_Im_zero"; +Addsimps [complex_Re_zero,complex_Im_zero]; + +Goalw [complex_one_def] "Re 1 = 1"; +by (Simp_tac 1); +qed "complex_Re_one"; +Addsimps [complex_Re_one]; + +Goalw [complex_one_def] "Im 1 = 0"; +by (Simp_tac 1); +qed "complex_Im_one"; +Addsimps [complex_Im_one]; + +Goalw [i_def] "Re(ii) = 0"; +by Auto_tac; +qed "complex_Re_i"; +Addsimps [complex_Re_i]; + +Goalw [i_def] "Im(ii) = 1"; +by Auto_tac; +qed "complex_Im_i"; +Addsimps [complex_Im_i]; + +Goalw [complex_of_real_def] "Re(complex_of_real 0) = 0"; +by (Simp_tac 1); +qed "Re_complex_of_real_zero"; +Addsimps [Re_complex_of_real_zero]; + +Goalw [complex_of_real_def] "Im(complex_of_real 0) = 0"; +by (Simp_tac 1); +qed "Im_complex_of_real_zero"; +Addsimps [Im_complex_of_real_zero]; + +Goalw [complex_of_real_def] "Re(complex_of_real 1) = 1"; +by (Simp_tac 1); +qed "Re_complex_of_real_one"; +Addsimps [Re_complex_of_real_one]; + +Goalw [complex_of_real_def] "Im(complex_of_real 1) = 0"; +by (Simp_tac 1); +qed "Im_complex_of_real_one"; +Addsimps [Im_complex_of_real_one]; + +Goalw [complex_of_real_def] "Re(complex_of_real z) = z"; +by Auto_tac; +qed "Re_complex_of_real"; +Addsimps [Re_complex_of_real]; + +Goalw [complex_of_real_def] "Im(complex_of_real z) = 0"; +by Auto_tac; +qed "Im_complex_of_real"; +Addsimps [Im_complex_of_real]; + +(*** negation ***) + +Goalw [complex_minus_def] "- Abs_complex(x,y) = Abs_complex(-x,-y)"; +by (Simp_tac 1); +qed "complex_minus"; + +Goalw [Re_def] "Re (-z) = - Re z"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_minus])); +qed "complex_Re_minus"; + +Goalw [Im_def] "Im (-z) = - Im z"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_minus])); +qed "complex_Im_minus"; + +Goalw [complex_minus_def] "- (- z) = (z::complex)"; +by (Simp_tac 1); +qed "complex_minus_minus"; +Addsimps [complex_minus_minus]; + +Goal "inj(%r::complex. -r)"; +by (rtac injI 1); +by (dres_inst_tac [("f","uminus")] arg_cong 1); +by (Asm_full_simp_tac 1); +qed "inj_complex_minus"; + +Goalw [complex_zero_def] "-(0::complex) = 0"; +by (simp_tac (simpset() addsimps [complex_minus]) 1); +qed "complex_minus_zero"; +Addsimps [complex_minus_zero]; + +Goal "(-x = 0) = (x = (0::complex))"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (auto_tac (claset() addDs [inj_Abs_complex RS injD],simpset() + addsimps [complex_zero_def,complex_minus])); +qed "complex_minus_zero_iff"; +Addsimps [complex_minus_zero_iff]; + +Goal "(0 = -x) = (x = (0::real))"; +by (auto_tac (claset() addDs [sym],simpset())); +qed "complex_minus_zero_iff2"; +Addsimps [complex_minus_zero_iff2]; + +Goal "(-x ~= 0) = (x ~= (0::complex))"; +by Auto_tac; +qed "complex_minus_not_zero_iff"; + +(*** addition ***) + +Goalw [complex_add_def] + "Abs_complex(x1,y1) + Abs_complex(x2,y2) = Abs_complex(x1+x2,y1+y2)"; +by (Simp_tac 1); +qed "complex_add"; + +Goalw [Re_def] "Re(x + y) = Re(x) + Re(y)"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (res_inst_tac [("z","y")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_add])); +qed "complex_Re_add"; + +Goalw [Im_def] "Im(x + y) = Im(x) + Im(y)"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (res_inst_tac [("z","y")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_add])); +qed "complex_Im_add"; + +Goalw [complex_add_def] "(u::complex) + v = v + u"; +by (simp_tac (simpset() addsimps [real_add_commute]) 1); +qed "complex_add_commute"; + +Goalw [complex_add_def] "((u::complex) + v) + w = u + (v + w)"; +by (simp_tac (simpset() addsimps [real_add_assoc]) 1); +qed "complex_add_assoc"; + +Goalw [complex_add_def] "(x::complex) + (y + z) = y + (x + z)"; +by (simp_tac (simpset() addsimps [real_add_left_commute]) 1); +qed "complex_add_left_commute"; + +val complex_add_ac = [complex_add_assoc,complex_add_commute, + complex_add_left_commute]; + +Goalw [complex_add_def,complex_zero_def] "(0::complex) + z = z"; +by (Simp_tac 1); +qed "complex_add_zero_left"; +Addsimps [complex_add_zero_left]; + +Goalw [complex_add_def,complex_zero_def] "z + (0::complex) = z"; +by (Simp_tac 1); +qed "complex_add_zero_right"; +Addsimps [complex_add_zero_right]; + +Goalw [complex_add_def,complex_minus_def,complex_zero_def] + "z + -z = (0::complex)"; +by (Simp_tac 1); +qed "complex_add_minus_right_zero"; +Addsimps [complex_add_minus_right_zero]; + +Goalw [complex_add_def,complex_minus_def,complex_zero_def] + "-z + z = (0::complex)"; +by (Simp_tac 1); +qed "complex_add_minus_left_zero"; +Addsimps [complex_add_minus_left_zero]; + +Goal "z + (- z + w) = (w::complex)"; +by (simp_tac (simpset() addsimps [complex_add_assoc RS sym]) 1); +qed "complex_add_minus_cancel"; + +Goal "(-z) + (z + w) = (w::complex)"; +by (simp_tac (simpset() addsimps [complex_add_assoc RS sym]) 1); +qed "complex_minus_add_cancel"; + +Addsimps [complex_add_minus_cancel, complex_minus_add_cancel]; + +Goal "x + y = (0::complex) ==> x = -y"; +by (auto_tac (claset(),simpset() addsimps [complex_Re_Im_cancel_iff, + complex_Re_add,complex_Im_add,complex_Re_minus,complex_Im_minus])); +qed "complex_add_minus_eq_minus"; + +Goal "-(x + y) = -x + -(y::complex)"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (res_inst_tac [("z","y")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_minus,complex_add])); +qed "complex_minus_add_distrib"; +Addsimps [complex_minus_add_distrib]; + +Goal "((x::complex) + y = x + z) = (y = z)"; +by (Step_tac 1); +by (dres_inst_tac [("f","%t.-x + t")] arg_cong 1); +by (asm_full_simp_tac (simpset() addsimps [complex_add_assoc RS sym]) 1); +qed "complex_add_left_cancel"; +AddIffs [complex_add_left_cancel]; + +Goal "(y + (x::complex)= z + x) = (y = z)"; +by (simp_tac (simpset() addsimps [complex_add_commute]) 1); +qed "complex_add_right_cancel"; +Addsimps [complex_add_right_cancel]; + +Goal "((x::complex) = y) = (0 = x + - y)"; +by (Step_tac 1); +by (res_inst_tac [("x1","-y")] + (complex_add_right_cancel RS iffD1) 2); +by (Auto_tac); +qed "complex_eq_minus_iff"; + +Goal "((x::complex) = y) = (x + - y = 0)"; +by (Step_tac 1); +by (res_inst_tac [("x1","-y")] + (complex_add_right_cancel RS iffD1) 2); +by (Auto_tac); +qed "complex_eq_minus_iff2"; + +Goal "(0::complex) - x = -x"; +by (simp_tac (simpset() addsimps [complex_diff_def]) 1); +qed "complex_diff_0"; + +Goal "x - (0::complex) = x"; +by (simp_tac (simpset() addsimps [complex_diff_def]) 1); +qed "complex_diff_0_right"; + +Goal "x - x = (0::complex)"; +by (simp_tac (simpset() addsimps [complex_diff_def]) 1); +qed "complex_diff_self"; + +Addsimps [complex_diff_0, complex_diff_0_right, complex_diff_self]; + +Goalw [complex_diff_def] + "Abs_complex(x1,y1) - Abs_complex(x2,y2) = Abs_complex(x1-x2,y1-y2)"; +by (simp_tac (simpset() addsimps [complex_add,complex_minus]) 1); +qed "complex_diff"; + +Goal "((x::complex) - y = z) = (x = z + y)"; +by (auto_tac (claset(),simpset() addsimps [complex_diff_def,complex_add_assoc])); +qed "complex_diff_eq_eq"; + +(*** complex multiplication ***) + +Goalw [complex_mult_def] + "Abs_complex(x1,y1) * Abs_complex(x2,y2) = \ +\ Abs_complex(x1*x2 - y1*y2,x1*y2 + y1*x2)"; +by (Simp_tac 1); +qed "complex_mult"; + +Goalw [complex_mult_def] "(w::complex) * z = z * w"; +by (simp_tac (simpset() addsimps [real_mult_commute,real_add_commute]) 1); +qed "complex_mult_commute"; + +Goalw [complex_mult_def] "((u::complex) * v) * w = u * (v * w)"; +by (simp_tac (simpset() addsimps [complex_Re_Im_cancel_iff, + real_mult_assoc,real_diff_mult_distrib2, + real_add_mult_distrib2,real_diff_mult_distrib, + real_add_mult_distrib,real_mult_left_commute]) 1); +qed "complex_mult_assoc"; + +Goalw [complex_mult_def] "(x::complex) * (y * z) = y * (x * z)"; +by (simp_tac (simpset() addsimps [complex_Re_Im_cancel_iff, + real_mult_left_commute,real_diff_mult_distrib2, + real_add_mult_distrib2]) 1); +qed "complex_mult_left_commute"; + +val complex_mult_ac = [complex_mult_assoc,complex_mult_commute, + complex_mult_left_commute]; + +Goalw [complex_one_def] "(1::complex) * z = z"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (asm_simp_tac (simpset() addsimps [complex_mult]) 1); +qed "complex_mult_one_left"; +Addsimps [complex_mult_one_left]; + +Goal "z * (1::complex) = z"; +by (simp_tac (simpset() addsimps [complex_mult_commute]) 1); +qed "complex_mult_one_right"; +Addsimps [complex_mult_one_right]; + +Goalw [complex_zero_def] "(0::complex) * z = 0"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (asm_simp_tac (simpset() addsimps [complex_mult]) 1); +qed "complex_mult_zero_left"; +Addsimps [complex_mult_zero_left]; + +Goal "z * 0 = (0::complex)"; +by (simp_tac (simpset() addsimps [complex_mult_commute]) 1); +qed "complex_mult_zero_right"; +Addsimps [complex_mult_zero_right]; + +Goalw [complex_divide_def] "0 / z = (0::complex)"; +by Auto_tac; +qed "complex_divide_zero"; +Addsimps [complex_divide_zero]; + +Goal "-(x * y) = -x * (y::complex)"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (res_inst_tac [("z","y")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult,complex_minus, + real_diff_def])); +qed "complex_minus_mult_eq1"; + +Goal "-(x * y) = x * -(y::complex)"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (res_inst_tac [("z","y")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult,complex_minus, + real_diff_def])); +qed "complex_minus_mult_eq2"; + +Addsimps [ complex_minus_mult_eq1 RS sym, complex_minus_mult_eq2 RS sym]; + +Goal "-(1::complex) * z = -z"; +by (Simp_tac 1); +qed "complex_mult_minus_one"; +Addsimps [complex_mult_minus_one]; + +Goal "z * -(1::complex) = -z"; +by (stac complex_mult_commute 1); +by (Simp_tac 1); +qed "complex_mult_minus_one_right"; +Addsimps [complex_mult_minus_one_right]; + +Goal "-x * -y = x * (y::complex)"; +by (Simp_tac 1); +qed "complex_minus_mult_cancel"; +Addsimps [complex_minus_mult_cancel]; + +Goal "-x * y = x * -(y::complex)"; +by (Simp_tac 1); +qed "complex_minus_mult_commute"; + +qed_goal "complex_add_assoc_cong" thy + "!!z. (z::complex) + v = z' + v' ==> z + (v + w) = z' + (v' + w)" + (fn _ => [(asm_simp_tac (simpset() addsimps [complex_add_assoc RS sym]) 1)]); + +qed_goal "complex_add_assoc_swap" thy "(z::complex) + (v + w) = v + (z + w)" + (fn _ => [(REPEAT (ares_tac [complex_add_commute RS complex_add_assoc_cong] 1))]); + +Goal "((z1::complex) + z2) * w = (z1 * w) + (z2 * w)"; +by (res_inst_tac [("z","z1")] eq_Abs_complex 1); +by (res_inst_tac [("z","z2")] eq_Abs_complex 1); +by (res_inst_tac [("z","w")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult,complex_add, + real_add_mult_distrib,real_diff_def] @ real_add_ac)); +qed "complex_add_mult_distrib"; + +Goal "(w::complex) * (z1 + z2) = (w * z1) + (w * z2)"; +by (res_inst_tac [("z1","z1 + z2")] (complex_mult_commute RS ssubst) 1); +by (simp_tac (simpset() addsimps [complex_add_mult_distrib]) 1); +by (simp_tac (simpset() addsimps [complex_mult_commute]) 1); +qed "complex_add_mult_distrib2"; + +Goalw [complex_zero_def,complex_one_def] "(0::complex) ~= 1"; +by (simp_tac (simpset() addsimps [complex_Re_Im_cancel_iff]) 1); +qed "complex_zero_not_eq_one"; +Addsimps [complex_zero_not_eq_one]; +Addsimps [complex_zero_not_eq_one RS not_sym]; + +(*** inverse ***) +Goalw [complex_inverse_def] "inverse (Abs_complex(x,y)) = \ +\ Abs_complex(x/(x ^ 2 + y ^ 2),-y/(x ^ 2 + y ^ 2))"; +by (Simp_tac 1); +qed "complex_inverse"; + +Goalw [complex_inverse_def,complex_zero_def] "inverse 0 = (0::complex)"; +by Auto_tac; +qed "COMPLEX_INVERSE_ZERO"; + +Goal "a / (0::complex) = 0"; +by (simp_tac (simpset() addsimps [complex_divide_def, COMPLEX_INVERSE_ZERO]) 1); +qed "COMPLEX_DIVISION_BY_ZERO"; (*NOT for adding to default simpset*) + +fun complex_div_undefined_case_tac s i = + case_tac s i THEN + asm_simp_tac (simpset() addsimps [COMPLEX_DIVISION_BY_ZERO, COMPLEX_INVERSE_ZERO]) i; + +(*REMOVE?: +lemmas:replace previous versions to accommodate new behaviour of simplification +Goal "x ^ 2 + y ^ 2 = 0 ==> x = (0::real)"; +by (auto_tac (claset() addIs [real_sum_squares_cancel], + simpset() addsimps [CLAIM "2 = Suc(Suc 0)"])); +qed "real_sum_squares_cancel"; + +Goal "x ^ 2 + y ^ 2 = 0 ==> y = (0::real)"; +by (auto_tac (claset() addIs [real_sum_squares_cancel2], + simpset() addsimps [CLAIM "2 = Suc(Suc 0)"])); +qed "real_sum_squares_cancel2"; +*) + +Goal "z ~= (0::complex) ==> inverse(z) * z = 1"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult,complex_inverse, + complex_one_def,complex_zero_def,real_add_divide_distrib RS sym, + realpow_two_eq_mult] @ real_mult_ac)); +by (dres_inst_tac [("y","y")] real_sum_squares_not_zero 1); +by (dres_inst_tac [("x","x")] real_sum_squares_not_zero2 2); +by Auto_tac; +qed "complex_mult_inv_left"; +Addsimps [complex_mult_inv_left]; + +Goal "z ~= (0::complex) ==> z * inverse(z) = 1"; +by (auto_tac (claset() addIs [complex_mult_commute RS subst],simpset())); +qed "complex_mult_inv_right"; +Addsimps [complex_mult_inv_right]; + +Goal "(c::complex) ~= 0 ==> (c*a=c*b) = (a=b)"; +by Auto_tac; +by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1); +by (asm_full_simp_tac (simpset() addsimps complex_mult_ac) 1); +qed "complex_mult_left_cancel"; + +Goal "(c::complex) ~= 0 ==> (a*c=b*c) = (a=b)"; +by (Step_tac 1); +by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1); +by (asm_full_simp_tac (simpset() addsimps complex_mult_ac) 1); +qed "complex_mult_right_cancel"; + +Goal "z ~= 0 ==> inverse(z::complex) ~= 0"; +by (Step_tac 1); +by (ftac (complex_mult_right_cancel RS iffD2) 1); +by (thin_tac "inverse z = 0" 2); +by (assume_tac 1 THEN Auto_tac); +qed "complex_inverse_not_zero"; +Addsimps [complex_inverse_not_zero]; + +Goal "!!x. [| x ~= 0; y ~= (0::complex) |] ==> x * y ~= 0"; +by (Step_tac 1); +by (dres_inst_tac [("f","%z. inverse x*z")] arg_cong 1); +by (asm_full_simp_tac (simpset() addsimps [complex_mult_assoc RS sym]) 1); +qed "complex_mult_not_zero"; + +bind_thm ("complex_mult_not_zeroE",complex_mult_not_zero RS notE); + +Goal "inverse(inverse (x::complex)) = x"; +by (complex_div_undefined_case_tac "x = 0" 1); +by (res_inst_tac [("c1","inverse x")] (complex_mult_right_cancel RS iffD1) 1); +by (etac complex_inverse_not_zero 1); +by (auto_tac (claset() addDs [complex_inverse_not_zero],simpset())); +qed "complex_inverse_inverse"; +Addsimps [complex_inverse_inverse]; + +Goalw [complex_one_def] "inverse(1::complex) = 1"; +by (simp_tac (simpset() addsimps [complex_inverse,realpow_num_two]) 1); +qed "complex_inverse_one"; +Addsimps [complex_inverse_one]; + +Goal "inverse(-x) = -inverse(x::complex)"; +by (complex_div_undefined_case_tac "x = 0" 1); +by (res_inst_tac [("c1","-x")] (complex_mult_right_cancel RS iffD1) 1); +by (stac complex_mult_inv_left 2); +by Auto_tac; +qed "complex_minus_inverse"; + +Goal "inverse(x*y) = inverse x * inverse (y::complex)"; +by (complex_div_undefined_case_tac "x = 0" 1); +by (complex_div_undefined_case_tac "y = 0" 1); +by (res_inst_tac [("c1","x*y")] (complex_mult_left_cancel RS iffD1) 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult_not_zero] + @ complex_mult_ac)); +by (auto_tac (claset(),simpset() addsimps [complex_mult_not_zero, + complex_mult_assoc RS sym])); +qed "complex_inverse_distrib"; + + +(*** division ***) + +(*adding some of these theorems to simpset as for reals: + not 100% convinced for some*) + +Goal "(x::complex) * (y/z) = (x*y)/z"; +by (simp_tac (simpset() addsimps [complex_divide_def, complex_mult_assoc]) 1); +qed "complex_times_divide1_eq"; + +Goal "(y/z) * (x::complex) = (y*x)/z"; +by (simp_tac (simpset() addsimps [complex_divide_def]@complex_mult_ac) 1); +qed "complex_times_divide2_eq"; + +Addsimps [complex_times_divide1_eq, complex_times_divide2_eq]; + +Goal "(x::complex) / (y/z) = (x*z)/y"; +by (simp_tac (simpset() addsimps [complex_divide_def, complex_inverse_distrib]@ + complex_mult_ac) 1); +qed "complex_divide_divide1_eq"; + +Goal "((x::complex) / y) / z = x/(y*z)"; +by (simp_tac (simpset() addsimps [complex_divide_def, complex_inverse_distrib, + complex_mult_assoc]) 1); +qed "complex_divide_divide2_eq"; + +Addsimps [complex_divide_divide1_eq, complex_divide_divide2_eq]; + +(** As with multiplication, pull minus signs OUT of the / operator **) + +Goal "(-x) / (y::complex) = - (x/y)"; +by (simp_tac (simpset() addsimps [complex_divide_def]) 1); +qed "complex_minus_divide_eq"; +Addsimps [complex_minus_divide_eq]; + +Goal "(x / -(y::complex)) = - (x/y)"; +by (simp_tac (simpset() addsimps [complex_divide_def, complex_minus_inverse]) 1); +qed "complex_divide_minus_eq"; +Addsimps [complex_divide_minus_eq]; + +Goal "(x+y)/(z::complex) = x/z + y/z"; +by (simp_tac (simpset() addsimps [complex_divide_def, complex_add_mult_distrib]) 1); +qed "complex_add_divide_distrib"; + +(*---------------------------------------------------------------------------*) +(* Embedding properties for complex_of_real map *) +(*---------------------------------------------------------------------------*) + +Goal "inj complex_of_real"; +by (rtac injI 1); +by (auto_tac (claset() addDs [inj_Abs_complex RS injD], + simpset() addsimps [complex_of_real_def])); +qed "inj_complex_of_real"; + +Goalw [complex_one_def,complex_of_real_def] + "complex_of_real 1 = 1"; +by (rtac refl 1); +qed "complex_of_real_one"; +Addsimps [complex_of_real_one]; + +Goalw [complex_zero_def,complex_of_real_def] + "complex_of_real 0 = 0"; +by (rtac refl 1); +qed "complex_of_real_zero"; +Addsimps [complex_of_real_zero]; + +Goal "(complex_of_real x = complex_of_real y) = (x = y)"; +by (auto_tac (claset() addDs [inj_complex_of_real RS injD],simpset())); +qed "complex_of_real_eq_iff"; +AddIffs [complex_of_real_eq_iff]; + +Goal "complex_of_real(-x) = - complex_of_real x"; +by (simp_tac (simpset() addsimps [complex_of_real_def,complex_minus]) 1); +qed "complex_of_real_minus"; + +Goal "complex_of_real(inverse x) = inverse(complex_of_real x)"; +by (real_div_undefined_case_tac "x=0" 1); +by (simp_tac (simpset() addsimps [DIVISION_BY_ZERO,COMPLEX_INVERSE_ZERO]) 1); +by (auto_tac (claset(),simpset() addsimps [complex_inverse, + complex_of_real_def,realpow_num_two,real_divide_def, + real_inverse_distrib])); +qed "complex_of_real_inverse"; + +Goal "complex_of_real x + complex_of_real y = complex_of_real (x + y)"; +by (simp_tac (simpset() addsimps [complex_add,complex_of_real_def]) 1); +qed "complex_of_real_add"; + +Goal "complex_of_real x - complex_of_real y = complex_of_real (x - y)"; +by (simp_tac (simpset() addsimps [complex_of_real_minus RS sym, + complex_diff_def,complex_of_real_add]) 1); +qed "complex_of_real_diff"; + +Goal "complex_of_real x * complex_of_real y = complex_of_real (x * y)"; +by (simp_tac (simpset() addsimps [complex_mult,complex_of_real_def]) 1); +qed "complex_of_real_mult"; + +Goalw [complex_divide_def] + "complex_of_real x / complex_of_real y = complex_of_real(x/y)"; +by (real_div_undefined_case_tac "y=0" 1); +by (simp_tac (simpset() addsimps [rename_numerals DIVISION_BY_ZERO, + COMPLEX_INVERSE_ZERO]) 1); +by (asm_simp_tac (simpset() addsimps [complex_of_real_mult RS sym, + complex_of_real_inverse,real_divide_def]) 1); +qed "complex_of_real_divide"; + +Goal "complex_of_real (x ^ n) = (complex_of_real x) ^ n"; +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps [complex_of_real_mult RS sym])); +qed "complex_of_real_pow"; + +Goalw [cmod_def] "cmod (Abs_complex(x,y)) = sqrt(x ^ 2 + y ^ 2)"; +by (Simp_tac 1); +qed "complex_mod"; + +Goalw [cmod_def] "cmod(0) = 0"; +by (Simp_tac 1); +qed "complex_mod_zero"; +Addsimps [complex_mod_zero]; + +Goalw [cmod_def] "cmod(1) = 1"; +by (simp_tac (simpset() addsimps [realpow_num_two]) 1); +qed "complex_mod_one"; +Addsimps [complex_mod_one]; + +Goalw [complex_of_real_def] "cmod(complex_of_real x) = abs x"; +by (simp_tac (simpset() addsimps [complex_mod,realpow_num_two]) 1); +qed "complex_mod_complex_of_real"; +Addsimps [complex_mod_complex_of_real]; + +Goal "complex_of_real (abs x) = complex_of_real(cmod(complex_of_real x))"; +by (Simp_tac 1); +qed "complex_of_real_abs"; + +(*---------------------------------------------------------------------------*) +(* conjugation is an automorphism *) +(*---------------------------------------------------------------------------*) + +Goalw [cnj_def] "cnj (Abs_complex(x,y)) = Abs_complex(x,-y)"; +by (Simp_tac 1); +qed "complex_cnj"; + +Goal "inj cnj"; +by (rtac injI 1); +by (auto_tac (claset(),simpset() addsimps [cnj_def, + Abs_complex_cancel_iff,complex_Re_Im_cancel_iff])); +qed "inj_cnj"; + +Goal "(cnj x = cnj y) = (x = y)"; +by (auto_tac (claset() addDs [inj_cnj RS injD],simpset())); +qed "complex_cnj_cancel_iff"; +Addsimps [complex_cnj_cancel_iff]; + +Goalw [cnj_def] "cnj (cnj z) = z"; +by (Simp_tac 1); +qed "complex_cnj_cnj"; +Addsimps [complex_cnj_cnj]; + +Goalw [complex_of_real_def] "cnj (complex_of_real x) = complex_of_real x"; +by (simp_tac (simpset() addsimps [complex_cnj]) 1); +qed "complex_cnj_complex_of_real"; +Addsimps [complex_cnj_complex_of_real]; + +Goal "cmod (cnj z) = cmod z"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (asm_simp_tac (simpset() addsimps [complex_cnj,complex_mod,realpow_num_two]) 1); +qed "complex_mod_cnj"; +Addsimps [complex_mod_cnj]; + +Goalw [cnj_def] "cnj (-z) = - cnj z"; +by (simp_tac (simpset() addsimps [complex_minus, + complex_Re_minus,complex_Im_minus]) 1); +qed "complex_cnj_minus"; + +Goal "cnj(inverse z) = inverse(cnj z)"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (asm_simp_tac (simpset() addsimps [complex_cnj,complex_inverse, + realpow_num_two]) 1); +qed "complex_cnj_inverse"; + +Goal "cnj(w + z) = cnj(w) + cnj(z)"; +by (res_inst_tac [("z","w")] eq_Abs_complex 1); +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (asm_simp_tac (simpset() addsimps [complex_cnj,complex_add]) 1); +qed "complex_cnj_add"; + +Goalw [complex_diff_def] "cnj(w - z) = cnj(w) - cnj(z)"; +by (simp_tac (simpset() addsimps [complex_cnj_add,complex_cnj_minus]) 1); +qed "complex_cnj_diff"; + +Goal "cnj(w * z) = cnj(w) * cnj(z)"; +by (res_inst_tac [("z","w")] eq_Abs_complex 1); +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (asm_simp_tac (simpset() addsimps [complex_cnj,complex_mult]) 1); +qed "complex_cnj_mult"; + +Goalw [complex_divide_def] "cnj(w / z) = (cnj w)/(cnj z)"; +by (simp_tac (simpset() addsimps [complex_cnj_mult,complex_cnj_inverse]) 1); +qed "complex_cnj_divide"; + +Goalw [cnj_def,complex_one_def] "cnj 1 = 1"; +by (Simp_tac 1); +qed "complex_cnj_one"; +Addsimps [complex_cnj_one]; + +Goal "cnj(z ^ n) = cnj(z) ^ n"; +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps [complex_cnj_mult])); +qed "complex_cnj_pow"; + +Goal "z + cnj z = complex_of_real (2 * Re(z))"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (asm_simp_tac (simpset() addsimps [complex_add,complex_cnj, + complex_of_real_def]) 1); +qed "complex_add_cnj"; + +Goal "z - cnj z = complex_of_real (2 * Im(z)) * ii"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (asm_simp_tac (simpset() addsimps [complex_add,complex_cnj, + complex_of_real_def,complex_diff_def,complex_minus, + i_def,complex_mult]) 1); +qed "complex_diff_cnj"; + +goalw Complex.thy [cnj_def,complex_zero_def] + "cnj 0 = 0"; +by Auto_tac; +qed "complex_cnj_zero"; +Addsimps [complex_cnj_zero]; + +goal Complex.thy "(cnj z = 0) = (z = 0)"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_zero_def, + complex_cnj])); +qed "complex_cnj_zero_iff"; +AddIffs [complex_cnj_zero_iff]; + +Goal "z * cnj z = complex_of_real (Re(z) ^ 2 + Im(z) ^ 2)"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_cnj,complex_mult, + complex_of_real_def,realpow_num_two])); +qed "complex_mult_cnj"; + +(*---------------------------------------------------------------------------*) +(* algebra *) +(*---------------------------------------------------------------------------*) + +Goal "(x*y = (0::complex)) = (x = 0 | y = 0)"; +by Auto_tac; +by (auto_tac (claset() addIs [ccontr] addDs + [complex_mult_not_zero],simpset())); +qed "complex_mult_zero_iff"; +AddIffs [complex_mult_zero_iff]; + +Goalw [complex_zero_def] "(x + y = x) = (y = (0::complex))"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (res_inst_tac [("z","y")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_add])); +qed "complex_add_left_cancel_zero"; +Addsimps [complex_add_left_cancel_zero]; + +Goalw [complex_diff_def] + "((z1::complex) - z2) * w = (z1 * w) - (z2 * w)"; +by (simp_tac (simpset() addsimps [complex_add_mult_distrib]) 1); +qed "complex_diff_mult_distrib"; + +Goalw [complex_diff_def] + "(w::complex) * (z1 - z2) = (w * z1) - (w * z2)"; +by (simp_tac (simpset() addsimps [complex_add_mult_distrib2]) 1); +qed "complex_diff_mult_distrib2"; + +(*---------------------------------------------------------------------------*) +(* modulus *) +(*---------------------------------------------------------------------------*) + +(* +Goal "[| sqrt(x) = 0; 0 <= x |] ==> x = 0"; +by (auto_tac (claset() addIs [real_sqrt_eq_zero_cancel], + simpset())); +qed "real_sqrt_eq_zero_cancel2"; +*) + +Goal "(cmod x = 0) = (x = 0)"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (auto_tac (claset() addIs + [real_sum_squares_cancel,real_sum_squares_cancel2], + simpset() addsimps [complex_mod,complex_zero_def, + realpow_num_two])); +qed "complex_mod_eq_zero_cancel"; +Addsimps [complex_mod_eq_zero_cancel]; + +Goal "cmod (complex_of_real(real (n::nat))) = real n"; +by (Simp_tac 1); +qed "complex_mod_complex_of_real_of_nat"; +Addsimps [complex_mod_complex_of_real_of_nat]; + +Goal "cmod (-x) = cmod(x)"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (asm_simp_tac (simpset() addsimps [complex_mod,complex_minus, + realpow_num_two]) 1); +qed "complex_mod_minus"; +Addsimps [complex_mod_minus]; + +Goal "cmod(z * cnj(z)) = cmod(z) ^ 2"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (asm_simp_tac (simpset() addsimps [complex_mod,complex_cnj, + complex_mult, CLAIM "0 ^ 2 = (0::real)"]) 1); +by (simp_tac (simpset() addsimps [realpow_two_eq_mult]) 1); +qed "complex_mod_mult_cnj"; + +Goalw [cmod_def] "cmod(Abs_complex(x,y)) ^ 2 = x ^ 2 + y ^ 2"; +by Auto_tac; +qed "complex_mod_squared"; + +Goalw [cmod_def] "0 <= cmod x"; +by (auto_tac (claset() addIs [real_sqrt_ge_zero],simpset())); +qed "complex_mod_ge_zero"; +Addsimps [complex_mod_ge_zero]; + +Goal "abs(cmod x) = cmod x"; +by (auto_tac (claset() addIs [abs_eqI1],simpset())); +qed "abs_cmod_cancel"; +Addsimps [abs_cmod_cancel]; + +Goal "cmod(x*y) = cmod(x) * cmod(y)"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (res_inst_tac [("z","y")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult, + complex_mod,real_sqrt_mult_distrib2 RS sym] delsimps [realpow_Suc])); +by (res_inst_tac [("n","1")] realpow_Suc_cancel_eq 1); +by (auto_tac (claset(),simpset() addsimps [realpow_num_two RS sym] + delsimps [realpow_Suc])); +by (auto_tac (claset(),simpset() addsimps [real_diff_def,realpow_num_two, + real_add_mult_distrib2,real_add_mult_distrib] @ real_add_ac @ + real_mult_ac)); +qed "complex_mod_mult"; + +Goal "cmod(x + y) ^ 2 = cmod(x) ^ 2 + cmod(y) ^ 2 + 2 * Re(x * cnj y)"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (res_inst_tac [("z","y")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_add, + complex_mod_squared,complex_mult,complex_cnj,real_diff_def] + delsimps [realpow_Suc])); +by (auto_tac (claset(),simpset() addsimps [real_add_mult_distrib2, + real_add_mult_distrib,realpow_num_two] @ real_mult_ac @ real_add_ac)); +qed "complex_mod_add_squared_eq"; + +Goal "Re(x * cnj y) <= cmod(x * cnj y)"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (res_inst_tac [("z","y")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mod, + complex_mult,complex_cnj,real_diff_def] delsimps [realpow_Suc])); +qed "complex_Re_mult_cnj_le_cmod"; +Addsimps [complex_Re_mult_cnj_le_cmod]; + +Goal "Re(x * cnj y) <= cmod(x * y)"; +by (cut_inst_tac [("x","x"),("y","y")] complex_Re_mult_cnj_le_cmod 1); +by (asm_full_simp_tac (simpset() addsimps [complex_mod_mult]) 1); +qed "complex_Re_mult_cnj_le_cmod2"; +Addsimps [complex_Re_mult_cnj_le_cmod2]; + +Goal "((x::real) + y) ^ 2 = x ^ 2 + y ^ 2 + 2 * x * y"; +by (simp_tac (simpset() addsimps [real_add_mult_distrib, + real_add_mult_distrib2,realpow_num_two]) 1); +qed "real_sum_squared_expand"; + +Goal "cmod (x + y) ^ 2 <= (cmod(x) + cmod(y)) ^ 2"; +by (simp_tac (simpset() addsimps [real_sum_squared_expand, + complex_mod_add_squared_eq,real_mult_assoc,complex_mod_mult RS sym]) 1); +qed "complex_mod_triangle_squared"; +Addsimps [complex_mod_triangle_squared]; + +Goal "- cmod x <= cmod x"; +by (rtac (complex_mod_ge_zero RSN (2,real_le_trans)) 1); +by (Simp_tac 1); +qed "complex_mod_minus_le_complex_mod"; +Addsimps [complex_mod_minus_le_complex_mod]; + +Goal "cmod (x + y) <= cmod(x) + cmod(y)"; +by (res_inst_tac [("n","1")] realpow_increasing 1); +by (auto_tac (claset() addIs [(complex_mod_ge_zero RSN (2,real_le_trans))], + simpset() addsimps [realpow_num_two RS sym])); +qed "complex_mod_triangle_ineq"; +Addsimps [complex_mod_triangle_ineq]; + +Goal "cmod(b + a) - cmod b <= cmod a"; +by (cut_inst_tac [("x1","b"),("y1","a"),("z","-cmod b")] + (complex_mod_triangle_ineq RS real_add_le_mono1) 1); +by (Simp_tac 1); +qed "complex_mod_triangle_ineq2"; +Addsimps [complex_mod_triangle_ineq2]; + +Goal "cmod (x - y) = cmod (y - x)"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (res_inst_tac [("z","y")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_diff, + complex_mod,real_diff_mult_distrib2,realpow_num_two, + real_diff_mult_distrib] @ real_add_ac @ real_mult_ac)); +qed "complex_mod_diff_commute"; + +Goal "[| cmod x < r; cmod y < s |] ==> cmod (x + y) < r + s"; +by (auto_tac (claset() addIs [order_le_less_trans, + complex_mod_triangle_ineq],simpset())); +qed "complex_mod_add_less"; + +Goal "[| cmod x < r; cmod y < s |] ==> cmod (x * y) < r * s"; +by (auto_tac (claset() addIs [real_mult_less_mono'],simpset() + addsimps [complex_mod_mult])); +qed "complex_mod_mult_less"; + +goal Complex.thy "cmod(a) - cmod(b) <= cmod(a + b)"; +by (res_inst_tac [("R1.0","cmod(a)"),("R2.0","cmod(b)")] + real_linear_less2 1); +by Auto_tac; +by (dtac (ARITH_PROVE "a < b ==> a - (b::real) < 0") 1); +by (rtac real_le_trans 1 THEN rtac order_less_imp_le 1); +by Auto_tac; +by (dtac (ARITH_PROVE "a < b ==> 0 < (b::real) - a") 1); +by (rtac (ARITH_PROVE "a <= b + c ==> a - c <= (b::real)") 1); +by (rtac (complex_mod_minus RS subst) 1); +by (rtac real_le_trans 1); +by (rtac complex_mod_triangle_ineq 2); +by (auto_tac (claset(),simpset() addsimps complex_add_ac)); +qed "complex_mod_diff_ineq"; +Addsimps [complex_mod_diff_ineq]; + +Goal "Re z <= cmod z"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mod] + delsimps [realpow_Suc])); +qed "complex_Re_le_cmod"; +Addsimps [complex_Re_le_cmod]; + +Goal "z ~= 0 ==> 0 < cmod z"; +by (cut_inst_tac [("x","z")] complex_mod_ge_zero 1); +by (dtac order_le_imp_less_or_eq 1); +by Auto_tac; +qed "complex_mod_gt_zero"; + + +(*---------------------------------------------------------------------------*) +(* a few more theorems *) +(*---------------------------------------------------------------------------*) + +Goal "cmod(x ^ n) = cmod(x) ^ n"; +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps [complex_mod_mult])); +qed "complex_mod_complexpow"; + +Goal "(-x::complex) ^ n = (if even n then (x ^ n) else -(x ^ n))"; +by (induct_tac "n" 1); +by Auto_tac; +qed "complexpow_minus"; + +Goal "inverse (-x) = - inverse (x::complex)"; +by (res_inst_tac [("z","x")] eq_Abs_complex 1); +by (asm_simp_tac (simpset() addsimps [complex_inverse,complex_minus, + realpow_num_two]) 1); +qed "complex_inverse_minus"; + +Goalw [complex_divide_def] "x / (1::complex) = x"; +by (Simp_tac 1); +qed "complex_divide_one"; +Addsimps [complex_divide_one]; + +Goal "cmod(inverse x) = inverse(cmod x)"; +by (complex_div_undefined_case_tac "x=0" 1); +by (res_inst_tac [("c1","cmod x")] (real_mult_left_cancel RS iffD1) 1); +by (auto_tac (claset(),simpset() addsimps [complex_mod_mult RS sym])); +qed "complex_mod_inverse"; + +Goalw [complex_divide_def,real_divide_def] + "cmod(x/y) = cmod(x)/(cmod y)"; +by (auto_tac (claset(),simpset() addsimps [complex_mod_mult, + complex_mod_inverse])); +qed "complex_mod_divide"; + +Goalw [complex_divide_def] + "inverse(x/y) = y/(x::complex)"; +by (auto_tac (claset(),simpset() addsimps [complex_inverse_distrib, + complex_mult_commute])); +qed "complex_inverse_divide"; +Addsimps [complex_inverse_divide]; + +Goal "((r::complex) * s) ^ n = (r ^ n) * (s ^ n)"; +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps complex_mult_ac)); +qed "complexpow_mult"; + +(*---------------------------------------------------------------------------*) +(* More exponentiation *) +(*---------------------------------------------------------------------------*) + +Goal "(0::complex) ^ (Suc n) = 0"; +by (Auto_tac); +qed "complexpow_zero"; +Addsimps [complexpow_zero]; + +Goal "r ~= (0::complex) --> r ^ n ~= 0"; +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult_not_zero])); +qed_spec_mp "complexpow_not_zero"; +Addsimps [complexpow_not_zero]; +AddIs [complexpow_not_zero]; + +Goal "r ^ n = (0::complex) ==> r = 0"; +by (blast_tac (claset() addIs [ccontr] + addDs [complexpow_not_zero]) 1); +qed "complexpow_zero_zero"; + +Goalw [i_def] "ii ^ 2 = -(1::complex)"; +by (auto_tac (claset(),simpset() addsimps + [complex_mult,complex_one_def,complex_minus,realpow_num_two])); +qed "complexpow_i_squared"; +Addsimps [complexpow_i_squared]; + +Goalw [i_def,complex_zero_def] "ii ~= 0"; +by Auto_tac; +qed "complex_i_not_zero"; +Addsimps [complex_i_not_zero]; + +Goal "x * y ~= (0::complex) ==> x ~= 0"; +by Auto_tac; +qed "complex_mult_eq_zero_cancel1"; + +Goal "x * y ~= 0 ==> y ~= (0::complex)"; +by Auto_tac; +qed "complex_mult_eq_zero_cancel2"; + +Goal "(x * y ~= 0) = (x ~= 0 & y ~= (0::complex))"; +by Auto_tac; +qed "complex_mult_not_eq_zero_iff"; +AddIffs [complex_mult_not_eq_zero_iff]; + +Goal "inverse ((r::complex) ^ n) = (inverse r) ^ n"; +by (induct_tac "n" 1); +by (auto_tac (claset(), simpset() addsimps [complex_inverse_distrib])); +qed "complexpow_inverse"; + +(*---------------------------------------------------------------------------*) +(* sgn *) +(*---------------------------------------------------------------------------*) + +Goalw [sgn_def] "sgn 0 = 0"; +by (Simp_tac 1); +qed "sgn_zero"; +Addsimps[sgn_zero]; + +Goalw [sgn_def] "sgn 1 = 1"; +by (Simp_tac 1); +qed "sgn_one"; +Addsimps [sgn_one]; + +Goalw [sgn_def] "sgn (-z) = - sgn(z)"; +by Auto_tac; +qed "sgn_minus"; + +Goalw [sgn_def] + "sgn z = z / complex_of_real (cmod z)"; +by (Simp_tac 1); +qed "sgn_eq"; + +Goal "EX x y. z = complex_of_real(x) + ii * complex_of_real(y)"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_of_real_def, + i_def,complex_mult,complex_add])); +qed "complex_split"; + +Goal "Re(complex_of_real(x) + ii * complex_of_real(y)) = x"; +by (auto_tac (claset(),simpset() addsimps [complex_of_real_def, + i_def,complex_mult,complex_add])); +qed "Re_complex_i"; +Addsimps [Re_complex_i]; + +Goal "Im(complex_of_real(x) + ii * complex_of_real(y)) = y"; +by (auto_tac (claset(),simpset() addsimps [complex_of_real_def, + i_def,complex_mult,complex_add])); +qed "Im_complex_i"; +Addsimps [Im_complex_i]; + +Goalw [i_def,complex_of_real_def] "ii * ii = complex_of_real (-1)"; +by (auto_tac (claset(),simpset() addsimps [complex_mult,complex_add])); +qed "i_mult_eq"; + +Goalw [i_def,complex_one_def] "ii * ii = -(1::complex)"; +by (simp_tac (simpset() addsimps [complex_mult,complex_minus]) 1); +qed "i_mult_eq2"; +Addsimps [i_mult_eq2]; + +Goal "cmod (complex_of_real(x) + ii * complex_of_real(y)) = \ +\ sqrt (x ^ 2 + y ^ 2)"; +by (auto_tac (claset(),simpset() addsimps [complex_mult,complex_add, + i_def,complex_of_real_def,cmod_def])); +qed "cmod_i"; + +Goalw [complex_of_real_def,i_def] + "complex_of_real xa + ii * complex_of_real ya = \ +\ complex_of_real xb + ii * complex_of_real yb \ +\ ==> xa = xb"; +by (auto_tac (claset(),simpset() addsimps [complex_mult,complex_add])); +qed "complex_eq_Re_eq"; + +Goalw [complex_of_real_def,i_def] + "complex_of_real xa + ii * complex_of_real ya = \ +\ complex_of_real xb + ii * complex_of_real yb \ +\ ==> ya = yb"; +by (auto_tac (claset(),simpset() addsimps [complex_mult,complex_add])); +qed "complex_eq_Im_eq"; + +Goal "(complex_of_real xa + ii * complex_of_real ya = \ +\ complex_of_real xb + ii * complex_of_real yb) = ((xa = xb) & (ya = yb))"; +by (auto_tac (claset() addIs [complex_eq_Im_eq,complex_eq_Re_eq],simpset())); +qed "complex_eq_cancel_iff"; +AddIffs [complex_eq_cancel_iff]; + +Goal "(complex_of_real xa + complex_of_real ya * ii = \ +\ complex_of_real xb + complex_of_real yb * ii ) = ((xa = xb) & (ya = yb))"; +by (auto_tac (claset(),simpset() addsimps [complex_mult_commute])); +qed "complex_eq_cancel_iffA"; +AddIffs [complex_eq_cancel_iffA]; + +Goal "(complex_of_real xa + complex_of_real ya * ii = \ +\ complex_of_real xb + ii * complex_of_real yb) = ((xa = xb) & (ya = yb))"; +by (auto_tac (claset(),simpset() addsimps [complex_mult_commute])); +qed "complex_eq_cancel_iffB"; +AddIffs [complex_eq_cancel_iffB]; + +Goal "(complex_of_real xa + ii * complex_of_real ya = \ +\ complex_of_real xb + complex_of_real yb * ii) = ((xa = xb) & (ya = yb))"; +by (auto_tac (claset(),simpset() addsimps [complex_mult_commute])); +qed "complex_eq_cancel_iffC"; +AddIffs [complex_eq_cancel_iffC]; + +Goal"(complex_of_real x + ii * complex_of_real y = \ +\ complex_of_real xa) = (x = xa & y = 0)"; +by (cut_inst_tac [("xa","x"),("ya","y"),("xb","xa"),("yb","0")] + complex_eq_cancel_iff 1); +by (asm_full_simp_tac (simpset() delsimps [complex_eq_cancel_iff]) 1); +qed "complex_eq_cancel_iff2"; +Addsimps [complex_eq_cancel_iff2]; + +Goal"(complex_of_real x + complex_of_real y * ii = \ +\ complex_of_real xa) = (x = xa & y = 0)"; +by (auto_tac (claset(),simpset() addsimps [complex_mult_commute])); +qed "complex_eq_cancel_iff2a"; +Addsimps [complex_eq_cancel_iff2a]; + +Goal "(complex_of_real x + ii * complex_of_real y = \ +\ ii * complex_of_real ya) = (x = 0 & y = ya)"; +by (cut_inst_tac [("xa","x"),("ya","y"),("xb","0"),("yb","ya")] + complex_eq_cancel_iff 1); +by (asm_full_simp_tac (simpset() delsimps [complex_eq_cancel_iff]) 1); +qed "complex_eq_cancel_iff3"; +Addsimps [complex_eq_cancel_iff3]; + +Goal "(complex_of_real x + complex_of_real y * ii = \ +\ ii * complex_of_real ya) = (x = 0 & y = ya)"; +by (auto_tac (claset(),simpset() addsimps [complex_mult_commute])); +qed "complex_eq_cancel_iff3a"; +Addsimps [complex_eq_cancel_iff3a]; + +Goalw [complex_of_real_def,i_def,complex_zero_def] + "complex_of_real x + ii * complex_of_real y = 0 \ +\ ==> x = 0"; +by (auto_tac (claset(),simpset() addsimps [complex_mult,complex_add])); +qed "complex_split_Re_zero"; + +Goalw [complex_of_real_def,i_def,complex_zero_def] + "complex_of_real x + ii * complex_of_real y = 0 \ +\ ==> y = 0"; +by (auto_tac (claset(),simpset() addsimps [complex_mult,complex_add])); +qed "complex_split_Im_zero"; + +Goalw [sgn_def,complex_divide_def] + "Re(sgn z) = Re(z)/cmod z"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_of_real_inverse RS sym])); +by (auto_tac (claset(),simpset() addsimps [complex_of_real_def, + complex_mult,real_divide_def])); +qed "Re_sgn"; +Addsimps [Re_sgn]; + +Goalw [sgn_def,complex_divide_def] + "Im(sgn z) = Im(z)/cmod z"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_of_real_inverse RS sym])); +by (auto_tac (claset(),simpset() addsimps [complex_of_real_def, + complex_mult,real_divide_def])); +qed "Im_sgn"; +Addsimps [Im_sgn]; + +Goalw [complex_of_real_def,i_def] + "inverse(complex_of_real x + ii * complex_of_real y) = \ +\ complex_of_real(x/(x ^ 2 + y ^ 2)) - \ +\ ii * complex_of_real(y/(x ^ 2 + y ^ 2))"; +by (auto_tac (claset(),simpset() addsimps [complex_mult,complex_add, + complex_diff_def,complex_minus,complex_inverse,real_divide_def])); +qed "complex_inverse_complex_split"; + +(*----------------------------------------------------------------------------*) +(* Many of the theorems below need to be moved elsewhere e.g. Transc.ML. Also *) +(* many of the theorems are not used - so should they be kept? *) +(*----------------------------------------------------------------------------*) + +Goalw [i_def,complex_of_real_def] + "Re (ii * complex_of_real y) = 0"; +by (auto_tac (claset(),simpset() addsimps [complex_mult])); +qed "Re_mult_i_eq"; +Addsimps [Re_mult_i_eq]; + +Goalw [i_def,complex_of_real_def] + "Im (ii * complex_of_real y) = y"; +by (auto_tac (claset(),simpset() addsimps [complex_mult])); +qed "Im_mult_i_eq"; +Addsimps [Im_mult_i_eq]; + +Goalw [i_def,complex_of_real_def] + "cmod (ii * complex_of_real y) = abs y"; +by (auto_tac (claset(),simpset() addsimps [complex_mult, + complex_mod,realpow_num_two])); +qed "complex_mod_mult_i"; +Addsimps [complex_mod_mult_i]; + +Goalw [arg_def] + "0 < y ==> cos (arg(ii * complex_of_real y)) = 0"; +by (auto_tac (claset(),simpset() addsimps [abs_eqI2])); +by (res_inst_tac [("a","pi/2")] someI2 1); +by Auto_tac; +by (res_inst_tac [("R2.0","0")] real_less_trans 1); +by Auto_tac; +qed "cos_arg_i_mult_zero"; +Addsimps [cos_arg_i_mult_zero]; + +Goalw [arg_def] + "y < 0 ==> cos (arg(ii * complex_of_real y)) = 0"; +by (auto_tac (claset(),simpset() addsimps [abs_minus_eqI2])); +by (res_inst_tac [("a","- pi/2")] someI2 1); +by Auto_tac; +by (res_inst_tac [("j","0")] real_le_trans 1); +by Auto_tac; +qed "cos_arg_i_mult_zero2"; +Addsimps [cos_arg_i_mult_zero2]; + +Goalw [complex_zero_def,complex_of_real_def] + "(complex_of_real y ~= 0) = (y ~= 0)"; +by Auto_tac; +qed "complex_of_real_not_zero_iff"; +Addsimps [complex_of_real_not_zero_iff]; + +Goal "(complex_of_real y = 0) = (y = 0)"; +by Auto_tac; +by (rtac ccontr 1 THEN dtac (complex_of_real_not_zero_iff RS iffD2) 1); +by (Asm_full_simp_tac 1); +qed "complex_of_real_zero_iff"; +Addsimps [complex_of_real_zero_iff]; + +Goal "y ~= 0 ==> cos (arg(ii * complex_of_real y)) = 0"; +by (cut_inst_tac [("R1.0","y"),("R2.0","0")] real_linear 1); +by Auto_tac; +qed "cos_arg_i_mult_zero3"; +Addsimps [cos_arg_i_mult_zero3]; + +(*---------------------------------------------------------------------------*) +(* Finally! Polar form for complex numbers *) +(*---------------------------------------------------------------------------*) + +Goal "EX r a. z = complex_of_real r * \ +\ (complex_of_real(cos a) + ii * complex_of_real(sin a))"; +by (cut_inst_tac [("z","z")] complex_split 1); +by (auto_tac (claset(),simpset() addsimps [polar_Ex, + complex_add_mult_distrib2,complex_of_real_mult] @ complex_mult_ac)); +qed "complex_split_polar"; + +Goalw [rcis_def,cis_def] "EX r a. z = rcis r a"; +by (rtac complex_split_polar 1); +qed "rcis_Ex"; + +Goal "Re(complex_of_real r * \ +\ (complex_of_real(cos a) + ii * complex_of_real(sin a))) = r * cos a"; +by (auto_tac (claset(),simpset() addsimps [complex_add_mult_distrib2, + complex_of_real_mult] @ complex_mult_ac)); +qed "Re_complex_polar"; +Addsimps [Re_complex_polar]; + +Goalw [rcis_def,cis_def] "Re(rcis r a) = r * cos a"; +by Auto_tac; +qed "Re_rcis"; +Addsimps [Re_rcis]; + +Goal "Im(complex_of_real r * \ +\ (complex_of_real(cos a) + ii * complex_of_real(sin a))) = r * sin a"; +by (auto_tac (claset(),simpset() addsimps [complex_add_mult_distrib2, + complex_of_real_mult] @ complex_mult_ac)); +qed "Im_complex_polar"; +Addsimps [Im_complex_polar]; + +Goalw [rcis_def,cis_def] "Im(rcis r a) = r * sin a"; +by Auto_tac; +qed "Im_rcis"; +Addsimps [Im_rcis]; + +Goal "cmod (complex_of_real r * \ +\ (complex_of_real(cos a) + ii * complex_of_real(sin a))) = abs r"; +by (auto_tac (claset(),simpset() addsimps [complex_add_mult_distrib2, + cmod_i,complex_of_real_mult,real_add_mult_distrib2 + RS sym,realpow_mult] @ complex_mult_ac@ real_mult_ac + delsimps [realpow_Suc])); +qed "complex_mod_complex_polar"; +Addsimps [complex_mod_complex_polar]; + +Goalw [rcis_def,cis_def] "cmod(rcis r a) = abs r"; +by Auto_tac; +qed "complex_mod_rcis"; +Addsimps [complex_mod_rcis]; + +Goalw [cmod_def] "cmod z = sqrt (Re (z * cnj z))"; +by (rtac (real_sqrt_eq_iff RS iffD2) 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult_cnj])); +qed "complex_mod_sqrt_Re_mult_cnj"; + +Goal "Re(cnj z) = Re z"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_cnj])); +qed "complex_Re_cnj"; +Addsimps [complex_Re_cnj]; + +Goal "Im(cnj z) = - Im z"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_cnj])); +qed "complex_Im_cnj"; +Addsimps [complex_Im_cnj]; + +Goal "Im (z * cnj z) = 0"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_cnj,complex_mult])); +qed "complex_In_mult_cnj_zero"; +Addsimps [complex_In_mult_cnj_zero]; + +Goal "[| Im w = 0; Im z = 0 |] ==> Re(w * z) = Re(w) * Re(z)"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (res_inst_tac [("z","w")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult])); +qed "complex_Re_mult"; + +Goalw [complex_of_real_def] "Re (z * complex_of_real c) = Re(z) * c"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult])); +qed "complex_Re_mult_complex_of_real"; +Addsimps [complex_Re_mult_complex_of_real]; + +Goalw [complex_of_real_def] "Im (z * complex_of_real c) = Im(z) * c"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult])); +qed "complex_Im_mult_complex_of_real"; +Addsimps [complex_Im_mult_complex_of_real]; + +Goalw [complex_of_real_def] "Re (complex_of_real c * z) = c * Re(z)"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult])); +qed "complex_Re_mult_complex_of_real2"; +Addsimps [complex_Re_mult_complex_of_real2]; + +Goalw [complex_of_real_def] "Im (complex_of_real c * z) = c * Im(z)"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult])); +qed "complex_Im_mult_complex_of_real2"; +Addsimps [complex_Im_mult_complex_of_real2]; + +(*---------------------------------------------------------------------------*) +(* (r1 * cis a) * (r2 * cis b) = r1 * r2 * cis (a + b) *) +(*---------------------------------------------------------------------------*) + +Goalw [rcis_def] "cis a = rcis 1 a"; +by (Simp_tac 1); +qed "cis_rcis_eq"; + +Goalw [rcis_def,cis_def] + "rcis r1 a * rcis r2 b = rcis (r1*r2) (a + b)"; +by (auto_tac (claset(),simpset() addsimps [cos_add,sin_add, + complex_add_mult_distrib2,complex_add_mult_distrib] + @ complex_mult_ac @ complex_add_ac)); +by (auto_tac (claset(),simpset() addsimps [complex_add_mult_distrib2 RS sym, + complex_mult_assoc RS sym,complex_of_real_mult,complex_of_real_add, + complex_add_assoc RS sym,i_mult_eq] delsimps [i_mult_eq2])); +by (auto_tac (claset(),simpset() addsimps complex_add_ac)); +by (auto_tac (claset(),simpset() addsimps [complex_add_assoc RS sym, + complex_of_real_add,real_add_mult_distrib2, + real_diff_def] @ real_mult_ac @ real_add_ac)); +qed "rcis_mult"; + +Goal "cis a * cis b = cis (a + b)"; +by (simp_tac (simpset() addsimps [cis_rcis_eq,rcis_mult]) 1); +qed "cis_mult"; + +Goalw [cis_def] "cis 0 = 1"; +by Auto_tac; +qed "cis_zero"; +Addsimps [cis_zero]; + +Goalw [cis_def] "cis 0 = complex_of_real 1"; +by Auto_tac; +qed "cis_zero2"; +Addsimps [cis_zero2]; + +Goalw [rcis_def] "rcis 0 a = 0"; +by (Simp_tac 1); +qed "rcis_zero_mod"; +Addsimps [rcis_zero_mod]; + +Goalw [rcis_def] "rcis r 0 = complex_of_real r"; +by (Simp_tac 1); +qed "rcis_zero_arg"; +Addsimps [rcis_zero_arg]; + +Goalw [complex_of_real_def,complex_one_def] + "complex_of_real (-(1::real)) = -(1::complex)"; +by (simp_tac (simpset() addsimps [complex_minus]) 1); +qed "complex_of_real_minus_one"; + +Goal "ii * (ii * x) = - x"; +by (simp_tac (simpset() addsimps [complex_mult_assoc RS sym]) 1); +qed "complex_i_mult_minus"; +Addsimps [complex_i_mult_minus]; + +Goal "ii * ii * x = - x"; +by (Simp_tac 1); +qed "complex_i_mult_minus2"; +Addsimps [complex_i_mult_minus2]; + +Goalw [cis_def] + "cis (real (Suc n) * a) = cis a * cis (real n * a)"; +by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc, + real_add_mult_distrib,cos_add,sin_add,complex_add_mult_distrib, + complex_add_mult_distrib2,complex_of_real_add,complex_of_real_mult] + @ complex_mult_ac @ complex_add_ac)); +by (auto_tac (claset(),simpset() addsimps [complex_add_mult_distrib2 RS sym, + complex_mult_assoc RS sym,i_mult_eq,complex_of_real_mult, + complex_of_real_add,complex_add_assoc RS sym,complex_of_real_minus + RS sym,real_diff_def] @ real_mult_ac delsimps [i_mult_eq2])); +qed "cis_real_of_nat_Suc_mult"; + +Goal "(cis a) ^ n = cis (real n * a)"; +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps [cis_real_of_nat_Suc_mult])); +qed "DeMoivre"; + +Goalw [rcis_def] + "(rcis r a) ^ n = rcis (r ^ n) (real n * a)"; +by (auto_tac (claset(),simpset() addsimps [complexpow_mult, + DeMoivre,complex_of_real_pow])); +qed "DeMoivre2"; + +Goalw [cis_def] "inverse(cis a) = cis (-a)"; +by (auto_tac (claset(),simpset() addsimps [complex_inverse_complex_split, + complex_of_real_minus,complex_diff_def])); +qed "cis_inverse"; +Addsimps [cis_inverse]; + +Goal "inverse(rcis r a) = rcis (1/r) (-a)"; +by (real_div_undefined_case_tac "r=0" 1); +by (simp_tac (simpset() addsimps [rename_numerals DIVISION_BY_ZERO, + COMPLEX_INVERSE_ZERO]) 1); +by (auto_tac (claset(),simpset() addsimps [complex_inverse_complex_split, + complex_add_mult_distrib2,complex_of_real_mult,rcis_def,cis_def, + realpow_num_two] @ complex_mult_ac @ real_mult_ac)); +by (auto_tac (claset(),simpset() addsimps [real_add_mult_distrib2 RS sym, + complex_of_real_minus,complex_diff_def])); +qed "rcis_inverse"; + +Goalw [complex_divide_def] "cis a / cis b = cis (a - b)"; +by (auto_tac (claset(),simpset() addsimps [cis_mult,real_diff_def])); +qed "cis_divide"; + +Goalw [complex_divide_def] + "rcis r1 a / rcis r2 b = rcis (r1/r2) (a - b)"; +by (real_div_undefined_case_tac "r2=0" 1); +by (simp_tac (simpset() addsimps [rename_numerals DIVISION_BY_ZERO, + COMPLEX_INVERSE_ZERO]) 1); +by (auto_tac (claset(),simpset() addsimps [rcis_inverse,rcis_mult, + real_diff_def])); +qed "rcis_divide"; + +Goalw [cis_def] "Re(cis a) = cos a"; +by Auto_tac; +qed "Re_cis"; +Addsimps [Re_cis]; + +Goalw [cis_def] "Im(cis a) = sin a"; +by Auto_tac; +qed "Im_cis"; +Addsimps [Im_cis]; + +Goal "cos (real n * a) = Re(cis a ^ n)"; +by (auto_tac (claset(),simpset() addsimps [DeMoivre])); +qed "cos_n_Re_cis_pow_n"; + +Goal "sin (real n * a) = Im(cis a ^ n)"; +by (auto_tac (claset(),simpset() addsimps [DeMoivre])); +qed "sin_n_Im_cis_pow_n"; + +Goalw [expi_def,cis_def] + "expi (ii * complex_of_real y) = \ +\ complex_of_real (cos y) + ii * complex_of_real (sin y)"; +by Auto_tac; +qed "expi_Im_split"; + +Goalw [expi_def] + "expi (ii * complex_of_real y) = cis y"; +by Auto_tac; +qed "expi_Im_cis"; + +Goalw [expi_def] "expi(a + b) = expi(a) * expi(b)"; +by (auto_tac (claset(),simpset() addsimps [complex_Re_add,exp_add, + complex_Im_add,cis_mult RS sym,complex_of_real_mult] @ + complex_mult_ac)); +qed "expi_add"; + +Goalw [expi_def] + "expi(complex_of_real x + ii * complex_of_real y) = \ +\ complex_of_real (exp(x)) * cis y"; +by Auto_tac; +qed "expi_complex_split"; + +Goalw [expi_def] "expi (0::complex) = 1"; +by Auto_tac; +qed "expi_zero"; +Addsimps [expi_zero]; + +goal Complex.thy + "Re (w * z) = Re w * Re z - Im w * Im z"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (res_inst_tac [("z","w")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult])); +qed "complex_Re_mult_eq"; + +goal Complex.thy + "Im (w * z) = Re w * Im z + Im w * Re z"; +by (res_inst_tac [("z","z")] eq_Abs_complex 1); +by (res_inst_tac [("z","w")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mult])); +qed "complex_Im_mult_eq"; + +goal Complex.thy + "EX a r. z = complex_of_real r * expi a"; +by (cut_inst_tac [("z","z")] rcis_Ex 1); +by (auto_tac (claset(),simpset() addsimps [expi_def,rcis_def, + complex_mult_assoc RS sym,complex_of_real_mult])); +by (res_inst_tac [("x","ii * complex_of_real a")] exI 1); +by Auto_tac; +qed "complex_expi_Ex"; + + +(**** +Goal "[| - pi < a; a <= pi |] ==> (-pi < a & a <= 0) | (0 <= a & a <= pi)"; +by Auto_tac; +qed "lemma_split_interval"; + +Goalw [arg_def] + "[| r ~= 0; - pi < a; a <= pi |] \ +\ ==> arg(complex_of_real r * \ +\ (complex_of_real(cos a) + ii * complex_of_real(sin a))) = a"; +by Auto_tac; +by (cut_inst_tac [("R1.0","0"),("R2.0","r")] real_linear 1); +by (auto_tac (claset(),simpset() addsimps (map (full_rename_numerals thy) + [rabs_eqI2,rabs_minus_eqI2,real_minus_rinv]) @ [real_divide_def, + real_minus_mult_eq2 RS sym] @ real_mult_ac)); +by (auto_tac (claset(),simpset() addsimps [real_mult_assoc RS sym])); +by (dtac lemma_split_interval 1 THEN Step_tac 1); +****) + diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/Complex.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/Complex.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,111 @@ +(* Title: Complex.thy + Author: Jacques D. Fleuriot + Copyright: 2001 University of Edinburgh + Description: Complex numbers +*) + +Complex = HLog + + +typedef complex = "{p::(real*real). True}" + +instance + complex :: {ord,zero,one,plus,minus,times,power,inverse} + +consts + "ii" :: complex ("ii") + +constdefs + + (*--- real and Imaginary parts ---*) + + Re :: complex => real + "Re(z) == fst(Rep_complex z)" + + Im :: complex => real + "Im(z) == snd(Rep_complex z)" + + (*----------- modulus ------------*) + + cmod :: complex => real + "cmod z == sqrt(Re(z) ^ 2 + Im(z) ^ 2)" + + (*----- injection from reals -----*) + + complex_of_real :: real => complex + "complex_of_real r == Abs_complex(r,0::real)" + + (*------- complex conjugate ------*) + + cnj :: complex => complex + "cnj z == Abs_complex(Re z, -Im z)" + + (*------------ Argand -------------*) + + sgn :: complex => complex + "sgn z == z / complex_of_real(cmod z)" + + arg :: complex => real + "arg z == @a. Re(sgn z) = cos a & Im(sgn z) = sin a & -pi < a & a <= pi" + +defs + + complex_zero_def + "0 == Abs_complex(0::real,0)" + + complex_one_def + "1 == Abs_complex(1,0::real)" + + (*------ imaginary unit ----------*) + + i_def + "ii == Abs_complex(0::real,1)" + + (*----------- negation -----------*) + + complex_minus_def + "- (z::complex) == Abs_complex(-Re z, -Im z)" + + + (*----------- inverse -----------*) + complex_inverse_def + "inverse (z::complex) == Abs_complex(Re(z)/(Re(z) ^ 2 + Im(z) ^ 2), + -Im(z)/(Re(z) ^ 2 + Im(z) ^ 2))" + + complex_add_def + "w + (z::complex) == Abs_complex(Re(w) + Re(z),Im(w) + Im(z))" + + complex_diff_def + "w - (z::complex) == w + -(z::complex)" + + complex_mult_def + "w * (z::complex) == Abs_complex(Re(w) * Re(z) - Im(w) * Im(z), + Re(w) * Im(z) + Im(w) * Re(z))" + + + (*----------- division ----------*) + complex_divide_def + "w / (z::complex) == w * inverse z" + + +primrec + complexpow_0 "z ^ 0 = complex_of_real 1" + complexpow_Suc "z ^ (Suc n) = (z::complex) * (z ^ n)" + + +constdefs + + (* abbreviation for (cos a + i sin a) *) + cis :: real => complex + "cis a == complex_of_real(cos a) + ii * complex_of_real(sin a)" + + (* abbreviation for r*(cos a + i sin a) *) + rcis :: [real, real] => complex + "rcis r a == complex_of_real r * cis a" + + (* e ^ (x + iy) *) + expi :: complex => complex + "expi z == complex_of_real(exp (Re z)) * cis (Im z)" + +end + + diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/ComplexArith0.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/ComplexArith0.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,321 @@ +(* Title: ComplexArith0.ML + Author: Jacques D. Fleuriot + Copyright: 2001 University of Edinburgh + Description: Assorted facts that need binary literals + Also, common factor cancellation (see e.g. HyperArith0) +*) + +Goal "x - - y = x + (y::complex)"; +by (Simp_tac 1); +qed "real_diff_minus_eq"; +Addsimps [real_diff_minus_eq]; + +(** Division and inverse **) + +Goal "0/x = (0::complex)"; +by (simp_tac (simpset() addsimps [complex_divide_def]) 1); +qed "complex_0_divide"; +Addsimps [complex_0_divide]; + +Goalw [complex_divide_def] "x/(0::complex) = 0"; +by (stac COMPLEX_INVERSE_ZERO 1); +by (Simp_tac 1); +qed "COMPLEX_DIVIDE_ZERO"; + +Goal "inverse (x::complex) = 1/x"; +by (simp_tac (simpset() addsimps [complex_divide_def]) 1); +qed "complex_inverse_eq_divide"; + +Goal "(inverse(x::complex) = 0) = (x = 0)"; +by (auto_tac (claset(), + simpset() addsimps [COMPLEX_INVERSE_ZERO])); +by (rtac ccontr 1); +by (blast_tac (claset() addDs [complex_inverse_not_zero]) 1); +qed "complex_inverse_zero_iff"; +Addsimps [complex_inverse_zero_iff]; + +Goal "(x/y = 0) = (x=0 | y=(0::complex))"; +by (auto_tac (claset(), simpset() addsimps [complex_divide_def])); +qed "complex_divide_eq_0_iff"; +Addsimps [complex_divide_eq_0_iff]; + +Goal "h ~= (0::complex) ==> h/h = 1"; +by (asm_simp_tac + (simpset() addsimps [complex_divide_def]) 1); +qed "complex_divide_self_eq"; +Addsimps [complex_divide_self_eq]; + +bind_thm ("complex_mult_minus_right", complex_minus_mult_eq2 RS sym); + +Goal "!!k::complex. (k*m = k*n) = (k = 0 | m=n)"; +by (case_tac "k=0" 1); +by (auto_tac (claset(), simpset() addsimps [complex_mult_left_cancel])); +qed "complex_mult_eq_cancel1"; + +Goal "!!k::complex. (m*k = n*k) = (k = 0 | m=n)"; +by (case_tac "k=0" 1); +by (auto_tac (claset(), simpset() addsimps [complex_mult_right_cancel])); +qed "complex_mult_eq_cancel2"; + +Goal "!!k::complex. k~=0 ==> (k*m) / (k*n) = (m/n)"; +by (asm_simp_tac + (simpset() addsimps [complex_divide_def, complex_inverse_distrib]) 1); +by (subgoal_tac "k * m * (inverse k * inverse n) = \ +\ (k * inverse k) * (m * inverse n)" 1); +by (Asm_full_simp_tac 1); +by (asm_full_simp_tac (HOL_ss addsimps complex_mult_ac) 1); +qed "complex_mult_div_cancel1"; + +(*For ExtractCommonTerm*) +Goal "(k*m) / (k*n) = (if k = (0::complex) then 0 else m/n)"; +by (simp_tac (simpset() addsimps [complex_mult_div_cancel1]) 1); +qed "complex_mult_div_cancel_disj"; + + +local + open Complex_Numeral_Simprocs +in + +val rel_complex_number_of = [eq_complex_number_of]; + + +structure CancelNumeralFactorCommon = + struct + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff 1 + val trans_tac = Real_Numeral_Simprocs.trans_tac + val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps complex_minus_from_mult_simps @ mult_1s)) + THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@complex_mult_minus_simps)) + THEN ALLGOALS (simp_tac (HOL_ss addsimps complex_mult_ac)) + val numeral_simp_tac = ALLGOALS (simp_tac (HOL_ss addsimps rel_complex_number_of@bin_simps)) + val simplify_meta_eq = simplify_meta_eq + end + + +structure DivCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Bin_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binop "HOL.divide" + val dest_bal = HOLogic.dest_bin "HOL.divide" complexT + val cancel = complex_mult_div_cancel1 RS trans + val neg_exchanges = false +) + + +structure EqCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Bin_Simprocs.prove_conv + val mk_bal = HOLogic.mk_eq + val dest_bal = HOLogic.dest_bin "op =" complexT + val cancel = complex_mult_eq_cancel1 RS trans + val neg_exchanges = false +) + +val complex_cancel_numeral_factors_relations = + map prep_simproc + [("complexeq_cancel_numeral_factor", + ["(l::complex) * m = n", "(l::complex) = m * n"], + EqCancelNumeralFactor.proc)]; + +val complex_cancel_numeral_factors_divide = prep_simproc + ("complexdiv_cancel_numeral_factor", + ["((l::complex) * m) / n", "(l::complex) / (m * n)", + "((number_of v)::complex) / (number_of w)"], + DivCancelNumeralFactor.proc); + +val complex_cancel_numeral_factors = + complex_cancel_numeral_factors_relations @ + [complex_cancel_numeral_factors_divide]; + +end; + + +Addsimprocs complex_cancel_numeral_factors; + + +(*examples: +print_depth 22; +set timing; +set trace_simp; +fun test s = (Goal s; by (Simp_tac 1)); + + +test "9*x = 12 * (y::complex)"; +test "(9*x) / (12 * (y::complex)) = z"; + +test "-99*x = 132 * (y::complex)"; + +test "999*x = -396 * (y::complex)"; +test "(999*x) / (-396 * (y::complex)) = z"; + +test "-99*x = -81 * (y::complex)"; +test "(-99*x) / (-81 * (y::complex)) = z"; + +test "-2 * x = -1 * (y::complex)"; +test "-2 * x = -(y::complex)"; +test "(-2 * x) / (-1 * (y::complex)) = z"; + +*) + + +(** Declarations for ExtractCommonTerm **) + +local + open Complex_Numeral_Simprocs +in + +structure CancelFactorCommon = + struct + val mk_sum = long_mk_prod + val dest_sum = dest_prod + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff + val find_first = find_first [] + val trans_tac = Real_Numeral_Simprocs.trans_tac + val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@complex_mult_ac)) + end; + + +structure EqCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Bin_Simprocs.prove_conv + val mk_bal = HOLogic.mk_eq + val dest_bal = HOLogic.dest_bin "op =" complexT + val simplify_meta_eq = cancel_simplify_meta_eq complex_mult_eq_cancel1 +); + + +structure DivideCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Bin_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binop "HOL.divide" + val dest_bal = HOLogic.dest_bin "HOL.divide" complexT + val simplify_meta_eq = cancel_simplify_meta_eq complex_mult_div_cancel_disj +); + +val complex_cancel_factor = + map prep_simproc + [("complex_eq_cancel_factor", ["(l::complex) * m = n", "(l::complex) = m * n"], + EqCancelFactor.proc), + ("complex_divide_cancel_factor", ["((l::complex) * m) / n", "(l::complex) / (m * n)"], + DivideCancelFactor.proc)]; + +end; + +Addsimprocs complex_cancel_factor; + + +(*examples: +print_depth 22; +set timing; +set trace_simp; +fun test s = (Goal s; by (Asm_simp_tac 1)); + +test "x*k = k*(y::complex)"; +test "k = k*(y::complex)"; +test "a*(b*c) = (b::complex)"; +test "a*(b*c) = d*(b::complex)*(x*a)"; + + +test "(x*k) / (k*(y::complex)) = (uu::complex)"; +test "(k) / (k*(y::complex)) = (uu::complex)"; +test "(a*(b*c)) / ((b::complex)) = (uu::complex)"; +test "(a*(b*c)) / (d*(b::complex)*(x*a)) = (uu::complex)"; + +(*FIXME: what do we do about this?*) +test "a*(b*c)/(y*z) = d*(b::complex)*(x*a)/z"; +*) + + +Goal "z~=0 ==> ((x::complex) = y/z) = (x*z = y)"; +by (subgoal_tac "(x*z = y) = (x*z = (y/z)*z)" 1); +by (asm_simp_tac (simpset() addsimps [complex_divide_def, complex_mult_assoc]) 2); +by (etac ssubst 1); +by (stac complex_mult_eq_cancel2 1); +by (Asm_simp_tac 1); +qed "complex_eq_divide_eq"; +Addsimps [inst "z" "number_of ?w" complex_eq_divide_eq]; + +Goal "z~=0 ==> (y/z = (x::complex)) = (y = x*z)"; +by (subgoal_tac "(y = x*z) = ((y/z)*z = x*z)" 1); +by (asm_simp_tac (simpset() addsimps [complex_divide_def, complex_mult_assoc]) 2); +by (etac ssubst 1); +by (stac complex_mult_eq_cancel2 1); +by (Asm_simp_tac 1); +qed "complex_divide_eq_eq"; +Addsimps [inst "z" "number_of ?w" complex_divide_eq_eq]; + +Goal "(m/k = n/k) = (k = 0 | m = (n::complex))"; +by (case_tac "k=0" 1); +by (asm_simp_tac (simpset() addsimps [COMPLEX_DIVIDE_ZERO]) 1); +by (asm_simp_tac (simpset() addsimps [complex_divide_eq_eq, complex_eq_divide_eq, + complex_mult_eq_cancel2]) 1); +qed "complex_divide_eq_cancel2"; + +Goal "(k/m = k/n) = (k = 0 | m = (n::complex))"; +by (case_tac "m=0 | n = 0" 1); +by (auto_tac (claset(), + simpset() addsimps [COMPLEX_DIVIDE_ZERO, complex_divide_eq_eq, + complex_eq_divide_eq, complex_mult_eq_cancel1])); +qed "complex_divide_eq_cancel1"; + +(** Division by 1, -1 **) + +Goal "(x::complex)/1 = x"; +by (simp_tac (simpset() addsimps [complex_divide_def]) 1); +qed "complex_divide_1"; +Addsimps [complex_divide_1]; + +Goal "x/-1 = -(x::complex)"; +by (Simp_tac 1); +qed "complex_divide_minus1"; +Addsimps [complex_divide_minus1]; + +Goal "-1/(x::complex) = - (1/x)"; +by (simp_tac (simpset() addsimps [complex_divide_def, complex_minus_inverse]) 1); +qed "complex_minus1_divide"; +Addsimps [complex_minus1_divide]; + + +Goal "(x = - y) = (y = - (x::complex))"; +by Auto_tac; +qed "complex_equation_minus"; + +Goal "(- x = y) = (- (y::complex) = x)"; +by Auto_tac; +qed "complex_minus_equation"; + +Goal "(x + - a = (0::complex)) = (x=a)"; +by (simp_tac (simpset() addsimps [complex_diff_eq_eq,symmetric complex_diff_def]) 1); +qed "complex_add_minus_iff"; +Addsimps [complex_add_minus_iff]; + +Goal "(-b = -a) = (b = (a::complex))"; +by Auto_tac; +by (etac ( inj_complex_minus RS injD) 1); +qed "complex_minus_eq_cancel"; +Addsimps [complex_minus_eq_cancel]; + +(*Distributive laws for literals*) +Addsimps (map (inst "w" "number_of ?v") + [complex_add_mult_distrib, complex_add_mult_distrib2, + complex_diff_mult_distrib, complex_diff_mult_distrib2]); + +Addsimps [inst "x" "number_of ?v" complex_equation_minus]; + +Addsimps [inst "y" "number_of ?v" complex_minus_equation]; + +Goal "(x+y = (0::complex)) = (y = -x)"; +by Auto_tac; +by (dtac (sym RS (complex_diff_eq_eq RS iffD2)) 1); +by Auto_tac; +qed "complex_add_eq_0_iff"; +AddIffs [complex_add_eq_0_iff]; + +Goalw [complex_diff_def]"-(x-y) = y - (x::complex)"; +by (auto_tac (claset(),simpset() addsimps [complex_add_commute])); +qed "complex_minus_diff_eq"; +Addsimps [complex_minus_diff_eq]; + +Addsimps [inst "x" "number_of ?w" complex_inverse_eq_divide]; + diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/ComplexArith0.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/ComplexArith0.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,2 @@ +ComplexArith0 = ComplexBin + diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/ComplexBin.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/ComplexBin.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,609 @@ +(* Title: ComplexBin.ML + Author: Jacques D. Fleuriot + Copyright: 2001 University of Edinburgh + Descrition: Binary arithmetic for the complex numbers +*) + +(** complex_of_real (coercion from real to complex) **) + +Goal "complex_of_real (number_of w) = number_of w"; +by (simp_tac (simpset() addsimps [complex_number_of_def]) 1); +qed "complex_number_of"; +Addsimps [complex_number_of]; + +Goalw [complex_number_of_def] "Numeral0 = (0::complex)"; +by (Simp_tac 1); +qed "complex_numeral_0_eq_0"; + +Goalw [complex_number_of_def] "Numeral1 = (1::complex)"; +by (Simp_tac 1); +qed "complex_numeral_1_eq_1"; + +(** Addition **) + +Goal "(number_of v :: complex) + number_of v' = number_of (bin_add v v')"; +by (simp_tac + (HOL_ss addsimps [complex_number_of_def, + complex_of_real_add, add_real_number_of]) 1); +qed "add_complex_number_of"; +Addsimps [add_complex_number_of]; + + +(** Subtraction **) + +Goalw [complex_number_of_def] + "- (number_of w :: complex) = number_of (bin_minus w)"; +by (simp_tac + (HOL_ss addsimps [minus_real_number_of, complex_of_real_minus RS sym]) 1); +qed "minus_complex_number_of"; +Addsimps [minus_complex_number_of]; + +Goalw [complex_number_of_def, complex_diff_def] + "(number_of v :: complex) - number_of w = number_of (bin_add v (bin_minus w))"; +by (Simp_tac 1); +qed "diff_complex_number_of"; +Addsimps [diff_complex_number_of]; + + +(** Multiplication **) + +Goal "(number_of v :: complex) * number_of v' = number_of (bin_mult v v')"; +by (simp_tac + (HOL_ss addsimps [complex_number_of_def, + complex_of_real_mult, mult_real_number_of]) 1); +qed "mult_complex_number_of"; +Addsimps [mult_complex_number_of]; + +Goal "(2::complex) = 1 + 1"; +by (simp_tac (simpset() addsimps [complex_numeral_1_eq_1 RS sym]) 1); +val lemma = result(); + +(*For specialist use: NOT as default simprules*) +Goal "2 * z = (z+z::complex)"; +by (simp_tac (simpset () addsimps [lemma, complex_add_mult_distrib]) 1); +qed "complex_mult_2"; + +Goal "z * 2 = (z+z::complex)"; +by (stac complex_mult_commute 1 THEN rtac complex_mult_2 1); +qed "complex_mult_2_right"; + +(** Equals (=) **) + +Goal "((number_of v :: complex) = number_of v') = \ +\ iszero (number_of (bin_add v (bin_minus v')))"; +by (simp_tac + (HOL_ss addsimps [complex_number_of_def, + complex_of_real_eq_iff, eq_real_number_of]) 1); +qed "eq_complex_number_of"; +Addsimps [eq_complex_number_of]; + +(*** New versions of existing theorems involving 0, 1 ***) + +Goal "- 1 = (-1::complex)"; +by (simp_tac (simpset() addsimps [complex_numeral_1_eq_1 RS sym]) 1); +qed "complex_minus_1_eq_m1"; + +Goal "-1 * z = -(z::complex)"; +by (simp_tac (simpset() addsimps [complex_minus_1_eq_m1 RS sym]) 1); +qed "complex_mult_minus1"; + +Goal "z * -1 = -(z::complex)"; +by (stac complex_mult_commute 1 THEN rtac complex_mult_minus1 1); +qed "complex_mult_minus1_right"; + +Addsimps [complex_mult_minus1,complex_mult_minus1_right]; + + +(*Maps 0 to Numeral0 and 1 to Numeral1 and -Numeral1 to -1*) +val complex_numeral_ss = + hypreal_numeral_ss addsimps [complex_numeral_0_eq_0 RS sym, complex_numeral_1_eq_1 RS sym, + complex_minus_1_eq_m1]; + +fun rename_numerals th = + asm_full_simplify complex_numeral_ss (Thm.transfer (the_context ()) th); + +(*Now insert some identities previously stated for 0 and 1c*) + +Addsimps [complex_numeral_0_eq_0,complex_numeral_1_eq_1]; + +Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::complex)"; +by (auto_tac (claset(),simpset() addsimps [complex_add_assoc RS sym])); +qed "complex_add_number_of_left"; + +Goal "number_of v *(number_of w * z) = (number_of(bin_mult v w) * z::complex)"; +by (simp_tac (simpset() addsimps [complex_mult_assoc RS sym]) 1); +qed "complex_mult_number_of_left"; + +Goalw [complex_diff_def] + "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::complex)"; +by (rtac complex_add_number_of_left 1); +qed "complex_add_number_of_diff1"; + +Goal "number_of v + (c - number_of w) = \ +\ number_of (bin_add v (bin_minus w)) + (c::complex)"; +by (auto_tac (claset(),simpset() addsimps [complex_diff_def]@ complex_add_ac)); +qed "complex_add_number_of_diff2"; + +Addsimps [complex_add_number_of_left, complex_mult_number_of_left, + complex_add_number_of_diff1, complex_add_number_of_diff2]; + + +(**** Simprocs for numeric literals ****) + +(** Combining of literal coefficients in sums of products **) + +Goal "(x = y) = (x-y = (0::complex))"; +by (simp_tac (simpset() addsimps [complex_diff_eq_eq]) 1); +qed "complex_eq_iff_diff_eq_0"; + +(** For combine_numerals **) + +Goal "i*u + (j*u + k) = (i+j)*u + (k::complex)"; +by (asm_simp_tac (simpset() addsimps [complex_add_mult_distrib] + @ complex_add_ac) 1); +qed "left_complex_add_mult_distrib"; + +(** For cancel_numerals **) + +Goal "((x::complex) = u + v) = (x - (u + v) = 0)"; +by (auto_tac (claset(),simpset() addsimps [complex_diff_eq_eq])); +qed "complex_eq_add_diff_eq_0"; + +Goal "((x::complex) = n) = (x - n = 0)"; +by (auto_tac (claset(),simpset() addsimps [complex_diff_eq_eq])); +qed "complex_eq_diff_eq_0"; + +val complex_rel_iff_rel_0_rls = [complex_eq_diff_eq_0,complex_eq_add_diff_eq_0]; + +Goal "!!i::complex. (i*u + m = j*u + n) = ((i-j)*u + m = n)"; +by (auto_tac (claset(), simpset() addsimps [complex_add_mult_distrib, + complex_diff_def] @ complex_add_ac)); +by (asm_simp_tac (simpset() addsimps [complex_add_assoc RS sym]) 1); +by (simp_tac (simpset() addsimps [complex_add_assoc]) 1); +qed "complex_eq_add_iff1"; + +Goal "!!i::complex. (i*u + m = j*u + n) = (m = (j-i)*u + n)"; +by (simp_tac (simpset() addsimps [ complex_eq_add_iff1]) 1); +by (auto_tac (claset(), simpset() addsimps [complex_diff_def, + complex_add_mult_distrib]@ complex_add_ac)); +qed "complex_eq_add_iff2"; + +structure Complex_Numeral_Simprocs = +struct + +(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in simprocs + isn't complicated by the abstract 0 and 1.*) +val numeral_syms = [complex_numeral_0_eq_0 RS sym, complex_numeral_1_eq_1 RS sym]; + + +(*Utilities*) + +val complexT = Type("Complex.complex",[]); + +fun mk_numeral n = HOLogic.number_of_const complexT $ HOLogic.mk_bin n; + +val dest_numeral = Real_Numeral_Simprocs.dest_numeral; +val find_first_numeral = Real_Numeral_Simprocs.find_first_numeral; + +val zero = mk_numeral 0; +val mk_plus = HOLogic.mk_binop "op +"; + +val uminus_const = Const ("uminus", complexT --> complexT); + +(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) +fun mk_sum [] = zero + | mk_sum [t,u] = mk_plus (t, u) + | mk_sum (t :: ts) = mk_plus (t, mk_sum ts); + +(*this version ALWAYS includes a trailing zero*) +fun long_mk_sum [] = zero + | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts); + +val dest_plus = HOLogic.dest_bin "op +" complexT; + +(*decompose additions AND subtractions as a sum*) +fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) = + dest_summing (pos, t, dest_summing (pos, u, ts)) + | dest_summing (pos, Const ("op -", _) $ t $ u, ts) = + dest_summing (pos, t, dest_summing (not pos, u, ts)) + | dest_summing (pos, t, ts) = + if pos then t::ts else uminus_const$t :: ts; + +fun dest_sum t = dest_summing (true, t, []); + +val mk_diff = HOLogic.mk_binop "op -"; +val dest_diff = HOLogic.dest_bin "op -" complexT; + +val one = mk_numeral 1; +val mk_times = HOLogic.mk_binop "op *"; + +fun mk_prod [] = one + | mk_prod [t] = t + | mk_prod (t :: ts) = if t = one then mk_prod ts + else mk_times (t, mk_prod ts); + +val dest_times = HOLogic.dest_bin "op *" complexT; + +fun dest_prod t = + let val (t,u) = dest_times t + in dest_prod t @ dest_prod u end + handle TERM _ => [t]; + +(*DON'T do the obvious simplifications; that would create special cases*) +fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts); + +(*Express t as a product of (possibly) a numeral with other sorted terms*) +fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t + | dest_coeff sign t = + let val ts = sort Term.term_ord (dest_prod t) + val (n, ts') = find_first_numeral [] ts + handle TERM _ => (1, ts) + in (sign*n, mk_prod ts') end; + +(*Find first coefficient-term THAT MATCHES u*) +fun find_first_coeff past u [] = raise TERM("find_first_coeff", []) + | find_first_coeff past u (t::terms) = + let val (n,u') = dest_coeff 1 t + in if u aconv u' then (n, rev past @ terms) + else find_first_coeff (t::past) u terms + end + handle TERM _ => find_first_coeff (t::past) u terms; + + +(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*) +val add_0s = map rename_numerals [complex_add_zero_left, complex_add_zero_right]; +val mult_plus_1s = map rename_numerals [complex_mult_one_left, complex_mult_one_right]; +val mult_minus_1s = map rename_numerals + [complex_mult_minus1, complex_mult_minus1_right]; +val mult_1s = mult_plus_1s @ mult_minus_1s; + +(*To perform binary arithmetic*) +val bin_simps = + [complex_numeral_0_eq_0 RS sym, complex_numeral_1_eq_1 RS sym, + add_complex_number_of, complex_add_number_of_left, + minus_complex_number_of, diff_complex_number_of, mult_complex_number_of, + complex_mult_number_of_left] @ bin_arith_simps @ bin_rel_simps; + +(*To evaluate binary negations of coefficients*) +val complex_minus_simps = NCons_simps @ + [complex_minus_1_eq_m1,minus_complex_number_of, + bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min, + bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min]; + +(*To let us treat subtraction as addition*) +val diff_simps = [complex_diff_def, complex_minus_add_distrib, + complex_minus_minus]; + +(* push the unary minus down: - x * y = x * - y *) +val complex_minus_mult_eq_1_to_2 = + [complex_minus_mult_eq1 RS sym, complex_minus_mult_eq2] MRS trans + |> standard; + +(*to extract again any uncancelled minuses*) +val complex_minus_from_mult_simps = + [complex_minus_minus, complex_minus_mult_eq1 RS sym, + complex_minus_mult_eq2 RS sym]; + +(*combine unary minus with numeric literals, however nested within a product*) +val complex_mult_minus_simps = + [complex_mult_assoc, complex_minus_mult_eq1, complex_minus_mult_eq_1_to_2]; + +(*Final simplification: cancel + and * *) +val simplify_meta_eq = + Int_Numeral_Simprocs.simplify_meta_eq + [complex_add_zero_left, complex_add_zero_right, + complex_mult_zero_left, complex_mult_zero_right, complex_mult_one_left, + complex_mult_one_right]; + +val prep_simproc = Real_Numeral_Simprocs.prep_simproc; + + +structure CancelNumeralsCommon = + struct + val mk_sum = mk_sum + val dest_sum = dest_sum + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff 1 + val find_first_coeff = find_first_coeff [] + val trans_tac = Real_Numeral_Simprocs.trans_tac + val norm_tac = + ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@ + complex_minus_simps@complex_add_ac)) + THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@complex_mult_minus_simps)) + THEN ALLGOALS + (simp_tac (HOL_ss addsimps complex_minus_from_mult_simps@ + complex_add_ac@complex_mult_ac)) + val numeral_simp_tac = ALLGOALS (simp_tac (HOL_ss addsimps add_0s@bin_simps)) + val simplify_meta_eq = simplify_meta_eq + end; + + +structure EqCancelNumerals = CancelNumeralsFun + (open CancelNumeralsCommon + val prove_conv = Bin_Simprocs.prove_conv + val mk_bal = HOLogic.mk_eq + val dest_bal = HOLogic.dest_bin "op =" complexT + val bal_add1 = complex_eq_add_iff1 RS trans + val bal_add2 = complex_eq_add_iff2 RS trans +); + + +val cancel_numerals = + map prep_simproc + [("complexeq_cancel_numerals", + ["(l::complex) + m = n", "(l::complex) = m + n", + "(l::complex) - m = n", "(l::complex) = m - n", + "(l::complex) * m = n", "(l::complex) = m * n"], + EqCancelNumerals.proc)]; + +structure CombineNumeralsData = + struct + val add = op + : int*int -> int + val mk_sum = long_mk_sum (*to work for e.g. #2*x + #3*x *) + val dest_sum = dest_sum + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff 1 + val left_distrib = left_complex_add_mult_distrib RS trans + val prove_conv = Bin_Simprocs.prove_conv_nohyps + val trans_tac = Real_Numeral_Simprocs.trans_tac + val norm_tac = + ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@ + complex_minus_simps@complex_add_ac)) + THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@complex_mult_minus_simps)) + THEN ALLGOALS (simp_tac (HOL_ss addsimps complex_minus_from_mult_simps@ + complex_add_ac@complex_mult_ac)) + val numeral_simp_tac = ALLGOALS + (simp_tac (HOL_ss addsimps add_0s@bin_simps)) + val simplify_meta_eq = simplify_meta_eq + end; + +structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); + +val combine_numerals = + prep_simproc ("complex_combine_numerals", + ["(i::complex) + j", "(i::complex) - j"], + CombineNumerals.proc); + + +(** Declarations for ExtractCommonTerm **) + +(*this version ALWAYS includes a trailing one*) +fun long_mk_prod [] = one + | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts); + +(*Find first term that matches u*) +fun find_first past u [] = raise TERM("find_first", []) + | find_first past u (t::terms) = + if u aconv t then (rev past @ terms) + else find_first (t::past) u terms + handle TERM _ => find_first (t::past) u terms; + +(*Final simplification: cancel + and * *) +fun cancel_simplify_meta_eq cancel_th th = + Int_Numeral_Simprocs.simplify_meta_eq + [complex_mult_one_left, complex_mult_one_right] + (([th, cancel_th]) MRS trans); + +(*** Making constant folding work for 0 and 1 too ***) + +structure ComplexAbstractNumeralsData = + struct + val dest_eq = HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of + val is_numeral = Bin_Simprocs.is_numeral + val numeral_0_eq_0 = complex_numeral_0_eq_0 + val numeral_1_eq_1 = complex_numeral_1_eq_1 + val prove_conv = Bin_Simprocs.prove_conv_nohyps_novars + fun norm_tac simps = ALLGOALS (simp_tac (HOL_ss addsimps simps)) + val simplify_meta_eq = Bin_Simprocs.simplify_meta_eq + end + +structure ComplexAbstractNumerals = AbstractNumeralsFun (ComplexAbstractNumeralsData) + +(*For addition, we already have rules for the operand 0. + Multiplication is omitted because there are already special rules for + both 0 and 1 as operands. Unary minus is trivial, just have - 1 = -1. + For the others, having three patterns is a compromise between just having + one (many spurious calls) and having nine (just too many!) *) +val eval_numerals = + map prep_simproc + [("complex_add_eval_numerals", + ["(m::complex) + 1", "(m::complex) + number_of v"], + ComplexAbstractNumerals.proc add_complex_number_of), + ("complex_diff_eval_numerals", + ["(m::complex) - 1", "(m::complex) - number_of v"], + ComplexAbstractNumerals.proc diff_complex_number_of), + ("complex_eq_eval_numerals", + ["(m::complex) = 0", "(m::complex) = 1", "(m::complex) = number_of v"], + ComplexAbstractNumerals.proc eq_complex_number_of)] + +end; + +Addsimprocs Complex_Numeral_Simprocs.eval_numerals; +Addsimprocs Complex_Numeral_Simprocs.cancel_numerals; +Addsimprocs [Complex_Numeral_Simprocs.combine_numerals]; + +(*The Abel_Cancel simprocs are now obsolete +Delsimprocs [Complex_Cancel.sum_conv, Complex_Cancel.rel_conv]; +*) + +(*examples: +print_depth 22; +set timing; +set trace_simp; +fun test s = (Goal s, by (Simp_tac 1)); + +test "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::complex)"; +test " 2*u = (u::complex)"; +test "(i + j + 12 + (k::complex)) - 15 = y"; +test "(i + j + 12 + (k::complex)) - 5 = y"; + +test "( 2*x - (u*v) + y) - v* 3*u = (w::complex)"; +test "( 2*x*u*v + (u*v)* 4 + y) - v*u* 4 = (w::complex)"; +test "( 2*x*u*v + (u*v)* 4 + y) - v*u = (w::complex)"; +test "u*v - (x*u*v + (u*v)* 4 + y) = (w::complex)"; + +test "(i + j + 12 + (k::complex)) = u + 15 + y"; +test "(i + j* 2 + 12 + (k::complex)) = j + 5 + y"; + +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::complex)"; + +test "a + -(b+c) + b = (d::complex)"; +test "a + -(b+c) - b = (d::complex)"; + +(*negative numerals*) +test "(i + j + -2 + (k::complex)) - (u + 5 + y) = zz"; + +test "(i + j + -12 + (k::complex)) - 15 = y"; +test "(i + j + 12 + (k::complex)) - -15 = y"; +test "(i + j + -12 + (k::complex)) - -15 = y"; + +*) + + +(** Constant folding for complex plus and times **) + +structure Complex_Times_Assoc_Data : ASSOC_FOLD_DATA = +struct + val ss = HOL_ss + val eq_reflection = eq_reflection + val sg_ref = Sign.self_ref (Theory.sign_of (the_context ())) + val T = Complex_Numeral_Simprocs.complexT + val plus = Const ("op *", [T,T] ---> T) + val add_ac = complex_mult_ac +end; + +structure Complex_Times_Assoc = Assoc_Fold (Complex_Times_Assoc_Data); + +Addsimprocs [Complex_Times_Assoc.conv]; + +Addsimps [complex_of_real_zero_iff]; + +(*Simplification of x-y = 0 *) + +AddIffs [complex_eq_iff_diff_eq_0 RS sym]; + +(*** Real and imaginary stuff ***) + +Goalw [complex_number_of_def] + "((number_of xa :: complex) + ii * number_of ya = \ +\ number_of xb + ii * number_of yb) = \ +\ (((number_of xa :: complex) = number_of xb) & \ +\ ((number_of ya :: complex) = number_of yb))"; +by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff])); +qed "complex_number_of_eq_cancel_iff"; +Addsimps [complex_number_of_eq_cancel_iff]; + +Goalw [complex_number_of_def] + "((number_of xa :: complex) + number_of ya * ii = \ +\ number_of xb + number_of yb * ii) = \ +\ (((number_of xa :: complex) = number_of xb) & \ +\ ((number_of ya :: complex) = number_of yb))"; +by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iffA])); +qed "complex_number_of_eq_cancel_iffA"; +Addsimps [complex_number_of_eq_cancel_iffA]; + +Goalw [complex_number_of_def] + "((number_of xa :: complex) + number_of ya * ii = \ +\ number_of xb + ii * number_of yb) = \ +\ (((number_of xa :: complex) = number_of xb) & \ +\ ((number_of ya :: complex) = number_of yb))"; +by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iffB])); +qed "complex_number_of_eq_cancel_iffB"; +Addsimps [complex_number_of_eq_cancel_iffB]; + +Goalw [complex_number_of_def] + "((number_of xa :: complex) + ii * number_of ya = \ +\ number_of xb + number_of yb * ii) = \ +\ (((number_of xa :: complex) = number_of xb) & \ +\ ((number_of ya :: complex) = number_of yb))"; +by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iffC])); +qed "complex_number_of_eq_cancel_iffC"; +Addsimps [complex_number_of_eq_cancel_iffC]; + +Goalw [complex_number_of_def] + "((number_of xa :: complex) + ii * number_of ya = \ +\ number_of xb) = \ +\ (((number_of xa :: complex) = number_of xb) & \ +\ ((number_of ya :: complex) = 0))"; +by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff2, + complex_of_real_zero_iff])); +qed "complex_number_of_eq_cancel_iff2"; +Addsimps [complex_number_of_eq_cancel_iff2]; + +Goalw [complex_number_of_def] + "((number_of xa :: complex) + number_of ya * ii = \ +\ number_of xb) = \ +\ (((number_of xa :: complex) = number_of xb) & \ +\ ((number_of ya :: complex) = 0))"; +by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff2a, + complex_of_real_zero_iff])); +qed "complex_number_of_eq_cancel_iff2a"; +Addsimps [complex_number_of_eq_cancel_iff2a]; + +Goalw [complex_number_of_def] + "((number_of xa :: complex) + ii * number_of ya = \ +\ ii * number_of yb) = \ +\ (((number_of xa :: complex) = 0) & \ +\ ((number_of ya :: complex) = number_of yb))"; +by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff3, + complex_of_real_zero_iff])); +qed "complex_number_of_eq_cancel_iff3"; +Addsimps [complex_number_of_eq_cancel_iff3]; + +Goalw [complex_number_of_def] + "((number_of xa :: complex) + number_of ya * ii= \ +\ ii * number_of yb) = \ +\ (((number_of xa :: complex) = 0) & \ +\ ((number_of ya :: complex) = number_of yb))"; +by (auto_tac (claset(), HOL_ss addsimps [complex_eq_cancel_iff3a, + complex_of_real_zero_iff])); +qed "complex_number_of_eq_cancel_iff3a"; +Addsimps [complex_number_of_eq_cancel_iff3a]; + +Goalw [complex_number_of_def] "cnj (number_of v :: complex) = number_of v"; +by (rtac complex_cnj_complex_of_real 1); +qed "complex_number_of_cnj"; +Addsimps [complex_number_of_cnj]; + +Goalw [complex_number_of_def] + "cmod(number_of v :: complex) = abs (number_of v :: real)"; +by (auto_tac (claset(), HOL_ss addsimps [complex_mod_complex_of_real])); +qed "complex_number_of_cmod"; +Addsimps [complex_number_of_cmod]; + +Goalw [complex_number_of_def] + "Re(number_of v :: complex) = number_of v"; +by (auto_tac (claset(), HOL_ss addsimps [Re_complex_of_real])); +qed "complex_number_of_Re"; +Addsimps [complex_number_of_Re]; + +Goalw [complex_number_of_def] + "Im(number_of v :: complex) = 0"; +by (auto_tac (claset(), HOL_ss addsimps [Im_complex_of_real])); +qed "complex_number_of_Im"; +Addsimps [complex_number_of_Im]; + +Goalw [expi_def] + "expi((2::complex) * complex_of_real pi * ii) = 1"; +by (auto_tac (claset(),simpset() addsimps [complex_Re_mult_eq, + complex_Im_mult_eq,cis_def])); +qed "expi_two_pi_i"; +Addsimps [expi_two_pi_i]; + +(*examples: +print_depth 22; +set timing; +set trace_simp; +fun test s = (Goal s, by (Simp_tac 1)); + +test "23 * ii + 45 * ii= (x::complex)"; + +test "5 * ii + 12 - 45 * ii= (x::complex)"; +test "5 * ii + 40 - 12 * ii + 9 = (x::complex) + 89 * ii"; +test "5 * ii + 40 - 12 * ii + 9 - 78 = (x::complex) + 89 * ii"; + +test "l + 10 * ii + 90 + 3*l + 9 + 45 * ii= (x::complex)"; +test "87 + 10 * ii + 90 + 3*7 + 9 + 45 * ii= (x::complex)"; + + +*) diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/ComplexBin.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/ComplexBin.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,22 @@ +(* Title: ComplexBin.thy + Author: Jacques D. Fleuriot + Copyright: 2001 University of Edinburgh + Descrition: Binary arithmetic for the complex numbers + This case is reduced to that for the reals. +*) + +ComplexBin = Complex + + + +instance + complex :: number + +instance complex :: plus_ac0(complex_add_commute,complex_add_assoc,complex_add_zero_left) + + +defs + complex_number_of_def + "number_of v == complex_of_real (number_of v)" + (*::bin=>complex ::bin=>complex*) + +end diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/NSCA.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/NSCA.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,1451 @@ +(* Title : NSCA.ML + Author : Jacques D. Fleuriot + Copyright : 2001,2002 University of Edinburgh + Description : Infinite, infinitesimal complex number etc! +*) + +(*--------------------------------------------------------------------------------------*) +(* Closure laws for members of (embedded) set standard complex SComplex *) +(* -------------------------------------------------------------------------------------*) + +Goalw [SComplex_def] "[| (x::hcomplex): SComplex; y: SComplex |] ==> x + y: SComplex"; +by (Step_tac 1); +by (res_inst_tac [("x","r + ra")] exI 1); +by (Simp_tac 1); +qed "SComplex_add"; + +Goalw [SComplex_def] "[| (x::hcomplex): SComplex; y: SComplex |] ==> x * y: SComplex"; +by (Step_tac 1); +by (res_inst_tac [("x","r * ra")] exI 1); +by (Simp_tac 1); +qed "SComplex_mult"; + +Goalw [SComplex_def] "x: SComplex ==> inverse x : SComplex"; +by (blast_tac (claset() addIs [hcomplex_of_complex_inverse RS sym]) 1); +qed "SComplex_inverse"; + +Goal "[| x: SComplex; y: SComplex |] ==> x/y: SComplex"; +by (asm_simp_tac (simpset() addsimps [SComplex_mult,SComplex_inverse, + hcomplex_divide_def]) 1); +qed "SComplex_divide"; + +Goalw [SComplex_def] "x: SComplex ==> -x : SComplex"; +by (blast_tac (claset() addIs [hcomplex_of_complex_minus RS sym]) 1); +qed "SComplex_minus"; + +Goal "(-x : SComplex) = (x: SComplex)"; +by Auto_tac; +by (etac SComplex_minus 2); +by (dtac SComplex_minus 1); +by Auto_tac; +qed "SComplex_minus_iff"; +Addsimps [SComplex_minus_iff]; + +Goal "[| x + y : SComplex; y: SComplex |] ==> x: SComplex"; +by (dres_inst_tac [("x","y")] SComplex_minus 1); +by (dtac SComplex_add 1); +by (assume_tac 1); +by Auto_tac; +qed "SComplex_add_cancel"; + +Goalw [hcomplex_of_complex_def] + "hcmod (hcomplex_of_complex r) : Reals"; +by (simp_tac (simpset() addsimps [hcmod,SReal_def, + hypreal_of_real_def]) 1); +qed "SReal_hcmod_hcomplex_of_complex"; +Addsimps [SReal_hcmod_hcomplex_of_complex]; + +Goalw [hcomplex_number_of_def] + "hcmod (number_of w ::hcomplex) : Reals"; +by (rtac SReal_hcmod_hcomplex_of_complex 1); +qed "SReal_hcmod_number_of"; +Addsimps [SReal_hcmod_number_of]; + +Goalw [SComplex_def] "x: SComplex ==> hcmod x : Reals"; +by Auto_tac; +qed "SReal_hcmod_SComplex"; + +Goalw [SComplex_def] "hcomplex_of_complex x: SComplex"; +by (Blast_tac 1); +qed "SComplex_hcomplex_of_complex"; +Addsimps [SComplex_hcomplex_of_complex]; + +Goalw [hcomplex_number_of_def] "(number_of w ::hcomplex) : SComplex"; +by (rtac SComplex_hcomplex_of_complex 1); +qed "SComplex_number_of"; +Addsimps [SComplex_number_of]; + +Goalw [hcomplex_divide_def] "r : SComplex ==> r/(number_of w::hcomplex) : SComplex"; +by (blast_tac (claset() addSIs [SComplex_number_of, SComplex_mult, + SComplex_inverse]) 1); +qed "SComplex_divide_number_of"; + +Goalw [SComplex_def] "{x. hcomplex_of_complex x : SComplex} = (UNIV::complex set)"; +by Auto_tac; +qed "SComplex_UNIV_complex"; + +Goalw [SComplex_def] "(x: SComplex) = (EX y. x = hcomplex_of_complex y)"; +by Auto_tac; +qed "SComplex_iff"; + +Goalw [SComplex_def] "hcomplex_of_complex `(UNIV::complex set) = SComplex"; +by Auto_tac; +qed "hcomplex_of_complex_image"; + +Goalw [SComplex_def] "inv hcomplex_of_complex `SComplex = (UNIV::complex set)"; +by Auto_tac; +by (rtac (inj_hcomplex_of_complex RS inv_f_f RS subst) 1); +by (Blast_tac 1); +qed "inv_hcomplex_of_complex_image"; + +Goalw [SComplex_def] + "[| EX x. x: P; P <= SComplex |] ==> EX Q. P = hcomplex_of_complex ` Q"; +by (Best_tac 1); +qed "SComplex_hcomplex_of_complex_image"; + +Goal "[| (x::hcomplex): SComplex; y: SComplex; hcmod x < hcmod y \ +\ |] ==> EX r: Reals. hcmod x< r & r < hcmod y"; +by (auto_tac (claset() addIs [SReal_dense], simpset() + addsimps [SReal_hcmod_SComplex])); +qed "SComplex_SReal_dense"; + +Goalw [SComplex_def,SReal_def] + "z : SComplex ==> hcmod z : Reals"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcmod,hypreal_of_real_def, + hcomplex_of_complex_def,cmod_def])); +by (res_inst_tac [("x","cmod r")] exI 1); +by (ultra_tac (claset(),simpset() addsimps [cmod_def]) 1); +qed "SComplex_hcmod_SReal"; + +Goal "0 : SComplex"; +by (auto_tac (claset(),simpset() addsimps [SComplex_def])); +qed "SComplex_zero"; +Addsimps [SComplex_zero]; + +Goal "1 : SComplex"; +by (auto_tac (claset(),simpset() addsimps [SComplex_def,hcomplex_of_complex_def, + hcomplex_one_def])); +qed "SComplex_one"; +Addsimps [SComplex_one]; + +(* +Goalw [SComplex_def,SReal_def] "hcmod z : Reals ==> z : SComplex"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcmod,hypreal_of_real_def, + hcomplex_of_complex_def,cmod_def])); +*) + +(*--------------------------------------------------------------------------------------------*) +(* Set of finite elements is a subring of the extended complex numbers *) +(* -------------------------------------------------------------------------------------------*) + +Goalw [CFinite_def] "[|x : CFinite; y : CFinite|] ==> (x+y) : CFinite"; +by (blast_tac (claset() addSIs [SReal_add,hcmod_add_less]) 1); +qed "CFinite_add"; + +Goalw [CFinite_def] "[|x : CFinite; y : CFinite|] ==> x*y : CFinite"; +by (blast_tac (claset() addSIs [SReal_mult,hcmod_mult_less]) 1); +qed "CFinite_mult"; + +Goalw [CFinite_def] "(-x : CFinite) = (x : CFinite)"; +by (Simp_tac 1); +qed "CFinite_minus_iff"; +Addsimps [CFinite_minus_iff]; + +Goalw [SComplex_def,CFinite_def] "SComplex <= CFinite"; +by Auto_tac; +by (res_inst_tac [("x","1 + hcmod(hcomplex_of_complex r)")] bexI 1); +by (auto_tac (claset() addIs [SReal_add],simpset())); +qed "SComplex_subset_CFinite"; +Addsimps [ SComplex_subset_CFinite]; + +Goal "hcmod (hcomplex_of_complex r) : HFinite"; +by (auto_tac (claset() addSIs [ SReal_subset_HFinite RS subsetD],simpset())); +qed "HFinite_hcmod_hcomplex_of_complex"; +Addsimps [HFinite_hcmod_hcomplex_of_complex]; + +Goal "hcomplex_of_complex x: CFinite"; +by (auto_tac (claset() addSIs [ SComplex_subset_CFinite RS subsetD],simpset())); +qed "CFinite_hcomplex_of_complex"; +Addsimps [CFinite_hcomplex_of_complex]; + +Goalw [CFinite_def] "x : CFinite ==> EX t: Reals. hcmod x < t"; +by Auto_tac; +qed "CFiniteD"; + +Goalw [CFinite_def] "(x : CFinite) = (hcmod x : HFinite)"; +by (auto_tac (claset(), simpset() addsimps [HFinite_def])); +qed "CFinite_hcmod_iff"; + +Goal "number_of w : CFinite"; +by (rtac (SComplex_number_of RS (SComplex_subset_CFinite RS subsetD)) 1); +qed "CFinite_number_of"; +Addsimps [CFinite_number_of]; + +Goal "[|x : CFinite; y <= hcmod x; 0 <= y |] ==> y: HFinite"; +by (auto_tac (claset() addIs [HFinite_bounded],simpset() addsimps + [CFinite_hcmod_iff])); +qed "CFinite_bounded"; + +(*--------------------------------------------------------------------------------------*) +(* Set of complex infinitesimals is a subring of the nonstandard complex numbers *) +(*--------------------------------------------------------------------------------------*) + +Goalw [CInfinitesimal_def] + "x : CInfinitesimal ==> ALL r: Reals. 0 < r --> hcmod x < r"; +by Auto_tac; +qed "CInfinitesimalD"; + +Goalw [CInfinitesimal_def] "0 : CInfinitesimal"; +by Auto_tac; +qed "CInfinitesimal_zero"; +AddIffs [CInfinitesimal_zero]; + +Goal "x/(2::hcomplex) + x/(2::hcomplex) = x"; +by Auto_tac; +qed "hcomplex_sum_of_halves"; + +Goalw [CInfinitesimal_def,Infinitesimal_def] + "(z : CInfinitesimal) = (hcmod z : Infinitesimal)"; +by Auto_tac; +qed "CInfinitesimal_hcmod_iff"; + +Goal "1 ~: CInfinitesimal"; +by (simp_tac (simpset() addsimps [CInfinitesimal_hcmod_iff]) 1); +qed "one_not_CInfinitesimal"; +Addsimps [one_not_CInfinitesimal]; + +Goal "[| x : CInfinitesimal; y : CInfinitesimal |] ==> (x+y) : CInfinitesimal"; +by (auto_tac (claset(),simpset() addsimps [CInfinitesimal_hcmod_iff])); +by (rtac hrabs_le_Infinitesimal 1); +by (res_inst_tac [("y","hcmod y")] Infinitesimal_add 1); +by Auto_tac; +qed "CInfinitesimal_add"; + +Goalw [CInfinitesimal_def] "(-x:CInfinitesimal) = (x:CInfinitesimal)"; +by (Full_simp_tac 1); +qed "CInfinitesimal_minus_iff"; +Addsimps [CInfinitesimal_minus_iff]; + +Goal "[| x : CInfinitesimal; y : CInfinitesimal |] ==> x-y : CInfinitesimal"; +by (asm_simp_tac + (simpset() addsimps [hcomplex_diff_def, CInfinitesimal_add]) 1); +qed "CInfinitesimal_diff"; + +Goal "[| x : CInfinitesimal; y : CInfinitesimal |] ==> (x * y) : CInfinitesimal"; +by (auto_tac (claset() addIs [Infinitesimal_mult],simpset() addsimps + [CInfinitesimal_hcmod_iff,hcmod_mult])); +qed "CInfinitesimal_mult"; + +Goal "[| x : CInfinitesimal; y : CFinite |] ==> (x * y) : CInfinitesimal"; +by (auto_tac (claset() addIs [Infinitesimal_HFinite_mult],simpset() + addsimps [CInfinitesimal_hcmod_iff,CFinite_hcmod_iff,hcmod_mult])); +qed "CInfinitesimal_CFinite_mult"; + +Goal "[| x : CInfinitesimal; y : CFinite |] ==> (y * x) : CInfinitesimal"; +by (auto_tac (claset() addDs [CInfinitesimal_CFinite_mult], + simpset() addsimps [hcomplex_mult_commute])); +qed "CInfinitesimal_CFinite_mult2"; + +Goalw [CInfinite_def,HInfinite_def] + "(z : CInfinite) = (hcmod z : HInfinite)"; +by Auto_tac; +qed "CInfinite_hcmod_iff"; + +Goal "x: CInfinite ==> inverse x: CInfinitesimal"; +by (auto_tac (claset() addIs [HInfinite_inverse_Infinitesimal], + simpset() addsimps [CInfinitesimal_hcmod_iff, + CInfinite_hcmod_iff,hcmod_hcomplex_inverse])); +qed "CInfinite_inverse_CInfinitesimal"; + +Goal "[|x: CInfinite; y: CInfinite|] ==> (x*y): CInfinite"; +by (auto_tac (claset() addIs [HInfinite_mult],simpset() addsimps + [CInfinite_hcmod_iff,hcmod_mult])); +qed "CInfinite_mult"; + +Goalw [CInfinite_def] "(-x : CInfinite) = (x : CInfinite)"; +by (Simp_tac 1); +qed "CInfinite_minus_iff"; +Addsimps [CInfinite_minus_iff]; + +Goal "[|a: CFinite; b: CFinite; c: CFinite|] \ +\ ==> a*a + b*b + c*c : CFinite"; +by (auto_tac (claset() addIs [CFinite_mult,CFinite_add], simpset())); +qed "CFinite_sum_squares"; + +Goal "x ~: CInfinitesimal ==> x ~= 0"; +by Auto_tac; +qed "not_CInfinitesimal_not_zero"; + +Goal "x: CFinite - CInfinitesimal ==> x ~= 0"; +by Auto_tac; +qed "not_CInfinitesimal_not_zero2"; + +Goal "x : CFinite - CInfinitesimal ==> hcmod x : HFinite - Infinitesimal"; +by (auto_tac (claset(),simpset() addsimps [CFinite_hcmod_iff,CInfinitesimal_hcmod_iff])); +qed "CFinite_diff_CInfinitesimal_hcmod"; + +Goal "[| e : CInfinitesimal; hcmod x < hcmod e |] ==> x : CInfinitesimal"; +by (auto_tac (claset() addIs [hrabs_less_Infinitesimal],simpset() + addsimps [CInfinitesimal_hcmod_iff])); +qed "hcmod_less_CInfinitesimal"; + +Goal "[| e : CInfinitesimal; hcmod x <= hcmod e |] ==> x : CInfinitesimal"; +by (auto_tac (claset() addIs [hrabs_le_Infinitesimal],simpset() + addsimps [CInfinitesimal_hcmod_iff])); +qed "hcmod_le_CInfinitesimal"; + +Goal "[| e : CInfinitesimal; \ +\ e' : CInfinitesimal; \ +\ hcmod e' < hcmod x ; hcmod x < hcmod e \ +\ |] ==> x : CInfinitesimal"; +by (auto_tac (claset() addIs [Infinitesimal_interval],simpset() + addsimps [CInfinitesimal_hcmod_iff])); +qed "CInfinitesimal_interval"; + +Goal "[| e : CInfinitesimal; \ +\ e' : CInfinitesimal; \ +\ hcmod e' <= hcmod x ; hcmod x <= hcmod e \ +\ |] ==> x : CInfinitesimal"; +by (auto_tac (claset() addIs [Infinitesimal_interval2],simpset() + addsimps [CInfinitesimal_hcmod_iff])); +qed "CInfinitesimal_interval2"; + +Goal "[| x ~: CInfinitesimal; y ~: CInfinitesimal|] ==> (x*y) ~: CInfinitesimal"; +by (auto_tac (claset(),simpset() addsimps [CInfinitesimal_hcmod_iff,hcmod_mult])); +by (dtac not_Infinitesimal_mult 1); +by Auto_tac; +qed "not_CInfinitesimal_mult"; + +Goal "x*y : CInfinitesimal ==> x : CInfinitesimal | y : CInfinitesimal"; +by (auto_tac (claset() addDs [Infinitesimal_mult_disj],simpset() addsimps + [CInfinitesimal_hcmod_iff,hcmod_mult])); +qed "CInfinitesimal_mult_disj"; + +Goal "[| x : CFinite - CInfinitesimal; \ +\ y : CFinite - CInfinitesimal \ +\ |] ==> (x*y) : CFinite - CInfinitesimal"; +by (Clarify_tac 1); +by (blast_tac (claset() addDs [CFinite_mult,not_CInfinitesimal_mult]) 1); +qed "CFinite_CInfinitesimal_diff_mult"; + +Goal "CInfinitesimal <= CFinite"; +by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD], + simpset() addsimps [CInfinitesimal_hcmod_iff,CFinite_hcmod_iff])); +qed "CInfinitesimal_subset_CFinite"; + +Goal "x: CInfinitesimal ==> x * hcomplex_of_complex r : CInfinitesimal"; +by (auto_tac (claset() addSIs [Infinitesimal_HFinite_mult], + simpset() addsimps [CInfinitesimal_hcmod_iff,hcmod_mult])); +qed "CInfinitesimal_hcomplex_of_complex_mult"; + +Goal "x: CInfinitesimal ==> hcomplex_of_complex r * x: CInfinitesimal"; +by (auto_tac (claset() addSIs [Infinitesimal_HFinite_mult2], + simpset() addsimps [CInfinitesimal_hcmod_iff,hcmod_mult])); +qed "CInfinitesimal_hcomplex_of_complex_mult2"; + + +(*--------------------------------------------------------------------------------------*) +(* Infinitely close relation @c= *) +(* -------------------------------------------------------------------------------------*) + +(* +Goalw [capprox_def,approx_def] "(z @c= w) = (hcmod z @= hcmod w)"; +by (auto_tac (claset(),simpset() addsimps [CInfinitesimal_hcmod_iff])); +*) + +Goal "x:CInfinitesimal = (x @c= 0)"; +by (simp_tac (simpset() addsimps [CInfinitesimal_hcmod_iff, + capprox_def]) 1); +qed "mem_cinfmal_iff"; + +Goalw [capprox_def,hcomplex_diff_def] + " (x @c= y) = (x + -y @c= 0)"; +by (Simp_tac 1); +qed "capprox_minus_iff"; + +Goalw [capprox_def,hcomplex_diff_def] + " (x @c= y) = (-y + x @c= 0)"; +by (simp_tac (simpset() addsimps [hcomplex_add_commute]) 1); +qed "capprox_minus_iff2"; + +Goalw [capprox_def] "x @c= x"; +by (Simp_tac 1); +qed "capprox_refl"; +Addsimps [capprox_refl]; + +Goalw [capprox_def,CInfinitesimal_def] + "x @c= y ==> y @c= x"; +by (auto_tac (claset() addSDs [bspec],simpset() addsimps + [hcmod_diff_commute])); +qed "capprox_sym"; + +Goalw [capprox_def] "[| x @c= y; y @c= z |] ==> x @c= z"; +by (dtac CInfinitesimal_add 1); +by (assume_tac 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_def])); +qed "capprox_trans"; + +Goal "[| r @c= x; s @c= x |] ==> r @c= s"; +by (blast_tac (claset() addIs [capprox_sym, capprox_trans]) 1); +qed "capprox_trans2"; + +Goal "[| x @c= r; x @c= s|] ==> r @c= s"; +by (blast_tac (claset() addIs [capprox_sym, capprox_trans]) 1); +qed "capprox_trans3"; + +Goal "(number_of w @c= x) = (x @c= number_of w)"; +by (blast_tac (claset() addIs [capprox_sym]) 1); +qed "number_of_capprox_reorient"; +Addsimps [number_of_capprox_reorient]; + +Goal "(x-y : CInfinitesimal) = (x @c= y)"; +by (auto_tac (claset(), + simpset() addsimps [hcomplex_diff_def, capprox_minus_iff RS sym, + mem_cinfmal_iff])); +qed "CInfinitesimal_capprox_minus"; + +Goalw [cmonad_def] "(x @c= y) = (cmonad(x)=cmonad(y))"; +by (auto_tac (claset() addDs [capprox_sym] + addSEs [capprox_trans,equalityCE], + simpset())); +qed "capprox_monad_iff"; + +Goal "[| x: CInfinitesimal; y: CInfinitesimal |] ==> x @c= y"; +by (asm_full_simp_tac (simpset() addsimps [mem_cinfmal_iff]) 1); +by (blast_tac (claset() addIs [capprox_trans, capprox_sym]) 1); +qed "Infinitesimal_capprox"; + +val prem1::prem2::rest = +goalw thy [capprox_def,hcomplex_diff_def] + "[| a @c= b; c @c= d |] ==> a+c @c= b+d"; +by (rtac (hcomplex_minus_add_distrib RS ssubst) 1); +by (rtac (hcomplex_add_assoc RS ssubst) 1); +by (res_inst_tac [("y1","c")] (hcomplex_add_left_commute RS subst) 1); +by (rtac (hcomplex_add_assoc RS subst) 1); +by (rtac ([prem1,prem2] MRS CInfinitesimal_add) 1); +qed "capprox_add"; + +Goal "a @c= b ==> -a @c= -b"; +by (rtac ((capprox_minus_iff RS iffD2) RS capprox_sym) 1); +by (dtac (capprox_minus_iff RS iffD1) 1); +by (simp_tac (simpset() addsimps [hcomplex_add_commute]) 1); +qed "capprox_minus"; + +Goal "-a @c= -b ==> a @c= b"; +by (auto_tac (claset() addDs [capprox_minus], simpset())); +qed "capprox_minus2"; + +Goal "(-a @c= -b) = (a @c= b)"; +by (blast_tac (claset() addIs [capprox_minus,capprox_minus2]) 1); +qed "capprox_minus_cancel"; +Addsimps [capprox_minus_cancel]; + +Goal "[| a @c= b; c @c= d |] ==> a + -c @c= b + -d"; +by (blast_tac (claset() addSIs [capprox_add,capprox_minus]) 1); +qed "capprox_add_minus"; + +Goalw [capprox_def,hcomplex_diff_def] + "[| a @c= b; c: CFinite|] ==> a*c @c= b*c"; +by (asm_full_simp_tac (simpset() addsimps [CInfinitesimal_CFinite_mult, + hcomplex_minus_mult_eq1,hcomplex_add_mult_distrib RS sym] + delsimps [hcomplex_minus_mult_eq1 RS sym]) 1); +qed "capprox_mult1"; + +Goal "[|a @c= b; c: CFinite|] ==> c*a @c= c*b"; +by (asm_simp_tac (simpset() addsimps [capprox_mult1,hcomplex_mult_commute]) 1); +qed "capprox_mult2"; + +Goal "[|u @c= v*x; x @c= y; v: CFinite|] ==> u @c= v*y"; +by (fast_tac (claset() addIs [capprox_mult2,capprox_trans]) 1); +qed "capprox_mult_subst"; + +Goal "[| u @c= x*v; x @c= y; v: CFinite |] ==> u @c= y*v"; +by (fast_tac (claset() addIs [capprox_mult1,capprox_trans]) 1); +qed "capprox_mult_subst2"; + +Goal "[| u @c= x*hcomplex_of_complex v; x @c= y |] ==> u @c= y*hcomplex_of_complex v"; +by (auto_tac (claset() addIs [capprox_mult_subst2], simpset())); +qed "capprox_mult_subst_SComplex"; + +Goalw [capprox_def] "a = b ==> a @c= b"; +by (Asm_simp_tac 1); +qed "capprox_eq_imp"; + +Goal "x: CInfinitesimal ==> -x @c= x"; +by (fast_tac (HOL_cs addIs [CInfinitesimal_minus_iff RS iffD2, + mem_cinfmal_iff RS iffD1,capprox_trans2]) 1); +qed "CInfinitesimal_minus_capprox"; + +Goalw [capprox_def] + "(EX y: CInfinitesimal. x - z = y) = (x @c= z)"; +by (Blast_tac 1); +qed "bex_CInfinitesimal_iff"; + +Goal "(EX y: CInfinitesimal. x = z + y) = (x @c= z)"; +by (asm_full_simp_tac (simpset() addsimps [bex_CInfinitesimal_iff RS sym]) 1); +by (Force_tac 1); +qed "bex_CInfinitesimal_iff2"; + +Goal "[| y: CInfinitesimal; x + y = z |] ==> x @c= z"; +by (rtac (bex_CInfinitesimal_iff RS iffD1) 1); +by (dtac (CInfinitesimal_minus_iff RS iffD2) 1); +by (auto_tac (claset(), simpset() addsimps [hcomplex_add_assoc RS sym])); +qed "CInfinitesimal_add_capprox"; + +Goal "y: CInfinitesimal ==> x @c= x + y"; +by (rtac (bex_CInfinitesimal_iff RS iffD1) 1); +by (dtac (CInfinitesimal_minus_iff RS iffD2) 1); +by (auto_tac (claset(), simpset() addsimps [hcomplex_add_assoc RS sym])); +qed "CInfinitesimal_add_capprox_self"; + +Goal "y: CInfinitesimal ==> x @c= y + x"; +by (auto_tac (claset() addDs [CInfinitesimal_add_capprox_self], + simpset() addsimps [hcomplex_add_commute])); +qed "CInfinitesimal_add_capprox_self2"; + +Goal "y: CInfinitesimal ==> x @c= x + -y"; +by (blast_tac (claset() addSIs [CInfinitesimal_add_capprox_self, + CInfinitesimal_minus_iff RS iffD2]) 1); +qed "CInfinitesimal_add_minus_capprox_self"; + +Goal "[| y: CInfinitesimal; x+y @c= z|] ==> x @c= z"; +by (dres_inst_tac [("x","x")] (CInfinitesimal_add_capprox_self RS capprox_sym) 1); +by (etac (capprox_trans3 RS capprox_sym) 1); +by (assume_tac 1); +qed "CInfinitesimal_add_cancel"; + +Goal "[| y: CInfinitesimal; x @c= z + y|] ==> x @c= z"; +by (dres_inst_tac [("x","z")] (CInfinitesimal_add_capprox_self2 RS capprox_sym) 1); +by (etac (capprox_trans3 RS capprox_sym) 1); +by (asm_full_simp_tac (simpset() addsimps [hcomplex_add_commute]) 1); +by (etac capprox_sym 1); +qed "CInfinitesimal_add_right_cancel"; + +Goal "d + b @c= d + c ==> b @c= c"; +by (dtac (capprox_minus_iff RS iffD1) 1); +by (asm_full_simp_tac (simpset() addsimps + [hcomplex_minus_add_distrib,capprox_minus_iff RS sym] + @ hcomplex_add_ac) 1); +qed "capprox_add_left_cancel"; + +Goal "b + d @c= c + d ==> b @c= c"; +by (rtac capprox_add_left_cancel 1); +by (asm_full_simp_tac (simpset() addsimps + [hcomplex_add_commute]) 1); +qed "capprox_add_right_cancel"; + +Goal "b @c= c ==> d + b @c= d + c"; +by (rtac (capprox_minus_iff RS iffD2) 1); +by (asm_full_simp_tac (simpset() addsimps + [capprox_minus_iff RS sym] @ hcomplex_add_ac) 1); +qed "capprox_add_mono1"; + +Goal "b @c= c ==> b + a @c= c + a"; +by (asm_simp_tac (simpset() addsimps + [hcomplex_add_commute,capprox_add_mono1]) 1); +qed "capprox_add_mono2"; + +Goal "(a + b @c= a + c) = (b @c= c)"; +by (fast_tac (claset() addEs [capprox_add_left_cancel, + capprox_add_mono1]) 1); +qed "capprox_add_left_iff"; + +AddIffs [capprox_add_left_iff]; + + +Goal "(b + a @c= c + a) = (b @c= c)"; +by (simp_tac (simpset() addsimps [hcomplex_add_commute]) 1); +qed "capprox_add_right_iff"; + +AddIffs [capprox_add_right_iff]; + +Goal "[| x: CFinite; x @c= y |] ==> y: CFinite"; +by (dtac (bex_CInfinitesimal_iff2 RS iffD2) 1); +by (Step_tac 1); +by (dtac (CInfinitesimal_subset_CFinite RS subsetD + RS (CFinite_minus_iff RS iffD2)) 1); +by (dtac CFinite_add 1); +by (assume_tac 1 THEN Auto_tac); +qed "capprox_CFinite"; + +Goal "x @c= hcomplex_of_complex D ==> x: CFinite"; +by (rtac (capprox_sym RSN (2,capprox_CFinite)) 1); +by Auto_tac; +qed "capprox_hcomplex_of_complex_CFinite"; + +Goal "[|a @c= b; c @c= d; b: CFinite; d: CFinite|] ==> a*c @c= b*d"; +by (rtac capprox_trans 1); +by (rtac capprox_mult2 2); +by (rtac capprox_mult1 1); +by (blast_tac (claset() addIs [capprox_CFinite, capprox_sym]) 2); +by Auto_tac; +qed "capprox_mult_CFinite"; + +Goal "[|a @c= hcomplex_of_complex b; c @c= hcomplex_of_complex d |] \ +\ ==> a*c @c= hcomplex_of_complex b * hcomplex_of_complex d"; +by (blast_tac (claset() addSIs [capprox_mult_CFinite, + capprox_hcomplex_of_complex_CFinite,CFinite_hcomplex_of_complex]) 1); +qed "capprox_mult_hcomplex_of_complex"; + +Goal "[| a: SComplex; a ~= 0; a*x @c= 0 |] ==> x @c= 0"; +by (dtac (SComplex_inverse RS (SComplex_subset_CFinite RS subsetD)) 1); +by (auto_tac (claset() addDs [capprox_mult2], + simpset() addsimps [hcomplex_mult_assoc RS sym])); +qed "capprox_SComplex_mult_cancel_zero"; + +Goal "[| a: SComplex; x @c= 0 |] ==> x*a @c= 0"; +by (auto_tac (claset() addDs [(SComplex_subset_CFinite RS subsetD), + capprox_mult1], simpset())); +qed "capprox_mult_SComplex1"; + +Goal "[| a: SComplex; x @c= 0 |] ==> a*x @c= 0"; +by (auto_tac (claset() addDs [(SComplex_subset_CFinite RS subsetD), + capprox_mult2], simpset())); +qed "capprox_mult_SComplex2"; + +Goal "[|a : SComplex; a ~= 0 |] ==> (a*x @c= 0) = (x @c= 0)"; +by (blast_tac (claset() addIs [capprox_SComplex_mult_cancel_zero, + capprox_mult_SComplex2]) 1); +qed "capprox_mult_SComplex_zero_cancel_iff"; +Addsimps [capprox_mult_SComplex_zero_cancel_iff]; + +Goal "[| a: SComplex; a ~= 0; a* w @c= a*z |] ==> w @c= z"; +by (dtac (SComplex_inverse RS (SComplex_subset_CFinite RS subsetD)) 1); +by (auto_tac (claset() addDs [capprox_mult2], + simpset() addsimps [hcomplex_mult_assoc RS sym])); +qed "capprox_SComplex_mult_cancel"; + +Goal "[| a: SComplex; a ~= 0|] ==> (a* w @c= a*z) = (w @c= z)"; +by (auto_tac (claset() addSIs [capprox_mult2,SComplex_subset_CFinite RS subsetD] + addIs [capprox_SComplex_mult_cancel], simpset())); +qed "capprox_SComplex_mult_cancel_iff1"; +Addsimps [capprox_SComplex_mult_cancel_iff1]; + +Goal "(x @c= y) = (hcmod (y - x) @= 0)"; +by (rtac (capprox_minus_iff RS ssubst) 1); +by (auto_tac (claset(),simpset() addsimps [capprox_def, + CInfinitesimal_hcmod_iff,mem_infmal_iff,symmetric hcomplex_diff_def, + hcmod_diff_commute])); +qed "capprox_hcmod_approx_zero"; + +Goal "(x @c= 0) = (hcmod x @= 0)"; +by (auto_tac (claset(),simpset() addsimps + [capprox_hcmod_approx_zero])); +qed "capprox_approx_zero_iff"; + +Goal "(-x @c= 0) = (x @c= 0)"; +by (auto_tac (claset(),simpset() addsimps + [capprox_hcmod_approx_zero])); +qed "capprox_minus_zero_cancel_iff"; +Addsimps [capprox_minus_zero_cancel_iff]; + +Goal "u @c= 0 ==> hcmod(x + u) - hcmod x : Infinitesimal"; +by (res_inst_tac [("e","hcmod u"),("e'","- hcmod u")] Infinitesimal_interval2 1); +by (auto_tac (claset() addDs [capprox_approx_zero_iff RS iffD1], + simpset() addsimps [mem_infmal_iff RS sym,hypreal_diff_def])); +by (res_inst_tac [("C","hcmod x")] hypreal_le_add_left_cancel 1); +by (auto_tac (claset(),simpset() addsimps [symmetric hypreal_diff_def])); +qed "Infinitesimal_hcmod_add_diff"; + +Goal "u @c= 0 ==> hcmod(x + u) @= hcmod x"; +by (rtac (approx_minus_iff RS iffD2) 1); +by (auto_tac (claset() addIs [Infinitesimal_hcmod_add_diff], + simpset() addsimps [mem_infmal_iff RS sym,symmetric hypreal_diff_def])); +qed "approx_hcmod_add_hcmod"; + +Goal "x @c= y ==> hcmod x @= hcmod y"; +by (auto_tac (claset() addIs [approx_hcmod_add_hcmod] + addSDs [bex_CInfinitesimal_iff2 RS iffD2],simpset() addsimps [mem_cinfmal_iff])); +qed "capprox_hcmod_approx"; + +(*--------------------------------------------------------------------------------------*) +(* zero is the only complex number that is also infinitesimal *) +(*--------------------------------------------------------------------------------------*) + +Goal "[| x: SComplex; y: CInfinitesimal; 0 < hcmod x |] ==> hcmod y < hcmod x"; +by (auto_tac (claset() addSIs [Infinitesimal_less_SReal,SComplex_hcmod_SReal], + simpset() addsimps [CInfinitesimal_hcmod_iff])); +qed "CInfinitesimal_less_SComplex"; + +Goal "y: CInfinitesimal ==> ALL r: SComplex. 0 < hcmod r --> hcmod y < hcmod r"; +by (blast_tac (claset() addIs [CInfinitesimal_less_SComplex]) 1); +qed "CInfinitesimal_less_SComplex2"; + +Goal "SComplex Int CInfinitesimal = {0}"; +by (auto_tac (claset(),simpset() addsimps [SComplex_def,CInfinitesimal_hcmod_iff])); +by (cut_inst_tac [("r","r")] SReal_hcmod_hcomplex_of_complex 1); +by (dres_inst_tac [("A","Reals")] IntI 1 THEN assume_tac 1); +by (subgoal_tac "hcmod (hcomplex_of_complex r) = 0" 1); +by (Asm_full_simp_tac 1); +by (cut_facts_tac [SReal_Int_Infinitesimal_zero] 1); +by (rotate_tac 2 1); +by (Asm_full_simp_tac 1); +qed "SComplex_Int_CInfinitesimal_zero"; + +Goal "[| x: SComplex; x: CInfinitesimal|] ==> x = 0"; +by (cut_facts_tac [SComplex_Int_CInfinitesimal_zero] 1); +by (Blast_tac 1); +qed "SComplex_CInfinitesimal_zero"; + +Goal "[| x : SComplex; x ~= 0 |] ==> x : CFinite - CInfinitesimal"; +by (auto_tac (claset() addDs [SComplex_CInfinitesimal_zero, + SComplex_subset_CFinite RS subsetD], + simpset())); +qed "SComplex_CFinite_diff_CInfinitesimal"; + +Goal "hcomplex_of_complex x ~= 0 ==> hcomplex_of_complex x : CFinite - CInfinitesimal"; +by (rtac SComplex_CFinite_diff_CInfinitesimal 1); +by Auto_tac; +qed "hcomplex_of_complex_CFinite_diff_CInfinitesimal"; + +Goal "(hcomplex_of_complex x : CInfinitesimal) = (x=0)"; +by (auto_tac (claset(), simpset() addsimps [hcomplex_of_complex_zero])); +by (rtac ccontr 1); +by (rtac (hcomplex_of_complex_CFinite_diff_CInfinitesimal RS DiffD2) 1); +by Auto_tac; +qed "hcomplex_of_complex_CInfinitesimal_iff_0"; +AddIffs [hcomplex_of_complex_CInfinitesimal_iff_0]; + +Goal "number_of w ~= (0::hcomplex) ==> number_of w ~: CInfinitesimal"; +by (fast_tac (claset() addDs [SComplex_number_of RS SComplex_CInfinitesimal_zero]) 1); +qed "number_of_not_CInfinitesimal"; +Addsimps [number_of_not_CInfinitesimal]; + +Goal "[| y: SComplex; x @c= y; y~= 0 |] ==> x ~= 0"; +by (auto_tac (claset() addDs [SComplex_CInfinitesimal_zero, + capprox_sym RS (mem_cinfmal_iff RS iffD2)],simpset())); +qed "capprox_SComplex_not_zero"; + +Goal "[| x @c= y; y : CFinite - CInfinitesimal |] \ +\ ==> x : CFinite - CInfinitesimal"; +by (auto_tac (claset() addIs [capprox_sym RSN (2,capprox_CFinite)], + simpset() addsimps [mem_cinfmal_iff])); +by (dtac capprox_trans3 1 THEN assume_tac 1); +by (blast_tac (claset() addDs [capprox_sym]) 1); +qed "CFinite_diff_CInfinitesimal_capprox"; + +Goal "[| y ~= 0; y: CInfinitesimal; x/y : CFinite |] ==> x : CInfinitesimal"; +by (dtac CInfinitesimal_CFinite_mult2 1); +by (assume_tac 1); +by (asm_full_simp_tac + (simpset() addsimps [hcomplex_divide_def, hcomplex_mult_assoc]) 1); +qed "CInfinitesimal_ratio"; + +Goal "[|x: SComplex; y: SComplex|] ==> (x @c= y) = (x = y)"; +by Auto_tac; +by (rewrite_goals_tac [capprox_def]); +by (dres_inst_tac [("x","y")] SComplex_minus 1); +by (dtac SComplex_add 1 THEN assume_tac 1); +by (rtac (CLAIM "x - y = 0 ==> x = (y::hcomplex)") 1); +by (rtac SComplex_CInfinitesimal_zero 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_def])); +qed "SComplex_capprox_iff"; + +Goal "(number_of v @c= number_of w) = (number_of v = (number_of w :: hcomplex))"; +by (rtac SComplex_capprox_iff 1); +by Auto_tac; +qed "number_of_capprox_iff"; +Addsimps [number_of_capprox_iff]; + +Goal "(number_of w : CInfinitesimal) = (number_of w = (0::hcomplex))"; +by (rtac iffI 1); +by (fast_tac (claset() addDs [SComplex_number_of RS SComplex_CInfinitesimal_zero]) 1); +by (Asm_simp_tac 1); +qed "number_of_CInfinitesimal_iff"; +Addsimps [number_of_CInfinitesimal_iff]; + +Goal "(hcomplex_of_complex k @c= hcomplex_of_complex m) = (k = m)"; +by Auto_tac; +by (rtac (inj_hcomplex_of_complex RS injD) 1); +by (rtac (SComplex_capprox_iff RS iffD1) 1); +by Auto_tac; +qed "hcomplex_of_complex_approx_iff"; +Addsimps [hcomplex_of_complex_approx_iff]; + +Goal "(hcomplex_of_complex k @c= number_of w) = (k = number_of w)"; +by (stac (hcomplex_of_complex_approx_iff RS sym) 1); +by Auto_tac; +qed "hcomplex_of_complex_capprox_number_of_iff"; +Addsimps [hcomplex_of_complex_capprox_number_of_iff]; + +Goal "[| r: SComplex; s: SComplex; r @c= x; s @c= x|] ==> r = s"; +by (blast_tac (claset() addIs [(SComplex_capprox_iff RS iffD1), + capprox_trans2]) 1); +qed "capprox_unique_complex"; + +Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) @c= Abs_hcomplex(hcomplexrel``{%n. Y n}) \ +\ ==> Abs_hypreal(hyprel `` {%n. Re(X n)}) @= \ +\ Abs_hypreal(hyprel `` {%n. Re(Y n)})"; +by (auto_tac (claset(),simpset() addsimps [approx_FreeUltrafilterNat_iff])); +by (dtac (capprox_minus_iff RS iffD1) 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_minus,hcomplex_add, + mem_cinfmal_iff RS sym,CInfinitesimal_hcmod_iff,hcmod, + Infinitesimal_FreeUltrafilterNat_iff2])); +by (dres_inst_tac [("x","m")] spec 1); +by (Ultra_tac 1); +by (res_inst_tac [("z","X x")] eq_Abs_complex 1); +by (res_inst_tac [("z","Y x")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_minus,complex_add, + complex_mod] delsimps [realpow_Suc])); +by (rtac order_le_less_trans 1 THEN assume_tac 2); +by (dres_inst_tac [("t","Ya x")] sym 1); +by (auto_tac (claset(),simpset() addsimps [abs_eqI1] delsimps [realpow_Suc])); +qed "hcomplex_capproxD1"; + +(* same proof *) +Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) @c= Abs_hcomplex(hcomplexrel``{%n. Y n}) \ +\ ==> Abs_hypreal(hyprel `` {%n. Im(X n)}) @= \ +\ Abs_hypreal(hyprel `` {%n. Im(Y n)})"; +by (auto_tac (claset(),simpset() addsimps [approx_FreeUltrafilterNat_iff])); +by (dtac (capprox_minus_iff RS iffD1) 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_minus,hcomplex_add, + mem_cinfmal_iff RS sym,CInfinitesimal_hcmod_iff,hcmod, + Infinitesimal_FreeUltrafilterNat_iff2])); +by (dres_inst_tac [("x","m")] spec 1); +by (Ultra_tac 1); +by (res_inst_tac [("z","X x")] eq_Abs_complex 1); +by (res_inst_tac [("z","Y x")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_minus,complex_add, + complex_mod] delsimps [realpow_Suc])); +by (rtac order_le_less_trans 1 THEN assume_tac 2); +by (dres_inst_tac [("t","Ya x")] sym 1); +by (auto_tac (claset(),simpset() addsimps [abs_eqI1] delsimps [realpow_Suc])); +qed "hcomplex_capproxD2"; + +Goal "[| Abs_hypreal(hyprel `` {%n. Re(X n)}) @= \ +\ Abs_hypreal(hyprel `` {%n. Re(Y n)}); \ +\ Abs_hypreal(hyprel `` {%n. Im(X n)}) @= \ +\ Abs_hypreal(hyprel `` {%n. Im(Y n)}) \ +\ |] ==> Abs_hcomplex(hcomplexrel ``{%n. X n}) @c= Abs_hcomplex(hcomplexrel``{%n. Y n})"; +by (dtac (approx_minus_iff RS iffD1) 1); +by (dtac (approx_minus_iff RS iffD1) 1); +by (rtac (capprox_minus_iff RS iffD2) 1); +by (auto_tac (claset(),simpset() addsimps [mem_cinfmal_iff RS sym, + mem_infmal_iff RS sym,hypreal_minus,hypreal_add,hcomplex_minus, + hcomplex_add,CInfinitesimal_hcmod_iff,hcmod,Infinitesimal_FreeUltrafilterNat_iff])); +by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2); +by Auto_tac; +by (dres_inst_tac [("x","u/2")] spec 1); +by (dres_inst_tac [("x","u/2")] spec 1); +by (Step_tac 1); +by (TRYALL(Force_tac)); +by (ultra_tac (claset(),HOL_ss) 1); +by (dtac sym 1 THEN dtac sym 1); +by (res_inst_tac [("z","X x")] eq_Abs_complex 1); +by (res_inst_tac [("z","Y x")] eq_Abs_complex 1); +by (auto_tac (claset(),HOL_ss addsimps [complex_minus,complex_add, + complex_mod,pair_mem_complex RS Abs_complex_inverse,snd_conv, + fst_conv,two_eq_Suc_Suc])); +by (rtac (realpow_two_abs RS subst) 1); +by (res_inst_tac [("x1","xa + - xb")] (realpow_two_abs RS subst) 1); +by (simp_tac (simpset() addsimps [two_eq_Suc_Suc RS sym]) 1); +by (rtac lemma_sqrt_hcomplex_capprox 1); +by Auto_tac; +qed "hcomplex_capproxI"; + +Goal "(Abs_hcomplex(hcomplexrel ``{%n. X n}) @c= Abs_hcomplex(hcomplexrel``{%n. Y n})) =\ +\ (Abs_hypreal(hyprel `` {%n. Re(X n)}) @= Abs_hypreal(hyprel `` {%n. Re(Y n)}) & \ +\ Abs_hypreal(hyprel `` {%n. Im(X n)}) @= Abs_hypreal(hyprel `` {%n. Im(Y n)}))"; +by (blast_tac (claset() addIs [hcomplex_capproxI,hcomplex_capproxD1,hcomplex_capproxD2]) 1); +qed "capprox_approx_iff"; + +Goal "(hcomplex_of_hypreal x @c= hcomplex_of_hypreal z) = (x @= z)"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + capprox_approx_iff])); +qed "hcomplex_of_hypreal_capprox_iff"; +Addsimps [hcomplex_of_hypreal_capprox_iff]; + +Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) : CFinite \ +\ ==> Abs_hypreal(hyprel `` {%n. Re(X n)}) : HFinite"; +by (auto_tac (claset(),simpset() addsimps [CFinite_hcmod_iff, + hcmod,HFinite_FreeUltrafilterNat_iff])); +by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2); +by (res_inst_tac [("x","u")] exI 1 THEN Auto_tac); +by (Ultra_tac 1); +by (dtac sym 1 THEN res_inst_tac [("z","X x")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mod,two_eq_Suc_Suc] delsimps [realpow_Suc])); +by (rtac ccontr 1 THEN dtac real_leI 1); +by (dtac order_less_le_trans 1 THEN assume_tac 1); +by (dtac (real_sqrt_ge_abs1 RSN (2,order_less_le_trans)) 1); +by (auto_tac (claset(),simpset() addsimps [two_eq_Suc_Suc RS sym])); +qed "CFinite_HFinite_Re"; + +Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) : CFinite \ +\ ==> Abs_hypreal(hyprel `` {%n. Im(X n)}) : HFinite"; +by (auto_tac (claset(),simpset() addsimps [CFinite_hcmod_iff, + hcmod,HFinite_FreeUltrafilterNat_iff])); +by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2); +by (res_inst_tac [("x","u")] exI 1 THEN Auto_tac); +by (Ultra_tac 1); +by (dtac sym 1 THEN res_inst_tac [("z","X x")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mod] delsimps [realpow_Suc])); +by (rtac ccontr 1 THEN dtac real_leI 1); +by (dtac order_less_le_trans 1 THEN assume_tac 1); +by (dtac (real_sqrt_ge_abs2 RSN (2,order_less_le_trans)) 1); +by Auto_tac; +qed "CFinite_HFinite_Im"; + +Goal "[| Abs_hypreal(hyprel `` {%n. Re(X n)}) : HFinite; \ +\ Abs_hypreal(hyprel `` {%n. Im(X n)}) : HFinite \ +\ |] ==> Abs_hcomplex(hcomplexrel ``{%n. X n}) : CFinite"; +by (auto_tac (claset(),simpset() addsimps [CFinite_hcmod_iff, + hcmod,HFinite_FreeUltrafilterNat_iff])); +by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2); +by (res_inst_tac [("x","2*(u + ua)")] exI 1); +by (Ultra_tac 1); +by (dtac sym 1 THEN res_inst_tac [("z","X x")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_mod,two_eq_Suc_Suc] delsimps [realpow_Suc])); +by (subgoal_tac "0 < u" 1 THEN arith_tac 2); +by (subgoal_tac "0 < ua" 1 THEN arith_tac 2); +by (rtac (realpow_two_abs RS subst) 1); +by (res_inst_tac [("x1","Y x")] (realpow_two_abs RS subst) 1); +by (simp_tac (simpset() addsimps [two_eq_Suc_Suc RS sym]) 1); +by (rtac lemma_sqrt_hcomplex_capprox 1); +by Auto_tac; +qed "HFinite_Re_Im_CFinite"; + +Goal "(Abs_hcomplex(hcomplexrel ``{%n. X n}) : CFinite) = \ +\ (Abs_hypreal(hyprel `` {%n. Re(X n)}) : HFinite & \ +\ Abs_hypreal(hyprel `` {%n. Im(X n)}) : HFinite)"; +by (blast_tac (claset() addIs [HFinite_Re_Im_CFinite,CFinite_HFinite_Im, + CFinite_HFinite_Re]) 1); +qed "CFinite_HFinite_iff"; + +Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) : SComplex \ +\ ==> Abs_hypreal(hyprel `` {%n. Re(X n)}) : Reals"; +by (auto_tac (claset(),simpset() addsimps [SComplex_def, + hcomplex_of_complex_def,SReal_def,hypreal_of_real_def])); +by (res_inst_tac [("x","Re r")] exI 1); +by (Ultra_tac 1); +qed "SComplex_Re_SReal"; + +Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) : SComplex \ +\ ==> Abs_hypreal(hyprel `` {%n. Im(X n)}) : Reals"; +by (auto_tac (claset(),simpset() addsimps [SComplex_def, + hcomplex_of_complex_def,SReal_def,hypreal_of_real_def])); +by (res_inst_tac [("x","Im r")] exI 1); +by (Ultra_tac 1); +qed "SComplex_Im_SReal"; + +Goal "[| Abs_hypreal(hyprel `` {%n. Re(X n)}) : Reals; \ +\ Abs_hypreal(hyprel `` {%n. Im(X n)}) : Reals \ +\ |] ==> Abs_hcomplex(hcomplexrel ``{%n. X n}) : SComplex"; +by (auto_tac (claset(),simpset() addsimps [SComplex_def, + hcomplex_of_complex_def,SReal_def,hypreal_of_real_def])); +by (res_inst_tac [("x","complex_of_real r + ii * complex_of_real ra")] exI 1); +by (Ultra_tac 1); +by (res_inst_tac [("z","X x")] eq_Abs_complex 1); +by (auto_tac (claset(),simpset() addsimps [complex_of_real_def,i_def, + complex_add,complex_mult])); +qed "Reals_Re_Im_SComplex"; + +Goal "(Abs_hcomplex(hcomplexrel ``{%n. X n}) : SComplex) = \ +\ (Abs_hypreal(hyprel `` {%n. Re(X n)}) : Reals & \ +\ Abs_hypreal(hyprel `` {%n. Im(X n)}) : Reals)"; +by (blast_tac (claset() addIs [SComplex_Re_SReal,SComplex_Im_SReal, + Reals_Re_Im_SComplex]) 1); +qed "SComplex_SReal_iff"; + +Goal "(Abs_hcomplex(hcomplexrel ``{%n. X n}) : CInfinitesimal) = \ +\ (Abs_hypreal(hyprel `` {%n. Re(X n)}) : Infinitesimal & \ +\ Abs_hypreal(hyprel `` {%n. Im(X n)}) : Infinitesimal)"; +by (auto_tac (claset(),simpset() addsimps [mem_cinfmal_iff, + mem_infmal_iff,hcomplex_zero_num,hypreal_zero_num,capprox_approx_iff])); +qed "CInfinitesimal_Infinitesimal_iff"; + +(*** more lemmas ****) +Goal "(EX t. P t) = (EX X. P (Abs_hcomplex(hcomplexrel `` {X})))"; +by Auto_tac; +by (res_inst_tac [("z","t")] eq_Abs_hcomplex 1); +by Auto_tac; +qed "eq_Abs_hcomplex_EX"; + +Goal "(EX t : A. P t) = (EX X. (Abs_hcomplex(hcomplexrel `` {X})) : A & \ +\ P (Abs_hcomplex(hcomplexrel `` {X})))"; +by Auto_tac; +by (res_inst_tac [("z","t")] eq_Abs_hcomplex 1); +by Auto_tac; +qed "eq_Abs_hcomplex_Bex"; + +(* Here we go - easy proof now!! *) +Goal "x:CFinite ==> EX t: SComplex. x @c= t"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [CFinite_HFinite_iff, + eq_Abs_hcomplex_Bex,SComplex_SReal_iff,capprox_approx_iff])); +by (REPEAT(dtac st_part_Ex 1 THEN Step_tac 1)); +by (res_inst_tac [("z","t")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","ta")] eq_Abs_hypreal 1); +by Auto_tac; +by (res_inst_tac [("x","%n. complex_of_real (xa n) + ii * complex_of_real (xb n)")] + exI 1); +by Auto_tac; +qed "stc_part_Ex"; + +Goal "x:CFinite ==> EX! t. t : SComplex & x @c= t"; +by (dtac stc_part_Ex 1 THEN Step_tac 1); +by (dtac capprox_sym 2 THEN dtac capprox_sym 2 + THEN dtac capprox_sym 2); +by (auto_tac (claset() addSIs [capprox_unique_complex], simpset())); +qed "stc_part_Ex1"; + +Goalw [CFinite_def,CInfinite_def] "CFinite Int CInfinite = {}"; +by Auto_tac; +qed "CFinite_Int_CInfinite_empty"; +Addsimps [CFinite_Int_CInfinite_empty]; + +Goal "x: CFinite ==> x ~: CInfinite"; +by (EVERY1[Step_tac, dtac IntI, assume_tac]); +by Auto_tac; +qed "CFinite_not_CInfinite"; + +Goal "x~: CFinite ==> x: CInfinite"; +by (auto_tac (claset() addIs [not_HFinite_HInfinite], + simpset() addsimps [CFinite_hcmod_iff,CInfinite_hcmod_iff])); +qed "not_CFinite_CInfinite"; + +Goal "x : CInfinite | x : CFinite"; +by (blast_tac (claset() addIs [not_CFinite_CInfinite]) 1); +qed "CInfinite_CFinite_disj"; + +Goal "(x : CInfinite) = (x ~: CFinite)"; +by (blast_tac (claset() addDs [CFinite_not_CInfinite, + not_CFinite_CInfinite]) 1); +qed "CInfinite_CFinite_iff"; + +Goal "(x : CFinite) = (x ~: CInfinite)"; +by (simp_tac (simpset() addsimps [CInfinite_CFinite_iff]) 1); +qed "CFinite_CInfinite_iff"; + +Goal "x ~: CInfinitesimal ==> x : CInfinite | x : CFinite - CInfinitesimal"; +by (fast_tac (claset() addIs [not_CFinite_CInfinite]) 1); +qed "CInfinite_diff_CFinite_CInfinitesimal_disj"; + +Goal "[| x : CFinite; x ~: CInfinitesimal |] ==> inverse x : CFinite"; +by (cut_inst_tac [("x","inverse x")] CInfinite_CFinite_disj 1); +by (auto_tac (claset() addSDs [CInfinite_inverse_CInfinitesimal], simpset())); +qed "CFinite_inverse"; + +Goal "x : CFinite - CInfinitesimal ==> inverse x : CFinite"; +by (blast_tac (claset() addIs [CFinite_inverse]) 1); +qed "CFinite_inverse2"; + +Goal "x ~: CInfinitesimal ==> inverse(x) : CFinite"; +by (dtac CInfinite_diff_CFinite_CInfinitesimal_disj 1); +by (blast_tac (claset() addIs [CFinite_inverse, + CInfinite_inverse_CInfinitesimal, + CInfinitesimal_subset_CFinite RS subsetD]) 1); +qed "CInfinitesimal_inverse_CFinite"; + + +Goal "x : CFinite - CInfinitesimal ==> inverse x : CFinite - CInfinitesimal"; +by (auto_tac (claset() addIs [CInfinitesimal_inverse_CFinite], simpset())); +by (dtac CInfinitesimal_CFinite_mult2 1); +by (assume_tac 1); +by (asm_full_simp_tac (simpset() addsimps [not_CInfinitesimal_not_zero]) 1); +qed "CFinite_not_CInfinitesimal_inverse"; + +Goal "[| x @c= y; y : CFinite - CInfinitesimal |] \ +\ ==> inverse x @c= inverse y"; +by (forward_tac [CFinite_diff_CInfinitesimal_capprox] 1); +by (assume_tac 1); +by (forward_tac [not_CInfinitesimal_not_zero2] 1); +by (forw_inst_tac [("x","x")] not_CInfinitesimal_not_zero2 1); +by (REPEAT(dtac CFinite_inverse2 1)); +by (dtac capprox_mult2 1 THEN assume_tac 1); +by Auto_tac; +by (dres_inst_tac [("c","inverse x")] capprox_mult1 1 + THEN assume_tac 1); +by (auto_tac (claset() addIs [capprox_sym], + simpset() addsimps [hcomplex_mult_assoc])); +qed "capprox_inverse"; + +bind_thm ("hcomplex_of_complex_capprox_inverse", + hcomplex_of_complex_CFinite_diff_CInfinitesimal RSN (2, capprox_inverse)); + +Goal "[| x: CFinite - CInfinitesimal; \ +\ h : CInfinitesimal |] ==> inverse(x + h) @c= inverse x"; +by (auto_tac (claset() addIs [capprox_inverse, capprox_sym, + CInfinitesimal_add_capprox_self], + simpset())); +qed "inverse_add_CInfinitesimal_capprox"; + +Goal "[| x: CFinite - CInfinitesimal; \ +\ h : CInfinitesimal |] ==> inverse(h + x) @c= inverse x"; +by (rtac (hcomplex_add_commute RS subst) 1); +by (blast_tac (claset() addIs [inverse_add_CInfinitesimal_capprox]) 1); +qed "inverse_add_CInfinitesimal_capprox2"; + +Goal "[| x : CFinite - CInfinitesimal; \ +\ h : CInfinitesimal |] ==> inverse(x + h) - inverse x @c= h"; +by (rtac capprox_trans2 1); +by (auto_tac (claset() addIs [inverse_add_CInfinitesimal_capprox], + simpset() addsimps [mem_cinfmal_iff,hcomplex_diff_def, + capprox_minus_iff RS sym])); +qed "inverse_add_CInfinitesimal_approx_CInfinitesimal"; + +Goal "(x*x : CInfinitesimal) = (x : CInfinitesimal)"; +by (auto_tac (claset(), simpset() addsimps [CInfinitesimal_hcmod_iff, + hcmod_mult])); +qed "CInfinitesimal_square_iff"; +AddIffs [CInfinitesimal_square_iff]; + +Goal "[| a: CFinite-CInfinitesimal; a*w @c= a*z |] ==> w @c= z"; +by (Step_tac 1); +by (ftac CFinite_inverse 1 THEN assume_tac 1); +by (dtac not_CInfinitesimal_not_zero 1); +by (auto_tac (claset() addDs [capprox_mult2], + simpset() addsimps [hcomplex_mult_assoc RS sym])); +qed "capprox_CFinite_mult_cancel"; + +Goal "a: CFinite-CInfinitesimal ==> (a * w @c= a * z) = (w @c= z)"; +by (auto_tac (claset() addIs [capprox_mult2, + capprox_CFinite_mult_cancel], simpset())); +qed "capprox_CFinite_mult_cancel_iff1"; + + +(*---------------------------------------------------------------------------*) +(* Theorems about monads *) +(*---------------------------------------------------------------------------*) + +Goalw [cmonad_def] "(x @c= y) = (cmonad(x)=cmonad(y))"; +by (auto_tac (claset() addDs [capprox_sym] + addSEs [capprox_trans,equalityCE], + simpset())); +qed "capprox_cmonad_iff"; + +Goal "e : CInfinitesimal ==> cmonad (x+e) = cmonad x"; +by (fast_tac (claset() addSIs [CInfinitesimal_add_capprox_self RS capprox_sym, + capprox_cmonad_iff RS iffD1]) 1); +qed "CInfinitesimal_cmonad_eq"; + +Goalw [cmonad_def] "(u:cmonad x) = (-u:cmonad (-x))"; +by Auto_tac; +qed "mem_cmonad_iff"; + +Goalw [cmonad_def] "(x:CInfinitesimal) = (x:cmonad 0)"; +by (auto_tac (claset() addIs [capprox_sym], + simpset() addsimps [mem_cinfmal_iff])); +qed "CInfinitesimal_cmonad_zero_iff"; + +Goal "(x:cmonad 0) = (-x:cmonad 0)"; +by (simp_tac (simpset() addsimps [CInfinitesimal_cmonad_zero_iff RS sym]) 1); +qed "cmonad_zero_minus_iff"; + +Goal "(x:cmonad 0) = (hcmod x:monad 0)"; +by (auto_tac (claset(), simpset() addsimps + [CInfinitesimal_cmonad_zero_iff RS sym, + CInfinitesimal_hcmod_iff,Infinitesimal_monad_zero_iff RS sym])); +qed "cmonad_zero_hcmod_iff"; + +Goalw [cmonad_def] "x:cmonad x"; +by (Simp_tac 1); +qed "mem_cmonad_self"; +Addsimps [mem_cmonad_self]; + +(*---------------------------------------------------------------------------*) +(* Theorems about standard part *) +(*---------------------------------------------------------------------------*) +Goalw [stc_def] "x: CFinite ==> stc x @c= x"; +by (forward_tac [stc_part_Ex] 1 THEN Step_tac 1); +by (rtac someI2 1); +by (auto_tac (claset() addIs [capprox_sym], simpset())); +qed "stc_capprox_self"; + +Goalw [stc_def] "x: CFinite ==> stc x: SComplex"; +by (forward_tac [stc_part_Ex] 1 THEN Step_tac 1); +by (rtac someI2 1); +by (auto_tac (claset() addIs [capprox_sym], simpset())); +qed "stc_SComplex"; + +Goal "x: CFinite ==> stc x: CFinite"; +by (etac (stc_SComplex RS (SComplex_subset_CFinite RS subsetD)) 1); +qed "stc_CFinite"; + +Goalw [stc_def] "x: SComplex ==> stc x = x"; +by (rtac some_equality 1); +by (auto_tac (claset() addIs [(SComplex_subset_CFinite RS subsetD)],simpset())); +by (blast_tac (claset() addDs [SComplex_capprox_iff RS iffD1]) 1); +qed "stc_SComplex_eq"; +Addsimps [stc_SComplex_eq]; + +Goal "stc (hcomplex_of_complex x) = hcomplex_of_complex x"; +by Auto_tac; +qed "stc_hcomplex_of_complex"; + +Goal "[| x: CFinite; y: CFinite; stc x = stc y |] ==> x @c= y"; +by (auto_tac (claset() addSDs [stc_capprox_self] + addSEs [capprox_trans3], simpset())); +qed "stc_eq_capprox"; + +Goal "[| x: CFinite; y: CFinite; x @c= y |] ==> stc x = stc y"; +by (EVERY1 [forward_tac [stc_capprox_self], + forw_inst_tac [("x","y")] stc_capprox_self, + dtac stc_SComplex,dtac stc_SComplex]); +by (fast_tac (claset() addEs [capprox_trans, + capprox_trans2,SComplex_capprox_iff RS iffD1]) 1); +qed "capprox_stc_eq"; + +Goal "[| x: CFinite; y: CFinite|] ==> (x @c= y) = (stc x = stc y)"; +by (blast_tac (claset() addIs [capprox_stc_eq,stc_eq_capprox]) 1); +qed "stc_eq_capprox_iff"; + +Goal "[| x: SComplex; e: CInfinitesimal |] ==> stc(x + e) = x"; +by (forward_tac [stc_SComplex_eq RS subst] 1); +by (assume_tac 2); +by (forward_tac [SComplex_subset_CFinite RS subsetD] 1); +by (forward_tac [CInfinitesimal_subset_CFinite RS subsetD] 1); +by (dtac stc_SComplex_eq 1); +by (rtac capprox_stc_eq 1); +by (auto_tac (claset() addIs [CFinite_add], + simpset() addsimps [CInfinitesimal_add_capprox_self + RS capprox_sym])); +qed "stc_CInfinitesimal_add_SComplex"; + +Goal "[| x: SComplex; e: CInfinitesimal |] ==> stc(e + x) = x"; +by (rtac (hcomplex_add_commute RS subst) 1); +by (blast_tac (claset() addSIs [stc_CInfinitesimal_add_SComplex]) 1); +qed "stc_CInfinitesimal_add_SComplex2"; + +Goal "x: CFinite ==> EX e: CInfinitesimal. x = stc(x) + e"; +by (blast_tac (claset() addSDs [(stc_capprox_self RS + capprox_sym),bex_CInfinitesimal_iff2 RS iffD2]) 1); +qed "CFinite_stc_CInfinitesimal_add"; + +Goal "[| x: CFinite; y: CFinite |] ==> stc (x + y) = stc(x) + stc(y)"; +by (forward_tac [CFinite_stc_CInfinitesimal_add] 1); +by (forw_inst_tac [("x","y")] CFinite_stc_CInfinitesimal_add 1); +by (Step_tac 1); +by (subgoal_tac "stc (x + y) = stc ((stc x + e) + (stc y + ea))" 1); +by (dtac sym 2 THEN dtac sym 2); +by (Asm_full_simp_tac 2); +by (asm_simp_tac (simpset() addsimps hcomplex_add_ac) 1); +by (REPEAT(dtac stc_SComplex 1)); +by (dtac SComplex_add 1 THEN assume_tac 1); +by (dtac CInfinitesimal_add 1 THEN assume_tac 1); +by (rtac (hcomplex_add_assoc RS subst) 1); +by (blast_tac (claset() addSIs [stc_CInfinitesimal_add_SComplex2]) 1); +qed "stc_add"; + +Goal "stc (number_of w) = number_of w"; +by (rtac (SComplex_number_of RS stc_SComplex_eq) 1); +qed "stc_number_of"; +Addsimps [stc_number_of]; + +Goal "stc 0 = 0"; +by (Simp_tac 1); +qed "stc_zero"; +Addsimps [stc_zero]; + +Goal "stc 1 = 1"; +by (Simp_tac 1); +qed "stc_one"; +Addsimps [stc_one]; + +Goal "y: CFinite ==> stc(-y) = -stc(y)"; +by (forward_tac [CFinite_minus_iff RS iffD2] 1); +by (rtac hcomplex_add_minus_eq_minus 1); +by (dtac (stc_add RS sym) 1 THEN assume_tac 1); +by Auto_tac; +qed "stc_minus"; + +Goalw [hcomplex_diff_def] + "[| x: CFinite; y: CFinite |] ==> stc (x-y) = stc(x) - stc(y)"; +by (forw_inst_tac [("y1","y")] (stc_minus RS sym) 1); +by (dres_inst_tac [("x1","y")] (CFinite_minus_iff RS iffD2) 1); +by (auto_tac (claset() addIs [stc_add],simpset())); +qed "stc_diff"; + +Goal "[| x: CFinite; y: CFinite; \ +\ e: CInfinitesimal; \ +\ ea: CInfinitesimal |] \ +\ ==> e*y + x*ea + e*ea: CInfinitesimal"; +by (forw_inst_tac [("x","e"),("y","y")] CInfinitesimal_CFinite_mult 1); +by (forw_inst_tac [("x","ea"),("y","x")] CInfinitesimal_CFinite_mult 2); +by (dtac CInfinitesimal_mult 3); +by (auto_tac (claset() addIs [CInfinitesimal_add], + simpset() addsimps hcomplex_add_ac @ hcomplex_mult_ac)); +qed "lemma_stc_mult"; + +Goal "[| x: CFinite; y: CFinite |] \ +\ ==> stc (x * y) = stc(x) * stc(y)"; +by (forward_tac [CFinite_stc_CInfinitesimal_add] 1); +by (forw_inst_tac [("x","y")] CFinite_stc_CInfinitesimal_add 1); +by (Step_tac 1); +by (subgoal_tac "stc (x * y) = stc ((stc x + e) * (stc y + ea))" 1); +by (dtac sym 2 THEN dtac sym 2); +by (Asm_full_simp_tac 2); +by (thin_tac "x = stc x + e" 1); +by (thin_tac "y = stc y + ea" 1); +by (asm_full_simp_tac (simpset() addsimps + [hcomplex_add_mult_distrib,hcomplex_add_mult_distrib2]) 1); +by (REPEAT(dtac stc_SComplex 1)); +by (full_simp_tac (simpset() addsimps [hcomplex_add_assoc]) 1); +by (rtac stc_CInfinitesimal_add_SComplex 1); +by (blast_tac (claset() addSIs [SComplex_mult]) 1); +by (REPEAT(dtac (SComplex_subset_CFinite RS subsetD) 1)); +by (rtac (hcomplex_add_assoc RS subst) 1); +by (blast_tac (claset() addSIs [lemma_stc_mult]) 1); +qed "stc_mult"; + +Goal "x: CInfinitesimal ==> stc x = 0"; +by (rtac (stc_zero RS subst) 1); +by (rtac capprox_stc_eq 1); +by (auto_tac (claset() addIs [CInfinitesimal_subset_CFinite RS subsetD], + simpset() addsimps [mem_cinfmal_iff RS sym])); +qed "stc_CInfinitesimal"; + +Goal "stc(x) ~= 0 ==> x ~: CInfinitesimal"; +by (fast_tac (claset() addIs [stc_CInfinitesimal]) 1); +qed "stc_not_CInfinitesimal"; + +Goal "[| x: CFinite; stc x ~= 0 |] \ +\ ==> stc(inverse x) = inverse (stc x)"; +by (res_inst_tac [("c1","stc x")] (hcomplex_mult_left_cancel RS iffD1) 1); +by (auto_tac (claset(), + simpset() addsimps [stc_mult RS sym, stc_not_CInfinitesimal, + CFinite_inverse])); +by (stac hcomplex_mult_inv_right 1); +by Auto_tac; +qed "stc_inverse"; + +Goal "[| x: CFinite; y: CFinite; stc y ~= 0 |] \ +\ ==> stc(x/y) = (stc x) / (stc y)"; +by (auto_tac (claset(), + simpset() addsimps [hcomplex_divide_def, stc_mult, stc_not_CInfinitesimal, + CFinite_inverse, stc_inverse])); +qed "stc_divide"; +Addsimps [stc_divide]; + +Goal "x: CFinite ==> stc(stc(x)) = stc(x)"; +by (blast_tac (claset() addIs [stc_CFinite, stc_capprox_self, + capprox_stc_eq]) 1); +qed "stc_idempotent"; +Addsimps [stc_idempotent]; + +Goal "z : HFinite ==> hcomplex_of_hypreal z : CFinite"; +by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + CFinite_HFinite_iff,symmetric hypreal_zero_def])); +qed "CFinite_HFinite_hcomplex_of_hypreal"; + +Goal "x : Reals ==> hcomplex_of_hypreal x : SComplex"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + SComplex_SReal_iff,symmetric hypreal_zero_def])); +qed "SComplex_SReal_hcomplex_of_hypreal"; + +Goalw [st_def,stc_def] + "z : HFinite ==> stc(hcomplex_of_hypreal z) = hcomplex_of_hypreal (st z)"; +by (ftac st_part_Ex 1 THEN Step_tac 1); +by (rtac someI2 1); +by (auto_tac (claset() addIs [approx_sym],simpset())); +by (dtac CFinite_HFinite_hcomplex_of_hypreal 1); +by (ftac stc_part_Ex 1 THEN Step_tac 1); +by (rtac someI2 1); +by (auto_tac (claset() addIs [capprox_sym] addSIs [capprox_unique_complex] + addDs [SComplex_SReal_hcomplex_of_hypreal],simpset())); +qed "stc_hcomplex_of_hypreal"; + +(* +Goal "x: CFinite ==> hcmod(stc x) = st(hcmod x)"; +by (dtac stc_capprox_self 1); +by (auto_tac (claset(),simpset() addsimps [bex_CInfinitesimal_iff2 RS sym])); + + +approx_hcmod_add_hcmod +*) + +(*---------------------------------------------------------------------------*) +(* More nonstandard complex specific theorems *) +(*---------------------------------------------------------------------------*) +Goal "(hcnj z : CInfinitesimal) = (z : CInfinitesimal)"; +by (auto_tac (claset(),simpset() addsimps [CInfinitesimal_hcmod_iff])); +qed "CInfinitesimal_hcnj_iff"; +Addsimps [CInfinitesimal_hcnj_iff]; + +Goal "(Abs_hcomplex(hcomplexrel ``{%n. X n}) : CInfinite) = \ +\ (Abs_hypreal(hyprel `` {%n. Re(X n)}) : HInfinite | \ +\ Abs_hypreal(hyprel `` {%n. Im(X n)}) : HInfinite)"; +by (auto_tac (claset(),simpset() addsimps [CInfinite_CFinite_iff, + HInfinite_HFinite_iff,CFinite_HFinite_iff])); +qed "CInfinite_HInfinite_iff"; + +Goal "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y : CInfinitesimal) = \ +\ (x : Infinitesimal & y : Infinitesimal)"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [iii_def,hcomplex_add,hcomplex_mult, + hcomplex_of_hypreal,CInfinitesimal_Infinitesimal_iff])); +qed "hcomplex_split_CInfinitesimal_iff"; + +Goal "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y : CFinite) = \ +\ (x : HFinite & y : HFinite)"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [iii_def,hcomplex_add,hcomplex_mult, + hcomplex_of_hypreal,CFinite_HFinite_iff])); +qed "hcomplex_split_CFinite_iff"; + +Goal "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y : SComplex) = \ +\ (x : Reals & y : Reals)"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [iii_def,hcomplex_add,hcomplex_mult, + hcomplex_of_hypreal,SComplex_SReal_iff])); +qed "hcomplex_split_SComplex_iff"; + +Goal "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y : CInfinite) = \ +\ (x : HInfinite | y : HInfinite)"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [iii_def,hcomplex_add,hcomplex_mult, + hcomplex_of_hypreal,CInfinite_HInfinite_iff])); +qed "hcomplex_split_CInfinite_iff"; + +Goal "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y @c= \ +\ hcomplex_of_hypreal x' + iii * hcomplex_of_hypreal y') = \ +\ (x @= x' & y @= y')"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","x'")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y'")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [iii_def,hcomplex_add,hcomplex_mult, + hcomplex_of_hypreal,capprox_approx_iff])); +qed "hcomplex_split_capprox_iff"; + +(*** More theorems ***) + +Goal "ALL n. cmod (X n - x) < inverse (real (Suc n)) ==> \ +\ Abs_hcomplex(hcomplexrel``{X}) - hcomplex_of_complex x : CInfinitesimal"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_diff, + CInfinitesimal_hcmod_iff,hcomplex_of_complex_def, + Infinitesimal_FreeUltrafilterNat_iff,hcmod])); +by (rtac bexI 1 THEN Auto_tac); +by (auto_tac (claset() addDs [FreeUltrafilterNat_inverse_real_of_posnat, + FreeUltrafilterNat_all,FreeUltrafilterNat_Int] + addIs [order_less_trans, FreeUltrafilterNat_subset], + simpset())); +qed "complex_seq_to_hcomplex_CInfinitesimal"; + +Goal "hcomplex_of_hypreal epsilon : CInfinitesimal"; +by (auto_tac (claset(),simpset() addsimps [CInfinitesimal_hcmod_iff])); +qed "CInfinitesimal_hcomplex_of_hypreal_epsilon"; +Addsimps [CInfinitesimal_hcomplex_of_hypreal_epsilon]; + +Goal "(hcomplex_of_complex z @c= 0) = (z = 0)"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_complex_zero RS sym] + delsimps [hcomplex_of_complex_zero])); +qed "hcomplex_of_complex_approx_zero_iff"; + +Goal "(0 @c= hcomplex_of_complex z) = (z = 0)"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_complex_zero RS sym] + delsimps [hcomplex_of_complex_zero])); +qed "hcomplex_of_complex_approx_zero_iff2"; + +Addsimps [hcomplex_of_complex_approx_zero_iff,hcomplex_of_complex_approx_zero_iff2]; + + diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/NSCA.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/NSCA.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,44 @@ +(* Title : NSCA.thy + Author : Jacques D. Fleuriot + Copyright : 2001,2002 University of Edinburgh + Description : Infinite, infinitesimal complex number etc! +*) + +NSCA = NSComplexArith0 + + +consts + + (* infinitely close *) + "@c=" :: [hcomplex,hcomplex] => bool (infixl 50) + + +constdefs + (* standard complex numbers reagarded as an embedded subset of NS complex *) + SComplex :: "hcomplex set" + "SComplex == {x. EX r. x = hcomplex_of_complex r}" + + CInfinitesimal :: "hcomplex set" + "CInfinitesimal == {x. ALL r: Reals. 0 < r --> hcmod x < r}" + + CFinite :: "hcomplex set" + "CFinite == {x. EX r: Reals. hcmod x < r}" + + CInfinite :: "hcomplex set" + "CInfinite == {x. ALL r: Reals. r < hcmod x}" + + (* standard part map *) + stc :: hcomplex => hcomplex + "stc x == (@r. x : CFinite & r:SComplex & r @c= x)" + + cmonad :: hcomplex => hcomplex set + "cmonad x == {y. x @c= y}" + + cgalaxy :: hcomplex => hcomplex set + "cgalaxy x == {y. (x - y) : CFinite}" + + +defs + + capprox_def "x @c= y == (x - y) : CInfinitesimal" + +end diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/NSComplex.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/NSComplex.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,1895 @@ +(* Title: NSComplex.ML + Author: Jacques D. Fleuriot + Copyhright: 2001 University of Edinburgh + Description: Nonstandard Complex numbers +*) + +Goalw [hcomplexrel_def] + "((X,Y): hcomplexrel) = ({n. X n = Y n}: FreeUltrafilterNat)"; +by (Fast_tac 1); +qed "hcomplexrel_iff"; + +Goalw [hcomplexrel_def] + "!!X. {n. X n = Y n}: FreeUltrafilterNat \ +\ ==> (X,Y): hcomplexrel"; +by (Fast_tac 1); +qed "hcomplexrelI"; + +Goalw [hcomplexrel_def] + "p: hcomplexrel --> (EX X Y. \ +\ p = (X,Y) & {n. X n = Y n} : FreeUltrafilterNat)"; +by (Fast_tac 1); +qed "hcomplexrelE_lemma"; + +val [major,minor] = goal thy + "[| p: hcomplexrel; \ +\ !!X Y. [| p = (X,Y); {n. X n = Y n}: FreeUltrafilterNat\ +\ |] ==> Q |] ==> Q"; +by (cut_facts_tac [major RS (hcomplexrelE_lemma RS mp)] 1); +by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1)); +qed "hcomplexrelE"; + +AddSIs [hcomplexrelI]; +AddSEs [hcomplexrelE]; + +Goalw [hcomplexrel_def] "(x,x): hcomplexrel"; +by (Auto_tac); +qed "hcomplexrel_refl"; + +Goalw [hcomplexrel_def] "(x,y): hcomplexrel --> (y,x):hcomplexrel"; +by (auto_tac (claset() addIs [lemma_perm RS subst],simpset())); +qed_spec_mp "hcomplexrel_sym"; + +Goalw [hcomplexrel_def] + "(x,y): hcomplexrel --> (y,z):hcomplexrel --> (x,z):hcomplexrel"; +by (Auto_tac); +by (Ultra_tac 1); +qed_spec_mp "hcomplexrel_trans"; + +Goalw [equiv_def, refl_def, sym_def, trans_def] + "equiv {x::nat=>complex. True} hcomplexrel"; +by (auto_tac (claset() addSIs [hcomplexrel_refl] addSEs + [hcomplexrel_sym,hcomplexrel_trans] delrules [hcomplexrelI,hcomplexrelE], + simpset())); +qed "equiv_hcomplexrel"; + +val equiv_hcomplexrel_iff = + [TrueI, TrueI] MRS + ([CollectI, CollectI] MRS + (equiv_hcomplexrel RS eq_equiv_class_iff)); + +Goalw [hcomplex_def,hcomplexrel_def,quotient_def] "hcomplexrel``{x}:hcomplex"; +by (Blast_tac 1); +qed "hcomplexrel_in_hcomplex"; + +Goal "inj_on Abs_hcomplex hcomplex"; +by (rtac inj_on_inverseI 1); +by (etac Abs_hcomplex_inverse 1); +qed "inj_on_Abs_hcomplex"; + +Addsimps [equiv_hcomplexrel_iff,inj_on_Abs_hcomplex RS inj_on_iff, + hcomplexrel_iff, hcomplexrel_in_hcomplex, Abs_hcomplex_inverse]; + +Addsimps [equiv_hcomplexrel RS eq_equiv_class_iff]; +val eq_hcomplexrelD = equiv_hcomplexrel RSN (2,eq_equiv_class); + +Goal "inj(Rep_hcomplex)"; +by (rtac inj_inverseI 1); +by (rtac Rep_hcomplex_inverse 1); +qed "inj_Rep_hcomplex"; + +Goalw [hcomplexrel_def] "x: hcomplexrel `` {x}"; +by (Step_tac 1); +by (Auto_tac); +qed "lemma_hcomplexrel_refl"; +Addsimps [lemma_hcomplexrel_refl]; + +Goalw [hcomplex_def] "{} ~: hcomplex"; +by (auto_tac (claset() addSEs [quotientE],simpset())); +qed "hcomplex_empty_not_mem"; +Addsimps [hcomplex_empty_not_mem]; + +Goal "Rep_hcomplex x ~= {}"; +by (cut_inst_tac [("x","x")] Rep_hcomplex 1); +by (Auto_tac); +qed "Rep_hcomplex_nonempty"; +Addsimps [Rep_hcomplex_nonempty]; + +val [prem] = goal thy + "(!!x. z = Abs_hcomplex(hcomplexrel `` {x}) ==> P) ==> P"; +by (res_inst_tac [("x1","z")] + (rewrite_rule [hcomplex_def] Rep_hcomplex RS quotientE) 1); +by (dres_inst_tac [("f","Abs_hcomplex")] arg_cong 1); +by (res_inst_tac [("x","x")] prem 1); +by (asm_full_simp_tac (simpset() addsimps [Rep_hcomplex_inverse]) 1); +qed "eq_Abs_hcomplex"; + +(*-----------------------------------------------------------------------*) +(* Properties of nonstandard real and imaginary parts *) +(*-----------------------------------------------------------------------*) + +Goalw [hRe_def] + "hRe(Abs_hcomplex (hcomplexrel `` {X})) = \ +\ Abs_hypreal(hyprel `` {%n. Re(X n)})"; +by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1); +by (Auto_tac THEN Ultra_tac 1); +qed "hRe"; + +Goalw [hIm_def] + "hIm(Abs_hcomplex (hcomplexrel `` {X})) = \ +\ Abs_hypreal(hyprel `` {%n. Im(X n)})"; +by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1); +by (Auto_tac THEN Ultra_tac 1); +qed "hIm"; + +Goal "(w=z) = (hRe(w) = hRe(z) & hIm(w) = hIm(z))"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hRe,hIm, + complex_Re_Im_cancel_iff])); +by (ALLGOALS(Ultra_tac)); +qed "hcomplex_hRe_hIm_cancel_iff"; + +Goalw [hcomplex_zero_def] "hRe 0 = 0"; +by (simp_tac (simpset() addsimps [hRe,hypreal_zero_num]) 1); +qed "hcomplex_hRe_zero"; +Addsimps [hcomplex_hRe_zero]; + +Goalw [hcomplex_zero_def] "hIm 0 = 0"; +by (simp_tac (simpset() addsimps [hIm,hypreal_zero_num]) 1); +qed "hcomplex_hIm_zero"; +Addsimps [hcomplex_hIm_zero]; + +Goalw [hcomplex_one_def] "hRe 1 = 1"; +by (simp_tac (simpset() addsimps [hRe,hypreal_one_num]) 1); +qed "hcomplex_hRe_one"; +Addsimps [hcomplex_hRe_one]; + +Goalw [hcomplex_one_def] "hIm 1 = 0"; +by (simp_tac (simpset() addsimps [hIm,hypreal_one_def,hypreal_zero_num]) 1); +qed "hcomplex_hIm_one"; +Addsimps [hcomplex_hIm_one]; + +(*-----------------------------------------------------------------------*) +(* hcomplex_of_complex: the injection from complex to hcomplex *) +(* ----------------------------------------------------------------------*) + +Goal "inj(hcomplex_of_complex)"; +by (rtac injI 1 THEN rtac ccontr 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_complex_def])); +qed "inj_hcomplex_of_complex"; + +Goalw [iii_def,hcomplex_of_complex_def] "iii = hcomplex_of_complex ii"; +by (Simp_tac 1); +qed "hcomplex_of_complex_i"; + +(*-----------------------------------------------------------------------*) +(* Addition for nonstandard complex numbers: hcomplex_add *) +(* ----------------------------------------------------------------------*) + +Goalw [congruent2_def] + "congruent2 hcomplexrel (%X Y. hcomplexrel `` {%n. X n + Y n})"; +by (safe_tac (claset())); +by (ALLGOALS(Ultra_tac)); +qed "hcomplex_add_congruent2"; + +Goalw [hcomplex_add_def] + "Abs_hcomplex(hcomplexrel``{%n. X n}) + Abs_hcomplex(hcomplexrel``{%n. Y n}) = \ +\ Abs_hcomplex(hcomplexrel``{%n. X n + Y n})"; +by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1); +by Auto_tac; +by (Ultra_tac 1); +qed "hcomplex_add"; + +Goal "(z::hcomplex) + w = w + z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1); +by (asm_simp_tac (simpset() addsimps (complex_add_ac @ [hcomplex_add])) 1); +qed "hcomplex_add_commute"; + +Goal "((z1::hcomplex) + z2) + z3 = z1 + (z2 + z3)"; +by (res_inst_tac [("z","z1")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","z2")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","z3")] eq_Abs_hcomplex 1); +by (asm_simp_tac (simpset() addsimps [hcomplex_add,complex_add_assoc]) 1); +qed "hcomplex_add_assoc"; + +(*For AC rewriting*) +Goal "(x::hcomplex)+(y+z)=y+(x+z)"; +by (rtac (hcomplex_add_commute RS trans) 1); +by (rtac (hcomplex_add_assoc RS trans) 1); +by (rtac (hcomplex_add_commute RS arg_cong) 1); +qed "hcomplex_add_left_commute"; + +(* hcomplex addition is an AC operator *) +val hcomplex_add_ac = [hcomplex_add_assoc,hcomplex_add_commute, + hcomplex_add_left_commute]; + +Goalw [hcomplex_zero_def] "(0::hcomplex) + z = z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (asm_full_simp_tac (simpset() addsimps + [hcomplex_add]) 1); +qed "hcomplex_add_zero_left"; + +Goal "z + (0::hcomplex) = z"; +by (simp_tac (simpset() addsimps + [hcomplex_add_zero_left,hcomplex_add_commute]) 1); +qed "hcomplex_add_zero_right"; +Addsimps [hcomplex_add_zero_left,hcomplex_add_zero_right]; + +Goal "hRe(x + y) = hRe(x) + hRe(y)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hRe,hcomplex_add, + hypreal_add,complex_Re_add])); +qed "hRe_add"; + +Goal "hIm(x + y) = hIm(x) + hIm(y)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hIm,hcomplex_add, + hypreal_add,complex_Im_add])); +qed "hIm_add"; + +(*-----------------------------------------------------------------------*) +(* hypreal_minus: additive inverse on nonstandard complex *) +(* ----------------------------------------------------------------------*) + +Goalw [congruent_def] + "congruent hcomplexrel (%X. hcomplexrel `` {%n. - (X n)})"; +by (safe_tac (claset())); +by (ALLGOALS(Ultra_tac)); +qed "hcomplex_minus_congruent"; + +Goalw [hcomplex_minus_def] + "- (Abs_hcomplex(hcomplexrel `` {%n. X n})) = \ +\ Abs_hcomplex(hcomplexrel `` {%n. -(X n)})"; +by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1); +by (Auto_tac THEN Ultra_tac 1); +qed "hcomplex_minus"; + +Goal "- (- z) = (z::hcomplex)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (asm_simp_tac (simpset() addsimps [hcomplex_minus]) 1); +qed "hcomplex_minus_minus"; +Addsimps [hcomplex_minus_minus]; + +Goal "inj(%z::hcomplex. -z)"; +by (rtac injI 1); +by (dres_inst_tac [("f","uminus")] arg_cong 1); +by (Asm_full_simp_tac 1); +qed "inj_hcomplex_minus"; + +Goalw [hcomplex_zero_def] "- 0 = (0::hcomplex)"; +by (simp_tac (simpset() addsimps [hcomplex_minus]) 1); +qed "hcomplex_minus_zero"; +Addsimps [hcomplex_minus_zero]; + +Goal "(-x = 0) = (x = (0::hcomplex))"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_zero_def, + hcomplex_minus] @ complex_add_ac)); +qed "hcomplex_minus_zero_iff"; +Addsimps [hcomplex_minus_zero_iff]; + +Goal "(0 = -x) = (x = (0::hcomplex))"; +by (auto_tac (claset() addDs [sym],simpset())); +qed "hcomplex_minus_zero_iff2"; +Addsimps [hcomplex_minus_zero_iff2]; + +Goal "(-x ~= 0) = (x ~= (0::hcomplex))"; +by Auto_tac; +qed "hcomplex_minus_not_zero_iff"; + +Goal "z + - z = (0::hcomplex)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_add,hcomplex_minus, + hcomplex_zero_def])); +qed "hcomplex_add_minus_right"; +Addsimps [hcomplex_add_minus_right]; + +Goal "-z + z = (0::hcomplex)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_add,hcomplex_minus, + hcomplex_zero_def])); +qed "hcomplex_add_minus_left"; +Addsimps [hcomplex_add_minus_left]; + +Goal "z + (- z + w) = (w::hcomplex)"; +by (simp_tac (simpset() addsimps [hcomplex_add_assoc RS sym]) 1); +qed "hcomplex_add_minus_cancel"; + +Goal "(-z) + (z + w) = (w::hcomplex)"; +by (simp_tac (simpset() addsimps [hcomplex_add_assoc RS sym]) 1); +qed "hcomplex_minus_add_cancel"; + +Goal "hRe(-z) = - hRe(z)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hRe,hcomplex_minus, + hypreal_minus,complex_Re_minus])); +qed "hRe_minus"; + +Goal "hIm(-z) = - hIm(z)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hIm,hcomplex_minus, + hypreal_minus,complex_Im_minus])); +qed "hIm_minus"; + +Goalw [hcomplex_zero_def] + "x + y = (0::hcomplex) ==> x = -y"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_add,hcomplex_minus])); +by (ultra_tac (claset() addIs [complex_add_minus_eq_minus],simpset()) 1); +qed "hcomplex_add_minus_eq_minus"; + +Goal "-(x + y) = -x + -(y::hcomplex)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_add,hcomplex_minus])); +qed "hcomplex_minus_add_distrib"; +Addsimps [hcomplex_minus_add_distrib]; + +Goal "((x::hcomplex) + y = x + z) = (y = z)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_add])); +qed "hcomplex_add_left_cancel"; +AddIffs [hcomplex_add_left_cancel]; + +Goal "(y + (x::hcomplex)= z + x) = (y = z)"; +by (simp_tac (simpset() addsimps [hcomplex_add_commute]) 1); +qed "hcomplex_add_right_cancel"; +AddIffs [hcomplex_add_right_cancel]; + +Goal "((x::hcomplex) = y) = ((0::hcomplex) = x + - y)"; +by (Step_tac 1); +by (res_inst_tac [("x1","-y")] + (hcomplex_add_right_cancel RS iffD1) 2); +by (Auto_tac); +qed "hcomplex_eq_minus_iff"; + +Goal "((x::hcomplex) = y) = (x + - y = (0::hcomplex))"; +by (Step_tac 1); +by (res_inst_tac [("x1","-y")] + (hcomplex_add_right_cancel RS iffD1) 2); +by (Auto_tac); +qed "hcomplex_eq_minus_iff2"; + +(*-----------------------------------------------------------------------*) +(* Subraction for nonstandard complex numbers: hcomplex_diff *) +(* ----------------------------------------------------------------------*) + +Goalw [hcomplex_diff_def] + "Abs_hcomplex(hcomplexrel``{%n. X n}) - Abs_hcomplex(hcomplexrel``{%n. Y n}) = \ +\ Abs_hcomplex(hcomplexrel``{%n. X n - Y n})"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_minus,hcomplex_add, + complex_diff_def])); +qed "hcomplex_diff"; + +Goalw [hcomplex_diff_def] "(z::hcomplex) - z = (0::hcomplex)"; +by (Simp_tac 1); +qed "hcomplex_diff_zero"; +Addsimps [hcomplex_diff_zero]; + +Goal "(0::hcomplex) - x = -x"; +by (simp_tac (simpset() addsimps [hcomplex_diff_def]) 1); +qed "hcomplex_diff_0"; + +Goal "x - (0::hcomplex) = x"; +by (simp_tac (simpset() addsimps [hcomplex_diff_def]) 1); +qed "hcomplex_diff_0_right"; + +Goal "x - x = (0::hcomplex)"; +by (simp_tac (simpset() addsimps [hcomplex_diff_def]) 1); +qed "hcomplex_diff_self"; + +Addsimps [hcomplex_diff_0, hcomplex_diff_0_right, hcomplex_diff_self]; + +Goal "((x::hcomplex) - y = z) = (x = z + y)"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_def,hcomplex_add_assoc])); +qed "hcomplex_diff_eq_eq"; + +(*-----------------------------------------------------------------------*) +(* Multiplication for nonstandard complex numbers: hcomplex_mult *) +(* ----------------------------------------------------------------------*) + +Goalw [hcomplex_mult_def] + "Abs_hcomplex(hcomplexrel``{%n. X n}) * Abs_hcomplex(hcomplexrel``{%n. Y n}) = \ +\ Abs_hcomplex(hcomplexrel``{%n. X n * Y n})"; +by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1); +by (Auto_tac THEN Ultra_tac 1); +qed "hcomplex_mult"; + +Goal "(w::hcomplex) * z = z * w"; +by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult, + complex_mult_commute])); +qed "hcomplex_mult_commute"; + +Goal "((u::hcomplex) * v) * w = u * (v * w)"; +by (res_inst_tac [("z","u")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","v")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult, + complex_mult_assoc])); +qed "hcomplex_mult_assoc"; + +Goal "(x::hcomplex) * (y * z) = y * (x * z)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult, + complex_mult_left_commute])); +qed "hcomplex_mult_left_commute"; + +val hcomplex_mult_ac = [hcomplex_mult_assoc,hcomplex_mult_commute, + hcomplex_mult_left_commute]; + +Goalw [hcomplex_one_def] "(1::hcomplex) * z = z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult])); +qed "hcomplex_mult_one_left"; +Addsimps [hcomplex_mult_one_left]; + +Goal "z * (1::hcomplex) = z"; +by (simp_tac (simpset() addsimps [hcomplex_mult_commute]) 1); +qed "hcomplex_mult_one_right"; +Addsimps [hcomplex_mult_one_right]; + +Goalw [hcomplex_zero_def] "(0::hcomplex) * z = 0"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult])); +qed "hcomplex_mult_zero_left"; +Addsimps [hcomplex_mult_zero_left]; + +Goal "z * (0::hcomplex) = 0"; +by (simp_tac (simpset() addsimps [hcomplex_mult_commute]) 1); +qed "hcomplex_mult_zero_right"; +Addsimps [hcomplex_mult_zero_right]; + +Goal "-(x * y) = -x * (y::hcomplex)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult, + hcomplex_minus])); +qed "hcomplex_minus_mult_eq1"; + +Goal "-(x * y) = x * -(y::hcomplex)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult, + hcomplex_minus])); +qed "hcomplex_minus_mult_eq2"; + +Addsimps [hcomplex_minus_mult_eq1 RS sym,hcomplex_minus_mult_eq2 RS sym]; + +Goal "- 1 * (z::hcomplex) = -z"; +by (Simp_tac 1); +qed "hcomplex_mult_minus_one"; +Addsimps [hcomplex_mult_minus_one]; + +Goal "(z::hcomplex) * - 1 = -z"; +by (stac hcomplex_mult_commute 1); +by (Simp_tac 1); +qed "hcomplex_mult_minus_one_right"; +Addsimps [hcomplex_mult_minus_one_right]; + +Goal "-x * -y = x * (y::hcomplex)"; +by Auto_tac; +qed "hcomplex_minus_mult_cancel"; +Addsimps [hcomplex_minus_mult_cancel]; + +Goal "-x * y = x * -(y::hcomplex)"; +by Auto_tac; +qed "hcomplex_minus_mult_commute"; + +qed_goal "hcomplex_add_assoc_cong" thy + "!!z. (z::hcomplex) + v = z' + v' ==> z + (v + w) = z' + (v' + w)" + (fn _ => [(asm_simp_tac (simpset() addsimps [hcomplex_add_assoc RS sym]) 1)]); + +qed_goal "hcomplex_add_assoc_swap" thy "(z::hcomplex) + (v + w) = v + (z + w)" + (fn _ => [(REPEAT (ares_tac [hcomplex_add_commute RS hcomplex_add_assoc_cong] 1))]); + +Goal "((z1::hcomplex) + z2) * w = (z1 * w) + (z2 * w)"; +by (res_inst_tac [("z","z1")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","z2")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult,hcomplex_add, + complex_add_mult_distrib])); +qed "hcomplex_add_mult_distrib"; + +Goal "(w::hcomplex) * (z1 + z2) = (w * z1) + (w * z2)"; +by (res_inst_tac [("z1","z1 + z2")] (hcomplex_mult_commute RS ssubst) 1); +by (simp_tac (simpset() addsimps [hcomplex_add_mult_distrib]) 1); +by (simp_tac (simpset() addsimps [hcomplex_mult_commute]) 1); +qed "hcomplex_add_mult_distrib2"; + +Goalw [hcomplex_zero_def,hcomplex_one_def] "(0::hcomplex) ~= (1::hcomplex)"; +by Auto_tac; +qed "hcomplex_zero_not_eq_one"; +Addsimps [hcomplex_zero_not_eq_one]; +Addsimps [hcomplex_zero_not_eq_one RS not_sym]; + +(*-----------------------------------------------------------------------*) +(* Inverse of nonstandard complex number *) +(*-----------------------------------------------------------------------*) + +Goalw [hcinv_def] + "inverse (Abs_hcomplex(hcomplexrel `` {%n. X n})) = \ +\ Abs_hcomplex(hcomplexrel `` {%n. inverse (X n)})"; +by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1); +by (Auto_tac THEN Ultra_tac 1); +qed "hcomplex_inverse"; + +Goalw [hcomplex_zero_def] "inverse (0::hcomplex) = 0"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse])); +qed "HCOMPLEX_INVERSE_ZERO"; + +Goal "a / (0::hcomplex) = 0"; +by (simp_tac (simpset() addsimps [hcomplex_divide_def, HCOMPLEX_INVERSE_ZERO]) 1); +qed "HCOMPLEX_DIVISION_BY_ZERO"; (*NOT for adding to default simpset*) + +fun hcomplex_div_undefined_case_tac s i = + case_tac s i THEN + asm_simp_tac (simpset() addsimps [HCOMPLEX_DIVISION_BY_ZERO, HCOMPLEX_INVERSE_ZERO]) i; + +Goalw [hcomplex_zero_def,hcomplex_one_def] + "z ~= (0::hcomplex) ==> inverse(z) * z = (1::hcomplex)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse,hcomplex_mult])); +by (Ultra_tac 1); +by (rtac ccontr 1 THEN dtac (complex_mult_inv_left) 1); +by Auto_tac; +qed "hcomplex_mult_inv_left"; +Addsimps [hcomplex_mult_inv_left]; + +Goal "z ~= (0::hcomplex) ==> z * inverse(z) = (1::hcomplex)"; +by (auto_tac (claset() addIs [hcomplex_mult_commute RS subst],simpset())); +qed "hcomplex_mult_inv_right"; +Addsimps [hcomplex_mult_inv_right]; + +Goal "(c::hcomplex) ~= (0::hcomplex) ==> (c*a=c*b) = (a=b)"; +by Auto_tac; +by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1); +by (asm_full_simp_tac (simpset() addsimps hcomplex_mult_ac) 1); +qed "hcomplex_mult_left_cancel"; + +Goal "(c::hcomplex) ~= (0::hcomplex) ==> (a*c=b*c) = (a=b)"; +by (Step_tac 1); +by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1); +by (asm_full_simp_tac (simpset() addsimps hcomplex_mult_ac) 1); +qed "hcomplex_mult_right_cancel"; + +Goal "z ~= (0::hcomplex) ==> inverse(z) ~= 0"; +by (Step_tac 1); +by (ftac (hcomplex_mult_right_cancel RS iffD2) 1); +by (thin_tac "inverse z = 0" 2); +by (assume_tac 1 THEN Auto_tac); +qed "hcomplex_inverse_not_zero"; +Addsimps [hcomplex_inverse_not_zero]; + +Goal "[| x ~= (0::hcomplex); y ~= 0 |] ==> x * y ~= 0"; +by (Step_tac 1); +by (dres_inst_tac [("f","%z. inverse x*z")] arg_cong 1); +by (asm_full_simp_tac (simpset() addsimps [hcomplex_mult_assoc RS sym]) 1); +qed "hcomplex_mult_not_zero"; + +bind_thm ("hcomplex_mult_not_zeroE",hcomplex_mult_not_zero RS notE); + +Goal "inverse(inverse x) = (x::hcomplex)"; +by (hcomplex_div_undefined_case_tac "x = 0" 1); +by (res_inst_tac [("c1","inverse x")] (hcomplex_mult_right_cancel RS iffD1) 1); +by (etac hcomplex_inverse_not_zero 1); +by (auto_tac (claset() addDs [hcomplex_inverse_not_zero],simpset())); +qed "hcomplex_inverse_inverse"; +Addsimps [hcomplex_inverse_inverse]; + +Goalw [hcomplex_one_def] "inverse((1::hcomplex)) = 1"; +by (simp_tac (simpset() addsimps [hcomplex_inverse]) 1); +qed "hcomplex_inverse_one"; +Addsimps [hcomplex_inverse_one]; + +Goal "inverse(-x) = -inverse(x::hcomplex)"; +by (hcomplex_div_undefined_case_tac "x = 0" 1); +by (res_inst_tac [("c1","-x")] (hcomplex_mult_right_cancel RS iffD1) 1); +by (stac hcomplex_mult_inv_left 2); +by Auto_tac; +qed "hcomplex_minus_inverse"; + +Goal "inverse(x*y) = inverse x * inverse (y::hcomplex)"; +by (hcomplex_div_undefined_case_tac "x = 0" 1); +by (hcomplex_div_undefined_case_tac "y = 0" 1); +by (res_inst_tac [("c1","x*y")] (hcomplex_mult_left_cancel RS iffD1) 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_not_zero] + @ hcomplex_mult_ac)); +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_not_zero, + hcomplex_mult_assoc RS sym])); +qed "hcomplex_inverse_distrib"; + +(*** division ***) + +(* adding some of these theorems to simpset as for reals: not 100% convinced for some*) + +Goal "(x::hcomplex) * (y/z) = (x*y)/z"; +by (simp_tac (simpset() addsimps [hcomplex_divide_def, hcomplex_mult_assoc]) 1); +qed "hcomplex_times_divide1_eq"; + +Goal "(y/z) * (x::hcomplex) = (y*x)/z"; +by (simp_tac (simpset() addsimps [hcomplex_divide_def] @ hcomplex_mult_ac) 1); +qed "hcomplex_times_divide2_eq"; + +Addsimps [hcomplex_times_divide1_eq, hcomplex_times_divide2_eq]; + +Goal "(x::hcomplex) / (y/z) = (x*z)/y"; +by (simp_tac (simpset() addsimps [hcomplex_divide_def, hcomplex_inverse_distrib]@ + hcomplex_mult_ac) 1); +qed "hcomplex_divide_divide1_eq"; + +Goal "((x::hcomplex) / y) / z = x/(y*z)"; +by (simp_tac (simpset() addsimps [hcomplex_divide_def, hcomplex_inverse_distrib, + hcomplex_mult_assoc]) 1); +qed "hcomplex_divide_divide2_eq"; + +Addsimps [hcomplex_divide_divide1_eq, hcomplex_divide_divide2_eq]; + +(** As with multiplication, pull minus signs OUT of the / operator **) + +Goal "(-x) / (y::hcomplex) = - (x/y)"; +by (simp_tac (simpset() addsimps [hcomplex_divide_def]) 1); +qed "hcomplex_minus_divide_eq"; +Addsimps [hcomplex_minus_divide_eq]; + +Goal "(x / -(y::hcomplex)) = - (x/y)"; +by (simp_tac (simpset() addsimps [hcomplex_divide_def, hcomplex_minus_inverse]) 1); +qed "hcomplex_divide_minus_eq"; +Addsimps [hcomplex_divide_minus_eq]; + +Goal "(x+y)/(z::hcomplex) = x/z + y/z"; +by (simp_tac (simpset() addsimps [hcomplex_divide_def, hcomplex_add_mult_distrib]) 1); +qed "hcomplex_add_divide_distrib"; + +(*---------------------------------------------------------------------------*) +(* Embedding properties for hcomplex_of_hypreal map *) +(*---------------------------------------------------------------------------*) + +Goalw [hcomplex_of_hypreal_def] + "hcomplex_of_hypreal (Abs_hypreal(hyprel `` {%n. X n})) = \ +\ Abs_hcomplex(hcomplexrel `` {%n. complex_of_real (X n)})"; +by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1); +by (Auto_tac THEN Ultra_tac 1); +qed "hcomplex_of_hypreal"; + +Goal "inj hcomplex_of_hypreal"; +by (rtac injI 1); +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal])); +qed "inj_hcomplex_of_hypreal"; + +Goal "(hcomplex_of_hypreal x = hcomplex_of_hypreal y) = (x = y)"; +by (auto_tac (claset() addDs [inj_hcomplex_of_hypreal RS injD],simpset())); +qed "hcomplex_of_hypreal_cancel_iff"; +AddIffs [hcomplex_of_hypreal_cancel_iff]; + +Goal "hcomplex_of_hypreal(-x) = - hcomplex_of_hypreal x"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + hcomplex_minus,hypreal_minus,complex_of_real_minus])); +qed "hcomplex_of_hypreal_minus"; + +Goal "hcomplex_of_hypreal(inverse x) = inverse(hcomplex_of_hypreal x)"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + hypreal_inverse,hcomplex_inverse,complex_of_real_inverse])); +qed "hcomplex_of_hypreal_inverse"; + +Goal "hcomplex_of_hypreal x + hcomplex_of_hypreal y = \ +\ hcomplex_of_hypreal (x + y)"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + hypreal_add,hcomplex_add,complex_of_real_add])); +qed "hcomplex_of_hypreal_add"; + +Goalw [hcomplex_diff_def] + "hcomplex_of_hypreal x - hcomplex_of_hypreal y = \ +\ hcomplex_of_hypreal (x - y)"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal_minus + RS sym,hcomplex_of_hypreal_add,hypreal_diff_def])); +qed "hcomplex_of_hypreal_diff"; + +Goal "hcomplex_of_hypreal x * hcomplex_of_hypreal y = \ +\ hcomplex_of_hypreal (x * y)"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + hypreal_mult,hcomplex_mult,complex_of_real_mult])); +qed "hcomplex_of_hypreal_mult"; + +Goalw [hcomplex_divide_def] + "hcomplex_of_hypreal x / hcomplex_of_hypreal y = hcomplex_of_hypreal(x/y)"; +by (hypreal_div_undefined_case_tac "y=0" 1); +by (simp_tac (simpset() addsimps [rename_numerals HYPREAL_DIVISION_BY_ZERO, + HCOMPLEX_INVERSE_ZERO]) 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal_mult, + hcomplex_of_hypreal_inverse RS sym])); +by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); +qed "hcomplex_of_hypreal_divide"; + +Goalw [hcomplex_one_def] + "hcomplex_of_hypreal 1 = 1"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,hypreal_one_num])); +qed "hcomplex_of_hypreal_one"; + +Goalw [hcomplex_zero_def,hypreal_zero_def] + "hcomplex_of_hypreal 0 = 0"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal])); +qed "hcomplex_of_hypreal_zero"; + +Addsimps [hcomplex_of_hypreal_one,hcomplex_of_hypreal_zero, + rename_numerals hcomplex_of_hypreal_zero]; + +Goal "hcomplex_of_hypreal (x ^ n) = (hcomplex_of_hypreal x) ^ n"; +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal_mult RS sym])); +qed "hcomplex_of_hypreal_pow"; + +Goal "hRe(hcomplex_of_hypreal z) = z"; +by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,hRe])); +qed "hRe_hcomplex_of_hypreal"; +Addsimps [hRe_hcomplex_of_hypreal]; + +Goal "hIm(hcomplex_of_hypreal z) = 0"; +by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,hIm, + hypreal_zero_num])); +qed "hIm_hcomplex_of_hypreal"; +Addsimps [hIm_hcomplex_of_hypreal]; + +Goal "hcomplex_of_hypreal epsilon ~= 0"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + epsilon_def,hcomplex_zero_def])); +qed "hcomplex_of_hypreal_epsilon_not_zero"; +Addsimps [hcomplex_of_hypreal_epsilon_not_zero]; + +(*---------------------------------------------------------------------------*) +(* Modulus (absolute value) of nonstandard complex number *) +(*---------------------------------------------------------------------------*) + +Goalw [hcmod_def] + "hcmod (Abs_hcomplex(hcomplexrel `` {%n. X n})) = \ +\ Abs_hypreal(hyprel `` {%n. cmod (X n)})"; +by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1); +by (Auto_tac THEN Ultra_tac 1); +qed "hcmod"; + +Goalw [hcomplex_zero_def,hypreal_zero_def] + "hcmod(0) = 0"; +by (auto_tac (claset(),simpset() addsimps [hcmod])); +qed "hcmod_zero"; +Addsimps [hcmod_zero,rename_numerals hcmod_zero]; + +Goalw [hcomplex_one_def] + "hcmod(1) = 1"; +by (auto_tac (claset(),simpset() addsimps [hcmod,hypreal_one_num])); +qed "hcmod_one"; +Addsimps [hcmod_one]; + +Goal "hcmod(hcomplex_of_hypreal x) = abs x"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_of_hypreal, + hypreal_hrabs])); +qed "hcmod_hcomplex_of_hypreal"; +Addsimps [hcmod_hcomplex_of_hypreal]; + +Goal "hcomplex_of_hypreal (abs x) = \ +\ hcomplex_of_hypreal(hcmod(hcomplex_of_hypreal x))"; +by (Simp_tac 1); +qed "hcomplex_of_hypreal_abs"; + +(*---------------------------------------------------------------------------*) +(* conjugation *) +(*---------------------------------------------------------------------------*) + +Goalw [hcnj_def] + "hcnj (Abs_hcomplex(hcomplexrel `` {%n. X n})) = \ +\ Abs_hcomplex(hcomplexrel `` {%n. cnj(X n)})"; +by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1); +by (Auto_tac THEN Ultra_tac 1); +qed "hcnj"; + +Goal "inj hcnj"; +by (rtac injI 1); +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcnj])); +qed "inj_hcnj"; + +Goal "(hcnj x = hcnj y) = (x = y)"; +by (auto_tac (claset() addDs [inj_hcnj RS injD],simpset())); +qed "hcomplex_hcnj_cancel_iff"; +Addsimps [hcomplex_hcnj_cancel_iff]; + +Goal "hcnj (hcnj z) = z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcnj])); +qed "hcomplex_hcnj_hcnj"; +Addsimps [hcomplex_hcnj_hcnj]; + +Goal "hcnj (hcomplex_of_hypreal x) = hcomplex_of_hypreal x"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_of_hypreal])); +qed "hcomplex_hcnj_hcomplex_of_hypreal"; +Addsimps [hcomplex_hcnj_hcomplex_of_hypreal]; + +Goal "hcmod (hcnj z) = hcmod z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcnj,hcmod])); +qed "hcomplex_hmod_hcnj"; +Addsimps [hcomplex_hmod_hcnj]; + +Goal "hcnj (-z) = - hcnj z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_minus, + complex_cnj_minus])); +qed "hcomplex_hcnj_minus"; + +Goal "hcnj(inverse z) = inverse(hcnj z)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_inverse, + complex_cnj_inverse])); +qed "hcomplex_hcnj_inverse"; + +Goal "hcnj(w + z) = hcnj(w) + hcnj(z)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_add, + complex_cnj_add])); +qed "hcomplex_hcnj_add"; + +Goal "hcnj(w - z) = hcnj(w) - hcnj(z)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_diff, + complex_cnj_diff])); +qed "hcomplex_hcnj_diff"; + +Goal "hcnj(w * z) = hcnj(w) * hcnj(z)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_mult, + complex_cnj_mult])); +qed "hcomplex_hcnj_mult"; + +Goalw [hcomplex_divide_def] "hcnj(w / z) = (hcnj w)/(hcnj z)"; +by (simp_tac (simpset() addsimps [hcomplex_hcnj_mult,hcomplex_hcnj_inverse]) 1); +qed "hcomplex_hcnj_divide"; + +Goalw [hcomplex_one_def] "hcnj 1 = 1"; +by (simp_tac (simpset() addsimps [hcnj]) 1); +qed "hcnj_one"; +Addsimps [hcnj_one]; + +Goal "hcnj(z ^ n) = hcnj(z) ^ n"; +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_hcnj_mult])); +qed "hcomplex_hcnj_pow"; + +(* MOVE to NSComplexBin +Goal "z + hcnj z = \ +\ hcomplex_of_hypreal (2 * hRe(z))"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),HOL_ss addsimps [hRe,hcnj,hcomplex_add, + hypreal_mult,hcomplex_of_hypreal,complex_add_cnj])); +qed "hcomplex_add_hcnj"; + +Goal "z - hcnj z = \ +\ hcomplex_of_hypreal (hypreal_of_real 2 * hIm(z)) * iii"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hIm,hcnj,hcomplex_diff, + hypreal_of_real_def,hypreal_mult,hcomplex_of_hypreal, + complex_diff_cnj,iii_def,hcomplex_mult])); +qed "hcomplex_diff_hcnj"; +*) + +Goalw [hcomplex_zero_def] + "hcnj 0 = 0"; +by (auto_tac (claset(),simpset() addsimps [hcnj])); +qed "hcomplex_hcnj_zero"; +Addsimps [hcomplex_hcnj_zero]; + +Goal "(hcnj z = 0) = (z = 0)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_zero_def, + hcnj])); +qed "hcomplex_hcnj_zero_iff"; +AddIffs [hcomplex_hcnj_zero_iff]; + +Goal "z * hcnj z = hcomplex_of_hypreal (hRe(z) ^ 2 + hIm(z) ^ 2)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_mult, + hcomplex_of_hypreal,hRe,hIm,hypreal_add,hypreal_mult, + complex_mult_cnj,two_eq_Suc_Suc])); +qed "hcomplex_mult_hcnj"; + + +(*---------------------------------------------------------------------------*) +(* some algebra etc. *) +(*---------------------------------------------------------------------------*) + +Goal "(x*y = (0::hcomplex)) = (x = 0 | y = 0)"; +by Auto_tac; +by (auto_tac (claset() addIs [ccontr] addDs + [hcomplex_mult_not_zero],simpset())); +qed "hcomplex_mult_zero_iff"; +Addsimps [hcomplex_mult_zero_iff]; + +Goal "(x + y = x) = (y = (0::hcomplex))"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_add, + hcomplex_zero_def])); +qed "hcomplex_add_left_cancel_zero"; +Addsimps [hcomplex_add_left_cancel_zero]; + +Goalw [hcomplex_diff_def] + "((z1::hcomplex) - z2) * w = (z1 * w) - (z2 * w)"; +by (simp_tac (simpset() addsimps [hcomplex_add_mult_distrib]) 1); +qed "hcomplex_diff_mult_distrib"; + +Goalw [hcomplex_diff_def] + "(w::hcomplex) * (z1 - z2) = (w * z1) - (w * z2)"; +by (simp_tac (simpset() addsimps [hcomplex_add_mult_distrib2]) 1); +qed "hcomplex_diff_mult_distrib2"; + +(*---------------------------------------------------------------------------*) +(* More theorems about hcmod *) +(*---------------------------------------------------------------------------*) + +Goal "(hcmod x = 0) = (x = 0)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_zero_def, + hypreal_zero_num])); +qed "hcomplex_hcmod_eq_zero_cancel"; +Addsimps [hcomplex_hcmod_eq_zero_cancel]; + +(* not proved already? strange! *) +Goalw [hypreal_le_def] + "(hypreal_of_nat n <= hypreal_of_nat m) = (n <= m)"; +by Auto_tac; +qed "hypreal_of_nat_le_iff"; +Addsimps [hypreal_of_nat_le_iff]; + +Goal "0 <= hypreal_of_nat n"; +by (simp_tac (simpset() addsimps [hypreal_of_nat_zero RS sym] + delsimps [hypreal_of_nat_zero]) 1); +qed "hypreal_of_nat_ge_zero"; +Addsimps [hypreal_of_nat_ge_zero]; + +Addsimps [hypreal_of_nat_ge_zero RS hrabs_eqI1]; + +Goal "0 <= hypreal_of_hypnat n"; +by (res_inst_tac [("z","n")] eq_Abs_hypnat 1); +by (asm_simp_tac (simpset() addsimps [hypreal_of_hypnat, + hypreal_zero_num,hypreal_le]) 1); +qed "hypreal_of_hypnat_ge_zero"; +Addsimps [hypreal_of_hypnat_ge_zero]; + +Addsimps [hypreal_of_hypnat_ge_zero RS hrabs_eqI1]; + +Goal "hcmod (hcomplex_of_hypreal(hypreal_of_nat n)) = hypreal_of_nat n"; +by Auto_tac; +qed "hcmod_hcomplex_of_hypreal_of_nat"; +Addsimps [hcmod_hcomplex_of_hypreal_of_nat]; + +Goal "hcmod (hcomplex_of_hypreal(hypreal_of_hypnat n)) = hypreal_of_hypnat n"; +by Auto_tac; +qed "hcmod_hcomplex_of_hypreal_of_hypnat"; +Addsimps [hcmod_hcomplex_of_hypreal_of_hypnat]; + +Goal "hcmod (-x) = hcmod(x)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_minus])); +qed "hcmod_minus"; +Addsimps [hcmod_minus]; + +Goal "hcmod(z * hcnj(z)) = hcmod(z) ^ 2"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_mult, + hcnj,hypreal_mult,complex_mod_mult_cnj,two_eq_Suc_Suc])); +qed "hcmod_mult_hcnj"; + +Goal "(0::hypreal) <= hcmod x"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcmod, + hypreal_zero_num,hypreal_le])); +qed "hcmod_ge_zero"; +Addsimps [hcmod_ge_zero]; + +Goal "abs(hcmod x) = hcmod x"; +by (auto_tac (claset() addIs [hrabs_eqI1],simpset())); +qed "hrabs_hcmod_cancel"; +Addsimps [hrabs_hcmod_cancel]; + +Goal "hcmod(x*y) = hcmod(x) * hcmod(y)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_mult, + hypreal_mult,complex_mod_mult])); +qed "hcmod_mult"; + +Goal "hcmod(x + y) ^ 2 = hcmod(x) ^ 2 + hcmod(y) ^ 2 + 2 * hRe(x * hcnj y)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_add, + hypreal_mult,hRe,hcnj,hcomplex_mult,two_eq_Suc_Suc, + realpow_two RS sym] delsimps [realpow_Suc])); +by (auto_tac (claset(),simpset() addsimps [two_eq_Suc_Suc RS sym, + complex_mod_add_squared_eq,hypreal_add RS sym,hypreal_mult RS sym, + symmetric hypreal_of_real_def])); +qed "hcmod_add_squared_eq"; + +Goal "hRe(x * hcnj y) <= hcmod(x * hcnj y)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcmod,hcnj, + hcomplex_mult,hRe,hypreal_le])); +qed "hcomplex_hRe_mult_hcnj_le_hcmod"; +Addsimps [hcomplex_hRe_mult_hcnj_le_hcmod]; + +Goal "hRe(x * hcnj y) <= hcmod(x * y)"; +by (cut_inst_tac [("x","x"),("y","y")] hcomplex_hRe_mult_hcnj_le_hcmod 1); +by (asm_full_simp_tac (simpset() addsimps [hcmod_mult]) 1); +qed "hcomplex_hRe_mult_hcnj_le_hcmod2"; +Addsimps [hcomplex_hRe_mult_hcnj_le_hcmod2]; + +Goal "hcmod (x + y) ^ 2 <= (hcmod(x) + hcmod(y)) ^ 2"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcmod,hcnj, + hcomplex_add,hypreal_mult,hypreal_add,hypreal_le, + realpow_two RS sym,two_eq_Suc_Suc] delsimps [realpow_Suc])); +by (simp_tac (simpset() addsimps [two_eq_Suc_Suc RS sym]) 1); +qed "hcmod_triangle_squared"; +Addsimps [hcmod_triangle_squared]; + +Goal "hcmod (x + y) <= hcmod(x) + hcmod(y)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcmod, + hcomplex_add,hypreal_add,hypreal_le])); +qed "hcmod_triangle_ineq"; +Addsimps [hcmod_triangle_ineq]; + +Goal "hcmod(b + a) - hcmod b <= hcmod a"; +by (cut_inst_tac [("x1","b"),("y1","a"),("x","-hcmod b")] + (hcmod_triangle_ineq RS hypreal_add_le_mono1) 1); +by (asm_full_simp_tac (simpset() addsimps hypreal_add_ac) 1); +qed "hcmod_triangle_ineq2"; +Addsimps [hcmod_triangle_ineq2]; + +Goal "hcmod (x - y) = hcmod (y - x)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcmod, + hcomplex_diff,complex_mod_diff_commute])); +qed "hcmod_diff_commute"; + +Goal "[| hcmod x < r; hcmod y < s |] ==> hcmod (x + y) < r + s"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","r")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","s")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_add, + hypreal_add,hypreal_less])); +by (ultra_tac (claset() addIs [complex_mod_add_less],simpset()) 1); +qed "hcmod_add_less"; + +Goal "[| hcmod x < r; hcmod y < s |] ==> hcmod (x * y) < r * s"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","r")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","s")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcmod,hypreal_mult, + hypreal_less,hcomplex_mult])); +by (ultra_tac (claset() addIs [complex_mod_mult_less],simpset()) 1); +qed "hcmod_mult_less"; + +goal NSComplex.thy "hcmod(a) - hcmod(b) <= hcmod(a + b)"; +by (res_inst_tac [("z","a")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","b")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_add, + hypreal_diff,hypreal_le])); +qed "hcmod_diff_ineq"; +Addsimps [hcmod_diff_ineq]; + + +(*---------------------------------------------------------------------------*) +(* a few nonlinear theorems *) +(*---------------------------------------------------------------------------*) + +Goalw [hcpow_def] + "Abs_hcomplex(hcomplexrel``{%n. X n}) hcpow \ +\ Abs_hypnat(hypnatrel``{%n. Y n}) = \ +\ Abs_hcomplex(hcomplexrel``{%n. X n ^ Y n})"; +by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1); +by (Auto_tac THEN Ultra_tac 1); +qed "hcpow"; + +Goal "hcomplex_of_hypreal (x pow n) = (hcomplex_of_hypreal x) hcpow n"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","n")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + hyperpow,hcpow,complex_of_real_pow])); +qed "hcomplex_of_hypreal_hyperpow"; + +Goal "hcmod(x ^ n) = hcmod(x) ^ n"; +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps [hcmod_mult])); +qed "hcmod_hcomplexpow"; + +Goal "hcmod(x hcpow n) = hcmod(x) pow n"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","n")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hcpow,hyperpow, + hcmod,complex_mod_complexpow])); +qed "hcmod_hcpow"; + +Goal "(-x::hcomplex) ^ n = (if even n then (x ^ n) else -(x ^ n))"; +by (induct_tac "n" 1); +by Auto_tac; +qed "hcomplexpow_minus"; + +Goal "(-x::hcomplex) hcpow n = \ +\ (if ( *pNat* even) n then (x hcpow n) else -(x hcpow n))"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","n")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hcpow,hyperpow,starPNat, + hcomplex_minus])); +by (ALLGOALS(ultra_tac (claset(),simpset() addsimps [complexpow_minus]))); +qed "hcpow_minus"; + +Goal "inverse(-x) = - inverse (x::hcomplex)"; +by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse,hcomplex_minus, + complex_inverse_minus])); +qed "hccomplex_inverse_minus"; + +Goalw [hcomplex_divide_def] "x / (1::hcomplex) = x"; +by (Simp_tac 1); +qed "hcomplex_div_one"; +Addsimps [hcomplex_div_one]; + +Goal "hcmod(inverse x) = inverse(hcmod x)"; +by (hcomplex_div_undefined_case_tac "x = 0" 1); +by (res_inst_tac [("c1","hcmod x")] (hypreal_mult_left_cancel RS iffD1) 1); +by (auto_tac (claset(),simpset() addsimps [hcmod_mult RS sym])); +qed "hcmod_hcomplex_inverse"; + +Goalw [hcomplex_divide_def,hypreal_divide_def] + "hcmod(x/y) = hcmod(x)/(hcmod y)"; +by (auto_tac (claset(),simpset() addsimps [hcmod_mult, + hcmod_hcomplex_inverse])); +qed "hcmod_divide"; + +Goalw [hcomplex_divide_def] + "inverse(x/y) = y/(x::hcomplex)"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse_distrib, + hcomplex_mult_commute])); +qed "hcomplex_inverse_divide"; +Addsimps [hcomplex_inverse_divide]; + +Goal "((r::hcomplex) * s) ^ n = (r ^ n) * (s ^ n)"; +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps hcomplex_mult_ac)); +qed "hcomplexpow_mult"; + +Goal "((r::hcomplex) * s) hcpow n = (r hcpow n) * (s hcpow n)"; +by (res_inst_tac [("z","r")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","s")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","n")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hcpow,hypreal_mult, + hcomplex_mult,complexpow_mult])); +qed "hcpow_mult"; + +Goal "(0::hcomplex) ^ (Suc n) = 0"; +by (Auto_tac); +qed "hcomplexpow_zero"; +Addsimps [hcomplexpow_zero]; + +Goalw [hcomplex_zero_def,hypnat_one_def] + "0 hcpow (n + 1) = 0"; +by (res_inst_tac [("z","n")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hcpow,hypnat_add])); +qed "hcpow_zero"; +Addsimps [hcpow_zero]; + +Goalw [hSuc_def] + "0 hcpow (hSuc n) = 0"; +by (Simp_tac 1); +qed "hcpow_zero2"; +Addsimps [hcpow_zero2]; + +Goal "r ~= (0::hcomplex) --> r ^ n ~= 0"; +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps + [hcomplex_mult_not_zero])); +qed_spec_mp "hcomplexpow_not_zero"; +Addsimps [hcomplexpow_not_zero]; +AddIs [hcomplexpow_not_zero]; + +Goal "r ~= 0 ==> r hcpow n ~= (0::hcomplex)"; +by (res_inst_tac [("z","r")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","n")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hcpow, + hcomplex_zero_def])); +by (ultra_tac (claset() addDs [complexpow_zero_zero],simpset()) 1); +qed "hcpow_not_zero"; +Addsimps [hcpow_not_zero]; +AddIs [hcpow_not_zero]; + +Goal "r ^ n = (0::hcomplex) ==> r = 0"; +by (blast_tac (claset() addIs [ccontr] + addDs [hcomplexpow_not_zero]) 1); +qed "hcomplexpow_zero_zero"; + +Goal "r hcpow n = (0::hcomplex) ==> r = 0"; +by (blast_tac (claset() addIs [ccontr] + addDs [hcpow_not_zero]) 1); +qed "hcpow_zero_zero"; + +Goalw [iii_def] "iii * iii = - 1"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult, + hcomplex_one_def,hcomplex_minus])); +qed "hcomplex_i_mult_eq"; +Addsimps [hcomplex_i_mult_eq]; + +Goal "iii ^ 2 = - 1"; +by (simp_tac (simpset() addsimps [two_eq_Suc_Suc]) 1); +qed "hcomplexpow_i_squared"; +Addsimps [hcomplexpow_i_squared]; + +Goalw [iii_def,hcomplex_zero_def] "iii ~= 0"; +by Auto_tac; +qed "hcomplex_i_not_zero"; +Addsimps [hcomplex_i_not_zero]; + +Goal "x * y ~= (0::hcomplex) ==> x ~= 0"; +by Auto_tac; +qed "hcomplex_mult_eq_zero_cancel1"; + +Goal "x * y ~= (0::hcomplex) ==> y ~= 0"; +by Auto_tac; +qed "hcomplex_mult_eq_zero_cancel2"; + +Goal "(x * y ~= (0::hcomplex)) = (x ~= 0 & y ~= 0)"; +by Auto_tac; +qed "hcomplex_mult_not_eq_zero_iff"; +AddIffs [hcomplex_mult_not_eq_zero_iff]; + +Goalw [hcomplex_divide_def,complex_divide_def] + "Abs_hcomplex(hcomplexrel``{%n. X n}) / Abs_hcomplex(hcomplexrel``{%n. Y n}) = \ +\ Abs_hcomplex(hcomplexrel``{%n. X n / Y n})"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse,hcomplex_mult])); +qed "hcomplex_divide"; + +(*---------------------------------------------------------------------------*) +(* hsgn *) +(*---------------------------------------------------------------------------*) + +Goalw [hsgn_def] + "hsgn (Abs_hcomplex(hcomplexrel `` {%n. X n})) = \ +\ Abs_hcomplex(hcomplexrel `` {%n. sgn (X n)})"; +by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1); +by (Auto_tac THEN Ultra_tac 1); +qed "hsgn"; + +Addsimps [rename_numerals sgn_zero]; +Goalw [hcomplex_zero_def] "hsgn 0 = 0"; +by (simp_tac (simpset() addsimps [hsgn]) 1); +qed "hsgn_zero"; +Addsimps[hsgn_zero]; + +Addsimps [rename_numerals sgn_one]; + +Goalw [hcomplex_one_def] "hsgn 1 = 1"; +by (simp_tac (simpset() addsimps [hsgn]) 1); +qed "hsgn_one"; +Addsimps[hsgn_one]; + +Goal "hsgn (-z) = - hsgn(z)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hsgn,hcomplex_minus, + sgn_minus])); +qed "hsgn_minus"; + +Goal "hsgn z = z / hcomplex_of_hypreal (hcmod z)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hsgn,hcomplex_divide, + hcomplex_of_hypreal,hcmod,sgn_eq])); +qed "hsgn_eq"; + +Goal "(EX (x::hypreal) y. P x y) = \ +\ (EX f g. P (Abs_hypreal(hyprel `` {f})) (Abs_hypreal(hyprel `` {g})))"; +by Auto_tac; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by Auto_tac; +qed "lemma_hypreal_P_EX2"; + +Goal "ALL (n::nat). EX x y. (z n) = complex_of_real(x) + ii * complex_of_real(y)"; +by (blast_tac (claset() addIs [complex_split]) 1); +qed "complex_split2"; + +(* Interesting proof! *) +Goal "EX x y. z = hcomplex_of_hypreal(x) + iii * hcomplex_of_hypreal(y)"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [lemma_hypreal_P_EX2, + hcomplex_of_hypreal,iii_def,hcomplex_add,hcomplex_mult])); +by (cut_inst_tac [("z","x")] complex_split2 1); +by (REPEAT(dtac choice 1 THEN Step_tac 1)); +by (res_inst_tac [("x","f")] exI 1); +by (res_inst_tac [("x","fa")] exI 1); +by Auto_tac; +qed "hcomplex_split"; + +Goal "hRe(hcomplex_of_hypreal(x) + iii * hcomplex_of_hypreal(y)) = x"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hRe,iii_def,hcomplex_add, + hcomplex_mult,hcomplex_of_hypreal])); +qed "hRe_hcomplex_i"; +Addsimps [hRe_hcomplex_i]; + +Goal "hIm(hcomplex_of_hypreal(x) + iii * hcomplex_of_hypreal(y)) = y"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hIm,iii_def,hcomplex_add, + hcomplex_mult,hcomplex_of_hypreal])); +qed "hIm_hcomplex_i"; +Addsimps [hIm_hcomplex_i]; + +Goal "hcmod (hcomplex_of_hypreal(x) + iii * hcomplex_of_hypreal(y)) = \ +\ ( *f* sqrt) (x ^ 2 + y ^ 2)"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + iii_def,hcomplex_add,hcomplex_mult,starfun,hypreal_mult, + hypreal_add,hcmod,cmod_i,two_eq_Suc_Suc])); +qed "hcmod_i"; + +Goalw [iii_def] + "hcomplex_of_hypreal xa + iii * hcomplex_of_hypreal ya = \ +\ hcomplex_of_hypreal xb + iii * hcomplex_of_hypreal yb \ +\ ==> xa = xb"; +by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","ya")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","xb")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","yb")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult,hcomplex_add, + hcomplex_of_hypreal])); +by (Ultra_tac 1); +qed "hcomplex_eq_hRe_eq"; + +Goalw [iii_def] + "hcomplex_of_hypreal xa + iii * hcomplex_of_hypreal ya = \ +\ hcomplex_of_hypreal xb + iii * hcomplex_of_hypreal yb \ +\ ==> ya = yb"; +by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","ya")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","xb")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","yb")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult,hcomplex_add, + hcomplex_of_hypreal])); +by (Ultra_tac 1); +qed "hcomplex_eq_hIm_eq"; + +Goal "(hcomplex_of_hypreal xa + iii * hcomplex_of_hypreal ya = \ +\ hcomplex_of_hypreal xb + iii * hcomplex_of_hypreal yb) = \ +\ ((xa = xb) & (ya = yb))"; +by (auto_tac (claset() addIs [hcomplex_eq_hIm_eq,hcomplex_eq_hRe_eq],simpset())); +qed "hcomplex_eq_cancel_iff"; +Addsimps [hcomplex_eq_cancel_iff]; + +Goal "(hcomplex_of_hypreal xa + hcomplex_of_hypreal ya * iii = \ +\ hcomplex_of_hypreal xb + hcomplex_of_hypreal yb * iii ) = ((xa = xb) & (ya = yb))"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_commute])); +qed "hcomplex_eq_cancel_iffA"; +AddIffs [hcomplex_eq_cancel_iffA]; + +Goal "(hcomplex_of_hypreal xa + hcomplex_of_hypreal ya * iii = \ +\ hcomplex_of_hypreal xb + iii * hcomplex_of_hypreal yb) = ((xa = xb) & (ya = yb))"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_commute])); +qed "hcomplex_eq_cancel_iffB"; +AddIffs [hcomplex_eq_cancel_iffB]; + +Goal "(hcomplex_of_hypreal xa + iii * hcomplex_of_hypreal ya = \ +\ hcomplex_of_hypreal xb + hcomplex_of_hypreal yb * iii) = ((xa = xb) & (ya = yb))"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_commute])); +qed "hcomplex_eq_cancel_iffC"; +AddIffs [hcomplex_eq_cancel_iffC]; + +Goal"(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y = \ +\ hcomplex_of_hypreal xa) = (x = xa & y = 0)"; +by (cut_inst_tac [("xa","x"),("ya","y"),("xb","xa"),("yb","0")] + hcomplex_eq_cancel_iff 1); +by (asm_full_simp_tac (simpset() delsimps [hcomplex_eq_cancel_iff]) 1); +qed "hcomplex_eq_cancel_iff2"; +Addsimps [hcomplex_eq_cancel_iff2]; + +Goal"(hcomplex_of_hypreal x + hcomplex_of_hypreal y * iii = \ +\ hcomplex_of_hypreal xa) = (x = xa & y = 0)"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_commute])); +qed "hcomplex_eq_cancel_iff2a"; +Addsimps [hcomplex_eq_cancel_iff2a]; + +Goal "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y = \ +\ iii * hcomplex_of_hypreal ya) = (x = 0 & y = ya)"; +by (cut_inst_tac [("xa","x"),("ya","y"),("xb","0"),("yb","ya")] + hcomplex_eq_cancel_iff 1); +by (asm_full_simp_tac (simpset() delsimps [hcomplex_eq_cancel_iff]) 1); +qed "hcomplex_eq_cancel_iff3"; +Addsimps [hcomplex_eq_cancel_iff3]; + +Goal "(hcomplex_of_hypreal x + hcomplex_of_hypreal y * iii = \ +\ iii * hcomplex_of_hypreal ya) = (x = 0 & y = ya)"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_commute])); +qed "hcomplex_eq_cancel_iff3a"; +Addsimps [hcomplex_eq_cancel_iff3a]; + +Goalw [iii_def] + "hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y = 0 \ +\ ==> x = 0"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + hcomplex_add,hcomplex_mult,hcomplex_zero_def,hypreal_zero_num])); +by (ultra_tac (claset(),simpset() addsimps [complex_split_Re_zero]) 1); +qed "hcomplex_split_hRe_zero"; + +Goalw [iii_def] + "hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y = 0 \ +\ ==> y = 0"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + hcomplex_add,hcomplex_mult,hcomplex_zero_def,hypreal_zero_num])); +by (ultra_tac (claset(),simpset() addsimps [complex_split_Im_zero]) 1); +qed "hcomplex_split_hIm_zero"; + +Goal "hRe(hsgn z) = hRe(z)/hcmod z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hsgn,hcmod,hRe,hypreal_divide])); +qed "hRe_hsgn"; +Addsimps [hRe_hsgn]; + +Goal "hIm(hsgn z) = hIm(z)/hcmod z"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hsgn,hcmod,hIm,hypreal_divide])); +qed "hIm_hsgn"; +Addsimps [hIm_hsgn]; + +Goal "(x*x + y*y = 0) = ((x::real) = 0 & y = 0)"; +by (auto_tac (claset() addIs [real_sum_squares_cancel],simpset())); +qed "real_two_squares_add_zero_iff"; +Addsimps [real_two_squares_add_zero_iff]; + +Goal "inverse(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y) = \ +\ hcomplex_of_hypreal(x/(x ^ 2 + y ^ 2)) - \ +\ iii * hcomplex_of_hypreal(y/(x ^ 2 + y ^ 2))"; +by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + hcomplex_mult,hcomplex_add,iii_def,starfun,hypreal_mult, + hypreal_add,hcomplex_inverse,hypreal_divide,hcomplex_diff, + complex_inverse_complex_split,two_eq_Suc_Suc])); +qed "hcomplex_inverse_complex_split"; + +Goalw [iii_def] + "hRe (iii * hcomplex_of_hypreal y) = 0"; +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + hcomplex_mult,hRe,hypreal_zero_num])); +qed "hRe_mult_i_eq"; +Addsimps [hRe_mult_i_eq]; + +Goalw [iii_def] + "hIm (iii * hcomplex_of_hypreal y) = y"; +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + hcomplex_mult,hIm,hypreal_zero_num])); +qed "hIm_mult_i_eq"; +Addsimps [hIm_mult_i_eq]; + +Goal "hcmod (iii * hcomplex_of_hypreal y) = abs y"; +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + hcmod,hypreal_hrabs,iii_def,hcomplex_mult])); +qed "hcmod_mult_i"; +Addsimps [hcmod_mult_i]; + +Goal "hcmod (hcomplex_of_hypreal y * iii) = abs y"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_commute])); +qed "hcmod_mult_i2"; +Addsimps [hcmod_mult_i2]; + +(*---------------------------------------------------------------------------*) +(* harg *) +(*---------------------------------------------------------------------------*) + +Goalw [harg_def] + "harg (Abs_hcomplex(hcomplexrel `` {%n. X n})) = \ +\ Abs_hypreal(hyprel `` {%n. arg (X n)})"; +by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1); +by (Auto_tac THEN Ultra_tac 1); +qed "harg"; + +Goal "0 < y ==> ( *f* cos) (harg(iii * hcomplex_of_hypreal y)) = 0"; +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + iii_def,hcomplex_mult,hypreal_zero_num,hypreal_less,starfun, + harg])); +by (Ultra_tac 1); +qed "cos_harg_i_mult_zero"; +Addsimps [cos_harg_i_mult_zero]; + +Goal "y < 0 ==> ( *f* cos) (harg(iii * hcomplex_of_hypreal y)) = 0"; +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + iii_def,hcomplex_mult,hypreal_zero_num,hypreal_less,starfun, + harg])); +by (Ultra_tac 1); +qed "cos_harg_i_mult_zero2"; +Addsimps [cos_harg_i_mult_zero2]; + +Goal "(hcomplex_of_hypreal y ~= 0) = (y ~= 0)"; +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + hypreal_zero_num,hcomplex_zero_def])); +qed "hcomplex_of_hypreal_not_zero_iff"; +Addsimps [hcomplex_of_hypreal_not_zero_iff]; + +Goal "(hcomplex_of_hypreal y = 0) = (y = 0)"; +by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal, + hypreal_zero_num,hcomplex_zero_def])); +qed "hcomplex_of_hypreal_zero_iff"; +Addsimps [hcomplex_of_hypreal_zero_iff]; + +Goal "y ~= 0 ==> ( *f* cos) (harg(iii * hcomplex_of_hypreal y)) = 0"; +by (cut_inst_tac [("x","y"),("y","0")] hypreal_linear 1); +by Auto_tac; +qed "cos_harg_i_mult_zero3"; +Addsimps [cos_harg_i_mult_zero3]; + +(*---------------------------------------------------------------------------*) +(* Polar form for nonstandard complex numbers *) +(*---------------------------------------------------------------------------*) + +Goal "ALL n. EX r a. (z n) = complex_of_real r * \ +\ (complex_of_real(cos a) + ii * complex_of_real(sin a))"; +by (blast_tac (claset() addIs [complex_split_polar]) 1); +qed "complex_split_polar2"; + +Goal + "EX r a. z = hcomplex_of_hypreal r * \ +\ (hcomplex_of_hypreal(( *f* cos) a) + iii * hcomplex_of_hypreal(( *f* sin) a))"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [lemma_hypreal_P_EX2, + hcomplex_of_hypreal,iii_def,starfun,hcomplex_add,hcomplex_mult])); +by (cut_inst_tac [("z","x")] complex_split_polar2 1); +by (REPEAT(dtac choice 1 THEN Step_tac 1)); +by (res_inst_tac [("x","f")] exI 1); +by (res_inst_tac [("x","fa")] exI 1); +by Auto_tac; +qed "hcomplex_split_polar"; + +Goalw [hcis_def] + "hcis (Abs_hypreal(hyprel `` {%n. X n})) = \ +\ Abs_hcomplex(hcomplexrel `` {%n. cis (X n)})"; +by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1); +by (Auto_tac THEN Ultra_tac 1); +qed "hcis"; + +Goal + "hcis a = \ +\ (hcomplex_of_hypreal(( *f* cos) a) + \ +\ iii * hcomplex_of_hypreal(( *f* sin) a))"; +by (res_inst_tac [("z","a")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [starfun, hcis, + hcomplex_of_hypreal,iii_def,hcomplex_mult,hcomplex_add, + cis_def])); +qed "hcis_eq"; + +Goalw [hrcis_def] + "hrcis (Abs_hypreal(hyprel `` {%n. X n})) (Abs_hypreal(hyprel `` {%n. Y n})) = \ +\ Abs_hcomplex(hcomplexrel `` {%n. rcis (X n) (Y n)})"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,hcomplex_mult, + hcis,rcis_def])); +qed "hrcis"; + +Goal "EX r a. z = hrcis r a"; +by (simp_tac (simpset() addsimps [hrcis_def,hcis_eq]) 1); +by (rtac hcomplex_split_polar 1); +qed "hrcis_Ex"; + +Goal "hRe(hcomplex_of_hypreal r * \ +\ (hcomplex_of_hypreal(( *f* cos) a) + \ +\ iii * hcomplex_of_hypreal(( *f* sin) a))) = r * ( *f* cos) a"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_add_mult_distrib2, + hcomplex_of_hypreal_mult] @ hcomplex_mult_ac)); +qed "hRe_hcomplex_polar"; +Addsimps [hRe_hcomplex_polar]; + +Goalw [hrcis_def] "hRe(hrcis r a) = r * ( *f* cos) a"; +by (auto_tac (claset(),simpset() addsimps [hcis_eq])); +qed "hRe_hrcis"; +Addsimps [hRe_hrcis]; + +Goal "hIm(hcomplex_of_hypreal r * \ +\ (hcomplex_of_hypreal(( *f* cos) a) + \ +\ iii * hcomplex_of_hypreal(( *f* sin) a))) = r * ( *f* sin) a"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_add_mult_distrib2, + hcomplex_of_hypreal_mult] @ hcomplex_mult_ac)); +qed "hIm_hcomplex_polar"; +Addsimps [hIm_hcomplex_polar]; + +Goalw [hrcis_def] "hIm(hrcis r a) = r * ( *f* sin) a"; +by (auto_tac (claset(),simpset() addsimps [hcis_eq])); +qed "hIm_hrcis"; +Addsimps [hIm_hrcis]; + +Goal "hcmod (hcomplex_of_hypreal r * \ +\ (hcomplex_of_hypreal(( *f* cos) a) + \ +\ iii * hcomplex_of_hypreal(( *f* sin) a))) = abs r"; +by (res_inst_tac [("z","r")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","a")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [iii_def,starfun, + hcomplex_of_hypreal,hcomplex_mult,hcmod,hcomplex_add, + hypreal_hrabs])); +qed "hcmod_complex_polar"; +Addsimps [hcmod_complex_polar]; + +Goalw [hrcis_def] "hcmod(hrcis r a) = abs r"; +by (auto_tac (claset(),simpset() addsimps [hcis_eq])); +qed "hcmod_hrcis"; +Addsimps [hcmod_hrcis]; + +(*---------------------------------------------------------------------------*) +(* (r1 * hrcis a) * (r2 * hrcis b) = r1 * r2 * hrcis (a + b) *) +(*---------------------------------------------------------------------------*) + +Goalw [hrcis_def] "hcis a = hrcis 1 a"; +by (Simp_tac 1); +qed "hcis_hrcis_eq"; +Addsimps [hcis_hrcis_eq RS sym]; + +Goalw [hrcis_def] + "hrcis r1 a * hrcis r2 b = hrcis (r1*r2) (a + b)"; +by (res_inst_tac [("z","r1")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","r2")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","a")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","b")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hrcis,hcis, + hypreal_add,hypreal_mult,hcomplex_of_hypreal, + hcomplex_mult,cis_mult RS sym,complex_of_real_mult + RS sym] addsimps complex_mult_ac)); +qed "hrcis_mult"; + +Goal "hcis a * hcis b = hcis (a + b)"; +by (res_inst_tac [("z","a")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","b")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcis,hcomplex_mult, + hypreal_add,cis_mult])); +qed "hcis_mult"; + +Goalw [hcomplex_one_def] + "hcis 0 = 1"; +by (auto_tac (claset(),simpset() addsimps [hcis,hypreal_zero_num])); +qed "hcis_zero"; +Addsimps [hcis_zero]; + +Goalw [hcomplex_zero_def] + "hrcis 0 a = 0"; +by (res_inst_tac [("z","a")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hrcis,hypreal_zero_num])); +qed "hrcis_zero_mod"; +Addsimps [hrcis_zero_mod]; + +Goal "hrcis r 0 = hcomplex_of_hypreal r"; +by (res_inst_tac [("z","r")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hrcis, + hypreal_zero_num,hcomplex_of_hypreal])); +qed "hrcis_zero_arg"; +Addsimps [hrcis_zero_arg]; + +Goal "iii * (iii * x) = - x"; +by (simp_tac (simpset() addsimps [hcomplex_mult_assoc RS sym]) 1); +qed "hcomplex_i_mult_minus"; +Addsimps [hcomplex_i_mult_minus]; + +Goal "iii * iii * x = - x"; +by (Simp_tac 1); +qed "hcomplex_i_mult_minus2"; +Addsimps [hcomplex_i_mult_minus2]; + +(* Move to one of the hyperreal theories *) +Goal "hypreal_of_nat m = Abs_hypreal(hyprel `` {%n. real m})"; +by (induct_tac "m" 1); +by (auto_tac (claset(),simpset() addsimps [hypreal_zero_def, + hypreal_of_nat_Suc,hypreal_zero_num, + hypreal_one_num,hypreal_add,real_of_nat_Suc])); +qed "hypreal_of_nat"; + +Goal + "hcis (hypreal_of_nat (Suc n) * a) = hcis a * hcis (hypreal_of_nat n * a)"; +by (res_inst_tac [("z","a")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hypreal_of_nat,hcis, + hypreal_mult,hcomplex_mult,cis_real_of_nat_Suc_mult])); +qed "hcis_hypreal_of_nat_Suc_mult"; + +Goal "(hcis a) ^ n = hcis (hypreal_of_nat n * a)"; +by (induct_tac "n" 1); +by (auto_tac (claset(),simpset() addsimps [hcis_hypreal_of_nat_Suc_mult])); +qed "NSDeMoivre"; + +Goal "hcis (hypreal_of_hypnat (n + 1) * a) = \ +\ hcis a * hcis (hypreal_of_hypnat n * a)"; +by (res_inst_tac [("z","a")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","n")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hcis,hypreal_of_hypnat, + hypnat_add,hypnat_one_def,hypreal_mult,hcomplex_mult, + cis_real_of_nat_Suc_mult])); +qed "hcis_hypreal_of_hypnat_Suc_mult"; + +Goal "(hcis a) hcpow n = hcis (hypreal_of_hypnat n * a)"; +by (res_inst_tac [("z","a")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","n")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hcis,hypreal_of_hypnat, + hypreal_mult,hcpow,DeMoivre])); +qed "NSDeMoivre_ext"; + +Goalw [hrcis_def] + "(hrcis r a) ^ n = hrcis (r ^ n) (hypreal_of_nat n * a)"; +by (auto_tac (claset(),simpset() addsimps [hcomplexpow_mult, + NSDeMoivre,hcomplex_of_hypreal_pow])); +qed "DeMoivre2"; + +Goalw [hrcis_def] + "(hrcis r a) hcpow n = hrcis (r pow n) (hypreal_of_hypnat n * a)"; +by (auto_tac (claset(),simpset() addsimps [hcpow_mult, + NSDeMoivre_ext,hcomplex_of_hypreal_hyperpow])); +qed "DeMoivre2_ext"; + +Goal "inverse(hcis a) = hcis (-a)"; +by (res_inst_tac [("z","a")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse,hcis,hypreal_minus])); +qed "hcis_inverse"; +Addsimps [hcis_inverse]; + +Goal "inverse(hrcis r a) = hrcis (inverse r) (-a)"; +by (res_inst_tac [("z","a")] eq_Abs_hypreal 1); +by (res_inst_tac [("z","r")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse,hrcis,hypreal_minus, + hypreal_inverse,rcis_inverse])); +by (Ultra_tac 1); +by (rewtac real_divide_def); +by (auto_tac (claset(),simpset() addsimps [INVERSE_ZERO])); +qed "hrcis_inverse"; + +Goal "hRe(hcis a) = ( *f* cos) a"; +by (res_inst_tac [("z","a")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcis,starfun,hRe])); +qed "hRe_hcis"; +Addsimps [hRe_hcis]; + +Goal "hIm(hcis a) = ( *f* sin) a"; +by (res_inst_tac [("z","a")] eq_Abs_hypreal 1); +by (auto_tac (claset(),simpset() addsimps [hcis,starfun,hIm])); +qed "hIm_hcis"; +Addsimps [hIm_hcis]; + +Goal "( *f* cos) (hypreal_of_nat n * a) = hRe(hcis a ^ n)"; +by (auto_tac (claset(),simpset() addsimps [NSDeMoivre])); +qed "cos_n_hRe_hcis_pow_n"; + +Goal "( *f* sin) (hypreal_of_nat n * a) = hIm(hcis a ^ n)"; +by (auto_tac (claset(),simpset() addsimps [NSDeMoivre])); +qed "sin_n_hIm_hcis_pow_n"; + +Goal "( *f* cos) (hypreal_of_hypnat n * a) = hRe(hcis a hcpow n)"; +by (auto_tac (claset(),simpset() addsimps [NSDeMoivre_ext])); +qed "cos_n_hRe_hcis_hcpow_n"; + +Goal "( *f* sin) (hypreal_of_hypnat n * a) = hIm(hcis a hcpow n)"; +by (auto_tac (claset(),simpset() addsimps [NSDeMoivre_ext])); +qed "sin_n_hIm_hcis_hcpow_n"; + +Goalw [hexpi_def] "hexpi(a + b) = hexpi(a) * hexpi(b)"; +by (res_inst_tac [("z","a")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","b")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hcis,hRe,hIm, + hcomplex_add,hcomplex_mult,hypreal_mult,starfun, + hcomplex_of_hypreal,cis_mult RS sym,complex_Im_add, + complex_Re_add,exp_add,complex_of_real_mult])); +qed "hexpi_add"; + +(*----------------------------------------------------------------------------------*) +(* hcomplex_of_complex preserves field and order properties *) +(*----------------------------------------------------------------------------------*) + +Goalw [hcomplex_of_complex_def] + "hcomplex_of_complex (z1 + z2) = hcomplex_of_complex z1 + hcomplex_of_complex z2"; +by (simp_tac (simpset() addsimps [hcomplex_add]) 1); +qed "hcomplex_of_complex_add"; +Addsimps [hcomplex_of_complex_add]; + +Goalw [hcomplex_of_complex_def] + "hcomplex_of_complex (z1 * z2) = hcomplex_of_complex z1 * hcomplex_of_complex z2"; +by (simp_tac (simpset() addsimps [hcomplex_mult]) 1); +qed "hcomplex_of_complex_mult"; +Addsimps [hcomplex_of_complex_mult]; + +Goalw [hcomplex_of_complex_def] + "(hcomplex_of_complex z1 = hcomplex_of_complex z2) = (z1 = z2)"; +by Auto_tac; +qed "hcomplex_of_complex_eq_iff"; +Addsimps [hcomplex_of_complex_eq_iff]; + +Goalw [hcomplex_of_complex_def] "hcomplex_of_complex (-r) = - hcomplex_of_complex r"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_minus])); +qed "hcomplex_of_complex_minus"; +Addsimps [hcomplex_of_complex_minus]; + +Goalw [hcomplex_of_complex_def,hcomplex_one_def] + "hcomplex_of_complex 1 = 1"; +by Auto_tac; +qed "hcomplex_of_complex_one"; + +Goalw [hcomplex_of_complex_def,hcomplex_zero_def] + "hcomplex_of_complex 0 = 0"; +by (Simp_tac 1); +qed "hcomplex_of_complex_zero"; + +Goal "(hcomplex_of_complex r = 0) = (r = 0)"; +by (auto_tac (claset() addIs [FreeUltrafilterNat_P], + simpset() addsimps [hcomplex_of_complex_def, + hcomplex_zero_def])); +qed "hcomplex_of_complex_zero_iff"; + +Goal "hcomplex_of_complex (inverse r) = inverse (hcomplex_of_complex r)"; +by (case_tac "r=0" 1); +by (asm_simp_tac (simpset() addsimps [COMPLEX_INVERSE_ZERO, + HCOMPLEX_INVERSE_ZERO, hcomplex_of_complex_zero, + COMPLEX_DIVIDE_ZERO]) 1); +by (res_inst_tac [("c1","hcomplex_of_complex r")] + (hcomplex_mult_left_cancel RS iffD1) 1); +by (stac (hcomplex_of_complex_mult RS sym) 2); +by (auto_tac (claset(), + simpset() addsimps [hcomplex_of_complex_one, hcomplex_of_complex_zero_iff])); +qed "hcomplex_of_complex_inverse"; +Addsimps [hcomplex_of_complex_inverse]; + +Goal "hcomplex_of_complex (z1 / z2) = hcomplex_of_complex z1 / hcomplex_of_complex z2"; +by (simp_tac (simpset() addsimps [hcomplex_divide_def, complex_divide_def]) 1); +qed "hcomplex_of_complex_divide"; +Addsimps [hcomplex_of_complex_divide]; + +Goalw [hcomplex_of_complex_def,hypreal_of_real_def] + "hRe (hcomplex_of_complex z) = hypreal_of_real (Re z)"; +by (auto_tac (claset(),simpset() addsimps [hRe])); +qed "hRe_hcomplex_of_complex"; + +Goalw [hcomplex_of_complex_def,hypreal_of_real_def] + "hIm (hcomplex_of_complex z) = hypreal_of_real (Im z)"; +by (auto_tac (claset(),simpset() addsimps [hIm])); +qed "hIm_hcomplex_of_complex"; + +Goalw [hypreal_of_real_def,hcomplex_of_complex_def] + "hcmod (hcomplex_of_complex x) = hypreal_of_real (cmod x)"; +by (auto_tac (claset(),simpset() addsimps [hcmod])); +qed "hcmod_hcomplex_of_complex"; diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/NSComplex.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/NSComplex.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,127 @@ +(* Title: NSComplex.thy + Author: Jacques D. Fleuriot + Copyright: 2001 University of Edinburgh + Description: Nonstandard Complex numbers +*) + +NSComplex = NSInduct + + +constdefs + hcomplexrel :: "((nat=>complex)*(nat=>complex)) set" + "hcomplexrel == {p. EX X Y. p = ((X::nat=>complex),Y) & + {n::nat. X(n) = Y(n)}: FreeUltrafilterNat}" + +typedef hcomplex = "{x::nat=>complex. True}//hcomplexrel" (quotient_def) + +instance + hcomplex :: {zero,one,plus,times,minus,power,inverse} + +defs + hcomplex_zero_def + "0 == Abs_hcomplex(hcomplexrel `` {%n. (0::complex)})" + + hcomplex_one_def + "1 == Abs_hcomplex(hcomplexrel `` {%n. (1::complex)})" + + + hcomplex_minus_def + "- z == Abs_hcomplex(UN X: Rep_hcomplex(z). hcomplexrel `` {%n::nat. - (X n)})" + + hcomplex_diff_def + "w - z == w + -(z::hcomplex)" + +constdefs + + hcomplex_of_complex :: complex => hcomplex + "hcomplex_of_complex z == Abs_hcomplex(hcomplexrel `` {%n. z})" + + hcinv :: hcomplex => hcomplex + "inverse(P) == Abs_hcomplex(UN X: Rep_hcomplex(P). + hcomplexrel `` {%n. inverse(X n)})" + + (*--- real and Imaginary parts ---*) + + hRe :: hcomplex => hypreal + "hRe(z) == Abs_hypreal(UN X:Rep_hcomplex(z). hyprel `` {%n. Re (X n)})" + + hIm :: hcomplex => hypreal + "hIm(z) == Abs_hypreal(UN X:Rep_hcomplex(z). hyprel `` {%n. Im (X n)})" + + + (*----------- modulus ------------*) + + hcmod :: hcomplex => hypreal + "hcmod z == Abs_hypreal(UN X: Rep_hcomplex(z). + hyprel `` {%n. cmod (X n)})" + + (*------ imaginary unit ----------*) + + iii :: hcomplex + "iii == Abs_hcomplex(hcomplexrel `` {%n. ii})" + + (*------- complex conjugate ------*) + + hcnj :: hcomplex => hcomplex + "hcnj z == Abs_hcomplex(UN X:Rep_hcomplex(z). hcomplexrel `` {%n. cnj (X n)})" + + (*------------ Argand -------------*) + + hsgn :: hcomplex => hcomplex + "hsgn z == Abs_hcomplex(UN X:Rep_hcomplex(z). hcomplexrel `` {%n. sgn(X n)})" + + harg :: hcomplex => hypreal + "harg z == Abs_hypreal(UN X:Rep_hcomplex(z). hyprel `` {%n. arg(X n)})" + + (* abbreviation for (cos a + i sin a) *) + hcis :: hypreal => hcomplex + "hcis a == Abs_hcomplex(UN X:Rep_hypreal(a). hcomplexrel `` {%n. cis (X n)})" + + (* abbreviation for r*(cos a + i sin a) *) + hrcis :: [hypreal, hypreal] => hcomplex + "hrcis r a == hcomplex_of_hypreal r * hcis a" + + (*----- injection from hyperreals -----*) + + hcomplex_of_hypreal :: hypreal => hcomplex + "hcomplex_of_hypreal r == Abs_hcomplex(UN X:Rep_hypreal(r). + hcomplexrel `` {%n. complex_of_real (X n)})" + + (*------------ e ^ (x + iy) ------------*) + + hexpi :: hcomplex => hcomplex + "hexpi z == hcomplex_of_hypreal(( *f* exp) (hRe z)) * hcis (hIm z)" + + +defs + + + (*----------- division ----------*) + + hcomplex_divide_def + "w / (z::hcomplex) == w * inverse z" + + hcomplex_add_def + "w + z == Abs_hcomplex(UN X:Rep_hcomplex(w). UN Y:Rep_hcomplex(z). + hcomplexrel `` {%n. X n + Y n})" + + hcomplex_mult_def + "w * z == Abs_hcomplex(UN X:Rep_hcomplex(w). UN Y:Rep_hcomplex(z). + hcomplexrel `` {%n. X n * Y n})" + + +primrec + hcomplexpow_0 "z ^ 0 = 1" + hcomplexpow_Suc "z ^ (Suc n) = (z::hcomplex) * (z ^ n)" + + +consts + "hcpow" :: [hcomplex,hypnat] => hcomplex (infixr 80) + +defs + (* hypernatural powers of nonstandard complex numbers *) + hcpow_def + "(z::hcomplex) hcpow (n::hypnat) + == Abs_hcomplex(UN X:Rep_hcomplex(z). UN Y: Rep_hypnat(n). + hcomplexrel `` {%n. (X n) ^ (Y n)})" + +end diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/NSComplexArith0.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/NSComplexArith0.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,331 @@ +(* Title: NSComplexArith0.ML + Author: Jacques D. Fleuriot + Copyright: 2001 University of Edinburgh + Description: Assorted facts that need binary literals + Also, common factor cancellation (see e.g. HyperArith0) +*) + +(**** +Goal "((x * y = #0) = (x = #0 | y = (#0::hcomplex)))"; +by (auto_tac (claset(),simpset() addsimps [rename_numerals hcomplex_mult_zero_iff])); +qed "hcomplex_mult_is_0"; +AddIffs [hcomplex_mult_is_0]; +****) + +(** Division and inverse **) + +Goal "0/x = (0::hcomplex)"; +by (simp_tac (simpset() addsimps [hcomplex_divide_def]) 1); +qed "hcomplex_0_divide"; +Addsimps [hcomplex_0_divide]; + +Goalw [hcomplex_divide_def] "x/(0::hcomplex) = 0"; +by (stac HCOMPLEX_INVERSE_ZERO 1); +by (Simp_tac 1); +qed "HCOMPLEX_DIVIDE_ZERO"; + +Goal "inverse (x::hcomplex) = 1/x"; +by (simp_tac (simpset() addsimps [hcomplex_divide_def]) 1); +qed "hcomplex_inverse_eq_divide"; + +Goal "(inverse(x::hcomplex) = 0) = (x = 0)"; +by (auto_tac (claset(), + simpset() addsimps [HCOMPLEX_INVERSE_ZERO])); +by (blast_tac (claset() addIs [ccontr] addDs [hcomplex_inverse_not_zero]) 1); +qed "hcomplex_inverse_zero_iff"; +Addsimps [hcomplex_inverse_zero_iff]; + +Goal "(x/y = 0) = (x=0 | y=(0::hcomplex))"; +by (auto_tac (claset(), simpset() addsimps [hcomplex_divide_def])); +qed "hcomplex_divide_eq_0_iff"; +Addsimps [hcomplex_divide_eq_0_iff]; + +Goal "h ~= (0::hcomplex) ==> h/h = 1"; +by (asm_simp_tac + (simpset() addsimps [hcomplex_divide_def]) 1); +qed "hcomplex_divide_self_eq"; +Addsimps [hcomplex_divide_self_eq]; + +bind_thm ("hcomplex_mult_minus_right", hcomplex_minus_mult_eq2 RS sym); + +Goal "!!k::hcomplex. (k*m = k*n) = (k = 0 | m=n)"; +by (case_tac "k=0" 1); +by (auto_tac (claset(), simpset() addsimps [hcomplex_mult_left_cancel])); +qed "hcomplex_mult_eq_cancel1"; + +Goal "!!k::hcomplex. (m*k = n*k) = (k = 0 | m=n)"; +by (case_tac "k=0" 1); +by (auto_tac (claset(), simpset() addsimps [hcomplex_mult_right_cancel])); +qed "hcomplex_mult_eq_cancel2"; + +Goal "!!k::hcomplex. k~=0 ==> (k*m) / (k*n) = (m/n)"; +by (asm_simp_tac + (simpset() addsimps [hcomplex_divide_def, hcomplex_inverse_distrib]) 1); +by (subgoal_tac "k * m * (inverse k * inverse n) = \ +\ (k * inverse k) * (m * inverse n)" 1); +by (Asm_full_simp_tac 1); +by (asm_full_simp_tac (HOL_ss addsimps hcomplex_mult_ac) 1); +qed "hcomplex_mult_div_cancel1"; + +(*For ExtractCommonTerm*) +Goal "(k*m) / (k*n) = (if k = (0::hcomplex) then 0 else m/n)"; +by (simp_tac (simpset() addsimps [hcomplex_mult_div_cancel1]) 1); +qed "hcomplex_mult_div_cancel_disj"; + + +local + open HComplex_Numeral_Simprocs +in + +val rel_hcomplex_number_of = [eq_hcomplex_number_of]; + + +structure CancelNumeralFactorCommon = + struct + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff 1 + val trans_tac = Real_Numeral_Simprocs.trans_tac + val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps hcomplex_minus_from_mult_simps @ mult_1s)) + THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@hcomplex_mult_minus_simps)) + THEN ALLGOALS (simp_tac (HOL_ss addsimps hcomplex_mult_ac)) + val numeral_simp_tac = + ALLGOALS (simp_tac (HOL_ss addsimps rel_hcomplex_number_of@bin_simps)) + val simplify_meta_eq = simplify_meta_eq + end + + +structure DivCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Bin_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binop "HOL.divide" + val dest_bal = HOLogic.dest_bin "HOL.divide" hcomplexT + val cancel = hcomplex_mult_div_cancel1 RS trans + val neg_exchanges = false +) + + +structure EqCancelNumeralFactor = CancelNumeralFactorFun + (open CancelNumeralFactorCommon + val prove_conv = Bin_Simprocs.prove_conv + val mk_bal = HOLogic.mk_eq + val dest_bal = HOLogic.dest_bin "op =" hcomplexT + val cancel = hcomplex_mult_eq_cancel1 RS trans + val neg_exchanges = false +) + + +val hcomplex_cancel_numeral_factors_relations = + map prep_simproc + [("hcomplexeq_cancel_numeral_factor", + ["(l::hcomplex) * m = n", "(l::hcomplex) = m * n"], + EqCancelNumeralFactor.proc)]; + +val hcomplex_cancel_numeral_factors_divide = prep_simproc + ("hcomplexdiv_cancel_numeral_factor", + ["((l::hcomplex) * m) / n", "(l::hcomplex) / (m * n)", + "((number_of v)::hcomplex) / (number_of w)"], + DivCancelNumeralFactor.proc); + +val hcomplex_cancel_numeral_factors = + hcomplex_cancel_numeral_factors_relations @ + [hcomplex_cancel_numeral_factors_divide]; + +end; + + +Addsimprocs hcomplex_cancel_numeral_factors; + + +(*examples: +print_depth 22; +set timing; +set trace_simp; +fun test s = (Goal s; by (Simp_tac 1)); + + +test "#9*x = #12 * (y::hcomplex)"; +test "(#9*x) / (#12 * (y::hcomplex)) = z"; + +test "#-99*x = #132 * (y::hcomplex)"; + +test "#999*x = #-396 * (y::hcomplex)"; +test "(#999*x) / (#-396 * (y::hcomplex)) = z"; + +test "#-99*x = #-81 * (y::hcomplex)"; +test "(#-99*x) / (#-81 * (y::hcomplex)) = z"; + +test "#-2 * x = #-1 * (y::hcomplex)"; +test "#-2 * x = -(y::hcomplex)"; +test "(#-2 * x) / (#-1 * (y::hcomplex)) = z"; + +*) + + +(** Declarations for ExtractCommonTerm **) + +local + open HComplex_Numeral_Simprocs +in + +structure CancelFactorCommon = + struct + val mk_sum = long_mk_prod + val dest_sum = dest_prod + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff + val find_first = find_first [] + val trans_tac = Real_Numeral_Simprocs.trans_tac + val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@hcomplex_mult_ac)) + end; + + +structure EqCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Bin_Simprocs.prove_conv + val mk_bal = HOLogic.mk_eq + val dest_bal = HOLogic.dest_bin "op =" hcomplexT + val simplify_meta_eq = cancel_simplify_meta_eq hcomplex_mult_eq_cancel1 +); + + +structure DivideCancelFactor = ExtractCommonTermFun + (open CancelFactorCommon + val prove_conv = Bin_Simprocs.prove_conv + val mk_bal = HOLogic.mk_binop "HOL.divide" + val dest_bal = HOLogic.dest_bin "HOL.divide" hcomplexT + val simplify_meta_eq = cancel_simplify_meta_eq hcomplex_mult_div_cancel_disj +); + +val hcomplex_cancel_factor = + map prep_simproc + [("hcomplex_eq_cancel_factor", ["(l::hcomplex) * m = n", "(l::hcomplex) = m * n"], + EqCancelFactor.proc), + ("hcomplex_divide_cancel_factor", ["((l::hcomplex) * m) / n", "(l::hcomplex) / (m * n)"], + DivideCancelFactor.proc)]; + +end; + +Addsimprocs hcomplex_cancel_factor; + + +(*examples: +print_depth 22; +set timing; +set trace_simp; +fun test s = (Goal s; by (Asm_simp_tac 1)); + +test "x*k = k*(y::hcomplex)"; +test "k = k*(y::hcomplex)"; +test "a*(b*c) = (b::hcomplex)"; +test "a*(b*c) = d*(b::hcomplex)*(x*a)"; + + +test "(x*k) / (k*(y::hcomplex)) = (uu::hcomplex)"; +test "(k) / (k*(y::hcomplex)) = (uu::hcomplex)"; +test "(a*(b*c)) / ((b::hcomplex)) = (uu::hcomplex)"; +test "(a*(b*c)) / (d*(b::hcomplex)*(x*a)) = (uu::hcomplex)"; + +(*FIXME: what do we do about this?*) +test "a*(b*c)/(y*z) = d*(b::hcomplex)*(x*a)/z"; +*) + + +Goal "z~=0 ==> ((x::hcomplex) = y/z) = (x*z = y)"; +by (subgoal_tac "(x*z = y) = (x*z = (y/z)*z)" 1); +by (asm_simp_tac (simpset() addsimps [hcomplex_divide_def, hcomplex_mult_assoc]) 2); +by (etac ssubst 1); +by (stac hcomplex_mult_eq_cancel2 1); +by (Asm_simp_tac 1); +qed "hcomplex_eq_divide_eq"; +Addsimps [inst "z" "number_of ?w" hcomplex_eq_divide_eq]; + +Goal "z~=0 ==> (y/z = (x::hcomplex)) = (y = x*z)"; +by (subgoal_tac "(y = x*z) = ((y/z)*z = x*z)" 1); +by (asm_simp_tac (simpset() addsimps [hcomplex_divide_def, hcomplex_mult_assoc]) 2); +by (etac ssubst 1); +by (stac hcomplex_mult_eq_cancel2 1); +by (Asm_simp_tac 1); +qed "hcomplex_divide_eq_eq"; +Addsimps [inst "z" "number_of ?w" hcomplex_divide_eq_eq]; + +Goal "(m/k = n/k) = (k = 0 | m = (n::hcomplex))"; +by (case_tac "k=0" 1); +by (asm_simp_tac (simpset() addsimps [HCOMPLEX_DIVIDE_ZERO]) 1); +by (asm_simp_tac (simpset() addsimps [hcomplex_divide_eq_eq, hcomplex_eq_divide_eq, + hcomplex_mult_eq_cancel2]) 1); +qed "hcomplex_divide_eq_cancel2"; + +Goal "(k/m = k/n) = (k = 0 | m = (n::hcomplex))"; +by (case_tac "m=0 | n = 0" 1); +by (auto_tac (claset(), + simpset() addsimps [HCOMPLEX_DIVIDE_ZERO, hcomplex_divide_eq_eq, + hcomplex_eq_divide_eq, hcomplex_mult_eq_cancel1])); +qed "hcomplex_divide_eq_cancel1"; + +(** Division by 1, -1 **) + +Goal "(x::hcomplex)/1 = x"; +by (simp_tac (simpset() addsimps [hcomplex_divide_def]) 1); +qed "hcomplex_divide_1"; +Addsimps [hcomplex_divide_1]; + +Goal "x/-1 = -(x::hcomplex)"; +by (Simp_tac 1); +qed "hcomplex_divide_minus1"; +Addsimps [hcomplex_divide_minus1]; + +Goal "-1/(x::hcomplex) = - (1/x)"; +by (simp_tac (simpset() addsimps [hcomplex_divide_def, hcomplex_minus_inverse]) 1); +qed "hcomplex_minus1_divide"; +Addsimps [hcomplex_minus1_divide]; + + +Goal "(x = - y) = (y = - (x::hcomplex))"; +by Auto_tac; +qed "hcomplex_equation_minus"; + +Goal "(- x = y) = (- (y::hcomplex) = x)"; +by Auto_tac; +qed "hcomplex_minus_equation"; + +Goal "(x + - a = (0::hcomplex)) = (x=a)"; +by (simp_tac (simpset() addsimps [hcomplex_diff_eq_eq,symmetric hcomplex_diff_def]) 1); +qed "hcomplex_add_minus_iff"; +Addsimps [hcomplex_add_minus_iff]; + +Goal "(-b = -a) = (b = (a::hcomplex))"; +by Auto_tac; +by (etac ( inj_hcomplex_minus RS injD) 1); +qed "hcomplex_minus_eq_cancel"; +Addsimps [hcomplex_minus_eq_cancel]; + +(*Distributive laws for literals*) +Addsimps (map (inst "w" "number_of ?v") + [hcomplex_add_mult_distrib, hcomplex_add_mult_distrib2, + hcomplex_diff_mult_distrib, hcomplex_diff_mult_distrib2]); + +Addsimps [inst "x" "number_of ?v" hcomplex_equation_minus]; + +Addsimps [inst "y" "number_of ?v" hcomplex_minus_equation]; + +Goal "(x+y = (0::hcomplex)) = (y = -x)"; +by Auto_tac; +by (dtac (sym RS (hcomplex_diff_eq_eq RS iffD2)) 1); +by Auto_tac; +qed "hcomplex_add_eq_0_iff"; +AddIffs [hcomplex_add_eq_0_iff]; + +Goalw [hcomplex_diff_def]"-(x-y) = y - (x::hcomplex)"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_add_commute])); +qed "hcomplex_minus_diff_eq"; +Addsimps [hcomplex_minus_diff_eq]; + +Addsimps [inst "x" "number_of ?w" hcomplex_inverse_eq_divide]; + +Goal "[|(x::hcomplex) ~= 0; y ~= 0 |] \ +\ ==> inverse(x) + inverse(y) = (x + y)*inverse(x*y)"; +by (asm_full_simp_tac (simpset() addsimps [hcomplex_inverse_distrib, + hcomplex_add_mult_distrib,hcomplex_mult_assoc RS sym]) 1); +qed "hcomplex_inverse_add"; + +Addsimps [hcomplex_of_complex_zero,hcomplex_of_complex_one]; diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/NSComplexArith0.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/NSComplexArith0.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,4 @@ +NSComplexArith0 = NSComplexBin + + + diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/NSComplexBin.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/NSComplexBin.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,637 @@ +(* Title: NSComplexBin.ML + Author: Jacques D. Fleuriot + Copyright: 2001 University of Edinburgh + Descrition: Binary arithmetic for the nonstandard complex numbers +*) + +(** hcomplex_of_complex (coercion from complex to nonstandard complex) **) + +Goal "hcomplex_of_complex (number_of w) = number_of w"; +by (simp_tac (simpset() addsimps [hcomplex_number_of_def]) 1); +qed "hcomplex_number_of"; +Addsimps [hcomplex_number_of]; + +Goalw [hypreal_of_real_def] + "hcomplex_of_hypreal (hypreal_of_real x) = \ +\ hcomplex_of_complex(complex_of_real x)"; +by (simp_tac (simpset() addsimps [hcomplex_of_hypreal, + hcomplex_of_complex_def,complex_of_real_def]) 1); +qed "hcomplex_of_hypreal_eq_hcomplex_of_complex"; + +Goalw [complex_number_of_def,hypreal_number_of_def] + "hcomplex_of_complex (number_of w) = hcomplex_of_hypreal(number_of w)"; +by (rtac (hcomplex_of_hypreal_eq_hcomplex_of_complex RS sym) 1); +qed "hcomplex_hypreal_number_of"; + +Goalw [hcomplex_number_of_def] "Numeral0 = (0::hcomplex)"; +by (simp_tac (simpset() addsimps [hcomplex_of_complex_zero RS sym]) 1); +qed "hcomplex_numeral_0_eq_0"; + +Goalw [hcomplex_number_of_def] "Numeral1 = (1::hcomplex)"; +by (simp_tac (simpset() addsimps [hcomplex_of_complex_one RS sym]) 1); +qed "hcomplex_numeral_1_eq_1"; + +(* +Goal "z + hcnj z = \ +\ hcomplex_of_hypreal (2 * hRe(z))"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),HOL_ss addsimps [hRe,hcnj,hcomplex_add, + hypreal_mult,hcomplex_of_hypreal,complex_add_cnj])); +qed "hcomplex_add_hcnj"; + +Goal "z - hcnj z = \ +\ hcomplex_of_hypreal (hypreal_of_real #2 * hIm(z)) * iii"; +by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); +by (auto_tac (claset(),simpset() addsimps [hIm,hcnj,hcomplex_diff, + hypreal_of_real_def,hypreal_mult,hcomplex_of_hypreal, + complex_diff_cnj,iii_def,hcomplex_mult])); +qed "hcomplex_diff_hcnj"; +*) + +(** Addition **) + +Goal "(number_of v :: hcomplex) + number_of v' = number_of (bin_add v v')"; +by (simp_tac + (HOL_ss addsimps [hcomplex_number_of_def, + hcomplex_of_complex_add RS sym, add_complex_number_of]) 1); +qed "add_hcomplex_number_of"; +Addsimps [add_hcomplex_number_of]; + + +(** Subtraction **) + +Goalw [hcomplex_number_of_def] + "- (number_of w :: hcomplex) = number_of (bin_minus w)"; +by (simp_tac + (HOL_ss addsimps [minus_complex_number_of, hcomplex_of_complex_minus RS sym]) 1); +qed "minus_hcomplex_number_of"; +Addsimps [minus_hcomplex_number_of]; + +Goalw [hcomplex_number_of_def, hcomplex_diff_def] + "(number_of v :: hcomplex) - number_of w = \ +\ number_of (bin_add v (bin_minus w))"; +by (Simp_tac 1); +qed "diff_hcomplex_number_of"; +Addsimps [diff_hcomplex_number_of]; + + +(** Multiplication **) + +Goal "(number_of v :: hcomplex) * number_of v' = number_of (bin_mult v v')"; +by (simp_tac + (HOL_ss addsimps [hcomplex_number_of_def, + hcomplex_of_complex_mult RS sym, mult_complex_number_of]) 1); +qed "mult_hcomplex_number_of"; +Addsimps [mult_hcomplex_number_of]; + +Goal "(2::hcomplex) = 1 + 1"; +by (simp_tac (simpset() addsimps [hcomplex_numeral_1_eq_1 RS sym]) 1); +val lemma = result(); + +(*For specialist use: NOT as default simprules*) +Goal "2 * z = (z+z::hcomplex)"; +by (simp_tac (simpset() addsimps [lemma, hcomplex_add_mult_distrib]) 1); +qed "hcomplex_mult_2"; + +Goal "z * 2 = (z+z::hcomplex)"; +by (stac hcomplex_mult_commute 1 THEN rtac hcomplex_mult_2 1); +qed "hcomplex_mult_2_right"; + +(** Equals (=) **) + +Goal "((number_of v :: hcomplex) = number_of v') = \ +\ iszero (number_of (bin_add v (bin_minus v')))"; +by (simp_tac + (HOL_ss addsimps [hcomplex_number_of_def, + hcomplex_of_complex_eq_iff, eq_complex_number_of]) 1); +qed "eq_hcomplex_number_of"; +Addsimps [eq_hcomplex_number_of]; + +(*** New versions of existing theorems involving 0, 1hc ***) + +Goal "- 1 = (-1::hcomplex)"; +by (simp_tac (simpset() addsimps [hcomplex_numeral_1_eq_1 RS sym]) 1); +qed "hcomplex_minus_1_eq_m1"; + +Goal "-1 * z = -(z::hcomplex)"; +by (simp_tac (simpset() addsimps [hcomplex_minus_1_eq_m1 RS sym]) 1); +qed "hcomplex_mult_minus1"; + +Goal "z * -1 = -(z::hcomplex)"; +by (stac hcomplex_mult_commute 1 THEN rtac hcomplex_mult_minus1 1); +qed "hcomplex_mult_minus1_right"; + +Addsimps [hcomplex_mult_minus1,hcomplex_mult_minus1_right]; + +(*Maps 0 to Numeral0 and 1 to Numeral1 and -Numeral1 to -1*) +val hcomplex_numeral_ss = + complex_numeral_ss addsimps [hcomplex_numeral_0_eq_0 RS sym, hcomplex_numeral_1_eq_1 RS sym, + hcomplex_minus_1_eq_m1]; + +fun rename_numerals th = + asm_full_simplify hcomplex_numeral_ss (Thm.transfer (the_context ()) th); + + +(*Now insert some identities previously stated for 0 and 1hc*) + + +Addsimps [hcomplex_numeral_0_eq_0,hcomplex_numeral_1_eq_1]; + +Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::hcomplex)"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_add_assoc RS sym])); +qed "hcomplex_add_number_of_left"; + +Goal "number_of v *(number_of w * z) = (number_of(bin_mult v w) * z::hcomplex)"; +by (simp_tac (simpset() addsimps [hcomplex_mult_assoc RS sym]) 1); +qed "hcomplex_mult_number_of_left"; + +Goalw [hcomplex_diff_def] + "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::hcomplex)"; +by (rtac hcomplex_add_number_of_left 1); +qed "hcomplex_add_number_of_diff1"; + +Goal "number_of v + (c - number_of w) = \ +\ number_of (bin_add v (bin_minus w)) + (c::hcomplex)"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_def]@ hcomplex_add_ac)); +qed "hcomplex_add_number_of_diff2"; + +Addsimps [hcomplex_add_number_of_left, hcomplex_mult_number_of_left, + hcomplex_add_number_of_diff1, hcomplex_add_number_of_diff2]; + + +(**** Simprocs for numeric literals ****) + +(** Combining of literal coefficients in sums of products **) + +Goal "(x = y) = (x-y = (0::hcomplex))"; +by (simp_tac (simpset() addsimps [hcomplex_diff_eq_eq]) 1); +qed "hcomplex_eq_iff_diff_eq_0"; + +(** For combine_numerals **) + +Goal "i*u + (j*u + k) = (i+j)*u + (k::hcomplex)"; +by (asm_simp_tac (simpset() addsimps [hcomplex_add_mult_distrib] + @ hcomplex_add_ac) 1); +qed "left_hcomplex_add_mult_distrib"; + +(** For cancel_numerals **) + +Goal "((x::hcomplex) = u + v) = (x - (u + v) = 0)"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_eq_eq])); +qed "hcomplex_eq_add_diff_eq_0"; + +Goal "((x::hcomplex) = n) = (x - n = 0)"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_eq_eq])); +qed "hcomplex_eq_diff_eq_0"; + +val hcomplex_rel_iff_rel_0_rls = [hcomplex_eq_diff_eq_0,hcomplex_eq_add_diff_eq_0]; + +Goal "!!i::hcomplex. (i*u + m = j*u + n) = ((i-j)*u + m = n)"; +by (auto_tac (claset(), simpset() addsimps [hcomplex_add_mult_distrib, + hcomplex_diff_def] @ hcomplex_add_ac)); +by (asm_simp_tac (simpset() addsimps [hcomplex_add_assoc RS sym]) 1); +by (simp_tac (simpset() addsimps [hcomplex_add_assoc]) 1); +qed "hcomplex_eq_add_iff1"; + +Goal "!!i::hcomplex. (i*u + m = j*u + n) = (m = (j-i)*u + n)"; +by (res_inst_tac [("z","i")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","j")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","u")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","m")] eq_Abs_hcomplex 1); +by (res_inst_tac [("z","n")] eq_Abs_hcomplex 1); +by (auto_tac (claset(), simpset() addsimps [hcomplex_diff,hcomplex_add, + hcomplex_mult,complex_eq_add_iff2])); +qed "hcomplex_eq_add_iff2"; + + +structure HComplex_Numeral_Simprocs = +struct + +(*Utilities*) + +val hcomplexT = Type("NSComplex.hcomplex",[]); + +fun mk_numeral n = HOLogic.number_of_const hcomplexT $ HOLogic.mk_bin n; + +val dest_numeral = Complex_Numeral_Simprocs.dest_numeral; + +val find_first_numeral = Complex_Numeral_Simprocs.find_first_numeral; + +val zero = mk_numeral 0; +val mk_plus = HOLogic.mk_binop "op +"; + +val uminus_const = Const ("uminus", hcomplexT --> hcomplexT); + +(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) +fun mk_sum [] = zero + | mk_sum [t,u] = mk_plus (t, u) + | mk_sum (t :: ts) = mk_plus (t, mk_sum ts); + +(*this version ALWAYS includes a trailing zero*) +fun long_mk_sum [] = zero + | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts); + +val dest_plus = HOLogic.dest_bin "op +" hcomplexT; + +(*decompose additions AND subtractions as a sum*) +fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) = + dest_summing (pos, t, dest_summing (pos, u, ts)) + | dest_summing (pos, Const ("op -", _) $ t $ u, ts) = + dest_summing (pos, t, dest_summing (not pos, u, ts)) + | dest_summing (pos, t, ts) = + if pos then t::ts else uminus_const$t :: ts; + +fun dest_sum t = dest_summing (true, t, []); + +val mk_diff = HOLogic.mk_binop "op -"; +val dest_diff = HOLogic.dest_bin "op -" hcomplexT; + +val one = mk_numeral 1; +val mk_times = HOLogic.mk_binop "op *"; + +fun mk_prod [] = one + | mk_prod [t] = t + | mk_prod (t :: ts) = if t = one then mk_prod ts + else mk_times (t, mk_prod ts); + +val dest_times = HOLogic.dest_bin "op *" hcomplexT; + +fun dest_prod t = + let val (t,u) = dest_times t + in dest_prod t @ dest_prod u end + handle TERM _ => [t]; + +(*DON'T do the obvious simplifications; that would create special cases*) +fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts); + +(*Express t as a product of (possibly) a numeral with other sorted terms*) +fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t + | dest_coeff sign t = + let val ts = sort Term.term_ord (dest_prod t) + val (n, ts') = find_first_numeral [] ts + handle TERM _ => (1, ts) + in (sign*n, mk_prod ts') end; + +(*Find first coefficient-term THAT MATCHES u*) +fun find_first_coeff past u [] = raise TERM("find_first_coeff", []) + | find_first_coeff past u (t::terms) = + let val (n,u') = dest_coeff 1 t + in if u aconv u' then (n, rev past @ terms) + else find_first_coeff (t::past) u terms + end + handle TERM _ => find_first_coeff (t::past) u terms; + + + +(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*) +val add_0s = map rename_numerals [hcomplex_add_zero_left, hcomplex_add_zero_right]; +val mult_plus_1s = map rename_numerals [hcomplex_mult_one_left, hcomplex_mult_one_right]; +val mult_minus_1s = map rename_numerals [hcomplex_mult_minus1, hcomplex_mult_minus1_right]; +val mult_1s = mult_plus_1s @ mult_minus_1s; + +(*To perform binary arithmetic*) +val bin_simps = + [hcomplex_numeral_0_eq_0 RS sym, hcomplex_numeral_1_eq_1 RS sym, + add_hcomplex_number_of, hcomplex_add_number_of_left, + minus_hcomplex_number_of, diff_hcomplex_number_of, mult_hcomplex_number_of, + hcomplex_mult_number_of_left] @ bin_arith_simps @ bin_rel_simps; + +(*To evaluate binary negations of coefficients*) +val hcomplex_minus_simps = NCons_simps @ + [hcomplex_minus_1_eq_m1,minus_hcomplex_number_of, + bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min, + bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min]; + + +(*To let us treat subtraction as addition*) +val diff_simps = [hcomplex_diff_def, hcomplex_minus_add_distrib, + hcomplex_minus_minus]; + +(*push the unary minus down: - x * y = x * - y *) +val hcomplex_minus_mult_eq_1_to_2 = + [hcomplex_minus_mult_eq1 RS sym, hcomplex_minus_mult_eq2] MRS trans + |> standard; + +(*to extract again any uncancelled minuses*) +val hcomplex_minus_from_mult_simps = + [hcomplex_minus_minus, hcomplex_minus_mult_eq1 RS sym, + hcomplex_minus_mult_eq2 RS sym]; + +(*combine unary minus with numeric literals, however nested within a product*) +val hcomplex_mult_minus_simps = + [hcomplex_mult_assoc, hcomplex_minus_mult_eq1, hcomplex_minus_mult_eq_1_to_2]; + +(*Final simplification: cancel + and * *) +val simplify_meta_eq = + Int_Numeral_Simprocs.simplify_meta_eq + [hcomplex_add_zero_left, hcomplex_add_zero_right, + hcomplex_mult_zero_left, hcomplex_mult_zero_right, hcomplex_mult_one_left, + hcomplex_mult_one_right]; + +val prep_simproc = Complex_Numeral_Simprocs.prep_simproc; + + +structure CancelNumeralsCommon = + struct + val mk_sum = mk_sum + val dest_sum = dest_sum + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff 1 + val find_first_coeff = find_first_coeff [] + val trans_tac = Real_Numeral_Simprocs.trans_tac + val norm_tac = + ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@ + hcomplex_minus_simps@hcomplex_add_ac)) + THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@hcomplex_mult_minus_simps)) + THEN ALLGOALS + (simp_tac (HOL_ss addsimps hcomplex_minus_from_mult_simps@ + hcomplex_add_ac@hcomplex_mult_ac)) + val numeral_simp_tac = ALLGOALS (simp_tac (HOL_ss addsimps add_0s@bin_simps)) + val simplify_meta_eq = simplify_meta_eq + end; + + +structure EqCancelNumerals = CancelNumeralsFun + (open CancelNumeralsCommon + val prove_conv = Bin_Simprocs.prove_conv + val mk_bal = HOLogic.mk_eq + val dest_bal = HOLogic.dest_bin "op =" hcomplexT + val bal_add1 = hcomplex_eq_add_iff1 RS trans + val bal_add2 = hcomplex_eq_add_iff2 RS trans +); + + +val cancel_numerals = + map prep_simproc + [("hcomplexeq_cancel_numerals", + ["(l::hcomplex) + m = n", "(l::hcomplex) = m + n", + "(l::hcomplex) - m = n", "(l::hcomplex) = m - n", + "(l::hcomplex) * m = n", "(l::hcomplex) = m * n"], + EqCancelNumerals.proc)]; + +structure CombineNumeralsData = + struct + val add = op + : int*int -> int + val mk_sum = long_mk_sum (*to work for e.g. #2*x + #3*x *) + val dest_sum = dest_sum + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff 1 + val left_distrib = left_hcomplex_add_mult_distrib RS trans + val prove_conv = Bin_Simprocs.prove_conv_nohyps + val trans_tac = Real_Numeral_Simprocs.trans_tac + val norm_tac = + ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@ + hcomplex_minus_simps@hcomplex_add_ac)) + THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@hcomplex_mult_minus_simps)) + THEN ALLGOALS (simp_tac (HOL_ss addsimps hcomplex_minus_from_mult_simps@ + hcomplex_add_ac@hcomplex_mult_ac)) + val numeral_simp_tac = ALLGOALS + (simp_tac (HOL_ss addsimps add_0s@bin_simps)) + val simplify_meta_eq = simplify_meta_eq + end; + +structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); + +val combine_numerals = + prep_simproc ("hcomplex_combine_numerals", + ["(i::hcomplex) + j", "(i::hcomplex) - j"], + CombineNumerals.proc); + +(** Declarations for ExtractCommonTerm **) + +(*this version ALWAYS includes a trailing one*) +fun long_mk_prod [] = one + | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts); + +(*Find first term that matches u*) +fun find_first past u [] = raise TERM("find_first", []) + | find_first past u (t::terms) = + if u aconv t then (rev past @ terms) + else find_first (t::past) u terms + handle TERM _ => find_first (t::past) u terms; + +(*Final simplification: cancel + and * *) +fun cancel_simplify_meta_eq cancel_th th = + Int_Numeral_Simprocs.simplify_meta_eq + [hcomplex_mult_one_left, hcomplex_mult_one_right] + (([th, cancel_th]) MRS trans); + +(*** Making constant folding work for 0 and 1 too ***) + +structure HComplexAbstractNumeralsData = + struct + val dest_eq = HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of + val is_numeral = Bin_Simprocs.is_numeral + val numeral_0_eq_0 = hcomplex_numeral_0_eq_0 + val numeral_1_eq_1 = hcomplex_numeral_1_eq_1 + val prove_conv = Bin_Simprocs.prove_conv_nohyps_novars + fun norm_tac simps = ALLGOALS (simp_tac (HOL_ss addsimps simps)) + val simplify_meta_eq = Bin_Simprocs.simplify_meta_eq + end + +structure HComplexAbstractNumerals = AbstractNumeralsFun (HComplexAbstractNumeralsData) + +(*For addition, we already have rules for the operand 0. + Multiplication is omitted because there are already special rules for + both 0 and 1 as operands. Unary minus is trivial, just have - 1 = -1. + For the others, having three patterns is a compromise between just having + one (many spurious calls) and having nine (just too many!) *) +val eval_numerals = + map prep_simproc + [("hcomplex_add_eval_numerals", + ["(m::hcomplex) + 1", "(m::hcomplex) + number_of v"], + HComplexAbstractNumerals.proc add_hcomplex_number_of), + ("hcomplex_diff_eval_numerals", + ["(m::hcomplex) - 1", "(m::hcomplex) - number_of v"], + HComplexAbstractNumerals.proc diff_hcomplex_number_of), + ("hcomplex_eq_eval_numerals", + ["(m::hcomplex) = 0", "(m::hcomplex) = 1", "(m::hcomplex) = number_of v"], + HComplexAbstractNumerals.proc eq_hcomplex_number_of)] + +end; + +Addsimprocs HComplex_Numeral_Simprocs.eval_numerals; +Addsimprocs HComplex_Numeral_Simprocs.cancel_numerals; +Addsimprocs [HComplex_Numeral_Simprocs.combine_numerals]; + + +(*examples: +print_depth 22; +set timing; +set trace_simp; +fun test s = (Goal s, by (Simp_tac 1)); + +test "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::hcomplex)"; +test " 2*u = (u::hcomplex)"; +test "(i + j + 12 + (k::hcomplex)) - 15 = y"; +test "(i + j + 12 + (k::hcomplex)) - 5 = y"; + +test "( 2*x - (u*v) + y) - v* 3*u = (w::hcomplex)"; +test "( 2*x*u*v + (u*v)* 4 + y) - v*u* 4 = (w::hcomplex)"; +test "( 2*x*u*v + (u*v)* 4 + y) - v*u = (w::hcomplex)"; +test "u*v - (x*u*v + (u*v)* 4 + y) = (w::hcomplex)"; + +test "(i + j + 12 + (k::hcomplex)) = u + 15 + y"; +test "(i + j* 2 + 12 + (k::hcomplex)) = j + 5 + y"; + +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)"; + +test "a + -(b+c) + b = (d::hcomplex)"; +test "a + -(b+c) - b = (d::hcomplex)"; + +(*negative numerals*) +test "(i + j + -2 + (k::hcomplex)) - (u + 5 + y) = zz"; + +test "(i + j + -12 + (k::hcomplex)) - 15 = y"; +test "(i + j + 12 + (k::hcomplex)) - -15 = y"; +test "(i + j + -12 + (k::hcomplex)) - -15 = y"; +*) + +(** Constant folding for hcomplex plus and times **) + +structure HComplex_Times_Assoc_Data : ASSOC_FOLD_DATA = +struct + val ss = HOL_ss + val eq_reflection = eq_reflection + val sg_ref = Sign.self_ref (Theory.sign_of (the_context ())) + val T = HComplex_Numeral_Simprocs.hcomplexT + val plus = Const ("op *", [T,T] ---> T) + val add_ac = hcomplex_mult_ac +end; + +structure HComplex_Times_Assoc = Assoc_Fold (HComplex_Times_Assoc_Data); + +Addsimprocs [HComplex_Times_Assoc.conv]; + +Addsimps [hcomplex_of_complex_zero_iff]; + +(*Simplification of x-y = 0 *) + +AddIffs [hcomplex_eq_iff_diff_eq_0 RS sym]; + +(** extra thms **) + +Goal "(hcnj z = 0) = (z = 0)"; +by (auto_tac (claset(),simpset() addsimps [hcomplex_hcnj_zero_iff])); +qed "hcomplex_hcnj_num_zero_iff"; +Addsimps [hcomplex_hcnj_num_zero_iff]; + +Goal "0 = Abs_hcomplex (hcomplexrel `` {%n. 0})"; +by (simp_tac (simpset() addsimps [hcomplex_zero_def RS meta_eq_to_obj_eq RS sym]) 1); +qed "hcomplex_zero_num"; + +Goal "1 = Abs_hcomplex (hcomplexrel `` {%n. 1})"; +by (simp_tac (simpset() addsimps [hcomplex_one_def RS meta_eq_to_obj_eq RS sym]) 1); +qed "hcomplex_one_num"; + +(*** Real and imaginary stuff ***) + +Goalw [hcomplex_number_of_def] + "((number_of xa :: hcomplex) + iii * number_of ya = \ +\ number_of xb + iii * number_of yb) = \ +\ (((number_of xa :: hcomplex) = number_of xb) & \ +\ ((number_of ya :: hcomplex) = number_of yb))"; +by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff, + hcomplex_hypreal_number_of])); +qed "hcomplex_number_of_eq_cancel_iff"; +Addsimps [hcomplex_number_of_eq_cancel_iff]; + +Goalw [hcomplex_number_of_def] + "((number_of xa :: hcomplex) + number_of ya * iii = \ +\ number_of xb + number_of yb * iii) = \ +\ (((number_of xa :: hcomplex) = number_of xb) & \ +\ ((number_of ya :: hcomplex) = number_of yb))"; +by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffA, + hcomplex_hypreal_number_of])); +qed "hcomplex_number_of_eq_cancel_iffA"; +Addsimps [hcomplex_number_of_eq_cancel_iffA]; + +Goalw [hcomplex_number_of_def] + "((number_of xa :: hcomplex) + number_of ya * iii = \ +\ number_of xb + iii * number_of yb) = \ +\ (((number_of xa :: hcomplex) = number_of xb) & \ +\ ((number_of ya :: hcomplex) = number_of yb))"; +by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffB, + hcomplex_hypreal_number_of])); +qed "hcomplex_number_of_eq_cancel_iffB"; +Addsimps [hcomplex_number_of_eq_cancel_iffB]; + +Goalw [hcomplex_number_of_def] + "((number_of xa :: hcomplex) + iii * number_of ya = \ +\ number_of xb + number_of yb * iii) = \ +\ (((number_of xa :: hcomplex) = number_of xb) & \ +\ ((number_of ya :: hcomplex) = number_of yb))"; +by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffC, + hcomplex_hypreal_number_of])); +qed "hcomplex_number_of_eq_cancel_iffC"; +Addsimps [hcomplex_number_of_eq_cancel_iffC]; + +Goalw [hcomplex_number_of_def] + "((number_of xa :: hcomplex) + iii * number_of ya = \ +\ number_of xb) = \ +\ (((number_of xa :: hcomplex) = number_of xb) & \ +\ ((number_of ya :: hcomplex) = 0))"; +by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff2, + hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff])); +qed "hcomplex_number_of_eq_cancel_iff2"; +Addsimps [hcomplex_number_of_eq_cancel_iff2]; + +Goalw [hcomplex_number_of_def] + "((number_of xa :: hcomplex) + number_of ya * iii = \ +\ number_of xb) = \ +\ (((number_of xa :: hcomplex) = number_of xb) & \ +\ ((number_of ya :: hcomplex) = 0))"; +by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff2a, + hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff])); +qed "hcomplex_number_of_eq_cancel_iff2a"; +Addsimps [hcomplex_number_of_eq_cancel_iff2a]; + +Goalw [hcomplex_number_of_def] + "((number_of xa :: hcomplex) + iii * number_of ya = \ +\ iii * number_of yb) = \ +\ (((number_of xa :: hcomplex) = 0) & \ +\ ((number_of ya :: hcomplex) = number_of yb))"; +by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff3, + hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff])); +qed "hcomplex_number_of_eq_cancel_iff3"; +Addsimps [hcomplex_number_of_eq_cancel_iff3]; + +Goalw [hcomplex_number_of_def] + "((number_of xa :: hcomplex) + number_of ya * iii= \ +\ iii * number_of yb) = \ +\ (((number_of xa :: hcomplex) = 0) & \ +\ ((number_of ya :: hcomplex) = number_of yb))"; +by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff3a, + hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff])); +qed "hcomplex_number_of_eq_cancel_iff3a"; +Addsimps [hcomplex_number_of_eq_cancel_iff3a]; + +Goalw [hcomplex_number_of_def] "hcnj (number_of v :: hcomplex) = number_of v"; +by (rtac (hcomplex_hypreal_number_of RS ssubst) 1); +by (rtac hcomplex_hcnj_hcomplex_of_hypreal 1); +qed "hcomplex_number_of_hcnj"; +Addsimps [hcomplex_number_of_hcnj]; + +Goalw [hcomplex_number_of_def] + "hcmod(number_of v :: hcomplex) = abs (number_of v :: hypreal)"; +by (rtac (hcomplex_hypreal_number_of RS ssubst) 1); +by (auto_tac (claset(), HOL_ss addsimps [hcmod_hcomplex_of_hypreal])); +qed "hcomplex_number_of_hcmod"; +Addsimps [hcomplex_number_of_hcmod]; + +Goalw [hcomplex_number_of_def] + "hRe(number_of v :: hcomplex) = number_of v"; +by (rtac (hcomplex_hypreal_number_of RS ssubst) 1); +by (auto_tac (claset(), HOL_ss addsimps [hRe_hcomplex_of_hypreal])); +qed "hcomplex_number_of_hRe"; +Addsimps [hcomplex_number_of_hRe]; + +Goalw [hcomplex_number_of_def] + "hIm(number_of v :: hcomplex) = 0"; +by (rtac (hcomplex_hypreal_number_of RS ssubst) 1); +by (auto_tac (claset(), HOL_ss addsimps [hIm_hcomplex_of_hypreal])); +qed "hcomplex_number_of_hIm"; +Addsimps [hcomplex_number_of_hIm]; + + + diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/NSComplexBin.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/NSComplexBin.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,19 @@ +(* Title: NSComplexBin.thy + Author: Jacques D. Fleuriot + Copyright: 2001 University of Edinburgh + Descrition: Binary arithmetic for the nonstandard complex numbers + This case is reduced to that for the complexes (hence reals). +*) + +NSComplexBin = NSComplex + + + +instance + hcomplex :: number + +defs + hcomplex_number_of_def + "number_of v == hcomplex_of_complex (number_of v)" + (*::bin=>complex ::bin=>complex*) + +end \ No newline at end of file diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/NSInduct.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/NSInduct.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,138 @@ +(* Title: NSInduct.ML + Author: Jacques D. Fleuriot + Copyright: 2001 University of Edinburgh + Description: Nonstandard characterization of induction etc. +*) + +Goalw [starPNat_def] +"(( *pNat* P) (Abs_hypnat(hypnatrel``{%n. X n}))) = \ +\ ({n. P (X n)} : FreeUltrafilterNat)"; +by (Auto_tac); +by (Ultra_tac 1); +qed "starPNat"; + +Goalw [hypnat_of_nat_def] "( *pNat* P) (hypnat_of_nat n) = P n"; +by (auto_tac (claset(),simpset() addsimps [starPNat])); +qed "starPNat_hypnat_of_nat"; +Addsimps [starPNat_hypnat_of_nat]; + +Goalw [hypnat_zero_def,hypnat_one_def] + "!!P. (( *pNat* P) 0 & \ +\ (ALL n. ( *pNat* P)(n) --> ( *pNat* P)(n + 1))) \ +\ --> ( *pNat* P)(n)"; +by (res_inst_tac [("z","n")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [starPNat])); +by (Ultra_tac 1); +by (etac nat_induct 1); +by (dres_inst_tac [("x","hypnat_of_nat n")] spec 1); +by (rtac ccontr 1); +by (auto_tac (claset(),simpset() addsimps [starPNat, + hypnat_of_nat_def,hypnat_add])); +qed "hypnat_induct_obj"; + +Goal + "!!P. [| ( *pNat* P) 0; \ +\ !!n. ( *pNat* P)(n) ==> ( *pNat* P)(n + 1)|] \ +\ ==> ( *pNat* P)(n)"; +by (cut_inst_tac [("P","P"),("n","n")] hypnat_induct_obj 1); +by (Auto_tac); +qed "hypnat_induct"; + +fun hypnat_ind_tac a i = + res_inst_tac [("n",a)] hypnat_induct i THEN rename_last_tac a [""] (i+1); + +Goalw [starPNat2_def] +"(( *pNat2* P) (Abs_hypnat(hypnatrel``{%n. X n})) \ +\ (Abs_hypnat(hypnatrel``{%n. Y n}))) = \ +\ ({n. P (X n) (Y n)} : FreeUltrafilterNat)"; +by (Auto_tac); +by (Ultra_tac 1); +qed "starPNat2"; + +Goalw [starPNat2_def] "( *pNat2* (op =)) = (op =)"; +by (rtac ext 1); +by (rtac ext 1); +by (res_inst_tac [("z","x")] eq_Abs_hypnat 1); +by (res_inst_tac [("z","y")] eq_Abs_hypnat 1); +by (Auto_tac THEN Ultra_tac 1); +qed "starPNat2_eq_iff"; + +Goal "( *pNat2* (%x y. x = y)) X Y = (X = Y)"; +by (simp_tac (simpset() addsimps [starPNat2_eq_iff]) 1); +qed "starPNat2_eq_iff2"; + +Goal "(EX h. P(h::hypnat)) = (EX x. P(Abs_hypnat(hypnatrel `` {x})))"; +by (Auto_tac); +by (res_inst_tac [("z","h")] eq_Abs_hypnat 1); +by (Auto_tac); +val lemma_hyp = result(); + +Goal "(EX (v::nat=>nat). ALL (x::nat). P (g v) (f x)) = \ +\ (EX (v::nat=>nat). ALL x. EX f': Rep_hypnat(f x). P (g v) (Abs_hypnat(hypnatrel `` {f'})))"; +by (Auto_tac); +by (ALLGOALS(res_inst_tac [("x","v")] exI)); +by (Step_tac 1); +by (ALLGOALS(res_inst_tac [("z","f x")] eq_Abs_hypnat)); +by (Auto_tac); +by (rtac bexI 1); +by (dres_inst_tac [("x","x")] spec 3); +by (dtac sym 1); +by (Auto_tac); +by (subgoal_tac + "Abs_hypnat (hypnatrel `` {X}) = Abs_hypnat (hypnatrel `` {Y})" 1); +by (Asm_simp_tac 1); +by (Auto_tac); +val lemma_hyp2 = result(); + +Goalw [hSuc_def] "hSuc m ~= 0"; +by Auto_tac; +qed "hSuc_not_zero"; + +bind_thm ("zero_not_hSuc", hSuc_not_zero RS not_sym); + +Goalw [hSuc_def,hypnat_one_def] + "(hSuc m = hSuc n) = (m = n)"; +by (res_inst_tac [("z","m")] eq_Abs_hypnat 1); +by (res_inst_tac [("z","n")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hypnat_add])); +qed "hSuc_hSuc_eq"; + +AddIffs [hSuc_not_zero,zero_not_hSuc,hSuc_hSuc_eq]; + +Goalw [hypnat_le_def] "m <= (n::hypnat) | n <= m"; +by (auto_tac (claset() addDs [hypnat_less_trans],simpset())); +qed "hypnat_le_linear"; + +val hypnat_less_le = hypnat_less_imp_le; + +Goal "c : (S :: nat set) ==> (LEAST n. n : S) : S"; +by (etac LeastI 1); +qed "nonempty_nat_set_Least_mem"; + +Goalw [InternalNatSets_def,starsetNat_n_def] + "[| S : InternalNatSets; S ~= {} |] ==> EX n: S. ALL m: S. n <= m"; +by (auto_tac (claset(),simpset() addsimps [lemma_hyp])); +by (res_inst_tac [("z","x")] eq_Abs_hypnat 1); +by (auto_tac (claset() addSDs [bspec],simpset() addsimps [hypnat_le])); +by (dres_inst_tac [("x","%m. LEAST n. n : As m")] spec 1); +by Auto_tac; +by (ultra_tac (claset() addDs [nonempty_nat_set_Least_mem],simpset()) 1); +by (dres_inst_tac [("x","x")] bspec 1 THEN Auto_tac); +by (ultra_tac (claset() addIs [Least_le],simpset()) 1); +qed "nonempty_InternalNatSet_has_least"; + +(* Goldblatt p 129 Thm 11.3.2 *) +Goal "[| X : InternalNatSets; 0 : X; ALL n. n : X --> n + 1 : X |] \ +\ ==> X = (UNIV:: hypnat set)"; +by (rtac ccontr 1); +by (ftac InternalNatSets_Compl 1); +by (dres_inst_tac [("S","- X")] nonempty_InternalNatSet_has_least 1); +by Auto_tac; +by (subgoal_tac "1 <= n" 1); +by (dres_inst_tac [("x","n - 1")] bspec 1); +by (Step_tac 1); +by (dres_inst_tac [("x","n - 1")] spec 1); +by (dres_inst_tac [("x","1"),("q1.0","n")] hypnat_add_le_mono1 2); +by (auto_tac (claset(),simpset() addsimps [hypnat_le_def])); +qed "internal_induct"; + diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/NSInduct.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/NSInduct.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,23 @@ +(* Title: NSInduct.thy + Author: Jacques D. Fleuriot + Copyright: 2001 University of Edinburgh + Description: Nonstandard characterization of induction etc. +*) + +NSInduct = ComplexArith0 + + +constdefs + + starPNat :: (nat => bool) => hypnat => bool ("*pNat* _" [80] 80) + "*pNat* P == (%x. EX X. (X: Rep_hypnat(x) & + {n. P(X n)} : FreeUltrafilterNat))" + + + starPNat2 :: (nat => nat => bool) => hypnat =>hypnat => bool ("*pNat2* _" [80] 80) + "*pNat2* P == (%x y. EX X. EX Y. (X: Rep_hypnat(x) & Y: Rep_hypnat(y) & + {n. P(X n) (Y n)} : FreeUltrafilterNat))" + + hSuc :: hypnat => hypnat + "hSuc n == n + 1" + +end \ No newline at end of file diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/ROOT.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,9 @@ +(* Title: HOL/Complex/ROOT.ML + ID: $Id$ + Author: Jacques Fleuriot + +The Complex Numbers +*) + +with_path "../Hyperreal" use_thy "Hyperreal"; +time_use_thy "CLim"; diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/ex/NSPrimes.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/ex/NSPrimes.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,492 @@ +(* Title : NSPrimes.ML + Author : Jacques D. Fleuriot + Copyright : 2002 University of Edinburgh + Description : The nonstandard primes as an extension of + the prime numbers +*) + +(*--------------------------------------------------------------*) +(* A "choice" theorem for ultrafilters cf. almost everywhere *) +(* quantification *) +(*--------------------------------------------------------------*) + +Goal "{n. EX m. Q n m} : FreeUltrafilterNat \ +\ ==> EX f. {n. Q n (f n)} : FreeUltrafilterNat"; +by (res_inst_tac [("x","%n. (@x. Q n x)")] exI 1); +by (Ultra_tac 1 THEN rtac someI 1 THEN Auto_tac); +qed "UF_choice"; + +Goal "({n. P n} : FreeUltrafilterNat --> {n. Q n} : FreeUltrafilterNat) = \ +\ ({n. P n --> Q n} : FreeUltrafilterNat)"; +by Auto_tac; +by (ALLGOALS(Ultra_tac)); +qed "UF_if"; + +Goal "({n. P n} : FreeUltrafilterNat & {n. Q n} : FreeUltrafilterNat) = \ +\ ({n. P n & Q n} : FreeUltrafilterNat)"; +by Auto_tac; +by (ALLGOALS(Ultra_tac)); +qed "UF_conj"; + +Goal "(ALL f. {n. Q n (f n)} : FreeUltrafilterNat) = \ +\ ({n. ALL m. Q n m} : FreeUltrafilterNat)"; +by Auto_tac; +by (Ultra_tac 2); +by (rtac ccontr 1); +by (rtac swap 1 THEN assume_tac 2); +by (simp_tac (simpset() addsimps [FreeUltrafilterNat_Compl_iff1, + CLAIM "-{n. Q n} = {n. ~ Q n}"]) 1); +by (rtac UF_choice 1); +by (Ultra_tac 1 THEN Auto_tac); +qed "UF_choice_ccontr"; + +Goal "ALL M. EX N. 0 < N & (ALL m. 0 < m & (m::nat) <= M --> m dvd N)"; +by (rtac allI 1); +by (induct_tac "M" 1); +by Auto_tac; +by (res_inst_tac [("x","N * (Suc n)")] exI 1); +by (Step_tac 1 THEN Force_tac 1); +by (dtac le_imp_less_or_eq 1 THEN etac disjE 1); +by (force_tac (claset() addSIs [dvd_mult2],simpset()) 1); +by (force_tac (claset() addSIs [dvd_mult],simpset()) 1); +qed "dvd_by_all"; + +bind_thm ("dvd_by_all2",dvd_by_all RS spec); + +Goal "(EX (x::hypnat). P x) = (EX f. P (Abs_hypnat(hypnatrel `` {f})))"; +by Auto_tac; +by (res_inst_tac [("z","x")] eq_Abs_hypnat 1); +by Auto_tac; +qed "lemma_hypnat_P_EX"; + +Goal "(ALL (x::hypnat). P x) = (ALL f. P (Abs_hypnat(hypnatrel `` {f})))"; +by Auto_tac; +by (res_inst_tac [("z","x")] eq_Abs_hypnat 1); +by Auto_tac; +qed "lemma_hypnat_P_ALL"; + +Goalw [hdvd_def] + "(Abs_hypnat(hypnatrel``{%n. X n}) hdvd \ +\ Abs_hypnat(hypnatrel``{%n. Y n})) = \ +\ ({n. X n dvd Y n} : FreeUltrafilterNat)"; +by (Auto_tac THEN Ultra_tac 1); +qed "hdvd"; + +Goal "(hypnat_of_nat n <= 0) = (n = 0)"; +by (stac (hypnat_of_nat_zero RS sym) 1); +by (auto_tac (claset(),simpset() addsimps [hypnat_of_nat_le_iff RS sym])); +qed "hypnat_of_nat_le_zero_iff"; +Addsimps [hypnat_of_nat_le_zero_iff]; + + +(* Goldblatt: Exercise 5.11(2) - p. 57 *) +Goal "ALL M. EX N. 0 < N & (ALL m. 0 < m & (m::hypnat) <= M --> m hdvd N)"; +by (Step_tac 1); +by (res_inst_tac [("z","M")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [lemma_hypnat_P_EX, + lemma_hypnat_P_ALL,hypnat_zero_def,hypnat_le,hypnat_less,hdvd])); +by (cut_facts_tac [dvd_by_all] 1); +by (dtac (CLAIM "(ALL (M::nat). EX N. 0 < N & (ALL m. 0 < m & m <= M \ +\ --> m dvd N)) \ +\ ==> ALL (n::nat). EX N. 0 < N & (ALL m. 0 < (m::nat) & m <= (x n) \ +\ --> m dvd N)") 1); +by (dtac choice 1); +by (Step_tac 1); +by (res_inst_tac [("x","f")] exI 1); +by Auto_tac; +by (ALLGOALS(Ultra_tac)); +by Auto_tac; +qed "hdvd_by_all"; + +bind_thm ("hdvd_by_all2",hdvd_by_all RS spec); + +(* Goldblatt: Exercise 5.11(2) - p. 57 *) +Goal "EX (N::hypnat). 0 < N & (ALL n : - {0::nat}. hypnat_of_nat(n) hdvd N)"; +by (cut_facts_tac [hdvd_by_all] 1); +by (dres_inst_tac [("x","whn")] spec 1); +by Auto_tac; +by (rtac exI 1 THEN Auto_tac); +by (dres_inst_tac [("x","hypnat_of_nat n")] spec 1); +by (auto_tac (claset(),simpset() addsimps [symmetric hypnat_le_def, + hypnat_of_nat_zero_iff])); +qed "hypnat_dvd_all_hypnat_of_nat"; + + +(*--------------------------------------------------------------*) +(* the nonstandard extension of the set prime numbers consists *) +(* of precisely those hypernaturals > 1 that have no nontrivial *) +(* factors *) +(*--------------------------------------------------------------*) + +(* Goldblatt: Exercise 5.11(3a) - p 57 *) +Goalw [starprime_def,thm "prime_def"] + "starprime = {p. 1 < p & (ALL m. m hdvd p --> m = 1 | m = p)}"; +by (auto_tac (claset(),simpset() addsimps + [CLAIM "{p. Q(p) & R(p)} = {p. Q(p)} Int {p. R(p)}",NatStar_Int])); +by (ALLGOALS(res_inst_tac [("z","x")] eq_Abs_hypnat)); +by (res_inst_tac [("z","m")] eq_Abs_hypnat 2); +by (auto_tac (claset(),simpset() addsimps [hdvd,hypnat_one_def,hypnat_less, + lemma_hypnat_P_ALL,starsetNat_def])); +by (dtac bspec 1 THEN dtac bspec 2 THEN Auto_tac); +by (Ultra_tac 1 THEN Ultra_tac 1); +by (rtac ccontr 1); +by (dtac FreeUltrafilterNat_Compl_mem 1); +by (auto_tac (claset(),simpset() addsimps [CLAIM "- {p. Q p} = {p. ~ Q p}"])); +by (dtac UF_choice 1 THEN Auto_tac); +by (dres_inst_tac [("x","f")] spec 1 THEN Auto_tac); +by (ALLGOALS(Ultra_tac)); +qed "starprime"; + +Goalw [thm "prime_def"] "2 : prime"; +by Auto_tac; +by (ftac dvd_imp_le 1); +by (auto_tac (claset() addDs [dvd_0_left],simpset() addsimps + [ARITH_PROVE "(m <= (2::nat)) = (m = 0 | m = 1 | m = 2)"])); +qed "prime_two"; +Addsimps [prime_two]; + +(* proof uses course-of-value induction *) +Goal "Suc 0 < n --> (EX k : prime. k dvd n)"; +by (res_inst_tac [("n","n")] nat_less_induct 1); +by Auto_tac; +by (case_tac "n : prime" 1); +by (res_inst_tac [("x","n")] bexI 1 THEN Auto_tac); +by (dtac (conjI RS (thm "not_prime_ex_mk")) 1 THEN Auto_tac); +by (dres_inst_tac [("x","m")] spec 1 THEN Auto_tac); +by (res_inst_tac [("x","ka")] bexI 1); +by (auto_tac (claset() addIs [dvd_mult2],simpset())); +qed_spec_mp "prime_factor_exists"; + +(* Goldblatt Exercise 5.11(3b) - p 57 *) +Goal "1 < n ==> (EX k : starprime. k hdvd n)"; +by (res_inst_tac [("z","n")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hypnat_one_def, + hypnat_less,starprime_def,lemma_hypnat_P_EX,lemma_hypnat_P_ALL, + hdvd,starsetNat_def,CLAIM "(ALL X: S. P X) = (ALL X. X : S --> P X)", + UF_if])); +by (res_inst_tac [("x","%n. @y. y : prime & y dvd x n")] exI 1); +by Auto_tac; +by (ALLGOALS (Ultra_tac)); +by (dtac sym 1 THEN Asm_simp_tac 1); +by (ALLGOALS(rtac someI2_ex)); +by (auto_tac (claset() addSDs [prime_factor_exists],simpset())); +qed_spec_mp "hyperprime_factor_exists"; + +(* behaves as expected! *) +Goal "( *sNat* insert x A) = insert (hypnat_of_nat x) ( *sNat* A)"; +by (auto_tac (claset(),simpset() addsimps [starsetNat_def, + hypnat_of_nat_def])); +by (ALLGOALS(res_inst_tac [("z","xa")] eq_Abs_hypnat)); +by Auto_tac; +by (TRYALL(dtac bspec)); +by Auto_tac; +by (ALLGOALS(Ultra_tac)); +qed "NatStar_insert"; + +(* Goldblatt Exercise 3.10(1) - p. 29 *) +Goal "finite A ==> *sNat* A = hypnat_of_nat ` A"; +by (res_inst_tac [("P","%x. *sNat* x = hypnat_of_nat ` x")] finite_induct 1); +by (auto_tac (claset(),simpset() addsimps [NatStar_insert])); +qed "NatStar_hypnat_of_nat"; + +(* proved elsewhere? *) +Goal "{x} ~: FreeUltrafilterNat"; +by (auto_tac (claset() addSIs [FreeUltrafilterNat_finite],simpset())); +qed "FreeUltrafilterNat_singleton_not_mem"; +Addsimps [FreeUltrafilterNat_singleton_not_mem]; + +(*-------------------------------------------------------------------------------*) +(* Another characterization of infinite set of natural numbers *) +(*-------------------------------------------------------------------------------*) + +Goal "finite N ==> EX n. (ALL i:N. i<(n::nat))"; +by (eres_inst_tac [("F","N")] finite_induct 1); +by Auto_tac; +by (res_inst_tac [("x","Suc n + x")] exI 1); +by Auto_tac; +qed "finite_nat_set_bounded"; + +Goal "finite N = (EX n. (ALL i:N. i<(n::nat)))"; +by (blast_tac (claset() addIs [finite_nat_set_bounded, + bounded_nat_set_is_finite]) 1); +qed "finite_nat_set_bounded_iff"; + +Goal "(~ finite N) = (ALL n. EX i:N. n <= (i::nat))"; +by (auto_tac (claset(),simpset() addsimps [finite_nat_set_bounded_iff, + le_def])); +qed "not_finite_nat_set_iff"; + +Goal "(ALL i:N. i<=(n::nat)) ==> finite N"; +by (rtac finite_subset 1); + by (rtac finite_atMost 2); +by Auto_tac; +qed "bounded_nat_set_is_finite2"; + +Goal "finite N ==> EX n. (ALL i:N. i<=(n::nat))"; +by (eres_inst_tac [("F","N")] finite_induct 1); +by Auto_tac; +by (res_inst_tac [("x","n + x")] exI 1); +by Auto_tac; +qed "finite_nat_set_bounded2"; + +Goal "finite N = (EX n. (ALL i:N. i<=(n::nat)))"; +by (blast_tac (claset() addIs [finite_nat_set_bounded2, + bounded_nat_set_is_finite2]) 1); +qed "finite_nat_set_bounded_iff2"; + +Goal "(~ finite N) = (ALL n. EX i:N. n < (i::nat))"; +by (auto_tac (claset(),simpset() addsimps [finite_nat_set_bounded_iff2, + le_def])); +qed "not_finite_nat_set_iff2"; + +(*-------------------------------------------------------------------------------*) +(* An injective function cannot define an embedded natural number *) +(*-------------------------------------------------------------------------------*) + +Goal "ALL m n. m ~= n --> f n ~= f m \ +\ ==> {n. f n = N} = {} | (EX m. {n. f n = N} = {m})"; +by Auto_tac; +by (dres_inst_tac [("x","x")] spec 1 THEN Auto_tac); +by (subgoal_tac "ALL n. (f n = f x) = (x = n)" 1); +by Auto_tac; +qed "lemma_infinite_set_singleton"; + +Goalw [SHNat_def,hypnat_of_nat_def] + "inj f ==> Abs_hypnat(hypnatrel `` {f}) ~: Nats"; +by Auto_tac; +by (subgoal_tac "ALL m n. m ~= n --> f n ~= f m" 1); +by Auto_tac; +by (dtac injD 2); +by (assume_tac 2 THEN Force_tac 2); +by (dres_inst_tac[("N","N")] lemma_infinite_set_singleton 1); +by Auto_tac; +qed "inj_fun_not_hypnat_in_SHNat"; + +Goalw [starsetNat_def] + "range f <= A ==> Abs_hypnat(hypnatrel `` {f}) : *sNat* A"; +by Auto_tac; +by (Ultra_tac 1); +by (dres_inst_tac [("c","f x")] subsetD 1); +by (rtac rangeI 1); +by Auto_tac; +qed "range_subset_mem_starsetNat"; + +(*--------------------------------------------------------------------------------*) +(* Gleason Proposition 11-5.5. pg 149, pg 155 (ex. 3) and pg. 360 *) +(* Let E be a nonvoid ordered set with no maximal elements (note: effectively an *) +(* infinite set if we take E = N (Nats)). Then there exists an order-preserving *) +(* injection from N to E. Of course, (as some doofus will undoubtedly point out! *) +(* :-)) can use notion of least element in proof (i.e. no need for choice) if *) +(* dealing with nats as we have well-ordering property *) +(*--------------------------------------------------------------------------------*) + +Goal "E ~= {} ==> EX x. EX X : (Pow E - {{}}). x: X"; +by Auto_tac; +val lemmaPow3 = result(); + +Goalw [choicefun_def] "E ~= {} ==> choicefun E : E"; +by (rtac (lemmaPow3 RS someI2_ex) 1); +by Auto_tac; +qed "choicefun_mem_set"; +Addsimps [choicefun_mem_set]; + +Goal "[| E ~={}; ALL x. EX y: E. x < y |] ==> injf_max n E : E"; +by (induct_tac "n" 1); +by (Force_tac 1); +by (simp_tac (simpset() addsimps [choicefun_def]) 1); +by (rtac (lemmaPow3 RS someI2_ex) 1); +by Auto_tac; +qed "injf_max_mem_set"; + +Goal "ALL x. EX y: E. x < y ==> injf_max n E < injf_max (Suc n) E"; +by (simp_tac (simpset() addsimps [choicefun_def]) 1); +by (rtac (lemmaPow3 RS someI2_ex) 1); +by Auto_tac; +qed "injf_max_order_preserving"; + +Goal "ALL x. EX y: E. x < y \ +\ ==> ALL n m. m < n --> injf_max m E < injf_max n E"; +by (rtac allI 1); +by (induct_tac "n" 1); +by Auto_tac; +by (simp_tac (simpset() addsimps [choicefun_def]) 1); +by (rtac (lemmaPow3 RS someI2_ex) 1); +by Auto_tac; +by (dtac (CLAIM "m < Suc n ==> m = n | m < n") 1); +by Auto_tac; +by (dres_inst_tac [("x","m")] spec 1); +by Auto_tac; +by (dtac subsetD 1 THEN Auto_tac); +by (dres_inst_tac [("x","injf_max m E")] order_less_trans 1); +by Auto_tac; +qed "injf_max_order_preserving2"; + +Goal "ALL x. EX y: E. x < y ==> inj (%n. injf_max n E)"; +by (rtac injI 1); +by (rtac ccontr 1); +by Auto_tac; +by (dtac injf_max_order_preserving2 1); +by (cut_inst_tac [("m","x"),("n","y")] less_linear 1); +by Auto_tac; +by (auto_tac (claset() addSDs [spec],simpset())); +qed "inj_injf_max"; + +Goal "[| (E::('a::{order} set)) ~= {}; ALL x. EX y: E. x < y |] \ +\ ==> EX f. range f <= E & inj (f::nat => 'a) & (ALL m. f m < f(Suc m))"; +by (res_inst_tac [("x","%n. injf_max n E")] exI 1); +by (Step_tac 1); +by (rtac injf_max_mem_set 1); +by (rtac inj_injf_max 3); +by (rtac injf_max_order_preserving 4); +by Auto_tac; +qed "infinite_set_has_order_preserving_inj"; + +(*-------------------------------------------------------------------------------*) +(* Only need fact that we can have an injective function from N to A for proof *) +(*-------------------------------------------------------------------------------*) + +Goal "~ finite A ==> hypnat_of_nat ` A < ( *sNat* A)"; +by Auto_tac; +by (rtac subsetD 1); +by (rtac NatStar_hypreal_of_real_image_subset 1); +by Auto_tac; +by (subgoal_tac "A ~= {}" 1); +by (Force_tac 2); +by (dtac infinite_set_has_order_preserving_inj 1); +by (etac (not_finite_nat_set_iff2 RS iffD1) 1); +by Auto_tac; +by (dtac inj_fun_not_hypnat_in_SHNat 1); +by (dtac range_subset_mem_starsetNat 1); +by (auto_tac (claset(),simpset() addsimps [SHNat_def])); +qed "hypnat_infinite_has_nonstandard"; + +Goal "*sNat* A = hypnat_of_nat ` A ==> finite A"; +by (rtac ccontr 1); +by (auto_tac (claset() addDs [hypnat_infinite_has_nonstandard], + simpset())); +qed "starsetNat_eq_hypnat_of_nat_image_finite"; + +Goal "( *sNat* A = hypnat_of_nat ` A) = (finite A)"; +by (blast_tac (claset() addSIs [starsetNat_eq_hypnat_of_nat_image_finite, + NatStar_hypnat_of_nat]) 1); +qed "finite_starsetNat_iff"; + +Goal "(~ finite A) = (hypnat_of_nat ` A < *sNat* A)"; +by (rtac iffI 1); +by (blast_tac (claset() addSIs [hypnat_infinite_has_nonstandard]) 1); +by (auto_tac (claset(),simpset() addsimps [finite_starsetNat_iff RS sym])); +qed "hypnat_infinite_has_nonstandard_iff"; + + +(*-------------------------------------------------------------------------------*) +(* Existence of infinitely many primes: a nonstandard proof *) +(*-------------------------------------------------------------------------------*) + +Goal "~ (ALL n:- {0}. hypnat_of_nat n hdvd 1)"; +by Auto_tac; +by (res_inst_tac [("x","2")] bexI 1); +by (auto_tac (claset(),simpset() addsimps [hypnat_of_nat_def, + hypnat_one_def,hdvd,dvd_def])); +val lemma_not_dvd_hypnat_one = result(); +Addsimps [lemma_not_dvd_hypnat_one]; + +Goal "EX n:- {0}. ~ hypnat_of_nat n hdvd 1"; +by (cut_facts_tac [lemma_not_dvd_hypnat_one] 1); +by (auto_tac (claset(),simpset() delsimps [lemma_not_dvd_hypnat_one])); +val lemma_not_dvd_hypnat_one2 = result(); +Addsimps [lemma_not_dvd_hypnat_one2]; + +(* not needed here *) +Goalw [hypnat_zero_def,hypnat_one_def] + "[| 0 < (N::hypnat); N ~= 1 |] ==> 1 < N"; +by (res_inst_tac [("z","N")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hypnat_less])); +by (Ultra_tac 1); +qed "hypnat_gt_zero_gt_one"; + +Goalw [hypnat_zero_def,hypnat_one_def] + "0 < N ==> 1 < (N::hypnat) + 1"; +by (res_inst_tac [("z","N")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hypnat_less,hypnat_add])); +qed "hypnat_add_one_gt_one"; + +Goal "0 ~: prime"; +by (Step_tac 1); +by (dtac (thm "prime_g_zero") 1); +by Auto_tac; +qed "zero_not_prime"; +Addsimps [zero_not_prime]; + +Goalw [starprime_def,starsetNat_def,hypnat_of_nat_def] + "hypnat_of_nat 0 ~: starprime"; +by (auto_tac (claset() addSIs [bexI],simpset())); +qed "hypnat_of_nat_zero_not_prime"; +Addsimps [hypnat_of_nat_zero_not_prime]; + +Goalw [starprime_def,starsetNat_def,hypnat_zero_def] + "0 ~: starprime"; +by (auto_tac (claset() addSIs [bexI],simpset())); +qed "hypnat_zero_not_prime"; +Addsimps [hypnat_zero_not_prime]; + +Goal "1 ~: prime"; +by (Step_tac 1); +by (dtac (thm "prime_g_one") 1); +by Auto_tac; +qed "one_not_prime"; +Addsimps [one_not_prime]; + +Goal "Suc 0 ~: prime"; +by (Step_tac 1); +by (dtac (thm "prime_g_one") 1); +by Auto_tac; +qed "one_not_prime2"; +Addsimps [one_not_prime2]; + +Goalw [starprime_def,starsetNat_def,hypnat_of_nat_def] + "hypnat_of_nat 1 ~: starprime"; +by (auto_tac (claset() addSIs [bexI],simpset())); +qed "hypnat_of_nat_one_not_prime"; +Addsimps [hypnat_of_nat_one_not_prime]; + +Goalw [starprime_def,starsetNat_def,hypnat_one_def] + "1 ~: starprime"; +by (auto_tac (claset() addSIs [bexI],simpset())); +qed "hypnat_one_not_prime"; +Addsimps [hypnat_one_not_prime]; + +Goal "[| k hdvd m; k hdvd n |] ==> k hdvd (m - n)"; +by (res_inst_tac [("z","k")] eq_Abs_hypnat 1); +by (res_inst_tac [("z","m")] eq_Abs_hypnat 1); +by (res_inst_tac [("z","n")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hdvd,hypnat_minus])); +by (Ultra_tac 1); +by (blast_tac (claset() addIs [dvd_diff]) 1); +qed "hdvd_diff"; + +Goalw [dvd_def] "x dvd (1::nat) ==> x = 1"; +by Auto_tac; +qed "dvd_one_eq_one"; + +Goalw [hypnat_one_def] "x hdvd 1 ==> x = 1"; +by (res_inst_tac [("z","x")] eq_Abs_hypnat 1); +by (auto_tac (claset(),simpset() addsimps [hdvd])); +qed "hdvd_one_eq_one"; + +Goal "~ finite prime"; +by (rtac (hypnat_infinite_has_nonstandard_iff RS iffD2) 1); +by (cut_facts_tac [hypnat_dvd_all_hypnat_of_nat] 1); +by (etac exE 1); +by (etac conjE 1); +by (subgoal_tac "1 < N + 1" 1); +by (blast_tac (claset() addIs [hypnat_add_one_gt_one]) 2); +by (dtac hyperprime_factor_exists 1); +by (auto_tac (claset() addIs [NatStar_mem],simpset())); +by (subgoal_tac "k ~: hypnat_of_nat ` prime" 1); +by (force_tac (claset(),simpset() addsimps [starprime_def]) 1); +by (Step_tac 1); +by (dres_inst_tac [("x","x")] bspec 1); +by (rtac ccontr 1 THEN Asm_full_simp_tac 1); +by (dtac hdvd_diff 1 THEN assume_tac 1); +by (auto_tac (claset() addDs [hdvd_one_eq_one],simpset())); +qed "not_finite_prime"; diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/ex/NSPrimes.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/ex/NSPrimes.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,36 @@ +(* Title : NSPrimes.thy + Author : Jacques D. Fleuriot + Copyright : 2002 University of Edinburgh + Description : The nonstandard primes as an extension of + the prime numbers + +These can be used to derive an alternative proof of the infinitude of primes by +considering a property of nonstandard sets. +*) + +NSPrimes = Factorization + HSeries + + +consts + hdvd :: [hypnat, hypnat] => bool (infixl 50) + +defs + hdvd_def "(M::hypnat) hdvd N == + EX X Y. X: Rep_hypnat(M) & Y: Rep_hypnat(N) & + {n::nat. X n dvd Y n} : FreeUltrafilterNat" + +constdefs + starprime :: hypnat set + "starprime == ( *sNat* prime)" + +constdefs + choicefun :: 'a set => 'a + "choicefun E == (@x. EX X: Pow(E) -{{}}. x : X)" + +consts injf_max :: "nat => ('a::{order} set) => 'a" +primrec + injf_max_zero "injf_max 0 E = choicefun E" + injf_max_Suc "injf_max (Suc n) E = choicefun({e. e : E & injf_max n E < e})" + +end + + diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/ex/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/ex/ROOT.ML Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,12 @@ +(* Title: HOL/Hyperreal/ex/ROOT.ML + ID: $Id$ + +Miscellaneous HOL-Hyperreal Examples. +*) + +no_document (with_path "../../NumberTheory" use_thy) "Factorization"; +(*These two examples just need the theory Primes.*) +use_thy "Sqrt"; +use_thy "Sqrt_Script"; +(*This example also needs the theory Factorization.*) +use_thy "NSPrimes"; diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/ex/Sqrt.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/ex/Sqrt.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,147 @@ +(* Title: HOL/Hyperreal/ex/Sqrt.thy + ID: $Id$ + Author: Markus Wenzel, TU Muenchen + License: GPL (GNU GENERAL PUBLIC LICENSE) +*) + +header {* Square roots of primes are irrational *} + +theory Sqrt = Primes + Hyperreal: + +subsection {* Preliminaries *} + +text {* + The set of rational numbers, including the key representation + theorem. +*} + +constdefs + rationals :: "real set" ("\") + "\ \ {x. \m n. n \ 0 \ \x\ = real (m::nat) / real (n::nat)}" + +theorem rationals_rep: "x \ \ \ + \m n. n \ 0 \ \x\ = real m / real n \ gcd (m, n) = 1" +proof - + assume "x \ \" + then obtain m n :: nat where + n: "n \ 0" and x_rat: "\x\ = real m / real n" + by (unfold rationals_def) blast + let ?gcd = "gcd (m, n)" + from n have gcd: "?gcd \ 0" by (simp add: gcd_zero) + let ?k = "m div ?gcd" + let ?l = "n div ?gcd" + let ?gcd' = "gcd (?k, ?l)" + have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m" + by (rule dvd_mult_div_cancel) + have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n" + by (rule dvd_mult_div_cancel) + + from n and gcd_l have "?l \ 0" + by (auto iff del: neq0_conv) + moreover + have "\x\ = real ?k / real ?l" + proof - + from gcd have "real ?k / real ?l = + real (?gcd * ?k) / real (?gcd * ?l)" + by (simp add: real_mult_div_cancel1) + also from gcd_k and gcd_l have "\ = real m / real n" by simp + also from x_rat have "\ = \x\" .. + finally show ?thesis .. + qed + moreover + have "?gcd' = 1" + proof - + have "?gcd * ?gcd' = gcd (?gcd * ?k, ?gcd * ?l)" + by (rule gcd_mult_distrib2) + with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp + with gcd show ?thesis by simp + qed + ultimately show ?thesis by blast +qed + +lemma [elim?]: "r \ \ \ + (\m n. n \ 0 \ \r\ = real m / real n \ gcd (m, n) = 1 \ C) + \ C" + using rationals_rep by blast + + +subsection {* Main theorem *} + +text {* + The square root of any prime number (including @{text 2}) is + irrational. +*} + +theorem sqrt_prime_irrational: "p \ prime \ sqrt (real p) \ \" +proof + assume p_prime: "p \ prime" + then have p: "1 < p" by (simp add: prime_def) + assume "sqrt (real p) \ \" + then obtain m n where + n: "n \ 0" and sqrt_rat: "\sqrt (real p)\ = real m / real n" + and gcd: "gcd (m, n) = 1" .. + have eq: "m\ = p * n\" + proof - + from n and sqrt_rat have "real m = \sqrt (real p)\ * real n" by simp + then have "real (m\) = (sqrt (real p))\ * real (n\)" + by (auto simp add: power_two real_power_two) + also have "(sqrt (real p))\ = real p" by simp + also have "\ * real (n\) = real (p * n\)" by simp + finally show ?thesis .. + qed + have "p dvd m \ p dvd n" + proof + from eq have "p dvd m\" .. + with p_prime show "p dvd m" by (rule prime_dvd_power_two) + then obtain k where "m = p * k" .. + with eq have "p * n\ = p\ * k\" by (auto simp add: power_two mult_ac) + with p have "n\ = p * k\" by (simp add: power_two) + then have "p dvd n\" .. + with p_prime show "p dvd n" by (rule prime_dvd_power_two) + qed + then have "p dvd gcd (m, n)" .. + with gcd have "p dvd 1" by simp + then have "p \ 1" by (simp add: dvd_imp_le) + with p show False by simp +qed + +corollary "sqrt (real (2::nat)) \ \" + by (rule sqrt_prime_irrational) (rule two_is_prime) + + +subsection {* Variations *} + +text {* + Here is an alternative version of the main proof, using mostly + linear forward-reasoning. While this results in less top-down + structure, it is probably closer to proofs seen in mathematics. +*} + +theorem "p \ prime \ sqrt (real p) \ \" +proof + assume p_prime: "p \ prime" + then have p: "1 < p" by (simp add: prime_def) + assume "sqrt (real p) \ \" + then obtain m n where + n: "n \ 0" and sqrt_rat: "\sqrt (real p)\ = real m / real n" + and gcd: "gcd (m, n) = 1" .. + from n and sqrt_rat have "real m = \sqrt (real p)\ * real n" by simp + then have "real (m\) = (sqrt (real p))\ * real (n\)" + by (auto simp add: power_two real_power_two) + also have "(sqrt (real p))\ = real p" by simp + also have "\ * real (n\) = real (p * n\)" by simp + finally have eq: "m\ = p * n\" .. + then have "p dvd m\" .. + with p_prime have dvd_m: "p dvd m" by (rule prime_dvd_power_two) + then obtain k where "m = p * k" .. + with eq have "p * n\ = p\ * k\" by (auto simp add: power_two mult_ac) + with p have "n\ = p * k\" by (simp add: power_two) + then have "p dvd n\" .. + with p_prime have "p dvd n" by (rule prime_dvd_power_two) + with dvd_m have "p dvd gcd (m, n)" by (rule gcd_greatest) + with gcd have "p dvd 1" by simp + then have "p \ 1" by (simp add: dvd_imp_le) + with p show False by simp +qed + +end diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/ex/Sqrt_Script.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/ex/Sqrt_Script.thy Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,78 @@ +(* Title: HOL/Hyperreal/ex/Sqrt_Script.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 2001 University of Cambridge +*) + +header {* Square roots of primes are irrational (script version) *} + +theory Sqrt_Script = Primes + Hyperreal: + +text {* + \medskip Contrast this linear Isabelle/Isar script with Markus + Wenzel's more mathematical version. +*} + +subsection {* Preliminaries *} + +lemma prime_nonzero: "p \ prime \ p \ 0" + by (force simp add: prime_def) + +lemma prime_dvd_other_side: + "n * n = p * (k * k) \ p \ prime \ p dvd n" + apply (subgoal_tac "p dvd n * n", blast dest: prime_dvd_mult) + apply (rule_tac j = "k * k" in dvd_mult_left, simp) + done + +lemma reduction: "p \ prime \ + 0 < k \ k * k = p * (j * j) \ k < p * j \ 0 < j" + apply (rule ccontr) + apply (simp add: linorder_not_less) + apply (erule disjE) + apply (frule mult_le_mono, assumption) + apply auto + apply (force simp add: prime_def) + done + +lemma rearrange: "(j::nat) * (p * j) = k * k \ k * k = p * (j * j)" + by (simp add: mult_ac) + +lemma prime_not_square: + "p \ prime \ (\k. 0 < k \ m * m \ p * (k * k))" + apply (induct m rule: nat_less_induct) + apply clarify + apply (frule prime_dvd_other_side, assumption) + apply (erule dvdE) + apply (simp add: nat_mult_eq_cancel_disj prime_nonzero) + apply (blast dest: rearrange reduction) + done + + +subsection {* The set of rational numbers *} + +constdefs + rationals :: "real set" ("\") + "\ \ {x. \m n. n \ 0 \ \x\ = real (m::nat) / real (n::nat)}" + + +subsection {* Main theorem *} + +text {* + The square root of any prime number (including @{text 2}) is + irrational. +*} + +theorem prime_sqrt_irrational: + "p \ prime \ x * x = real p \ 0 \ x \ x \ \" + apply (simp add: rationals_def real_abs_def) + apply clarify + apply (erule_tac P = "real m / real n * ?x = ?y" in rev_mp) + apply (simp del: real_of_nat_mult + add: real_divide_eq_eq prime_not_square + real_of_nat_mult [symmetric]) + done + +lemmas two_sqrt_irrational = + prime_sqrt_irrational [OF two_is_prime] + +end diff -r 8fe7e12290e1 -r 10dbf16be15f src/HOL/Complex/ex/document/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complex/ex/document/root.tex Mon May 05 18:22:31 2003 +0200 @@ -0,0 +1,20 @@ + +\documentclass[11pt,a4paper]{article} +\usepackage[latin1]{inputenc} +\usepackage{isabelle,isabellesym} +\usepackage{pdfsetup} + +\urlstyle{rm} +\isabellestyle{it} + +\begin{document} + +\title{Miscellaneous HOL-Hyperreal Examples} +\maketitle + +\tableofcontents + +\parindent 0pt\parskip 0.5ex +\input{session} + +\end{document}