diff -r 5b067c681989 -r b4b11391c676 src/HOL/NSA/CLim.thy --- a/src/HOL/NSA/CLim.thy Wed Dec 30 11:32:56 2015 +0100 +++ b/src/HOL/NSA/CLim.thy Wed Dec 30 11:37:29 2015 +0100 @@ -4,7 +4,7 @@ Conversion to Isar and new proofs by Lawrence C Paulson, 2004 *) -section{*Limits, Continuity and Differentiation for Complex Functions*} +section\Limits, Continuity and Differentiation for Complex Functions\ theory CLim imports CStar @@ -18,7 +18,7 @@ "x \ (0::complex) \ x * (inverse x)\<^sup>2 = inverse x" by (simp add: numeral_2_eq_2) -text{*Changing the quantified variable. Install earlier?*} +text\Changing the quantified variable. Install earlier?\ lemma all_shift: "(\x::'a::comm_ring_1. P x) = (\x. P (x-a))" apply auto apply (drule_tac x="x+a" in spec) @@ -34,7 +34,7 @@ done -subsection{*Limit of Complex to Complex Function*} +subsection\Limit of Complex to Complex Function\ lemma NSLIM_Re: "f \a\\<^sub>N\<^sub>S L ==> (%x. Re(f x)) \a\\<^sub>N\<^sub>S Re(L)" by (simp add: NSLIM_def starfunC_approx_Re_Im_iff @@ -108,13 +108,13 @@ by (simp add: LIM_NSLIM_iff NSLIM_NSCRLIM_Re_Im_iff) -subsection{*Continuity*} +subsection\Continuity\ lemma NSLIM_isContc_iff: "(f \a\\<^sub>N\<^sub>S f a) = ((%h. f(a + h)) \0\\<^sub>N\<^sub>S f a)" by (rule NSLIM_h_iff) -subsection{*Functions from Complex to Reals*} +subsection\Functions from Complex to Reals\ lemma isNSContCR_cmod [simp]: "isNSCont cmod (a)" by (auto intro: approx_hnorm @@ -134,7 +134,7 @@ shows "isCont f a ==> isCont (%x. Im (f x)) a" by (simp add: isCont_def LIM_Im) -subsection{* Differentiation of Natural Number Powers*} +subsection\Differentiation of Natural Number Powers\ lemma CDERIV_pow [simp]: "DERIV (%x. x ^ n) x :> (complex_of_real (real n)) * (x ^ (n - Suc 0))" @@ -145,11 +145,11 @@ apply (auto simp add: ac_simps) done -text{*Nonstandard version*} +text\Nonstandard version\ lemma NSCDERIV_pow: "NSDERIV (%x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - 1))" by (metis CDERIV_pow NSDERIV_DERIV_iff One_nat_def) -text{*Can't relax the premise @{term "x \ 0"}: it isn't continuous at zero*} +text\Can't relax the premise @{term "x \ 0"}: it isn't continuous at zero\ lemma NSCDERIV_inverse: "(x::complex) \ 0 ==> NSDERIV (%x. inverse(x)) x :> (- ((inverse x)\<^sup>2))" unfolding numeral_2_eq_2 @@ -161,7 +161,7 @@ by (rule DERIV_inverse) -subsection{*Derivative of Reciprocals (Function @{term inverse})*} +subsection\Derivative of Reciprocals (Function @{term inverse})\ lemma CDERIV_inverse_fun: "[| DERIV f x :> d; f(x) \ (0::complex) |] @@ -176,7 +176,7 @@ by (rule NSDERIV_inverse_fun) -subsection{* Derivative of Quotient*} +subsection\Derivative of Quotient\ lemma CDERIV_quotient: "[| DERIV f x :> d; DERIV g x :> e; g(x) \ (0::complex) |] @@ -191,7 +191,7 @@ by (rule NSDERIV_quotient) -subsection{*Caratheodory Formulation of Derivative at a Point: Standard Proof*} +subsection\Caratheodory Formulation of Derivative at a Point: Standard Proof\ lemma CARAT_CDERIVD: "(\z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l