# HG changeset patch # User blanchet # Date 1461617993 -7200 # Node ID 2cc4e85b46d450dcb8032c37eb1e0f8eee3d4090 # Parent 1836456b7d8279fbf786a475b77a2b680ef6b015# Parent 9478f0dff636cbd3b703cde4ca3be0f2fcccd3b6 merge diff -r 1836456b7d82 -r 2cc4e85b46d4 NEWS --- a/NEWS Mon Apr 25 21:09:02 2016 +0200 +++ b/NEWS Mon Apr 25 22:59:53 2016 +0200 @@ -59,6 +59,11 @@ *** Isar *** +* Command 'define' introduces a local (non-polymorphic) definition, with +optional abstraction over local parameters. The syntax resembles +'definition' and 'obtain'. It fits better into the Isar language than +old 'def', which is now a legacy feature. + * Command '\' is an alias for 'sorry', with different typesetting. E.g. to produce proof holes in examples and documentation. diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Doc/Isar_Ref/Framework.thy --- a/src/Doc/Isar_Ref/Framework.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Doc/Isar_Ref/Framework.thy Mon Apr 25 22:59:53 2016 +0200 @@ -577,8 +577,8 @@ i.e.\ it may only occur internally when derived commands are defined in ML. The default inference for \<^theory_text>\assume\ is @{inference export} as given below. - The derived element \<^theory_text>\def x \ a\ is defined as \<^theory_text>\fix x assume \expand\ x \ - a\, with the subsequent inference @{inference expand}. + The derived element \<^theory_text>\define x where "x \ a"\ is defined as \<^theory_text>\fix x assume + \expand\ x \ a\, with the subsequent inference @{inference expand}. \[ \infer[(@{inference_def export})]{\\\ - A \ A \ B\}{\\\ \ B\} @@ -649,7 +649,7 @@ note \A \ B\ text_raw \\end{minipage}\\[3ex]\begin{minipage}[t]{0.45\textwidth}\(*<*)next(*>*) { - def x \ a + define x where "x \ a" have "B x" \ } note \B a\ diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Doc/Isar_Ref/Proof.thy --- a/src/Doc/Isar_Ref/Proof.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Doc/Isar_Ref/Proof.thy Mon Apr 25 22:59:53 2016 +0200 @@ -127,6 +127,7 @@ @{command_def "fix"} & : & \proof(state) \ proof(state)\ \\ @{command_def "assume"} & : & \proof(state) \ proof(state)\ \\ @{command_def "presume"} & : & \proof(state) \ proof(state)\ \\ + @{command_def "define"} & : & \proof(state) \ proof(state)\ \\ @{command_def "def"} & : & \proof(state) \ proof(state)\ \\ \end{matharray} @@ -151,7 +152,7 @@ "presume"} leaves the subgoal unchanged in order to be proved later by the user. - Local definitions, introduced by ``@{command "def"}~\x \ t\'', are achieved + Local definitions, introduced by ``\<^theory_text>\define x where x = t\'', are achieved by combining ``@{command "fix"}~\x\'' with another version of assumption that causes any hypothetical equation \x \ t\ to be eliminated by the reflexivity rule. Thus, exporting some result \x \ t \ \[x]\ yields \\ @@ -166,6 +167,9 @@ ; prems: (@'if' (@{syntax props'} + @'and'))? ; + @@{command define} (@{syntax "fixes"} + @'and') + @'where' (@{syntax props} + @'and') @{syntax for_fixes} + ; @@{command def} (def + @'and') ; def: @{syntax thmdecl}? \ @@ -189,14 +193,23 @@ to \<^theory_text>\assume "\x. A x \ B x"\, but vacuous quantification is avoided: a for-context only effects propositions according to actual use of variables. - \<^descr> @{command "def"}~\x \ t\ introduces a local (non-polymorphic) definition. - In results exported from the context, \x\ is replaced by \t\. Basically, - ``@{command "def"}~\x \ t\'' abbreviates ``@{command "fix"}~\x\~@{command - "assume"}~\x \ t\'', with the resulting hypothetical equation solved by - reflexivity. + \<^descr> \<^theory_text>\define x where "x = t"\ introduces a local (non-polymorphic) definition. + In results that are exported from the context, \x\ is replaced by \t\. + + Internally, equational assumptions are added to the context in Pure form, + using \x \ t\ instead of \x = t\ or \x \ t\ from the object-logic. When + exporting results from the context, \x\ is generalized and the assumption + discharged by reflexivity, causing the replacement by \t\. - The default name for the definitional equation is \x_def\. Several - simultaneous definitions may be given at the same time. + The default name for the definitional fact is \x_def\. Several simultaneous + definitions may be given as well, with a collective default name. + + \<^medskip> + It is also possible to abstract over local parameters as follows: \<^theory_text>\define f + :: "'a \ 'b" where "f x = t" for x :: 'a\. + + \<^descr> \<^theory_text>\def x \ t\ introduces a local (non-polymorphic) definition. This is an + old form of \<^theory_text>\define x where "x = t"\. \ @@ -227,10 +240,10 @@ \<^medskip> Term abbreviations are quite different from local definitions as introduced - via @{command "def"} (see \secref{sec:proof-context}). The latter are + via @{command "define"} (see \secref{sec:proof-context}). The latter are visible within the logic as actual equations, while abbreviations disappear during the input process just after type checking. Also note that @{command - "def"} does not support polymorphism. + "define"} does not support polymorphism. @{rail \ @@{command let} ((@{syntax term} + @'and') '=' @{syntax term} + @'and') diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Doc/Isar_Ref/Quick_Reference.thy --- a/src/Doc/Isar_Ref/Quick_Reference.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Doc/Isar_Ref/Quick_Reference.thy Mon Apr 25 22:59:53 2016 +0200 @@ -88,7 +88,7 @@ \<^theory_text>\moreover\ & \\\ & \<^theory_text>\note calculation = calculation this\ \\ \<^theory_text>\ultimately\ & \\\ & \<^theory_text>\moreover from calculation\ \\[0.5ex] \<^theory_text>\presume a: A\ & \\\ & \<^theory_text>\assume a: A\ \\ - \<^theory_text>\def "x \ t"\ & \\\ & \<^theory_text>\fix x assume x_def: "x \ t"\ \\ + \<^theory_text>\define x where "x = t"\ & \\\ & \<^theory_text>\fix x assume x_def: "x = t"\ \\ \<^theory_text>\consider x where A | \\ & \\\ & \<^theory_text>\have thesis\ \\ & & \quad \<^theory_text>\if "\x. A \ thesis" and \ for thesis\ \\ \<^theory_text>\obtain x where a: A \\ & \\\ & \<^theory_text>\consider x where A \\ \\ diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Doc/Isar_Ref/Spec.thy --- a/src/Doc/Isar_Ref/Spec.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Doc/Isar_Ref/Spec.thy Mon Apr 25 22:59:53 2016 +0200 @@ -538,11 +538,12 @@ to \<^theory_text>\assume\ within a proof (cf.\ \secref{sec:proof-context}). \<^descr> @{element "defines"}~\a: x \ t\ defines a previously declared parameter. - This is similar to \<^theory_text>\def\ within a proof (cf.\ - \secref{sec:proof-context}), but @{element "defines"} takes an equational - proposition instead of variable-term pair. The left-hand side of the - equation may have additional arguments, e.g.\ ``@{element "defines"}~\f - x\<^sub>1 \ x\<^sub>n \ t\''. + This is similar to \<^theory_text>\define\ within a proof (cf.\ + \secref{sec:proof-context}), but @{element "defines"} is restricted to + Pure equalities and the defined variable needs to be declared beforehand + via @{element "fixes"}. The left-hand side of the equation may have + additional arguments, e.g.\ ``@{element "defines"}~\f x\<^sub>1 \ x\<^sub>n \ t\'', + which need to be free in the context. \<^descr> @{element "notes"}~\a = b\<^sub>1 \ b\<^sub>n\ reconsiders facts within a local context. Most notably, this may include arbitrary declarations in any diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Doc/Isar_Ref/Synopsis.thy --- a/src/Doc/Isar_Ref/Synopsis.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Doc/Isar_Ref/Synopsis.thy Mon Apr 25 22:59:53 2016 +0200 @@ -26,7 +26,7 @@ assume "a = b" \ \type assignment at first occurrence in concrete term\ txt \Definitions (non-polymorphic):\ - def x \ "t::'a" + define x :: 'a where "x = t" txt \Abbreviations (polymorphic):\ let ?f = "\x. x" @@ -635,7 +635,7 @@ next { - def x \ t + define x where "x = t" have "B x" \ } have "B t" by fact diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Algebra/Ideal.thy --- a/src/HOL/Algebra/Ideal.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Algebra/Ideal.thy Mon Apr 25 22:59:53 2016 +0200 @@ -787,7 +787,7 @@ abI: "a \ b \ I" and anI: "a \ I" and bnI: "b \ I" by fast - def J \ "{x\carrier R. a \ x \ I}" + define J where "J = {x\carrier R. a \ x \ I}" from is_cring and acarr have idealJ: "ideal J R" unfolding J_def by (rule helper_max_prime) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Algebra/QuotRing.thy --- a/src/HOL/Algebra/QuotRing.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Algebra/QuotRing.thy Mon Apr 25 22:59:53 2016 +0200 @@ -238,7 +238,7 @@ assume IanI: "I +> a \ I" and acarr: "a \ carrier R" --\Helper ideal @{text "J"}\ - def J \ "(carrier R #> a) <+> I :: 'a set" + define J :: "'a set" where "J = (carrier R #> a) <+> I" have idealJ: "ideal J R" apply (unfold J_def, rule add_ideals) apply (simp only: cgenideal_eq_rcos[symmetric], rule cgenideal_ideal, rule acarr) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/BNF_Cardinal_Order_Relation.thy --- a/src/HOL/BNF_Cardinal_Order_Relation.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/BNF_Cardinal_Order_Relation.thy Mon Apr 25 22:59:53 2016 +0200 @@ -470,10 +470,15 @@ fixes A :: "'a set" and B :: "'b set" and C :: "'c set" shows "|(A <+> B) <+> C| =o |A <+> B <+> C|" proof - - def f \ "\(k::('a + 'b) + 'c). - case k of Inl ab \ (case ab of Inl a \ Inl a - |Inr b \ Inr (Inl b)) - |Inr c \ Inr (Inr c)" + define f :: "('a + 'b) + 'c \ 'a + 'b + 'c" + where [abs_def]: "f k = + (case k of + Inl ab \ + (case ab of + Inl a \ Inl a + | Inr b \ Inr (Inl b)) + | Inr c \ Inr (Inr c))" + for k have "A <+> B <+> C \ f ` ((A <+> B) <+> C)" proof fix x assume x: "x \ A <+> B <+> C" @@ -499,7 +504,7 @@ qed qed hence "bij_betw f ((A <+> B) <+> C) (A <+> B <+> C)" - unfolding bij_betw_def inj_on_def f_def by fastforce + unfolding bij_betw_def inj_on_def f_def by fastforce thus ?thesis using card_of_ordIso by blast qed @@ -1626,8 +1631,8 @@ lemma card_of_Pow_Func: "|Pow A| =o |Func A (UNIV::bool set)|" proof- - def F \ "\ A' a. if a \ A then (if a \ A' then True else False) - else undefined" + define F where [abs_def]: "F A' a = + (if a \ A then (if a \ A' then True else False) else undefined)" for A' a have "bij_betw F (Pow A) (Func A (UNIV::bool set))" unfolding bij_betw_def inj_on_def proof (intro ballI impI conjI) fix A1 A2 assume "A1 \ Pow A" "A2 \ Pow A" "F A1 = F A2" @@ -1638,8 +1643,9 @@ fix f assume f: "f \ Func A (UNIV::bool set)" show "f \ F ` Pow A" unfolding image_def mem_Collect_eq proof(intro bexI) let ?A1 = "{a \ A. f a = True}" - show "f = F ?A1" unfolding F_def apply(rule ext) - using f unfolding Func_def mem_Collect_eq by auto + show "f = F ?A1" + unfolding F_def apply(rule ext) + using f unfolding Func_def mem_Collect_eq by auto qed auto qed(unfold Func_def mem_Collect_eq F_def, auto) qed diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/BNF_Wellorder_Constructions.thy --- a/src/HOL/BNF_Wellorder_Constructions.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/BNF_Wellorder_Constructions.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1616,7 +1616,7 @@ show ?thesis proof safe fix h assume h: "h \ Func B2 B1" - def j1 \ "inv_into A1 f1" + define j1 where "j1 = inv_into A1 f1" have "\ a2 \ f2 ` B2. \ b2. b2 \ B2 \ f2 b2 = a2" by blast then obtain k where k: "\ a2 \ f2 ` B2. k a2 \ B2 \ f2 (k a2) = a2" by atomize_elim (rule bchoice) @@ -1626,11 +1626,11 @@ ultimately have "k (f2 b2) = b2" using b2 A2(1) unfolding inj_on_def by blast } note kk = this obtain b22 where b22: "b22 \ B2" using B2 by auto - def j2 \ "\ a2. if a2 \ f2 ` B2 then k a2 else b22" + define j2 where [abs_def]: "j2 a2 = (if a2 \ f2 ` B2 then k a2 else b22)" for a2 have j2A2: "j2 ` A2 \ B2" unfolding j2_def using k b22 by auto have j2: "\ b2. b2 \ B2 \ j2 (f2 b2) = b2" using kk unfolding j2_def by auto - def g \ "Func_map A2 j1 j2 h" + define g where "g = Func_map A2 j1 j2 h" have "Func_map B2 f1 f2 g = h" proof (rule ext) fix b2 show "Func_map B2 f1 f2 g b2 = h b2" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Binomial.thy --- a/src/HOL/Binomial.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Binomial.thy Mon Apr 25 22:59:53 2016 +0200 @@ -643,7 +643,7 @@ shows "pochhammer z (Suc n) * pochhammer (z + 1/2) (Suc n) = (\k=0..2*n+1. z + of_nat k / 2)" proof (induction n) case (Suc n) - def n' \ "Suc n" + define n' where "n' = Suc n" have "pochhammer z (Suc n') * pochhammer (z + 1 / 2) (Suc n') = (pochhammer z n' * pochhammer (z + 1 / 2) n') * ((z + of_nat n') * (z + 1/2 + of_nat n'))" (is "_ = _ * ?A") @@ -1161,7 +1161,8 @@ (\k\m. (-r gchoose k) * (-x)^k * (x + y)^(m-k))" (is "?lhs m = ?rhs m") proof (induction m) case (Suc mm) - def G \ "\i k. (of_nat i + r gchoose k) * x^k * y^(i-k)" and S \ ?lhs + define G where "G i k = (of_nat i + r gchoose k) * x^k * y^(i-k)" for i k + define S where "S = ?lhs" have SG_def: "S = (\i. (\k\i. (G i k)))" unfolding S_def G_def .. have "S (Suc mm) = G (Suc mm) 0 + (\k=Suc 0..Suc mm. G (Suc mm) k)" @@ -1388,7 +1389,7 @@ proof(rule setsum.cong[OF refl]) fix x assume x: "x \ \A" - def K \ "{X \ A. x \ X}" + define K where "K = {X \ A. x \ X}" with \finite A\ have K: "finite K" by auto let ?I = "\i. {I. I \ A \ card I = i \ x \ \I}" have "inj_on snd (SIGMA i:{1..card A}. ?I i)" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Cardinals/Bounded_Set.thy --- a/src/HOL/Cardinals/Bounded_Set.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Cardinals/Bounded_Set.thy Mon Apr 25 22:59:53 2016 +0200 @@ -50,8 +50,9 @@ BNF_Def.Grp {a. set_bset a \ {(a, b). R a b}} (map_bset snd)) a b" (is "?L \ ?R") proof assume ?L - def R' \ "the_inv set_bset (Collect (case_prod R) \ (set_bset a \ set_bset b)) :: ('a \ 'b) set['k]" - (is "the_inv set_bset ?L'") + define R' :: "('a \ 'b) set['k]" + where "R' = the_inv set_bset (Collect (case_prod R) \ (set_bset a \ set_bset b))" + (is "_ = the_inv set_bset ?L'") have "|?L'| {x}| =o |A|" + fixes A :: "'a set" + shows "|A \ {x}| =o |A|" proof - - def f \ "\(a::'a,b::'b). (a)" + define f :: "'a \ 'b \ 'a" where "f = (\(a, b). a)" have "A \ f ` (A \ {x})" unfolding f_def by (auto simp: image_iff) hence "bij_betw f (A \ {x}) A" unfolding bij_betw_def inj_on_def f_def by fastforce thus ?thesis using card_of_ordIso by blast @@ -223,7 +223,7 @@ fixes x :: 'b and A :: "'a set" shows "|Func A {x}| =o |{x}|" proof (rule ordIso_symmetric) - def f \ "\y a. if y = x \ a \ A then x else undefined" + define f where [abs_def]: "f y a = (if y = x \ a \ A then x else undefined)" for y a have "Func A {x} \ f ` {x}" unfolding f_def Func_def by (force simp: fun_eq_iff) hence "bij_betw f {x} (Func A {x})" unfolding bij_betw_def inj_on_def f_def Func_def by (auto split: if_split_asm) @@ -237,7 +237,7 @@ fixes A :: "'a set" shows "|Func (UNIV :: bool set) A| =o |A \ A|" proof (rule ordIso_symmetric) - def f \ "\(x::'a,y) b. if A = {} then undefined else if b then x else y" + define f where "f = (\(x::'a,y) b. if A = {} then undefined else if b then x else y)" have "Func (UNIV :: bool set) A \ f ` (A \ A)" unfolding f_def Func_def by (auto simp: image_iff fun_eq_iff split: option.splits if_split_asm) blast hence "bij_betw f (A \ A) (Func (UNIV :: bool set) A)" @@ -252,8 +252,8 @@ fixes A :: "'a set" and B :: "'b set" and C :: "'c set" shows "|Func (A <+> B) C| =o |Func A C \ Func B C|" proof (rule ordIso_symmetric) - def f \ "\(g :: 'a => 'c, h::'b \ 'c) ab. case ab of Inl a \ g a | Inr b \ h b" - def f' \ "\(f :: ('a + 'b) \ 'c). (\a. f (Inl a), \b. f (Inr b))" + define f where "f = (\(g :: 'a => 'c, h::'b \ 'c) ab. case ab of Inl a \ g a | Inr b \ h b)" + define f' where "f' = (\(f :: ('a + 'b) \ 'c). (\a. f (Inl a), \b. f (Inr b)))" have "f ` (Func A C \ Func B C) \ Func (A <+> B) C" unfolding Func_def f_def by (force split: sum.splits) moreover have "f' ` Func (A <+> B) C \ Func A C \ Func B C" unfolding Func_def f'_def by force diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Cardinals/Cardinal_Order_Relation.thy --- a/src/HOL/Cardinals/Cardinal_Order_Relation.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Cardinals/Cardinal_Order_Relation.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1431,7 +1431,7 @@ show "?F ` Func A B = Func_option A B" proof safe fix f assume f: "f \ Func_option A B" - def g \ "\ a. case f a of Some b \ b | None \ undefined" + define g where [abs_def]: "g a = (case f a of Some b \ b | None \ undefined)" for a have "g \ Func A B" using f unfolding g_def Func_def Func_option_def by force+ moreover have "f = ?F g" @@ -1475,8 +1475,8 @@ shows "|Func A1 B| \o |Func A2 B|" proof- obtain bb where bb: "bb \ B" using B by auto - def F \ "\ (f1::'a \ 'b) a. if a \ A2 then (if a \ A1 then f1 a else bb) - else undefined" + define F where [abs_def]: "F f1 a = + (if a \ A2 then (if a \ A1 then f1 a else bb) else undefined)" for f1 :: "'a \ 'b" and a show ?thesis unfolding card_of_ordLeq[symmetric] proof(intro exI[of _ F] conjI) show "inj_on F (Func A1 B)" unfolding inj_on_def proof safe fix f g assume f: "f \ Func A1 B" and g: "g \ Func A1 B" and eq: "F f = F g" @@ -1523,7 +1523,7 @@ hence "|X| \o |Field r|" by (metis Field_card_of card_of_mono2) then obtain F where 1: "X = F ` (Field r)" using card_of_ordLeq2[OF X] by metis - def f \ "\ i. if i \ Field r then F i else undefined" + define f where [abs_def]: "f i = (if i \ Field r then F i else undefined)" for i have "\ f \ Func (Field r) A. X = ?F f" apply (intro bexI[of _ f]) using 1 XA unfolding Func_def f_def by auto } @@ -1586,7 +1586,7 @@ hence rr: "Refl r" by (metis wo_rel.REFL) show ?thesis unfolding wo_rel.isLimOrd_def[OF 0] wo_rel.isSuccOrd_def[OF 0] proof safe fix j assume j: "j \ Field r" and jm: "\i\Field r. (i, j) \ r" - def p \ "Restr r (Field r - {j})" + define p where "p = Restr r (Field r - {j})" have fp: "Field p = Field r - {j}" unfolding p_def apply(rule Refl_Field_Restr2[OF rr]) by auto have of: "ofilter r (Field p)" unfolding wo_rel.ofilter_def[OF 0] proof safe @@ -1703,7 +1703,7 @@ ultimately obtain f where f: "f ` Sigma A F = Field r" using card_of_ordLeq2 by metis have r: "wo_rel r" using cr unfolding card_order_on_def wo_rel_def by auto {fix a assume a: "a \ A" - def L \ "{(a,u) | u. u \ F a}" + define L where "L = {(a,u) | u. u \ F a}" have fL: "f ` L \ Field r" using f a unfolding L_def by auto have "|L| =o |F a|" unfolding L_def card_of_ordIso[symmetric] apply(rule exI[of _ snd]) unfolding bij_betw_def inj_on_def by (auto simp: image_def) @@ -1717,7 +1717,7 @@ } then obtain gg where gg: "\ a \ A. \ u \ F a. (f (a,u), gg a) \ r" by metis obtain j0 where j0: "j0 \ Field r" using Fi by auto - def g \ "\ a. if F a \ {} then gg a else j0" + define g where [abs_def]: "g a = (if F a \ {} then gg a else j0)" for a have g: "\ a \ A. \ u \ F a. (f (a,u),g a) \ r" using gg unfolding g_def by auto hence 1: "Field r \ (\a \ A. under r (g a))" using f[symmetric] unfolding under_def image_def by auto diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Cardinals/Ordinal_Arithmetic.thy --- a/src/HOL/Cardinals/Ordinal_Arithmetic.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Cardinals/Ordinal_Arithmetic.thy Mon Apr 25 22:59:53 2016 +0200 @@ -751,12 +751,12 @@ with F have maxF: "\f \ F. s.isMaxim (SUPP f) (s.maxim (SUPP f))" and SUPPF: "\f \ F. finite (SUPP f) \ SUPP f \ {} \ SUPP f \ Field s" using maxim_isMaxim_support support_not_const by auto - def z \ "s.minim {s.maxim (SUPP f) | f. f \ F}" + define z where "z = s.minim {s.maxim (SUPP f) | f. f \ F}" from F SUPPF maxF have zmin: "s.isMinim {s.maxim (SUPP f) | f. f \ F} z" unfolding z_def by (intro s.minim_isMinim) (auto simp: s.isMaxim_def) with F have zy: "(z, y) \ s" unfolding s.isMinim_def by auto hence zField: "z \ Field s" unfolding Field_def by auto - def x0 \ "r.minim {f z | f. f \ F \ z = s.maxim (SUPP f)}" + define x0 where "x0 = r.minim {f z | f. f \ F \ z = s.maxim (SUPP f)}" from F(1,2) maxF(1) SUPPF zmin have x0min: "r.isMinim {f z | f. f \ F \ z = s.maxim (SUPP f)} x0" unfolding x0_def s.isMaxim_def s.isMinim_def @@ -766,7 +766,7 @@ from x0min maxF(1) SUPPF F(1) have x0notzero: "x0 \ r.zero" unfolding r.isMinim_def s.isMaxim_def FinFunc_def Func_def support_def by fastforce - def G \ "{f(z := r.zero) | f. f \ F \ z = s.maxim (SUPP f) \ f z = x0}" + define G where "G = {f(z := r.zero) | f. f \ F \ z = s.maxim (SUPP f) \ f z = x0}" from zmin x0min have "G \ {}" unfolding G_def z_def s.isMinim_def r.isMinim_def by blast have GF: "G \ (\f. f(z := r.zero)) ` F" unfolding G_def by auto have "G \ fin_support r.zero (Field s)" @@ -786,7 +786,7 @@ with G have maxG: "\g \ G. s.isMaxim (SUPP g) (s.maxim (SUPP g))" and SUPPG: "\g \ G. finite (SUPP g) \ SUPP g \ {} \ SUPP g \ Field s" using maxim_isMaxim_support support_not_const by auto - def y' \ "s.minim {s.maxim (SUPP f) | f. f \ G}" + define y' where "y' = s.minim {s.maxim (SUPP f) | f. f \ G}" from G SUPPG maxG `G \ {}` have y'min: "s.isMinim {s.maxim (SUPP f) | f. f \ G} y'" unfolding y'_def by (intro s.minim_isMinim) (auto simp: s.isMaxim_def) moreover @@ -810,7 +810,7 @@ qed simp then obtain g0 where g0: "g0 \ G" "\g \ G. (g0, g) \ oexp" by blast hence g0z: "g0 z = r.zero" unfolding G_def by auto - def f0 \ "g0(z := x0)" + define f0 where "f0 = g0(z := x0)" with x0notzero zField have SUPP: "SUPP f0 = SUPP g0 \ {z}" unfolding support_def by auto from g0z have f0z: "f0(z := r.zero) = g0" unfolding f0_def fun_upd_upd by auto have f0: "f0 \ F" using x0min g0(1) @@ -1297,8 +1297,9 @@ using embed_iff_compat_inj_on_ofilter[OF s t, of f] embedS_iff[OF s, of t f] unfolding embedS_def by auto note invff = the_inv_into_f_f[OF *(1)] and injfD = inj_onD[OF *(1)] - def F \ "\g z. if z \ f ` Field s then g (the_inv_into (Field s) f z) else - if z \ Field t then r.zero else undefined" + define F where [abs_def]: "F g z = + (if z \ f ` Field s then g (the_inv_into (Field s) f z) + else if z \ Field t then r.zero else undefined)" for g z from *(4) x(2) the_inv_into_f_eq[OF *(1)] have FLR: "F ` Field ?L \ Field ?R" unfolding rt.Field_oexp rs.Field_oexp FinFunc_def Func_def fin_support_def support_def F_def by (fastforce split: option.splits if_split_asm elim!: finite_surj[of _ _ f]) @@ -1368,7 +1369,8 @@ from FLR have "F ` Field ?L \ Field ?R" proof (intro psubsetI) from *(4) obtain z where z: "z \ Field t" "z \ f ` Field s" by auto - def h \ "\z'. if z' \ Field t then if z' = z then x else r.zero else undefined" + define h where [abs_def]: "h z' = + (if z' \ Field t then if z' = z then x else r.zero else undefined)" for z' from z x(3) have "rt.SUPP h = {z}" unfolding support_def h_def by simp with x have "h \ Field ?R" unfolding h_def rt.Field_oexp FinFunc_def Func_def fin_support_def by auto diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Cardinals/Wellorder_Constructions.thy --- a/src/HOL/Cardinals/Wellorder_Constructions.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Cardinals/Wellorder_Constructions.thy Mon Apr 25 22:59:53 2016 +0200 @@ -658,7 +658,7 @@ lemma isSucc_pred_in: assumes "isSucc i" shows "(pred i, i) \ r" proof- - def j \ "pred i" + define j where "j = pred i" have i: "i = succ j" using assms unfolding j_def by simp have a: "aboveS j \ {}" unfolding j_def using assms by auto show ?thesis unfolding j_def[symmetric] unfolding i using succ_in[OF a] . @@ -766,7 +766,7 @@ assumes a: "adm_woL L" and i: "aboveS j \ {}" shows "worecSL S L (succ j) = S j (worecSL S L j)" proof- - def i \ "succ j" + define i where "i = succ j" have i: "isSucc i" by (metis i i_def isSucc_succ) have ij: "j = pred i" unfolding i_def using assms by simp have 0: "succ (pred i) = i" using i by simp @@ -935,7 +935,7 @@ interpret r: wo_rel r by unfold_locales (rule r) interpret s: wo_rel s by unfold_locales (rule s) let ?G = "\ g a. suc s (g ` underS r a)" - def g \ "worec r ?G" + define g where "g = worec r ?G" have adm: "adm_wo r ?G" unfolding r.adm_wo_def image_def by auto (* *) {fix a assume "a \ Field r" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Cardinals/Wellorder_Extension.thy --- a/src/HOL/Cardinals/Wellorder_Extension.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Cardinals/Wellorder_Extension.thy Mon Apr 25 22:59:53 2016 +0200 @@ -66,7 +66,7 @@ shows "\w. p \ w \ Well_order w" proof - let ?K = "{r. Well_order r \ downset_on (Field r) p \ extension_on (Field r) r p}" - def I \ "init_seg_of \ ?K \ ?K" + define I where "I = init_seg_of \ ?K \ ?K" have I_init: "I \ init_seg_of" by (simp add: I_def) then have subch: "\R. R \ Chains I \ chain\<^sub>\ R" by (auto simp: init_seg_of_def chain_subset_def Chains_def) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Complex.thy --- a/src/HOL/Complex.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Complex.thy Mon Apr 25 22:59:53 2016 +0200 @@ -817,7 +817,8 @@ from assms have "z \ 0" by auto have "(SOME a. sgn z = cis a \ -pi < a \ a \ pi) = x" proof - fix a def d \ "a - x" + fix a + define d where "d = a - x" assume a: "sgn z = cis a \ - pi < a \ a \ pi" from a assms have "- (2*pi) < d \ d < 2*pi" unfolding d_def by simp @@ -839,7 +840,7 @@ proof (simp add: arg_def assms, rule someI_ex) obtain r a where z: "z = rcis r a" using rcis_Ex by fast with assms have "r \ 0" by auto - def b \ "if 0 < r then a else a + pi" + define b where "b = (if 0 < r then a else a + pi)" have b: "sgn z = cis b" unfolding z b_def rcis_def using \r \ 0\ by (simp add: of_real_def sgn_scaleR sgn_if complex_eq_iff) @@ -848,7 +849,7 @@ have cis_2pi_int: "\x. cis (2 * pi * real_of_int x) = 1" by (case_tac x rule: int_diff_cases) (simp add: right_diff_distrib cis_divide [symmetric] cis_2pi_nat) - def c \ "b - 2*pi * of_int \(b - pi) / (2*pi)\" + define c where "c = b - 2 * pi * of_int \(b - pi) / (2 * pi)\" have "sgn z = cis c" unfolding b c_def by (simp add: cis_divide [symmetric] cis_2pi_int) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Conditionally_Complete_Lattices.thy --- a/src/HOL/Conditionally_Complete_Lattices.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Conditionally_Complete_Lattices.thy Mon Apr 25 22:59:53 2016 +0200 @@ -590,7 +590,7 @@ S = {a.. S = {a..b}" proof - - def lower \ "{x. \s\S. s \ x}" and upper \ "{x. \s\S. x \ s}" + define lower upper where "lower = {x. \s\S. s \ x}" and "upper = {x. \s\S. x \ s}" with ivl have "S = lower \ upper" by auto moreover diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Datatype_Examples/Derivation_Trees/Gram_Lang.thy --- a/src/HOL/Datatype_Examples/Derivation_Trees/Gram_Lang.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Datatype_Examples/Derivation_Trees/Gram_Lang.thy Mon Apr 25 22:59:53 2016 +0200 @@ -665,7 +665,7 @@ assume ?C then obtain tr where tr_tr0: "Inr tr \ cont tr0" and t_tr: "inFr (ns - {root tr0}) tr t" unfolding inFrr_def by auto - def tr1 \ "hsubst tr" + define tr1 where "tr1 = hsubst tr" have 1: "inFr ns tr1 t" using t_tr unfolding tr1_def using inFr_hsubst_minus by auto have "Inr tr1 \ cont (hsubst tr0)" unfolding tr1_def using tr_tr0 by auto thus ?A using 1 inFr.Ind assms by (metis root_hsubst) @@ -689,8 +689,9 @@ lemma regular_def2: "regular tr \ (\ f. reg f tr \ (\ n. root (f n) = n))" unfolding regular_def proof safe - fix f assume f: "reg f tr" - def g \ "\ n. if inItr UNIV tr n then f n else deftr n" + fix f + assume f: "reg f tr" + define g where "g n = (if inItr UNIV tr n then f n else deftr n)" for n show "\g. reg g tr \ (\n. root (g n) = n)" apply(rule exI[of _ g]) using f deftr_simps(1) unfolding g_def reg_def apply safe @@ -1169,8 +1170,8 @@ fix ts assume "ts \ Lr ns n" then obtain tr where dtr: "wf tr" and r: "root tr = n" and tr: "regular tr" and ts: "ts = Fr ns tr" unfolding Lr_def by auto - def tns \ "(id \ root) ` (cont tr)" - def K \ "\ n'. Fr (ns - {n}) (subtrOf tr n')" + define tns where "tns = (id \ root) ` (cont tr)" + define K where "K n' = Fr (ns - {n}) (subtrOf tr n')" for n' show "\tns K. ts = ?F tns K \ (n, tns) \ P \ ?\ tns K" apply(rule exI[of _ tns], rule exI[of _ K]) proof(intro conjI allI impI) show "ts = Inl -` tns \ \{K n' |n'. Inr n' \ tns}" @@ -1222,7 +1223,8 @@ then obtain ftr where 0: "\ n'. Inr n' \ tns \ K n' = Fr (ns - {n}) (ftr n') \ wf (ftr n') \ root (ftr n') = n'" by metis - def tr \ "Node n ((id \ ftr) ` tns)" def tr' \ "hsubst tr tr" + define tr where "tr = Node n ((id \ ftr) ` tns)" + define tr' where "tr' = hsubst tr tr" have rtr: "root tr = n" and ctr: "cont tr = (id \ ftr) ` tns" unfolding tr_def by (simp, metis P cont_Node finite_imageI finite_in_P) have prtr: "prodOf tr = tns" apply(rule Inl_Inr_image_cong) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Decision_Procs/Approximation.thy --- a/src/HOL/Decision_Procs/Approximation.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Decision_Procs/Approximation.thy Mon Apr 25 22:59:53 2016 +0200 @@ -2399,7 +2399,7 @@ shows "ln (Float m e) = ln 2 * (e + (bitlen m - 1)) + ln (Float m (- (bitlen m - 1)))" proof - let ?B = "2^nat (bitlen m - 1)" - def bl \ "bitlen m - 1" + define bl where "bl = bitlen m - 1" have "0 < real_of_int m" and "\X. (0 :: real) < 2^X" and "0 < (2 :: real)" and "m \ 0" using assms by auto hence "0 \ bl" by (simp add: bitlen_def bl_def) @@ -2528,8 +2528,8 @@ using \1 \ x\ by auto show ?thesis proof - - def m \ "mantissa x" - def e \ "exponent x" + define m where "m = mantissa x" + define e where "e = exponent x" from Float_mantissa_exponent[of x] have Float: "x = Float m e" by (simp add: m_def e_def) let ?s = "Float (e + (bitlen m - 1)) 0" @@ -2539,7 +2539,7 @@ apply (auto simp add: zero_less_mult_iff) using not_le powr_ge_pzero apply blast done - def bl \ "bitlen m - 1" + define bl where "bl = bitlen m - 1" hence "bl \ 0" using \m > 0\ by (simp add: bitlen_def) have "1 \ Float m e" @@ -2746,7 +2746,7 @@ with x lu show ?thesis by (auto simp: bnds_powr_def) next assume A: "l1 = 0" "u1 \ 0" "l2 \ 1" - def uln \ "the (ub_ln prec u1)" + define uln where "uln = the (ub_ln prec u1)" show ?thesis proof (cases "x = 0") case False @@ -3969,7 +3969,7 @@ have setsum_move0: "\k F. setsum F {0.. k. F (Suc k)) {0.. "xs!x - c" + define C where "C = xs!x - c" { fix t::real assume t: "t \ {lx .. ux}" @@ -4006,7 +4006,8 @@ and ate: "Some (l, u) = approx_tse prec x s c 1 f vs" shows "interpret_floatarith f xs \ {l .. u}" proof - - def F \ "\n z. interpret_floatarith ((DERIV_floatarith x ^^ n) f) (xs[x := z])" + define F where [abs_def]: "F n z = + interpret_floatarith ((DERIV_floatarith x ^^ n) f) (xs[x := z])" for n z hence F0: "F 0 = (\ z. interpret_floatarith f (xs[x := z]))" by auto hence "bounded_by (xs[x := c]) vs" and "x < length vs" "x < length xs" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Deriv.thy Mon Apr 25 22:59:53 2016 +0200 @@ -235,8 +235,8 @@ let ?D = "\f f' x y. (f y - f x) - f' (y - x)" let ?N = "\f f' x y. norm (?D f f' x y) / norm (y - x)" let ?gf = "\x. g (f x)" and ?gf' = "\x. g' (f' x)" - def Nf \ "?N f f' x" - def Ng \ "\y. ?N g g' (f x) (f y)" + define Nf where "Nf = ?N f f' x" + define Ng where [abs_def]: "Ng y = ?N g g' (f x) (f y)" for y show ?thesis proof (rule has_derivativeI_sandwich[of 1]) @@ -297,7 +297,8 @@ "\a b. norm (a ** b) \ norm a * norm b * K" by fast let ?D = "\f f' y. f y - f x - f' (y - x)" let ?N = "\f f' y. norm (?D f f' y) / norm (y - x)" - def Ng =="?N g g'" and Nf =="?N f f'" + define Ng where "Ng = ?N g g'" + define Nf where "Nf = ?N f f'" let ?fun1 = "\y. norm (f y ** g y - f x ** g x - (f x ** g' (y - x) + f' (y - x) ** g x)) / norm (y - x)" let ?fun2 = "\y. norm (f x) * Ng y * K + Nf y * norm (g y) * K + KF * norm (g y - g x) * K" @@ -1612,10 +1613,10 @@ assumes lim: "((\ x. (f' x / g' x)) \ x) (at_right 0)" shows "((\ x. f0 x / g0 x) \ x) (at_right 0)" proof - - def f \ "\x. if x \ 0 then 0 else f0 x" + define f where [abs_def]: "f x = (if x \ 0 then 0 else f0 x)" for x then have "f 0 = 0" by simp - def g \ "\x. if x \ 0 then 0 else g0 x" + define g where [abs_def]: "g x = (if x \ 0 then 0 else g0 x)" for x then have "g 0 = 0" by simp have "eventually (\x. g0 x \ 0 \ g' x \ 0 \ diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Divides.thy --- a/src/HOL/Divides.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Divides.thy Mon Apr 25 22:59:53 2016 +0200 @@ -623,7 +623,8 @@ with \0 < b\ have "0 < a div b" by (auto intro: div_positive) then have [simp]: "1 \ a div b" by (simp add: discrete) with \0 < b\ have mod_less: "a mod b < b" by (simp add: pos_mod_bound) - def w \ "a div b mod 2" with parity have w_exhaust: "w = 0 \ w = 1" by auto + define w where "w = a div b mod 2" + with parity have w_exhaust: "w = 0 \ w = 1" by auto have mod_w: "a mod (2 * b) = a mod b + b * w" by (simp add: w_def mod_mult2_eq ac_simps) from assms w_exhaust have "w = 1" @@ -641,7 +642,8 @@ shows "2 * (a div (2 * b)) = a div b" (is "?P") and "a mod (2 * b) = a mod b" (is "?Q") proof - - def w \ "a div b mod 2" with parity have w_exhaust: "w = 0 \ w = 1" by auto + define w where "w = a div b mod 2" + with parity have w_exhaust: "w = 0 \ w = 1" by auto have mod_w: "a mod (2 * b) = a mod b + b * w" by (simp add: w_def mod_mult2_eq ac_simps) moreover have "b \ a mod b + b" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Finite_Set.thy Mon Apr 25 22:59:53 2016 +0200 @@ -919,7 +919,7 @@ case 0 then show ?case by simp next case (Suc n g) - def h \ "\z. g z - 1" + define h where "h z = g z - 1" for z with Suc have "n = h y" by simp with Suc have hyp: "f y ^^ h y \ f x = f x \ f y ^^ h y" by auto @@ -927,7 +927,7 @@ then show ?case by (simp add: comp_assoc hyp) (simp add: o_assoc comp_fun_commute) qed - def h \ "\z. if z = x then g x - 1 else g z" + define h where "h z = (if z = x then g x - 1 else g z)" for z with Suc have "n = h x" by simp with Suc have "f y ^^ h y \ f x ^^ h x = f x ^^ h x \ f y ^^ h y" by auto diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Groups_Big.thy --- a/src/HOL/Groups_Big.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Groups_Big.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1137,7 +1137,7 @@ case True with insert show ?thesis by simp next case False with insert have "a \ B" by simp - def C \ "B - {a}" + define C where "C = B - {a}" with \finite B\ \a \ B\ have *: "B = insert a C" "finite C" "a \ C" by auto with insert show ?thesis by (auto simp add: insert_commute ac_simps) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/HOLCF/Adm.thy --- a/src/HOL/HOLCF/Adm.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/HOLCF/Adm.thy Mon Apr 25 22:59:53 2016 +0200 @@ -61,7 +61,7 @@ assumes P: "\i. \j\i. P (Y j)" shows "P (\i. Y i)" proof - - def f \ "\i. LEAST j. i \ j \ P (Y j)" + define f where "f i = (LEAST j. i \ j \ P (Y j))" for i have chain': "chain (\i. Y (f i))" unfolding f_def apply (rule chainI) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/HOLCF/Completion.thy --- a/src/HOL/HOLCF/Completion.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/HOLCF/Completion.thy Mon Apr 25 22:59:53 2016 +0200 @@ -191,14 +191,14 @@ proof - obtain count :: "'a \ nat" where inj: "inj count" using countable .. - def enum \ "\i. THE a. count a = i" + define enum where "enum i = (THE a. count a = i)" for i have enum_count [simp]: "\x. enum (count x) = x" unfolding enum_def by (simp add: inj_eq [OF inj]) - def a \ "LEAST i. enum i \ rep x" - def b \ "\i. LEAST j. enum j \ rep x \ \ enum j \ enum i" - def c \ "\i j. LEAST k. enum k \ rep x \ enum i \ enum k \ enum j \ enum k" - def P \ "\i. \j. enum j \ rep x \ \ enum j \ enum i" - def X \ "rec_nat a (\n i. if P i then c i (b i) else i)" + define a where "a = (LEAST i. enum i \ rep x)" + define b where "b i = (LEAST j. enum j \ rep x \ \ enum j \ enum i)" for i + define c where "c i j = (LEAST k. enum k \ rep x \ enum i \ enum k \ enum j \ enum k)" for i j + define P where "P i \ (\j. enum j \ rep x \ \ enum j \ enum i)" for i + define X where "X = rec_nat a (\n i. if P i then c i (b i) else i)" have X_0: "X 0 = a" unfolding X_def by simp have X_Suc: "\n. X (Suc n) = (if P (X n) then c (X n) (b (X n)) else X n)" unfolding X_def by simp diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/HOLCF/Representable.thy --- a/src/HOL/HOLCF/Representable.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/HOLCF/Representable.thy Mon Apr 25 22:59:53 2016 +0200 @@ -97,7 +97,7 @@ obtain Y where Y: "\i. Y i \ Y (Suc i)" and t: "t = (\i. defl_principal (Y i))" by (rule defl.obtain_principal_chain) - def approx \ "\i. (p oo cast\(defl_principal (Y i)) oo e) :: 'a \ 'a" + define approx where "approx i = (p oo cast\(defl_principal (Y i)) oo e)" for i have "approx_chain approx" proof (rule approx_chain.intro) show "chain (\i. approx i)" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/HOLCF/Up.thy --- a/src/HOL/HOLCF/Up.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/HOLCF/Up.thy Mon Apr 25 22:59:53 2016 +0200 @@ -73,7 +73,7 @@ proof (cases "\k. Y k \ Ibottom") case True then obtain k where k: "Y k \ Ibottom" .. - def A \ "\i. THE a. Iup a = Y (i + k)" + define A where "A i = (THE a. Iup a = Y (i + k))" for i have Iup_A: "\i. Iup (A i) = Y (i + k)" proof fix i :: nat diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Hahn_Banach/Function_Norm.thy --- a/src/HOL/Hahn_Banach/Function_Norm.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Hahn_Banach/Function_Norm.thy Mon Apr 25 22:59:53 2016 +0200 @@ -108,7 +108,7 @@ that \y \ b\ for all \y \ B\. Due to the definition of \B\ there are two cases.\ - def b \ "max c 0" + define b where "b = max c 0" have "\y \ B V f. y \ b" proof fix y assume y: "y \ B V f" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Hahn_Banach/Hahn_Banach.thy --- a/src/HOL/Hahn_Banach/Hahn_Banach.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Hahn_Banach/Hahn_Banach.thy Mon Apr 25 22:59:53 2016 +0200 @@ -55,7 +55,7 @@ interpret subspace F E by fact interpret seminorm E p by fact interpret linearform F f by fact - def M \ "norm_pres_extensions E p F f" + define M where "M = norm_pres_extensions E p F f" then have M: "M = \" by (simp only:) from E have F: "vectorspace F" .. note FE = \F \ E\ @@ -140,7 +140,7 @@ qed qed - def H' \ "H + lin x'" + define H' where "H' = H + lin x'" \ \Define \H'\ as the direct sum of \H\ and the linear closure of \x'\. \<^smallskip>\ have HH': "H \ H'" proof (unfold H'_def) @@ -178,8 +178,8 @@ then show thesis by (blast intro: that) qed - def h' \ "\x. let (y, a) = - SOME (y, a). x = y + a \ x' \ y \ H in h y + a * xi" + define h' where "h' x = (let (y, a) = + SOME (y, a). x = y + a \ x' \ y \ H in h y + a * xi)" for x \ \Define the extension \h'\ of \h\ to \H'\ using \\\. \<^smallskip>\ have "g \ graph H' h' \ g \ graph H' h'" @@ -365,7 +365,7 @@ OF F_norm \continuous F f norm\ , folded B_def fn_norm_def]) txt \We define a function \p\ on \E\ as follows: \p x = \f\ \ \x\\\ - def p \ "\x. \f\\F * \x\" + define p where "p x = \f\\F * \x\" for x txt \\p\ is a seminorm on \E\:\ have q: "seminorm E p" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy --- a/src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy Mon Apr 25 22:59:53 2016 +0200 @@ -86,9 +86,9 @@ \ lemma h'_lf: - assumes h'_def: "h' \ \x. let (y, a) = - SOME (y, a). x = y + a \ x0 \ y \ H in h y + a * xi" - and H'_def: "H' \ H + lin x0" + assumes h'_def: "\x. h' x = (let (y, a) = + SOME (y, a). x = y + a \ x0 \ y \ H in h y + a * xi)" + and H'_def: "H' = H + lin x0" and HE: "H \ E" assumes "linearform H h" assumes x0: "x0 \ H" "x0 \ E" "x0 \ 0" @@ -192,9 +192,9 @@ \ lemma h'_norm_pres: - assumes h'_def: "h' \ \x. let (y, a) = - SOME (y, a). x = y + a \ x0 \ y \ H in h y + a * xi" - and H'_def: "H' \ H + lin x0" + assumes h'_def: "\x. h' x = (let (y, a) = + SOME (y, a). x = y + a \ x0 \ y \ H in h y + a * xi)" + and H'_def: "H' = H + lin x0" and x0: "x0 \ H" "x0 \ E" "x0 \ 0" assumes E: "vectorspace E" and HE: "subspace H E" and "seminorm E p" and "linearform H h" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Hahn_Banach/Hahn_Banach_Sup_Lemmas.thy --- a/src/HOL/Hahn_Banach/Hahn_Banach_Sup_Lemmas.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Hahn_Banach/Hahn_Banach_Sup_Lemmas.thy Mon Apr 25 22:59:53 2016 +0200 @@ -158,7 +158,7 @@ \ lemma sup_definite: - assumes M_def: "M \ norm_pres_extensions E p F f" + assumes M_def: "M = norm_pres_extensions E p F f" and cM: "c \ chains M" and xy: "(x, y) \ \c" and xz: "(x, z) \ \c" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Hahn_Banach/Subspace.thy --- a/src/HOL/Hahn_Banach/Subspace.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Hahn_Banach/Subspace.thy Mon Apr 25 22:59:53 2016 +0200 @@ -467,9 +467,9 @@ lemma h'_definite: fixes H assumes h'_def: - "h' \ \x. - let (y, a) = SOME (y, a). (x = y + a \ x' \ y \ H) - in (h y) + a * xi" + "\x. h' x = + (let (y, a) = SOME (y, a). (x = y + a \ x' \ y \ H) + in (h y) + a * xi)" and x: "x = y + a \ x'" assumes "vectorspace E" "subspace H E" assumes y: "y \ H" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Hilbert_Choice.thy --- a/src/HOL/Hilbert_Choice.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Hilbert_Choice.thy Mon Apr 25 22:59:53 2016 +0200 @@ -110,7 +110,8 @@ 2: "\x n. P n x \ \y. P (Suc n) y \ Q n x y" shows "\f. \n. P n (f n) \ Q n (f n) (f (Suc n))" proof (intro exI allI conjI) - fix n def f \ "rec_nat (SOME x. P 0 x) (\n x. SOME y. P (Suc n) y \ Q n x y)" + fix n + define f where "f = rec_nat (SOME x. P 0 x) (\n x. SOME y. P (Suc n) y \ Q n x y)" have "P 0 (f 0)" "\n. P n (f n) \ P (Suc n) (f (Suc n)) \ Q n (f n) (f (Suc n))" using someI_ex[OF 1] someI_ex[OF 2] by (simp_all add: f_def) then show "P n (f n)" "Q n (f n) (f (Suc n))" @@ -309,8 +310,8 @@ shows "\f. inj (f::nat \ 'a) \ range f \ S" \ \Courtesy of Stephan Merz\ proof - - def Sseq \ "rec_nat S (\n T. T - {SOME e. e \ T})" - def pick \ "\n. (SOME e. e \ Sseq n)" + define Sseq where "Sseq = rec_nat S (\n T. T - {SOME e. e \ T})" + define pick where "pick n = (SOME e. e \ Sseq n)" for n { fix n have "Sseq n \ S" "\ finite (Sseq n)" by (induct n) (auto simp add: Sseq_def inf) } moreover then have *: "\n. pick n \ Sseq n" unfolding pick_def by (subst (asm) finite.simps) (auto simp add: ex_in_conv intro: someI_ex) @@ -746,7 +747,7 @@ then show False by simp qed then obtain n where n: "f n = f (Suc n)" .. - def N \ "LEAST n. f n = f (Suc n)" + define N where "N = (LEAST n. f n = f (Suc n))" have N: "f N = f (Suc N)" unfolding N_def using n by (rule LeastI) show ?thesis diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Isar_Examples/Schroeder_Bernstein.thy --- a/src/HOL/Isar_Examples/Schroeder_Bernstein.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Isar_Examples/Schroeder_Bernstein.thy Mon Apr 25 22:59:53 2016 +0200 @@ -21,8 +21,8 @@ assumes "inj f" and "inj g" shows "\h :: 'a \ 'b. inj h \ surj h" proof - def A \ "lfp (\X. - (g ` (- (f ` X))))" - def g' \ "inv g" + define A where "A = lfp (\X. - (g ` (- (f ` X))))" + define g' where "g' = inv g" let ?h = "\z. if z \ A then f z else g' z" have "A = - (g ` (- (f ` A)))" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Complete_Partial_Order2.thy --- a/src/HOL/Library/Complete_Partial_Order2.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Complete_Partial_Order2.thy Mon Apr 25 22:59:53 2016 +0200 @@ -133,7 +133,7 @@ fix x assume "x \ f y' ` Y" then obtain y where "y \ Y" and x: "x = f y' y" by blast - def y'' \ "if y \ y' then y' else y" + define y'' where "y'' = (if y \ y' then y' else y)" from chain \y \ Y\ \y' \ Y\ have "y \ y' \ y' \ y" by(rule chainD) hence "f y' y \ f y'' y''" using \y \ Y\ \y' \ Y\ by(auto simp add: y''_def intro: mono2 monotoneD[OF mono1]) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/ContNotDenum.thy --- a/src/HOL/Library/ContNotDenum.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/ContNotDenum.thy Mon Apr 25 22:59:53 2016 +0200 @@ -46,8 +46,9 @@ "\a b c. a < b \ c \ {i a b c .. j a b c}" by metis - def ivl \ "rec_nat (f 0 + 1, f 0 + 2) (\n x. (i (fst x) (snd x) (f n), j (fst x) (snd x) (f n)))" - def I \ "\n. {fst (ivl n) .. snd (ivl n)}" + define ivl where "ivl = + rec_nat (f 0 + 1, f 0 + 2) (\n x. (i (fst x) (snd x) (f n), j (fst x) (snd x) (f n)))" + define I where "I n = {fst (ivl n) .. snd (ivl n)}" for n have ivl[simp]: "ivl 0 = (f 0 + 1, f 0 + 2)" @@ -95,7 +96,7 @@ assumes "a < b" "c < d" shows "\f. bij_betw f {a<.. "\a b c d x::real. (d - c)/(b - a) * (x - a) + c" + define f where "f a b c d x = (d - c)/(b - a) * (x - a) + c" for a b c d x :: real { fix a b c d x :: real assume *: "a < b" "c < d" "a < x" "x < b" moreover from * have "(d - c) * (x - a) < (d - c) * (b - a)" by (intro mult_strict_left_mono) simp_all diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Convex.thy --- a/src/HOL/Library/Convex.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Convex.thy Mon Apr 25 22:59:53 2016 +0200 @@ -664,7 +664,7 @@ shows "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" and "(f x - f y) / (x - y) \ (f t - f y) / (t - y)" proof - - def a \ "(t - y) / (x - y)" + define a where "a \ (t - y) / (x - y)" with t have "0 \ a" "0 \ 1 - a" by (auto simp: field_simps) with f \x \ I\ \y \ I\ have cvx: "f (a * x + (1 - a) * y) \ a * f x + (1 - a) * f y" @@ -882,7 +882,7 @@ proof (rule convex_on_linorderI) fix t x y :: real assume t: "t > 0" "t < 1" and xy: "x \ A" "y \ A" "x < y" - def z \ "(1 - t) * x + t * y" + define z where "z = (1 - t) * x + t * y" with \connected A\ and xy have ivl: "{x..y} \ A" using connected_contains_Icc by blast from xy t have xz: "z > x" by (simp add: z_def algebra_simps) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Countable.thy --- a/src/HOL/Library/Countable.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Countable.thy Mon Apr 25 22:59:53 2016 +0200 @@ -142,7 +142,7 @@ assumes finite_item: "\x. rep_set x \ finite_item x" shows "OFCLASS('b, countable_class)" proof - def f \ "\y. LEAST n. nth_item n = Rep y" + define f where "f y = (LEAST n. nth_item n = Rep y)" for y { fix y :: 'b have "rep_set (Rep y)" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Countable_Set_Type.thy --- a/src/HOL/Library/Countable_Set_Type.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Countable_Set_Type.thy Mon Apr 25 22:59:53 2016 +0200 @@ -567,8 +567,8 @@ BNF_Def.Grp {x. rcset x \ {(a, b). R a b}} (cimage snd)) a b" (is "?L = ?R") proof assume ?L - def R' \ "the_inv rcset (Collect (case_prod R) \ (rcset a \ rcset b))" - (is "the_inv rcset ?L'") + define R' where "R' = the_inv rcset (Collect (case_prod R) \ (rcset a \ rcset b))" + (is "_ = the_inv rcset ?L'") have L: "countable ?L'" by auto hence *: "rcset R' = ?L'" unfolding R'_def by (intro rcset_to_rcset) thus ?R unfolding Grp_def relcompp.simps conversep.simps including cset.lifting diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/DAList_Multiset.thy --- a/src/HOL/Library/DAList_Multiset.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/DAList_Multiset.thy Mon Apr 25 22:59:53 2016 +0200 @@ -313,7 +313,7 @@ have inv: "ys \ ?inv" using Cons(2) by auto note IH = Cons(1)[OF inv] - def Ys \ "Abs_multiset (count_of ys)" + define Ys where "Ys = Abs_multiset (count_of ys)" have id: "Abs_multiset (count_of ((a, n) # ys)) = ((op + {# a #}) ^^ n) Ys" unfolding Ys_def proof (rule multiset_eqI, unfold count) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Extended_Real.thy --- a/src/HOL/Library/Extended_Real.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Extended_Real.thy Mon Apr 25 22:59:53 2016 +0200 @@ -3900,7 +3900,7 @@ shows Liminf_inverse_ereal: "Liminf F (\x. inverse (f x)) = inverse (Limsup F f)" and Limsup_inverse_ereal: "Limsup F (\x. inverse (f x)) = inverse (Liminf F f)" proof - - def inv \ "\x. if x \ 0 then \ else inverse x :: ereal" + define inv where [abs_def]: "inv x = (if x \ 0 then \ else inverse x)" for x :: ereal have "continuous_on ({..0} \ {0..}) inv" unfolding inv_def by (intro continuous_on_If) (auto intro!: continuous_intros) also have "{..0} \ {0..} = (UNIV :: ereal set)" by auto diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/FSet.thy --- a/src/HOL/Library/FSet.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/FSet.thy Mon Apr 25 22:59:53 2016 +0200 @@ -939,7 +939,8 @@ BNF_Def.Grp {a. fset a \ {(a, b). R a b}} (fimage snd)) a b" (is "?L = ?R") proof assume ?L - def R' \ "the_inv fset (Collect (case_prod R) \ (fset a \ fset b))" (is "the_inv fset ?L'") + define R' where "R' = + the_inv fset (Collect (case_prod R) \ (fset a \ fset b))" (is "_ = the_inv fset ?L'") have "finite ?L'" by (intro finite_Int[OF disjI2] finite_cartesian_product) (transfer, simp)+ hence *: "fset R' = ?L'" unfolding R'_def by (intro fset_to_fset) show ?R unfolding Grp_def relcompp.simps conversep.simps diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Float.thy --- a/src/HOL/Library/Float.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Float.thy Mon Apr 25 22:59:53 2016 +0200 @@ -953,7 +953,7 @@ then show ?thesis by simp next case False - def n \ "\log 2 (real_of_int x)\" + define n where "n = \log 2 (real_of_int x)\" then have "0 \ n" using \2 \ x\ by simp from \2 \ x\ False have "x mod 2 = 1" "\ 2 dvd x" @@ -1306,7 +1306,7 @@ show ?thesis proof (cases "0 \ l") case True - def x' \ "x * 2 ^ nat l" + define x' where "x' = x * 2 ^ nat l" have "int x * 2 ^ nat l = x'" by (simp add: x'_def of_nat_mult of_nat_power) moreover have "real x * 2 powr l = real x'" @@ -1319,7 +1319,7 @@ by (metis floor_divide_of_int_eq of_int_of_nat_eq) next case False - def y' \ "y * 2 ^ nat (- l)" + define y' where "y' = y * 2 ^ nat (- l)" from \y \ 0\ have "y' \ 0" by (simp add: y'_def) have "int y * 2 ^ nat (- l) = y'" by (simp add: y'_def of_nat_mult of_nat_power) moreover have "real x * real_of_int (2::int) powr real_of_int l / real y = x / real y'" @@ -1683,14 +1683,14 @@ then show ?thesis by simp next case False - def k \ "\log 2 \ai\\" + define k where "k = \log 2 \ai\\" then have "\log 2 \ai\\ = k" by simp then have k: "2 powr k \ \ai\" "\ai\ < 2 powr (k + 1)" by (simp_all add: floor_log_eq_powr_iff \ai \ 0\) have "k \ 0" using assms by (auto simp: k_def) - def r \ "\ai\ - 2 ^ nat k" + define r where "r = \ai\ - 2 ^ nat k" have r: "0 \ r" "r < 2 powr k" using \k \ 0\ k by (auto simp: r_def k_def algebra_simps powr_add abs_if powr_int) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Formal_Power_Series.thy --- a/src/HOL/Library/Formal_Power_Series.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Formal_Power_Series.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1157,8 +1157,8 @@ shows "f div g * g = f" proof (cases "f = 0") assume nz: "f \ 0" - def n \ "subdegree g" - def h \ "fps_shift n g" + define n where "n = subdegree g" + define h where "h = fps_shift n g" from assms have [simp]: "h $ 0 \ 0" unfolding h_def by (simp add: n_def) from assms nz have "f div g * g = fps_shift n (f * inverse h) * g" @@ -1217,8 +1217,8 @@ instance proof fix f g :: "'a fps" - def n \ "subdegree g" - def h \ "fps_shift n g" + define n where "n = subdegree g" + define h where "h = fps_shift n g" show "f div g * g + f mod g = f" proof (cases "g = 0 \ f = 0") @@ -1248,8 +1248,8 @@ assume "h \ 0" show "(h * f) div (h * g) = f div g" proof - - def m \ "subdegree h" - def h' \ "fps_shift m h" + define m where "m = subdegree h" + define h' where "h' = fps_shift m h" have h_decomp: "h = h' * X ^ m" unfolding h'_def m_def by (rule subdegree_decompose) from \h \ 0\ have [simp]: "h'$0 \ 0" by (simp add: h'_def m_def) have "(h * f) div (h * g) = (h' * f * X^m) div (h' * g * X^m)" @@ -1261,9 +1261,7 @@ next fix f g h :: "'a fps" assume [simp]: "h \ 0" - def n \ "subdegree h" - def h' \ "fps_shift n h" - note dfs = n_def h'_def + define n h' where dfs: "n = subdegree h" "h' = fps_shift n h" have "(f + g * h) div h = fps_shift n (f * inverse h') + fps_shift n (g * (h * inverse h'))" by (simp add: fps_divide_def Let_def dfs[symmetric] algebra_simps fps_shift_add) also have "h * inverse h' = (inverse h' * h') * X^n" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Fun_Lexorder.thy --- a/src/HOL/Library/Fun_Lexorder.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Fun_Lexorder.thy Mon Apr 25 22:59:53 2016 +0200 @@ -76,12 +76,12 @@ assumes "finite {k. f k \ g k}" shows "less_fun f g \ f = g \ less_fun g f" proof - - { def K \ "{k. f k \ g k}" + { define K where "K = {k. f k \ g k}" assume "f \ g" then obtain k' where "f k' \ g k'" by auto then have [simp]: "K \ {}" by (auto simp add: K_def) with assms have [simp]: "finite K" by (simp add: K_def) - def q \ "Min K" + define q where "q = Min K" then have "q \ K" and "\k. k \ K \ k \ q" by auto then have "\k. \ k \ q \ k \ K" by blast then have *: "\k. k < q \ f k = g k" by (simp add: K_def) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Function_Growth.thy --- a/src/HOL/Library/Function_Growth.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Function_Growth.thy Mon Apr 25 22:59:53 2016 +0200 @@ -35,7 +35,7 @@ assumes "m \ n" shows "a ^ (m - n) = (a ^ m) div (a ^ n)" proof - - def q == "m - n" + define q where "q = m - n" moreover with assms have "m = q + n" by (simp add: q_def) ultimately show ?thesis using \a \ 0\ by (simp add: power_add) qed diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/More_List.thy --- a/src/HOL/Library/More_List.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/More_List.thy Mon Apr 25 22:59:53 2016 +0200 @@ -165,7 +165,7 @@ lemma strip_while_idem_iff: "strip_while P xs = xs \ no_trailing P xs" proof - - def ys \ "rev xs" + define ys where "ys = rev xs" moreover have "strip_while P (rev ys) = rev ys \ no_trailing P (rev ys)" by (simp add: dropWhile_idem_iff) ultimately show ?thesis by simp diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Multiset.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1071,7 +1071,7 @@ by (simp add: fold_mset_def del: count_union) next case True - def N \ "set_mset M - {x}" + define N where "N = set_mset M - {x}" from N_def True have *: "set_mset M = insert x N" "x \ N" "finite N" by auto then have "Finite_Set.fold (\y. f y ^^ count (M + {#x#}) y) s N = Finite_Set.fold (\y. f y ^^ count M y) s N" @@ -2617,7 +2617,7 @@ obtain j where j_len: "j < length xs'" and nth_j: "xs' ! j = x" by (metis Cons.prems in_set_conv_nth list.set_intros(1) mset_eq_setD) - def xsa \ "take j xs' @ drop (Suc j) xs'" + define xsa where "xsa = take j xs' @ drop (Suc j) xs'" have "mset xs' = {#x#} + mset xsa" unfolding xsa_def using j_len nth_j by (metis (no_types) ab_semigroup_add_class.add_ac(1) append_take_drop_id Cons_nth_drop_Suc @@ -2628,7 +2628,7 @@ len_a: "length ysa = length xsa" and ms_a: "mset (zip xsa ysa) = mset (zip xs ys)" using Cons.hyps(2) by blast - def ys' \ "take j ysa @ y # drop j ysa" + define ys' where "ys' = take j ysa @ y # drop j ysa" have xs': "xs' = take j xsa @ x # drop j xsa" using ms_x j_len nth_j Cons.prems xsa_def by (metis append_eq_append_conv append_take_drop_id diff_Suc_Suc Cons_nth_drop_Suc length_Cons diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Multiset_Order.thy --- a/src/HOL/Library/Multiset_Order.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Multiset_Order.thy Mon Apr 25 22:59:53 2016 +0200 @@ -157,8 +157,8 @@ assume "less_multiset\<^sub>H\<^sub>O M N" then obtain z where z: "count M z < count N z" unfolding less_multiset\<^sub>H\<^sub>O_def by (auto simp: multiset_eq_iff nat_neq_iff) - def X \ "N - M" - def Y \ "M - N" + define X where "X = N - M" + define Y where "Y = M - N" from z show "X \ {#}" unfolding X_def by (auto simp: multiset_eq_iff not_less_eq_eq Suc_le_eq) from z show "X \# N" unfolding X_def by auto show "M = (N - X) + Y" unfolding X_def Y_def multiset_eq_iff count_union count_diff by force diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Omega_Words_Fun.thy --- a/src/HOL/Library/Omega_Words_Fun.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Omega_Words_Fun.thy Mon Apr 25 22:59:53 2016 +0200 @@ -355,7 +355,7 @@ assumes "A \ range w \ {}" obtains u a v where "w = u \ [a] \ v" "A \ set u = {}" "a \ A" proof - - def i \ "LEAST i. w i \ A" + define i where "i = (LEAST i. w i \ A)" show ?thesis proof show "w = prefix i w \ [w i] \ suffix (Suc i) w" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Polynomial.thy --- a/src/HOL/Library/Polynomial.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Polynomial.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1529,7 +1529,7 @@ case (Suc n q r dr) let ?rr = "smult lc r" let ?qq = "coeff r dr" - def [simp]: b \ "monom ?qq n" + define b where [simp]: "b = monom ?qq n" let ?rrr = "?rr - b * d" let ?qqq = "smult lc q + b" note res = Suc(3) @@ -1610,7 +1610,7 @@ proof - let ?cg = "coeff g (degree g)" let ?cge = "?cg ^ (Suc (degree f) - degree g)" - def a \ ?cge + define a where "a = ?cge" obtain q where pdm: "pseudo_divmod f g = (q,r)" using r_def[unfolded pseudo_mod_def] by (cases "pseudo_divmod f g", auto) from pseudo_divmod[OF g pdm] have id: "smult a f = g * q + r" and "r = 0 \ degree r < degree g" @@ -1649,7 +1649,7 @@ let ?rr = "d * r" let ?a = "coeff ?rr dr" let ?qq = "?a div lc" - def [simp]: b \ "monom ?qq n" + define b where [simp]: "b = monom ?qq n" let ?rrr = "d * (r - b)" let ?qqq = "q + b" note res = Suc(3) @@ -2434,8 +2434,8 @@ in (smult ilc q, r))" (is "?l = ?r") proof (cases "g = 0") case False - def lc \ "inverse (coeff g (degree g))" - def h \ "smult lc g" + define lc where "lc = inverse (coeff g (degree g))" + define h where "h = smult lc g" from False have h1: "coeff h (degree h) = 1" and lc: "lc \ 0" unfolding h_def lc_def by auto hence h0: "h \ 0" by auto obtain q r where p: "pseudo_divmod f h = (q,r)" by force @@ -2468,7 +2468,7 @@ case True thus ?thesis unfolding d by auto next case False - def ilc \ "inverse (coeff g (degree g))" + define ilc where "ilc = inverse (coeff g (degree g))" from False have ilc: "ilc \ 0" unfolding ilc_def by auto with False have id: "(g = 0) = False" "(coeffs g = []) = False" "last (coeffs g) = coeff g (degree g)" @@ -2756,9 +2756,9 @@ assumes "P 0 0" "\a p b q. P p q \ P (pCons a p) (pCons b q)" shows "P p q" proof - - def n \ "max (length (coeffs p)) (length (coeffs q))" - def xs \ "coeffs p @ (replicate (n - length (coeffs p)) 0)" - def ys \ "coeffs q @ (replicate (n - length (coeffs q)) 0)" + define n where "n = max (length (coeffs p)) (length (coeffs q))" + define xs where "xs = coeffs p @ (replicate (n - length (coeffs p)) 0)" + define ys where "ys = coeffs q @ (replicate (n - length (coeffs q)) 0)" have "length xs = length ys" by (simp add: xs_def ys_def n_def) hence "P (Poly xs) (Poly ys)" @@ -3284,7 +3284,7 @@ assume "p\0" then obtain n1 where gte_lcoeff:"\x\n1. lead_coeff p \ poly p x" using that pCons by auto have gt_0:"lead_coeff p >0" using pCons(3) \p\0\ by auto - def n\"max n1 (1+ \a\/(lead_coeff p))" + define n where "n = max n1 (1+ \a\/(lead_coeff p))" show ?thesis proof (rule_tac x=n in exI,rule,rule) fix x assume "n \ x" @@ -3351,13 +3351,13 @@ shows "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" proof safe fix p assume rat: "\i. coeff p i \ \" and root: "poly p x = 0" and nz: "p \ 0" - def cs \ "coeffs p" + define cs where "cs = coeffs p" from rat have "\c\range (coeff p). \c'. c = of_rat c'" unfolding Rats_def by blast then obtain f where f: "\i. coeff p i = of_rat (f (coeff p i))" by (subst (asm) bchoice_iff) blast - def cs' \ "map (quotient_of \ f) (coeffs p)" - def d \ "Lcm (set (map snd cs'))" - def p' \ "smult (of_int d) p" + define cs' where "cs' = map (quotient_of \ f) (coeffs p)" + define d where "d = Lcm (set (map snd cs'))" + define p' where "p' = smult (of_int d) p" have "\n. coeff p' n \ \" proof @@ -3365,8 +3365,9 @@ show "coeff p' n \ \" proof (cases "n \ degree p") case True - def c \ "coeff p n" - def a \ "fst (quotient_of (f (coeff p n)))" and b \ "snd (quotient_of (f (coeff p n)))" + define c where "c = coeff p n" + define a where "a = fst (quotient_of (f (coeff p n)))" + define b where "b = snd (quotient_of (f (coeff p n)))" have b_pos: "b > 0" unfolding b_def using quotient_of_denom_pos' by simp have "coeff p' n = of_int d * coeff p n" by (simp add: p'_def) also have "coeff p n = of_rat (of_int a / of_int b)" unfolding a_def b_def @@ -3483,9 +3484,9 @@ lemma order_mult: "p * q \ 0 \ order a (p * q) = order a p + order a q" proof - - def i \ "order a p" - def j \ "order a q" - def t \ "[:-a, 1:]" + define i where "i = order a p" + define j where "j = order a q" + define t where "t = [:-a, 1:]" have t_dvd_iff: "\u. t dvd u \ poly u a = 0" unfolding t_def by (simp add: dvd_iff_poly_eq_0) assume "p * q \ 0" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/Product_Vector.thy --- a/src/HOL/Library/Product_Vector.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/Product_Vector.thy Mon Apr 25 22:59:53 2016 +0200 @@ -125,12 +125,14 @@ fix x assume "x \ S" then obtain e where "0 < e" and S: "\y. dist y x < e \ y \ S" using * by fast - def r \ "e / sqrt 2" and s \ "e / sqrt 2" + define r where "r = e / sqrt 2" + define s where "s = e / sqrt 2" from \0 < e\ have "0 < r" and "0 < s" unfolding r_def s_def by simp_all from \0 < e\ have "e = sqrt (r\<^sup>2 + s\<^sup>2)" unfolding r_def s_def by (simp add: power_divide) - def A \ "{y. dist (fst x) y < r}" and B \ "{y. dist (snd x) y < s}" + define A where "A = {y. dist (fst x) y < r}" + define B where "B = {y. dist (snd x) y < s}" have "open A" and "open B" unfolding A_def B_def by (simp_all add: open_ball) moreover have "x \ A \ B" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/RBT_Impl.thy --- a/src/HOL/Library/RBT_Impl.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/RBT_Impl.thy Mon Apr 25 22:59:53 2016 +0200 @@ -906,7 +906,7 @@ qed simp+ next case (3 xx lta zz vv rta yy ss bb) - def mt[simp]: mt == "Branch B lta zz vv rta" + define mt where [simp]: "mt = Branch B lta zz vv rta" from 3 have "inv2 mt \ inv1 mt" by simp hence "inv2 (rbt_del xx mt) \ (color_of mt = R \ bheight (rbt_del xx mt) = bheight mt \ inv1 (rbt_del xx mt) \ color_of mt = B \ bheight (rbt_del xx mt) = bheight mt - 1 \ inv1l (rbt_del xx mt))" by (blast dest: rbt_del_inv1_inv2) with 3 have 4: "entry_in_tree k v (rbt_del_from_left xx mt yy ss bb) = (False \ xx \ k \ entry_in_tree k v mt \ (k = yy \ v = ss) \ entry_in_tree k v bb)" by (simp add: balance_left_in_tree) @@ -936,7 +936,7 @@ qed auto next case (5 xx aa yy ss lta zz vv rta) - def mt[simp]: mt == "Branch B lta zz vv rta" + define mt where [simp]: "mt = Branch B lta zz vv rta" from 5 have "inv2 mt \ inv1 mt" by simp hence "inv2 (rbt_del xx mt) \ (color_of mt = R \ bheight (rbt_del xx mt) = bheight mt \ inv1 (rbt_del xx mt) \ color_of mt = B \ bheight (rbt_del xx mt) = bheight mt - 1 \ inv1l (rbt_del xx mt))" by (blast dest: rbt_del_inv1_inv2) with 5 have 3: "entry_in_tree k v (rbt_del_from_right xx aa yy ss mt) = (entry_in_tree k v aa \ (k = yy \ v = ss) \ False \ xx \ k \ entry_in_tree k v mt)" by (simp add: balance_right_in_tree) @@ -1924,7 +1924,7 @@ assumes "is_rbt t1" shows "is_rbt (fold (rbt_insert_with_key f) t2 t1)" proof - - def xs \ "entries t2" + define xs where "xs = entries t2" from assms show ?thesis unfolding fold_def xs_def[symmetric] by(induct xs rule: rev_induct)(auto simp add: rbt_insertwk_is_rbt) qed @@ -1936,7 +1936,7 @@ | Some v \ case rbt_lookup t2 k of None \ Some v | Some w \ Some (f k w v))" proof - - def xs \ "entries t1" + define xs where "xs = entries t1" hence dt1: "distinct (map fst xs)" using t1 by(simp add: distinct_entries) with t2 show ?thesis unfolding fold_def map_of_entries[OF t1, symmetric] diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Library/While_Combinator.thy --- a/src/HOL/Library/While_Combinator.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Library/While_Combinator.thy Mon Apr 25 22:59:53 2016 +0200 @@ -60,12 +60,12 @@ by(metis while_option_stop2) theorem while_option_rule: -assumes step: "!!s. P s ==> b s ==> P (c s)" -and result: "while_option b c s = Some t" -and init: "P s" -shows "P t" + assumes step: "!!s. P s ==> b s ==> P (c s)" + and result: "while_option b c s = Some t" + and init: "P s" + shows "P t" proof - - def k == "LEAST k. ~ b ((c ^^ k) s)" + define k where "k = (LEAST k. ~ b ((c ^^ k) s))" from assms have t: "t = (c ^^ k) s" by (simp add: while_option_def k_def split: if_splits) have 1: "ALL ix\X. R x (f x)" proof - def f \ "\x. THE y. R x y" + define f where "f x = (THE y. R x y)" for x { fix x assume "x \ X" with \rel_set R X Y\ \bi_unique R\ have "R x (f x)" by (simp add: bi_unique_def rel_set_def f_def) (metis theI) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Limits.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1070,8 +1070,9 @@ lemma tendsto_of_nat [tendsto_intros]: "filterlim (of_nat :: nat \ 'a :: real_normed_algebra_1) at_infinity sequentially" proof (subst filterlim_at_infinity[OF order.refl], intro allI impI) - fix r :: real assume r: "r > 0" - def n \ "nat \r\" + fix r :: real + assume r: "r > 0" + define n where "n = nat \r\" from r have n: "\m\n. of_nat m \ r" unfolding n_def by linarith from eventually_ge_at_top[of n] show "eventually (\m. norm (of_nat m :: 'a) \ r) sequentially" by eventually_elim (insert n, simp_all) @@ -2199,8 +2200,9 @@ assumes local: "\x. a \ x \ x \ b \ \d>0. \a b. a \ x \ x \ b \ b - a < d \ P a b" shows "P a b" proof - - def bisect \ "rec_nat (a, b) (\n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))" - def l \ "\n. fst (bisect n)" and u \ "\n. snd (bisect n)" + define bisect where "bisect = + rec_nat (a, b) (\n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))" + define l u where "l n = fst (bisect n)" and "u n = snd (bisect n)" for n have l[simp]: "l 0 = a" "\n. l (Suc n) = (if P (l n) ((l n + u n) / 2) then (l n + u n) / 2 else l n)" and u[simp]: "u 0 = b" "\n. u (Suc n) = (if P (l n) ((l n + u n) / 2) then u n else (l n + u n) / 2)" by (simp_all add: l_def u_def bisect_def split: prod.split) @@ -2243,7 +2245,7 @@ lemma compact_Icc[simp, intro]: "compact {a .. b::real}" proof (cases "a \ b", rule compactI) fix C assume C: "a \ b" "\t\C. open t" "{a..b} \ \C" - def T == "{a .. b}" + define T where "T = {a .. b}" from C(1,3) show "\C'\C. finite C' \ {a..b} \ \C'" proof (induct rule: Bolzano) case (trans a b c) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/List.thy --- a/src/HOL/List.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/List.thy Mon Apr 25 22:59:53 2016 +0200 @@ -3468,7 +3468,7 @@ case (3 x1 x2 xs ys) let ?xs = "x1 # x2 # xs" let ?cond = "x1 = x2" - def zs \ "remdups_adj (x2 # xs)" + define zs where "zs = remdups_adj (x2 # xs)" from 3(1-2)[of zs] obtain f where p: "?p f (x2 # xs) zs" unfolding zs_def by (cases ?cond) auto then have f0: "f 0 = 0" @@ -3585,7 +3585,7 @@ apply (rename_tac [2] ys', case_tac [2] ys') by (auto simp: \f 0 = 0\ \f (Suc 0) = Suc 0\) - def f' \ "\x. f (Suc x) - 1" + define f' where "f' x = f (Suc x) - 1" for x { fix i have "Suc 0 \ f (Suc 0)" using f_nth[of 0] \x1 \ x2\ \f 0 = 0\ by auto @@ -4046,8 +4046,8 @@ proof (induct "length (xs @ ys)" arbitrary: xs ys rule: less_induct) case less - def xs' \ "if (length xs \ length ys) then xs else ys" - and ys' \ "if (length xs \ length ys) then ys else xs" + define xs' ys' where "xs' = (if (length xs \ length ys) then xs else ys)" + and "ys' = (if (length xs \ length ys) then ys else xs)" then have prems': "length xs' \ length ys'" "xs' @ ys' = ys' @ xs'" @@ -4499,7 +4499,7 @@ shows "transpose xs ! i = map (\xs. xs ! i) [ys \ xs. i < length ys]" using assms proof (induct arbitrary: i rule: transpose.induct) case (3 x xs xss) - def XS == "(x # xs) # xss" + define XS where "XS = (x # xs) # xss" hence [simp]: "XS \ []" by auto thus ?case proof (cases i) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/MacLaurin.thy --- a/src/HOL/MacLaurin.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/MacLaurin.thy Mon Apr 25 22:59:53 2016 +0200 @@ -85,22 +85,22 @@ (\m "(\t. f t - - (setsum (\m. (diff m 0 / (fact m)) * t^m) {..m. (diff m 0 / (fact m)) * t^m) {.. "(%m t. diff m t - - (setsum (%p. (diff (m + p) 0 / (fact p)) * (t ^ p)) {..(m::nat) t::real. m < n \ (0::real) \ t \ t \ h \ DERIV (difg m) t :> difg (Suc m) t" - using diff_Suc m unfolding difg_def by (rule Maclaurin_lemma2) + using diff_Suc m unfolding difg_def [abs_def] by (rule Maclaurin_lemma2) have difg_eq_0: "\m order cl to (x, f x)"}, *} --- {* because of the def of @{text H} *} +-- {* because of the definition of @{text H} *} apply fast -- {* so it remains to show @{text "(f x, f (lub H cl)) \ r"} *} apply (rule_tac f = "f" in monotoneE) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/MicroJava/Comp/Index.thy --- a/src/HOL/MicroJava/Comp/Index.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/MicroJava/Comp/Index.thy Mon Apr 25 22:59:53 2016 +0200 @@ -98,7 +98,7 @@ -(* The following def should replace the conditions in WellType.thy / wf_java_mdecl +(* The following definition should replace the conditions in WellType.thy / wf_java_mdecl *) definition disjoint_varnames :: "[vname list, (vname \ ty) list] \ bool" where "disjoint_varnames pns lvars \ diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Bounded_Linear_Function.thy --- a/src/HOL/Multivariate_Analysis/Bounded_Linear_Function.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Bounded_Linear_Function.thy Mon Apr 25 22:59:53 2016 +0200 @@ -244,7 +244,7 @@ hence "convergent (\n. X n x)" by (metis Cauchy_convergent_iff) } note convergent_norm1 = this - def y \ "x /\<^sub>R norm x" + define y where "y = x /\<^sub>R norm x" have y: "norm y \ 1" and xy: "x = norm x *\<^sub>R y" by (simp_all add: y_def inverse_eq_divide) have "convergent (\n. norm x *\<^sub>R X n y)" @@ -301,7 +301,7 @@ have "X \ Blinfun v" proof (rule LIMSEQ_I) fix r::real assume "r > 0" - def r' \ "r / 2" + define r' where "r' = r / 2" have "0 < r'" "r' < r" using \r > 0\ by (simp_all add: r'_def) from CauchyD[OF \Cauchy X\ \r' > 0\] obtain M where M: "\m n. m \ M \ n \ M \ norm (X m - X n) < r'" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Mon Apr 25 22:59:53 2016 +0200 @@ -602,7 +602,8 @@ have n_in_upd: "\i. n \ upd ` {..< Suc i}" using \upd 0 = n\ by auto - def f' \ "\i j. if j \ (upd\Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j" + define f' where "f' i j = + (if j \ (upd\Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j)" for i j { fix x i assume i[arith]: "i \ n" then have "enum (Suc i) x = f' i x" unfolding f'_def enum_def using \a n < p\ \a = enum 0\ \upd 0 = n\ \a n = p - 1\ by (simp add: upd_Suc enum_0 n_in_upd) } @@ -616,8 +617,8 @@ proof cases case (ksimplex base upd) then interpret kuhn_simplex p n base upd s' . - def b \ "base (n := p - 1)" - def u \ "\i. case i of 0 \ n | Suc i \ upd i" + define b where "b = base (n := p - 1)" + define u where "u i = (case i of 0 \ n | Suc i \ upd i)" for i have "ksimplex p (Suc n) (s' \ {b})" proof (rule ksimplex.intros, standard) @@ -632,7 +633,7 @@ then show "bij_betw u {.. "\i j. if j \ u`{..< i} then Suc (b j) else b j" + define f' where "f' i j = (if j \ u`{..< i} then Suc (b j) else b j)" for i j have u_eq: "\i. i \ n \ u ` {..< Suc i} = upd ` {..< i} \ { n }" by (auto simp: u_def image_iff upd_inj Ball_def split: nat.split) arith @@ -724,7 +725,7 @@ then have "\!s'. s' \ s \ ksimplex p n s' \ (\b\s'. s - {a} = s'- {b})" proof (elim disjE conjE) assume "i = 0" - def rot \ "\i. if i + 1 = n then 0 else i + 1" + define rot where [abs_def]: "rot i = (if i + 1 = n then 0 else i + 1)" for i let ?upd = "upd \ rot" have rot: "bij_betw rot {..< n} {..< n}" @@ -733,7 +734,8 @@ from rot upd have "bij_betw ?upd {.. "\i j. if j \ ?upd`{..< i} then Suc (enum (Suc 0) j) else enum (Suc 0) j" + define f' where [abs_def]: "f' i j = + (if j \ ?upd`{..< i} then Suc (enum (Suc 0) j) else enum (Suc 0) j)" for i j interpret b: kuhn_simplex p n "enum (Suc 0)" "upd \ rot" "f' ` {.. n}" proof @@ -832,7 +834,7 @@ from \n \ 0\ obtain n' where n': "n = Suc n'" by (cases n) auto - def rot \ "\i. case i of 0 \ n' | Suc i \ i" + define rot where "rot i = (case i of 0 \ n' | Suc i \ i)" for i let ?upd = "upd \ rot" have rot: "bij_betw rot {..< n} {..< n}" @@ -841,8 +843,8 @@ from rot upd have "bij_betw ?upd {.. "base (upd n' := base (upd n') - 1)" - def f' \ "\i j. if j \ ?upd`{..< i} then Suc (b j) else b j" + define b where "b = base (upd n' := base (upd n') - 1)" + define f' where [abs_def]: "f' i j = (if j \ ?upd`{..< i} then Suc (b j) else b j)" for i j interpret b: kuhn_simplex p n b "upd \ rot" "f' ` {.. n}" proof @@ -945,7 +947,7 @@ done next assume i: "0 < i" "i < n" - def i' \ "i - 1" + define i' where "i' = i - 1" with i have "Suc i' < n" by simp with i have Suc_i': "Suc i' = i" @@ -955,7 +957,8 @@ from i upd have "bij_betw ?upd {..< n} {..< n}" by (subst bij_betw_swap_iff) (auto simp: i'_def) - def f' \ "\i j. if j \ ?upd`{..< i} then Suc (base j) else base j" + define f' where [abs_def]: "f' i j = (if j \ ?upd`{..< i} then Suc (base j) else base j)" + for i j interpret b: kuhn_simplex p n base ?upd "f' ` {.. n}" proof show "base \ {.. {..a \ s\ \c \ t\ by auto next assume u: "u l = upd (Suc i')" - def B \ "b.enum ` {..n}" + define B where "B = b.enum ` {..n}" have "b.enum i' = enum i'" using enum_eq_benum[of i'] i by (auto simp add: i'_def gr0_conv_Suc) have "c = t.enum (Suc l)" unfolding c_eq .. @@ -1447,7 +1450,7 @@ and "f ` unit_cube \ unit_cube" shows "\x\unit_cube. f x = x" proof (rule ccontr) - def n \ "DIM('a)" + define n where "n = DIM('a)" have n: "1 \ n" "0 < n" "n \ 0" unfolding n_def by (auto simp add: Suc_le_eq DIM_positive) assume "\ ?thesis" @@ -1625,7 +1628,7 @@ obtain b :: "nat \ 'a" where b: "bij_betw b {..< n} Basis" by atomize_elim (auto simp: n_def intro!: finite_same_card_bij) - def b' \ "inv_into {..< n} b" + define b' where "b' = inv_into {..< n} b" then have b': "bij_betw b' Basis {..< n}" using bij_betw_inv_into[OF b] by auto then have b'_Basis: "\i. i \ Basis \ b' i \ {..< n}" @@ -1668,7 +1671,7 @@ (label (\i\Basis. (real (r (b' i)) / real p) *\<^sub>R i) \ b) i \ (label (\i\Basis. (real (s (b' i)) / real p) *\<^sub>R i) \ b) i" by (rule kuhn_lemma[OF q1 q2 q3]) - def z \ "(\i\Basis. (real (q (b' i)) / real p) *\<^sub>R i)::'a" + define z :: 'a where "z = (\i\Basis. (real (q (b' i)) / real p) *\<^sub>R i)" have "\i\Basis. d / real n \ \(f z - z)\i\" proof (rule ccontr) have "\i\Basis. q (b' i) \ {0..p}" @@ -1710,7 +1713,7 @@ using q(2)[rule_format,OF *] by blast have b'_im: "\i. i \ Basis \ b' i < n" using b' unfolding bij_betw_def by auto - def r' \ "(\i\Basis. (real (r (b' i)) / real p) *\<^sub>R i)::'a" + define r' ::'a where "r' = (\i\Basis. (real (r (b' i)) / real p) *\<^sub>R i)" have "\i. i \ Basis \ r (b' i) \ p" apply (rule order_trans) apply (rule rs(1)[OF b'_im,THEN conjunct2]) @@ -1721,7 +1724,7 @@ unfolding r'_def mem_unit_cube using b'_Basis by (auto simp add: bij_betw_def zero_le_divide_iff divide_le_eq_1) - def s' \ "(\i\Basis. (real (s (b' i)) / real p) *\<^sub>R i)::'a" + define s' :: 'a where "s' = (\i\Basis. (real (s (b' i)) / real p) *\<^sub>R i)" have "\i. i \ Basis \ s (b' i) \ p" apply (rule order_trans) apply (rule rs(2)[OF b'_im, THEN conjunct2]) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy --- a/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1941,9 +1941,9 @@ e * (K/2)^2 \ norm(contour_integral(linepath a' b') f + contour_integral(linepath b' c') f + contour_integral(linepath c' a') f)" proof - note divide_le_eq_numeral1 [simp del] - def a' \ "midpoint b c" - def b' \ "midpoint c a" - def c' \ "midpoint a b" + define a' where "a' = midpoint b c" + define b' where "b' = midpoint c a" + define c' where "c' = midpoint a b" have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" using f continuous_on_subset segments_subset_convex_hull by metis+ have fcont': "continuous_on (closed_segment c' b') f" @@ -2145,17 +2145,19 @@ { fix y::complex assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" and ynz: "y \ 0" - def K \ "1 + max (dist a b) (max (dist b c) (dist c a))" - def e \ "norm y / K^2" + define K where "K = 1 + max (dist a b) (max (dist b c) (dist c a))" + define e where "e = norm y / K^2" have K1: "K \ 1" by (simp add: K_def max.coboundedI1) then have K: "K > 0" by linarith have [iff]: "dist a b \ K" "dist b c \ K" "dist c a \ K" by (simp_all add: K_def) have e: "e > 0" unfolding e_def using ynz K1 by simp - def At \ "\x y z n. convex hull {x,y,z} \ convex hull {a,b,c} \ - dist x y \ K/2^n \ dist y z \ K/2^n \ dist z x \ K/2^n \ - norm(?pathint x y + ?pathint y z + ?pathint z x) \ e*(K/2^n)^2" + define At where "At x y z n \ + convex hull {x,y,z} \ convex hull {a,b,c} \ + dist x y \ K/2^n \ dist y z \ K/2^n \ dist z x \ K/2^n \ + norm(?pathint x y + ?pathint y z + ?pathint z x) \ e*(K/2^n)^2" + for x y z n have At0: "At a b c 0" using fy by (simp add: At_def e_def has_chain_integral_chain_integral3) @@ -2358,8 +2360,8 @@ assume d1_pos: "0 < d1" and d1: "\x x'. \x\convex hull {a, b, c}; x'\convex hull {a, b, c}; cmod (x' - x) < d1\ \ cmod (f x' - f x) < cmod y / (24 * C)" - def e \ "min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))" - def shrink \ "\x. x - e *\<^sub>R (x - d)" + define e where "e = min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))" + define shrink where "shrink x = x - e *\<^sub>R (x - d)" for x let ?pathint = "\x y. contour_integral(linepath x y) f" have e: "0 < e" "e \ 1" "e \ d1 / (4 * C)" "e \ cmod y / 24 / C / B" using d1_pos \C>0\ \B>0\ ynz by (simp_all add: e_def) @@ -2982,7 +2984,7 @@ using open_contains_ball os p(2) by blast then obtain ee where ee: "\z. z \ path_image p \ 0 < ee z \ ball z (ee z) \ s" by metis - def cover \ "(\z. ball z (ee z/3)) ` (path_image p)" + define cover where "cover = (\z. ball z (ee z/3)) ` (path_image p)" have "compact (path_image p)" by (metis p(1) compact_path_image) moreover have "path_image p \ (\c\path_image p. ball c (ee c / 3))" @@ -3001,7 +3003,7 @@ using k by (auto simp: path_image_def) then have eepi: "\i. i \ k \ 0 < ee((p i))" by (metis ee) - def e \ "Min((ee o p) ` k)" + define e where "e = Min((ee o p) ` k)" have fin_eep: "finite ((ee o p) ` k)" using k by blast have enz: "0 < e" @@ -3475,7 +3477,7 @@ then obtain h where h: "polynomial_function h \ pathstart h = pathstart \ \ pathfinish h = pathfinish \ \ (\t \ {0..1}. norm(h t - \ t) < d/2)" using path_approx_polynomial_function [OF \path \\, of "d/2"] d by auto - def nn \ "1/(2* pi*ii) * contour_integral h (\w. 1/(w - z))" + define nn where "nn = 1/(2* pi*ii) * contour_integral h (\w. 1/(w - z))" have "\n. \e > 0. \p. valid_path p \ z \ path_image p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t \ {0..1}. norm(\ t - p t) < e) \ @@ -3924,7 +3926,7 @@ using path_image_def z by auto have gpd: "\ piecewise_C1_differentiable_on {0..1}" using \ valid_path_def by blast - def r \ "(w - z) / (\ 0 - z)" + define r where "r = (w - z) / (\ 0 - z)" have [simp]: "r \ 0" using w z by (auto simp: r_def) have "Arg r \ 2*pi" @@ -3944,7 +3946,7 @@ then obtain t where t: "t \ {0..1}" and eqArg: "Im (integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg r" by blast - def i \ "integral {0..t} (\x. vector_derivative \ (at x) / (\ x - z))" + define i where "i = integral {0..t} (\x. vector_derivative \ (at x) / (\ x - z))" have iArg: "Arg r = Im i" using eqArg by (simp add: i_def) have gpdt: "\ piecewise_C1_differentiable_on {0..t}" @@ -4433,7 +4435,7 @@ by (metis continuous_at_winding_number valid_path_imp_path \ z) then obtain d where "d>0" and d: "\x'. dist x' z < d \ dist (winding_number \ x') (winding_number \ z) < \Re(winding_number \ z)\ - 1/2" using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast - def z' \ "z - (d / (2 * cmod a)) *\<^sub>R a" + define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a" have *: "a \ z' \ b - d / 3 * cmod a" unfolding z'_def inner_mult_right' divide_inverse apply (simp add: divide_simps algebra_simps dot_square_norm power2_eq_square anz) @@ -4612,15 +4614,15 @@ \ contour_integral g2 f = contour_integral g1 f)))" by metis note ee_rule = ee [THEN conjunct2, rule_format] - def C \ "(\t. ball t (ee t / 3)) ` {0..1}" + define C where "C = (\t. ball t (ee t / 3)) ` {0..1}" have "\t \ C. open t" by (simp add: C_def) moreover have "{0..1} \ \C" using ee [THEN conjunct1] by (auto simp: C_def dist_norm) ultimately obtain C' where C': "C' \ C" "finite C'" and C'01: "{0..1} \ \C'" by (rule compactE [OF compact_interval]) - def kk \ "{t \ {0..1}. ball t (ee t / 3) \ C'}" + define kk where "kk = {t \ {0..1}. ball t (ee t / 3) \ C'}" have kk01: "kk \ {0..1}" by (auto simp: kk_def) - def e \ "Min (ee ` kk)" + define e where "e = Min (ee ` kk)" have C'_eq: "C' = (\t. ball t (ee t / 3)) ` kk" using C' by (auto simp: kk_def C_def) have ee_pos[simp]: "\t. t \ {0..1} \ ee t > 0" @@ -4985,7 +4987,7 @@ by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff) have "finite (part_circlepath z r s t -` {y} \ {0..1})" if "y \ k" for y proof - - def w \ "(y - z)/of_real r / exp(ii * of_real s)" + define w where "w = (y - z)/of_real r / exp(ii * of_real s)" have fin: "finite (of_real -` {z. cmod z \ 1 \ exp (\ * complex_of_real (t - s) * z) = w})" apply (rule finite_vimageI [OF finite_bounded_log2]) using \s < t\ apply (auto simp: inj_of_real) @@ -5205,7 +5207,7 @@ case True then show ?thesis by force next case False - def w \ "x - z" + define w where "w = x - z" then have "w \ 0" by (simp add: False) have **: "\t. \Re w = cos t * cmod w; Im w = sin t * cmod w\ \ w = of_real (cmod w) * exp (\ * t)" using cis_conv_exp complex_eq_iff by auto @@ -5279,7 +5281,7 @@ case False have [simp]: "r > 0" using assms le_less_trans norm_ge_zero by blast - def r' \ "norm(w - z)" + define r' where "r' = norm(w - z)" have "r' < r" by (simp add: assms r'_def) have disjo: "cball z r' \ sphere z r = {}" @@ -5379,7 +5381,7 @@ apply (blast intro: integrable_uniform_limit_real) done { fix e::real - def B' \ "B+1" + define B' where "B' = B + 1" have B': "B' > 0" "B' > B" using \0 \ B\ by (auto simp: B'_def) assume "0 < e" then have ev_no': "\\<^sub>F n in F. \x\path_image \. 2 * cmod (f n x - l x) < e / B'" @@ -5463,9 +5465,11 @@ and uw_less: "cmod (u - w) < e * (d / 2) ^ (k+2) / (1 + real k)" for u x proof - - def ff \ "\n::nat. \w. if n = 0 then inverse(x - w)^k - else if n = 1 then k / (x - w)^(Suc k) - else (k * of_real(Suc k)) / (x - w)^(k + 2)" + define ff where [abs_def]: + "ff n w = + (if n = 0 then inverse(x - w)^k + else if n = 1 then k / (x - w)^(Suc k) + else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w have km1: "\z::complex. z \ 0 \ z ^ (k - Suc 0) = z ^ k / z" by (simp add: field_simps) (metis Suc_pred \k \ 0\ neq0_conv power_Suc) have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d / 2))" @@ -6285,7 +6289,7 @@ with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"] obtain B where B: "\x. B \ cmod x \ norm (f x) * 2 < cmod (f z)" by (auto simp: dist_norm) - def R \ "1 + \B\ + norm z" + define R where "R = 1 + \B\ + norm z" have "R > 0" unfolding R_def proof - have "0 \ cmod z + \B\" @@ -6384,7 +6388,7 @@ have 2: "((\u. 1 / (2 * of_real pi * \) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)" if w: "w \ ball z r" for w proof - - def d \ "(r - norm(w - z))" + define d where "d = (r - norm(w - z))" have "0 < d" "d \ r" using w by (auto simp: norm_minus_commute d_def dist_norm) have dle: "\u. cmod (z - u) = r \ d \ cmod (u - w)" unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) @@ -6468,7 +6472,7 @@ show ?thesis by (simp add: f' Cauchy_contour_integral_circlepath_2 [OF g w] derg [OF w] divide_simps) qed - def d \ "(r - norm(w - z))^2" + define d where "d = (r - norm(w - z))^2" have "d > 0" using w by (simp add: dist_commute dist_norm d_def) have dle: "\y. r = cmod (z - y) \ d \ cmod ((y - w)\<^sup>2)" @@ -7082,8 +7086,8 @@ by (simp add: path_image_def compact_imp_bounded compact_continuous_image continuous_on_polymonial_function) then obtain B where "B>0" and B: "\x. x \ path_image \' \ norm x \ B" using bounded_pos by force - def d \ "\z w. if w = z then deriv f z else (f w - f z)/(w - z)" - def v \ "{w. w \ path_image \ \ winding_number \ w = 0}" + define d where [abs_def]: "d z w = (if w = z then deriv f z else (f w - f z)/(w - z))" for z w + define v where "v = {w. w \ path_image \ \ winding_number \ w = 0}" have "path \" "valid_path \" using \ by (auto simp: path_polynomial_function valid_path_polynomial_function) then have ov: "open v" @@ -7120,7 +7124,8 @@ apply (rule continuous_intros holomorphic_intros continuous_on_subset [OF conf] holomorphic_on_subset [OF holf] | force simp add: that)+ done - def h \ "\z. if z \ u then contour_integral \ (d z) else contour_integral \ (\w. f w/(w - z))" + define h where + "h z = (if z \ u then contour_integral \ (d z) else contour_integral \ (\w. f w/(w - z)))" for z have U: "\z. z \ u \ ((d z) has_contour_integral h z) \" apply (simp add: h_def) apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]]) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Complex_Transcendental.thy --- a/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1367,7 +1367,7 @@ proof (rule has_field_derivative_zero_constant) fix z :: complex assume z': "z \ ball 0 1" hence z: "norm z < 1" by (simp add: dist_0_norm) - def t \ "of_real (1 + norm z) / 2 :: complex" + define t :: complex where "t = of_real (1 + norm z) / 2" from z have t: "norm z < norm t" "norm t < 1" unfolding t_def by (simp_all add: field_simps norm_divide del: of_real_add) @@ -2049,7 +2049,7 @@ shows "(\n. g n * z^n) sums Arctan z" and "h z sums Arctan z" proof - - def G \ "\z. (\n. g n * z^n)" + define G where [abs_def]: "G z = (\n. g n * z^n)" for z have summable: "summable (\n. g n * u^n)" if "norm u < 1" for u proof (cases "u = 0") assume u: "u \ 0" @@ -2089,7 +2089,7 @@ proof (rule has_field_derivative_zero_constant) fix u :: complex assume "u \ ball 0 1" hence u: "norm u < 1" by (simp add: dist_0_norm) - def K \ "(norm u + 1) / 2" + define K where "K = (norm u + 1) / 2" from u and abs_Im_le_cmod[of u] have Im_u: "\Im u\ < 1" by linarith from u have K: "0 \ K" "norm u < K" "K < 1" by (simp_all add: K_def) hence "(G has_field_derivative (\n. diffs g n * u ^ n)) (at u)" unfolding G_def @@ -2120,7 +2120,7 @@ assumes x: "x > (0::real)" shows "(\n. (2*((x - 1) / (x + 1)) ^ (2*n+1) / of_nat (2*n+1))) sums ln x" proof - - def y \ "of_real ((x-1)/(x+1)) :: complex" + define y :: complex where "y = of_real ((x-1)/(x+1))" from x have x': "complex_of_real x \ of_real (-1)" by (subst of_real_eq_iff) auto from x have "\x - 1\ < \x + 1\" by linarith hence "norm (complex_of_real (x - 1) / complex_of_real (x + 1)) < 1" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Conformal_Mappings.thy --- a/src/HOL/Multivariate_Analysis/Conformal_Mappings.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Conformal_Mappings.thy Mon Apr 25 22:59:53 2016 +0200 @@ -35,7 +35,7 @@ by (rule contf continuous_intros)+ have holf': "(\u. (f u - y)) holomorphic_on (ball z r)" by (simp add: holf holomorphic_on_diff) - def a \ "(2 * pi)/(fact n)" + define a where "a = (2 * pi)/(fact n)" have "0 < a" by (simp add: a_def) have "B0/r^(Suc n)*2 * pi * r = a*((fact n)*B0/r^n)" using \0 < r\ by (simp add: a_def divide_simps) @@ -105,44 +105,44 @@ case True then show ?thesis by simp next case False - def w \ "complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" - have "1 \ abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" - using \0 < B\ by simp - then have wge1: "1 \ norm w" - by (metis norm_of_real w_def) - then have "w \ 0" by auto - have kB: "0 < fact k * B" - using \0 < B\ by simp - then have "0 \ fact k * B / cmod ((deriv ^^ k) f 0)" - by simp - then have wgeA: "A \ cmod w" - by (simp only: w_def norm_of_real) - have "fact k * B / cmod ((deriv ^^ k) f 0) < abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" - using \0 < B\ by simp - then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w" - by (metis norm_of_real w_def) - then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)" - using False by (simp add: divide_simps mult.commute split: if_split_asm) - also have "... \ fact k * (B * norm w ^ n) / norm w ^ k" - apply (rule Cauchy_inequality) - using holf holomorphic_on_subset apply force - using holf holomorphic_on_imp_continuous_on holomorphic_on_subset apply blast - using \w \ 0\ apply (simp add:) - by (metis nof wgeA dist_0_norm dist_norm) - also have "... = fact k * (B * 1 / cmod w ^ (k-n))" - apply (simp only: mult_cancel_left times_divide_eq_right [symmetric]) - using \k>n\ \w \ 0\ \0 < B\ apply (simp add: divide_simps semiring_normalization_rules) - done - also have "... = fact k * B / cmod w ^ (k-n)" - by simp - finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" . - then have "1 / cmod w < 1 / cmod w ^ (k - n)" - by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos) - then have "cmod w ^ (k - n) < cmod w" - by (metis frac_le le_less_trans norm_ge_zero norm_one not_less order_refl wge1 zero_less_one) - with self_le_power [OF wge1] have False - by (meson diff_is_0_eq not_gr0 not_le that) - then show ?thesis by blast + define w where "w = complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" + have "1 \ abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" + using \0 < B\ by simp + then have wge1: "1 \ norm w" + by (metis norm_of_real w_def) + then have "w \ 0" by auto + have kB: "0 < fact k * B" + using \0 < B\ by simp + then have "0 \ fact k * B / cmod ((deriv ^^ k) f 0)" + by simp + then have wgeA: "A \ cmod w" + by (simp only: w_def norm_of_real) + have "fact k * B / cmod ((deriv ^^ k) f 0) < abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" + using \0 < B\ by simp + then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w" + by (metis norm_of_real w_def) + then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)" + using False by (simp add: divide_simps mult.commute split: if_split_asm) + also have "... \ fact k * (B * norm w ^ n) / norm w ^ k" + apply (rule Cauchy_inequality) + using holf holomorphic_on_subset apply force + using holf holomorphic_on_imp_continuous_on holomorphic_on_subset apply blast + using \w \ 0\ apply (simp add:) + by (metis nof wgeA dist_0_norm dist_norm) + also have "... = fact k * (B * 1 / cmod w ^ (k-n))" + apply (simp only: mult_cancel_left times_divide_eq_right [symmetric]) + using \k>n\ \w \ 0\ \0 < B\ apply (simp add: divide_simps semiring_normalization_rules) + done + also have "... = fact k * B / cmod w ^ (k-n)" + by simp + finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" . + then have "1 / cmod w < 1 / cmod w ^ (k - n)" + by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos) + then have "cmod w ^ (k - n) < cmod w" + by (metis frac_le le_less_trans norm_ge_zero norm_one not_less order_refl wge1 zero_less_one) + with self_le_power [OF wge1] have False + by (meson diff_is_0_eq not_gr0 not_le that) + then show ?thesis by blast qed then have "(deriv ^^ (k + Suc n)) f 0 / fact (k + Suc n) * \ ^ (k + Suc n) = 0" for k using not_less_eq by blast @@ -185,8 +185,8 @@ apply (rule LeastI2) apply (fastforce intro: dest!: not_less_Least)+ done - def b \ "\i. a (i+m) / a m" - def g \ "\x. suminf (\i. b i * (x - \) ^ i)" + define b where "b i = a (i+m) / a m" for i + define g where "g x = suminf (\i. b i * (x - \) ^ i)" for x have [simp]: "b 0 = 1" by (simp add: am b_def) { fix x::'a @@ -258,7 +258,7 @@ proof - obtain e where "0 < e" and e: "cball \ e \ S" using \open S\ \\ \ S\ open_contains_cball_eq by blast - def T \ "cball \ e \ U" + define T where "T = cball \ e \ U" have contf: "continuous_on (closure T) f" by (metis T_def closed_cball closure_minimal e holf holomorphic_on_imp_continuous_on holomorphic_on_subset inf.cobounded1) @@ -295,7 +295,7 @@ by (metis norm_less dist_norm fnz less_eq_real_def mem_ball mem_cball norm_not_less_zero norm_zero) have "frontier(cball \ r) \ {}" using \0 < r\ by simp - def g \ "\z. inverse (f z)" + define g where [abs_def]: "g z = inverse (f z)" for z have contg: "continuous_on (cball \ r) g" unfolding g_def using contf continuous_on_inverse fnz' by blast have holg: "g holomorphic_on ball \ r" @@ -377,7 +377,7 @@ apply (rule bexE [OF continuous_attains_inf [OF cof frne contfr]]) apply (simp add: dist_norm) done - moreover def \ \ "norm (f w - f \) / 3" + moreover define \ where "\ \ norm (f w - f \) / 3" ultimately have "0 < \" using \0 < r\ dist_complex_def r sne by auto have "ball (f \) \ \ f ` U" @@ -601,12 +601,12 @@ obtain r where "r>0" and r: "ball \ r \ S" using assms by (blast elim!: openE) then have holfb: "f holomorphic_on ball \ r" using holf holomorphic_on_subset by blast - def g \ "\w. suminf (\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i)" + define g where "g w = suminf (\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i)" for w have sumsg: "(\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i) sums g w" and feq: "f w - f \ = (w - \)^n * g w" if w: "w \ ball \ r" for w proof - - def powf \ "(\i. (deriv ^^ i) f \/(fact i) * (w - \)^i)" + define powf where "powf = (\i. (deriv ^^ i) f \/(fact i) * (w - \)^i)" have sing: "{.. = 0 then {} else {0})" unfolding powf_def using \0 < n\ dfz by (auto simp: dfz; metis funpow_0 not_gr0) have "powf sums f w" @@ -737,7 +737,7 @@ case False then obtain n0 where "n0 > 0" and n0: "(deriv ^^ n0) f \ \ 0" by blast obtain r0 where "r0 > 0" "ball \ r0 \ S" using S openE \\ \ S\ by auto - def n \ "LEAST n. (deriv ^^ n) f \ \ 0" + define n where "n \ LEAST n. (deriv ^^ n) f \ \ 0" have n_ne: "(deriv ^^ n) f \ \ 0" by (rule def_LeastI [OF n_def]) (rule n0) then have "0 < n" using \f \ = 0\ @@ -766,7 +766,7 @@ "ball \ r \ S" "\w. w \ ball \ r \ k * norm(w - \)^n \ norm(f w - f \)" proof - - def n \ "LEAST n. 0 < n \ (deriv ^^ n) f \ \ 0" + define n where "n = (LEAST n. 0 < n \ (deriv ^^ n) f \ \ 0)" obtain n0 where "0 < n0" and n0: "(deriv ^^ n0) f \ \ 0" using fne holomorphic_fun_eq_const_on_connected [OF holf S] \\ \ S\ \\ \ S\ by blast then have "0 < n" and n_ne: "(deriv ^^ n) f \ \ 0" @@ -781,7 +781,7 @@ obtain e where "e>0" and e: "ball \ e \ S" using assms by (blast elim!: openE) then have holfb: "f holomorphic_on ball \ e" using holf holomorphic_on_subset by blast - def d \ "(min e r) / 2" + define d where "d = (min e r) / 2" have "0 < d" using \0 < r\ \0 < e\ by (simp add: d_def) have "d < r" using \0 < r\ by (auto simp: d_def) @@ -847,7 +847,7 @@ by (rule lim_null_mult_right_bounded [OF _ that]) (simp add: LIM_zero) moreover have "?P" if "(\z. (z - \) * f z) \\\ 0" proof - - def h \ "\z. (z - \)^2 * f z" + define h where [abs_def]: "h z = (z - \)^2 * f z" for z have h0: "(h has_field_derivative 0) (at \)" apply (simp add: h_def Derivative.DERIV_within_iff) apply (rule Lim_transform_within [OF that, of 1]) @@ -870,7 +870,7 @@ by (simp add: h_def power2_eq_square derivative_intros) qed qed - def g \ "\z. if z = \ then deriv h \ else (h z - h \) / (z - \)" + define g where [abs_def]: "g z = (if z = \ then deriv h \ else (h z - h \) / (z - \))" for z have holg: "g holomorphic_on S" unfolding g_def by (rule pole_lemma [OF holh \]) show ?thesis @@ -1058,7 +1058,7 @@ next case False then obtain k where k: "0 < k" "k\n" "a k \ 0" by force - def m \ "GREATEST k. k\n \ a k \ 0" + define m where "m = (GREATEST k. k\n \ a k \ 0)" have m: "m\n \ a m \ 0" unfolding m_def apply (rule GreatestI [where b = "Suc n"]) @@ -1185,7 +1185,7 @@ next case False then obtain n0 where n0: "n0 > 0 \ (deriv ^^ n0) f \ \ 0" by blast - def n \ "LEAST n. n > 0 \ (deriv ^^ n) f \ \ 0" + define n where [abs_def]: "n = (LEAST n. n > 0 \ (deriv ^^ n) f \ \ 0)" have n_ne: "n > 0" "(deriv ^^ n) f \ \ 0" using def_LeastI [OF n_def n0] by auto have n_min: "\k. 0 < k \ k < n \ (deriv ^^ k) f \ = 0" @@ -1220,7 +1220,7 @@ using \0 < r\ \0 < \\ apply (simp_all add:) by (meson Topology_Euclidean_Space.open_ball centre_in_ball) - def U \ "(\w. (w - \) * g w) ` T" + define U where "U = (\w. (w - \) * g w) ` T" have "open U" by (metis oimT U_def) have "0 \ U" apply (auto simp: U_def) @@ -1350,7 +1350,7 @@ assumes holf: "f holomorphic_on (ball 0 r)" and [simp]: "f 0 = 0" obtains h where "h holomorphic_on (ball 0 r)" and "\z. norm z < r \ f z = z * (h z)" and "deriv f 0 = h 0" proof - - def h \ "\z. if z = 0 then deriv f 0 else f z / z" + define h where "h z = (if z = 0 then deriv f 0 else f z / z)" for z have d0: "deriv f 0 = h 0" by (simp add: h_def) moreover have "h holomorphic_on (ball 0 r)" @@ -1742,7 +1742,7 @@ case True then show ?thesis by simp next case False - def C \ "2 * norm(deriv f 0)" + define C where "C = 2 * norm(deriv f 0)" have "0 < C" using False by (simp add: C_def) have holf': "f holomorphic_on ball 0 r" using holf using ball_subset_cball holomorphic_on_subset by blast @@ -1932,9 +1932,9 @@ assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1" obtains b r where "1/12 < r" "ball b r \ f ` (ball a 1)" proof - - def r \ "249/256::real" + define r :: real where "r = 249/256" have "0 < r" "r < 1" by (auto simp: r_def) - def g \ "\z. deriv f z * of_real(r - norm(z - a))" + define g where "g z = deriv f z * of_real(r - norm(z - a))" for z have "deriv f holomorphic_on ball a 1" by (rule holomorphic_deriv [OF holf open_ball]) then have "continuous_on (ball a 1) (deriv f)" @@ -1950,7 +1950,7 @@ obtain p where pr: "p \ cball a r" and pge: "\y. y \ cball a r \ norm (g y) \ norm (g p)" using distance_attains_sup [OF 1 2, of 0] by force - def t \ "(r - norm(p - a)) / 2" + define t where "t = (r - norm(p - a)) / 2" have "norm (p - a) \ r" using pge [of a] \r > 0\ by (auto simp: g_def norm_mult) then have "norm (p - a) < r" using pr @@ -2028,73 +2028,73 @@ using ball_eq_empty that by fastforce next case False - def C \ "deriv f a" - have "0 < norm C" using False by (simp add: C_def) - have dfa: "f field_differentiable at a" - apply (rule holomorphic_on_imp_differentiable_at [OF holf]) - using \0 < r\ by auto - have fo: "(\z. f (a + of_real r * z)) = f o (\z. (a + of_real r * z))" - by (simp add: o_def) - have holf': "f holomorphic_on (\z. a + complex_of_real r * z) ` ball 0 1" - apply (rule holomorphic_on_subset [OF holf]) - using \0 < r\ apply (force simp: dist_norm norm_mult) - done - have 1: "(\z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1" - apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+ - using \0 < r\ by (simp add: C_def False) - have "((\z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative - (deriv f (a + of_real r * z) / C)) (at z)" - if "norm z < 1" for z - proof - + define C where "C = deriv f a" + have "0 < norm C" using False by (simp add: C_def) + have dfa: "f field_differentiable at a" + apply (rule holomorphic_on_imp_differentiable_at [OF holf]) + using \0 < r\ by auto + have fo: "(\z. f (a + of_real r * z)) = f o (\z. (a + of_real r * z))" + by (simp add: o_def) + have holf': "f holomorphic_on (\z. a + complex_of_real r * z) ` ball 0 1" + apply (rule holomorphic_on_subset [OF holf]) + using \0 < r\ apply (force simp: dist_norm norm_mult) + done + have 1: "(\z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1" + apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+ + using \0 < r\ by (simp add: C_def False) + have "((\z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative + (deriv f (a + of_real r * z) / C)) (at z)" + if "norm z < 1" for z + proof - have *: "((\x. f (a + of_real r * x)) has_field_derivative - (deriv f (a + of_real r * z) * of_real r)) (at z)" - apply (simp add: fo) - apply (rule DERIV_chain [OF field_differentiable_derivI]) - apply (rule holomorphic_on_imp_differentiable_at [OF holf], simp) - using \0 < r\ apply (simp add: dist_norm norm_mult that) - apply (rule derivative_eq_intros | simp)+ - done - show ?thesis - apply (rule derivative_eq_intros * | simp)+ - using \0 < r\ by (auto simp: C_def False) - qed - have 2: "deriv (\z. f (a + of_real r * z) / (C * of_real r)) 0 = 1" - apply (subst deriv_cdivide_right) - apply (simp add: field_differentiable_def fo) - apply (rule exI) + (deriv f (a + of_real r * z) * of_real r)) (at z)" + apply (simp add: fo) apply (rule DERIV_chain [OF field_differentiable_derivI]) - apply (simp add: dfa) - apply (rule derivative_eq_intros | simp add: C_def False fo)+ - using \0 < r\ - apply (simp add: C_def False fo) - apply (simp add: derivative_intros dfa complex_derivative_chain) + apply (rule holomorphic_on_imp_differentiable_at [OF holf], simp) + using \0 < r\ apply (simp add: dist_norm norm_mult that) + apply (rule derivative_eq_intros | simp)+ done - have sb1: "op * (C * r) ` (\z. f (a + of_real r * z) / (C * r)) ` ball 0 1 - \ f ` ball a r" - using \0 < r\ by (auto simp: dist_norm norm_mult C_def False) - have sb2: "ball (C * r * b) r' \ op * (C * r) ` ball b t" - if "1 / 12 < t" for b t - proof - - have *: "r * cmod (deriv f a) / 12 \ r * (t * cmod (deriv f a))" - using that \0 < r\ less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right - by auto - show ?thesis - apply clarify - apply (rule_tac x="x / (C * r)" in image_eqI) - using \0 < r\ - apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps) - apply (erule less_le_trans) - apply (rule order_trans [OF r' *]) - done - qed show ?thesis - apply (rule Bloch_unit [OF 1 2]) - apply (rename_tac t) - apply (rule_tac b="(C * of_real r) * b" in that) - apply (drule image_mono [where f = "\z. (C * of_real r) * z"]) - using sb1 sb2 - apply force + apply (rule derivative_eq_intros * | simp)+ + using \0 < r\ by (auto simp: C_def False) + qed + have 2: "deriv (\z. f (a + of_real r * z) / (C * of_real r)) 0 = 1" + apply (subst deriv_cdivide_right) + apply (simp add: field_differentiable_def fo) + apply (rule exI) + apply (rule DERIV_chain [OF field_differentiable_derivI]) + apply (simp add: dfa) + apply (rule derivative_eq_intros | simp add: C_def False fo)+ + using \0 < r\ + apply (simp add: C_def False fo) + apply (simp add: derivative_intros dfa complex_derivative_chain) + done + have sb1: "op * (C * r) ` (\z. f (a + of_real r * z) / (C * r)) ` ball 0 1 + \ f ` ball a r" + using \0 < r\ by (auto simp: dist_norm norm_mult C_def False) + have sb2: "ball (C * r * b) r' \ op * (C * r) ` ball b t" + if "1 / 12 < t" for b t + proof - + have *: "r * cmod (deriv f a) / 12 \ r * (t * cmod (deriv f a))" + using that \0 < r\ less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right + by auto + show ?thesis + apply clarify + apply (rule_tac x="x / (C * r)" in image_eqI) + using \0 < r\ + apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps) + apply (erule less_le_trans) + apply (rule order_trans [OF r' *]) done + qed + show ?thesis + apply (rule Bloch_unit [OF 1 2]) + apply (rename_tac t) + apply (rule_tac b="(C * of_real r) * b" in that) + apply (drule image_mono [where f = "\z. (C * of_real r) * z"]) + using sb1 sb2 + apply force + done qed corollary Bloch_general: @@ -2157,8 +2157,8 @@ then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ \ (x' - z) ^ (n - m) * g x' = h x'" for x' by auto next - def F\"at z within ball z r" - and f'\"\x. (x - z) ^ (n-m)" + define F where "F = at z within ball z r" + define f' where [abs_def]: "f' x = (x - z) ^ (n-m)" for x have "f' z=0" using \n>m\ unfolding f'_def by auto moreover have "continuous F f'" unfolding f'_def F_def by (intro continuous_intros) @@ -2187,8 +2187,8 @@ then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ \ (x' - z) ^ (m - n) * h x' = g x'" for x' by auto next - def F\"at z within ball z r" - and f'\"\x. (x - z) ^ (m-n)" + define F where "F = at z within ball z r" + define f' where [abs_def]: "f' x = (x - z) ^ (m-n)" for x have "f' z=0" using \m>n\ unfolding f'_def by auto moreover have "continuous F f'" unfolding f'_def F_def by (intro continuous_intros) @@ -2224,7 +2224,7 @@ "\w. w \ ball z r \ g w \ 0" using holomorphic_factor_zero_nonconstant[OF holo \open s\ \connected s\ \z\s\ \f z=0\] by (metis assms(3) assms(5) assms(6)) - def r'\"r/2" + define r' where "r' = r/2" have "cball z r' \ ball z r" unfolding r'_def by (simp add: \0 < r\ cball_subset_ball_iff) hence "cball z r' \ s" "g holomorphic_on cball z r'" "(\w\cball z r'. f w = (w - z) ^ n * g w \ g w \ 0)" @@ -2236,14 +2236,14 @@ by (simp add:\0 < n\) next fix m n - def fac\"\n g r. \w\cball z r. f w = (w - z) ^ n * g w \ g w \ 0" + define fac where "fac n g r \ (\w\cball z r. f w = (w - z) ^ n * g w \ g w \ 0)" for n g r assume n_asm:"\g r1. 0 < n \ 0 < r1 \ g holomorphic_on cball z r1 \ fac n g r1" and m_asm:"\h r2. 0 < m \ 0 < r2 \ h holomorphic_on cball z r2 \ fac m h r2" obtain g r1 where "0 < n" "0 < r1" and g_holo: "g holomorphic_on cball z r1" and "fac n g r1" using n_asm by auto obtain h r2 where "0 < m" "0 < r2" and h_holo: "h holomorphic_on cball z r2" and "fac m h r2" using m_asm by auto - def r\"min r1 r2" + define r where "r = min r1 r2" have "r>0" using \r1>0\ \r2>0\ unfolding r_def by auto moreover have "\w\ball z r. f w = (w-z)^n * g w \ g w\0 \ f w = (w - z)^m * h w \ h w\0" using \fac m h r2\ \fac n g r1\ unfolding fac_def r_def @@ -2276,7 +2276,7 @@ fix p assume "p\s" then obtain e1 where "e1>0" and e1: "\w\ball p e1. w\s \ (w\p \ w\pts)" using finite_ball_avoid[OF assms] by auto - def e2\"e1/2" + define e2 where "e2 = e1/2" have "e2>0" and "e2e1>0\ by auto then have "cball p e2 \ ball p e1" by (subst cball_subset_ball_iff,auto) then show "\e>0. \w\cball p e. w \ s \ (w \ p \ w \ pts)" using \e2>0\ e1 by auto @@ -2302,39 +2302,39 @@ using finite_ball_avoid[OF \open s\ \finite (insert p pts)\,rule_format,of a] \a \ s - insert p pts\ by auto - def a'\"a+e/2" + define a' where "a' = a+e/2" have "a'\s-{p} -pts" using e[rule_format,of "a+e/2"] \e>0\ by (auto simp add:dist_complex_def a'_def) then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b" "path_image g' \ s - {p} - pts" "f contour_integrable_on g'" using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1) by (metis Diff_insert2 open_delete) - def g\"linepath a a' +++ g'" + define g where "g = linepath a a' +++ g'" have "valid_path g" unfolding g_def by (auto intro: valid_path_join) moreover have "pathstart g = a" and "pathfinish g = b" unfolding g_def by auto moreover have "path_image g \ s - insert p pts" unfolding g_def - proof (rule subset_path_image_join) - have "closed_segment a a' \ ball a e" using \e>0\ - by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) - then show "path_image (linepath a a') \ s - insert p pts" using e idt(9) - by auto - next - show "path_image g' \ s - insert p pts" using g'(4) by blast - qed + proof (rule subset_path_image_join) + have "closed_segment a a' \ ball a e" using \e>0\ + by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) + then show "path_image (linepath a a') \ s - insert p pts" using e idt(9) + by auto + next + show "path_image g' \ s - insert p pts" using g'(4) by blast + qed moreover have "f contour_integrable_on g" - proof - - have "closed_segment a a' \ ball a e" using \e>0\ - by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) - then have "continuous_on (closed_segment a a') f" - using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)] - apply (elim continuous_on_subset) - by auto - then have "f contour_integrable_on linepath a a'" - using contour_integrable_continuous_linepath by auto - then show ?thesis unfolding g_def - apply (rule contour_integrable_joinI) - by (auto simp add: \e>0\) - qed + proof - + have "closed_segment a a' \ ball a e" using \e>0\ + by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) + then have "continuous_on (closed_segment a a') f" + using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)] + apply (elim continuous_on_subset) + by auto + then have "f contour_integrable_on linepath a a'" + using contour_integrable_continuous_linepath by auto + then show ?thesis unfolding g_def + apply (rule contour_integrable_joinI) + by (auto simp add: \e>0\) + qed ultimately show ?case using idt.prems(1)[of g] by auto qed @@ -2376,9 +2376,10 @@ obtain n::int where "n=winding_number g p" using integer_winding_number[OF _ g_loop,of p] valid path_img by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path) - def p_circ\"circlepath p (h p)" and p_circ_pt\"linepath (p+h p) (p+h p)" - def n_circ\"\n. (op +++ p_circ ^^ n) p_circ_pt" - def cp\"if n\0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))" + define p_circ where "p_circ = circlepath p (h p)" + define p_circ_pt where "p_circ_pt = linepath (p+h p) (p+h p)" + define n_circ where "n_circ n = (op +++ p_circ ^^ n) p_circ_pt" for n + define cp where "cp = (if n\0 then reversepath (n_circ (nat n)) else n_circ (nat (- n)))" have n_circ:"valid_path (n_circ k)" "winding_number (n_circ k) p = k" "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p" @@ -2509,7 +2510,7 @@ unfolding cp_def using contour_integral_reversepath[OF n_circ(1)] n_circ(9) by auto qed - def g'\"g +++ pg +++ cp +++ (reversepath pg)" + define g' where "g' = g +++ pg +++ cp +++ (reversepath pg)" have "contour_integral g' f = (\p\pts. winding_number g' p * contour_integral (circlepath p (h p)) f)" proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \finite pts\ ]) show "connected (s - {p} - pts)" using connected by (metis Diff_insert2) @@ -2523,8 +2524,8 @@ unfolding g'_def cp_def using pg(2) by simp show "path_image g' \ s - {p} - pts" proof - - def s'\"s - {p} - pts" - have s':"s' = s-insert p pts " unfolding s'_def by auto + define s' where "s' = s - {p} - pts" + have s': "s' = s-insert p pts " unfolding s'_def by auto then show ?thesis using path_img pg(4) cp(4) unfolding g'_def apply (fold s'_def s') @@ -2648,9 +2649,10 @@ shows "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" (is "?L=?R") proof - - def circ\"\p. winding_number g p * contour_integral (circlepath p (h p)) f" - def pts1\"pts \ s" - def pts2\"pts - pts1" + define circ + where [abs_def]: "circ p = winding_number g p * contour_integral (circlepath p (h p)) f" for p + define pts1 where "pts1 = pts \ s" + define pts2 where "pts2 = pts - pts1" have "pts=pts1 \ pts2" "pts1 \ pts2 = {}" "pts2 \ s={}" "pts1\s" unfolding pts1_def pts2_def by auto have "contour_integral g f = (\p\pts1. circ p)" unfolding circ_def @@ -2680,8 +2682,6 @@ qed - - (*order of the zero of f at z*) definition zorder::"(complex \ complex) \ complex \ nat" where "zorder f z = (THE n. n>0 \ (\h r. r>0 \ h holomorphic_on cball z r @@ -2714,8 +2714,8 @@ shows "\r. n>0 \ r>0 \ cball z r \ s \ h holomorphic_on cball z r \ (\w\cball z r. f w = h w * (w-z)^n \ h w \0) " proof - - def P\"\h r n. r>0 \ h holomorphic_on cball z r - \ (\w\cball z r. ( f w = h w * (w-z)^n) \ h w \0)" + define P where "P h r n \ r>0 \ h holomorphic_on cball z r + \ (\w\cball z r. ( f w = h w * (w-z)^n) \ h w \0)" for h r n have "(\!n. n>0 \ (\ h r. P h r n))" proof - have "\!n. \h r. n>0 \ P h r n" @@ -2739,7 +2739,7 @@ then obtain r1 where "P h r1 n" by auto obtain r2 where "r2>0" "cball z r2 \ s" using assms(3) assms(5) open_contains_cball_eq by blast - def r3\"min r1 r2" + define r3 where "r3 = min r1 r2" have "P h r3 n" using \P h r1 n\ \r2>0\ unfolding P_def r3_def by auto moreover have "cball z r3 \ s" using \cball z r2 \ s\ unfolding r3_def by auto @@ -2755,7 +2755,8 @@ shows "\r. n>0 \ r>0 \ cball z r \ s \ h holomorphic_on cball z r \ (\w\cball z r. f w = h w / (w-z)^n \ h w \0)" proof - - def zo\"zorder (inverse \ f) z" and zp\"zer_poly (inverse \ f) z" + define zo where "zo = zorder (inverse \ f) z" + define zp where "zp = zer_poly (inverse \ f) z" obtain r where "0 < zo" "0 < r" "cball z r \ s" and zp_holo: "zp holomorphic_on cball z r" and zp_fac: "\w\cball z r. (inverse \ f) w = zp w * (w - z) ^ zo \ zp w \ 0" using zorder_exist[OF \open s\ \connected s\ \z\s\ holo,folded zo_def zp_def] @@ -2782,14 +2783,15 @@ shows "\r>0. cball z r \ s \ (f has_contour_integral complex_of_real (2 * pi) * \ * residue f z) (circlepath z r)" proof - - def n\"porder f z" and h\"pol_poly f z" + define n where "n = porder f z" + define h where "h = pol_poly f z" obtain r where "n>0" "0 < r" and r_b:"cball z r \ s" and h_holo:"h holomorphic_on cball z r" and h:"(\w\cball z r. f w = h w / (w - z) ^ n \ h w \ 0)" using porder_exist[OF \open s\ \connected s\ \z\s\ holo \f z=0\ non_c] unfolding n_def h_def by auto - def c\"complex_of_real (2 * pi) * \" + define c where "c = complex_of_real (2 * pi) * \" have "residue f z = (deriv ^^ (n - 1)) h z / fact (n-1)" unfolding residue_def apply (fold n_def h_def) @@ -2819,10 +2821,11 @@ poles:"\p\poles. is_pole f p" shows "contour_integral \ f = 2 * pi * \ *(\p\poles. winding_number \ p * residue f p)" proof - - def pts\"{p. f p=0}" - def c\"2 * complex_of_real pi * \ " - def contour\"\p e. (f has_contour_integral c * residue f p) (circlepath p e)" - def avoid\"\p e. \w\cball p e. w \ s \ (w \ p \ w \ pts)" + define pts where "pts = {p. f p=0}" + define c where "c = 2 * complex_of_real pi * \ " + define contour + where "contour p e = (f has_contour_integral c * residue f p) (circlepath p e)" for p e + define avoid where "avoid p e \ (\w\cball p e. w \ s \ (w \ p \ w \ pts))" for p e have "poles \ pts" and "finite pts" using poles \finite {p. f p=0}\ unfolding pts_def is_pole_def by auto have "\e>0. avoid p e \ (p\poles \ contour p e)" @@ -2832,35 +2835,35 @@ using finite_cball_avoid[OF \open s\ \finite pts\] \p\s\ unfolding avoid_def by auto have "\e2>0. cball p e2 \ ball p e1 \ contour p e2" when "p\poles" unfolding c_def contour_def - proof (rule base_residue[of "ball p e1" p f,simplified,OF \e1>0\]) - show "inverse \ f holomorphic_on ball p e1" - proof - - def f'\"inverse o f" - have "f holomorphic_on ball p e1 - {p}" - using holo e \poles \ pts\ unfolding avoid_def - apply (elim holomorphic_on_subset) - by auto - then have f'_holo:"f' holomorphic_on ball p e1 - {p}" unfolding f'_def comp_def - apply (elim holomorphic_on_inverse) - using e pts_def ball_subset_cball unfolding avoid_def by blast - moreover have "isCont f' p" using \p\poles\ poles unfolding f'_def is_pole_def by auto - ultimately show "f' holomorphic_on ball p e1" - apply (elim no_isolated_singularity[rotated]) - apply (auto simp add:continuous_on_eq_continuous_at[of "ball p e1",simplified]) - using field_differentiable_imp_continuous_at f'_holo - holomorphic_on_imp_differentiable_at by fastforce - qed - next - show "f p = 0" using \p\poles\ poles unfolding is_pole_def by auto - next - def p'\"p+e1/2" - have "p'\ball p e1" and "p'\p" using \e1>0\ unfolding p'_def by (auto simp add:dist_norm) - then show "\w\ball p e1. f w \ 0" using e unfolding avoid_def - apply (rule_tac x=p' in bexI) - by (auto simp add:pts_def) + proof (rule base_residue[of "ball p e1" p f,simplified,OF \e1>0\]) + show "inverse \ f holomorphic_on ball p e1" + proof - + define f' where "f' = inverse o f" + have "f holomorphic_on ball p e1 - {p}" + using holo e \poles \ pts\ unfolding avoid_def + apply (elim holomorphic_on_subset) + by auto + then have f'_holo:"f' holomorphic_on ball p e1 - {p}" unfolding f'_def comp_def + apply (elim holomorphic_on_inverse) + using e pts_def ball_subset_cball unfolding avoid_def by blast + moreover have "isCont f' p" using \p\poles\ poles unfolding f'_def is_pole_def by auto + ultimately show "f' holomorphic_on ball p e1" + apply (elim no_isolated_singularity[rotated]) + apply (auto simp add:continuous_on_eq_continuous_at[of "ball p e1",simplified]) + using field_differentiable_imp_continuous_at f'_holo + holomorphic_on_imp_differentiable_at by fastforce qed + next + show "f p = 0" using \p\poles\ poles unfolding is_pole_def by auto + next + define p' where "p' = p+e1/2" + have "p'\ball p e1" and "p'\p" using \e1>0\ unfolding p'_def by (auto simp add:dist_norm) + then show "\w\ball p e1. f w \ 0" using e unfolding avoid_def + apply (rule_tac x=p' in bexI) + by (auto simp add:pts_def) + qed then obtain e2 where e2:"p\poles \ e2>0 \ cball p e2 \ ball p e1 \ contour p e2" by auto - def e3\"if p\poles then e2 else e1" + define e3 where "e3 = (if p\poles then e2 else e1)" have "avoid p e3" using e2 e that avoid_def e3_def by auto moreover have "e3>0" using \e1>0\ e2 unfolding e3_def by auto @@ -2869,33 +2872,33 @@ qed then obtain h where h:"\p\s. h p>0 \ (avoid p (h p) \ (p\poles \ contour p (h p)))" by metis - def cont\"\p. contour_integral (circlepath p (h p)) f" + define cont where "cont p = contour_integral (circlepath p (h p)) f" for p have "contour_integral \ f = (\p\poles. winding_number \ p * cont p)" unfolding cont_def - proof (rule Cauchy_theorem_singularities[OF \open s\ \connected (s-poles)\ _ holo \valid_path \\ - loop \path_image \ \ s-poles\ homo]) - show "finite poles" using \poles \ pts\ and \finite pts\ by (simp add: finite_subset) - next - show "\p\s. 0 < h p \ (\w\cball p (h p). w \ s \ (w \ p \ w \ poles))" - using \poles \ pts\ h unfolding avoid_def by blast - qed + proof (rule Cauchy_theorem_singularities[OF \open s\ \connected (s-poles)\ _ holo \valid_path \\ + loop \path_image \ \ s-poles\ homo]) + show "finite poles" using \poles \ pts\ and \finite pts\ by (simp add: finite_subset) + next + show "\p\s. 0 < h p \ (\w\cball p (h p). w \ s \ (w \ p \ w \ poles))" + using \poles \ pts\ h unfolding avoid_def by blast + qed also have "... = (\p\poles. c * (winding_number \ p * residue f p))" - proof (rule setsum.cong[of poles poles,simplified]) - fix p assume "p \ poles" - show "winding_number \ p * cont p = c * (winding_number \ p * residue f p)" - proof (cases "p\s") - assume "p \ s" - then have "cont p=c*residue f p" - unfolding cont_def - apply (intro contour_integral_unique) - using h[unfolded contour_def] \p \ poles\ by blast - then show ?thesis by auto - next - assume "p\s" - then have "winding_number \ p=0" using homo by auto - then show ?thesis by auto - qed + proof (rule setsum.cong[of poles poles,simplified]) + fix p assume "p \ poles" + show "winding_number \ p * cont p = c * (winding_number \ p * residue f p)" + proof (cases "p\s") + assume "p \ s" + then have "cont p=c*residue f p" + unfolding cont_def + apply (intro contour_integral_unique) + using h[unfolded contour_def] \p \ poles\ by blast + then show ?thesis by auto + next + assume "p\s" + then have "winding_number \ p=0" using homo by auto + then show ?thesis by auto qed + qed also have "... = c * (\p\poles. winding_number \ p * residue f p)" apply (subst setsum_right_distrib) by simp @@ -2921,11 +2924,13 @@ - (\p\poles. winding_number g p * h p * porder f p))" (is "?L=?R") proof - - def c\"2 * complex_of_real pi * \ " - def ff\"(\x. deriv f x * h x / f x)" - def cont_pole\"\ff p e. (ff has_contour_integral - c * porder f p * h p) (circlepath p e)" - def cont_zero\"\ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)" - def avoid\"\p e. \w\cball p e. w \ s \ (w \ p \ w \ pts)" + define c where "c = 2 * complex_of_real pi * \ " + define ff where [abs_def]: "ff x = deriv f x * h x / f x" for x + define cont_pole where "cont_pole ff p e \ + (ff has_contour_integral - c * porder f p * h p) (circlepath p e)" for ff p e + define cont_zero where "cont_zero ff p e \ + (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)" for ff p e + define avoid where "avoid p e \ (\w\cball p e. w \ s \ (w \ p \ w \ pts))" for p e have "poles \ pts" and "zeros \ pts" and "finite zeros" and "pts=zeros \ poles" using poles \finite pts\ unfolding pts_def zeros_def is_pole_def by auto have "\e>0. avoid p e \ (p\poles \ cont_pole ff p e) \ (p\zeros \ cont_zero ff p e)" @@ -2936,10 +2941,10 @@ have "\e2>0. cball p e2 \ ball p e1 \ cont_pole ff p e2" when "p\poles" proof - - def po\"porder f p" - def pp\"pol_poly f p" - def f'\"\w. pp w / (w - p) ^ po" - def ff'\"(\x. deriv f' x * h x / f' x)" + define po where "po = porder f p" + define pp where "pp = pol_poly f p" + define f' where [abs_def]: "f' w = pp w / (w - p) ^ po" for w + define ff' where [abs_def]: "ff' x = deriv f' x * h x / f' x" for x have "inverse \ f holomorphic_on ball p e1" proof - have "f holomorphic_on ball p e1 - {p}" @@ -2962,7 +2967,7 @@ moreover have "f p = 0" using \p\poles\ poles unfolding is_pole_def by auto moreover have "\w\ball p e1. f w \ 0" proof - - def p'\"p+e1/2" + define p' where "p' = p+e1/2" have "p'\ball p e1" and "p'\p" using \e1>0\ unfolding p'_def by (auto simp add:dist_norm) then show "\w\ball p e1. f w \ 0" using e unfolding avoid_def apply (rule_tac x=p' in bexI) @@ -2975,40 +2980,41 @@ pp_po:"(\w\cball p r. f w = pp w / (w - p) ^ po \ pp w \ 0)" using porder_exist[of "ball p e1" p f,simplified,OF \e1>0\] unfolding po_def pp_def by auto - def e2\"r/2" + define e2 where "e2 = r/2" have "e2>0" using \r>0\ unfolding e2_def by auto - def anal\"\w. deriv pp w * h w / pp w" and prin\"\w. - of_nat po * h w / (w - p)" + define anal where [abs_def]: "anal w = deriv pp w * h w / pp w" for w + define prin where [abs_def]: "prin w = - of_nat po * h w / (w - p)" for w have "((\w. prin w + anal w) has_contour_integral - c * po * h p) (circlepath p e2)" - proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) - have "ball p r \ s" - using \cball p r \ ball p e1\ avoid_def ball_subset_cball e by blast - then have "cball p e2 \ s" - using \r>0\ unfolding e2_def by auto - then have "(\w. - of_nat po * h w) holomorphic_on cball p e2" - using h_holo - by (auto intro!: holomorphic_intros) - then show "(prin has_contour_integral - c * of_nat po * h p ) (circlepath p e2)" - using Cauchy_integral_circlepath_simple[folded c_def, of "\w. - of_nat po * h w"] - \e2>0\ - unfolding prin_def - by (auto simp add: mult.assoc) - have "anal holomorphic_on ball p r" unfolding anal_def - using pp_holo h_holo pp_po \ball p r \ s\ - by (auto intro!: holomorphic_intros) - then show "(anal has_contour_integral 0) (circlepath p e2)" - using e2_def \r>0\ - by (auto elim!: Cauchy_theorem_disc_simple) - qed + proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) + have "ball p r \ s" + using \cball p r \ ball p e1\ avoid_def ball_subset_cball e by blast + then have "cball p e2 \ s" + using \r>0\ unfolding e2_def by auto + then have "(\w. - of_nat po * h w) holomorphic_on cball p e2" + using h_holo + by (auto intro!: holomorphic_intros) + then show "(prin has_contour_integral - c * of_nat po * h p ) (circlepath p e2)" + using Cauchy_integral_circlepath_simple[folded c_def, of "\w. - of_nat po * h w"] + \e2>0\ + unfolding prin_def + by (auto simp add: mult.assoc) + have "anal holomorphic_on ball p r" unfolding anal_def + using pp_holo h_holo pp_po \ball p r \ s\ + by (auto intro!: holomorphic_intros) + then show "(anal has_contour_integral 0) (circlepath p e2)" + using e2_def \r>0\ + by (auto elim!: Cauchy_theorem_disc_simple) + qed then have "cont_pole ff' p e2" unfolding cont_pole_def po_def proof (elim has_contour_integral_eq) fix w assume "w \ path_image (circlepath p e2)" then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto - def wp\"w-p" + define wp where "wp = w-p" have "wp\0" and "pp w \0" unfolding wp_def using \w\p\ \w\ball p r\ pp_po by auto moreover have der_f':"deriv f' w = - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po" proof (rule DERIV_imp_deriv) - def der \ "- po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po" + define der where "der = - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po" have po:"po = Suc (po - Suc 0) " using \po>0\ by auto have "(pp has_field_derivative (deriv pp w)) (at w)" using DERIV_deriv_iff_field_differentiable \w\ball p r\ @@ -3067,26 +3073,27 @@ have "\e3>0. cball p e3 \ ball p e1 \ cont_zero ff p e3" when "p\zeros" proof - - def zo\"zorder f p" - def zp\"zer_poly f p" - def f'\"\w. zp w * (w - p) ^ zo" - def ff'\"(\x. deriv f' x * h x / f' x)" + define zo where "zo = zorder f p" + define zp where "zp = zer_poly f p" + define f' where [abs_def]: "f' w = zp w * (w - p) ^ zo" for w + define ff' where [abs_def]: "ff' x = deriv f' x * h x / f' x" for x have "f holomorphic_on ball p e1" - proof - - have "ball p e1 \ s - poles" - using \poles \ pts\ avoid_def ball_subset_cball e that zeros_def by fastforce - thus ?thesis using f_holo by auto - qed + proof - + have "ball p e1 \ s - poles" + using \poles \ pts\ avoid_def ball_subset_cball e that zeros_def by fastforce + thus ?thesis using f_holo by auto + qed moreover have "f p = 0" using \p\zeros\ using DiffD1 mem_Collect_eq pts_def zeros_def by blast moreover have "\w\ball p e1. f w \ 0" - proof - - def p'\"p+e1/2" - have "p'\ball p e1" and "p'\p" using \e1>0\ unfolding p'_def by (auto simp add:dist_norm) - then show "\w\ball p e1. f w \ 0" using e unfolding avoid_def - apply (rule_tac x=p' in bexI) - by (auto simp add:pts_def) - qed + proof - + define p' where "p' = p+e1/2" + have "p'\ball p e1" and "p'\p" + using \e1>0\ unfolding p'_def by (auto simp add:dist_norm) + then show "\w\ball p e1. f w \ 0" using e unfolding avoid_def + apply (rule_tac x=p' in bexI) + by (auto simp add:pts_def) + qed ultimately obtain r where "0 < zo" "r>0" "cball p r \ ball p e1" and @@ -3094,58 +3101,59 @@ pp_po:"(\w\cball p r. f w = zp w * (w - p) ^ zo \ zp w \ 0)" using zorder_exist[of "ball p e1" p f,simplified,OF \e1>0\] unfolding zo_def zp_def by auto - def e2\"r/2" + define e2 where "e2 = r/2" have "e2>0" using \r>0\ unfolding e2_def by auto - def anal\"\w. deriv zp w * h w / zp w" and prin\"\w. of_nat zo * h w / (w - p)" + define anal where [abs_def]: "anal w = deriv zp w * h w / zp w" for w + define prin where [abs_def]: "prin w = of_nat zo * h w / (w - p)" for w have "((\w. prin w + anal w) has_contour_integral c * zo * h p) (circlepath p e2)" - proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) - have "ball p r \ s" - using \cball p r \ ball p e1\ avoid_def ball_subset_cball e by blast - then have "cball p e2 \ s" - using \r>0\ unfolding e2_def by auto - then have "(\w. of_nat zo * h w) holomorphic_on cball p e2" - using h_holo - by (auto intro!: holomorphic_intros) - then show "(prin has_contour_integral c * of_nat zo * h p ) (circlepath p e2)" - using Cauchy_integral_circlepath_simple[folded c_def, of "\w. of_nat zo * h w"] - \e2>0\ - unfolding prin_def - by (auto simp add: mult.assoc) - have "anal holomorphic_on ball p r" unfolding anal_def - using pp_holo h_holo pp_po \ball p r \ s\ - by (auto intro!: holomorphic_intros) - then show "(anal has_contour_integral 0) (circlepath p e2)" - using e2_def \r>0\ - by (auto elim!: Cauchy_theorem_disc_simple) - qed + proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) + have "ball p r \ s" + using \cball p r \ ball p e1\ avoid_def ball_subset_cball e by blast + then have "cball p e2 \ s" + using \r>0\ unfolding e2_def by auto + then have "(\w. of_nat zo * h w) holomorphic_on cball p e2" + using h_holo + by (auto intro!: holomorphic_intros) + then show "(prin has_contour_integral c * of_nat zo * h p ) (circlepath p e2)" + using Cauchy_integral_circlepath_simple[folded c_def, of "\w. of_nat zo * h w"] + \e2>0\ + unfolding prin_def + by (auto simp add: mult.assoc) + have "anal holomorphic_on ball p r" unfolding anal_def + using pp_holo h_holo pp_po \ball p r \ s\ + by (auto intro!: holomorphic_intros) + then show "(anal has_contour_integral 0) (circlepath p e2)" + using e2_def \r>0\ + by (auto elim!: Cauchy_theorem_disc_simple) + qed then have "cont_zero ff' p e2" unfolding cont_zero_def zo_def - proof (elim has_contour_integral_eq) - fix w assume "w \ path_image (circlepath p e2)" - then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto - def wp\"w-p" - have "wp\0" and "zp w \0" - unfolding wp_def using \w\p\ \w\ball p r\ pp_po by auto - moreover have der_f':"deriv f' w = zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo" - proof (rule DERIV_imp_deriv) - def der\"zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo" - have po:"zo = Suc (zo - Suc 0) " using \zo>0\ by auto - have "(zp has_field_derivative (deriv zp w)) (at w)" - using DERIV_deriv_iff_field_differentiable \w\ball p r\ - holomorphic_on_imp_differentiable_at[of _ "ball p r"] - holomorphic_on_subset [OF pp_holo ball_subset_cball] - by (metis open_ball) - then show "(f' has_field_derivative der) (at w)" - using \w\p\ \zo>0\ unfolding der_def f'_def - by (auto intro!: derivative_eq_intros simp add:field_simps) - qed - ultimately show "prin w + anal w = ff' w" - unfolding ff'_def prin_def anal_def - apply simp - apply (unfold f'_def) - apply (fold wp_def) - apply (auto simp add:field_simps) - by (metis Suc_diff_Suc minus_nat.diff_0 power_Suc) - qed + proof (elim has_contour_integral_eq) + fix w assume "w \ path_image (circlepath p e2)" + then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto + define wp where "wp = w-p" + have "wp\0" and "zp w \0" + unfolding wp_def using \w\p\ \w\ball p r\ pp_po by auto + moreover have der_f':"deriv f' w = zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo" + proof (rule DERIV_imp_deriv) + define der where "der = zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo" + have po:"zo = Suc (zo - Suc 0) " using \zo>0\ by auto + have "(zp has_field_derivative (deriv zp w)) (at w)" + using DERIV_deriv_iff_field_differentiable \w\ball p r\ + holomorphic_on_imp_differentiable_at[of _ "ball p r"] + holomorphic_on_subset [OF pp_holo ball_subset_cball] + by (metis open_ball) + then show "(f' has_field_derivative der) (at w)" + using \w\p\ \zo>0\ unfolding der_def f'_def + by (auto intro!: derivative_eq_intros simp add:field_simps) + qed + ultimately show "prin w + anal w = ff' w" + unfolding ff'_def prin_def anal_def + apply simp + apply (unfold f'_def) + apply (fold wp_def) + apply (auto simp add:field_simps) + by (metis Suc_diff_Suc minus_nat.diff_0 power_Suc) + qed then have "cont_zero ff p e2" unfolding cont_zero_def proof (elim has_contour_integral_eq) fix w assume "w \ path_image (circlepath p e2)" @@ -3181,7 +3189,7 @@ qed then obtain e3 where e3:"p\zeros \ e3>0 \ cball p e3 \ ball p e1 \ cont_zero ff p e3" by auto - def e4\"if p\poles then e2 else if p\zeros then e3 else e1" + define e4 where "e4 = (if p\poles then e2 else if p\zeros then e3 else e1)" have "e4>0" using e2 e3 \e1>0\ unfolding e4_def by auto moreover have "avoid p e4" using e2 e3 \e1>0\ e unfolding e4_def avoid_def by auto moreover have "p\poles \ cont_pole ff p e4" and "p\zeros \ cont_zero ff p e4" @@ -3191,19 +3199,19 @@ then obtain get_e where get_e:"\p\s. get_e p>0 \ avoid p (get_e p) \ (p\poles \ cont_pole ff p (get_e p)) \ (p\zeros \ cont_zero ff p (get_e p))" by metis - def cont\"\p. contour_integral (circlepath p (get_e p)) ff" - def w\"\p. winding_number g p" + define cont where "cont p = contour_integral (circlepath p (get_e p)) ff" for p + define w where "w p = winding_number g p" for p have "contour_integral g ff = (\p\pts. w p * cont p)" unfolding cont_def w_def - proof (rule Cauchy_theorem_singularities[OF \open s\ connected finite _ \valid_path g\ loop - path_img homo]) - have "open (s - pts)" using open_Diff[OF _ finite_imp_closed[OF finite]] \open s\ by auto - then show "ff holomorphic_on s - pts" unfolding ff_def using f_holo \poles \ pts\ h_holo - by (auto intro!: holomorphic_intros simp add:pts_def) - next - show "\p\s. 0 < get_e p \ (\w\cball p (get_e p). w \ s \ (w \ p \ w \ pts))" - using get_e using avoid_def by blast - qed + proof (rule Cauchy_theorem_singularities[OF \open s\ connected finite _ \valid_path g\ loop + path_img homo]) + have "open (s - pts)" using open_Diff[OF _ finite_imp_closed[OF finite]] \open s\ by auto + then show "ff holomorphic_on s - pts" unfolding ff_def using f_holo \poles \ pts\ h_holo + by (auto intro!: holomorphic_intros simp add:pts_def) + next + show "\p\s. 0 < get_e p \ (\w\cball p (get_e p). w \ s \ (w \ p \ w \ pts))" + using get_e using avoid_def by blast + qed also have "... = (\p\zeros. w p * cont p) + (\p\poles. w p * cont p)" using \finite pts\ unfolding \pts=zeros \ poles\ apply (subst setsum.union_disjoint) @@ -3309,9 +3317,9 @@ qed then show ?thesis unfolding zeros_f_def using path_img by auto qed - def w\"\p. winding_number \ p" - def c\"2 * complex_of_real pi * \" - def h\"\p. g p / f p + 1" + define w where "w p = winding_number \ p" for p + define c where "c = 2 * complex_of_real pi * \" + define h where [abs_def]: "h p = g p / f p + 1" for p obtain spikes where "finite spikes" and spikes: "\x\{0..1} - spikes. \ differentiable at x" using \valid_path \\ @@ -3366,8 +3374,8 @@ proof (rule vector_derivative_chain_at_general) show "\ differentiable at x" using that \valid_path \\ spikes by auto next - def der\"\p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" - def t\"\ x" + define der where "der p = (deriv g p * f p - g p * deriv f p)/(f p * f p)" for p + define t where "t = \ x" have "f t\0" unfolding zeros_f_def t_def by (metis DiffD1 image_eqI norm_not_less_zero norm_zero path_defs(4) path_less that) moreover have "t\s" @@ -3412,29 +3420,29 @@ have "fg p\0" and "f p\0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def by auto have "h p\0" - proof (rule ccontr) - assume "\ h p \ 0" - then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2) - then have "cmod (g p/f p) = 1" by auto - moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that] - apply (cases "cmod (f p) = 0") - by (auto simp add: norm_divide) - ultimately show False by auto - qed + proof (rule ccontr) + assume "\ h p \ 0" + then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2) + then have "cmod (g p/f p) = 1" by auto + moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that] + apply (cases "cmod (f p) = 0") + by (auto simp add: norm_divide) + ultimately show False by auto + qed have der_fg:"deriv fg p = deriv f p + deriv g p" unfolding fg_def using f_holo g_holo holomorphic_on_imp_differentiable_at[OF _ \open s\] path_img that by (auto intro!: deriv_add) have der_h:"deriv h p = (deriv g p * f p - g p * deriv f p)/(f p * f p)" - proof - - def der\"\p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" - have "p\s" using path_img that by auto - then have "(h has_field_derivative der p) (at p)" - unfolding h_def der_def using g_holo f_holo \f p\0\ - apply (auto intro!: derivative_eq_intros) - by (auto simp add: DERIV_deriv_iff_field_differentiable - holomorphic_on_imp_differentiable_at[OF _ \open s\]) - then show ?thesis unfolding der_def using DERIV_imp_deriv by auto - qed + proof - + define der where "der p = (deriv g p * f p - g p * deriv f p)/(f p * f p)" for p + have "p\s" using path_img that by auto + then have "(h has_field_derivative der p) (at p)" + unfolding h_def der_def using g_holo f_holo \f p\0\ + apply (auto intro!: derivative_eq_intros) + by (auto simp add: DERIV_deriv_iff_field_differentiable + holomorphic_on_imp_differentiable_at[OF _ \open s\]) + then show ?thesis unfolding der_def using DERIV_imp_deriv by auto + qed show ?thesis apply (simp only:der_fg der_h) apply (auto simp add:field_simps \h p\0\ \f p\0\ \fg p\0\) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Mon Apr 25 22:59:53 2016 +0200 @@ -105,7 +105,7 @@ proof - { fix x :: "'n::euclidean_space" - def y \ "(e / norm x) *\<^sub>R x" + define y where "y = (e / norm x) *\<^sub>R x" then have "y \ cball 0 e" using assms by auto moreover have *: "x = (norm x / e) *\<^sub>R y" @@ -463,7 +463,7 @@ fix s u assume as: "\x\V. \y\V. \u v. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ V" "finite s" "s \ {}" "s \ V" "setsum u s = (1::real)" - def n \ "card s" + define n where "n = card s" have "card s = 0 \ card s = 1 \ card s = 2 \ card s > 2" by auto then show "(\x\s. u x *\<^sub>R x) \ V" proof (auto simp only: disjE) @@ -838,7 +838,7 @@ assume "y = a + v" "v \ span {x - a |x. x \ s}" then obtain t u where obt: "finite t" "t \ {x - a |x. x \ s}" "a + (\v\t. u v *\<^sub>R v) = y" unfolding span_explicit by auto - def f \ "(\x. x + a) ` t" + define f where "f = (\x. x + a) ` t" have f: "finite f" "f \ s" "(\v\f. u (v - a) *\<^sub>R (v - a)) = y - a" unfolding f_def using obt by (auto simp add: setsum.reindex[unfolded inj_on_def]) have *: "f \ {a} = {}" "f \ - {a} = f" @@ -989,7 +989,7 @@ { fix x y assume xy: "x \ S" "y \ S" - def u == "(1 :: real)/2" + define u where "u = (1 :: real)/2" have "(1/2) *\<^sub>R (x+y) = (1/2) *\<^sub>R (x+y)" by auto moreover @@ -1370,7 +1370,7 @@ moreover assume "A \ s \ {}" "B \ s \ {}" then obtain a b where a: "a \ A" "a \ s" and b: "b \ B" "b \ s" by auto - def f \ "\u. u *\<^sub>R a + (1 - u) *\<^sub>R b" + define f where [abs_def]: "f u = u *\<^sub>R a + (1 - u) *\<^sub>R b" for u then have "continuous_on {0 .. 1} f" by (auto intro!: continuous_intros) then have "connected (f ` {0 .. 1})" @@ -1802,7 +1802,7 @@ using uv(1) x(1)[THEN bspec[where x=i]] by auto next case False - def j \ "i - k1" + define j where "j = i - k1" from i False have "j \ {1..k2}" unfolding j_def by auto then show ?thesis @@ -2149,7 +2149,7 @@ from assms(1)[unfolded dependent_explicit] obtain S u v where obt: "finite S" "S \ {x - a |x. x \ s}" "v\S" "u v \ 0" "(\v\S. u v *\<^sub>R v) = 0" by auto - def t \ "(\x. x + a) ` S" + define t where "t = (\x. x + a) ` S" have inj: "inj_on (\x. x + a) S" unfolding inj_on_def by auto @@ -2469,7 +2469,7 @@ "(\x. -a+x) ` (S - {a}) \ B \ B \ (\x. -a+x) ` V \ independent B \ (\x. -a+x) ` V \ span B" using maximal_independent_subset_extend[of "(\x. -a+x) ` (S-{a})" "(\x. -a + x) ` V"] assms by blast - def T \ "(\x. a+x) ` insert 0 B" + define T where "T = (\x. a+x) ` insert 0 B" then have "T = insert a ((\x. a+x) ` B)" by auto then have "affine hull T = (\x. a+x) ` span B" @@ -2601,7 +2601,7 @@ using assms B affine_hull_nonempty[of V] affine_hull_nonempty[of B] by auto then obtain a where a: "a \ B" by auto - def Lb \ "span ((\x. -a+x) ` (B-{a}))" + define Lb where "Lb = span ((\x. -a+x) ` (B-{a}))" moreover have "affine_parallel (affine hull B) Lb" using Lb_def B assms affine_hull_span2[of a B] a affine_parallel_commut[of "Lb" "(affine hull B)"] @@ -2765,7 +2765,7 @@ next case False then obtain a where a: "a \ B" by auto - def Lb \ "span ((\x. -a+x) ` (B-{a}))" + define Lb where "Lb = span ((\x. -a+x) ` (B-{a}))" have "affine_parallel (affine hull B) Lb" using Lb_def affine_hull_span2[of a B] a affine_parallel_commut[of "Lb" "(affine hull B)"] @@ -3002,13 +3002,13 @@ proof - obtain a where "a \ S" using assms by auto then have "a \ T" using assms by auto - def LS \ "{y. \x \ S. (-a) + x = y}" + define LS where "LS = {y. \x \ S. (-a) + x = y}" then have ls: "subspace LS" "affine_parallel S LS" using assms parallel_subspace_explicit[of S a LS] \a \ S\ by auto then have h1: "int(dim LS) = aff_dim S" using assms aff_dim_affine[of S LS] by auto have "T \ {}" using assms by auto - def LT \ "{y. \x \ T. (-a) + x = y}" + define LT where "LT = {y. \x \ T. (-a) + x = y}" then have lt: "subspace LT \ affine_parallel T LT" using assms parallel_subspace_explicit[of T a LT] \a \ T\ by auto then have "int(dim LT) = aff_dim T" @@ -3062,7 +3062,7 @@ where a1 [simp]: "setsum a t = 1" and [simp]: "setsum (\v. a v *\<^sub>R v) t = y" and [simp]: "setsum b u = 1" "setsum (\v. b v *\<^sub>R v) u = y" by (auto simp: affine_hull_finite \finite t\ \finite u\) - def c \ "\x. if x \ t then a x else if x \ u then -(b x) else 0" + define c where "c x = (if x \ t then a x else if x \ u then -(b x) else 0)" for x have [simp]: "s \ t = t" "s \ - t \ u = u" using assms by auto have "setsum c s = 0" by (simp add: c_def comm_monoid_add_class.setsum.If_cases \finite s\ setsum_negf) @@ -3172,8 +3172,8 @@ by blast then obtain w v where wv: "setsum w s = 0" "v\s" "w v \ 0" "(\v\s. w v *\<^sub>R v) = 0" using affine_dependent_explicit_finite[OF obt(1)] by auto - def i \ "(\v. (u v) / (- w v)) ` {v\s. w v < 0}" - def t \ "Min i" + define i where "i = (\v. (u v) / (- w v)) ` {v\s. w v < 0}" + define t where "t = Min i" have "\x\s. w x < 0" proof (rule ccontr, simp add: not_less) assume as:"\x\s. 0 \ w x" @@ -3428,7 +3428,7 @@ { fix x assume x: "x \ affine hull S" - def e \ "1::real" + define e :: real where "e = 1" then have "e > 0" "ball x e \ affine hull (affine hull S) \ affine hull S" using hull_hull[of _ S] by auto then have "x \ rel_interior (affine hull S)" @@ -3727,7 +3727,7 @@ qed then obtain y where "y \ S" and y: "norm (y - x) * (1 - e) < e * d" by auto - def z \ "c + ((1 - e) / e) *\<^sub>R (x - y)" + define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) @@ -3846,10 +3846,10 @@ using z rel_interior_cball[of "f ` S"] by auto obtain K where K: "K > 0" "\x. norm (f x) \ norm x * K" using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto - def e1 \ "1 / K" + define e1 where "e1 = 1 / K" then have e1: "e1 > 0" "\x. e1 * norm (f x) \ norm x" using K pos_le_divide_eq[of e1] by auto - def e \ "e1 * e2" + define e where "e = e1 * e2" then have "e > 0" using e1 e2 by auto { fix y @@ -3892,7 +3892,7 @@ using assms injective_imp_isometric[of "span S" f] subspace_span[of S] closed_subspace[of "span S"] by auto - def e \ "e1 * e2" + define e where "e = e1 * e2" hence "e > 0" using e1 e2 by auto { fix y @@ -3969,7 +3969,7 @@ using bchoice[of s "\x e. e > 0 \ cball x e \ s"] by auto have "b ` t \ {}" using obt by auto - def i \ "b ` t" + define i where "i = b ` t" show "\e > 0. cball a e \ {y. \sa u. finite sa \ sa \ s \ (\x\sa. 0 \ u x) \ setsum u sa = 1 \ (\v\sa. u v *\<^sub>R v) = y}" @@ -4726,7 +4726,7 @@ apply (erule_tac x="x - y" in ballE) apply (auto simp add: inner_diff) done - def k \ "SUP x:t. a \ x" + define k where "k = (SUP x:t. a \ x)" show ?thesis apply (rule_tac x="-a" in exI) apply (rule_tac x="-(k + b / 2)" in exI) @@ -5064,7 +5064,7 @@ using radon_ex_lemma[OF assms] by auto have fin: "finite {x \ c. 0 < u x}" "finite {x \ c. 0 > u x}" using assms(1) by auto - def z \ "inverse (setsum u {x\c. u x > 0}) *\<^sub>R setsum (\x. u x *\<^sub>R x) {x\c. u x > 0}" + define z where "z = inverse (setsum u {x\c. u x > 0}) *\<^sub>R setsum (\x. u x *\<^sub>R x) {x\c. u x > 0}" have "setsum u {x \ c. 0 < u x} \ 0" proof (cases "u v \ 0") case False @@ -5356,7 +5356,7 @@ using compact_imp_closed[OF assms(1)] apply simp done - def pi \ "\x::'a. inverse (norm x) *\<^sub>R x" + define pi where [abs_def]: "pi x = inverse (norm x) *\<^sub>R x" for x :: 'a have "0 \ frontier s" unfolding frontier_straddle apply (rule notI) @@ -5373,7 +5373,7 @@ apply (intro continuous_intros) apply simp done - def sphere \ "{x::'a. norm x = 1}" + define sphere :: "'a set" where "sphere = {x. norm x = 1}" have pi: "\x. x \ 0 \ pi x \ sphere" "\x u. u>0 \ pi (u *\<^sub>R x) = pi x" unfolding pi_def sphere_def by auto @@ -6292,7 +6292,7 @@ proof (rule,rule,rule) fix x and e :: real assume "x \ s" "e > 0" - def B \ "\b\ + 1" + define B where "B = \b\ + 1" have B: "0 < B" "\x. x\s \ \f x\ \ B" unfolding B_def defer @@ -6312,7 +6312,7 @@ show "\f y - f x\ < e" proof (cases "y = x") case False - def t \ "k / norm (y - x)" + define t where "t = k / norm (y - x)" have "2 < t" "0k>0\ by (auto simp add:field_simps) @@ -6323,7 +6323,7 @@ using as by (auto simp add: field_simps norm_minus_commute) { - def w \ "x + t *\<^sub>R (y - x)" + define w where "w = x + t *\<^sub>R (y - x)" have "w \ s" unfolding w_def apply (rule k[unfolded subset_eq,rule_format]) @@ -6356,7 +6356,7 @@ } moreover { - def w \ "x - t *\<^sub>R (y - x)" + define w where "w = x - t *\<^sub>R (y - x)" have "w \ s" unfolding w_def apply (rule k[unfolded subset_eq,rule_format]) @@ -6408,7 +6408,7 @@ case True fix y assume y: "y \ cball x e" - def z \ "2 *\<^sub>R x - y" + define z where "z = 2 *\<^sub>R x - y" have *: "x - (2 *\<^sub>R x - y) = y - x" by (simp add: scaleR_2) have z: "z \ cball x e" @@ -6445,7 +6445,7 @@ assume "x \ s" then obtain e where e: "cball x e \ s" "e > 0" using assms(1) unfolding open_contains_cball by auto - def d \ "e / real DIM('a)" + define d where "d = e / real DIM('a)" have "0 < d" unfolding d_def using \e > 0\ dimge1 by auto let ?d = "(\i\Basis. d *\<^sub>R i)::'a" @@ -6454,7 +6454,7 @@ and c1: "convex hull c \ cball x e" and c2: "cball x d \ convex hull c" proof - def c \ "\i\Basis. (\a. a *\<^sub>R i) ` {x\i - d, x\i + d}" + define c where "c = (\i\Basis. (\a. a *\<^sub>R i) ` {x\i - d, x\i + d})" show "finite c" unfolding c_def by (simp add: finite_set_setsum) have 1: "convex hull c = {a. \i\Basis. a \ i \ cbox (x \ i - d) (x \ i + d)}" @@ -6490,7 +6490,7 @@ apply simp done qed - def k \ "Max (f ` c)" + define k where "k = Max (f ` c)" have "convex_on (convex hull c) f" apply(rule convex_on_subset[OF assms(2)]) apply(rule subset_trans[OF _ e(1)]) @@ -7302,7 +7302,7 @@ qed then obtain y where "y \ s" and y: "norm (y - x) * (1 - e) < e * d" by auto - def z \ "c + ((1 - e) / e) *\<^sub>R (x - y)" + define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) @@ -7908,7 +7908,7 @@ { fix e :: real assume "e > 0" - def e1 \ "min 1 (e/norm (x - a))" + define e1 where "e1 = min 1 (e/norm (x - a))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (x - a) \ e" using \x \ a\ \e > 0\ le_divide_eq[of e1 e "norm (x - a)"] by simp_all @@ -7973,7 +7973,7 @@ assumes "0 < a" "0 < b" "(a + b) *\<^sub>R z = a *\<^sub>R x + b *\<^sub>R y" obtains e where "0 < e" "e \ 1" "z = y - e *\<^sub>R (y - x)" proof - - def e \ "a / (a + b)" + define e where "e = a / (a + b)" have "z = (1 / (a + b)) *\<^sub>R ((a + b) *\<^sub>R z)" apply auto using assms @@ -8023,7 +8023,7 @@ obtain e where e: "e > 0" "cball z e \ affine hull closure S \ closure S" using z rel_interior_cball[of "closure S"] by auto hence *: "0 < e/norm(z-x)" using e False by auto - def y \ "z + (e/norm(z-x)) *\<^sub>R (z-x)" + define y where "y = z + (e/norm(z-x)) *\<^sub>R (z-x)" have yball: "y \ cball z e" using mem_cball y_def dist_norm[of z y] e by auto have "x \ affine hull closure S" @@ -8239,7 +8239,7 @@ assume x: "x \ affine hull S" { assume "x \ z" - def m \ "1 + e1/norm(x-z)" + define m where "m = 1 + e1/norm(x-z)" hence "m > 1" using e1 \x \ z\ by auto { fix e @@ -8274,7 +8274,7 @@ moreover { assume "x = z" - def m \ "1 + e1" + define m where "m = 1 + e1" then have "m > 1" using e1 by auto { @@ -8314,9 +8314,9 @@ using rel_interior_subset by auto then obtain e where e: "e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using assms by auto - def y \ "(1 - e) *\<^sub>R x + e *\<^sub>R z" + define y where [abs_def]: "y = (1 - e) *\<^sub>R x + e *\<^sub>R z" then have "y \ S" using e by auto - def e1 \ "1/e" + define e1 where "e1 = 1/e" then have "0 < e1 \ e1 < 1" using e by auto then have "z =y - (1 - e1) *\<^sub>R (y - x)" using e1_def y_def by (auto simp add: algebra_simps) @@ -8362,10 +8362,10 @@ using r by auto obtain e2 where e2: "e2 > 0 \ z + e2 *\<^sub>R (z - x) \ S" using r by auto - def x1 \ "z + e1 *\<^sub>R (x - z)" + define x1 where [abs_def]: "x1 = z + e1 *\<^sub>R (x - z)" then have x1: "x1 \ affine hull S" using e1 hull_subset[of S] by auto - def x2 \ "z + e2 *\<^sub>R (z - x)" + define x2 where [abs_def]: "x2 = z + e2 *\<^sub>R (z - x)" then have x2: "x2 \ affine hull S" using e2 hull_subset[of S] by auto have *: "e1/(e1+e2) + e2/(e1+e2) = 1" @@ -8412,7 +8412,7 @@ fix x obtain e1 where e1: "e1 > 1" "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z \ S" using **[rule_format, of "z-x"] by auto - def e \ "e1 - 1" + define e where [abs_def]: "e = e1 - 1" then have "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z = z + e *\<^sub>R x" by (simp add: algebra_simps) then have "e > 0" "z + e *\<^sub>R x \ S" @@ -8427,7 +8427,7 @@ fix x obtain e1 where e1: "e1 > 0" "z + e1 *\<^sub>R (z - x) \ S" using r[rule_format, of "z-x"] by auto - def e \ "e1 + 1" + define e where "e = e1 + 1" then have "z + e1 *\<^sub>R (z - x) = (1 - e) *\<^sub>R x + e *\<^sub>R z" by (simp add: algebra_simps) then have "e > 1" "(1 - e)*\<^sub>R x + e *\<^sub>R z \ S" @@ -8508,11 +8508,11 @@ assume "y \ x" { fix e :: real assume e: "e > 0" - def e1 \ "min 1 (e/norm (y - x))" + define e1 where "e1 = min 1 (e/norm (y - x))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (y - x) \ e" using \y \ x\ \e > 0\ le_divide_eq[of e1 e "norm (y - x)"] by simp_all - def z \ "y - e1 *\<^sub>R (y - x)" + define z where "z = y - e1 *\<^sub>R (y - x)" { fix S assume "S \ I" @@ -8624,7 +8624,7 @@ } then obtain mS where mS: "\S\I. mS S > 1 \ (\e. e > 1 \ e \ mS S \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" by metis - def e \ "Min (mS ` I)" + define e where "e = Min (mS ` I)" then have "e \ mS ` I" using assms \I \ {}\ by simp then have "e > 1" using mS by auto moreover have "\S\I. e \ mS S" @@ -9112,7 +9112,7 @@ have conv: "convex ({(1 :: real)} \ S)" using convex_Times[of "{(1 :: real)}" S] assms convex_singleton[of "1 :: real"] by auto - def f \ "\y. {z. (y, z) \ cone hull ({1 :: real} \ S)}" + define f where "f y = {z. (y, z) \ cone hull ({1 :: real} \ S)}" for y then have *: "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) = (c \ rel_interior {y. f y \ {}} \ x \ rel_interior (f c))" apply (subst rel_interior_projection[of "cone hull ({(1 :: real)} \ S)" f c x]) @@ -9210,11 +9210,11 @@ { fix x assume "x \ S i" - def c \ "\j. if j = i then 1::real else 0" + define c where "c j = (if j = i then 1::real else 0)" for j then have *: "setsum c I = 1" using \finite I\ \i \ I\ setsum.delta[of I i "\j::'a. 1::real"] by auto - def s \ "\j. if j = i then x else p j" + define s where "s j = (if j = i then x else p j)" for j then have "\j. c j *\<^sub>R s j = (if j = i then x else 0)" using c_def by (auto simp add: algebra_simps) then have "x = setsum (\i. c i *\<^sub>R s i) I" @@ -9243,7 +9243,7 @@ from xy obtain d t where yc: "y = setsum (\i. d i *\<^sub>R t i) I \ (\i\I. d i \ 0) \ setsum d I = 1 \ (\i\I. t i \ S i)" by auto - def e \ "\i. u * c i + v * d i" + define e where "e i = u * c i + v * d i" for i have ge0: "\i\I. e i \ 0" using e_def xc yc uv by simp have "setsum (\i. u * c i) I = u * setsum c I" @@ -9252,7 +9252,8 @@ by (simp add: setsum_right_distrib) ultimately have sum1: "setsum e I = 1" using e_def xc yc uv by (simp add: setsum.distrib) - def q \ "\i. if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i" + define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)" + for i { fix i assume i: "i \ I" @@ -9318,8 +9319,8 @@ {u *\<^sub>R s + v *\<^sub>R t | u v s t. u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T}" (is "?lhs = ?rhs") proof - def I \ "{1::nat, 2}" - def s \ "\i. if i = (1::nat) then S else T" + define I :: "nat set" where "I = {1, 2}" + define s where "s i = (if i = (1::nat) then S else T)" for i have "\(s ` I) = S \ T" using s_def I_def by auto then have "convex hull (\(s ` I)) = convex hull (S \ T)" @@ -9414,7 +9415,7 @@ then obtain c xs where x: "x = setsum (\i. c i *\<^sub>R xs i) I \ (\i\I. c i \ 0) \ setsum c I = 1 \ (\i\I. xs i \ S i)" using convex_hull_finite_union[of I S] assms by auto - def s \ "\i. c i *\<^sub>R xs i" + define s where "s i = c i *\<^sub>R xs i" for i { fix i assume "i \ I" @@ -9432,7 +9433,7 @@ assume "x \ ?rhs" then obtain s where x: "x = setsum s I \ (\i\I. s i \ S i)" using set_setsum_alt[of I S] assms by auto - def xs \ "\i. of_nat(card I) *\<^sub>R s i" + define xs where "xs i = of_nat(card I) *\<^sub>R s i" for i then have "x = setsum (\i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I" using x assms by auto moreover have "\i\I. xs i \ S i" @@ -9465,8 +9466,8 @@ and "T \ {}" shows "convex hull (S \ T) = S + T" proof - - def I \ "{1::nat, 2}" - def A \ "(\i. if i = (1::nat) then S else T)" + define I :: "nat set" where "I = {1, 2}" + define A where "A i = (if i = (1::nat) then S else T)" for i have "\(A ` I) = S \ T" using A_def I_def by auto then have "convex hull (\(A ` I)) = convex hull (S \ T)" @@ -9500,11 +9501,11 @@ using convex_hull_empty rel_interior_empty by auto next case False - def C0 \ "convex hull (\(S ` I))" + define C0 where "C0 = convex hull (\(S ` I))" have "\i\I. C0 \ S i" unfolding C0_def using hull_subset[of "\(S ` I)"] by auto - def K0 \ "cone hull ({1 :: real} \ C0)" - def K \ "\i. cone hull ({1 :: real} \ S i)" + define K0 where "K0 = cone hull ({1 :: real} \ C0)" + define K where "K i = cone hull ({1 :: real} \ S i)" for i have "\i\I. K i \ {}" unfolding K_def using assms by (simp add: cone_hull_empty_iff[symmetric]) @@ -9619,7 +9620,7 @@ then obtain c s where cs: "x = setsum (\i. c i *\<^sub>R s i) I \ (\i\I. c i > 0) \ setsum c I = 1 \ (\i\I. s i \ rel_interior (S i))" by auto - def k \ "\i. (c i, c i *\<^sub>R s i)" + define k where "k i = (c i, c i *\<^sub>R s i)" for i { fix i assume "i:I" then have "k i \ rel_interior (K i)" @@ -9653,7 +9654,7 @@ have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . - moreover def t \ "min (x + e / 2) ((x + y) / 2)" + moreover define t where "t = min (x + e / 2) ((x + y) / 2)" ultimately have "x < t" "t < y" "t \ ball x e" by (auto simp: dist_real_def field_simps split: split_min) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto @@ -9661,7 +9662,7 @@ have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where "0 < e" "ball x e \ interior I" . - moreover def K \ "x - e / 2" + moreover define K where "K = x - e / 2" with \0 < e\ have "K \ ball x e" "K < x" by (auto simp: dist_real_def) ultimately have "K \ I" "K < x" "x \ I" @@ -9688,7 +9689,7 @@ have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . - moreover def t \ "x + e / 2" + moreover define t where "t = x + e / 2" ultimately have "x < t" "t \ ball x e" by (auto simp: dist_real_def field_simps) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto @@ -10767,7 +10768,7 @@ by blast next case False - def k \ "min (1/2) (e / norm (x-a))" + define k where "k = min (1/2) (e / norm (x-a))" have k: "0 < k" "k < 1" using \e > 0\ False by (auto simp: k_def) then have xa: "(x-a) = inverse k *\<^sub>R k *\<^sub>R (x-a)" @@ -11233,7 +11234,7 @@ assumes "convex S" "S \ {}" "0 \ S" obtains a where "a \ span S" "a \ 0" "\x. x \ S \ 0 \ a \ x" proof - - def k \ "\c::'a. {x. 0 \ c \ x}" + define k where [abs_def]: "k c = {x. 0 \ c \ x}" for c :: 'a have *: "span S \ frontier (cball 0 1) \ \f' \ {}" if f': "finite f'" "f' \ k ` S" for f' proof - @@ -11346,7 +11347,7 @@ proof - obtain e where "0 < e" "y \ S" and e: "cball y e \ affine hull S \ S" using \y \ rel_interior S\ by (force simp: rel_interior_cball) - def y' \ "y - (e / norm a) *\<^sub>R ((x + a) - x)" + define y' where "y' = y - (e / norm a) *\<^sub>R ((x + a) - x)" have "y' \ cball y e" unfolding y'_def using \0 < e\ by force moreover have "y' \ affine hull S" @@ -11392,7 +11393,7 @@ "setsum v t = 1" and eq: "(\x\t. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x)" for u v proof - - def f \ "\x. (if x \ s then u x else 0) - (if x \ t then v x else 0)" + define f where "f x = (if x \ s then u x else 0) - (if x \ t then v x else 0)" for x have "setsum f (s \ t) = 0" apply (simp add: f_def setsum_Un setsum_subtractf) apply (simp add: setsum.inter_restrict [symmetric] Int_commute) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Mon Apr 25 22:59:53 2016 +0200 @@ -379,7 +379,7 @@ proof (rule, rule ccontr) fix i :: 'a assume i: "i \ Basis" - def e \ "norm (f' i - f'' i)" + define e where "e = norm (f' i - f'' i)" assume "f' i \ f'' i" then have "e > 0" unfolding e_def by auto @@ -613,7 +613,7 @@ then show ?thesis by (metis c(2) d(2) box_subset_cbox subset_iff) next - def e \ "(a + b) /2" + define e where "e = (a + b) /2" case False then have "f d = f c" using d c assms(2) by auto @@ -781,9 +781,10 @@ note phi_tendsto = assms(3)[simplified continuous_on_def, rule_format] { fix e::real assume "e > 0" - def e2 \ "e / 2" with \e > 0\ have "e2 > 0" by simp + define e2 where "e2 = e / 2" + with \e > 0\ have "e2 > 0" by simp let ?le = "\x1. norm (f x1 - f a) \ \ x1 - \ a + e * (x1 - a) + e" - def A \ "{x2. a \ x2 \ x2 \ b \ (\x1\{a ..< x2}. ?le x1)}" + define A where "A = {x2. a \ x2 \ x2 \ b \ (\x1\{a ..< x2}. ?le x1)}" have A_subset: "A \ {a .. b}" by (auto simp: A_def) { fix x2 @@ -817,7 +818,7 @@ have A_ivl: "\x1 x2. x2 \ A \ x1 \ {a ..x2} \ x1 \ A" by (simp add: A_def) have [simp]: "bdd_above A" by (auto simp: A_def) - def y \ "Sup A" + define y where "y = Sup A" have "y \ b" unfolding y_def by (simp add: cSup_le_iff) (simp add: A_def) @@ -878,7 +879,7 @@ by metis from \open S\ obtain d where d: "\x. dist x y < d \ x \ S" "d > 0" by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \y \ S\]) - def d' \ "min ((y + b)/2) (y + (d/2))" + define d' where "d' = min ((y + b)/2) (y + (d/2))" have "d' \ A" unfolding A_def proof safe @@ -942,7 +943,7 @@ by metis from \open S\ obtain d where d: "\x. dist x y < d \ x \ S" "d > 0" by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \y \ S\]) - def d' \ "min b (y + (d/2))" + define d' where "d' = min b (y + (d/2))" have "d' \ A" unfolding A_def proof safe @@ -1077,7 +1078,7 @@ assumes "x0 \ S" shows "norm (f b - f a - f' x0 (b - a)) \ norm (b - a) * B" proof - - def g \ "\x. f x - f' x0 x" + define g where [abs_def]: "g x = f x - f' x0 x" for x have g: "\x. x \ S \ (g has_derivative (\i. f' x i - f' x0 i)) (at x within S)" unfolding g_def using assms by (auto intro!: derivative_eq_intros @@ -1232,7 +1233,7 @@ "0 < d" "\z. norm (z - y) < d \ norm (g z - g y - g' (z - y)) \ 1 / 2 * norm (g z - g y)" using lem1 * by blast - def B \ "C * 2" + define B where "B = C * 2" have "B > 0" unfolding B_def using C by auto have lem2: "norm (g z - g y) \ B * norm (z - y)" if z: "norm(z - y) < d" for z @@ -1722,7 +1723,7 @@ then have *: "0 < onorm g'" unfolding onorm_pos_lt[OF assms(3)] by fastforce - def k \ "1 / onorm g' / 2" + define k where "k = 1 / onorm g' / 2" have *: "k > 0" unfolding k_def using * by auto obtain d1 where d1: @@ -1749,7 +1750,7 @@ proof (intro strip) fix x y assume as: "x \ ball a d" "y \ ball a d" "f x = f y" - def ph \ "\w. w - g' (f w - f x)" + define ph where [abs_def]: "ph w = w - g' (f w - f x)" for w have ph':"ph = g' \ (\w. f' a w - (f w - f x))" unfolding ph_def o_def unfolding diff @@ -2207,7 +2208,7 @@ shows "summable (\n. f n x)" "((\x. \n. f n x) has_field_derivative (\n. f' n x)) (at x)" proof - from \x \ interior s\ have "x \ s" using interior_subset by blast - def g' \ "\x. \i. f' i x" + define g' where [abs_def]: "g' x = (\i. f' i x)" for x from assms(3) have "uniform_limit s (\n x. \i(x, y). f x y) has_derivative (\(tx, ty). fx x y tx + fy x y ty)) (at (x, y) within X \ Y)" proof (safe intro!: has_derivativeI tendstoI, goal_cases) case (2 e') - def e\"e' / 9" + define e where "e = e' / 9" have "e > 0" using \e' > 0\ by (simp add: e_def) have "(x, y) \ X \ Y" using assms by auto @@ -2720,7 +2721,7 @@ for x' y' using \0 < e\ by (cases "(x', y') = (x, y)") auto - def d \ "d' / sqrt 2" + define d where "d = d' / sqrt 2" have "d > 0" using \0 < d'\ by (simp add: d_def) have d: "x' \ X \ y' \ Y \ dist x' x < d \ dist y' y < d \ dist (fy x' y') (fy x y) < e" for x' y' diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Mon Apr 25 22:59:53 2016 +0200 @@ -119,7 +119,7 @@ assume "open {x}" then obtain e where "0 < e" and e: "\y. dist y x < e \ y = x" unfolding open_dist by fast - def y \ "x + scaleR (e/2) (SOME b. b \ Basis)" + define y where "y = x + scaleR (e/2) (SOME b. b \ Basis)" have [simp]: "(SOME b. b \ Basis) \ Basis" by (rule someI_ex) (auto simp: ex_in_conv) from \0 < e\ have "y \ x" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Fashoda.thy --- a/src/HOL/Multivariate_Analysis/Fashoda.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Fashoda.thy Mon Apr 25 22:59:53 2016 +0200 @@ -102,8 +102,10 @@ proof (rule ccontr) assume "\ ?thesis" note as = this[unfolded bex_simps,rule_format] - def sqprojection \ "\z::real^2. (inverse (infnorm z)) *\<^sub>R z" - def negatex \ "\x::real^2. (vector [-(x$1), x$2])::real^2" + define sqprojection + where [abs_def]: "sqprojection z = (inverse (infnorm z)) *\<^sub>R z" for z :: "real^2" + define negatex :: "real^2 \ real^2" + where "negatex x = (vector [-(x$1), x$2])" for x have lem1: "\z::real^2. infnorm (negatex z) = infnorm z" unfolding negatex_def infnorm_2 vector_2 by auto have lem2: "\z. z \ 0 \ infnorm (sqprojection z) = 1" @@ -349,7 +351,7 @@ obtains z where "z \ path_image f" and "z \ path_image g" proof - note assms=assms[unfolded path_def pathstart_def pathfinish_def path_image_def] - def iscale \ "\z::real. inverse 2 *\<^sub>R (z + 1)" + define iscale where [abs_def]: "iscale z = inverse 2 *\<^sub>R (z + 1)" for z :: real have isc: "iscale ` {- 1..1} \ {0..1}" unfolding iscale_def by auto have "\s\{- 1..1}. \t\{- 1..1}. (f \ iscale) s = (g \ iscale) t" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy --- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Mon Apr 25 22:59:53 2016 +0200 @@ -328,12 +328,12 @@ fix x assume "x \ S" then obtain e where "0 < e" and S: "\y. dist y x < e \ y \ S" using * by fast - def r \ "\i::'b. e / sqrt (of_nat CARD('b))" + define r where [abs_def]: "r i = e / sqrt (of_nat CARD('b))" for i :: 'b from \0 < e\ have r: "\i. 0 < r i" unfolding r_def by simp_all from \0 < e\ have e: "e = setL2 r UNIV" unfolding r_def by (simp add: setL2_constant) - def A \ "\i. {y. dist (x $ i) y < r i}" + define A where "A i = {y. dist (x $ i) y < r i}" for i have "\i. open (A i) \ x $ i \ A i" unfolding A_def by (simp add: open_ball r) moreover have "\y. (\i. y $ i \ A i) \ y \ S" @@ -360,8 +360,8 @@ proof (rule metric_CauchyI) fix r :: real assume "0 < r" hence "0 < r / of_nat CARD('n)" (is "0 < ?s") by simp - def N \ "\i. LEAST N. \m\N. \n\N. dist (X m $ i) (X n $ i) < ?s" - def M \ "Max (range N)" + define N where "N i = (LEAST N. \m\N. \n\N. dist (X m $ i) (X n $ i) < ?s)" for i + define M where "M = Max (range N)" have "\i. \N. \m\N. \n\N. dist (X m $ i) (X n $ i) < ?s" using X \0 < ?s\ by (rule metric_CauchyD) hence "\i. \m\N i. \n\N i. dist (X m $ i) (X n $ i) < ?s" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Gamma.thy --- a/src/HOL/Multivariate_Analysis/Gamma.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Gamma.thy Mon Apr 25 22:59:53 2016 +0200 @@ -344,8 +344,8 @@ shows "uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n :: complex)" proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI') fix e :: real assume e: "e > 0" - def e'' \ "SUP t:ball z d. norm t + norm t^2" - def e' \ "e / (2*e'')" + define e'' where "e'' = (SUP t:ball z d. norm t + norm t^2)" + define e' where "e' = e / (2*e'')" have "bounded ((\t. norm t + norm t^2) ` cball z d)" by (intro compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros) hence "bounded ((\t. norm t + norm t^2) ` ball z d)" by (rule bounded_subset) auto @@ -362,7 +362,7 @@ by (rule inverse_power_summable) simp from summable_partial_sum_bound[OF this e'] guess M . note M = this - def N \ "max 2 (max (nat \2 * (norm z + d)\) M)" + define N where "N = max 2 (max (nat \2 * (norm z + d)\) M)" { from d have "\2 * (cmod z + d)\ \ \0::real\" by (intro ceiling_mono mult_nonneg_nonneg add_nonneg_nonneg) simp_all @@ -434,8 +434,8 @@ assumes z: "(z :: complex) \ \\<^sub>\\<^sub>0" shows "\d>0. uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n)" proof - - def d' \ "Re z" - def d \ "if d' > 0 then d' / 2 else norm (z - of_int (round d')) / 2" + define d' where "d' = Re z" + define d where "d = (if d' > 0 then d' / 2 else norm (z - of_int (round d')) / 2)" have "of_int (round d') \ \\<^sub>\\<^sub>0" if "d' \ 0" using that by (intro nonpos_Ints_of_int) (simp_all add: round_def) with assms have d_pos: "d > 0" unfolding d_def by (force simp: not_less) @@ -633,8 +633,8 @@ assumes z: "z \ 0" and n: "n \ 2" shows "uniformly_convergent_on (ball z d) (\k z. \i "(1 + d / norm z)" - def m \ "nat \norm z * e\" + define e where "e = (1 + d / norm z)" + define m where "m = nat \norm z * e\" { fix t assume t: "t \ ball z d" have "norm t = norm (z + (t - z))" by simp @@ -683,7 +683,7 @@ by blast+ let ?f' = "\z k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))" let ?f = "\z k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))" and ?F' = "\z. \n. ?f' z n" - def d \ "min (norm z/2) (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" + define d where "d = min (norm z/2) (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" from z have d: "d > 0" "norm z/2 \ d" by (auto simp add: complex_nonpos_Reals_iff d_def) have ball: "Im t = 0 \ Re t > 0" if "dist z t < d" for t using no_nonpos_Real_in_ball[OF z, of t] that unfolding d_def by (force simp add: complex_nonpos_Reals_iff) @@ -860,7 +860,7 @@ assume n: "n \ 0" from z have z': "z \ 0" by auto from no_nonpos_Int_in_ball'[OF z] guess d . note d = this - def n' \ "Suc n" + define n' where "n' = Suc n" from n have n': "n' \ 2" by (simp add: n'_def) have "((\z. \k. inverse ((z + of_nat k) ^ n')) has_field_derivative (\k. - of_nat n' * inverse ((z + of_nat k) ^ (n'+1)))) (at z)" @@ -1763,10 +1763,10 @@ defines "a \ complex_of_real pi" obtains h' where "continuous_on UNIV h'" "\z. (h has_field_derivative (h' z)) (at z)" proof - - def f \ "\n. a * of_real (cos_coeff (n+1) - sin_coeff (n+2))" - def F \ "\z. if z = 0 then 0 else (cos (a*z) - sin (a*z)/(a*z)) / z" - def g \ "\n. complex_of_real (sin_coeff (n+1))" - def G \ "\z. if z = 0 then 1 else sin (a*z)/(a*z)" + define f where "f n = a * of_real (cos_coeff (n+1) - sin_coeff (n+2))" for n + define F where "F z = (if z = 0 then 0 else (cos (a*z) - sin (a*z)/(a*z)) / z)" for z + define g where "g n = complex_of_real (sin_coeff (n+1))" for n + define G where "G z = (if z = 0 then 1 else sin (a*z)/(a*z))" for z have a_nz: "a \ 0" unfolding a_def by simp have "(\n. f n * (a*z)^n) sums (F z) \ (\n. g n * (a*z)^n) sums (G z)" @@ -1803,13 +1803,13 @@ qed note sums = conjunct1[OF this] conjunct2[OF this] - def h2 \ "\z. (\n. f n * (a*z)^n) / (\n. g n * (a*z)^n) + - Digamma (1 + z) - Digamma (1 - z)" - def POWSER \ "\f z. (\n. f n * (z^n :: complex))" - def POWSER' \ "\f z. (\n. diffs f n * (z^n :: complex))" - - def h2' \ "\z. a * (POWSER g (a*z) * POWSER' f (a*z) - POWSER f (a*z) * POWSER' g (a*z)) / - (POWSER g (a*z))^2 + Polygamma 1 (1 + z) + Polygamma 1 (1 - z)" + define h2 where [abs_def]: + "h2 z = (\n. f n * (a*z)^n) / (\n. g n * (a*z)^n) + Digamma (1 + z) - Digamma (1 - z)" for z + define POWSER where [abs_def]: "POWSER f z = (\n. f n * (z^n :: complex))" for f z + define POWSER' where [abs_def]: "POWSER' f z = (\n. diffs f n * (z^n))" for f and z :: complex + define h2' where [abs_def]: + "h2' z = a * (POWSER g (a*z) * POWSER' f (a*z) - POWSER f (a*z) * POWSER' g (a*z)) / + (POWSER g (a*z))^2 + Polygamma 1 (1 + z) + Polygamma 1 (1 - z)" for z have h_eq: "h t = h2 t" if "abs (Re t) < 1" for t proof - @@ -1850,7 +1850,7 @@ { fix z :: complex assume z: "abs (Re z) < 1" - def d \ "\ * of_real (norm z + 1)" + define d where "d = \ * of_real (norm z + 1)" have d: "abs (Re d) < 1" "norm z < norm d" by (simp_all add: d_def norm_mult del: of_real_add) have "eventually (\z. h z = h2 z) (nhds z)" using eventually_nhds_in_nhd[of z ?A] using h_eq z @@ -1902,7 +1902,7 @@ ultimately show "h2' (z - 1) = h2' z" by (rule DERIV_unique) qed - def h2'' \ "\z. h2' (z - of_int \Re z\)" + define h2'' where "h2'' z = h2' (z - of_int \Re z\)" for z have deriv: "(h has_field_derivative h2'' z) (at z)" for z proof - fix z :: complex @@ -1916,8 +1916,8 @@ have cont: "continuous_on UNIV h2''" proof (intro continuous_at_imp_continuous_on ballI) fix z :: complex - def r \ "\Re z\" - def A \ "{t. of_int r - 1 < Re t \ Re t < of_int r + 1}" + define r where "r = \Re z\" + define A where "A = {t. of_int r - 1 < Re t \ Re t < of_int r + 1}" have "continuous_on A (\t. h2' (t - of_int r))" unfolding A_def by (intro continuous_at_imp_continuous_on isCont_o2[OF _ A(2)] ballI continuous_intros) (simp_all add: abs_real_def) @@ -1957,9 +1957,9 @@ shows "Gamma z * Gamma (1 - z) = of_real pi / sin (of_real pi * z)" proof - let ?g = "\z::complex. Gamma z * Gamma (1 - z) * sin (of_real pi * z)" - def g \ "\z::complex. if z \ \ then of_real pi else ?g z" + define g where [abs_def]: "g z = (if z \ \ then of_real pi else ?g z)" for z :: complex let ?h = "\z::complex. (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))" - def h \ "\z::complex. if z \ \ then 0 else ?h z" + define h where [abs_def]: "h z = (if z \ \ then 0 else ?h z)" for z :: complex \ \@{term g} is periodic with period 1.\ interpret g: periodic_fun_simple' g @@ -2073,7 +2073,7 @@ qed have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" for z proof - - def r \ "\Re z / 2\" + define r where "r = \Re z / 2\" have "Gamma (1/2)^2 * g z = Gamma (1/2)^2 * g (z - of_int (2*r))" by (simp only: g.minus_of_int) also have "of_int (2*r) = 2 * of_int r" by simp also have "Re z - 2 * of_int r > -1" "Re z - 2 * of_int r < 2" unfolding r_def by linarith+ @@ -2126,8 +2126,8 @@ have h'_zero: "h' z = 0" for z proof - - def m \ "max 1 \Re z\" - def B \ "{t. abs (Re t) \ m \ abs (Im t) \ abs (Im z)}" + define m where "m = max 1 \Re z\" + define B where "B = {t. abs (Re t) \ m \ abs (Im t) \ abs (Im z)}" have "closed ({t. Re t \ -m} \ {t. Re t \ m} \ {t. Im t \ -\Im z\} \ {t. Im t \ \Im z\})" (is "closed ?B") by (intro closed_Int closed_halfspace_Re_ge closed_halfspace_Re_le @@ -2144,7 +2144,7 @@ qed ultimately have compact: "compact B" by (subst compact_eq_bounded_closed) blast - def M \ "SUP z:B. norm (h' z)" + define M where "M = (SUP z:B. norm (h' z))" have "compact (h' ` B)" by (intro compact_continuous_image continuous_on_subset[OF h'_cont] compact) blast+ hence bdd: "bdd_above ((\z. norm (h' z)) ` B)" @@ -2560,10 +2560,10 @@ theorem inverse_squares_sums: "(\n. 1 / (n + 1)\<^sup>2) sums (pi\<^sup>2 / 6)" proof - - def P \ "\x n. (\k=1..n. 1 - x^2 / of_nat k^2 :: real)" - def K \ "\n. inverse (real_of_nat (Suc n))^2" - def f \ "\x. \n. P x n / of_nat (Suc n)^2" - def g \ "\x. (1 - sin (pi * x) / (pi * x))" + define P where "P x n = (\k=1..n. 1 - x^2 / of_nat k^2)" for x :: real and n + define K where "K = (\n. inverse (real_of_nat (Suc n))^2)" + define f where [abs_def]: "f x = (\n. P x n / of_nat (Suc n)^2)" for x + define g where [abs_def]: "g x = (1 - sin (pi * x) / (pi * x))" for x have sums: "(\n. P x n / of_nat (Suc n)^2) sums (if x = 0 then K else g x / x^2)" for x proof (cases "x = 0") @@ -2606,7 +2606,7 @@ moreover have "f \ 0 \ pi^2 / 6" proof (rule Lim_transform_eventually) - def f' \ "\x. \n. - sin_coeff (n+3) * pi ^ (n+2) * x^n" + define f' where [abs_def]: "f' x = (\n. - sin_coeff (n+3) * pi ^ (n+2) * x^n)" for x have "eventually (\x. x \ (0::real)) (at 0)" by (auto simp add: eventually_at intro!: exI[of _ 1]) thus "eventually (\x. f' x = f x) (at 0)" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy --- a/src/HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy Mon Apr 25 22:59:53 2016 +0200 @@ -73,7 +73,7 @@ assumes "norm z < 1" shows "(\n. (a gchoose n) * z^n) sums (1 + z) powr a" proof - - def K \ "1 - (1 - norm z) / 2" + define K where "K = 1 - (1 - norm z) / 2" from assms have K: "K > 0" "K < 1" "norm z < K" unfolding K_def by (auto simp: field_simps intro!: add_pos_nonneg) let ?f = "\n. a gchoose n" and ?f' = "diffs (\n. a gchoose n)" @@ -83,7 +83,7 @@ hence summable': "summable (\n. ?f' n * z ^ n)" if "norm z < K" for z using that by (intro termdiff_converges[of _ K]) simp_all - def f \ "\z. \n. ?f n * z ^ n" and f' \ "\z. \n. ?f' n * z ^ n" + define f f' where [abs_def]: "f z = (\n. ?f n * z ^ n)" "f' z = (\n. ?f' n * z ^ n)" for z { fix z :: complex assume z: "norm z < K" from summable_mult2[OF summable'[OF z], of z] @@ -92,7 +92,7 @@ unfolding diffs_def by (subst (asm) summable_Suc_iff) have "(1 + z) * f' z = (\n. ?f' n * z^n) + (\n. ?f' n * z^Suc n)" - unfolding f'_def using summable' z by (simp add: algebra_simps suminf_mult) + unfolding f_f'_def using summable' z by (simp add: algebra_simps suminf_mult) also have "(\n. ?f' n * z^n) = (\n. of_nat (Suc n) * ?f (Suc n) * z^n)" by (intro suminf_cong) (simp add: diffs_def) also have "(\n. ?f' n * z^Suc n) = (\n. of_nat n * ?f n * z ^ n)" @@ -103,15 +103,15 @@ by (subst gbinomial_mult_1, subst suminf_add) (insert summable'[OF z] summable2, simp_all add: summable_powser_split_head algebra_simps diffs_def) - also have "\ = a * f z" unfolding f_def + also have "\ = a * f z" unfolding f_f'_def by (subst suminf_mult[symmetric]) (simp_all add: summable[OF z] mult_ac) finally have "a * f z = (1 + z) * f' z" by simp } note deriv = this have [derivative_intros]: "(f has_field_derivative f' z) (at z)" if "norm z < of_real K" for z - unfolding f_def f'_def using K that + unfolding f_f'_def using K that by (intro termdiffs_strong[of "?f" K z] summable_strong) simp_all - have "f 0 = (\n. if n = 0 then 1 else 0)" unfolding f_def by (intro suminf_cong) simp + have "f 0 = (\n. if n = 0 then 1 else 0)" unfolding f_f'_def by (intro suminf_cong) simp also have "\ = 1" using sums_single[of 0 "\_. 1::complex"] unfolding sums_iff by simp finally have [simp]: "f 0 = 1" . @@ -133,7 +133,7 @@ from c[of 0] and K have "c = 1" by simp with c[of z] have "f z = (1 + z) powr a" using K by (simp add: powr_minus_complex field_simps dist_complex_def) - with summable K show ?thesis unfolding f_def by (simp add: sums_iff) + with summable K show ?thesis unfolding f_f'_def by (simp add: sums_iff) qed lemma gen_binomial_complex': diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Harmonic_Numbers.thy --- a/src/HOL/Multivariate_Analysis/Harmonic_Numbers.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Harmonic_Numbers.thy Mon Apr 25 22:59:53 2016 +0200 @@ -250,7 +250,7 @@ assumes "(x::real) > 0" "a > 0" shows "ln (x + a) - ln x \ a * (inverse x + inverse (x + a))/2" (is "_ \ ?A") proof - - def f' \ "(inverse (x + a) - inverse x)/a" + define f' where "f' = (inverse (x + a) - inverse x)/a" have f'_nonpos: "f' \ 0" using assms by (simp add: f'_def divide_simps) let ?f = "\t. (t - x) * f' + inverse x" let ?F = "\t. (t - x)^2 * f' / 2 + t * inverse x" @@ -297,8 +297,8 @@ assumes "(x::real) > 0" "x < y" shows "ln y - ln x \ 2 * (y - x) / (x + y)" (is "_ \ ?A") proof - - def m \ "(x+y)/2" - def f' \ "-inverse (m^2)" + define m where "m = (x+y)/2" + define f' where "f' = -inverse (m^2)" from assms have m: "m > 0" by (simp add: m_def) let ?F = "\t. (t - m)^2 * f' / 2 + t / m" from assms have "((\t. (t - m) * f' + inverse m) has_integral (?F y - ?F x)) {x..y}" @@ -337,9 +337,10 @@ and euler_mascheroni_upper: "euler_mascheroni \ harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 1))" proof - - def D \ "\n. inverse (of_nat (n+1)) + ln (of_nat (n+1)) - ln (of_nat (n+2)) :: real" + define D :: "_ \ real" + where "D n = inverse (of_nat (n+1)) + ln (of_nat (n+1)) - ln (of_nat (n+2))" for n let ?g = "\n. ln (of_nat (n+2)) - ln (of_nat (n+1)) - inverse (of_nat (n+1)) :: real" - def inv \ "\n. inverse (real_of_nat n)" + define inv where [abs_def]: "inv n = inverse (real_of_nat n)" for n fix n :: nat note summable = sums_summable[OF euler_mascheroni_sum, folded D_def] have sums: "(\k. (inv (Suc (k + (n+1))) - inv (Suc (Suc k + (n+1))))/2) sums ((inv (Suc (0 + (n+1))) - 0)/2)" @@ -358,7 +359,7 @@ also have "\ \ -(\k. (inv (k + Suc n + 1) - inv (k + Suc n + 2)) / 2)" proof (intro le_imp_neg_le suminf_le allI summable_ignore_initial_segment[OF summable]) fix k' :: nat - def k \ "k' + Suc n" + define k where "k = k' + Suc n" hence k: "k > 0" by (simp add: k_def) have "real_of_nat (k+1) > 0" by (simp add: k_def) with ln_inverse_approx_le[OF this zero_less_one] @@ -386,7 +387,7 @@ also have "-(\k. D (k + Suc n)) \ -(\k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2)" proof (intro le_imp_neg_le suminf_le allI summable_ignore_initial_segment[OF summable]) fix k' :: nat - def k \ "k' + Suc n" + define k where "k = k' + Suc n" hence k: "k > 0" by (simp add: k_def) have "real_of_nat (k+1) > 0" by (simp add: k_def) from ln_inverse_approx_ge[of "of_nat k + 1" "of_nat k + 2"] @@ -435,7 +436,7 @@ let ?f = "\k. 2 * y ^ (2*k+1) / of_nat (2*k+1)" note sums = ln_series_quadratic[OF x(1)] - def c \ "inverse (2*y^(2*n+1))" + define c where "c = inverse (2*y^(2*n+1))" let ?d = "c * (ln x - (\kk. y\<^sup>2^k / of_nat (2*(k+n)+1) \ y\<^sup>2 ^ k / of_nat (2*n+1)" by (intro allI divide_left_mono mult_right_mono mult_pos_pos zero_le_power[of "y^2"]) simp_all @@ -466,7 +467,7 @@ shows ln_approx_bounds: "ln x \ {approx..approx + 2*d}" and ln_approx_abs: "abs (ln x - (approx + d)) \ d" proof - - def c \ "2*y^(2*n+1)" + define c where "c = 2*y^(2*n+1)" from x have c_pos: "c > 0" unfolding c_def y_def by (intro mult_pos_pos zero_less_power) simp_all have A: "inverse c * (ln x - (\k diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Mon Apr 25 22:59:53 2016 +0200 @@ -928,7 +928,7 @@ proof let ?B = "\f::'a\'a \ 'a. cbox (\i\Basis. (fst (f i) \ i) *\<^sub>R i) (\i\Basis. (snd (f i) \ i) *\<^sub>R i)" - def p \ "?B ` (Basis \\<^sub>E {(a, c), (c, d), (d, b)})" + define p where "p = ?B ` (Basis \\<^sub>E {(a, c), (c, d), (d, b)})" show "cbox c d \ p" unfolding p_def @@ -2004,10 +2004,7 @@ apply (drule choice) apply blast done - def AB \ "\n. (f ^^ n) (a,b)" - def A \ "\n. fst(AB n)" - def B \ "\n. snd(AB n)" - note ab_def = A_def B_def AB_def + define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n have "A 0 = a" "B 0 = b" "\n. \ P (cbox (A(Suc n)) (B(Suc n))) \ (\i\Basis. A(n)\i \ A(Suc n)\i \ A(Suc n)\i \ B(Suc n)\i \ B(Suc n)\i \ B(n)\i \ 2 * (B(Suc n)\i - A(Suc n)\i) \ B(n)\i - A(n)\i)" (is "\n. ?P n") @@ -3453,8 +3450,8 @@ and "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?t2) proof - guess y using assms(1) unfolding integrable_on_def .. note y=this - def b' \ "\i\Basis. (if i = k then min (b\k) c else b\i)*\<^sub>R i::'a" - def a' \ "\i\Basis. (if i = k then max (a\k) c else a\i)*\<^sub>R i::'a" + define b' where "b' = (\i\Basis. (if i = k then min (b\k) c else b\i)*\<^sub>R i)" + define a' where "a' = (\i\Basis. (if i = k then max (a\k) c else a\i)*\<^sub>R i)" show ?t1 ?t2 unfolding interval_split[OF k] integrable_cauchy unfolding interval_split[symmetric,OF k] @@ -3954,7 +3951,7 @@ by (simp add: "*" iterate_expand_cases) next case True - def su \ "support opp f s" + define su where "su = support opp f s" have fsu: "finite su" using True by (simp add: su_def) moreover @@ -3976,7 +3973,7 @@ and "d division_of (cbox a b)" shows "iterate opp d f = f (cbox a b)" proof - - def C \ "card (division_points (cbox a b) d)" + define C where [abs_def]: "C = card (division_points (cbox a b) d)" then show ?thesis using assms proof (induct C arbitrary: a b d rule: full_nat_induct) @@ -4079,10 +4076,10 @@ done note this(2-4,1) note kc=this[unfolded interval_bounds[OF ab']] from this(3) guess j .. note j=this - def d1 \ "{l \ {x. x\k \ c} | l. l \ d \ l \ {x. x\k \ c} \ {}}" - def d2 \ "{l \ {x. x\k \ c} | l. l \ d \ l \ {x. x\k \ c} \ {}}" - def cb \ "(\i\Basis. (if i = k then c else b\i) *\<^sub>R i)::'a" - def ca \ "(\i\Basis. (if i = k then c else a\i) *\<^sub>R i)::'a" + define d1 where "d1 = {l \ {x. x\k \ c} | l. l \ d \ l \ {x. x\k \ c} \ {}}" + define d2 where "d2 = {l \ {x. x\k \ c} | l. l \ d \ l \ {x. x\k \ c} \ {}}" + define cb where "cb = (\i\Basis. (if i = k then c else b\i) *\<^sub>R i)" + define ca where "ca = (\i\Basis. (if i = k then c else a\i) *\<^sub>R i)" note division_points_psubset[OF \d division_of cbox a b\ ab kc(1-2) j] note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)] then have *: "(iterate opp d1 f) = f (cbox a b \ {x. x\k \ c})" @@ -4734,7 +4731,7 @@ done next case False - def d \ "e / 3 / setprod (\i. b\i - a\i) (Basis - {k})" + define d where "d = e / 3 / setprod (\i. b\i - a\i) (Basis - {k})" note False[unfolded content_eq_0 not_ex not_le, rule_format] then have "\x. x \ Basis \ b\x > a\x" by (auto simp add:not_le) @@ -5227,7 +5224,7 @@ next fix x k assume xk: "(x, k) \ p" - def n \ "nat \norm (f x)\" + define n where "n = nat \norm (f x)\" have *: "norm (f x) \ (\(x, k). norm (f x)) ` p" using xk by auto have nfx: "real n \ norm (f x)" "norm (f x) \ real n + 1" @@ -6085,7 +6082,7 @@ using assms proof cases assume p: "p \ 1" - def p' \ "p - 2" + define p' where "p' = p - 2" from assms p have p': "{..i. i \ p' \ Suc (Suc p' - i) = (Suc (Suc p') - i)" @@ -6125,8 +6122,9 @@ case 1 interpret bounded_bilinear "scaleR::real\'a\'a" by (rule bounded_bilinear_scaleR) - def g \ "\s. (b - s)^(p - 1)/fact (p - 1)" - def Dg \ "\n s. if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0" + define g where "g s = (b - s)^(p - 1)/fact (p - 1)" for s + define Dg where [abs_def]: + "Dg n s = (if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0)" for n s have g0: "Dg 0 = g" using \p > 0\ by (auto simp add: Dg_def divide_simps g_def split: if_split_asm) @@ -6253,8 +6251,8 @@ using s(3)[OF k(1),unfolded k] unfolding box_ne_empty by auto then have xi: "x\i = d\i" using as unfolding k mem_box by (metis antisym) - def y \ "\j\Basis. (if j = i then if c\i \ (a\i + b\i) / 2 then c\i + - min e (b\i - c\i) / 2 else c\i - min e (c\i - a\i) / 2 else x\j) *\<^sub>R j" + define y where "y = (\j\Basis. (if j = i then if c\i \ (a\i + b\i) / 2 then c\i + + min e (b\i - c\i) / 2 else c\i - min e (c\i - a\i) / 2 else x\j) *\<^sub>R j)" show "\x'\\(s - {k}). x' \ x \ dist x' x < e" apply (rule_tac x=y in bexI) proof @@ -6619,7 +6617,7 @@ assume e: "e > 0" with assms(1) have "e * r > 0" by simp from assms(8)[unfolded has_integral,rule_format,OF this] guess d by (elim exE conjE) note d=this[rule_format] - def d' \ "\x. {y. g y \ d (g x)}" + define d' where "d' x = {y. g y \ d (g x)}" for x have d': "\x. d' x = {y. g y \ (d (g x))}" unfolding d'_def .. show "\d. gauge d \ (\p. p tagged_division_of h ` cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)" @@ -7590,7 +7588,7 @@ done from integrable_integral[OF this,unfolded has_integral_real,rule_format,OF *] guess d1 .. note d1 = conjunctD2[OF this,rule_format] - def d \ "\x. ball x w \ d1 x" + define d where [abs_def]: "d x = ball x w \ d1 x" for x have "gauge d" unfolding d_def using w(1) d1 by auto note this[unfolded gauge_def,rule_format,of c] @@ -7628,7 +7626,7 @@ done from integrable_integral[OF this,unfolded has_integral_real,rule_format,OF *] guess d2 .. note d2 = conjunctD2[OF this,rule_format] - def d3 \ "\x. if x \ t then d1 x \ d2 x else d1 x" + define d3 where "d3 x = (if x \ t then d1 x \ d2 x else d1 x)" for x have "gauge d3" using d2(1) d1(1) unfolding d3_def gauge_def by auto from fine_division_exists_real[OF this, of a t] guess p . note p=this @@ -8072,7 +8070,7 @@ and "cbox c d \ cbox a b" shows "((\x. if x \ box c d then f x else 0) has_integral i) (cbox a b)" proof - - def g \ "\x. if x \box c d then f x else 0" + define g where [abs_def]: "g x = (if x \box c d then f x else 0)" for x { presume *: "cbox c d \ {} \ ?thesis" show ?thesis @@ -8264,8 +8262,8 @@ next assume as: "\e>0. ?r e" from this[rule_format,OF zero_less_one] guess C .. note C=conjunctD2[OF this,rule_format] - def c \ "(\i\Basis. (- max B C) *\<^sub>R i)::'n" - def d \ "(\i\Basis. max B C *\<^sub>R i)::'n" + define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" + define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" have c_d: "cbox a b \ cbox c d" apply safe apply (drule B(2)) @@ -8300,8 +8298,8 @@ then have "0 < norm (y - i)" by auto from as[rule_format,OF this] guess C .. note C=conjunctD2[OF this,rule_format] - def c \ "(\i\Basis. (- max B C) *\<^sub>R i)::'n" - def d \ "(\i\Basis. max B C *\<^sub>R i)::'n" + define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" + define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" have c_d: "cbox a b \ cbox c d" apply safe apply (drule B(2)) @@ -8895,8 +8893,8 @@ note conjunctD2[OF this, rule_format] note h = this(1) and this(2)[OF *] from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format] - def c \ "\i\Basis. min (a\i) (- (max B1 B2)) *\<^sub>R i::'n" - def d \ "\i\Basis. max (b\i) (max B1 B2) *\<^sub>R i::'n" + define c :: 'n where "c = (\i\Basis. min (a\i) (- (max B1 B2)) *\<^sub>R i)" + define d :: 'n where "d = (\i\Basis. max (b\i) (max B1 B2) *\<^sub>R i)" have *: "ball 0 B1 \ cbox c d" "ball 0 B2 \ cbox c d" apply safe unfolding mem_ball mem_box dist_norm @@ -9283,7 +9281,7 @@ using p'(3) by fastforce note partial_division_of_tagged_division[OF p(1)] this from partial_division_extend_interval[OF this] guess q . note q=this and q' = division_ofD[OF this(2)] - def r \ "q - snd ` p" + define r where "r = q - snd ` p" have "snd ` p \ r = {}" unfolding r_def by auto have r: "finite r" @@ -9567,7 +9565,7 @@ fixes x :: "'a::ring_1" shows "(1 - x) * setsum (\i. x^i) {0 .. n} = (1 - x^(Suc n))" proof - - def y \ "1 - x" + define y where "y = 1 - x" have "y * (\i=0..n. (1 - y) ^ i) = 1 - (1 - y) ^ Suc n" by (induct n) (simp_all add: algebra_simps) then show ?thesis @@ -9745,8 +9743,7 @@ done qed from bchoice[OF this] guess m .. note m=conjunctD2[OF this[rule_format],rule_format] - def d \ "\x. c (m x) x" - + define d where "d x = c (m x) x" for x show ?case apply (rule_tac x=d in exI) proof safe @@ -10548,7 +10545,7 @@ assume "p tagged_division_of (cbox a b)" and "?g fine p" note p = this(1) conjunctD2[OF this(2)[unfolded fine_inter]] note p' = tagged_division_ofD[OF p(1)] - def p' \ "{(x,k) | x k. \i l. x \ i \ i \ d \ (x,l) \ p \ k = i \ l}" + define p' where "p' = {(x,k) | x k. \i l. x \ i \ i \ d \ (x,l) \ p \ k = i \ l}" have gp': "g fine p'" using p(2) unfolding p'_def fine_def @@ -10705,7 +10702,7 @@ proof (rule setsum_mono, goal_cases) case k: (1 k) from d'(4)[OF this] guess u v by (elim exE) note uv=this - def d' \ "{cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" + define d' where "d' = {cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" note uvab = d'(2)[OF k[unfolded uv]] have "d' division_of cbox u v" apply (subst d'_def) @@ -11609,8 +11606,8 @@ obtains X0 where "x0 \ X0" "open X0" "\x\X0 \ U. \t \ C. dist (fx (x, t)) (fx (x0, t)) \ e" proof - - def psi \ "\(x, t). dist (fx (x, t)) (fx (x0, t))" - def W0 \ "{(x, t) \ U \ C. psi (x, t) < e}" + define psi where "psi = (\(x, t). dist (fx (x, t)) (fx (x0, t)))" + define W0 where "W0 = {(x, t) \ U \ C. psi (x, t) < e}" have W0_eq: "W0 = psi -` {.. U \ C" by (auto simp: vimage_def W0_def) have "open {.. 0" - def e \ "e' / (content (cbox a b) + 1)" + define e where "e = e' / (content (cbox a b) + 1)" have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) assume "x \ U" from continuous_on_prod_compactE[OF cont_fx compact_cbox \x \ U\ \0 < e\] @@ -11744,7 +11741,7 @@ note [continuous_intros] = continuous_on_compose2[OF cont_f1] fix e'::real assume "e' > 0" - def e \ "e' / (content (cbox a b) + 1)" + define e where "e = e' / (content (cbox a b) + 1)" have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) from continuous_on_prod_compactE[OF cont_fx compact_cbox \x0 \ U\ \e > 0\] obtain X0 where X0: "x0 \ X0" "open X0" @@ -11916,7 +11913,7 @@ proof (rule tendstoI) fix e::real assume "e > 0" - def e' \ "e / 2" + define e' where "e' = e / 2" with \e > 0\ have "e' > 0" by simp then have "\\<^sub>F n in F. \x\cbox a b. norm (f n x - g x) < e' / content (cbox a b)" using u content_nonzero content_pos_le[of a b] diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Path_Connected.thy --- a/src/HOL/Multivariate_Analysis/Path_Connected.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy Mon Apr 25 22:59:53 2016 +0200 @@ -563,7 +563,7 @@ assumes "path(g1 +++ g2)" "path g2" shows "pathfinish g1 = pathstart g2" proof (rule ccontr) - def e \ "dist (g1 1) (g2 0)" + define e where "e = dist (g1 1) (g2 0)" assume Neg: "pathfinish g1 \ pathstart g2" then have "0 < dist (pathfinish g1) (pathstart g2)" by auto @@ -1934,7 +1934,7 @@ then obtain B::real where B: "0 < B" and Bx: "norm (a - x) < B" and By: "norm (a - y) < B" and "s \ ball a B" using bounded_subset_ballD [OF bxy, of a] by (auto simp: dist_norm) - def C == "B / norm(x - a)" + define C where "C = B / norm(x - a)" { fix u assume u: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + C *\<^sub>R (x - a)) \ s" and "0 \ u" "u \ 1" have CC: "1 \ 1 + (C - 1) * u" @@ -1969,7 +1969,7 @@ } then have pcx: "path_component (- s) x (a + C *\<^sub>R (x - a))" by (force simp: closed_segment_def intro!: path_connected_linepath) - def D == "B / norm(y - a)" \\massive duplication with the proof above\ + define D where "D = B / norm(y - a)" \\massive duplication with the proof above\ { fix u assume u: "(1 - u) *\<^sub>R y + u *\<^sub>R (a + D *\<^sub>R (y - a)) \ s" and "0 \ u" "u \ 1" have DD: "1 \ 1 + (D - 1) * u" @@ -2531,7 +2531,7 @@ { assume "bounded (connected_component_set (- s) z)" with bounded_pos_less obtain B where "B>0" and B: "\x. connected_component (- s) z x \ norm x < B" by (metis mem_Collect_eq) - def C \ "((B + 1 + norm z) / norm (z-a))" + define C where "C = (B + 1 + norm z) / norm (z-a)" have "C > 0" using \0 < B\ zna by (simp add: C_def divide_simps add_strict_increasing) have "\norm (z + C *\<^sub>R (z-a)) - norm (C *\<^sub>R (z-a))\ \ norm z" @@ -2751,7 +2751,7 @@ { fix \ assume "path \" and pimg_sbs: "path_image \ - {pathfinish \} \ interior (- t)" and pf: "pathfinish \ \ frontier t" and ps: "pathstart \ = a" - def c \ "pathfinish \" + define c where "c = pathfinish \" have "c \ -s" unfolding c_def using front pf by blast moreover have "open (-s)" using s compact_imp_closed by blast ultimately obtain \::real where "\ > 0" and \: "cball c \ \ -s" @@ -3260,8 +3260,10 @@ and geq: "\x. k1 (1, x) = g x" "\x. k2 (0, x) = g x" and k12: "\x. k1 (0, x) = f x" "\x. k2 (1, x) = h x" and P: "\t\{0..1}. P (\x. k1 (t, x))" "\t\{0..1}. P (\x. k2 (t, x))" - def k \ "\y. if fst y \ 1 / 2 then (k1 o (\x. (2 *\<^sub>R fst x, snd x))) y - else (k2 o (\x. (2 *\<^sub>R fst x -1, snd x))) y" + define k where "k y = + (if fst y \ 1 / 2 + then (k1 o (\x. (2 *\<^sub>R fst x, snd x))) y + else (k2 o (\x. (2 *\<^sub>R fst x -1, snd x))) y)" for y have keq: "k1 (2 * u, v) = k2 (2 * u - 1, v)" if "u = 1/2" for u v by (simp add: geq that) have "continuous_on ({0..1} \ X) k" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Summation.thy --- a/src/HOL/Multivariate_Analysis/Summation.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Summation.thy Mon Apr 25 22:59:53 2016 +0200 @@ -142,7 +142,7 @@ finally have "l \ 0" by simp with l obtain l' where l': "l = ereal l'" by (cases l) simp_all - def c \ "(1 - l') / 2" + define c where "c = (1 - l') / 2" from l and \l \ 0\ have c: "l + c > l" "l' + c \ 0" "l' + c < 1" unfolding c_def by (simp_all add: field_simps l') have "\C>l. eventually (\n. ereal (root n (norm (f n))) < C) sequentially" @@ -176,7 +176,7 @@ also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) finally have l_nonneg: "l \ 0" by simp - def c \ "if l = \ then 2 else 1 + (real_of_ereal l - 1) / 2" + define c where "c = (if l = \ then 2 else 1 + (real_of_ereal l - 1) / 2)" from l l_nonneg consider "l = \" | "\l'. l = ereal l'" by (cases l) simp_all hence c: "c > 1 \ ereal c < l" by cases (insert l, auto simp: c_def field_simps) @@ -193,7 +193,7 @@ qed from bounded obtain K where K: "K > 0" "\n. norm (f n) \ K" using BseqE by blast - def n \ "nat \log c K\" + define n where "n = nat \log c K\" from unbounded have "\m>n. c < root m (norm (f m))" unfolding bdd_above_def by (auto simp: not_le) then guess m by (elim exE conjE) note m = this @@ -253,7 +253,7 @@ assumes nonneg: "\n. f n \ 0" shows "summable f \ summable (\n. 2^n * f (2^n))" proof - - def f' \ "\n. if n = 0 then 0 else f n" + define f' where "f' n = (if n = 0 then 0 else f n)" for n from mono have mono': "decseq (\n. f (Suc n))" by (intro decseq_SucI) simp hence mono': "f n \ f m" if "m \ n" "m > 0" for m n using that decseqD[OF mono', of "m - 1" "n - 1"] by simp @@ -402,7 +402,7 @@ shows "summable f" unfolding summable_iff_convergent' proof - - def r \ "(if l = \ then 1 else real_of_ereal l / 2)" + define r where "r = (if l = \ then 1 else real_of_ereal l / 2)" from l have "r > 0 \ of_real r < l" by (cases l) (simp_all add: r_def) hence r: "r > 0" "of_real r < l" by simp_all hence "eventually (\n. p n * f n / f (Suc n) - p (Suc n) > r) sequentially" @@ -420,7 +420,7 @@ have "Bseq (\n. (\k\n + Suc m. f k))" proof (rule BseqI') fix k :: nat - def n \ "k + Suc m" + define n where "n = k + Suc m" have n: "n > m" by (simp add: n_def) from r have "r * norm (\k\n. f k) = norm (\k\n. r * f k)" @@ -588,7 +588,7 @@ assumes "ereal (norm z) < conv_radius f" shows "summable (\n. norm (f n * z ^ n))" proof (rule root_test_convergence') - def l \ "limsup (\n. ereal (root n (norm (f n))))" + define l where "l = limsup (\n. ereal (root n (norm (f n))))" have "0 = limsup (\n. 0)" by (simp add: Limsup_const) also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) finally have l_nonneg: "l \ 0" . @@ -626,7 +626,7 @@ assumes "ereal (norm z) > conv_radius f" shows "\summable (\n. f n * z ^ n)" proof (rule root_test_divergence) - def l \ "limsup (\n. ereal (root n (norm (f n))))" + define l where "l = limsup (\n. ereal (root n (norm (f n))))" have "0 = limsup (\n. 0)" by (simp add: Limsup_const) also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) finally have l_nonneg: "l \ 0" . @@ -676,7 +676,7 @@ with conv_radius_nonneg[of f] obtain conv_radius' where [simp]: "conv_radius f = ereal conv_radius'" by (cases "conv_radius f") simp_all - def r \ "if R = \ then conv_radius' + 1 else (real_of_ereal R + conv_radius') / 2" + define r where "r = (if R = \ then conv_radius' + 1 else (real_of_ereal R + conv_radius') / 2)" from R conv_radius_nonneg[of f] have "0 < r \ ereal r < R \ ereal r > conv_radius f" unfolding r_def by (cases R) (auto simp: r_def field_simps) with assms(1)[of r] obtain z where "norm z > conv_radius f" "summable (\n. f n * z^n)" by auto @@ -701,7 +701,8 @@ proof (rule linorder_cases[of "conv_radius f" R]) assume R: "conv_radius f > R" from R assms(1) obtain R' where R': "R = ereal R'" by (cases R) simp_all - def r \ "if conv_radius f = \ then R' + 1 else (R' + real_of_ereal (conv_radius f)) / 2" + define r where + "r = (if conv_radius f = \ then R' + 1 else (R' + real_of_ereal (conv_radius f)) / 2)" from R conv_radius_nonneg[of f] have "r > R \ r < conv_radius f" unfolding r_def by (cases "conv_radius f") (auto simp: r_def field_simps R') with assms(1) assms(2)[of r] R' @@ -749,7 +750,7 @@ proof (rule ccontr) assume "conv_radius f \ 0" with conv_radius_nonneg[of f] have pos: "conv_radius f > 0" by simp - def r \ "if conv_radius f = \ then 1 else real_of_ereal (conv_radius f) / 2" + define r where "r = (if conv_radius f = \ then 1 else real_of_ereal (conv_radius f) / 2)" from pos have r: "ereal r > 0 \ ereal r < conv_radius f" by (cases "conv_radius f") (simp_all add: r_def) hence "summable (\n. f n * of_real r ^ n)" by (intro summable_in_conv_radius) simp diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Apr 25 22:59:53 2016 +0200 @@ -228,7 +228,8 @@ "\a. a \ A' \ open a" "\S. open S \ x \ S \ \a\A'. a \ S" by (rule first_countable_basisE) blast - def A \ "(\N. \((\n. from_nat_into A' n) ` N)) ` (Collect finite::nat set set)" + define A where [abs_def]: + "A = (\N. \((\n. from_nat_into A' n) ` N)) ` (Collect finite::nat set set)" then show "\A. countable A \ (\a. a \ A \ x \ a) \ (\a. a \ A \ open a) \ (\S. open S \ x \ S \ (\a\A. a \ S)) \ (\a b. a \ A \ b \ A \ a \ b \ A)" proof (safe intro!: exI[where x=A]) @@ -385,7 +386,7 @@ instance second_countable_topology \ first_countable_topology proof fix x :: 'a - def B \ "SOME B::'a set set. countable B \ topological_basis B" + define B :: "'a set set" where "B = (SOME B. countable B \ topological_basis B)" then have B: "countable B" "topological_basis B" using countable_basis is_basis by (auto simp: countable_basis is_basis) @@ -723,7 +724,7 @@ then show ?rhs unfolding openin_open open_dist by blast next - def T \ "{x. \a\S. \d>0. (\y\U. dist y a < d \ y \ S) \ dist x a < d}" + define T where "T = {x. \a\S. \d>0. (\y\U. dist y a < d \ y \ S) \ dist x a < d}" have 1: "\x\T. \e>0. \y. dist y x < e \ y \ T" unfolding T_def apply clarsimp @@ -1030,7 +1031,7 @@ assumes "e > 0" shows "\a b. (\i\Basis. a \ i \ \ \ b \ i \ \ ) \ x \ box a b \ box a b \ ball x e" proof - - def e' \ "e / (2 * sqrt (real (DIM ('a))))" + define e' where "e' = e / (2 * sqrt (real (DIM ('a))))" then have e: "e' > 0" using assms by (auto simp: DIM_positive) have "\i. \y. y \ \ \ y < x \ i \ x \ i - y < e'" (is "\i. ?th i") @@ -1385,9 +1386,8 @@ bs: "set bs = Basis" "distinct bs" by (metis finite_distinct_list) from nonempty_Basis s obtain j where j: "j \ Basis" "s j \ S" by blast - def y \ "rec_list - (s j) - (\j _ Y. (\i\Basis. (if i = j then s i \ i else Y \ i) *\<^sub>R i))" + define y where + "y = rec_list (s j) (\j _ Y. (\i\Basis. (if i = j then s i \ i else Y \ i) *\<^sub>R i))" have "x = (\i\Basis. (if i \ set bs then s i \ i else s j \ i) *\<^sub>R i)" using bs by (auto simp add: s(1)[symmetric] euclidean_representation) also have [symmetric]: "y bs = \" @@ -2534,7 +2534,7 @@ "\i. x \ A i" "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" by blast - def f \ "\n. SOME y. y \ S \ y \ A n \ x \ y" + define f where "f n = (SOME y. y \ S \ y \ A n \ x \ y)" for n { fix n from \?lhs\ have "\y. y \ S \ y \ A n \ x \ y" @@ -3137,8 +3137,8 @@ then obtain r where "0 < r" "\z. dist z y < r \ z \ T" unfolding open_dist by fast (* choose point between x and y, within distance r of y. *) - def k \ "min 1 (r / (2 * dist x y))" - def z \ "y + scaleR k (x - y)" + define k where "k = min 1 (r / (2 * dist x y))" + define z where "z = y + scaleR k (x - y)" have z_def2: "z = x + scaleR (1 - k) (y - x)" unfolding z_def by (simp add: algebra_simps) have "dist z y < r" @@ -3676,7 +3676,7 @@ "\i. l \ A i" "\S. open S \ l \ S \ eventually (\i. A i \ S) sequentially" by blast - def s \ "\n i. SOME j. i < j \ f j \ A (Suc n)" + define s where "s n i = (SOME j. i < j \ f j \ A (Suc n))" for n i { fix n i have "infinite (A (Suc n) \ range f - f`{.. i})" @@ -3691,7 +3691,7 @@ unfolding s_def by (auto intro: someI2_ex) } note s = this - def r \ "rec_nat (s 0 0) s" + define r where "r = rec_nat (s 0 0) s" have "subseq r" by (auto simp: r_def s subseq_Suc_iff) moreover @@ -3984,7 +3984,7 @@ then have "U \ {}" by (auto simp: eventually_False) - def Z \ "closure ` {A. eventually (\x. x \ A) F}" + define Z where "Z = closure ` {A. eventually (\x. x \ A) F}" then have "\z\Z. closed z" by auto moreover @@ -4018,7 +4018,7 @@ next fix A assume A: "\a\A. closed a" "\B\A. finite B \ U \ \B \ {}" "U \ \A = {}" - def F \ "INF a:insert U A. principal a" + define F where "F = (INF a:insert U A. principal a)" have "F \ bot" unfolding F_def proof (rule INF_filter_not_bot) @@ -4089,8 +4089,7 @@ fix A assume A: "\a\A. open a" "U \ \A" assume *: "\A. countable A \ (\a\A. open a) \ U \ \A \ (\T\A. finite T \ U \ \T)" - - moreover def C \ "{b\B. \a\A. b \ U \ a}" + moreover define C where "C = {b\B. \a\A. b \ U \ a}" ultimately have "countable C" "\a\C. open a" unfolding C_def using ccover by auto moreover @@ -4202,7 +4201,7 @@ by auto then obtain X' where T: "\T. T \ A \ finite T \ X' T \ U - \T" by metis - def X \ "\n. X' (from_nat_into A ` {.. n})" + define X where "X n = X' (from_nat_into A ` {.. n})" for n have X: "\n. X n \ U - (\i\n. from_nat_into A i)" using \A \ {}\ unfolding X_def by (intro T) (auto intro: from_nat_into) then have "range X \ U" @@ -4249,7 +4248,7 @@ "\i. x \ A i" "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" by blast - def s \ "\n i. SOME j. i < j \ X j \ A (Suc n)" + define s where "s n i = (SOME j. i < j \ X j \ A (Suc n))" for n i { fix n i have "\a. i < a \ X a \ A (Suc n)" @@ -4270,7 +4269,7 @@ unfolding s_def by (auto intro: someI2_ex) } note s = this - def r \ "rec_nat (s 0 0) s" + define r where "r = rec_nat (s 0 0) s" have "subseq r" by (auto simp: r_def s subseq_Suc_iff) moreover @@ -4303,7 +4302,7 @@ and "t \ s" shows "\x\s. \U. x\U \ open U \ infinite (U \ t)" proof (rule ccontr) - def C \ "(\F. interior (F \ (- t))) ` {F. finite F \ F \ t }" + define C where "C = (\F. interior (F \ (- t))) ` {F. finite F \ F \ t }" note \countably_compact s\ moreover have "\t\C. open t" by (auto simp: C_def) @@ -4442,7 +4441,7 @@ from seq_compact_imp_totally_bounded[OF \seq_compact s\] obtain f where f: "\e>0. finite (f e) \ f e \ s \ s \ (\x\f e. ball x e)" unfolding choice_iff' .. - def K \ "(\(x, r). ball x r) ` ((\e \ \ \ {0 <..}. f e) \ \)" + define K where "K = (\(x, r). ball x r) ` ((\e \ \ \ {0 <..}. f e) \ \)" have "countably_compact s" using \seq_compact s\ by (rule seq_compact_imp_countably_compact) then show "compact s" @@ -4628,11 +4627,11 @@ then obtain l2 r2 where r2:"subseq r2" and lr2:"((\i. f (r1 (r2 i)) proj k) \ l2) sequentially" using bounded_imp_convergent_subsequence[of "\i. f (r1 i) proj k"] by (auto simp: o_def) - def r \ "r1 \ r2" + define r where "r = r1 \ r2" have r:"subseq r" using r1 and r2 unfolding r_def o_def subseq_def by auto moreover - def l \ "unproj (\i. if i = k then l2 else l1 proj i)::'a" + define l where "l = unproj (\i. if i = k then l2 else l1 proj i)" { fix e::real assume "e > 0" @@ -4822,9 +4821,9 @@ fix f :: "nat \ 'a" assume f: "\n. f n \ s" - def e \ "\n. 1 / (2 * Suc n)" + define e where "e n = 1 / (2 * Suc n)" for n then have [simp]: "\n. 0 < e n" by auto - def B \ "\n U. SOME b. infinite {n. f n \ b} \ (\x. b \ ball x (e n) \ U)" + define B where "B n U = (SOME b. infinite {n. f n \ b} \ (\x. b \ ball x (e n) \ U))" for n U { fix n U assume "infinite {n. f n \ U}" @@ -4841,7 +4840,7 @@ } note B = this - def F \ "rec_nat (B 0 UNIV) B" + define F where "F = rec_nat (B 0 UNIV) B" { fix n have "infinite {i. f i \ F n}" @@ -4862,7 +4861,7 @@ by (simp add: set_eq_iff not_le conj_commute) qed - def t \ "rec_nat (sel 0 0) (\n i. sel (Suc n) i)" + define t where "t = rec_nat (sel 0 0) (\n i. sel (Suc n) i)" have "subseq t" unfolding subseq_Suc_iff by (simp add: t_def sel) moreover have "\i. (f \ t) i \ s" @@ -5587,8 +5586,8 @@ using choice[of "\d x. d>0 \ fst x \ s \ snd x \ s \ dist (snd x) (fst x) < d \ \ dist (f (snd x)) (f (fst x)) < e"] unfolding Bex_def by (auto simp add: dist_commute) - def x \ "\n::nat. fst (fa (inverse (real n + 1)))" - def y \ "\n::nat. snd (fa (inverse (real n + 1)))" + define x where "x n = fst (fa (inverse (real n + 1)))" for n + define y where "y n = snd (fa (inverse (real n + 1)))" for n have xyn: "\n. x n \ s \ y n \ s" and xy0: "\n. dist (x n) (y n) < inverse (real n + 1)" and fxy:"\n. \ dist (f (x n)) (f (y n)) < e" @@ -6552,7 +6551,8 @@ proof (cases, safe) fix e :: real assume "0 < e" "s \ {}" - def [simp]: R \ "{(y, d). y \ s \ 0 < d \ ball y d \ s \ {x \ s. f x \ ball (f y) (e/2) } }" + define R where [simp]: + "R = {(y, d). y \ s \ 0 < d \ ball y d \ s \ {x \ s. f x \ ball (f y) (e/2)}}" let ?b = "(\(y, d). ball y (d/2))" have "(\r\R. open (?b r))" "s \ (\r\R. ?b r)" proof safe @@ -7451,13 +7451,13 @@ instance euclidean_space \ second_countable_topology proof - def a \ "\f :: 'a \ (real \ real). \i\Basis. fst (f i) *\<^sub>R i" + define a where "a f = (\i\Basis. fst (f i) *\<^sub>R i)" for f :: "'a \ real \ real" then have a: "\f. (\i\Basis. fst (f i) *\<^sub>R i) = a f" by simp - def b \ "\f :: 'a \ (real \ real). \i\Basis. snd (f i) *\<^sub>R i" + define b where "b f = (\i\Basis. snd (f i) *\<^sub>R i)" for f :: "'a \ real \ real" then have b: "\f. (\i\Basis. snd (f i) *\<^sub>R i) = b f" by simp - def B \ "(\f. box (a f) (b f)) ` (Basis \\<^sub>E (\ \ \))" + define B where "B = (\f. box (a f) (b f)) ` (Basis \\<^sub>E (\ \ \))" have "Ball B open" by (simp add: B_def open_box) moreover have "(\A. open A \ (\B'\B. \B' = A))" @@ -7655,7 +7655,7 @@ { fix x assume as:"x \ cbox a b" - def f \ "\n::nat. x + (inverse (real n + 1)) *\<^sub>R (?c - x)" + define f where [abs_def]: "f n = x + (inverse (real n + 1)) *\<^sub>R (?c - x)" for n { fix n assume fn: "f n a f n = x" and xc: "x \ ?c" @@ -7712,7 +7712,7 @@ proof - obtain b where "b>0" and b: "\x\s. norm x \ b" using assms[unfolded bounded_pos] by auto - def a \ "(\i\Basis. (b + 1) *\<^sub>R i)::'a" + define a :: 'a where "a = (\i\Basis. (b + 1) *\<^sub>R i)" { fix x assume "x \ s" @@ -8062,7 +8062,7 @@ assumes "compact s" "continuous_on s f" "f ` s = t" "inj_on f s" shows "\g. homeomorphism s t f g" proof - - def g \ "\x. SOME y. y\s \ f y = x" + define g where "g x = (SOME y. y\s \ f y = x)" for x have g: "\x\s. g (f x) = x" using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto { @@ -8486,7 +8486,7 @@ have "1 - c > 0" using c by auto from s(2) obtain z0 where "z0 \ s" by auto - def z \ "\n. (f ^^ n) z0" + define z where "z n = (f ^^ n) z0" for n { fix n :: nat have "z n \ s" unfolding z_def @@ -8498,7 +8498,7 @@ then show ?case using f by auto qed } note z_in_s = this - def d \ "dist (z 0) (z 1)" + define d where "d = dist (z 0) (z 1)" have fzn:"\n. f (z n) = z (Suc n)" unfolding z_def by auto { @@ -8608,7 +8608,7 @@ then obtain x where "x\s" and x:"(z \ x) sequentially" using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto - def e \ "dist (f x) x" + define e where "e = dist (f x) x" have "e = 0" proof (rule ccontr) assume "e \ 0" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Uniform_Limit.thy --- a/src/HOL/Multivariate_Analysis/Uniform_Limit.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Uniform_Limit.thy Mon Apr 25 22:59:53 2016 +0200 @@ -72,7 +72,7 @@ shows "(h \ l) (at x within S)" proof (rule tendstoI) fix e :: real - def e' \ "e/3" + define e' where "e' = e/3" assume "0 < e" then have "0 < e'" by (simp add: e'_def) from uniform_limitD[OF uc \0 < e'\] diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/Weierstrass.thy --- a/src/HOL/Multivariate_Analysis/Weierstrass.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Weierstrass.thy Mon Apr 25 22:59:53 2016 +0200 @@ -218,7 +218,7 @@ have "t \ t0" using t t0 by auto then obtain g where g: "g \ R" "g t \ g t0" using separable t0 by (metis Diff_subset subset_eq t) - def h \ "\x. g x - g t0" + define h where [abs_def]: "h x = g x - g t0" for x have "h \ R" unfolding h_def by (fast intro: g const diff) then have hsq: "(\w. (h w)\<^sup>2) \ R" @@ -232,7 +232,7 @@ also have "... \ normf (\w. (h w)\<^sup>2)" using t normf_upper [where x=t] continuous [OF hsq] by force finally have nfp: "0 < normf (\w. (h w)\<^sup>2)" . - def p \ "\x. (1 / normf (\w. (h w)\<^sup>2)) * (h x)^2" + define p where [abs_def]: "p x = (1 / normf (\w. (h w)\<^sup>2)) * (h x)^2" for x have "p \ R" unfolding p_def by (fast intro: hsq const mult) moreover have "p t0 = 0" @@ -266,7 +266,7 @@ using t1 by auto then have cardp: "card subU > 0" using subU by (simp add: card_gt_0_iff) - def p \ "\x. (1 / card subU) * (\t \ subU. pf t x)" + define p where [abs_def]: "p x = (1 / card subU) * (\t \ subU. pf t x)" for x have pR: "p \ R" unfolding p_def using subU pf by (fast intro: pf const mult setsum) have pt0 [simp]: "p t0 = 0" @@ -307,7 +307,7 @@ by (auto simp: elim!: openE) then have pt_delta: "\x. x \ s-U \ p x \ delta0" by (force simp: ball_def dist_norm dest: p01) - def \ \ "delta0/2" + define \ where "\ = delta0/2" have "delta0 \ 1" using delta0 p01 [of t1] t1 by (force simp: ball_def dist_norm dest: p01) with delta0 have \01: "0 < \" "\ < 1" @@ -318,7 +318,7 @@ by (rule open_Collect_less_Int [OF continuous [OF pR] continuous_on_const]) then obtain V where V: "open V" "V \ s = {x\s. p x < \/2}" by blast - def k \ "nat\1/\\ + 1" + define k where "k = nat\1/\\ + 1" have "k>0" by (simp add: k_def) have "k-1 \ 1/\" using \01 by (simp add: k_def) @@ -331,7 +331,7 @@ using \01 unfolding k_def by linarith with \01 k2\ have k\: "1 < k*\" "k*\ < 2" by (auto simp: divide_simps) - def q \ "\n t. (1 - p t ^ n) ^ (k^n)" + define q where [abs_def]: "q n t = (1 - p t ^ n) ^ (k^n)" for n t have qR: "q n \ R" for n by (simp add: q_def const diff power pR) have q01: "\n t. t \ s \ q n t \ {0..1}" @@ -387,7 +387,8 @@ by (fastforce simp: field_simps) finally have "q n t \ (1 / (real k * \)) ^ n " . } note limitNonU = this - def NN \ "\e. 1 + nat \max (ln e / ln (real k * \ / 2)) (- ln e / ln (real k * \))\" + define NN + where "NN e = 1 + nat \max (ln e / ln (real k * \ / 2)) (- ln e / ln (real k * \))\" for e have NN: "of_nat (NN e) > ln e / ln (real k * \ / 2)" "of_nat (NN e) > - ln e / ln (real k * \)" if "0w. w \ A \ ff w \ R \ ff w ` s \ {0..1} \ (\x \ s \ Vf w. ff w x < e / card subA) \ (\x \ s \ B. ff w x > 1 - e / card subA)" by metis - def pff \ "\x. (\w \ subA. ff w x)" + define pff where [abs_def]: "pff x = (\w \ subA. ff w x)" for x have pffR: "pff \ R" unfolding pff_def using subA ff by (auto simp: intro: setprod) moreover @@ -559,9 +560,9 @@ and e: "0 < e" "e < 1/3" shows "\g \ R. \x\s. \f x - g x\ < 2*e" proof - - def n \ "1 + nat \normf f / e\" - def A \ "\j::nat. {x \ s. f x \ (j - 1/3)*e}" - def B \ "\j::nat. {x \ s. f x \ (j + 1/3)*e}" + define n where "n = 1 + nat \normf f / e\" + define A where "A j = {x \ s. f x \ (j - 1/3)*e}" for j :: nat + define B where "B j = {x \ s. f x \ (j + 1/3)*e}" for j :: nat have ngt: "(n-1) * e \ normf f" "n\1" using e apply (simp_all add: n_def field_simps of_nat_Suc) @@ -591,7 +592,7 @@ and xfA: "\x j. x \ A j \ xf j x < e/n" and xfB: "\x j. x \ B j \ xf j x > 1 - e/n" by metis - def g \ "\x. e * (\i\n. xf i x)" + define g where [abs_def]: "g x = e * (\i\n. xf i x)" for x have gR: "g \ R" unfolding g_def by (fast intro: mult const setsum xfR) have gge0: "\x. x \ s \ g x \ 0" @@ -606,7 +607,7 @@ done { fix t assume t: "t \ s" - def j \ "LEAST j. t \ A j" + define j where "j = (LEAST j. t \ A j)" have jn: "j \ n" using t An by (simp add: Least_le j_def) have Aj: "t \ A j" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Multivariate_Analysis/ex/Approximations.thy --- a/src/HOL/Multivariate_Analysis/ex/Approximations.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/ex/Approximations.thy Mon Apr 25 22:59:53 2016 +0200 @@ -89,7 +89,7 @@ assumes "n > 0" "0 \ x" "x < 2" shows "exp (x::real) - (\k {0..(2 * x^n / (2 - x)) / fact n}" proof (unfold atLeastAtMost_iff, safe) - def approx \ "(\kkk. x^k / fact k) sums exp x" using exp_converges[of x] by (simp add: field_simps) from sums_split_initial_segment[OF this, of n] @@ -323,10 +323,12 @@ shows "abs (euler_mascheroni - approx :: real) < e" (is "abs (_ - ?approx) < ?e") proof - - def l \ "47388813395531028639296492901910937/82101866951584879688289000000000000 :: real" - def u \ "142196984054132045946501548559032969 / 246305600854754639064867000000000000 :: real" + define l :: real + where "l = 47388813395531028639296492901910937/82101866951584879688289000000000000" + define u :: real + where "u = 142196984054132045946501548559032969 / 246305600854754639064867000000000000" have impI: "P \ Q" if Q for P Q using that by blast - have hsum_63: "harm 63 = (310559566510213034489743057 / 65681493561267903750631200 ::real)" + have hsum_63: "harm 63 = (310559566510213034489743057 / 65681493561267903750631200 :: real)" by (simp add: harm_expand) from harm_Suc[of 63] have hsum_64: "harm 64 = 623171679694215690971693339 / (131362987122535807501262400::real)" @@ -368,7 +370,7 @@ assumes x: "0 \ x" "x < 1" and n: "even n" shows "arctan x - arctan_approx n x \ {0..x^(2*n+1) / (1-x^4)}" proof - - def c \ "\k. 1 / (1+(4*real k + 2*real n)) - x\<^sup>2 / (3+(4*real k + 2*real n))" + define c where "c k = 1 / (1+(4*real k + 2*real n)) - x\<^sup>2 / (3+(4*real k + 2*real n))" for k from assms have "(\k. (-1) ^ k * (1 / real (k * 2 + 1) * x^(k*2+1))) sums arctan x" using arctan_series' by simp also have "(\k. (-1) ^ k * (1 / real (k * 2 + 1) * x^(k*2+1))) = @@ -559,11 +561,13 @@ text \We can now approximate pi to 22 decimals within a fraction of a second.\ lemma pi_approx_75: "abs (pi - 3.1415926535897932384626 :: real) \ inverse (10^22)" proof - - def a \ "8295936325956147794769600190539918304 / 2626685325478320010006427764892578125 :: real" - def b \ "8428294561696506782041394632 / 503593538783547230635598424135 :: real" + define a :: real + where "a = 8295936325956147794769600190539918304 / 2626685325478320010006427764892578125" + define b :: real + where "b = 8428294561696506782041394632 / 503593538783547230635598424135" \ \The introduction of this constant prevents the simplifier from applying solvers that we don't want. We want it to simply evaluate the terms to rational constants.}\ - def eq \ "op = :: real \ real \ bool" + define eq :: "real \ real \ bool" where "eq = op =" \ \Splitting the computation into several steps has the advantage that simplification can be done in parallel\ @@ -664,15 +668,18 @@ "abs (pi - 3.141592653589793238462643383279502884197169399375105821 :: real) \ inverse (10^54)" (is "abs (pi - ?pi') \ _") proof - - def a \ "2829469759662002867886529831139137601191652261996513014734415222704732791803 / - 1062141879292765061960538947347721564047051545995266466660439319087625011200 :: real" - def b \ "13355545553549848714922837267299490903143206628621657811747118592 / - 23792006023392488526789546722992491355941103837356113731091180925 :: real" - def c \ "28274063397213534906669125255762067746830085389618481175335056 / - 337877029279505250241149903214554249587517250716358486542628059 :: real" + define a :: real + where "a = 2829469759662002867886529831139137601191652261996513014734415222704732791803 / + 1062141879292765061960538947347721564047051545995266466660439319087625011200" + define b :: real + where "b = 13355545553549848714922837267299490903143206628621657811747118592 / + 23792006023392488526789546722992491355941103837356113731091180925" + define c :: real + where "c = 28274063397213534906669125255762067746830085389618481175335056 / + 337877029279505250241149903214554249587517250716358486542628059" let ?pi'' = "3882327391761098513316067116522233897127356523627918964967729040413954225768920394233198626889767468122598417405434625348404038165437924058179155035564590497837027530349 / 1235783190199688165469648572769847552336447197542738425378629633275352407743112409829873464564018488572820294102599160968781449606552922108667790799771278860366957772800" - def eq \ "op = :: real \ real \ bool" + define eq :: "real \ real \ bool" where "eq = op =" have "abs (pi - pi_approx2 4) \ inverse (2^183)" by (rule pi_approx2') simp_all also have "pi_approx2 4 = 48 * arctan_approx 24 (1 / 18) + diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Nat.thy --- a/src/HOL/Nat.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Nat.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1832,7 +1832,7 @@ proof - from assms obtain q where "m = n + Suc q" by (auto dest: less_imp_Suc_add) - moreover def r \ "Suc q" + moreover define r where "r = Suc q" ultimately have "Suc (m - Suc n) = r" and "m = n + r" by simp_all then show ?thesis by simp diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Nonstandard_Analysis/HLim.thy --- a/src/HOL/Nonstandard_Analysis/HLim.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Nonstandard_Analysis/HLim.thy Mon Apr 25 22:59:53 2016 +0200 @@ -209,7 +209,7 @@ done text\NS continuity can be defined using NS Limit in - similar fashion to standard def of continuity\ + similar fashion to standard definition of continuity\ lemma isNSCont_NSLIM_iff: "(isNSCont f a) = (f \a\\<^sub>N\<^sub>S (f a))" by (blast intro: isNSCont_NSLIM NSLIM_isNSCont) @@ -236,7 +236,7 @@ text\Alternative definition of continuity\ (* Prove equivalence between NS limits - *) -(* seems easier than using standard def *) +(* seems easier than using standard definition *) lemma NSLIM_h_iff: "(f \a\\<^sub>N\<^sub>S L) = ((%h. f(a + h)) \0\\<^sub>N\<^sub>S L)" apply (simp add: NSLIM_def, auto) apply (drule_tac x = "star_of a + x" in spec) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/NthRoot.thy --- a/src/HOL/NthRoot.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/NthRoot.thy Mon Apr 25 22:59:53 2016 +0200 @@ -646,7 +646,7 @@ lemma LIMSEQ_root: "(\n. root n n) \ 1" proof - - def x \ "\n. root n n - 1" + define x where "x n = root n n - 1" for n have "x \ sqrt 0" proof (rule tendsto_sandwich[OF _ _ tendsto_const]) show "(\x. sqrt (2 / x)) \ sqrt 0" @@ -684,7 +684,7 @@ shows "(\n. root n c) \ 1" proof - { fix c :: real assume "1 \ c" - def x \ "\n. root n c - 1" + define x where "x n = root n c - 1" for n have "x \ 0" proof (rule tendsto_sandwich[OF _ _ tendsto_const]) show "(\n. c / n) \ 0" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Number_Theory/Eratosthenes.thy --- a/src/HOL/Number_Theory/Eratosthenes.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Number_Theory/Eratosthenes.thy Mon Apr 25 22:59:53 2016 +0200 @@ -96,7 +96,7 @@ with A B C show ?thesis by simp next case False - def m \ "Suc n" + define m where "m = Suc n" then have "m > 0" by simp from False have "n > 0" by simp from A obtain q where q: "Suc (Suc a) = Suc n * q" by (rule dvdE) @@ -133,7 +133,8 @@ enumerate_Suc_eq in_set_enumerate_eq less_eq_dvd_minus) next case 3 - { def v \ "Suc m" and w \ "Suc n" + { define v where "v = Suc m" + define w where "w = Suc n" fix q assume "m + n \ q" then obtain r where q: "q = m + n + r" by (auto simp add: le_iff_add) @@ -359,7 +360,7 @@ have "List.find (\p. p \ m) (primes_upto n) = Some (smallest_prime_beyond m)" if assms: "m \ p" "prime p" "p \ n" for p proof - - def A \ "{p. p \ n \ prime p \ m \ p}" + define A where "A = {p. p \ n \ prime p \ m \ p}" from assms have "smallest_prime_beyond m \ p" by (auto intro: smallest_prime_beyond_smallest) from this \p \ n\ have *: "smallest_prime_beyond m \ n" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Number_Theory/Euclidean_Algorithm.thy --- a/src/HOL/Number_Theory/Euclidean_Algorithm.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Number_Theory/Euclidean_Algorithm.thy Mon Apr 25 22:59:53 2016 +0200 @@ -159,8 +159,8 @@ next case True then obtain l\<^sub>0 where l\<^sub>0_props: "l\<^sub>0 \ 0 \ (\a\A. a dvd l\<^sub>0)" by blast - def n \ "LEAST n. \l. l \ 0 \ (\a\A. a dvd l) \ euclidean_size l = n" - def l \ "SOME l. l \ 0 \ (\a\A. a dvd l) \ euclidean_size l = n" + define n where "n = (LEAST n. \l. l \ 0 \ (\a\A. a dvd l) \ euclidean_size l = n)" + define l where "l = (SOME l. l \ 0 \ (\a\A. a dvd l) \ euclidean_size l = n)" have "\l. l \ 0 \ (\a\A. a dvd l) \ euclidean_size l = n" apply (subst n_def) apply (rule LeastI[of _ "euclidean_size l\<^sub>0"]) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Number_Theory/Factorial_Ring.thy --- a/src/HOL/Number_Theory/Factorial_Ring.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Number_Theory/Factorial_Ring.thy Mon Apr 25 22:59:53 2016 +0200 @@ -258,7 +258,7 @@ by (rule dvd_mult_left) with Suc \is_prime p\ \\ p dvd a\ have "p dvd b" by (simp add: prime_dvd_mult_iff) - moreover def c \ "b div p" + moreover define c where "c = b div p" ultimately have b: "b = p * c" by simp with * have "p * p ^ n dvd p * (a * c)" by (simp add: ac_simps) @@ -371,10 +371,10 @@ by simp then have "p dvd a" and "is_prime p" and "normalize p = p" and "p \ 0" by (auto intro!: is_prime_not_zeroI) - def n \ "Max {n. p ^ n dvd a}" + define n where "n = Max {n. p ^ n dvd a}" then have "n > 0" and "p ^ n dvd a" and "\ p ^ Suc n dvd a" proof - - def N \ "{n. p ^ n dvd a}" + define N where "N = {n. p ^ n dvd a}" then have n_M: "n = Max N" by (simp add: n_def) from is_prime_inj_power \is_prime p\ have "inj (op ^ p)" . then have "inj_on (op ^ p) U" for U @@ -402,7 +402,7 @@ then show "\ p ^ Suc n dvd a" by (simp add: n_M) qed - def b \ "a div p ^ n" + define b where "b = a div p ^ n" with \p ^ n dvd a\ have a: "a = p ^ n * b" by simp with \\ p ^ Suc n dvd a\ have "\ p dvd b" and "b \ 0" @@ -466,7 +466,7 @@ "replicate_mset (Max {n. q ^ n dvd a}) q = replicate_mset (Max {n. q ^ n dvd b}) q" by simp qed - def Q \ "the (factorization b)" + define Q where "Q = the (factorization b)" with \b \ 0\ have [simp]: "factorization b = Some Q" by simp from \a \ 0\ have "factorization a = diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Power.thy --- a/src/HOL/Power.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Power.thy Mon Apr 25 22:59:53 2016 +0200 @@ -90,7 +90,7 @@ case 0 then show ?case by (simp add: fun_eq_iff) next case (Suc n) - def g \ "\x. f x - 1" + define g where "g x = f x - 1" for x with Suc have "n = g x" by simp with Suc have "times x ^^ g x = times (x ^ g x)" by simp moreover from Suc g_def have "f x = g x + 1" by simp diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Binary_Product_Measure.thy --- a/src/HOL/Probability/Binary_Product_Measure.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Binary_Product_Measure.thy Mon Apr 25 22:59:53 2016 +0200 @@ -995,7 +995,7 @@ finally show ?thesis by (simp add: top_unique) next case False - def C' \ "fst ` C" + define C' where "C' = fst ` C" have "\ = \\<^sup>+ x. ennreal (1 / Suc n) * indicator C' x \count_space UNIV" using C'_def False by(simp add: nn_integral_cmult ennreal_mult_top) also have "\ = \\<^sup>+ x. \\<^sup>+ y. ennreal (1 / Suc n) * indicator C' x * indicator {SOME y. (x, y) \ C} y \count_space UNIV \count_space UNIV" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Bochner_Integration.thy --- a/src/HOL/Probability/Bochner_Integration.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Bochner_Integration.thy Mon Apr 25 22:59:53 2016 +0200 @@ -24,7 +24,7 @@ obtain D :: "'b set" where "countable D" and D: "\X. open X \ X \ {} \ \d\D. d \ X" by (erule countable_dense_setE) - def e \ "from_nat_into D" + define e where "e = from_nat_into D" { fix n x obtain d where "d \ D" and d: "d \ ball x (1 / Suc n)" using D[of "ball x (1 / Suc n)"] by auto @@ -34,12 +34,14 @@ by auto } note e = this - def A \ "\m n. {x\space M. dist (f x) (e n) < 1 / (Suc m) \ 1 / (Suc m) \ dist (f x) z}" - def B \ "\m. disjointed (A m)" + define A where [abs_def]: "A m n = + {x\space M. dist (f x) (e n) < 1 / (Suc m) \ 1 / (Suc m) \ dist (f x) z}" for m n + define B where [abs_def]: "B m = disjointed (A m)" for m - def m \ "\N x. Max {m::nat. m \ N \ x \ (\n\N. B m n)}" - def F \ "\N::nat. \x. if (\m\N. x \ (\n\N. B m n)) \ (\n\N. x \ B (m N x) n) - then e (LEAST n. x \ B (m N x) n) else z" + define m where [abs_def]: "m N x = Max {m. m \ N \ x \ (\n\N. B m n)}" for N x + define F where [abs_def]: "F N x = + (if (\m\N. x \ (\n\N. B m n)) \ (\n\N. x \ B (m N x) n) + then e (LEAST n. x \ B (m N x) n) else z)" for N x have B_imp_A[intro, simp]: "\x m n. x \ B m n \ x \ A m n" using disjointed_subset[of "A m" for m] unfolding B_def by auto @@ -86,7 +88,7 @@ then have 1: "\m\N. x \ (\n\N. B m n)" by auto from m[OF this] obtain n where n: "m N x \ N" "n \ N" "x \ B (m N x) n" by auto moreover - def L \ "LEAST n. x \ B (m N x) n" + define L where "L = (LEAST n. x \ B (m N x) n)" have "dist (f x) (e L) < 1 / Suc (m N x)" proof - have "x \ B (m N x) L" @@ -175,7 +177,7 @@ sup: "\x. (SUP i. U i x) = ennreal (u x)" by blast - def U' \ "\i x. indicator (space M) x * enn2real (U i x)" + define U' where [abs_def]: "U' i x = indicator (space M) x * enn2real (U i x)" for i x then have U'_sf[measurable]: "\i. simple_function M (U' i)" using U by (auto intro!: simple_function_compose1[where g=enn2real]) @@ -268,7 +270,7 @@ assume non_empty: "\x\space M. f x \ 0" - def m \ "Min (f`space M - {0})" + define m where "m = Min (f`space M - {0})" have "m \ f`space M - {0}" unfolding m_def using f non_empty by (intro Min_in) (auto simp: simple_function_def) then have m: "0 < m" @@ -730,7 +732,7 @@ have [measurable]: "\i. s i \ borel_measurable M" using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) - def m \ "if space M = {} then 0 else Max ((\x. norm (s i x))`space M)" + define m where "m = (if space M = {} then 0 else Max ((\x. norm (s i x))`space M))" have "finite (s i ` space M)" using s by (auto simp: simple_function_def simple_bochner_integrable.simps) then have "finite (norm ` s i ` space M)" @@ -1821,7 +1823,7 @@ by (induct A rule: infinite_finite_induct) (auto intro!: add) } note setsum = this - def s' \ "\i z. indicator (space M) z *\<^sub>R s i z" + define s' where [abs_def]: "s' i z = indicator (space M) z *\<^sub>R s i z" for i z then have s'_eq_s: "\i x. x \ space M \ s' i x = s i x" by simp @@ -2600,7 +2602,8 @@ have "\i. s i \ measurable (N \\<^sub>M M) (count_space UNIV)" by (rule measurable_simple_function) fact - def f' \ "\i x. if integrable M (f x) then simple_bochner_integral M (\y. s i (x, y)) else 0" + define f' where [abs_def]: "f' i x = + (if integrable M (f x) then simple_bochner_integral M (\y. s i (x, y)) else 0)" for i x { fix i x assume "x \ space N" then have "simple_bochner_integral M (\y. s i (x, y)) = diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Borel_Space.thy Mon Apr 25 22:59:53 2016 +0200 @@ -282,34 +282,35 @@ assumes "a \ b" "(c::real) \ d" "{c..d} \ {g a..g b}" obtains c' d' where "{a..b} \ g -` {c..d} = {c'..d'}" "c' \ d'" "g c' = c" "g d' = d" proof- - let ?A = "{a..b} \ g -` {c..d}" - from IVT'[of g a c b, OF _ _ \a \ b\ assms(1)] assms(4,5) - obtain c'' where c'': "c'' \ ?A" "g c'' = c" by auto - from IVT'[of g a d b, OF _ _ \a \ b\ assms(1)] assms(4,5) - obtain d'' where d'': "d'' \ ?A" "g d'' = d" by auto - hence [simp]: "?A \ {}" by blast + let ?A = "{a..b} \ g -` {c..d}" + from IVT'[of g a c b, OF _ _ \a \ b\ assms(1)] assms(4,5) + obtain c'' where c'': "c'' \ ?A" "g c'' = c" by auto + from IVT'[of g a d b, OF _ _ \a \ b\ assms(1)] assms(4,5) + obtain d'' where d'': "d'' \ ?A" "g d'' = d" by auto + hence [simp]: "?A \ {}" by blast - def c' \ "Inf ?A" and d' \ "Sup ?A" - have "?A \ {c'..d'}" unfolding c'_def d'_def - by (intro subsetI) (auto intro: cInf_lower cSup_upper) - moreover from assms have "closed ?A" - using continuous_on_closed_vimage[of "{a..b}" g] by (subst Int_commute) simp - hence c'd'_in_set: "c' \ ?A" "d' \ ?A" unfolding c'_def d'_def - by ((intro closed_contains_Inf closed_contains_Sup, simp_all)[])+ - hence "{c'..d'} \ ?A" using assms - by (intro subsetI) - (auto intro!: order_trans[of c "g c'" "g x" for x] order_trans[of "g x" "g d'" d for x] - intro!: mono) - moreover have "c' \ d'" using c'd'_in_set(2) unfolding c'_def by (intro cInf_lower) auto - moreover have "g c' \ c" "g d' \ d" - apply (insert c'' d'' c'd'_in_set) - apply (subst c''(2)[symmetric]) - apply (auto simp: c'_def intro!: mono cInf_lower c'') [] - apply (subst d''(2)[symmetric]) - apply (auto simp: d'_def intro!: mono cSup_upper d'') [] - done - with c'd'_in_set have "g c' = c" "g d' = d" by auto - ultimately show ?thesis using that by blast + define c' where "c' = Inf ?A" + define d' where "d' = Sup ?A" + have "?A \ {c'..d'}" unfolding c'_def d'_def + by (intro subsetI) (auto intro: cInf_lower cSup_upper) + moreover from assms have "closed ?A" + using continuous_on_closed_vimage[of "{a..b}" g] by (subst Int_commute) simp + hence c'd'_in_set: "c' \ ?A" "d' \ ?A" unfolding c'_def d'_def + by ((intro closed_contains_Inf closed_contains_Sup, simp_all)[])+ + hence "{c'..d'} \ ?A" using assms + by (intro subsetI) + (auto intro!: order_trans[of c "g c'" "g x" for x] order_trans[of "g x" "g d'" d for x] + intro!: mono) + moreover have "c' \ d'" using c'd'_in_set(2) unfolding c'_def by (intro cInf_lower) auto + moreover have "g c' \ c" "g d' \ d" + apply (insert c'' d'' c'd'_in_set) + apply (subst c''(2)[symmetric]) + apply (auto simp: c'_def intro!: mono cInf_lower c'') [] + apply (subst d''(2)[symmetric]) + apply (auto simp: d'_def intro!: mono cSup_upper d'') [] + done + with c'd'_in_set have "g c' = c" "g d' = d" by auto + ultimately show ?thesis using that by blast qed subsection \Generic Borel spaces\ @@ -540,7 +541,7 @@ by (auto simp: topological_basis_def) from B(2)[OF K] obtain m where m: "\k. k \ K \ m k \ B" "\k. k \ K \ (\m k) = k" by metis - def U \ "(\k\K. m k)" + define U where "U = (\k\K. m k)" with m have "countable U" by (intro countable_subset[OF _ \countable B\]) auto have "\U = (\A\U. A)" by simp @@ -1811,7 +1812,7 @@ assumes f[measurable]: "\i. f i \ borel_measurable M" shows "(\x. lim (\i. f i x)) \ borel_measurable M" proof - - def u' \ "\x. lim (\i. if Cauchy (\i. f i x) then f i x else 0)" + define u' where "u' x = lim (\i. if Cauchy (\i. f i x) then f i x else 0)" for x then have *: "\x. lim (\i. f i x) = (if Cauchy (\i. f i x) then u' x else (THE x. False))" by (auto simp: lim_def convergent_eq_cauchy[symmetric]) have "u' \ borel_measurable M" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Caratheodory.thy --- a/src/HOL/Probability/Caratheodory.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Caratheodory.thy Mon Apr 25 22:59:53 2016 +0200 @@ -372,7 +372,7 @@ and Ble: "\n. (\i. f (B n i)) \ ?O (A n) + e * (1/2)^(Suc n)" by (metis less_imp_le outer_measure_close[OF *]) - def C \ "case_prod B o prod_decode" + define C where "C = case_prod B o prod_decode" from B have B_in_M: "\i j. B i j \ M" by (rule range_subsetD) then have C: "range C \ M" @@ -469,7 +469,7 @@ have inc: "increasing M f" by (metis additive_increasing ca countably_additive_additive posf) let ?O = "outer_measure M f" - def ls \ "lambda_system \ (Pow \) ?O" + define ls where "ls = lambda_system \ (Pow \) ?O" have mls: "measure_space \ ls ?O" using sigma_algebra.caratheodory_lemma [OF sigma_algebra_Pow outer_measure_space_outer_measure [OF posf inc]] @@ -660,7 +660,7 @@ have "F' i \ F' j = {}" by auto } note F'_disj = this - def F \ "\i. if i < card C then F' i else {}" + define F where "F i = (if i < card C then F' i else {})" for i then have "disjoint_family F" using F'_disj by (auto simp: disjoint_family_on_def) moreover from F' have "(\i. F i) = \C" @@ -704,7 +704,7 @@ from generated_ringE[OF Un_A] guess C' . note C' = this { fix c assume "c \ C'" - moreover def A \ "\i. A' i \ c" + moreover define A where [abs_def]: "A i = A' i \ c" for i ultimately have A: "range A \ generated_ring" "disjoint_family A" and Un_A: "(\i. A i) \ generated_ring" using A' C' @@ -722,7 +722,7 @@ have "\F'. bij_betw F' {..finite C\]) auto then guess F .. note F = this - def f \ "\i. if i < card C then F i else {}" + define f where [abs_def]: "f i = (if i < card C then F i else {})" for i then have f: "bij_betw f {..< card C} C" by (intro bij_betw_cong[THEN iffD1, OF _ F]) auto with C have "\j. f j \ M" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Central_Limit_Theorem.thy --- a/src/HOL/Probability/Central_Limit_Theorem.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Central_Limit_Theorem.thy Mon Apr 25 22:59:53 2016 +0200 @@ -24,8 +24,8 @@ shows "weak_conv_m (\n. distr M borel (\x. S n x / sqrt (n * \\<^sup>2))) std_normal_distribution" proof - let ?S' = "\n x. S n x / sqrt (real n * \\<^sup>2)" - def \ \ "\n. char (distr M borel (?S' n))" - def \ \ "\n t. char \ (t / sqrt (\\<^sup>2 * n))" + define \ where "\ n = char (distr M borel (?S' n))" for n + define \ where "\ n t = char \ (t / sqrt (\\<^sup>2 * n))" for n t have X_rv [simp, measurable]: "\n. random_variable borel (X n)" using X_indep unfolding indep_vars_def2 by simp @@ -48,7 +48,7 @@ hence n: "n \ t^2 / 4" by (subst nat_ceiling_le_eq [symmetric]) let ?t = "t / sqrt (\\<^sup>2 * n)" - def \' \ "\n i. char (distr M borel (\x. X i x / sqrt (\\<^sup>2 * n)))" + define \' where "\' n i = char (distr M borel (\x. X i x / sqrt (\\<^sup>2 * n)))" for n i have *: "\n i t. \' n i t = \ n t" unfolding \_def \'_def char_def by (subst X_distrib [symmetric]) (auto simp: integral_distr) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Characteristic_Functions.thy --- a/src/HOL/Probability/Characteristic_Functions.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Characteristic_Functions.thy Mon Apr 25 22:59:53 2016 +0200 @@ -308,7 +308,7 @@ have integ_iexp: "integrable M (\x. iexp (t * x))" by (intro integrable_const_bound) auto - def c \ "\k x. (ii * t)^k / fact k * complex_of_real (x^k)" + define c where [abs_def]: "c k x = (ii * t)^k / fact k * complex_of_real (x^k)" for k x have integ_c: "\k. k \ n \ integrable M (\x. c k x)" unfolding c_def by (intro integrable_mult_right integrable_of_real integrable_moments) @@ -343,7 +343,7 @@ have integ_iexp: "integrable M (\x. iexp (t * x))" by (intro integrable_const_bound) auto - def c \ "\k x. (ii * t)^k / fact k * complex_of_real (x^k)" + define c where [abs_def]: "c k x = (ii * t)^k / fact k * complex_of_real (x^k)" for k x have integ_c: "\k. k \ n \ integrable M (\x. c k x)" unfolding c_def by (intro integrable_mult_right integrable_of_real integrable_moments) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Distributions.thy --- a/src/HOL/Probability/Distributions.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Distributions.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1216,7 +1216,7 @@ shows "(\x. \\<^sup>+y. ennreal (normal_density 0 \ (x - y) * normal_density 0 \ y) \lborel) = normal_density 0 (sqrt (\\<^sup>2 + \\<^sup>2))" (is "?LHS = ?RHS") proof - - def \' \ "\\<^sup>2" and \' \ "\\<^sup>2" + define \' \' where "\' = \\<^sup>2" and "\' = \\<^sup>2" then have [simp, arith]: "0 < \'" "0 < \'" by simp_all let ?\ = "sqrt ((\' * \') / (\' + \'))" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Fin_Map.thy --- a/src/HOL/Probability/Fin_Map.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Fin_Map.thy Mon Apr 25 22:59:53 2016 +0200 @@ -467,12 +467,12 @@ assume "Cauchy P" then obtain Nd where Nd: "\n. n \ Nd \ dist (P n) (P Nd) < 1" by (force simp: cauchy) - def d \ "domain (P Nd)" + define d where "d = domain (P Nd)" with Nd have dim: "\n. n \ Nd \ domain (P n) = d" using dist_le_1_imp_domain_eq by auto have [simp]: "finite d" unfolding d_def by simp - def p \ "\i n. (P n) i" - def q \ "\i. lim (p i)" - def Q \ "finmap_of d q" + define p where "p i n = P n i" for i n + define q where "q i = lim (p i)" for i + define Q where "Q = finmap_of d q" have q: "\i. i \ d \ q i = Q i" by (auto simp add: Q_def Abs_finmap_inverse) { fix i assume "i \ d" @@ -505,7 +505,7 @@ from p[OF \i \ d\, THEN metric_LIMSEQ_D, OF \0 < e\] show "\no. \n\no. dist (p i n) (q i) < e" . qed then guess ni .. note ni = this - def N \ "max Nd (Max (ni ` d))" + define N where "N = max Nd (Max (ni ` d))" show "\N. \n\N. dist (P n) Q < e" proof (safe intro!: exI[where x="N"]) fix n assume "N \ n" @@ -602,7 +602,7 @@ thus "\y. x i \ y \ y \ a i \ y \ basis_proj" by auto qed then guess B .. note B = this - def B' \ "Pi' (domain x) (\i. (B i)::'b set)" + define B' where "B' = Pi' (domain x) (\i. (B i)::'b set)" have "B' \ Pi' (domain x) a" using B by (auto intro!: Pi'_mono simp: B'_def) also note \\ \ O'\ finally show "\B'\basis_finmap. x \ B' \ B' \ O'" using B @@ -1097,7 +1097,7 @@ proof (rule sigma_fprod_algebra_sigma_eq) show "finite I" by simp show "I \ {}" by fact - def S'\"from_nat_into S" + define S' where "S' = from_nat_into S" show "(\j. S' j) = space borel" using S apply (auto simp add: from_nat_into countable_basis_proj S'_def basis_proj_def) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Finite_Product_Measure.thy --- a/src/HOL/Probability/Finite_Product_Measure.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Finite_Product_Measure.thy Mon Apr 25 22:59:53 2016 +0200 @@ -440,7 +440,7 @@ obtain S where S: "\i. i \ j \ S i \ E i" "\i. i \ j \ countable (S i)" "\i. i \ j \ \ i = \(S i)" by (metis subset_eq \_cover \j \ I\) - def A' \ "\n. n(i := A)" + define A' where "A' n = n(i := A)" for n then have A'_i: "\n. A' n i = A" by simp { fix n assume "n \ Pi\<^sub>E (j - {i}) S" @@ -804,10 +804,12 @@ shows "P = Q" proof (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space]) interpret finite_measure P by fact - def i \ "SOME i. i \ I" + define i where "i = (SOME i. i \ I)" have i: "I \ {} \ i \ I" unfolding i_def by (rule someI_ex) auto - def A \ "\n::nat. if I = {} then prod_emb I M {} (\\<^sub>E i\{}. {}) else prod_emb I M {i} (\\<^sub>E i\{i}. space (M i))" + define A where "A n = + (if I = {} then prod_emb I M {} (\\<^sub>E i\{}. {}) else prod_emb I M {i} (\\<^sub>E i\{i}. space (M i)))" + for n :: nat then show "range A \ prod_algebra I M" using prod_algebraI[of "{}" I "\i. space (M i)" M] by (auto intro!: prod_algebraI i) have "\i. A i = space (PiM I M)" @@ -954,7 +956,7 @@ proof (rule measure_eqI_PiM_finite[OF I refl P, symmetric]) show "(\i. i \ I \ A i \ sets (M i)) \ (Pi\<^sub>M I M) (Pi\<^sub>E I A) = P (Pi\<^sub>E I A)" for A by (simp add: eq emeasure_PiM) - def A \ "\n. \\<^sub>E i\I. C i n" + define A where "A n = (\\<^sub>E i\I. C i n)" for n with C show "range A \ prod_algebra I M" "\i. emeasure (Pi\<^sub>M I M) (A i) \ \" "(\i. A i) = space (PiM I M)" by (auto intro!: prod_algebraI_finite simp: emeasure_PiM subset_eq ennreal_setprod_eq_top) qed diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Giry_Monad.thy --- a/src/HOL/Probability/Giry_Monad.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Giry_Monad.thy Mon Apr 25 22:59:53 2016 +0200 @@ -349,7 +349,8 @@ have [measurable]: "F i \ N \\<^sub>M count_space UNIV" for i using F(1) by (rule measurable_simple_function) - def F' \ "\M i. if integrable M f then integral\<^sup>L M (F i) else 0" + define F' where [abs_def]: + "F' M i = (if integrable M f then integral\<^sup>L M (F i) else 0)" for M i have "(\M. F' M i) \ subprob_algebra N \\<^sub>M borel" for i proof (rule measurable_cong[THEN iffD2]) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Helly_Selection.thy --- a/src/HOL/Probability/Helly_Selection.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Helly_Selection.thy Mon Apr 25 22:59:53 2016 +0200 @@ -43,7 +43,7 @@ thus "\s'. subseq s' \ (\l. (\k. f (s (s' k)) (r n)) \ l)" by (auto simp: comp_def) qed - def d \ "nat.diagseq" + define d where "d = nat.diagseq" have subseq: "subseq d" unfolding d_def using nat.subseq_diagseq by auto have rat_cnv: "?P n d" for n @@ -86,7 +86,7 @@ then have nonempty: "{lim (?f n) |n. x < r n} \ {}" for x by auto - def F \ "\x. Inf {lim (?f n) |n. x < r n}" + define F where "F x = Inf {lim (?f n) |n. x < r n}" for x have F_eq: "ereal (F x) = (INF n:{n. x < r n}. ereal (lim (?f n)))" for x unfolding F_def by (subst ereal_Inf'[OF bdd_below nonempty]) (simp add: setcompr_eq_image) have mono_F: "mono F" @@ -160,7 +160,7 @@ assumes \: "tight \" "subseq s" shows "\r M. subseq r \ real_distribution M \ weak_conv_m (\ \ s \ r) M" proof - - def f \ "\k. cdf (\ (s k))" + define f where "f k = cdf (\ (s k))" for k interpret \: real_distribution "\ k" for k using \ unfolding tight_def by auto @@ -275,7 +275,8 @@ subseq: "\s \. subseq s \ real_distribution \ \ weak_conv_m (\ \ s) \ \ weak_conv_m (\ \ s) M" shows "weak_conv_m \ M" proof (rule ccontr) - def f \ "\n. cdf (\ n)" and F \ "cdf M" + define f where "f n = cdf (\ n)" for n + define F where "F = cdf M" assume "\ weak_conv_m \ M" then obtain x where x: "isCont F x" "\ (\n. f n x) \ F x" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Independent_Family.thy --- a/src/HOL/Probability/Independent_Family.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Independent_Family.thy Mon Apr 25 22:59:53 2016 +0200 @@ -132,7 +132,7 @@ then have "indep_sets (?G J) K" proof induct case (insert j J) - moreover def G \ "?G J" + moreover define G where "G = ?G J" ultimately have G: "indep_sets G K" "\i. i \ K \ G i \ events" and "j \ K" by (auto simp: indep_sets_def) let ?D = "{E\events. indep_sets (G(j := {E})) K }" @@ -464,7 +464,7 @@ qed } note L_inj = this - def k \ "\l. (SOME k. k \ K \ l \ L k)" + define k where "k l = (SOME k. k \ K \ l \ L k)" for l { fix x j l assume *: "j \ K" "l \ L j" have "k l = j" unfolding k_def proof (rule some_equality) @@ -1265,7 +1265,7 @@ shows "(\\<^sup>+\. (\i\I. X i \) \M) = (\i\I. \\<^sup>+\. X i \ \M)" proof cases assume "I \ {}" - def Y \ "\i \. if i \ I then X i \ else 0" + define Y where [abs_def]: "Y i \ = (if i \ I then X i \ else 0)" for i \ { fix i have "i \ I \ random_variable borel (X i)" using I(2) by (cases "i\I") (auto simp: indep_vars_def) } note rv_X = this @@ -1302,7 +1302,7 @@ and indep_vars_integrable: "integrable M (\\. (\i\I. X i \))" (is ?int) proof (induct rule: case_split) assume "I \ {}" - def Y \ "\i \. if i \ I then X i \ else 0" + define Y where [abs_def]: "Y i \ = (if i \ I then X i \ else 0)" for i \ { fix i have "i \ I \ random_variable borel (X i)" using I(2) by (cases "i\I") (auto simp: indep_vars_def) } note rv_X = this[measurable] diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Information.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1671,9 +1671,14 @@ assumes X: "simple_function M X" and Y: "simple_function M Y" shows "\(X ; X | Y) = \(X | Y)" proof - - def Py \ "\x. if x \ Y`space M then measure M (Y -` {x} \ space M) else 0" - def Pxy \ "\x. if x \ (\x. (X x, Y x))`space M then measure M ((\x. (X x, Y x)) -` {x} \ space M) else 0" - def Pxxy \ "\x. if x \ (\x. (X x, X x, Y x))`space M then measure M ((\x. (X x, X x, Y x)) -` {x} \ space M) else 0" + define Py where "Py x = (if x \ Y`space M then measure M (Y -` {x} \ space M) else 0)" for x + define Pxy where "Pxy x = + (if x \ (\x. (X x, Y x))`space M then measure M ((\x. (X x, Y x)) -` {x} \ space M) else 0)" + for x + define Pxxy where "Pxxy x = + (if x \ (\x. (X x, X x, Y x))`space M then measure M ((\x. (X x, X x, Y x)) -` {x} \ space M) + else 0)" + for x let ?M = "X`space M \ X`space M \ Y`space M" note XY = simple_function_Pair[OF X Y] diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Interval_Integral.thy --- a/src/HOL/Probability/Interval_Integral.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Interval_Integral.thy Mon Apr 25 22:59:53 2016 +0200 @@ -81,7 +81,7 @@ (auto simp: incseq_def PInf) next case (real b') - def d \ "b' - (if a = -\ then b' - 1 else real_of_ereal a)" + define d where "d = b' - (if a = -\ then b' - 1 else real_of_ereal a)" with \a < b\ have a': "0 < d" by (cases a) (auto simp: real) moreover diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Mon Apr 25 22:59:53 2016 +0200 @@ -195,7 +195,7 @@ then obtain a' where a'lea [arith]: "a' > a" and a_prop: "F a' - F a < epsilon / 2" by auto - def S' \ "{i. l i < r i}" + define S' where "S' = {i. l i < r i}" obtain S :: "nat set" where "S \ S'" and finS: "finite S" and Sprop: "{a'..b} \ (\i \ S. {l i<.. "\x. THE i. P i x" def undef \ "THE i::'a. False" + define f where "f x = (THE i. P i x)" for x + define undef where "undef = (THE i::'a. False)" { fix i x assume "x \ space M" "P i x" then have "f x = i" unfolding f_def using unique by auto } note f_eq = this diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Measure_Space.thy --- a/src/HOL/Probability/Measure_Space.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Measure_Space.thy Mon Apr 25 22:59:53 2016 +0200 @@ -567,7 +567,7 @@ by simp next assume "infinite I" - def L \ "\n. LEAST i. i \ I \ i \ n" + define L where "L n = (LEAST i. i \ I \ i \ n)" for n have L: "L n \ I \ n \ L n" for n unfolding L_def proof (rule LeastI_ex) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy --- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Mon Apr 25 22:59:53 2016 +0200 @@ -210,7 +210,8 @@ assumes u[measurable]: "u \ borel_measurable M" shows "\f. incseq f \ (\i. (\x. f i x < top) \ simple_function M (f i)) \ u = (SUP i. f i)" proof - - def f \ "\i x. real_of_int (floor (enn2real (min i (u x)) * 2^i)) / 2^i" + define f where [abs_def]: + "f i x = real_of_int (floor (enn2real (min i (u x)) * 2^i)) / 2^i" for i x have [simp]: "0 \ f i x" for i x by (auto simp: f_def intro!: divide_nonneg_nonneg mult_nonneg_nonneg enn2real_nonneg) @@ -424,7 +425,8 @@ assumes sf: "simple_function M f" "simple_function M g" and A: "A \ space M \ sets M" shows "simple_function M (\x. if x \ A then f x else g x)" (is "simple_function M ?IF") proof - - def F \ "\x. f -` {x} \ space M" and G \ "\x. g -` {x} \ space M" + define F where "F x = f -` {x} \ space M" for x + define G where "G x = g -` {x} \ space M" for x show ?thesis unfolding simple_function_def proof safe have "?IF ` space M \ f ` space M \ g ` space M" by auto @@ -1452,7 +1454,7 @@ assumes "f \ measurable M (count_space UNIV)" shows "(\\<^sup>+x. of_nat (f x) \M) = (\t. emeasure M {x\space M. t < f x})" proof - - def F \ "\i. {x\space M. i < f x}" + define F where "F i = {x\space M. i < f x}" for i with assms have [measurable]: "\i. F i \ sets M" by auto @@ -1637,7 +1639,7 @@ assumes f: "f \ measurable M (count_space UNIV)" shows "(\\<^sup>+ x. ennreal_of_enat (f x) \M) = (\t. emeasure M {x \ space M. t < f x})" proof - - def F \ "\i::nat. {x\space M. i < f x}" + define F where "F i = {x\space M. i < f x}" for i :: nat with assms have [measurable]: "\i. F i \ sets M" by auto diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Probability_Mass_Function.thy --- a/src/HOL/Probability/Probability_Mass_Function.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Probability_Mass_Function.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1138,7 +1138,9 @@ from qr obtain qr where qr: "\y z. (y, z) \ set_pmf qr \ S y z" and q': "q = map_pmf fst qr" and r: "r = map_pmf snd qr" by cases auto - def pr \ "bind_pmf pq (\xy. bind_pmf (cond_pmf qr {yz. fst yz = snd xy}) (\yz. return_pmf (fst xy, snd yz)))" + define pr where "pr = + bind_pmf pq (\xy. bind_pmf (cond_pmf qr {yz. fst yz = snd xy}) + (\yz. return_pmf (fst xy, snd yz)))" have pr_welldefined: "\y. y \ q \ qr \ {yz. fst yz = y} \ {}" by (force simp: q') diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Probability_Measure.thy --- a/src/HOL/Probability/Probability_Measure.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Probability_Measure.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1147,7 +1147,7 @@ show "Int_stable (range atMost :: real set set)" by (auto simp: Int_stable_def) have vimage_eq: "\a. (X -` {..a} \ space M) = {x\space M. X x \ a}" by auto - def A \ "\i::nat. {.. real i}" + define A where "A i = {.. real i}" for i :: nat then show "range A \ range atMost" "(\i. A i) = space lborel" "\i. emeasure (distr M lborel X) (A i) \ \" by (auto simp: real_arch_simple emeasure_distr vimage_eq M_eq) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Projective_Family.thy --- a/src/HOL/Probability/Projective_Family.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Projective_Family.thy Mon Apr 25 22:59:53 2016 +0200 @@ -441,7 +441,8 @@ proof (rule PF.emeasure_lim[OF J subset_UNIV X]) fix J X' assume J[simp]: "\i. finite (J i)" and X'[measurable]: "\i. X' i \ sets (Pi\<^sub>M (J i) M)" and dec: "decseq (\i. PF.emb UNIV (J i) (X' i))" - def X \ "\n. (\i\{i. J i \ {0..< n}}. PF.emb {0.. space (PiM {0..i\{i. J i \ {0..< n}}. PF.emb {0.. space (PiM {0.. sets (PiM {0.. {0..< n}} = {}") @@ -605,7 +606,8 @@ assume inf: "infinite (\i. J i)" moreover have count: "countable (\i. J i)" using 1(3) by (auto intro: countable_finite) - def f \ "from_nat_into (\i. J i)" and t \ "to_nat_on (\i. J i)" + define f where "f = from_nat_into (\i. J i)" + define t where "t = to_nat_on (\i. J i)" have ft[simp]: "x \ J i \ f (t x) = x" for x i unfolding f_def t_def using inf count by (intro from_nat_into_to_nat_on) auto have tf[simp]: "t (f i) = i" for i diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Projective_Limit.thy --- a/src/HOL/Probability/Projective_Limit.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Projective_Limit.thy Mon Apr 25 22:59:53 2016 +0200 @@ -117,7 +117,7 @@ using J by (auto intro!: INF_lower2[of 0] prob_space_P[THEN prob_space.measure_le_1]) ultimately obtain r where r: "?a = ennreal r" "0 < r" "r \ 1" by (cases "?a") (auto simp: top_unique) - def Z \ "\n. emb I (J n) (B n)" + define Z where "Z n = emb I (J n) (B n)" for n have Z_mono: "n \ m \ Z m \ Z n" for n m unfolding Z_def using B[THEN antimonoD, of n m] . have J_mono: "\n m. n \ m \ J n \ J m" @@ -133,13 +133,13 @@ unfolding Z_def by (auto intro!: generator.intros J) have countable_UN_J: "countable (\n. J n)" by (simp add: countable_finite) - def Utn \ "to_nat_on (\n. J n)" + define Utn where "Utn = to_nat_on (\n. J n)" interpret function_to_finmap "J n" Utn "from_nat_into (\n. J n)" for n by unfold_locales (auto simp: Utn_def intro: from_nat_into_to_nat_on[OF countable_UN_J]) have inj_on_Utn: "inj_on Utn (\n. J n)" unfolding Utn_def using countable_UN_J by (rule inj_on_to_nat_on) hence inj_on_Utn_J: "\n. inj_on Utn (J n)" by (rule subset_inj_on) auto - def P' \ "\n. mapmeasure n (P (J n)) (\_. borel)" + define P' where "P' n = mapmeasure n (P (J n)) (\_. borel)" for n interpret P': prob_space "P' n" for n unfolding P'_def mapmeasure_def using J by (auto intro!: prob_space_distr fm_measurable simp: measurable_cong_sets[OF sets_P]) @@ -189,7 +189,7 @@ "\n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) \ ennreal (2 powr - real n) * ?a" "\n. compact (K' n)" "\n. K' n \ fm n ` B n" unfolding choice_iff by blast - def K \ "\n. fm n -` K' n \ space (Pi\<^sub>M (J n) (\_. borel))" + define K where "K n = fm n -` K' n \ space (Pi\<^sub>M (J n) (\_. borel))" for n have K_sets: "\n. K n \ sets (Pi\<^sub>M (J n) (\_. borel))" unfolding K_def using compact_imp_closed[OF \compact (K' _)\] @@ -204,7 +204,7 @@ using \x \ K n\ K_sets sets.sets_into_space J(1,2,3)[of n] inj_on_image_mem_iff[OF inj_on_fm] by (metis (no_types) Int_iff K_def fm_in space_borel) qed - def Z' \ "\n. emb I (J n) (K n)" + define Z' where "Z' n = emb I (J n) (K n)" for n have Z': "\n. Z' n \ Z n" unfolding Z'_def Z_def proof (rule prod_emb_mono, safe) @@ -228,7 +228,7 @@ have "\n. Z' n \ generator" using J K'(2) unfolding Z'_def by (auto intro!: generator.intros measurable_sets[OF fm_measurable[of _ "Collect finite"]] simp: K_def borel_eq_PiF_borel[symmetric] compact_imp_closed) - def Y \ "\n. \i\{1..n}. Z' i" + define Y where "Y n = (\i\{1..n}. Z' i)" for n hence "\n k. Y (n + k) \ Y n" by (induct_tac k) (auto simp: Y_def) hence Y_mono: "\n m. n \ m \ Y m \ Y n" by (auto simp: le_iff_add) have Y_Z': "\n. n \ 1 \ Y n \ Z' n" by (auto simp: Y_def) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Radon_Nikodym.thy --- a/src/HOL/Probability/Radon_Nikodym.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Radon_Nikodym.thy Mon Apr 25 22:59:53 2016 +0200 @@ -153,7 +153,7 @@ let ?A = "\A. if (\B\sets M. B \ space M - A \ -e < ?d B) then {} else (SOME B. B \ sets M \ B \ space M - A \ ?d B \ -e)" - def A \ "\n. ((\B. B \ ?A B) ^^ n) {}" + define A where "A n = ((\B. B \ ?A B) ^^ n) {}" for n have A_simps[simp]: "A 0 = {}" "\n. A (Suc n) = (A n \ ?A (A n))" unfolding A_def by simp_all @@ -296,7 +296,8 @@ shows "\f \ borel_measurable M. (\x. 0 \ f x) \ density M f = N" proof - interpret N: finite_measure N by fact - def G \ "{g \ borel_measurable M. (\x. 0 \ g x) \ (\A\sets M. (\\<^sup>+x. g x * indicator A x \M) \ N A)}" + define G where "G = + {g \ borel_measurable M. (\x. 0 \ g x) \ (\A\sets M. (\\<^sup>+x. g x * indicator A x \M) \ N A)}" { fix f have "f \ G \ f \ borel_measurable M" by (auto simp: G_def) } note this[measurable_dest] have "(\x. 0) \ G" unfolding G_def by auto @@ -366,7 +367,7 @@ from choice[OF this] obtain gs where "\i. gs i \ G" "\n. integral\<^sup>N M (gs n) = ys n" by auto hence y_eq: "?y = (SUP i. integral\<^sup>N M (gs i))" using ys by auto let ?g = "\i x. Max ((\n. gs n x) ` {..i})" - def f \ "\x. SUP i. ?g i x" + define f where [abs_def]: "f x = (SUP i. ?g i x)" for x let ?F = "\A x. f x * indicator A x" have gs_not_empty: "\i x. (\n. gs n x) ` {..i} \ {}" by auto { fix i have "?g i \ G" @@ -443,7 +444,7 @@ hence "(\\<^sup>+x. f x * indicator (space M) x \M) \ \" using M'.finite_emeasure_space by (auto simp: top_unique) moreover - def b \ "?M (space M) / emeasure M (space M) / 2" + define b where "b = ?M (space M) / emeasure M (space M) / 2" ultimately have b: "b \ 0 \ 0 \ b \ b \ \" by (auto simp: ennreal_divide_eq_top_iff) then have b: "b \ 0" "0 \ b" "0 < b" "b \ \" @@ -584,7 +585,7 @@ qed let ?O_0 = "(\i. ?O i)" have "?O_0 \ sets M" using Q' by auto - def Q \ "\i. case i of 0 \ Q' 0 | Suc n \ ?O (Suc n) - ?O n" + define Q where "Q i = (case i of 0 \ Q' 0 | Suc n \ ?O (Suc n) - ?O n)" for i { fix i have "Q i \ sets M" unfolding Q_def using Q'[of 0] by (cases i) (auto intro: O_sets) } note Q_sets = this show ?thesis @@ -990,7 +991,8 @@ next assume AE: "AE x in M. f x \ \" from sigma_finite guess Q . note Q = this - def A \ "\i. f -` (case i of 0 \ {\} | Suc n \ {.. ennreal(of_nat (Suc n))}) \ space M" + define A where "A i = + f -` (case i of 0 \ {\} | Suc n \ {.. ennreal(of_nat (Suc n))}) \ space M" for i { fix i j have "A i \ Q j \ sets M" unfolding A_def using f Q apply (rule_tac sets.Int) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Regularity.thy --- a/src/HOL/Probability/Regularity.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Regularity.thy Mon Apr 25 22:59:53 2016 +0200 @@ -76,7 +76,8 @@ "\K \ {K. K \ space M \ compact K}. emeasure M (space M) \ emeasure M K + ennreal e" (is "?thesis e") if "0 < e" for e :: real proof - - def B \ "\n. \i\{0..k e (Suc n)}. cball (from_nat_into X i) (1 / Suc n)" + define B where [abs_def]: + "B n = (\i\{0..k e (Suc n)}. cball (from_nat_into X i) (1 / Suc n))" for n have "\n. closed (B n)" by (auto simp: B_def) hence [simp]: "\n. B n \ sets M" by (simp add: sb) from k[OF \e > 0\ zero_less_Suc] @@ -84,7 +85,7 @@ by (simp add: algebra_simps B_def finite_measure_compl) hence B_compl_le: "\n::nat. measure M (space M - B n) \ e * 2 powr - real (Suc n)" by (simp add: finite_measure_compl) - def K \ "\n. B n" + define K where "K = (\n. B n)" from \closed (B _)\ have "closed K" by (auto simp: K_def) hence [simp]: "K \ sets M" by (simp add: sb) have "measure M (space M) - measure M K = measure M (space M - K)" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Sigma_Algebra.thy Mon Apr 25 22:59:53 2016 +0200 @@ -846,7 +846,7 @@ proof - from a guess Ca .. note Ca = this from b guess Cb .. note Cb = this - def C \ "(\(a,b). a \ b)` (Ca\Cb)" + define C where "C = (\(a,b). a \ b)` (Ca\Cb)" show ?thesis proof show "disjoint C" @@ -1971,7 +1971,7 @@ with \i\I\ * show ?thesis by simp next - def P \ "\\'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \'" + define P where "P \' \ (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \'" for \' assume "\ (\i\I. \ i = 0)" moreover have "measure_space (space M) (sets M) \'" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/Weak_Convergence.thy --- a/src/HOL/Probability/Weak_Convergence.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/Weak_Convergence.thy Mon Apr 25 22:59:53 2016 +0200 @@ -212,11 +212,11 @@ using emeasure_lborel_countable[OF D_countable] by (subst emeasure_restrict_space) auto - def Y' \ "\\. if \ \ ?D then 0 else M.I \" + define Y' where "Y' \ = (if \ \ ?D then 0 else M.I \)" for \ have Y'_AE: "AE \ in ?\. Y' \ = M.I \" by (rule AE_I [OF _ D]) (auto simp: space_restrict_space sets_restrict_space_iff Y'_def) - def Y_seq' \ "\n \. if \ \ ?D then 0 else \.I n \" + define Y_seq' where "Y_seq' n \ = (if \ \ ?D then 0 else \.I n \)" for n \ have Y_seq'_AE: "\n. AE \ in ?\. Y_seq' n \ = \.I n \" by (rule AE_I [OF _ D]) (auto simp: space_restrict_space sets_restrict_space_iff Y_seq'_def) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/ex/Dining_Cryptographers.thy --- a/src/HOL/Probability/ex/Dining_Cryptographers.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy Mon Apr 25 22:59:53 2016 +0200 @@ -64,7 +64,7 @@ have foldl_coin: "\ ?XOR (\c. coin dc c \ coin dc (c + 1)) n" proof - - def n' \ n \ "Need to hide n, as it is hidden in coin" + define n' where "n' = n" \ "Need to hide n, as it is hidden in coin" have "?XOR (\c. coin dc c \ coin dc (c + 1)) n' = (coin dc 0 \ coin dc n')" by (induct n') auto @@ -81,7 +81,7 @@ next assume "\k n \ "Need to hide n, as it is hidden in coin, payer etc." + define l where "l = n" \ "Need to hide n, as it is hidden in coin, payer etc." have "?XOR (\c. (payer dc = Some c) \ (coin dc c \ coin dc (c + 1))) l = ((k < l) \ ?XOR (\c. (coin dc c \ coin dc (c + 1))) l)" using \payer dc = Some k\ by (induct l) auto @@ -156,7 +156,7 @@ @{term i}. \ - def zs \ "map (\p. if p \ {min i j<..max i j} then \ ys ! p else ys ! p) [0..p. if p \ {min i j<..max i j} then \ ys ! p else ys ! p) [0.. inversion`dc_crypto" and i: "Some i \ Some ` {0.. "inversion (Some i, xs)" + moreover define ys where "ys = inversion (Some i, xs)" ultimately have ys: "ys \ inversion`dc_crypto" and "Some i \ Some ` {0..x. (inversion x, payer x)) -` {x} \ space (uniform_count_measure dc_crypto) = diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Probability/ex/Measure_Not_CCC.thy --- a/src/HOL/Probability/ex/Measure_Not_CCC.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Probability/ex/Measure_Not_CCC.thy Mon Apr 25 22:59:53 2016 +0200 @@ -106,7 +106,7 @@ unfolding fun_eq_iff using Union.IH by metis show ?case proof (intro exI conjI) - def G \ "\j. (\i. if j \ J i then F i j else X i)" + define G where "G j = (\i. if j \ J i then F i j else X i)" for j show "(\i. X i) \ sets M" "countable (\i. J i)" "G \ (\i. J i) \ sets M" using XFJ by (auto simp: G_def Pi_iff) show "UNION UNIV A = (UNIV - (\i. J i)) \ (\i. X i) \ (SIGMA j:\i. J i. \i. if j \ J i then F i j else X i)" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Real.thy --- a/src/HOL/Real.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Real.thy Mon Apr 25 22:59:53 2016 +0200 @@ -775,7 +775,7 @@ obtain x where x: "x \ S" using assms(1) .. obtain z where z: "\x\S. x \ z" using assms(2) .. - def P \ "\x. \y\S. y \ of_rat x" + define P where "P x \ (\y\S. y \ of_rat x)" for x obtain a where a: "\ P a" proof have "of_int \x - 1\ \ x - 1" by (rule of_int_floor_le) @@ -797,11 +797,11 @@ qed qed - def avg \ "\x y :: rat. x/2 + y/2" - def bisect \ "\(x, y). if P (avg x y) then (x, avg x y) else (avg x y, y)" - def A \ "\n. fst ((bisect ^^ n) (a, b))" - def B \ "\n. snd ((bisect ^^ n) (a, b))" - def C \ "\n. avg (A n) (B n)" + define avg where "avg x y = x/2 + y/2" for x y :: rat + define bisect where "bisect = (\(x, y). if P (avg x y) then (x, avg x y) else (avg x y, y))" + define A where "A n = fst ((bisect ^^ n) (a, b))" for n + define B where "B n = snd ((bisect ^^ n) (a, b))" for n + define C where "C n = avg (A n) (B n)" for n have A_0 [simp]: "A 0 = a" unfolding A_def by simp have B_0 [simp]: "B 0 = b" unfolding B_def by simp have A_Suc [simp]: "\n. A (Suc n) = (if P (C n) then A n else C n)" @@ -1251,8 +1251,8 @@ from \x have "0 < y-x" by simp with reals_Archimedean obtain q::nat where q: "inverse (real q) < y-x" and "0 < q" by blast - def p \ "\y * real q\ - 1" - def r \ "of_int p / real q" + define p where "p = \y * real q\ - 1" + define r where "r = of_int p / real q" from q have "x < y - inverse (real q)" by simp also have "y - inverse (real q) \ r" unfolding r_def p_def diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Real_Vector_Spaces.thy Mon Apr 25 22:59:53 2016 +0200 @@ -2070,7 +2070,7 @@ assumes X: "Cauchy X" shows "convergent X" proof - - def S \ "{x::real. \N. \n\N. x < X n}" + define S :: "real set" where "S = {x. \N. \n\N. x < X n}" then have mem_S: "\N x. \n\N. x < X n \ x \ S" by auto { fix N x assume N: "\n\N. X n < x" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Rings.thy --- a/src/HOL/Rings.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Rings.thy Mon Apr 25 22:59:53 2016 +0200 @@ -847,7 +847,7 @@ and "is_unit b" and "1 div a = b" and "1 div b = a" and "a * b = 1" and "c div a = c * b" proof (rule that) - def b \ "1 div a" + define b where "b = 1 div a" then show "1 div a = b" by simp from b_def \is_unit a\ show "is_unit b" by simp from \is_unit a\ and \is_unit b\ show "a \ 0" and "b \ 0" by auto diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Series.thy --- a/src/HOL/Series.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Series.thy Mon Apr 25 22:59:53 2016 +0200 @@ -84,7 +84,7 @@ shows "summable f = summable g" proof - from assms obtain N where N: "\n\N. f n = g n" by (auto simp: eventually_at_top_linorder) - def C \ "(\kkn. setsum f {..n. \k c" . next assume lim: "(\n. \k c" - def g_inv \ "\n. LEAST m. g m \ n" + define g_inv where "g_inv n = (LEAST m. g m \ n)" for n from filterlim_subseq[OF subseq] have g_inv_ex: "\m. g m \ n" for n by (auto simp: filterlim_at_top eventually_at_top_linorder) hence g_inv: "g (g_inv n) \ n" for n unfolding g_inv_def by (rule LeastI_ex) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Tools/Lifting/lifting_setup.ML --- a/src/HOL/Tools/Lifting/lifting_setup.ML Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Tools/Lifting/lifting_setup.ML Mon Apr 25 22:59:53 2016 +0200 @@ -93,7 +93,7 @@ ((Binding.empty, []), definition_term)) lthy |>> (snd #> snd); fun raw_def lthy = let - val ((_, rhs), prove) = Local_Defs.derived_def lthy true definition_term; + val ((_, rhs), prove) = Local_Defs.derived_def lthy {conditional = true} definition_term; val ((_, (_, raw_th)), lthy) = lthy |> Local_Theory.define ((Binding.concealed pcrel_name, NoSyn), ((Binding.empty, []), rhs)); val th = prove lthy raw_th; diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Tools/inductive_set.ML --- a/src/HOL/Tools/inductive_set.ML Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Tools/inductive_set.ML Mon Apr 25 22:59:53 2016 +0200 @@ -469,7 +469,7 @@ val (defs, lthy2) = lthy1 |> fold_map Local_Theory.define (map (fn (((b, mx), (fs, U, _)), p) => - ((b, mx), ((Binding.reset_pos (Thm.def_binding b), []), + ((b, mx), ((Thm.def_binding b, []), fold_rev lambda params (HOLogic.Collect_const U $ HOLogic.mk_ptupleabs fs U HOLogic.boolT (list_comb (p, params3)))))) (cnames_syn ~~ cs_info ~~ preds)); diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Topological_Spaces.thy --- a/src/HOL/Topological_Spaces.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Topological_Spaces.thy Mon Apr 25 22:59:53 2016 +0200 @@ -1183,7 +1183,7 @@ from first_countable_basis[of x] obtain A :: "nat \ 'a set" where nhds: "\i. open (A i)" "\i. x \ A i" and incl: "\S. open S \ x \ S \ \i. A i \ S" by auto - def F \ "\n. \i\n. A i" + define F where "F n = (\i\n. A i)" for n show "\A. (\i. open (A i)) \ (\i. x \ A i) \ (\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially)" proof (safe intro!: exI[of _ F]) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Transcendental.thy Mon Apr 25 22:59:53 2016 +0200 @@ -893,7 +893,7 @@ from \0 < r\ have "0 < ?r" by simp let ?s = "\n. SOME s. 0 < s \ (\ x. x \ 0 \ \ x \ < s \ \ ?diff n x - f' x0 n \ < ?r)" - def S' \ "Min (?s ` {..< ?N })" + define S' where "S' = Min (?s ` {..< ?N })" have "0 < S'" unfolding S'_def proof (rule iffD2[OF Min_gr_iff]) @@ -911,7 +911,7 @@ qed qed auto - def S \ "min (min (x0 - a) (b - x0)) S'" + define S where "S = min (min (x0 - a) (b - x0)) S'" hence "0 < S" and S_a: "S \ x0 - a" and S_b: "S \ b - x0" and "S \ S'" using x0_in_I and \0 < S'\ by auto @@ -2248,7 +2248,7 @@ assumes "x > 0" shows "DERIV (\y. log b y) x :> 1 / (ln b * x)" proof - - def lb \ "1 / ln b" + define lb where "lb = 1 / ln b" moreover have "DERIV (\y. lb * ln y) x :> lb / x" using \x > 0\ by (auto intro!: derivative_eq_intros) ultimately show ?thesis @@ -4769,7 +4769,7 @@ proof (rule tendstoI) fix e :: real assume "0 < e" - def y \ "pi/2 - min (pi/2) e" + define y where "y = pi/2 - min (pi/2) e" then have y: "0 \ y" "y < pi/2" "pi/2 \ e + y" using \0 < e\ by auto diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Wfrec.thy --- a/src/HOL/Wfrec.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Wfrec.thy Mon Apr 25 22:59:53 2016 +0200 @@ -37,7 +37,7 @@ lemma wfrec_unique: assumes "adm_wf R F" "wf R" shows "\!y. wfrec_rel R F x y" using \wf R\ proof induct - def f \ "\y. THE z. wfrec_rel R F y z" + define f where "f y = (THE z. wfrec_rel R F y z)" for y case (less x) then have "\y z. (y, x) \ R \ wfrec_rel R F y z \ z = f y" unfolding f_def by (rule theI_unique) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/Zorn.thy --- a/src/HOL/Zorn.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/Zorn.thy Mon Apr 25 22:59:53 2016 +0200 @@ -609,7 +609,7 @@ proof - \ \The initial segment relation on well-orders:\ let ?WO = "{r::'a rel. Well_order r}" - def I \ "init_seg_of \ ?WO \ ?WO" + define I where "I = init_seg_of \ ?WO \ ?WO" have I_init: "I \ init_seg_of" by (auto simp: I_def) hence subch: "\R. R \ Chains I \ chain\<^sub>\ R" unfolding init_seg_of_def chain_subset_def Chains_def by blast @@ -732,7 +732,7 @@ shows "\f. \x. P f x (f x)" proof (intro exI allI) fix x - def f \ "wfrec R (\f x. SOME r. P f x r)" + define f where "f \ wfrec R (\f x. SOME r. P f x r)" from \wf R\ show "P f x (f x)" proof (induct x) fix x assume "\y. (y, x) \ R \ P f y (f y)" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/ex/Ballot.thy --- a/src/HOL/ex/Ballot.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/ex/Ballot.thy Mon Apr 25 22:59:53 2016 +0200 @@ -77,13 +77,14 @@ card {V\Pow {0.. (\m\{1..a+b}. card ({0.. V) > m - card ({0.. V))}" (is "_ = card ?V") proof - - def P \ "\j i. i < a + b \ j = Suc i" + define P where "P j i \ i < a + b \ j = Suc i" for j i have unique_P: "bi_unique P" and total_P: "\m. m \ a + b \ rel_set P {1..m} {0..R f g. (\i. i < a+b \ R (f (Suc i)) (g i)) \ rel_fun P R f g" by (simp add: rel_fun_def P_def) - def R \ "\f V. V \ {0.. f \ extensional {1..a+b} \ (\i V \ f (Suc i) = A)" + define R where "R f V \ + V \ {0.. f \ extensional {1..a+b} \ (\i V \ f (Suc i) = A)" for f V { fix f g :: "nat \ vote" assume "f \ extensional {1..a + b}" "g \ extensional {1..a + b}" moreover assume "\iiS. S \ {0..m * n} \ card S = n + 1 \ mono_on f (op \) S)" proof (rule ccontr) let ?max_subseq = "\R k. Max (card ` {S. S \ {0..k} \ mono_on f R S \ k \ S})" - def phi == "\k. (?max_subseq (op \) k, ?max_subseq (op \) k)" + define phi where "phi k = (?max_subseq (op \) k, ?max_subseq (op \) k)" for k have one_member: "\R k. reflp R \ {k} \ {S. S \ {0..k} \ mono_on f R S \ k \ S}" by auto diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/ex/Gauge_Integration.thy --- a/src/HOL/ex/Gauge_Integration.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/ex/Gauge_Integration.thy Mon Apr 25 22:59:53 2016 +0200 @@ -379,9 +379,10 @@ and I2: "\ D. fine \2 (b,c) D \ \ rsum D f - x2 \ < (\ / 2)" using IntegralE [OF \Integral (b, c) f x2\ \0 < \/2\] by auto - def \ \ "\ x. if x < b then min (\1 x) (b - x) - else if x = b then min (\1 b) (\2 b) - else min (\2 x) (x - b)" + define \ where "\ x = + (if x < b then min (\1 x) (b - x) + else if x = b then min (\1 b) (\2 b) + else min (\2 x) (x - b))" for x have "gauge {a..c} \" using \1_gauge \2_gauge unfolding \_def gauge_def by auto @@ -410,8 +411,8 @@ let ?D1 = "take N D" let ?D2 = "drop N D" - def D1 \ "take N D @ [(d, t, b)]" - def D2 \ "(if b = e then [] else [(b, t, e)]) @ drop (Suc N) D" + define D1 where "D1 = take N D @ [(d, t, b)]" + define D2 where "D2 = (if b = e then [] else [(b, t, e)]) @ drop (Suc N) D" from hd_drop_conv_nth[OF \N < length D\] have "fst (hd ?D2) = d" using \D ! N = (d, t, e)\ by auto diff -r 1836456b7d82 -r 2cc4e85b46d4 src/HOL/ex/HarmonicSeries.thy --- a/src/HOL/ex/HarmonicSeries.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/HOL/ex/HarmonicSeries.thy Mon Apr 25 22:59:53 2016 +0200 @@ -279,7 +279,7 @@ assume sf: "summable ?f" then obtain n::nat where ndef: "n = nat \2 * ?s\" by simp then have ngt: "1 + real n/2 > ?s" by linarith - def j \ "(2::nat)^n" + define j where "j = (2::nat)^n" have "\m\j. 0 < ?f m" by simp with sf have "(\ii\{Suc 0.. thm list -> tactic val fold: Proof.context -> thm list -> thm -> thm val fold_tac: Proof.context -> thm list -> tactic - val derived_def: Proof.context -> bool -> term -> + val derived_def: Proof.context -> {conditional: bool} -> term -> ((string * typ) * term) * (Proof.context -> thm -> thm) end; @@ -46,7 +46,11 @@ quote (Syntax.string_of_term ctxt eq)); val ((lhs, _), eq') = eq |> Sign.no_vars ctxt - |> Primitive_Defs.dest_def ctxt Term.is_Free (Variable.is_fixed ctxt) (K true) + |> Primitive_Defs.dest_def ctxt + {check_head = Term.is_Free, + check_free_lhs = not o Variable.is_fixed ctxt, + check_free_rhs = if Variable.is_body ctxt then K true else Variable.is_fixed ctxt, + check_tfree = K true} handle TERM (msg, _) => err msg | ERROR msg => err msg; in (Term.dest_Free (Term.head_of lhs), eq') end; @@ -206,7 +210,7 @@ (* derived defs -- potentially within the object-logic *) -fun derived_def ctxt conditional prop = +fun derived_def ctxt {conditional} prop = let val ((c, T), rhs) = prop |> Thm.cterm_of ctxt @@ -215,7 +219,8 @@ |> conditional ? Logic.strip_imp_concl |> (abs_def o #2 o cert_def ctxt); fun prove ctxt' def = - Goal.prove ctxt' (Variable.add_free_names ctxt' prop []) [] prop + Goal.prove ctxt' + ((not (Variable.is_body ctxt') ? Variable.add_free_names ctxt' prop) []) [] prop (fn {context = ctxt'', ...} => ALLGOALS (CONVERSION (meta_rewrite_conv ctxt'') THEN' diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Pure/Isar/obtain.ML --- a/src/Pure/Isar/obtain.ML Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Pure/Isar/obtain.ML Mon Apr 25 22:59:53 2016 +0200 @@ -193,47 +193,28 @@ local -fun gen_obtain prep_att prep_var prep_propp - that_binding raw_vars raw_asms int state = +fun gen_obtain prep_spec prep_att that_binding raw_vars raw_asms int state = let val _ = Proof.assert_forward_or_chain state; - val ctxt = Proof.context_of state; - (*vars*) - val ((_, thesis), thesis_ctxt) = obtain_thesis ctxt; - val ((xs', vars), fix_ctxt) = thesis_ctxt - |> fold_map prep_var raw_vars - |-> (fn vars => Proof_Context.add_fixes vars ##>> pair vars); - val xs = map (Variable.check_name o #1) vars; + val ((_, thesis), thesis_ctxt) = obtain_thesis (Proof.context_of state); - (*asms*) - val (propss, binds) = prep_propp fix_ctxt (map snd raw_asms); - val props = flat propss; - val declare_asms = - fold Variable.declare_term props #> - fold Variable.bind_term binds; + val (((vars, xs, params), propss, binds, binds'), params_ctxt) = + prep_spec raw_vars (map #2 raw_asms) thesis_ctxt; + val cparams = map (Thm.cterm_of params_ctxt) params; val asms = - map (fn ((b, raw_atts), _) => (b, map (prep_att fix_ctxt) raw_atts)) raw_asms ~~ + map (fn ((b, raw_atts), _) => (b, map (prep_att params_ctxt) raw_atts)) raw_asms ~~ map (map (rpair [])) propss; - (*params*) - val (params, params_ctxt) = - declare_asms fix_ctxt |> fold_map Proof_Context.inferred_param xs' |>> map Free; - val cparams = map (Thm.cterm_of params_ctxt) params; - val binds' = (map o apsnd) (fold_rev Term.dependent_lambda_name (xs ~~ params)) binds; - - val _ = Variable.warn_extra_tfrees fix_ctxt params_ctxt; - - (*statements*) val that_prop = Logic.list_rename_params xs - (fold_rev Logic.all params (Logic.list_implies (props, thesis))); + (fold_rev Logic.all params (Logic.list_implies (flat propss, thesis))); fun after_qed (result_ctxt, results) state' = let val [rule] = Proof_Context.export result_ctxt (Proof.context_of state') (flat results) in state' |> Proof.fix vars - |> Proof.map_context declare_asms + |> Proof.map_context (fold Variable.bind_term binds) |> Proof.assm (obtain_export params_ctxt rule cparams) [] [] asms end; in @@ -248,8 +229,8 @@ in -val obtain = gen_obtain (K I) Proof_Context.cert_var Proof_Context.cert_propp; -val obtain_cmd = gen_obtain Attrib.attribute_cmd Proof_Context.read_var Proof_Context.read_propp; +val obtain = gen_obtain Proof_Context.cert_spec (K I); +val obtain_cmd = gen_obtain Proof_Context.read_spec Attrib.attribute_cmd; end; diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Pure/Isar/proof.ML --- a/src/Pure/Isar/proof.ML Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Pure/Isar/proof.ML Mon Apr 25 22:59:53 2016 +0200 @@ -86,6 +86,12 @@ val unfolding_cmd: ((Facts.ref * Token.src list) list) list -> state -> state val case_: Thm.binding * ((string * Position.T) * binding option list) -> state -> state val case_cmd: Attrib.binding * ((string * Position.T) * binding option list) -> state -> state + val define: (binding * typ option * mixfix) list -> + (binding * typ option * mixfix) list -> + (Thm.binding * (term * term list) list) list -> state -> state + val define_cmd: (binding * string option * mixfix) list -> + (binding * string option * mixfix) list -> + (Attrib.binding * (string * string list) list) list -> state -> state val begin_block: state -> state val next_block: state -> state val end_block: state -> state @@ -740,7 +746,7 @@ #-> (fn rhss => let val defs = (vars ~~ (name_atts ~~ rhss)) |> map (fn ((x, mx), ((a, atts), rhs)) => - ((x, mx), ((Binding.reset_pos (Thm.def_binding_optional x a), atts), rhs))); + ((x, mx), ((Thm.def_binding_optional x a, atts), rhs))); in map_context_result (Local_Defs.add_defs defs) end)) |-> (set_facts o map (#2 o #2)) end; @@ -883,6 +889,70 @@ end; +(* define *) + +local + +fun match_defs (((b, _, mx), x) :: more_decls) ((((a, _), t), _) :: more_defs) = + if x = a then ((b, mx), (Thm.empty_binding, t)) :: match_defs more_decls more_defs + else error ("Mismatch of declaration " ^ quote x ^ " wrt. definition " ^ quote a) + | match_defs [] [] = [] + | match_defs more_decls more_defs = + error ("Mismatch of declarations " ^ commas_quote (map #2 more_decls) ^ + (if null more_decls then "" else " ") ^ + "wrt. definitions " ^ commas_quote (map (#1 o #1 o #1) more_defs)); + +fun invisible_fixes vars ctxt = ctxt + |> Context_Position.set_visible false + |> Proof_Context.add_fixes vars |> #2 + |> Context_Position.restore_visible ctxt; + +fun gen_define prep_spec prep_att raw_vars raw_fixes raw_defs state = + let + val _ = assert_forward state; + val ctxt = context_of state; + + (*vars*) + val n_vars = length raw_vars; + val (((all_vars, _, all_params), defss, _, binds'), _) = + prep_spec (raw_vars @ raw_fixes) (map snd raw_defs) ctxt; + val (vars, fixes) = chop n_vars all_vars; + val (params, _) = chop n_vars all_params; + + (*defs*) + val derived_def = Local_Defs.derived_def (invisible_fixes vars ctxt) {conditional = false}; + val defs1 = map derived_def (flat defss); + val defs2 = match_defs (vars ~~ map (#1 o dest_Free) params) defs1; + val (defs3, defs_ctxt) = Local_Defs.add_defs defs2 ctxt; + + (*fixes*) + val fixes_ctxt = invisible_fixes fixes defs_ctxt; + val export = singleton (Variable.export fixes_ctxt defs_ctxt); + + (*notes*) + val def_thmss = + map (fn (((_, prove), ((b, _), _)), (_, (_, th))) => (b, export (prove fixes_ctxt th))) + (defs1 ~~ defs2 ~~ defs3) + |> unflat (map snd raw_defs); + val notes = + map2 (fn ((a, raw_atts), _) => fn def_thms => + ((Thm.def_binding_optional (Binding.conglomerate (map #1 def_thms)) a, + map (prep_att defs_ctxt) raw_atts), map (fn (_, th) => ([th], [])) def_thms)) + raw_defs def_thmss; + in + state + |> map_context (K defs_ctxt #> fold Variable.bind_term binds') + |> note_thmss notes + end; + +in + +val define = gen_define Proof_Context.cert_spec (K I); +val define_cmd = gen_define Proof_Context.read_spec Attrib.attribute_cmd; + +end; + + (** proof structure **) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Pure/Isar/proof_context.ML --- a/src/Pure/Isar/proof_context.ML Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Pure/Isar/proof_context.ML Mon Apr 25 22:59:53 2016 +0200 @@ -166,6 +166,12 @@ val generic_add_abbrev: string -> binding * term -> Context.generic -> (term * term) * Context.generic val generic_revert_abbrev: string -> string -> Context.generic -> Context.generic + val cert_spec: (binding * typ option * mixfix) list -> (term * term list) list list -> + Proof.context -> (((binding * typ option * mixfix) list * string list * term list) * + term list list * (indexname * term) list * (indexname * term) list) * Proof.context + val read_spec: (binding * string option * mixfix) list -> (string * string list) list list -> + Proof.context -> (((binding * typ option * mixfix) list * string list * term list) * + term list list * (indexname * term) list * (indexname * term) list) * Proof.context val print_syntax: Proof.context -> unit val print_abbrevs: bool -> Proof.context -> unit val pretty_term_bindings: Proof.context -> Pretty.T list @@ -1320,6 +1326,42 @@ end; +(* specification with parameters *) + +local + +fun prep_spec prep_var prep_propp raw_vars raw_propps ctxt = + let + (*vars*) + val ((xs', vars), vars_ctxt) = ctxt + |> fold_map prep_var raw_vars + |-> (fn vars => add_fixes vars ##>> pair vars); + val xs = map (Variable.check_name o #1) vars; + + (*propps*) + val (propss, binds) = prep_propp vars_ctxt raw_propps; + val props = flat propss; + + (*params*) + val (ps, params_ctxt) = vars_ctxt + |> fold Variable.declare_term props + |> fold Variable.bind_term binds + |> fold_map inferred_param xs'; + val params = map Free ps; + val binds' = (map o apsnd) (fold_rev Term.dependent_lambda_name (xs ~~ params)) binds; + val vars' = map2 (fn (b, _, mx) => fn (_, T) => (b, SOME T, mx)) vars ps; + + val _ = Variable.warn_extra_tfrees vars_ctxt params_ctxt; + in (((vars', xs, params), propss, binds, binds'), params_ctxt) end; + +in + +val cert_spec = prep_spec cert_var cert_propp; +val read_spec = prep_spec read_var read_propp; + +end; + + (** print context information **) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Pure/Isar/specification.ML --- a/src/Pure/Isar/specification.ML Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Pure/Isar/specification.ML Mon Apr 25 22:59:53 2016 +0200 @@ -237,7 +237,7 @@ let val ((vars, [((raw_name, atts), prop)]), get_pos) = fst (prep (the_list raw_var) [raw_spec] lthy); - val (((x, T), rhs), prove) = Local_Defs.derived_def lthy true prop; + val (((x, T), rhs), prove) = Local_Defs.derived_def lthy {conditional = true} prop; val _ = Name.reject_internal (x, []); val (b, mx) = (case vars of @@ -249,7 +249,7 @@ error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote y ^ Position.here (Binding.pos_of b)); in (b, mx) end); - val name = Binding.reset_pos (Thm.def_binding_optional b raw_name); + val name = Thm.def_binding_optional b raw_name; val ((lhs, (_, raw_th)), lthy2) = lthy |> Local_Theory.define_internal ((b, mx), ((Binding.suffix_name "_raw" name, []), rhs)); diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Pure/Pure.thy --- a/src/Pure/Pure.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Pure/Pure.thy Mon Apr 25 22:59:53 2016 +0200 @@ -54,7 +54,7 @@ and "note" :: prf_decl % "proof" and "supply" :: prf_script % "proof" and "using" "unfolding" :: prf_decl % "proof" - and "fix" "assume" "presume" "def" :: prf_asm % "proof" + and "fix" "assume" "presume" "define" "def" :: prf_asm % "proof" and "consider" :: prf_goal % "proof" and "obtain" :: prf_asm_goal % "proof" and "guess" :: prf_script_asm_goal % "proof" @@ -739,12 +739,18 @@ (structured_statement >> (fn (a, b, c) => Toplevel.proof (Proof.presume_cmd a b c))); val _ = + Outer_Syntax.command @{command_keyword define} "local definition (non-polymorphic)" + ((Parse.fixes --| Parse.where_) -- Parse_Spec.statement -- Parse.for_fixes + >> (fn ((a, b), c) => Toplevel.proof (Proof.define_cmd a c b))); + +val _ = Outer_Syntax.command @{command_keyword def} "local definition (non-polymorphic)" (Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- ((Parse.binding -- Parse.opt_mixfix) -- ((@{keyword "\"} || @{keyword "=="}) |-- Parse.!!! Parse.termp))) - >> (Toplevel.proof o Proof.def_cmd)); + >> (fn args => Toplevel.proof (fn state => + (legacy_feature "Old 'def' command -- use 'define' instead"; Proof.def_cmd args state)))); val _ = Outer_Syntax.command @{command_keyword consider} "state cases rule" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Pure/more_thm.ML --- a/src/Pure/more_thm.ML Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Pure/more_thm.ML Mon Apr 25 22:59:53 2016 +0200 @@ -566,13 +566,9 @@ fun def_name_optional c "" = def_name c | def_name_optional _ name = name; -val def_binding = Binding.map_name def_name; - -fun def_binding_optional b name = - if Binding.is_empty name then def_binding b else name; - -fun make_def_binding cond b = - if cond then Binding.reset_pos (def_binding b) else Binding.empty; +val def_binding = Binding.map_name def_name #> Binding.reset_pos; +fun def_binding_optional b name = if Binding.is_empty name then def_binding b else name; +fun make_def_binding cond b = if cond then def_binding b else Binding.empty; (* unofficial theorem names *) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Pure/primitive_defs.ML --- a/src/Pure/primitive_defs.ML Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Pure/primitive_defs.ML Mon Apr 25 22:59:53 2016 +0200 @@ -6,7 +6,11 @@ signature PRIMITIVE_DEFS = sig - val dest_def: Proof.context -> (term -> bool) -> (string -> bool) -> (string -> bool) -> + val dest_def: Proof.context -> + {check_head: term -> bool, + check_free_lhs: string -> bool, + check_free_rhs: string -> bool, + check_tfree: string -> bool} -> term -> (term * term) * term val abs_def: term -> term * term end; @@ -20,7 +24,7 @@ | term_kind _ = ""; (*c x == t[x] to !!x. c x == t[x]*) -fun dest_def ctxt check_head is_fixed is_fixedT eq = +fun dest_def ctxt {check_head, check_free_lhs, check_free_rhs, check_tfree} eq = let fun err msg = raise TERM (msg, [eq]); val eq_vars = Term.strip_all_vars eq; @@ -36,7 +40,7 @@ val head_tfrees = Term.add_tfrees head []; fun check_arg (Bound _) = true - | check_arg (Free (x, _)) = not (is_fixed x) + | check_arg (Free (x, _)) = check_free_lhs x | check_arg (Const ("Pure.type", Type ("itself", [TFree _]))) = true | check_arg _ = false; fun close_arg (Bound _) t = t @@ -45,10 +49,10 @@ val lhs_bads = filter_out check_arg args; val lhs_dups = duplicates (op aconv) args; val rhs_extras = Term.fold_aterms (fn v as Free (x, _) => - if is_fixed x orelse member (op aconv) args v then I + if check_free_rhs x orelse member (op aconv) args v then I else insert (op aconv) v | _ => I) rhs []; val rhs_extrasT = Term.fold_aterms (Term.fold_types (fn v as TFree (a, S) => - if is_fixedT a orelse member (op =) head_tfrees (a, S) then I + if check_tfree a orelse member (op =) head_tfrees (a, S) then I else insert (op =) v | _ => I)) rhs []; in if not (check_head head) then diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Pure/theory.ML --- a/src/Pure/theory.ML Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Pure/theory.ML Mon Apr 25 22:59:53 2016 +0200 @@ -291,7 +291,12 @@ fun check_def (context as (ctxt, _)) thy unchecked overloaded (b, tm) defs = let val name = Sign.full_name thy b; - val ((lhs, rhs), _) = Primitive_Defs.dest_def ctxt Term.is_Const (K false) (K false) tm + val ((lhs, rhs), _) = + Primitive_Defs.dest_def ctxt + {check_head = Term.is_Const, + check_free_lhs = K true, + check_free_rhs = K false, + check_tfree = K false} tm handle TERM (msg, _) => error msg; val lhs_const = Term.dest_Const (Term.head_of lhs); diff -r 1836456b7d82 -r 2cc4e85b46d4 src/Tools/jEdit/src/rich_text_area.scala --- a/src/Tools/jEdit/src/rich_text_area.scala Mon Apr 25 21:09:02 2016 +0200 +++ b/src/Tools/jEdit/src/rich_text_area.scala Mon Apr 25 22:59:53 2016 +0200 @@ -86,7 +86,7 @@ painter_clip = gfx.getClip caret_focus = JEdit_Lib.visible_range(text_area) match { - case Some(visible_range) if caret_enabled => + case Some(visible_range) if caret_enabled && !painter_rendering.snapshot.is_outdated => painter_rendering.caret_focus(JEdit_Lib.caret_range(text_area), visible_range) case _ => Set.empty[Long] } diff -r 1836456b7d82 -r 2cc4e85b46d4 src/ZF/CardinalArith.thy --- a/src/ZF/CardinalArith.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/ZF/CardinalArith.thy Mon Apr 25 22:59:53 2016 +0200 @@ -28,7 +28,7 @@ definition jump_cardinal :: "i=>i" where - \\This def is more complex than Kunen's but it more easily proved to + \\This definition is more complex than Kunen's but it more easily proved to be a cardinal\ "jump_cardinal(K) == \X\Pow(K). {z. r \ Pow(K*K), well_ord(X,r) & z = ordertype(X,r)}" diff -r 1836456b7d82 -r 2cc4e85b46d4 src/ZF/Induct/Multiset.thy --- a/src/ZF/Induct/Multiset.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/ZF/Induct/Multiset.thy Mon Apr 25 22:59:53 2016 +0200 @@ -895,7 +895,7 @@ apply (rule multirel1_mono, auto) done -(* Equivalence of multirel with the usual (closure-free) def *) +(* Equivalence of multirel with the usual (closure-free) definition *) lemma add_diff_eq: "k \ nat ==> 0 < k \ n #+ k #- 1 = n #+ (k #- 1)" by (erule nat_induct, auto) diff -r 1836456b7d82 -r 2cc4e85b46d4 src/ZF/QUniv.thy --- a/src/ZF/QUniv.thy Mon Apr 25 21:09:02 2016 +0200 +++ b/src/ZF/QUniv.thy Mon Apr 25 22:59:53 2016 +0200 @@ -58,7 +58,7 @@ apply (rule Transset_eclose [THEN Transset_univ]) done -(*Key property for proving A_subset_quniv; requires eclose in def of quniv*) +(*Key property for proving A_subset_quniv; requires eclose in definition of quniv*) lemma univ_subset_quniv: "univ(A) \ quniv(A)" apply (rule arg_subset_eclose [THEN univ_mono, THEN subset_trans]) apply (rule univ_eclose_subset_quniv)