--- a/src/HOL/Algebra/Ideal.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Algebra/Ideal.thy Mon Apr 25 16:09:26 2016 +0200
@@ -787,7 +787,7 @@
abI: "a \<otimes> b \<in> I" and
anI: "a \<notin> I" and
bnI: "b \<notin> I" by fast
- def J \<equiv> "{x\<in>carrier R. a \<otimes> x \<in> I}"
+ define J where "J = {x\<in>carrier R. a \<otimes> x \<in> I}"
from is_cring and acarr have idealJ: "ideal J R"
unfolding J_def by (rule helper_max_prime)
--- a/src/HOL/Algebra/QuotRing.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Algebra/QuotRing.thy Mon Apr 25 16:09:26 2016 +0200
@@ -238,7 +238,7 @@
assume IanI: "I +> a \<noteq> I" and acarr: "a \<in> carrier R"
--\<open>Helper ideal @{text "J"}\<close>
- def J \<equiv> "(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)
--- a/src/HOL/BNF_Cardinal_Order_Relation.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/BNF_Cardinal_Order_Relation.thy Mon Apr 25 16:09:26 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 \<equiv> "\<lambda>(k::('a + 'b) + 'c).
- case k of Inl ab \<Rightarrow> (case ab of Inl a \<Rightarrow> Inl a
- |Inr b \<Rightarrow> Inr (Inl b))
- |Inr c \<Rightarrow> Inr (Inr c)"
+ define f :: "('a + 'b) + 'c \<Rightarrow> 'a + 'b + 'c"
+ where [abs_def]: "f k =
+ (case k of
+ Inl ab \<Rightarrow>
+ (case ab of
+ Inl a \<Rightarrow> Inl a
+ | Inr b \<Rightarrow> Inr (Inl b))
+ | Inr c \<Rightarrow> Inr (Inr c))"
+ for k
have "A <+> B <+> C \<subseteq> f ` ((A <+> B) <+> C)"
proof
fix x assume x: "x \<in> 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 \<equiv> "\<lambda> A' a. if a \<in> A then (if a \<in> A' then True else False)
- else undefined"
+ define F where [abs_def]: "F A' a =
+ (if a \<in> A then (if a \<in> 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 \<in> Pow A" "A2 \<in> Pow A" "F A1 = F A2"
@@ -1638,8 +1643,9 @@
fix f assume f: "f \<in> Func A (UNIV::bool set)"
show "f \<in> F ` Pow A" unfolding image_def mem_Collect_eq proof(intro bexI)
let ?A1 = "{a \<in> 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
--- a/src/HOL/BNF_Wellorder_Constructions.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/BNF_Wellorder_Constructions.thy Mon Apr 25 16:09:26 2016 +0200
@@ -1616,7 +1616,7 @@
show ?thesis
proof safe
fix h assume h: "h \<in> Func B2 B1"
- def j1 \<equiv> "inv_into A1 f1"
+ define j1 where "j1 = inv_into A1 f1"
have "\<forall> a2 \<in> f2 ` B2. \<exists> b2. b2 \<in> B2 \<and> f2 b2 = a2" by blast
then obtain k where k: "\<forall> a2 \<in> f2 ` B2. k a2 \<in> B2 \<and> 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 \<in> B2" using B2 by auto
- def j2 \<equiv> "\<lambda> a2. if a2 \<in> f2 ` B2 then k a2 else b22"
+ define j2 where [abs_def]: "j2 a2 = (if a2 \<in> f2 ` B2 then k a2 else b22)" for a2
have j2A2: "j2 ` A2 \<subseteq> B2" unfolding j2_def using k b22 by auto
have j2: "\<And> b2. b2 \<in> B2 \<Longrightarrow> j2 (f2 b2) = b2"
using kk unfolding j2_def by auto
- def g \<equiv> "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"
--- a/src/HOL/Binomial.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Binomial.thy Mon Apr 25 16:09:26 2016 +0200
@@ -643,7 +643,7 @@
shows "pochhammer z (Suc n) * pochhammer (z + 1/2) (Suc n) = (\<Prod>k=0..2*n+1. z + of_nat k / 2)"
proof (induction n)
case (Suc n)
- def n' \<equiv> "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 @@
(\<Sum>k\<le>m. (-r gchoose k) * (-x)^k * (x + y)^(m-k))" (is "?lhs m = ?rhs m")
proof (induction m)
case (Suc mm)
- def G \<equiv> "\<lambda>i k. (of_nat i + r gchoose k) * x^k * y^(i-k)" and S \<equiv> ?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 = (\<lambda>i. (\<Sum>k\<le>i. (G i k)))" unfolding S_def G_def ..
have "S (Suc mm) = G (Suc mm) 0 + (\<Sum>k=Suc 0..Suc mm. G (Suc mm) k)"
@@ -1388,7 +1389,7 @@
proof(rule setsum.cong[OF refl])
fix x
assume x: "x \<in> \<Union>A"
- def K \<equiv> "{X \<in> A. x \<in> X}"
+ define K where "K = {X \<in> A. x \<in> X}"
with \<open>finite A\<close> have K: "finite K" by auto
let ?I = "\<lambda>i. {I. I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I}"
have "inj_on snd (SIGMA i:{1..card A}. ?I i)"
--- a/src/HOL/Cardinals/Bounded_Set.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Cardinals/Bounded_Set.thy Mon Apr 25 16:09:26 2016 +0200
@@ -50,8 +50,9 @@
BNF_Def.Grp {a. set_bset a \<subseteq> {(a, b). R a b}} (map_bset snd)) a b" (is "?L \<longleftrightarrow> ?R")
proof
assume ?L
- def R' \<equiv> "the_inv set_bset (Collect (case_prod R) \<inter> (set_bset a \<times> set_bset b)) :: ('a \<times> 'b) set['k]"
- (is "the_inv set_bset ?L'")
+ define R' :: "('a \<times> 'b) set['k]"
+ where "R' = the_inv set_bset (Collect (case_prod R) \<inter> (set_bset a \<times> set_bset b))"
+ (is "_ = the_inv set_bset ?L'")
have "|?L'| <o natLeq +c |UNIV :: 'k set|"
unfolding csum_def Field_natLeq
by (intro ordLeq_ordLess_trans[OF card_of_mono1[OF Int_lower2]]
--- a/src/HOL/Cardinals/Cardinal_Arithmetic.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Cardinals/Cardinal_Arithmetic.thy Mon Apr 25 16:09:26 2016 +0200
@@ -47,10 +47,10 @@
by (simp only: cprod_def Field_card_of card_of_refl)
lemma card_of_Times_singleton:
-fixes A :: "'a set"
-shows "|A \<times> {x}| =o |A|"
+ fixes A :: "'a set"
+ shows "|A \<times> {x}| =o |A|"
proof -
- def f \<equiv> "\<lambda>(a::'a,b::'b). (a)"
+ define f :: "'a \<times> 'b \<Rightarrow> 'a" where "f = (\<lambda>(a, b). a)"
have "A \<subseteq> f ` (A \<times> {x})" unfolding f_def by (auto simp: image_iff)
hence "bij_betw f (A \<times> {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 \<equiv> "\<lambda>y a. if y = x \<and> a \<in> A then x else undefined"
+ define f where [abs_def]: "f y a = (if y = x \<and> a \<in> A then x else undefined)" for y a
have "Func A {x} \<subseteq> 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 \<times> A|"
proof (rule ordIso_symmetric)
- def f \<equiv> "\<lambda>(x::'a,y) b. if A = {} then undefined else if b then x else y"
+ define f where "f = (\<lambda>(x::'a,y) b. if A = {} then undefined else if b then x else y)"
have "Func (UNIV :: bool set) A \<subseteq> f ` (A \<times> 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 \<times> 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 \<times> Func B C|"
proof (rule ordIso_symmetric)
- def f \<equiv> "\<lambda>(g :: 'a => 'c, h::'b \<Rightarrow> 'c) ab. case ab of Inl a \<Rightarrow> g a | Inr b \<Rightarrow> h b"
- def f' \<equiv> "\<lambda>(f :: ('a + 'b) \<Rightarrow> 'c). (\<lambda>a. f (Inl a), \<lambda>b. f (Inr b))"
+ define f where "f = (\<lambda>(g :: 'a => 'c, h::'b \<Rightarrow> 'c) ab. case ab of Inl a \<Rightarrow> g a | Inr b \<Rightarrow> h b)"
+ define f' where "f' = (\<lambda>(f :: ('a + 'b) \<Rightarrow> 'c). (\<lambda>a. f (Inl a), \<lambda>b. f (Inr b)))"
have "f ` (Func A C \<times> Func B C) \<subseteq> Func (A <+> B) C"
unfolding Func_def f_def by (force split: sum.splits)
moreover have "f' ` Func (A <+> B) C \<subseteq> Func A C \<times> Func B C" unfolding Func_def f'_def by force
--- a/src/HOL/Cardinals/Cardinal_Order_Relation.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Cardinals/Cardinal_Order_Relation.thy Mon Apr 25 16:09:26 2016 +0200
@@ -1431,7 +1431,7 @@
show "?F ` Func A B = Func_option A B"
proof safe
fix f assume f: "f \<in> Func_option A B"
- def g \<equiv> "\<lambda> a. case f a of Some b \<Rightarrow> b | None \<Rightarrow> undefined"
+ define g where [abs_def]: "g a = (case f a of Some b \<Rightarrow> b | None \<Rightarrow> undefined)" for a
have "g \<in> 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| \<le>o |Func A2 B|"
proof-
obtain bb where bb: "bb \<in> B" using B by auto
- def F \<equiv> "\<lambda> (f1::'a \<Rightarrow> 'b) a. if a \<in> A2 then (if a \<in> A1 then f1 a else bb)
- else undefined"
+ define F where [abs_def]: "F f1 a =
+ (if a \<in> A2 then (if a \<in> A1 then f1 a else bb) else undefined)" for f1 :: "'a \<Rightarrow> '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 \<in> Func A1 B" and g: "g \<in> Func A1 B" and eq: "F f = F g"
@@ -1523,7 +1523,7 @@
hence "|X| \<le>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 \<equiv> "\<lambda> i. if i \<in> Field r then F i else undefined"
+ define f where [abs_def]: "f i = (if i \<in> Field r then F i else undefined)" for i
have "\<exists> f \<in> 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 \<in> Field r" and jm: "\<forall>i\<in>Field r. (i, j) \<in> r"
- def p \<equiv> "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 \<in> A"
- def L \<equiv> "{(a,u) | u. u \<in> F a}"
+ define L where "L = {(a,u) | u. u \<in> F a}"
have fL: "f ` L \<subseteq> 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: "\<forall> a \<in> A. \<forall> u \<in> F a. (f (a,u), gg a) \<in> r" by metis
obtain j0 where j0: "j0 \<in> Field r" using Fi by auto
- def g \<equiv> "\<lambda> a. if F a \<noteq> {} then gg a else j0"
+ define g where [abs_def]: "g a = (if F a \<noteq> {} then gg a else j0)" for a
have g: "\<forall> a \<in> A. \<forall> u \<in> F a. (f (a,u),g a) \<in> r" using gg unfolding g_def by auto
hence 1: "Field r \<subseteq> (\<Union>a \<in> A. under r (g a))"
using f[symmetric] unfolding under_def image_def by auto
--- a/src/HOL/Cardinals/Ordinal_Arithmetic.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Cardinals/Ordinal_Arithmetic.thy Mon Apr 25 16:09:26 2016 +0200
@@ -751,12 +751,12 @@
with F have maxF: "\<forall>f \<in> F. s.isMaxim (SUPP f) (s.maxim (SUPP f))"
and SUPPF: "\<forall>f \<in> F. finite (SUPP f) \<and> SUPP f \<noteq> {} \<and> SUPP f \<subseteq> Field s"
using maxim_isMaxim_support support_not_const by auto
- def z \<equiv> "s.minim {s.maxim (SUPP f) | f. f \<in> F}"
+ define z where "z = s.minim {s.maxim (SUPP f) | f. f \<in> F}"
from F SUPPF maxF have zmin: "s.isMinim {s.maxim (SUPP f) | f. f \<in> F} z"
unfolding z_def by (intro s.minim_isMinim) (auto simp: s.isMaxim_def)
with F have zy: "(z, y) \<in> s" unfolding s.isMinim_def by auto
hence zField: "z \<in> Field s" unfolding Field_def by auto
- def x0 \<equiv> "r.minim {f z | f. f \<in> F \<and> z = s.maxim (SUPP f)}"
+ define x0 where "x0 = r.minim {f z | f. f \<in> F \<and> z = s.maxim (SUPP f)}"
from F(1,2) maxF(1) SUPPF zmin
have x0min: "r.isMinim {f z | f. f \<in> F \<and> 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 \<noteq> r.zero"
unfolding r.isMinim_def s.isMaxim_def FinFunc_def Func_def support_def
by fastforce
- def G \<equiv> "{f(z := r.zero) | f. f \<in> F \<and> z = s.maxim (SUPP f) \<and> f z = x0}"
+ define G where "G = {f(z := r.zero) | f. f \<in> F \<and> z = s.maxim (SUPP f) \<and> f z = x0}"
from zmin x0min have "G \<noteq> {}" unfolding G_def z_def s.isMinim_def r.isMinim_def by blast
have GF: "G \<subseteq> (\<lambda>f. f(z := r.zero)) ` F" unfolding G_def by auto
have "G \<subseteq> fin_support r.zero (Field s)"
@@ -786,7 +786,7 @@
with G have maxG: "\<forall>g \<in> G. s.isMaxim (SUPP g) (s.maxim (SUPP g))"
and SUPPG: "\<forall>g \<in> G. finite (SUPP g) \<and> SUPP g \<noteq> {} \<and> SUPP g \<subseteq> Field s"
using maxim_isMaxim_support support_not_const by auto
- def y' \<equiv> "s.minim {s.maxim (SUPP f) | f. f \<in> G}"
+ define y' where "y' = s.minim {s.maxim (SUPP f) | f. f \<in> G}"
from G SUPPG maxG `G \<noteq> {}` have y'min: "s.isMinim {s.maxim (SUPP f) | f. f \<in> 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 \<in> G" "\<forall>g \<in> G. (g0, g) \<in> oexp" by blast
hence g0z: "g0 z = r.zero" unfolding G_def by auto
- def f0 \<equiv> "g0(z := x0)"
+ define f0 where "f0 = g0(z := x0)"
with x0notzero zField have SUPP: "SUPP f0 = SUPP g0 \<union> {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 \<in> 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 \<equiv> "\<lambda>g z. if z \<in> f ` Field s then g (the_inv_into (Field s) f z) else
- if z \<in> Field t then r.zero else undefined"
+ define F where [abs_def]: "F g z =
+ (if z \<in> f ` Field s then g (the_inv_into (Field s) f z)
+ else if z \<in> 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 \<subseteq> 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 \<subset> Field ?R"
proof (intro psubsetI)
from *(4) obtain z where z: "z \<in> Field t" "z \<notin> f ` Field s" by auto
- def h \<equiv> "\<lambda>z'. if z' \<in> Field t then if z' = z then x else r.zero else undefined"
+ define h where [abs_def]: "h z' =
+ (if z' \<in> 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 \<in> Field ?R" unfolding h_def rt.Field_oexp FinFunc_def Func_def fin_support_def
by auto
--- a/src/HOL/Cardinals/Wellorder_Constructions.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Cardinals/Wellorder_Constructions.thy Mon Apr 25 16:09:26 2016 +0200
@@ -658,7 +658,7 @@
lemma isSucc_pred_in:
assumes "isSucc i" shows "(pred i, i) \<in> r"
proof-
- def j \<equiv> "pred i"
+ define j where "j = pred i"
have i: "i = succ j" using assms unfolding j_def by simp
have a: "aboveS j \<noteq> {}" 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 \<noteq> {}"
shows "worecSL S L (succ j) = S j (worecSL S L j)"
proof-
- def i \<equiv> "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 = "\<lambda> g a. suc s (g ` underS r a)"
- def g \<equiv> "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 \<in> Field r"
--- a/src/HOL/Cardinals/Wellorder_Extension.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Cardinals/Wellorder_Extension.thy Mon Apr 25 16:09:26 2016 +0200
@@ -66,7 +66,7 @@
shows "\<exists>w. p \<subseteq> w \<and> Well_order w"
proof -
let ?K = "{r. Well_order r \<and> downset_on (Field r) p \<and> extension_on (Field r) r p}"
- def I \<equiv> "init_seg_of \<inter> ?K \<times> ?K"
+ define I where "I = init_seg_of \<inter> ?K \<times> ?K"
have I_init: "I \<subseteq> init_seg_of" by (simp add: I_def)
then have subch: "\<And>R. R \<in> Chains I \<Longrightarrow> chain\<^sub>\<subseteq> R"
by (auto simp: init_seg_of_def chain_subset_def Chains_def)
--- a/src/HOL/Complex.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Complex.thy Mon Apr 25 16:09:26 2016 +0200
@@ -817,7 +817,8 @@
from assms have "z \<noteq> 0" by auto
have "(SOME a. sgn z = cis a \<and> -pi < a \<and> a \<le> pi) = x"
proof
- fix a def d \<equiv> "a - x"
+ fix a
+ define d where "d = a - x"
assume a: "sgn z = cis a \<and> - pi < a \<and> a \<le> pi"
from a assms have "- (2*pi) < d \<and> 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 \<noteq> 0" by auto
- def b \<equiv> "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 \<open>r \<noteq> 0\<close>
by (simp add: of_real_def sgn_scaleR sgn_if complex_eq_iff)
@@ -848,7 +849,7 @@
have cis_2pi_int: "\<And>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 \<equiv> "b - 2*pi * of_int \<lceil>(b - pi) / (2*pi)\<rceil>"
+ define c where "c = b - 2 * pi * of_int \<lceil>(b - pi) / (2 * pi)\<rceil>"
have "sgn z = cis c"
unfolding b c_def
by (simp add: cis_divide [symmetric] cis_2pi_int)
--- a/src/HOL/Conditionally_Complete_Lattices.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Conditionally_Complete_Lattices.thy Mon Apr 25 16:09:26 2016 +0200
@@ -590,7 +590,7 @@
S = {a..<b} \<or>
S = {a..b}"
proof -
- def lower \<equiv> "{x. \<exists>s\<in>S. s \<le> x}" and upper \<equiv> "{x. \<exists>s\<in>S. x \<le> s}"
+ define lower upper where "lower = {x. \<exists>s\<in>S. s \<le> x}" and "upper = {x. \<exists>s\<in>S. x \<le> s}"
with ivl have "S = lower \<inter> upper"
by auto
moreover
--- a/src/HOL/Datatype_Examples/Derivation_Trees/Gram_Lang.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Derivation_Trees/Gram_Lang.thy Mon Apr 25 16:09:26 2016 +0200
@@ -665,7 +665,7 @@
assume ?C then obtain tr where
tr_tr0: "Inr tr \<in> cont tr0" and t_tr: "inFr (ns - {root tr0}) tr t"
unfolding inFrr_def by auto
- def tr1 \<equiv> "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 \<in> 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 \<longleftrightarrow> (\<exists> f. reg f tr \<and> (\<forall> n. root (f n) = n))"
unfolding regular_def proof safe
- fix f assume f: "reg f tr"
- def g \<equiv> "\<lambda> 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 "\<exists>g. reg g tr \<and> (\<forall>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 \<in> 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 \<equiv> "(id \<oplus> root) ` (cont tr)"
- def K \<equiv> "\<lambda> n'. Fr (ns - {n}) (subtrOf tr n')"
+ define tns where "tns = (id \<oplus> root) ` (cont tr)"
+ define K where "K n' = Fr (ns - {n}) (subtrOf tr n')" for n'
show "\<exists>tns K. ts = ?F tns K \<and> (n, tns) \<in> P \<and> ?\<phi> tns K"
apply(rule exI[of _ tns], rule exI[of _ K]) proof(intro conjI allI impI)
show "ts = Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns}"
@@ -1222,7 +1223,8 @@
then obtain ftr where 0: "\<And> n'. Inr n' \<in> tns \<Longrightarrow>
K n' = Fr (ns - {n}) (ftr n') \<and> wf (ftr n') \<and> root (ftr n') = n'"
by metis
- def tr \<equiv> "Node n ((id \<oplus> ftr) ` tns)" def tr' \<equiv> "hsubst tr tr"
+ define tr where "tr = Node n ((id \<oplus> ftr) ` tns)"
+ define tr' where "tr' = hsubst tr tr"
have rtr: "root tr = n" and ctr: "cont tr = (id \<oplus> 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)
--- a/src/HOL/Decision_Procs/Approximation.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy Mon Apr 25 16:09:26 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 \<equiv> "bitlen m - 1"
+ define bl where "bl = bitlen m - 1"
have "0 < real_of_int m" and "\<And>X. (0 :: real) < 2^X" and "0 < (2 :: real)" and "m \<noteq> 0"
using assms by auto
hence "0 \<le> bl" by (simp add: bitlen_def bl_def)
@@ -2528,8 +2528,8 @@
using \<open>1 \<le> x\<close> by auto
show ?thesis
proof -
- def m \<equiv> "mantissa x"
- def e \<equiv> "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 \<equiv> "bitlen m - 1"
+ define bl where "bl = bitlen m - 1"
hence "bl \<ge> 0"
using \<open>m > 0\<close> by (simp add: bitlen_def)
have "1 \<le> Float m e"
@@ -2746,7 +2746,7 @@
with x lu show ?thesis by (auto simp: bnds_powr_def)
next
assume A: "l1 = 0" "u1 \<noteq> 0" "l2 \<ge> 1"
- def uln \<equiv> "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: "\<And>k F. setsum F {0..<Suc k} = F 0 + setsum (\<lambda> k. F (Suc k)) {0..<k}"
unfolding setsum_shift_bounds_Suc_ivl[symmetric]
unfolding setsum_head_upt_Suc[OF zero_less_Suc] ..
- def C \<equiv> "xs!x - c"
+ define C where "C = xs!x - c"
{
fix t::real assume t: "t \<in> {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 \<in> {l .. u}"
proof -
- def F \<equiv> "\<lambda>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 = (\<lambda> z. interpret_floatarith f (xs[x := z]))" by auto
hence "bounded_by (xs[x := c]) vs" and "x < length vs" "x < length xs"
--- a/src/HOL/Deriv.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Deriv.thy Mon Apr 25 16:09:26 2016 +0200
@@ -235,8 +235,8 @@
let ?D = "\<lambda>f f' x y. (f y - f x) - f' (y - x)"
let ?N = "\<lambda>f f' x y. norm (?D f f' x y) / norm (y - x)"
let ?gf = "\<lambda>x. g (f x)" and ?gf' = "\<lambda>x. g' (f' x)"
- def Nf \<equiv> "?N f f' x"
- def Ng \<equiv> "\<lambda>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 @@
"\<And>a b. norm (a ** b) \<le> norm a * norm b * K" by fast
let ?D = "\<lambda>f f' y. f y - f x - f' (y - x)"
let ?N = "\<lambda>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 = "\<lambda>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 = "\<lambda>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: "((\<lambda> x. (f' x / g' x)) \<longlongrightarrow> x) (at_right 0)"
shows "((\<lambda> x. f0 x / g0 x) \<longlongrightarrow> x) (at_right 0)"
proof -
- def f \<equiv> "\<lambda>x. if x \<le> 0 then 0 else f0 x"
+ define f where [abs_def]: "f x = (if x \<le> 0 then 0 else f0 x)" for x
then have "f 0 = 0" by simp
- def g \<equiv> "\<lambda>x. if x \<le> 0 then 0 else g0 x"
+ define g where [abs_def]: "g x = (if x \<le> 0 then 0 else g0 x)" for x
then have "g 0 = 0" by simp
have "eventually (\<lambda>x. g0 x \<noteq> 0 \<and> g' x \<noteq> 0 \<and>
--- a/src/HOL/Divides.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Divides.thy Mon Apr 25 16:09:26 2016 +0200
@@ -623,7 +623,8 @@
with \<open>0 < b\<close> have "0 < a div b" by (auto intro: div_positive)
then have [simp]: "1 \<le> a div b" by (simp add: discrete)
with \<open>0 < b\<close> have mod_less: "a mod b < b" by (simp add: pos_mod_bound)
- def w \<equiv> "a div b mod 2" with parity have w_exhaust: "w = 0 \<or> w = 1" by auto
+ define w where "w = a div b mod 2"
+ with parity have w_exhaust: "w = 0 \<or> 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 \<equiv> "a div b mod 2" with parity have w_exhaust: "w = 0 \<or> w = 1" by auto
+ define w where "w = a div b mod 2"
+ with parity have w_exhaust: "w = 0 \<or> 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 \<le> a mod b + b"
--- a/src/HOL/Finite_Set.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Finite_Set.thy Mon Apr 25 16:09:26 2016 +0200
@@ -919,7 +919,7 @@
case 0 then show ?case by simp
next
case (Suc n g)
- def h \<equiv> "\<lambda>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 \<circ> f x = f x \<circ> 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 \<equiv> "\<lambda>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 \<circ> f x ^^ h x = f x ^^ h x \<circ> f y ^^ h y"
by auto
--- a/src/HOL/Groups_Big.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Groups_Big.thy Mon Apr 25 16:09:26 2016 +0200
@@ -1137,7 +1137,7 @@
case True with insert show ?thesis by simp
next
case False with insert have "a \<in> B" by simp
- def C \<equiv> "B - {a}"
+ define C where "C = B - {a}"
with \<open>finite B\<close> \<open>a \<in> B\<close>
have *: "B = insert a C" "finite C" "a \<notin> C" by auto
with insert show ?thesis by (auto simp add: insert_commute ac_simps)
--- a/src/HOL/HOLCF/Adm.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/HOLCF/Adm.thy Mon Apr 25 16:09:26 2016 +0200
@@ -61,7 +61,7 @@
assumes P: "\<forall>i. \<exists>j\<ge>i. P (Y j)"
shows "P (\<Squnion>i. Y i)"
proof -
- def f \<equiv> "\<lambda>i. LEAST j. i \<le> j \<and> P (Y j)"
+ define f where "f i = (LEAST j. i \<le> j \<and> P (Y j))" for i
have chain': "chain (\<lambda>i. Y (f i))"
unfolding f_def
apply (rule chainI)
--- a/src/HOL/HOLCF/Completion.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/HOLCF/Completion.thy Mon Apr 25 16:09:26 2016 +0200
@@ -191,14 +191,14 @@
proof -
obtain count :: "'a \<Rightarrow> nat" where inj: "inj count"
using countable ..
- def enum \<equiv> "\<lambda>i. THE a. count a = i"
+ define enum where "enum i = (THE a. count a = i)" for i
have enum_count [simp]: "\<And>x. enum (count x) = x"
unfolding enum_def by (simp add: inj_eq [OF inj])
- def a \<equiv> "LEAST i. enum i \<in> rep x"
- def b \<equiv> "\<lambda>i. LEAST j. enum j \<in> rep x \<and> \<not> enum j \<preceq> enum i"
- def c \<equiv> "\<lambda>i j. LEAST k. enum k \<in> rep x \<and> enum i \<preceq> enum k \<and> enum j \<preceq> enum k"
- def P \<equiv> "\<lambda>i. \<exists>j. enum j \<in> rep x \<and> \<not> enum j \<preceq> enum i"
- def X \<equiv> "rec_nat a (\<lambda>n i. if P i then c i (b i) else i)"
+ define a where "a = (LEAST i. enum i \<in> rep x)"
+ define b where "b i = (LEAST j. enum j \<in> rep x \<and> \<not> enum j \<preceq> enum i)" for i
+ define c where "c i j = (LEAST k. enum k \<in> rep x \<and> enum i \<preceq> enum k \<and> enum j \<preceq> enum k)" for i j
+ define P where "P i \<longleftrightarrow> (\<exists>j. enum j \<in> rep x \<and> \<not> enum j \<preceq> enum i)" for i
+ define X where "X = rec_nat a (\<lambda>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: "\<And>n. X (Suc n) = (if P (X n) then c (X n) (b (X n)) else X n)"
unfolding X_def by simp
--- a/src/HOL/HOLCF/Representable.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/HOLCF/Representable.thy Mon Apr 25 16:09:26 2016 +0200
@@ -97,7 +97,7 @@
obtain Y where Y: "\<forall>i. Y i \<sqsubseteq> Y (Suc i)"
and t: "t = (\<Squnion>i. defl_principal (Y i))"
by (rule defl.obtain_principal_chain)
- def approx \<equiv> "\<lambda>i. (p oo cast\<cdot>(defl_principal (Y i)) oo e) :: 'a \<rightarrow> 'a"
+ define approx where "approx i = (p oo cast\<cdot>(defl_principal (Y i)) oo e)" for i
have "approx_chain approx"
proof (rule approx_chain.intro)
show "chain (\<lambda>i. approx i)"
--- a/src/HOL/HOLCF/Up.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/HOLCF/Up.thy Mon Apr 25 16:09:26 2016 +0200
@@ -73,7 +73,7 @@
proof (cases "\<exists>k. Y k \<noteq> Ibottom")
case True
then obtain k where k: "Y k \<noteq> Ibottom" ..
- def A \<equiv> "\<lambda>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: "\<forall>i. Iup (A i) = Y (i + k)"
proof
fix i :: nat
--- a/src/HOL/Hahn_Banach/Function_Norm.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Hahn_Banach/Function_Norm.thy Mon Apr 25 16:09:26 2016 +0200
@@ -108,7 +108,7 @@
that \<open>y \<le> b\<close> for all \<open>y \<in> B\<close>. Due to the definition of \<open>B\<close> there are
two cases.\<close>
- def b \<equiv> "max c 0"
+ define b where "b = max c 0"
have "\<forall>y \<in> B V f. y \<le> b"
proof
fix y assume y: "y \<in> B V f"
--- a/src/HOL/Hahn_Banach/Hahn_Banach.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Hahn_Banach/Hahn_Banach.thy Mon Apr 25 16:09:26 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 \<equiv> "norm_pres_extensions E p F f"
+ define M where "M = norm_pres_extensions E p F f"
then have M: "M = \<dots>" by (simp only:)
from E have F: "vectorspace F" ..
note FE = \<open>F \<unlhd> E\<close>
@@ -140,7 +140,7 @@
qed
qed
- def H' \<equiv> "H + lin x'"
+ define H' where "H' = H + lin x'"
\<comment> \<open>Define \<open>H'\<close> as the direct sum of \<open>H\<close> and the linear closure of \<open>x'\<close>. \<^smallskip>\<close>
have HH': "H \<unlhd> H'"
proof (unfold H'_def)
@@ -178,8 +178,8 @@
then show thesis by (blast intro: that)
qed
- def h' \<equiv> "\<lambda>x. let (y, a) =
- SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H in h y + a * xi"
+ define h' where "h' x = (let (y, a) =
+ SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H in h y + a * xi)" for x
\<comment> \<open>Define the extension \<open>h'\<close> of \<open>h\<close> to \<open>H'\<close> using \<open>\<xi>\<close>. \<^smallskip>\<close>
have "g \<subseteq> graph H' h' \<and> g \<noteq> graph H' h'"
@@ -365,7 +365,7 @@
OF F_norm \<open>continuous F f norm\<close> , folded B_def fn_norm_def])
txt \<open>We define a function \<open>p\<close> on \<open>E\<close> as follows:
\<open>p x = \<parallel>f\<parallel> \<cdot> \<parallel>x\<parallel>\<close>\<close>
- def p \<equiv> "\<lambda>x. \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>"
+ define p where "p x = \<parallel>f\<parallel>\<hyphen>F * \<parallel>x\<parallel>" for x
txt \<open>\<open>p\<close> is a seminorm on \<open>E\<close>:\<close>
have q: "seminorm E p"
--- a/src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy Mon Apr 25 16:09:26 2016 +0200
@@ -86,9 +86,9 @@
\<close>
lemma h'_lf:
- assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
- SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
- and H'_def: "H' \<equiv> H + lin x0"
+ assumes h'_def: "\<And>x. h' x = (let (y, a) =
+ SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi)"
+ and H'_def: "H' = H + lin x0"
and HE: "H \<unlhd> E"
assumes "linearform H h"
assumes x0: "x0 \<notin> H" "x0 \<in> E" "x0 \<noteq> 0"
@@ -192,9 +192,9 @@
\<close>
lemma h'_norm_pres:
- assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
- SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
- and H'_def: "H' \<equiv> H + lin x0"
+ assumes h'_def: "\<And>x. h' x = (let (y, a) =
+ SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi)"
+ and H'_def: "H' = H + lin x0"
and x0: "x0 \<notin> H" "x0 \<in> E" "x0 \<noteq> 0"
assumes E: "vectorspace E" and HE: "subspace H E"
and "seminorm E p" and "linearform H h"
--- a/src/HOL/Hahn_Banach/Hahn_Banach_Sup_Lemmas.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Hahn_Banach/Hahn_Banach_Sup_Lemmas.thy Mon Apr 25 16:09:26 2016 +0200
@@ -158,7 +158,7 @@
\<close>
lemma sup_definite:
- assumes M_def: "M \<equiv> norm_pres_extensions E p F f"
+ assumes M_def: "M = norm_pres_extensions E p F f"
and cM: "c \<in> chains M"
and xy: "(x, y) \<in> \<Union>c"
and xz: "(x, z) \<in> \<Union>c"
--- a/src/HOL/Hahn_Banach/Subspace.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Hahn_Banach/Subspace.thy Mon Apr 25 16:09:26 2016 +0200
@@ -467,9 +467,9 @@
lemma h'_definite:
fixes H
assumes h'_def:
- "h' \<equiv> \<lambda>x.
- let (y, a) = SOME (y, a). (x = y + a \<cdot> x' \<and> y \<in> H)
- in (h y) + a * xi"
+ "\<And>x. h' x =
+ (let (y, a) = SOME (y, a). (x = y + a \<cdot> x' \<and> y \<in> H)
+ in (h y) + a * xi)"
and x: "x = y + a \<cdot> x'"
assumes "vectorspace E" "subspace H E"
assumes y: "y \<in> H"
--- a/src/HOL/Hilbert_Choice.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Hilbert_Choice.thy Mon Apr 25 16:09:26 2016 +0200
@@ -110,7 +110,8 @@
2: "\<And>x n. P n x \<Longrightarrow> \<exists>y. P (Suc n) y \<and> Q n x y"
shows "\<exists>f. \<forall>n. P n (f n) \<and> Q n (f n) (f (Suc n))"
proof (intro exI allI conjI)
- fix n def f \<equiv> "rec_nat (SOME x. P 0 x) (\<lambda>n x. SOME y. P (Suc n) y \<and> Q n x y)"
+ fix n
+ define f where "f = rec_nat (SOME x. P 0 x) (\<lambda>n x. SOME y. P (Suc n) y \<and> Q n x y)"
have "P 0 (f 0)" "\<And>n. P n (f n) \<Longrightarrow> P (Suc n) (f (Suc n)) \<and> 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 "\<exists>f. inj (f::nat \<Rightarrow> 'a) \<and> range f \<subseteq> S"
\<comment> \<open>Courtesy of Stephan Merz\<close>
proof -
- def Sseq \<equiv> "rec_nat S (\<lambda>n T. T - {SOME e. e \<in> T})"
- def pick \<equiv> "\<lambda>n. (SOME e. e \<in> Sseq n)"
+ define Sseq where "Sseq = rec_nat S (\<lambda>n T. T - {SOME e. e \<in> T})"
+ define pick where "pick n = (SOME e. e \<in> Sseq n)" for n
{ fix n have "Sseq n \<subseteq> S" "\<not> finite (Sseq n)" by (induct n) (auto simp add: Sseq_def inf) }
moreover then have *: "\<And>n. pick n \<in> 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 \<equiv> "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
--- a/src/HOL/Isar_Examples/Schroeder_Bernstein.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Isar_Examples/Schroeder_Bernstein.thy Mon Apr 25 16:09:26 2016 +0200
@@ -21,8 +21,8 @@
assumes "inj f" and "inj g"
shows "\<exists>h :: 'a \<Rightarrow> 'b. inj h \<and> surj h"
proof
- def A \<equiv> "lfp (\<lambda>X. - (g ` (- (f ` X))))"
- def g' \<equiv> "inv g"
+ define A where "A = lfp (\<lambda>X. - (g ` (- (f ` X))))"
+ define g' where "g' = inv g"
let ?h = "\<lambda>z. if z \<in> A then f z else g' z"
have "A = - (g ` (- (f ` A)))"
--- a/src/HOL/Library/Complete_Partial_Order2.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Complete_Partial_Order2.thy Mon Apr 25 16:09:26 2016 +0200
@@ -133,7 +133,7 @@
fix x
assume "x \<in> f y' ` Y"
then obtain y where "y \<in> Y" and x: "x = f y' y" by blast
- def y'' \<equiv> "if y \<sqsubseteq> y' then y' else y"
+ define y'' where "y'' = (if y \<sqsubseteq> y' then y' else y)"
from chain \<open>y \<in> Y\<close> \<open>y' \<in> Y\<close> have "y \<sqsubseteq> y' \<or> y' \<sqsubseteq> y" by(rule chainD)
hence "f y' y \<le> f y'' y''" using \<open>y \<in> Y\<close> \<open>y' \<in> Y\<close>
by(auto simp add: y''_def intro: mono2 monotoneD[OF mono1])
--- a/src/HOL/Library/ContNotDenum.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/ContNotDenum.thy Mon Apr 25 16:09:26 2016 +0200
@@ -46,8 +46,9 @@
"\<And>a b c. a < b \<Longrightarrow> c \<notin> {i a b c .. j a b c}"
by metis
- def ivl \<equiv> "rec_nat (f 0 + 1, f 0 + 2) (\<lambda>n x. (i (fst x) (snd x) (f n), j (fst x) (snd x) (f n)))"
- def I \<equiv> "\<lambda>n. {fst (ivl n) .. snd (ivl n)}"
+ define ivl where "ivl =
+ rec_nat (f 0 + 1, f 0 + 2) (\<lambda>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 "\<exists>f. bij_betw f {a<..<b} {c<..<d}"
proof -
- def f \<equiv> "\<lambda>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
--- a/src/HOL/Library/Convex.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Convex.thy Mon Apr 25 16:09:26 2016 +0200
@@ -664,7 +664,7 @@
shows "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)"
and "(f x - f y) / (x - y) \<le> (f t - f y) / (t - y)"
proof -
- def a \<equiv> "(t - y) / (x - y)"
+ define a where "a \<equiv> (t - y) / (x - y)"
with t have "0 \<le> a" "0 \<le> 1 - a"
by (auto simp: field_simps)
with f \<open>x \<in> I\<close> \<open>y \<in> I\<close> have cvx: "f (a * x + (1 - a) * y) \<le> 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 \<in> A" "y \<in> A" "x < y"
- def z \<equiv> "(1 - t) * x + t * y"
+ define z where "z = (1 - t) * x + t * y"
with \<open>connected A\<close> and xy have ivl: "{x..y} \<subseteq> A" using connected_contains_Icc by blast
from xy t have xz: "z > x" by (simp add: z_def algebra_simps)
--- a/src/HOL/Library/Countable.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Countable.thy Mon Apr 25 16:09:26 2016 +0200
@@ -142,7 +142,7 @@
assumes finite_item: "\<And>x. rep_set x \<Longrightarrow> finite_item x"
shows "OFCLASS('b, countable_class)"
proof
- def f \<equiv> "\<lambda>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)"
--- a/src/HOL/Library/Countable_Set_Type.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Countable_Set_Type.thy Mon Apr 25 16:09:26 2016 +0200
@@ -567,8 +567,8 @@
BNF_Def.Grp {x. rcset x \<subseteq> {(a, b). R a b}} (cimage snd)) a b" (is "?L = ?R")
proof
assume ?L
- def R' \<equiv> "the_inv rcset (Collect (case_prod R) \<inter> (rcset a \<times> rcset b))"
- (is "the_inv rcset ?L'")
+ define R' where "R' = the_inv rcset (Collect (case_prod R) \<inter> (rcset a \<times> 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
--- a/src/HOL/Library/DAList_Multiset.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/DAList_Multiset.thy Mon Apr 25 16:09:26 2016 +0200
@@ -313,7 +313,7 @@
have inv: "ys \<in> ?inv"
using Cons(2) by auto
note IH = Cons(1)[OF inv]
- def Ys \<equiv> "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)
--- a/src/HOL/Library/Extended_Real.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Extended_Real.thy Mon Apr 25 16:09:26 2016 +0200
@@ -3900,7 +3900,7 @@
shows Liminf_inverse_ereal: "Liminf F (\<lambda>x. inverse (f x)) = inverse (Limsup F f)"
and Limsup_inverse_ereal: "Limsup F (\<lambda>x. inverse (f x)) = inverse (Liminf F f)"
proof -
- def inv \<equiv> "\<lambda>x. if x \<le> 0 then \<infinity> else inverse x :: ereal"
+ define inv where [abs_def]: "inv x = (if x \<le> 0 then \<infinity> else inverse x)" for x :: ereal
have "continuous_on ({..0} \<union> {0..}) inv" unfolding inv_def
by (intro continuous_on_If) (auto intro!: continuous_intros)
also have "{..0} \<union> {0..} = (UNIV :: ereal set)" by auto
--- a/src/HOL/Library/FSet.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/FSet.thy Mon Apr 25 16:09:26 2016 +0200
@@ -939,7 +939,8 @@
BNF_Def.Grp {a. fset a \<subseteq> {(a, b). R a b}} (fimage snd)) a b" (is "?L = ?R")
proof
assume ?L
- def R' \<equiv> "the_inv fset (Collect (case_prod R) \<inter> (fset a \<times> fset b))" (is "the_inv fset ?L'")
+ define R' where "R' =
+ the_inv fset (Collect (case_prod R) \<inter> (fset a \<times> 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
--- a/src/HOL/Library/Float.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Float.thy Mon Apr 25 16:09:26 2016 +0200
@@ -953,7 +953,7 @@
then show ?thesis by simp
next
case False
- def n \<equiv> "\<lfloor>log 2 (real_of_int x)\<rfloor>"
+ define n where "n = \<lfloor>log 2 (real_of_int x)\<rfloor>"
then have "0 \<le> n"
using \<open>2 \<le> x\<close> by simp
from \<open>2 \<le> x\<close> False have "x mod 2 = 1" "\<not> 2 dvd x"
@@ -1306,7 +1306,7 @@
show ?thesis
proof (cases "0 \<le> l")
case True
- def x' \<equiv> "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' \<equiv> "y * 2 ^ nat (- l)"
+ define y' where "y' = y * 2 ^ nat (- l)"
from \<open>y \<noteq> 0\<close> have "y' \<noteq> 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 \<equiv> "\<lfloor>log 2 \<bar>ai\<bar>\<rfloor>"
+ define k where "k = \<lfloor>log 2 \<bar>ai\<bar>\<rfloor>"
then have "\<lfloor>log 2 \<bar>ai\<bar>\<rfloor> = k"
by simp
then have k: "2 powr k \<le> \<bar>ai\<bar>" "\<bar>ai\<bar> < 2 powr (k + 1)"
by (simp_all add: floor_log_eq_powr_iff \<open>ai \<noteq> 0\<close>)
have "k \<ge> 0"
using assms by (auto simp: k_def)
- def r \<equiv> "\<bar>ai\<bar> - 2 ^ nat k"
+ define r where "r = \<bar>ai\<bar> - 2 ^ nat k"
have r: "0 \<le> r" "r < 2 powr k"
using \<open>k \<ge> 0\<close> k
by (auto simp: r_def k_def algebra_simps powr_add abs_if powr_int)
--- a/src/HOL/Library/Formal_Power_Series.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Formal_Power_Series.thy Mon Apr 25 16:09:26 2016 +0200
@@ -1157,8 +1157,8 @@
shows "f div g * g = f"
proof (cases "f = 0")
assume nz: "f \<noteq> 0"
- def n \<equiv> "subdegree g"
- def h \<equiv> "fps_shift n g"
+ define n where "n = subdegree g"
+ define h where "h = fps_shift n g"
from assms have [simp]: "h $ 0 \<noteq> 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 \<equiv> "subdegree g"
- def h \<equiv> "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 \<or> f = 0")
@@ -1248,8 +1248,8 @@
assume "h \<noteq> 0"
show "(h * f) div (h * g) = f div g"
proof -
- def m \<equiv> "subdegree h"
- def h' \<equiv> "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 \<open>h \<noteq> 0\<close> have [simp]: "h'$0 \<noteq> 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 \<noteq> 0"
- def n \<equiv> "subdegree h"
- def h' \<equiv> "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"
--- a/src/HOL/Library/Fun_Lexorder.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Fun_Lexorder.thy Mon Apr 25 16:09:26 2016 +0200
@@ -76,12 +76,12 @@
assumes "finite {k. f k \<noteq> g k}"
shows "less_fun f g \<or> f = g \<or> less_fun g f"
proof -
- { def K \<equiv> "{k. f k \<noteq> g k}"
+ { define K where "K = {k. f k \<noteq> g k}"
assume "f \<noteq> g"
then obtain k' where "f k' \<noteq> g k'" by auto
then have [simp]: "K \<noteq> {}" by (auto simp add: K_def)
with assms have [simp]: "finite K" by (simp add: K_def)
- def q \<equiv> "Min K"
+ define q where "q = Min K"
then have "q \<in> K" and "\<And>k. k \<in> K \<Longrightarrow> k \<ge> q" by auto
then have "\<And>k. \<not> k \<ge> q \<Longrightarrow> k \<notin> K" by blast
then have *: "\<And>k. k < q \<Longrightarrow> f k = g k" by (simp add: K_def)
--- a/src/HOL/Library/Function_Growth.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Function_Growth.thy Mon Apr 25 16:09:26 2016 +0200
@@ -35,7 +35,7 @@
assumes "m \<ge> 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 \<open>a \<noteq> 0\<close> by (simp add: power_add)
qed
--- a/src/HOL/Library/More_List.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/More_List.thy Mon Apr 25 16:09:26 2016 +0200
@@ -165,7 +165,7 @@
lemma strip_while_idem_iff:
"strip_while P xs = xs \<longleftrightarrow> no_trailing P xs"
proof -
- def ys \<equiv> "rev xs"
+ define ys where "ys = rev xs"
moreover have "strip_while P (rev ys) = rev ys \<longleftrightarrow> no_trailing P (rev ys)"
by (simp add: dropWhile_idem_iff)
ultimately show ?thesis by simp
--- a/src/HOL/Library/Multiset.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Multiset.thy Mon Apr 25 16:09:26 2016 +0200
@@ -1071,7 +1071,7 @@
by (simp add: fold_mset_def del: count_union)
next
case True
- def N \<equiv> "set_mset M - {x}"
+ define N where "N = set_mset M - {x}"
from N_def True have *: "set_mset M = insert x N" "x \<notin> N" "finite N" by auto
then have "Finite_Set.fold (\<lambda>y. f y ^^ count (M + {#x#}) y) s N =
Finite_Set.fold (\<lambda>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 \<equiv> "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' \<equiv> "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
--- a/src/HOL/Library/Multiset_Order.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Multiset_Order.thy Mon Apr 25 16:09:26 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 \<equiv> "N - M"
- def Y \<equiv> "M - N"
+ define X where "X = N - M"
+ define Y where "Y = M - N"
from z show "X \<noteq> {#}" unfolding X_def by (auto simp: multiset_eq_iff not_less_eq_eq Suc_le_eq)
from z show "X \<le># 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
--- a/src/HOL/Library/Omega_Words_Fun.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Omega_Words_Fun.thy Mon Apr 25 16:09:26 2016 +0200
@@ -355,7 +355,7 @@
assumes "A \<inter> range w \<noteq> {}"
obtains u a v where "w = u \<frown> [a] \<frown> v" "A \<inter> set u = {}" "a \<in> A"
proof -
- def i \<equiv> "LEAST i. w i \<in> A"
+ define i where "i = (LEAST i. w i \<in> A)"
show ?thesis
proof
show "w = prefix i w \<frown> [w i] \<frown> suffix (Suc i) w"
--- a/src/HOL/Library/Polynomial.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Polynomial.thy Mon Apr 25 16:09:26 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 \<equiv> "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 \<equiv> ?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 \<or> 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 \<equiv> "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 \<equiv> "inverse (coeff g (degree g))"
- def h \<equiv> "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 \<noteq> 0" unfolding h_def lc_def by auto
hence h0: "h \<noteq> 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 \<equiv> "inverse (coeff g (degree g))"
+ define ilc where "ilc = inverse (coeff g (degree g))"
from False have ilc: "ilc \<noteq> 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" "\<And>a p b q. P p q \<Longrightarrow> P (pCons a p) (pCons b q)"
shows "P p q"
proof -
- def n \<equiv> "max (length (coeffs p)) (length (coeffs q))"
- def xs \<equiv> "coeffs p @ (replicate (n - length (coeffs p)) 0)"
- def ys \<equiv> "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\<noteq>0"
then obtain n1 where gte_lcoeff:"\<forall>x\<ge>n1. lead_coeff p \<le> poly p x" using that pCons by auto
have gt_0:"lead_coeff p >0" using pCons(3) \<open>p\<noteq>0\<close> by auto
- def n\<equiv>"max n1 (1+ \<bar>a\<bar>/(lead_coeff p))"
+ define n where "n = max n1 (1+ \<bar>a\<bar>/(lead_coeff p))"
show ?thesis
proof (rule_tac x=n in exI,rule,rule)
fix x assume "n \<le> x"
@@ -3351,13 +3351,13 @@
shows "algebraic x \<longleftrightarrow> (\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0)"
proof safe
fix p assume rat: "\<forall>i. coeff p i \<in> \<rat>" and root: "poly p x = 0" and nz: "p \<noteq> 0"
- def cs \<equiv> "coeffs p"
+ define cs where "cs = coeffs p"
from rat have "\<forall>c\<in>range (coeff p). \<exists>c'. c = of_rat c'" unfolding Rats_def by blast
then obtain f where f: "\<And>i. coeff p i = of_rat (f (coeff p i))"
by (subst (asm) bchoice_iff) blast
- def cs' \<equiv> "map (quotient_of \<circ> f) (coeffs p)"
- def d \<equiv> "Lcm (set (map snd cs'))"
- def p' \<equiv> "smult (of_int d) p"
+ define cs' where "cs' = map (quotient_of \<circ> f) (coeffs p)"
+ define d where "d = Lcm (set (map snd cs'))"
+ define p' where "p' = smult (of_int d) p"
have "\<forall>n. coeff p' n \<in> \<int>"
proof
@@ -3365,8 +3365,9 @@
show "coeff p' n \<in> \<int>"
proof (cases "n \<le> degree p")
case True
- def c \<equiv> "coeff p n"
- def a \<equiv> "fst (quotient_of (f (coeff p n)))" and b \<equiv> "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 \<noteq> 0 \<Longrightarrow> order a (p * q) = order a p + order a q"
proof -
- def i \<equiv> "order a p"
- def j \<equiv> "order a q"
- def t \<equiv> "[:-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: "\<And>u. t dvd u \<longleftrightarrow> poly u a = 0"
unfolding t_def by (simp add: dvd_iff_poly_eq_0)
assume "p * q \<noteq> 0"
--- a/src/HOL/Library/Product_Vector.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/Product_Vector.thy Mon Apr 25 16:09:26 2016 +0200
@@ -125,12 +125,14 @@
fix x assume "x \<in> S"
then obtain e where "0 < e" and S: "\<forall>y. dist y x < e \<longrightarrow> y \<in> S"
using * by fast
- def r \<equiv> "e / sqrt 2" and s \<equiv> "e / sqrt 2"
+ define r where "r = e / sqrt 2"
+ define s where "s = e / sqrt 2"
from \<open>0 < e\<close> have "0 < r" and "0 < s"
unfolding r_def s_def by simp_all
from \<open>0 < e\<close> have "e = sqrt (r\<^sup>2 + s\<^sup>2)"
unfolding r_def s_def by (simp add: power_divide)
- def A \<equiv> "{y. dist (fst x) y < r}" and B \<equiv> "{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 \<in> A \<times> B"
--- a/src/HOL/Library/RBT_Impl.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/RBT_Impl.thy Mon Apr 25 16:09:26 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 \<and> inv1 mt" by simp
hence "inv2 (rbt_del xx mt) \<and> (color_of mt = R \<and> bheight (rbt_del xx mt) = bheight mt \<and> inv1 (rbt_del xx mt) \<or> color_of mt = B \<and> bheight (rbt_del xx mt) = bheight mt - 1 \<and> 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 \<or> xx \<noteq> k \<and> entry_in_tree k v mt \<or> (k = yy \<and> v = ss) \<or> 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 \<and> inv1 mt" by simp
hence "inv2 (rbt_del xx mt) \<and> (color_of mt = R \<and> bheight (rbt_del xx mt) = bheight mt \<and> inv1 (rbt_del xx mt) \<or> color_of mt = B \<and> bheight (rbt_del xx mt) = bheight mt - 1 \<and> 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 \<or> (k = yy \<and> v = ss) \<or> False \<or> xx \<noteq> k \<and> 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 \<equiv> "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 \<Rightarrow> case rbt_lookup t2 k of None \<Rightarrow> Some v
| Some w \<Rightarrow> Some (f k w v))"
proof -
- def xs \<equiv> "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]
--- a/src/HOL/Library/While_Combinator.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Library/While_Combinator.thy Mon Apr 25 16:09:26 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 i<k. b ((c ^^ i) s)"
--- a/src/HOL/Lifting_Set.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Lifting_Set.thy Mon Apr 25 16:09:26 2016 +0200
@@ -80,7 +80,7 @@
assumes "bi_unique R" and "rel_set R X Y"
obtains f where "Y = image f X" and "inj_on f X" and "\<forall>x\<in>X. R x (f x)"
proof
- def f \<equiv> "\<lambda>x. THE y. R x y"
+ define f where "f x = (THE y. R x y)" for x
{ fix x assume "x \<in> X"
with \<open>rel_set R X Y\<close> \<open>bi_unique R\<close> have "R x (f x)"
by (simp add: bi_unique_def rel_set_def f_def) (metis theI)
--- a/src/HOL/Limits.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Limits.thy Mon Apr 25 16:09:26 2016 +0200
@@ -1070,8 +1070,9 @@
lemma tendsto_of_nat [tendsto_intros]:
"filterlim (of_nat :: nat \<Rightarrow> '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 \<equiv> "nat \<lceil>r\<rceil>"
+ fix r :: real
+ assume r: "r > 0"
+ define n where "n = nat \<lceil>r\<rceil>"
from r have n: "\<forall>m\<ge>n. of_nat m \<ge> r" unfolding n_def by linarith
from eventually_ge_at_top[of n] show "eventually (\<lambda>m. norm (of_nat m :: 'a) \<ge> r) sequentially"
by eventually_elim (insert n, simp_all)
@@ -2199,8 +2200,9 @@
assumes local: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<exists>d>0. \<forall>a b. a \<le> x \<and> x \<le> b \<and> b - a < d \<longrightarrow> P a b"
shows "P a b"
proof -
- def bisect \<equiv> "rec_nat (a, b) (\<lambda>n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))"
- def l \<equiv> "\<lambda>n. fst (bisect n)" and u \<equiv> "\<lambda>n. snd (bisect n)"
+ define bisect where "bisect =
+ rec_nat (a, b) (\<lambda>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" "\<And>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" "\<And>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 \<le> b", rule compactI)
fix C assume C: "a \<le> b" "\<forall>t\<in>C. open t" "{a..b} \<subseteq> \<Union>C"
- def T == "{a .. b}"
+ define T where "T = {a .. b}"
from C(1,3) show "\<exists>C'\<subseteq>C. finite C' \<and> {a..b} \<subseteq> \<Union>C'"
proof (induct rule: Bolzano)
case (trans a b c)
--- a/src/HOL/List.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/List.thy Mon Apr 25 16:09:26 2016 +0200
@@ -3468,7 +3468,7 @@
case (3 x1 x2 xs ys)
let ?xs = "x1 # x2 # xs"
let ?cond = "x1 = x2"
- def zs \<equiv> "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: \<open>f 0 = 0\<close> \<open>f (Suc 0) = Suc 0\<close>)
- def f' \<equiv> "\<lambda>x. f (Suc x) - 1"
+ define f' where "f' x = f (Suc x) - 1" for x
{ fix i
have "Suc 0 \<le> f (Suc 0)" using f_nth[of 0] \<open>x1 \<noteq> x2\<close> \<open>f 0 = 0\<close> by auto
@@ -4046,8 +4046,8 @@
proof (induct "length (xs @ ys)" arbitrary: xs ys rule: less_induct)
case less
- def xs' \<equiv> "if (length xs \<le> length ys) then xs else ys"
- and ys' \<equiv> "if (length xs \<le> length ys) then ys else xs"
+ define xs' ys' where "xs' = (if (length xs \<le> length ys) then xs else ys)"
+ and "ys' = (if (length xs \<le> length ys) then ys else xs)"
then have
prems': "length xs' \<le> length ys'"
"xs' @ ys' = ys' @ xs'"
@@ -4499,7 +4499,7 @@
shows "transpose xs ! i = map (\<lambda>xs. xs ! i) [ys \<leftarrow> 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 \<noteq> []" by auto
thus ?case
proof (cases i)
--- a/src/HOL/MacLaurin.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/MacLaurin.thy Mon Apr 25 16:09:26 2016 +0200
@@ -85,22 +85,22 @@
(\<Sum>m<n. diff m (0::real) / (fact m) * h ^ m) + B * (h ^ n / (fact n))"
using Maclaurin_lemma [OF h] ..
- def g \<equiv> "(\<lambda>t. f t -
- (setsum (\<lambda>m. (diff m 0 / (fact m)) * t^m) {..<n} + (B * (t^n / (fact n)))))"
+ define g where [abs_def]: "g t =
+ f t - (setsum (\<lambda>m. (diff m 0 / (fact m)) * t^m) {..<n} + (B * (t^n / (fact n))))" for t
have g2: "g 0 = 0 & g h = 0"
by (simp add: m f_h g_def lessThan_Suc_eq_insert_0 image_iff diff_0 setsum.reindex)
- def difg \<equiv> "(%m t. diff m t -
- (setsum (%p. (diff (m + p) 0 / (fact p)) * (t ^ p)) {..<n-m}
- + (B * ((t ^ (n - m)) / (fact (n - m))))))"
+ define difg where [abs_def]: "difg m t =
+ diff m t - (setsum (%p. (diff (m + p) 0 / (fact p)) * (t ^ p)) {..<n-m}
+ + (B * ((t ^ (n - m)) / (fact (n - m)))))" for m t
have difg_0: "difg 0 = g"
unfolding difg_def g_def by (simp add: diff_0)
have difg_Suc: "\<forall>(m::nat) t::real.
m < n \<and> (0::real) \<le> t \<and> t \<le> h \<longrightarrow> 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: "\<forall>m<n. difg m 0 = 0"
by (auto simp: difg_def m Suc_diff_le lessThan_Suc_eq_insert_0 image_iff setsum.reindex)
--- a/src/HOL/Metis_Examples/Tarski.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Metis_Examples/Tarski.thy Mon Apr 25 16:09:26 2016 +0200
@@ -461,7 +461,7 @@
(*never proved, 2007-01-22*)
apply (rule transE)
-- {* instantiates @{text "(x, ?z) \<in> 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)) \<in> r"} *}
apply (rule_tac f = "f" in monotoneE)
--- a/src/HOL/MicroJava/Comp/Index.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/MicroJava/Comp/Index.thy Mon Apr 25 16:09:26 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 \<times> ty) list] \<Rightarrow> bool" where
"disjoint_varnames pns lvars \<equiv>
--- a/src/HOL/Multivariate_Analysis/Bounded_Linear_Function.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Bounded_Linear_Function.thy Mon Apr 25 16:09:26 2016 +0200
@@ -244,7 +244,7 @@
hence "convergent (\<lambda>n. X n x)"
by (metis Cauchy_convergent_iff)
} note convergent_norm1 = this
- def y \<equiv> "x /\<^sub>R norm x"
+ define y where "y = x /\<^sub>R norm x"
have y: "norm y \<le> 1" and xy: "x = norm x *\<^sub>R y"
by (simp_all add: y_def inverse_eq_divide)
have "convergent (\<lambda>n. norm x *\<^sub>R X n y)"
@@ -301,7 +301,7 @@
have "X \<longlonglongrightarrow> Blinfun v"
proof (rule LIMSEQ_I)
fix r::real assume "r > 0"
- def r' \<equiv> "r / 2"
+ define r' where "r' = r / 2"
have "0 < r'" "r' < r" using \<open>r > 0\<close> by (simp_all add: r'_def)
from CauchyD[OF \<open>Cauchy X\<close> \<open>r' > 0\<close>]
obtain M where M: "\<And>m n. m \<ge> M \<Longrightarrow> n \<ge> M \<Longrightarrow> norm (X m - X n) < r'"
--- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Mon Apr 25 16:09:26 2016 +0200
@@ -602,7 +602,8 @@
have n_in_upd: "\<And>i. n \<in> upd ` {..< Suc i}"
using \<open>upd 0 = n\<close> by auto
- def f' \<equiv> "\<lambda>i j. if j \<in> (upd\<circ>Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j"
+ define f' where "f' i j =
+ (if j \<in> (upd\<circ>Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j)" for i j
{ fix x i assume i[arith]: "i \<le> n" then have "enum (Suc i) x = f' i x"
unfolding f'_def enum_def using \<open>a n < p\<close> \<open>a = enum 0\<close> \<open>upd 0 = n\<close> \<open>a n = p - 1\<close>
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 \<equiv> "base (n := p - 1)"
- def u \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> n | Suc i \<Rightarrow> upd i"
+ define b where "b = base (n := p - 1)"
+ define u where "u i = (case i of 0 \<Rightarrow> n | Suc i \<Rightarrow> upd i)" for i
have "ksimplex p (Suc n) (s' \<union> {b})"
proof (rule ksimplex.intros, standard)
@@ -632,7 +633,7 @@
then show "bij_betw u {..<Suc n} {..<Suc n}"
by (simp add: u_def lessThan_Suc[symmetric] lessThan_Suc_eq_insert_0)
- def f' \<equiv> "\<lambda>i j. if j \<in> u`{..< i} then Suc (b j) else b j"
+ define f' where "f' i j = (if j \<in> u`{..< i} then Suc (b j) else b j)" for i j
have u_eq: "\<And>i. i \<le> n \<Longrightarrow> u ` {..< Suc i} = upd ` {..< i} \<union> { n }"
by (auto simp: u_def image_iff upd_inj Ball_def split: nat.split) arith
@@ -724,7 +725,7 @@
then have "\<exists>!s'. s' \<noteq> s \<and> ksimplex p n s' \<and> (\<exists>b\<in>s'. s - {a} = s'- {b})"
proof (elim disjE conjE)
assume "i = 0"
- def rot \<equiv> "\<lambda>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 \<circ> rot"
have rot: "bij_betw rot {..< n} {..< n}"
@@ -733,7 +734,8 @@
from rot upd have "bij_betw ?upd {..<n} {..<n}"
by (rule bij_betw_trans)
- def f' \<equiv> "\<lambda>i j. if j \<in> ?upd`{..< i} then Suc (enum (Suc 0) j) else enum (Suc 0) j"
+ define f' where [abs_def]: "f' i j =
+ (if j \<in> ?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 \<circ> rot" "f' ` {.. n}"
proof
@@ -832,7 +834,7 @@
from \<open>n \<noteq> 0\<close> obtain n' where n': "n = Suc n'"
by (cases n) auto
- def rot \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> n' | Suc i \<Rightarrow> i"
+ define rot where "rot i = (case i of 0 \<Rightarrow> n' | Suc i \<Rightarrow> i)" for i
let ?upd = "upd \<circ> rot"
have rot: "bij_betw rot {..< n} {..< n}"
@@ -841,8 +843,8 @@
from rot upd have "bij_betw ?upd {..<n} {..<n}"
by (rule bij_betw_trans)
- def b \<equiv> "base (upd n' := base (upd n') - 1)"
- def f' \<equiv> "\<lambda>i j. if j \<in> ?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 \<in> ?upd`{..< i} then Suc (b j) else b j)" for i j
interpret b: kuhn_simplex p n b "upd \<circ> rot" "f' ` {.. n}"
proof
@@ -945,7 +947,7 @@
done
next
assume i: "0 < i" "i < n"
- def i' \<equiv> "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' \<equiv> "\<lambda>i j. if j \<in> ?upd`{..< i} then Suc (base j) else base j"
+ define f' where [abs_def]: "f' i j = (if j \<in> ?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 \<in> {..<n} \<rightarrow> {..<p}" by fact
@@ -1056,7 +1059,7 @@
using eq_sma \<open>a \<in> s\<close> \<open>c \<in> t\<close> by auto
next
assume u: "u l = upd (Suc i')"
- def B \<equiv> "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 \<subseteq> unit_cube"
shows "\<exists>x\<in>unit_cube. f x = x"
proof (rule ccontr)
- def n \<equiv> "DIM('a)"
+ define n where "n = DIM('a)"
have n: "1 \<le> n" "0 < n" "n \<noteq> 0"
unfolding n_def by (auto simp add: Suc_le_eq DIM_positive)
assume "\<not> ?thesis"
@@ -1625,7 +1628,7 @@
obtain b :: "nat \<Rightarrow> 'a" where b: "bij_betw b {..< n} Basis"
by atomize_elim (auto simp: n_def intro!: finite_same_card_bij)
- def b' \<equiv> "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: "\<And>i. i \<in> Basis \<Longrightarrow> b' i \<in> {..< n}"
@@ -1668,7 +1671,7 @@
(label (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i) \<circ> b) i \<noteq>
(label (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i) \<circ> b) i"
by (rule kuhn_lemma[OF q1 q2 q3])
- def z \<equiv> "(\<Sum>i\<in>Basis. (real (q (b' i)) / real p) *\<^sub>R i)::'a"
+ define z :: 'a where "z = (\<Sum>i\<in>Basis. (real (q (b' i)) / real p) *\<^sub>R i)"
have "\<exists>i\<in>Basis. d / real n \<le> \<bar>(f z - z)\<bullet>i\<bar>"
proof (rule ccontr)
have "\<forall>i\<in>Basis. q (b' i) \<in> {0..p}"
@@ -1710,7 +1713,7 @@
using q(2)[rule_format,OF *] by blast
have b'_im: "\<And>i. i \<in> Basis \<Longrightarrow> b' i < n"
using b' unfolding bij_betw_def by auto
- def r' \<equiv> "(\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i)::'a"
+ define r' ::'a where "r' = (\<Sum>i\<in>Basis. (real (r (b' i)) / real p) *\<^sub>R i)"
have "\<And>i. i \<in> Basis \<Longrightarrow> r (b' i) \<le> 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' \<equiv> "(\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i)::'a"
+ define s' :: 'a where "s' = (\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i)"
have "\<And>i. i \<in> Basis \<Longrightarrow> s (b' i) \<le> p"
apply (rule order_trans)
apply (rule rs(2)[OF b'_im, THEN conjunct2])
--- a/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Mon Apr 25 16:09:26 2016 +0200
@@ -1941,9 +1941,9 @@
e * (K/2)^2 \<le> 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' \<equiv> "midpoint b c"
- def b' \<equiv> "midpoint c a"
- def c' \<equiv> "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 \<noteq> 0"
- def K \<equiv> "1 + max (dist a b) (max (dist b c) (dist c a))"
- def e \<equiv> "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 \<ge> 1" by (simp add: K_def max.coboundedI1)
then have K: "K > 0" by linarith
have [iff]: "dist a b \<le> K" "dist b c \<le> K" "dist c a \<le> K"
by (simp_all add: K_def)
have e: "e > 0"
unfolding e_def using ynz K1 by simp
- def At \<equiv> "\<lambda>x y z n. convex hull {x,y,z} \<subseteq> convex hull {a,b,c} \<and>
- dist x y \<le> K/2^n \<and> dist y z \<le> K/2^n \<and> dist z x \<le> K/2^n \<and>
- norm(?pathint x y + ?pathint y z + ?pathint z x) \<ge> e*(K/2^n)^2"
+ define At where "At x y z n \<longleftrightarrow>
+ convex hull {x,y,z} \<subseteq> convex hull {a,b,c} \<and>
+ dist x y \<le> K/2^n \<and> dist y z \<le> K/2^n \<and> dist z x \<le> K/2^n \<and>
+ norm(?pathint x y + ?pathint y z + ?pathint z x) \<ge> 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: "\<And>x x'. \<lbrakk>x\<in>convex hull {a, b, c}; x'\<in>convex hull {a, b, c}; cmod (x' - x) < d1\<rbrakk>
\<Longrightarrow> cmod (f x' - f x) < cmod y / (24 * C)"
- def e \<equiv> "min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))"
- def shrink \<equiv> "\<lambda>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 = "\<lambda>x y. contour_integral(linepath x y) f"
have e: "0 < e" "e \<le> 1" "e \<le> d1 / (4 * C)" "e \<le> cmod y / 24 / C / B"
using d1_pos \<open>C>0\<close> \<open>B>0\<close> 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: "\<And>z. z \<in> path_image p \<Longrightarrow> 0 < ee z \<and> ball z (ee z) \<subseteq> s"
by metis
- def cover \<equiv> "(\<lambda>z. ball z (ee z/3)) ` (path_image p)"
+ define cover where "cover = (\<lambda>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 \<subseteq> (\<Union>c\<in>path_image p. ball c (ee c / 3))"
@@ -3001,7 +3003,7 @@
using k by (auto simp: path_image_def)
then have eepi: "\<And>i. i \<in> k \<Longrightarrow> 0 < ee((p i))"
by (metis ee)
- def e \<equiv> "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 \<and> pathstart h = pathstart \<gamma> \<and> pathfinish h = pathfinish \<gamma> \<and>
(\<forall>t \<in> {0..1}. norm(h t - \<gamma> t) < d/2)"
using path_approx_polynomial_function [OF \<open>path \<gamma>\<close>, of "d/2"] d by auto
- def nn \<equiv> "1/(2* pi*ii) * contour_integral h (\<lambda>w. 1/(w - z))"
+ define nn where "nn = 1/(2* pi*ii) * contour_integral h (\<lambda>w. 1/(w - z))"
have "\<exists>n. \<forall>e > 0. \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
(\<forall>t \<in> {0..1}. norm(\<gamma> t - p t) < e) \<and>
@@ -3924,7 +3926,7 @@
using path_image_def z by auto
have gpd: "\<gamma> piecewise_C1_differentiable_on {0..1}"
using \<gamma> valid_path_def by blast
- def r \<equiv> "(w - z) / (\<gamma> 0 - z)"
+ define r where "r = (w - z) / (\<gamma> 0 - z)"
have [simp]: "r \<noteq> 0"
using w z by (auto simp: r_def)
have "Arg r \<le> 2*pi"
@@ -3944,7 +3946,7 @@
then obtain t where t: "t \<in> {0..1}"
and eqArg: "Im (integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x)/(\<gamma> x - z))) = Arg r"
by blast
- def i \<equiv> "integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
+ define i where "i = integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
have iArg: "Arg r = Im i"
using eqArg by (simp add: i_def)
have gpdt: "\<gamma> piecewise_C1_differentiable_on {0..t}"
@@ -4433,7 +4435,7 @@
by (metis continuous_at_winding_number valid_path_imp_path \<gamma> z)
then obtain d where "d>0" and d: "\<And>x'. dist x' z < d \<Longrightarrow> dist (winding_number \<gamma> x') (winding_number \<gamma> z) < \<bar>Re(winding_number \<gamma> z)\<bar> - 1/2"
using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast
- def z' \<equiv> "z - (d / (2 * cmod a)) *\<^sub>R a"
+ define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a"
have *: "a \<bullet> z' \<le> 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 @@
\<longrightarrow> contour_integral g2 f = contour_integral g1 f)))"
by metis
note ee_rule = ee [THEN conjunct2, rule_format]
- def C \<equiv> "(\<lambda>t. ball t (ee t / 3)) ` {0..1}"
+ define C where "C = (\<lambda>t. ball t (ee t / 3)) ` {0..1}"
have "\<forall>t \<in> C. open t" by (simp add: C_def)
moreover have "{0..1} \<subseteq> \<Union>C"
using ee [THEN conjunct1] by (auto simp: C_def dist_norm)
ultimately obtain C' where C': "C' \<subseteq> C" "finite C'" and C'01: "{0..1} \<subseteq> \<Union>C'"
by (rule compactE [OF compact_interval])
- def kk \<equiv> "{t \<in> {0..1}. ball t (ee t / 3) \<in> C'}"
+ define kk where "kk = {t \<in> {0..1}. ball t (ee t / 3) \<in> C'}"
have kk01: "kk \<subseteq> {0..1}" by (auto simp: kk_def)
- def e \<equiv> "Min (ee ` kk)"
+ define e where "e = Min (ee ` kk)"
have C'_eq: "C' = (\<lambda>t. ball t (ee t / 3)) ` kk"
using C' by (auto simp: kk_def C_def)
have ee_pos[simp]: "\<And>t. t \<in> {0..1} \<Longrightarrow> 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} \<inter> {0..1})" if "y \<in> k" for y
proof -
- def w \<equiv> "(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 \<le> 1 \<and> exp (\<i> * complex_of_real (t - s) * z) = w})"
apply (rule finite_vimageI [OF finite_bounded_log2])
using \<open>s < t\<close> apply (auto simp: inj_of_real)
@@ -5205,7 +5207,7 @@
case True then show ?thesis by force
next
case False
- def w \<equiv> "x - z"
+ define w where "w = x - z"
then have "w \<noteq> 0" by (simp add: False)
have **: "\<And>t. \<lbrakk>Re w = cos t * cmod w; Im w = sin t * cmod w\<rbrakk> \<Longrightarrow> w = of_real (cmod w) * exp (\<i> * 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' \<equiv> "norm(w - z)"
+ define r' where "r' = norm(w - z)"
have "r' < r"
by (simp add: assms r'_def)
have disjo: "cball z r' \<inter> sphere z r = {}"
@@ -5379,7 +5381,7 @@
apply (blast intro: integrable_uniform_limit_real)
done
{ fix e::real
- def B' \<equiv> "B+1"
+ define B' where "B' = B + 1"
have B': "B' > 0" "B' > B" using \<open>0 \<le> B\<close> by (auto simp: B'_def)
assume "0 < e"
then have ev_no': "\<forall>\<^sub>F n in F. \<forall>x\<in>path_image \<gamma>. 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 \<equiv> "\<lambda>n::nat. \<lambda>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: "\<And>z::complex. z \<noteq> 0 \<Longrightarrow> z ^ (k - Suc 0) = z ^ k / z"
by (simp add: field_simps) (metis Suc_pred \<open>k \<noteq> 0\<close> 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: "\<And>x. B \<le> cmod x \<Longrightarrow> norm (f x) * 2 < cmod (f z)"
by (auto simp: dist_norm)
- def R \<equiv> "1 + \<bar>B\<bar> + norm z"
+ define R where "R = 1 + \<bar>B\<bar> + norm z"
have "R > 0" unfolding R_def
proof -
have "0 \<le> cmod z + \<bar>B\<bar>"
@@ -6384,7 +6388,7 @@
have 2: "((\<lambda>u. 1 / (2 * of_real pi * \<i>) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)"
if w: "w \<in> ball z r" for w
proof -
- def d \<equiv> "(r - norm(w - z))"
+ define d where "d = (r - norm(w - z))"
have "0 < d" "d \<le> r" using w by (auto simp: norm_minus_commute d_def dist_norm)
have dle: "\<And>u. cmod (z - u) = r \<Longrightarrow> d \<le> 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 \<equiv> "(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: "\<And>y. r = cmod (z - y) \<Longrightarrow> d \<le> 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: "\<And>x. x \<in> path_image \<gamma>' \<Longrightarrow> norm x \<le> B"
using bounded_pos by force
- def d \<equiv> "\<lambda>z w. if w = z then deriv f z else (f w - f z)/(w - z)"
- def v \<equiv> "{w. w \<notin> path_image \<gamma> \<and> winding_number \<gamma> 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 \<notin> path_image \<gamma> \<and> winding_number \<gamma> w = 0}"
have "path \<gamma>" "valid_path \<gamma>" using \<gamma>
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 \<equiv> "\<lambda>z. if z \<in> u then contour_integral \<gamma> (d z) else contour_integral \<gamma> (\<lambda>w. f w/(w - z))"
+ define h where
+ "h z = (if z \<in> u then contour_integral \<gamma> (d z) else contour_integral \<gamma> (\<lambda>w. f w/(w - z)))" for z
have U: "\<And>z. z \<in> u \<Longrightarrow> ((d z) has_contour_integral h z) \<gamma>"
apply (simp add: h_def)
apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]])
--- a/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy Mon Apr 25 16:09:26 2016 +0200
@@ -1367,7 +1367,7 @@
proof (rule has_field_derivative_zero_constant)
fix z :: complex assume z': "z \<in> ball 0 1"
hence z: "norm z < 1" by (simp add: dist_0_norm)
- def t \<equiv> "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 "(\<lambda>n. g n * z^n) sums Arctan z"
and "h z sums Arctan z"
proof -
- def G \<equiv> "\<lambda>z. (\<Sum>n. g n * z^n)"
+ define G where [abs_def]: "G z = (\<Sum>n. g n * z^n)" for z
have summable: "summable (\<lambda>n. g n * u^n)" if "norm u < 1" for u
proof (cases "u = 0")
assume u: "u \<noteq> 0"
@@ -2089,7 +2089,7 @@
proof (rule has_field_derivative_zero_constant)
fix u :: complex assume "u \<in> ball 0 1"
hence u: "norm u < 1" by (simp add: dist_0_norm)
- def K \<equiv> "(norm u + 1) / 2"
+ define K where "K = (norm u + 1) / 2"
from u and abs_Im_le_cmod[of u] have Im_u: "\<bar>Im u\<bar> < 1" by linarith
from u have K: "0 \<le> K" "norm u < K" "K < 1" by (simp_all add: K_def)
hence "(G has_field_derivative (\<Sum>n. diffs g n * u ^ n)) (at u)" unfolding G_def
@@ -2120,7 +2120,7 @@
assumes x: "x > (0::real)"
shows "(\<lambda>n. (2*((x - 1) / (x + 1)) ^ (2*n+1) / of_nat (2*n+1))) sums ln x"
proof -
- def y \<equiv> "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 \<noteq> of_real (-1)" by (subst of_real_eq_iff) auto
from x have "\<bar>x - 1\<bar> < \<bar>x + 1\<bar>" by linarith
hence "norm (complex_of_real (x - 1) / complex_of_real (x + 1)) < 1"
--- a/src/HOL/Multivariate_Analysis/Conformal_Mappings.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Conformal_Mappings.thy Mon Apr 25 16:09:26 2016 +0200
@@ -35,7 +35,7 @@
by (rule contf continuous_intros)+
have holf': "(\<lambda>u. (f u - y)) holomorphic_on (ball z r)"
by (simp add: holf holomorphic_on_diff)
- def a \<equiv> "(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 \<open>0 < r\<close> by (simp add: a_def divide_simps)
@@ -105,44 +105,44 @@
case True then show ?thesis by simp
next
case False
- def w \<equiv> "complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))"
- have "1 \<le> abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))"
- using \<open>0 < B\<close> by simp
- then have wge1: "1 \<le> norm w"
- by (metis norm_of_real w_def)
- then have "w \<noteq> 0" by auto
- have kB: "0 < fact k * B"
- using \<open>0 < B\<close> by simp
- then have "0 \<le> fact k * B / cmod ((deriv ^^ k) f 0)"
- by simp
- then have wgeA: "A \<le> 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) + (\<bar>A\<bar> + 1))"
- using \<open>0 < B\<close> 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 "... \<le> 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 \<open>w \<noteq> 0\<close> 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 \<open>k>n\<close> \<open>w \<noteq> 0\<close> \<open>0 < B\<close> 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) + (\<bar>A\<bar> + 1))"
+ have "1 \<le> abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))"
+ using \<open>0 < B\<close> by simp
+ then have wge1: "1 \<le> norm w"
+ by (metis norm_of_real w_def)
+ then have "w \<noteq> 0" by auto
+ have kB: "0 < fact k * B"
+ using \<open>0 < B\<close> by simp
+ then have "0 \<le> fact k * B / cmod ((deriv ^^ k) f 0)"
+ by simp
+ then have wgeA: "A \<le> 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) + (\<bar>A\<bar> + 1))"
+ using \<open>0 < B\<close> 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 "... \<le> 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 \<open>w \<noteq> 0\<close> 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 \<open>k>n\<close> \<open>w \<noteq> 0\<close> \<open>0 < B\<close> 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) * \<xi> ^ (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 \<equiv> "\<lambda>i. a (i+m) / a m"
- def g \<equiv> "\<lambda>x. suminf (\<lambda>i. b i * (x - \<xi>) ^ i)"
+ define b where "b i = a (i+m) / a m" for i
+ define g where "g x = suminf (\<lambda>i. b i * (x - \<xi>) ^ 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 \<xi> e \<subseteq> S"
using \<open>open S\<close> \<open>\<xi> \<in> S\<close> open_contains_cball_eq by blast
- def T \<equiv> "cball \<xi> e \<inter> U"
+ define T where "T = cball \<xi> e \<inter> 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 \<xi> r) \<noteq> {}"
using \<open>0 < r\<close> by simp
- def g \<equiv> "\<lambda>z. inverse (f z)"
+ define g where [abs_def]: "g z = inverse (f z)" for z
have contg: "continuous_on (cball \<xi> r) g"
unfolding g_def using contf continuous_on_inverse fnz' by blast
have holg: "g holomorphic_on ball \<xi> r"
@@ -377,7 +377,7 @@
apply (rule bexE [OF continuous_attains_inf [OF cof frne contfr]])
apply (simp add: dist_norm)
done
- moreover def \<epsilon> \<equiv> "norm (f w - f \<xi>) / 3"
+ moreover define \<epsilon> where "\<epsilon> \<equiv> norm (f w - f \<xi>) / 3"
ultimately have "0 < \<epsilon>"
using \<open>0 < r\<close> dist_complex_def r sne by auto
have "ball (f \<xi>) \<epsilon> \<subseteq> f ` U"
@@ -601,12 +601,12 @@
obtain r where "r>0" and r: "ball \<xi> r \<subseteq> S" using assms by (blast elim!: openE)
then have holfb: "f holomorphic_on ball \<xi> r"
using holf holomorphic_on_subset by blast
- def g \<equiv> "\<lambda>w. suminf (\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i)"
+ define g where "g w = suminf (\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i)" for w
have sumsg: "(\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i) sums g w"
and feq: "f w - f \<xi> = (w - \<xi>)^n * g w"
if w: "w \<in> ball \<xi> r" for w
proof -
- def powf \<equiv> "(\<lambda>i. (deriv ^^ i) f \<xi>/(fact i) * (w - \<xi>)^i)"
+ define powf where "powf = (\<lambda>i. (deriv ^^ i) f \<xi>/(fact i) * (w - \<xi>)^i)"
have sing: "{..<n} - {i. powf i = 0} = (if f \<xi> = 0 then {} else {0})"
unfolding powf_def using \<open>0 < n\<close> 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 \<xi> \<noteq> 0" by blast
obtain r0 where "r0 > 0" "ball \<xi> r0 \<subseteq> S" using S openE \<open>\<xi> \<in> S\<close> by auto
- def n \<equiv> "LEAST n. (deriv ^^ n) f \<xi> \<noteq> 0"
+ define n where "n \<equiv> LEAST n. (deriv ^^ n) f \<xi> \<noteq> 0"
have n_ne: "(deriv ^^ n) f \<xi> \<noteq> 0"
by (rule def_LeastI [OF n_def]) (rule n0)
then have "0 < n" using \<open>f \<xi> = 0\<close>
@@ -766,7 +766,7 @@
"ball \<xi> r \<subseteq> S"
"\<And>w. w \<in> ball \<xi> r \<Longrightarrow> k * norm(w - \<xi>)^n \<le> norm(f w - f \<xi>)"
proof -
- def n \<equiv> "LEAST n. 0 < n \<and> (deriv ^^ n) f \<xi> \<noteq> 0"
+ define n where "n = (LEAST n. 0 < n \<and> (deriv ^^ n) f \<xi> \<noteq> 0)"
obtain n0 where "0 < n0" and n0: "(deriv ^^ n0) f \<xi> \<noteq> 0"
using fne holomorphic_fun_eq_const_on_connected [OF holf S] \<open>\<xi> \<in> S\<close> \<open>\<phi> \<in> S\<close> by blast
then have "0 < n" and n_ne: "(deriv ^^ n) f \<xi> \<noteq> 0"
@@ -781,7 +781,7 @@
obtain e where "e>0" and e: "ball \<xi> e \<subseteq> S" using assms by (blast elim!: openE)
then have holfb: "f holomorphic_on ball \<xi> e"
using holf holomorphic_on_subset by blast
- def d \<equiv> "(min e r) / 2"
+ define d where "d = (min e r) / 2"
have "0 < d" using \<open>0 < r\<close> \<open>0 < e\<close> by (simp add: d_def)
have "d < r"
using \<open>0 < r\<close> 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 "(\<lambda>z. (z - \<xi>) * f z) \<midarrow>\<xi>\<rightarrow> 0"
proof -
- def h \<equiv> "\<lambda>z. (z - \<xi>)^2 * f z"
+ define h where [abs_def]: "h z = (z - \<xi>)^2 * f z" for z
have h0: "(h has_field_derivative 0) (at \<xi>)"
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 \<equiv> "\<lambda>z. if z = \<xi> then deriv h \<xi> else (h z - h \<xi>) / (z - \<xi>)"
+ define g where [abs_def]: "g z = (if z = \<xi> then deriv h \<xi> else (h z - h \<xi>) / (z - \<xi>))" for z
have holg: "g holomorphic_on S"
unfolding g_def by (rule pole_lemma [OF holh \<xi>])
show ?thesis
@@ -1058,7 +1058,7 @@
next
case False
then obtain k where k: "0 < k" "k\<le>n" "a k \<noteq> 0" by force
- def m \<equiv> "GREATEST k. k\<le>n \<and> a k \<noteq> 0"
+ define m where "m = (GREATEST k. k\<le>n \<and> a k \<noteq> 0)"
have m: "m\<le>n \<and> a m \<noteq> 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 \<and> (deriv ^^ n0) f \<xi> \<noteq> 0" by blast
- def n \<equiv> "LEAST n. n > 0 \<and> (deriv ^^ n) f \<xi> \<noteq> 0"
+ define n where [abs_def]: "n = (LEAST n. n > 0 \<and> (deriv ^^ n) f \<xi> \<noteq> 0)"
have n_ne: "n > 0" "(deriv ^^ n) f \<xi> \<noteq> 0"
using def_LeastI [OF n_def n0] by auto
have n_min: "\<And>k. 0 < k \<Longrightarrow> k < n \<Longrightarrow> (deriv ^^ k) f \<xi> = 0"
@@ -1220,7 +1220,7 @@
using \<open>0 < r\<close> \<open>0 < \<delta>\<close>
apply (simp_all add:)
by (meson Topology_Euclidean_Space.open_ball centre_in_ball)
- def U \<equiv> "(\<lambda>w. (w - \<xi>) * g w) ` T"
+ define U where "U = (\<lambda>w. (w - \<xi>) * g w) ` T"
have "open U" by (metis oimT U_def)
have "0 \<in> 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 "\<And>z. norm z < r \<Longrightarrow> f z = z * (h z)" and "deriv f 0 = h 0"
proof -
- def h \<equiv> "\<lambda>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 \<equiv> "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 \<subseteq> f ` (ball a 1)"
proof -
- def r \<equiv> "249/256::real"
+ define r :: real where "r = 249/256"
have "0 < r" "r < 1" by (auto simp: r_def)
- def g \<equiv> "\<lambda>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 \<in> cball a r"
and pge: "\<And>y. y \<in> cball a r \<Longrightarrow> norm (g y) \<le> norm (g p)"
using distance_attains_sup [OF 1 2, of 0] by force
- def t \<equiv> "(r - norm(p - a)) / 2"
+ define t where "t = (r - norm(p - a)) / 2"
have "norm (p - a) \<noteq> r"
using pge [of a] \<open>r > 0\<close> 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 \<equiv> "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 \<open>0 < r\<close> by auto
- have fo: "(\<lambda>z. f (a + of_real r * z)) = f o (\<lambda>z. (a + of_real r * z))"
- by (simp add: o_def)
- have holf': "f holomorphic_on (\<lambda>z. a + complex_of_real r * z) ` ball 0 1"
- apply (rule holomorphic_on_subset [OF holf])
- using \<open>0 < r\<close> apply (force simp: dist_norm norm_mult)
- done
- have 1: "(\<lambda>z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1"
- apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+
- using \<open>0 < r\<close> by (simp add: C_def False)
- have "((\<lambda>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 \<open>0 < r\<close> by auto
+ have fo: "(\<lambda>z. f (a + of_real r * z)) = f o (\<lambda>z. (a + of_real r * z))"
+ by (simp add: o_def)
+ have holf': "f holomorphic_on (\<lambda>z. a + complex_of_real r * z) ` ball 0 1"
+ apply (rule holomorphic_on_subset [OF holf])
+ using \<open>0 < r\<close> apply (force simp: dist_norm norm_mult)
+ done
+ have 1: "(\<lambda>z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1"
+ apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+
+ using \<open>0 < r\<close> by (simp add: C_def False)
+ have "((\<lambda>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 *: "((\<lambda>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 \<open>0 < r\<close> apply (simp add: dist_norm norm_mult that)
- apply (rule derivative_eq_intros | simp)+
- done
- show ?thesis
- apply (rule derivative_eq_intros * | simp)+
- using \<open>0 < r\<close> by (auto simp: C_def False)
- qed
- have 2: "deriv (\<lambda>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 \<open>0 < r\<close>
- 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 \<open>0 < r\<close> apply (simp add: dist_norm norm_mult that)
+ apply (rule derivative_eq_intros | simp)+
done
- have sb1: "op * (C * r) ` (\<lambda>z. f (a + of_real r * z) / (C * r)) ` ball 0 1
- \<subseteq> f ` ball a r"
- using \<open>0 < r\<close> by (auto simp: dist_norm norm_mult C_def False)
- have sb2: "ball (C * r * b) r' \<subseteq> op * (C * r) ` ball b t"
- if "1 / 12 < t" for b t
- proof -
- have *: "r * cmod (deriv f a) / 12 \<le> r * (t * cmod (deriv f a))"
- using that \<open>0 < r\<close> 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 \<open>0 < r\<close>
- 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 = "\<lambda>z. (C * of_real r) * z"])
- using sb1 sb2
- apply force
+ apply (rule derivative_eq_intros * | simp)+
+ using \<open>0 < r\<close> by (auto simp: C_def False)
+ qed
+ have 2: "deriv (\<lambda>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 \<open>0 < r\<close>
+ apply (simp add: C_def False fo)
+ apply (simp add: derivative_intros dfa complex_derivative_chain)
+ done
+ have sb1: "op * (C * r) ` (\<lambda>z. f (a + of_real r * z) / (C * r)) ` ball 0 1
+ \<subseteq> f ` ball a r"
+ using \<open>0 < r\<close> by (auto simp: dist_norm norm_mult C_def False)
+ have sb2: "ball (C * r * b) r' \<subseteq> op * (C * r) ` ball b t"
+ if "1 / 12 < t" for b t
+ proof -
+ have *: "r * cmod (deriv f a) / 12 \<le> r * (t * cmod (deriv f a))"
+ using that \<open>0 < r\<close> 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 \<open>0 < r\<close>
+ 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 = "\<lambda>z. (C * of_real r) * z"])
+ using sb1 sb2
+ apply force
+ done
qed
corollary Bloch_general:
@@ -2157,8 +2157,8 @@
then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk>
\<Longrightarrow> (x' - z) ^ (n - m) * g x' = h x'" for x' by auto
next
- def F\<equiv>"at z within ball z r"
- and f'\<equiv>"\<lambda>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 \<open>n>m\<close> 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 "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk>
\<Longrightarrow> (x' - z) ^ (m - n) * h x' = g x'" for x' by auto
next
- def F\<equiv>"at z within ball z r"
- and f'\<equiv>"\<lambda>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 \<open>m>n\<close> unfolding f'_def by auto
moreover have "continuous F f'" unfolding f'_def F_def
by (intro continuous_intros)
@@ -2224,7 +2224,7 @@
"\<And>w. w \<in> ball z r \<Longrightarrow> g w \<noteq> 0"
using holomorphic_factor_zero_nonconstant[OF holo \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close> \<open>f z=0\<close>]
by (metis assms(3) assms(5) assms(6))
- def r'\<equiv>"r/2"
+ define r' where "r' = r/2"
have "cball z r' \<subseteq> ball z r" unfolding r'_def by (simp add: \<open>0 < r\<close> cball_subset_ball_iff)
hence "cball z r' \<subseteq> s" "g holomorphic_on cball z r'"
"(\<forall>w\<in>cball z r'. f w = (w - z) ^ n * g w \<and> g w \<noteq> 0)"
@@ -2236,14 +2236,14 @@
by (simp add:\<open>0 < n\<close>)
next
fix m n
- def fac\<equiv>"\<lambda>n g r. \<forall>w\<in>cball z r. f w = (w - z) ^ n * g w \<and> g w \<noteq> 0"
+ define fac where "fac n g r \<longleftrightarrow> (\<forall>w\<in>cball z r. f w = (w - z) ^ n * g w \<and> g w \<noteq> 0)" for n g r
assume n_asm:"\<exists>g r1. 0 < n \<and> 0 < r1 \<and> g holomorphic_on cball z r1 \<and> fac n g r1"
and m_asm:"\<exists>h r2. 0 < m \<and> 0 < r2 \<and> h holomorphic_on cball z r2 \<and> 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\<equiv>"min r1 r2"
+ define r where "r = min r1 r2"
have "r>0" using \<open>r1>0\<close> \<open>r2>0\<close> unfolding r_def by auto
moreover have "\<forall>w\<in>ball z r. f w = (w-z)^n * g w \<and> g w\<noteq>0 \<and> f w = (w - z)^m * h w \<and> h w\<noteq>0"
using \<open>fac m h r2\<close> \<open>fac n g r1\<close> unfolding fac_def r_def
@@ -2276,7 +2276,7 @@
fix p assume "p\<in>s"
then obtain e1 where "e1>0" and e1: "\<forall>w\<in>ball p e1. w\<in>s \<and> (w\<noteq>p \<longrightarrow> w\<notin>pts)"
using finite_ball_avoid[OF assms] by auto
- def e2\<equiv>"e1/2"
+ define e2 where "e2 = e1/2"
have "e2>0" and "e2<e1" unfolding e2_def using \<open>e1>0\<close> by auto
then have "cball p e2 \<subseteq> ball p e1" by (subst cball_subset_ball_iff,auto)
then show "\<exists>e>0. \<forall>w\<in>cball p e. w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts)" using \<open>e2>0\<close> e1 by auto
@@ -2302,39 +2302,39 @@
using finite_ball_avoid[OF \<open>open s\<close> \<open>finite (insert p pts)\<close>,rule_format,of a]
\<open>a \<in> s - insert p pts\<close>
by auto
- def a'\<equiv>"a+e/2"
+ define a' where "a' = a+e/2"
have "a'\<in>s-{p} -pts" using e[rule_format,of "a+e/2"] \<open>e>0\<close>
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' \<subseteq> 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\<equiv>"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 \<subseteq> s - insert p pts" unfolding g_def
- proof (rule subset_path_image_join)
- have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
- by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
- then show "path_image (linepath a a') \<subseteq> s - insert p pts" using e idt(9)
- by auto
- next
- show "path_image g' \<subseteq> s - insert p pts" using g'(4) by blast
- qed
+ proof (rule subset_path_image_join)
+ have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
+ by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
+ then show "path_image (linepath a a') \<subseteq> s - insert p pts" using e idt(9)
+ by auto
+ next
+ show "path_image g' \<subseteq> s - insert p pts" using g'(4) by blast
+ qed
moreover have "f contour_integrable_on g"
- proof -
- have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
- 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: \<open>e>0\<close>)
- qed
+ proof -
+ have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
+ 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: \<open>e>0\<close>)
+ 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\<equiv>"circlepath p (h p)" and p_circ_pt\<equiv>"linepath (p+h p) (p+h p)"
- def n_circ\<equiv>"\<lambda>n. (op +++ p_circ ^^ n) p_circ_pt"
- def cp\<equiv>"if n\<ge>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\<ge>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'\<equiv>"g +++ pg +++ cp +++ (reversepath pg)"
+ define g' where "g' = g +++ pg +++ cp +++ (reversepath pg)"
have "contour_integral g' f = (\<Sum>p\<in>pts. winding_number g' p * contour_integral (circlepath p (h p)) f)"
proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \<open>finite pts\<close> ])
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' \<subseteq> s - {p} - pts"
proof -
- def s'\<equiv>"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 = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
(is "?L=?R")
proof -
- def circ\<equiv>"\<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f"
- def pts1\<equiv>"pts \<inter> s"
- def pts2\<equiv>"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 \<inter> s"
+ define pts2 where "pts2 = pts - pts1"
have "pts=pts1 \<union> pts2" "pts1 \<inter> pts2 = {}" "pts2 \<inter> s={}" "pts1\<subseteq>s"
unfolding pts1_def pts2_def by auto
have "contour_integral g f = (\<Sum>p\<in>pts1. circ p)" unfolding circ_def
@@ -2680,8 +2682,6 @@
qed
-
-
(*order of the zero of f at z*)
definition zorder::"(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> nat" where
"zorder f z = (THE n. n>0 \<and> (\<exists>h r. r>0 \<and> h holomorphic_on cball z r
@@ -2714,8 +2714,8 @@
shows "\<exists>r. n>0 \<and> r>0 \<and> cball z r \<subseteq> s \<and> h holomorphic_on cball z r
\<and> (\<forall>w\<in>cball z r. f w = h w * (w-z)^n \<and> h w \<noteq>0) "
proof -
- def P\<equiv>"\<lambda>h r n. r>0 \<and> h holomorphic_on cball z r
- \<and> (\<forall>w\<in>cball z r. ( f w = h w * (w-z)^n) \<and> h w \<noteq>0)"
+ define P where "P h r n \<longleftrightarrow> r>0 \<and> h holomorphic_on cball z r
+ \<and> (\<forall>w\<in>cball z r. ( f w = h w * (w-z)^n) \<and> h w \<noteq>0)" for h r n
have "(\<exists>!n. n>0 \<and> (\<exists> h r. P h r n))"
proof -
have "\<exists>!n. \<exists>h r. n>0 \<and> 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 \<subseteq> s"
using assms(3) assms(5) open_contains_cball_eq by blast
- def r3\<equiv>"min r1 r2"
+ define r3 where "r3 = min r1 r2"
have "P h r3 n" using \<open>P h r1 n\<close> \<open>r2>0\<close> unfolding P_def r3_def
by auto
moreover have "cball z r3 \<subseteq> s" using \<open>cball z r2 \<subseteq> s\<close> unfolding r3_def by auto
@@ -2755,7 +2755,8 @@
shows "\<exists>r. n>0 \<and> r>0 \<and> cball z r \<subseteq> s \<and> h holomorphic_on cball z r
\<and> (\<forall>w\<in>cball z r. f w = h w / (w-z)^n \<and> h w \<noteq>0)"
proof -
- def zo\<equiv>"zorder (inverse \<circ> f) z" and zp\<equiv>"zer_poly (inverse \<circ> f) z"
+ define zo where "zo = zorder (inverse \<circ> f) z"
+ define zp where "zp = zer_poly (inverse \<circ> f) z"
obtain r where "0 < zo" "0 < r" "cball z r \<subseteq> s" and zp_holo: "zp holomorphic_on cball z r" and
zp_fac: "\<forall>w\<in>cball z r. (inverse \<circ> f) w = zp w * (w - z) ^ zo \<and> zp w \<noteq> 0"
using zorder_exist[OF \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close> holo,folded zo_def zp_def]
@@ -2782,14 +2783,15 @@
shows "\<exists>r>0. cball z r \<subseteq> s
\<and> (f has_contour_integral complex_of_real (2 * pi) * \<i> * residue f z) (circlepath z r)"
proof -
- def n\<equiv>"porder f z" and h\<equiv>"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 \<subseteq> s" and
h_holo:"h holomorphic_on cball z r"
and h:"(\<forall>w\<in>cball z r. f w = h w / (w - z) ^ n \<and> h w \<noteq> 0)"
using porder_exist[OF \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close> holo \<open>f z=0\<close> non_c]
unfolding n_def h_def by auto
- def c\<equiv>"complex_of_real (2 * pi) * \<i>"
+ define c where "c = complex_of_real (2 * pi) * \<i>"
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:"\<forall>p\<in>poles. is_pole f p"
shows "contour_integral \<gamma> f = 2 * pi * \<i> *(\<Sum>p\<in>poles. winding_number \<gamma> p * residue f p)"
proof -
- def pts\<equiv>"{p. f p=0}"
- def c\<equiv>"2 * complex_of_real pi * \<i> "
- def contour\<equiv>"\<lambda>p e. (f has_contour_integral c * residue f p) (circlepath p e)"
- def avoid\<equiv>"\<lambda>p e. \<forall>w\<in>cball p e. w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts)"
+ define pts where "pts = {p. f p=0}"
+ define c where "c = 2 * complex_of_real pi * \<i> "
+ 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 \<longleftrightarrow> (\<forall>w\<in>cball p e. w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts))" for p e
have "poles \<subseteq> pts" and "finite pts"
using poles \<open>finite {p. f p=0}\<close> unfolding pts_def is_pole_def by auto
have "\<exists>e>0. avoid p e \<and> (p\<in>poles \<longrightarrow> contour p e)"
@@ -2832,35 +2835,35 @@
using finite_cball_avoid[OF \<open>open s\<close> \<open>finite pts\<close>] \<open>p\<in>s\<close> unfolding avoid_def by auto
have "\<exists>e2>0. cball p e2 \<subseteq> ball p e1 \<and> contour p e2"
when "p\<in>poles" unfolding c_def contour_def
- proof (rule base_residue[of "ball p e1" p f,simplified,OF \<open>e1>0\<close>])
- show "inverse \<circ> f holomorphic_on ball p e1"
- proof -
- def f'\<equiv>"inverse o f"
- have "f holomorphic_on ball p e1 - {p}"
- using holo e \<open>poles \<subseteq> pts\<close> 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 \<open>p\<in>poles\<close> 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 \<open>p\<in>poles\<close> poles unfolding is_pole_def by auto
- next
- def p'\<equiv>"p+e1/2"
- have "p'\<in>ball p e1" and "p'\<noteq>p" using \<open>e1>0\<close> unfolding p'_def by (auto simp add:dist_norm)
- then show "\<exists>w\<in>ball p e1. f w \<noteq> 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 \<open>e1>0\<close>])
+ show "inverse \<circ> 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 \<open>poles \<subseteq> pts\<close> 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 \<open>p\<in>poles\<close> 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 \<open>p\<in>poles\<close> poles unfolding is_pole_def by auto
+ next
+ define p' where "p' = p+e1/2"
+ have "p'\<in>ball p e1" and "p'\<noteq>p" using \<open>e1>0\<close> unfolding p'_def by (auto simp add:dist_norm)
+ then show "\<exists>w\<in>ball p e1. f w \<noteq> 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\<in>poles \<longrightarrow> e2>0 \<and> cball p e2 \<subseteq> ball p e1 \<and> contour p e2" by auto
- def e3\<equiv>"if p\<in>poles then e2 else e1"
+ define e3 where "e3 = (if p\<in>poles then e2 else e1)"
have "avoid p e3"
using e2 e that avoid_def e3_def by auto
moreover have "e3>0" using \<open>e1>0\<close> e2 unfolding e3_def by auto
@@ -2869,33 +2872,33 @@
qed
then obtain h where h:"\<forall>p\<in>s. h p>0 \<and> (avoid p (h p) \<and> (p\<in>poles \<longrightarrow> contour p (h p)))"
by metis
- def cont\<equiv>"\<lambda>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 \<gamma> f = (\<Sum>p\<in>poles. winding_number \<gamma> p * cont p)"
unfolding cont_def
- proof (rule Cauchy_theorem_singularities[OF \<open>open s\<close> \<open>connected (s-poles)\<close> _ holo \<open>valid_path \<gamma>\<close>
- loop \<open>path_image \<gamma> \<subseteq> s-poles\<close> homo])
- show "finite poles" using \<open>poles \<subseteq> pts\<close> and \<open>finite pts\<close> by (simp add: finite_subset)
- next
- show "\<forall>p\<in>s. 0 < h p \<and> (\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> poles))"
- using \<open>poles \<subseteq> pts\<close> h unfolding avoid_def by blast
- qed
+ proof (rule Cauchy_theorem_singularities[OF \<open>open s\<close> \<open>connected (s-poles)\<close> _ holo \<open>valid_path \<gamma>\<close>
+ loop \<open>path_image \<gamma> \<subseteq> s-poles\<close> homo])
+ show "finite poles" using \<open>poles \<subseteq> pts\<close> and \<open>finite pts\<close> by (simp add: finite_subset)
+ next
+ show "\<forall>p\<in>s. 0 < h p \<and> (\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> poles))"
+ using \<open>poles \<subseteq> pts\<close> h unfolding avoid_def by blast
+ qed
also have "... = (\<Sum>p\<in>poles. c * (winding_number \<gamma> p * residue f p))"
- proof (rule setsum.cong[of poles poles,simplified])
- fix p assume "p \<in> poles"
- show "winding_number \<gamma> p * cont p = c * (winding_number \<gamma> p * residue f p)"
- proof (cases "p\<in>s")
- assume "p \<in> s"
- then have "cont p=c*residue f p"
- unfolding cont_def
- apply (intro contour_integral_unique)
- using h[unfolded contour_def] \<open>p \<in> poles\<close> by blast
- then show ?thesis by auto
- next
- assume "p\<notin>s"
- then have "winding_number \<gamma> p=0" using homo by auto
- then show ?thesis by auto
- qed
+ proof (rule setsum.cong[of poles poles,simplified])
+ fix p assume "p \<in> poles"
+ show "winding_number \<gamma> p * cont p = c * (winding_number \<gamma> p * residue f p)"
+ proof (cases "p\<in>s")
+ assume "p \<in> s"
+ then have "cont p=c*residue f p"
+ unfolding cont_def
+ apply (intro contour_integral_unique)
+ using h[unfolded contour_def] \<open>p \<in> poles\<close> by blast
+ then show ?thesis by auto
+ next
+ assume "p\<notin>s"
+ then have "winding_number \<gamma> p=0" using homo by auto
+ then show ?thesis by auto
qed
+ qed
also have "... = c * (\<Sum>p\<in>poles. winding_number \<gamma> p * residue f p)"
apply (subst setsum_right_distrib)
by simp
@@ -2921,11 +2924,13 @@
- (\<Sum>p\<in>poles. winding_number g p * h p * porder f p))"
(is "?L=?R")
proof -
- def c\<equiv>"2 * complex_of_real pi * \<i> "
- def ff\<equiv>"(\<lambda>x. deriv f x * h x / f x)"
- def cont_pole\<equiv>"\<lambda>ff p e. (ff has_contour_integral - c * porder f p * h p) (circlepath p e)"
- def cont_zero\<equiv>"\<lambda>ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)"
- def avoid\<equiv>"\<lambda>p e. \<forall>w\<in>cball p e. w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts)"
+ define c where "c = 2 * complex_of_real pi * \<i> "
+ 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 \<longleftrightarrow>
+ (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 \<longleftrightarrow>
+ (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)" for ff p e
+ define avoid where "avoid p e \<longleftrightarrow> (\<forall>w\<in>cball p e. w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts))" for p e
have "poles \<subseteq> pts" and "zeros \<subseteq> pts" and "finite zeros" and "pts=zeros \<union> poles"
using poles \<open>finite pts\<close> unfolding pts_def zeros_def is_pole_def by auto
have "\<exists>e>0. avoid p e \<and> (p\<in>poles \<longrightarrow> cont_pole ff p e) \<and> (p\<in>zeros \<longrightarrow> cont_zero ff p e)"
@@ -2936,10 +2941,10 @@
have "\<exists>e2>0. cball p e2 \<subseteq> ball p e1 \<and> cont_pole ff p e2"
when "p\<in>poles"
proof -
- def po\<equiv>"porder f p"
- def pp\<equiv>"pol_poly f p"
- def f'\<equiv>"\<lambda>w. pp w / (w - p) ^ po"
- def ff'\<equiv>"(\<lambda>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 \<circ> 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 \<open>p\<in>poles\<close> poles unfolding is_pole_def by auto
moreover have "\<exists>w\<in>ball p e1. f w \<noteq> 0"
proof -
- def p'\<equiv>"p+e1/2"
+ define p' where "p' = p+e1/2"
have "p'\<in>ball p e1" and "p'\<noteq>p" using \<open>e1>0\<close> unfolding p'_def by (auto simp add:dist_norm)
then show "\<exists>w\<in>ball p e1. f w \<noteq> 0" using e unfolding avoid_def
apply (rule_tac x=p' in bexI)
@@ -2975,40 +2980,41 @@
pp_po:"(\<forall>w\<in>cball p r. f w = pp w / (w - p) ^ po \<and> pp w \<noteq> 0)"
using porder_exist[of "ball p e1" p f,simplified,OF \<open>e1>0\<close>] unfolding po_def pp_def
by auto
- def e2\<equiv>"r/2"
+ define e2 where "e2 = r/2"
have "e2>0" using \<open>r>0\<close> unfolding e2_def by auto
- def anal\<equiv>"\<lambda>w. deriv pp w * h w / pp w" and prin\<equiv>"\<lambda>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 "((\<lambda>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 \<subseteq> s"
- using \<open>cball p r \<subseteq> ball p e1\<close> avoid_def ball_subset_cball e by blast
- then have "cball p e2 \<subseteq> s"
- using \<open>r>0\<close> unfolding e2_def by auto
- then have "(\<lambda>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 "\<lambda>w. - of_nat po * h w"]
- \<open>e2>0\<close>
- 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 \<open>ball p r \<subseteq> s\<close>
- by (auto intro!: holomorphic_intros)
- then show "(anal has_contour_integral 0) (circlepath p e2)"
- using e2_def \<open>r>0\<close>
- by (auto elim!: Cauchy_theorem_disc_simple)
- qed
+ proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified])
+ have "ball p r \<subseteq> s"
+ using \<open>cball p r \<subseteq> ball p e1\<close> avoid_def ball_subset_cball e by blast
+ then have "cball p e2 \<subseteq> s"
+ using \<open>r>0\<close> unfolding e2_def by auto
+ then have "(\<lambda>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 "\<lambda>w. - of_nat po * h w"]
+ \<open>e2>0\<close>
+ 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 \<open>ball p r \<subseteq> s\<close>
+ by (auto intro!: holomorphic_intros)
+ then show "(anal has_contour_integral 0) (circlepath p e2)"
+ using e2_def \<open>r>0\<close>
+ 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 \<in> path_image (circlepath p e2)"
then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto
- def wp\<equiv>"w-p"
+ define wp where "wp = w-p"
have "wp\<noteq>0" and "pp w \<noteq>0"
unfolding wp_def using \<open>w\<noteq>p\<close> \<open>w\<in>ball p r\<close> 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 \<equiv> "- 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 \<open>po>0\<close> by auto
have "(pp has_field_derivative (deriv pp w)) (at w)"
using DERIV_deriv_iff_field_differentiable \<open>w\<in>ball p r\<close>
@@ -3067,26 +3073,27 @@
have "\<exists>e3>0. cball p e3 \<subseteq> ball p e1 \<and> cont_zero ff p e3"
when "p\<in>zeros"
proof -
- def zo\<equiv>"zorder f p"
- def zp\<equiv>"zer_poly f p"
- def f'\<equiv>"\<lambda>w. zp w * (w - p) ^ zo"
- def ff'\<equiv>"(\<lambda>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 \<subseteq> s - poles"
- using \<open>poles \<subseteq> pts\<close> avoid_def ball_subset_cball e that zeros_def by fastforce
- thus ?thesis using f_holo by auto
- qed
+ proof -
+ have "ball p e1 \<subseteq> s - poles"
+ using \<open>poles \<subseteq> pts\<close> 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 \<open>p\<in>zeros\<close>
using DiffD1 mem_Collect_eq pts_def zeros_def by blast
moreover have "\<exists>w\<in>ball p e1. f w \<noteq> 0"
- proof -
- def p'\<equiv>"p+e1/2"
- have "p'\<in>ball p e1" and "p'\<noteq>p" using \<open>e1>0\<close> unfolding p'_def by (auto simp add:dist_norm)
- then show "\<exists>w\<in>ball p e1. f w \<noteq> 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'\<in>ball p e1" and "p'\<noteq>p"
+ using \<open>e1>0\<close> unfolding p'_def by (auto simp add:dist_norm)
+ then show "\<exists>w\<in>ball p e1. f w \<noteq> 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 \<subseteq> ball p e1" and
@@ -3094,58 +3101,59 @@
pp_po:"(\<forall>w\<in>cball p r. f w = zp w * (w - p) ^ zo \<and> zp w \<noteq> 0)"
using zorder_exist[of "ball p e1" p f,simplified,OF \<open>e1>0\<close>] unfolding zo_def zp_def
by auto
- def e2\<equiv>"r/2"
+ define e2 where "e2 = r/2"
have "e2>0" using \<open>r>0\<close> unfolding e2_def by auto
- def anal\<equiv>"\<lambda>w. deriv zp w * h w / zp w" and prin\<equiv>"\<lambda>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 "((\<lambda>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 \<subseteq> s"
- using \<open>cball p r \<subseteq> ball p e1\<close> avoid_def ball_subset_cball e by blast
- then have "cball p e2 \<subseteq> s"
- using \<open>r>0\<close> unfolding e2_def by auto
- then have "(\<lambda>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 "\<lambda>w. of_nat zo * h w"]
- \<open>e2>0\<close>
- 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 \<open>ball p r \<subseteq> s\<close>
- by (auto intro!: holomorphic_intros)
- then show "(anal has_contour_integral 0) (circlepath p e2)"
- using e2_def \<open>r>0\<close>
- by (auto elim!: Cauchy_theorem_disc_simple)
- qed
+ proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified])
+ have "ball p r \<subseteq> s"
+ using \<open>cball p r \<subseteq> ball p e1\<close> avoid_def ball_subset_cball e by blast
+ then have "cball p e2 \<subseteq> s"
+ using \<open>r>0\<close> unfolding e2_def by auto
+ then have "(\<lambda>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 "\<lambda>w. of_nat zo * h w"]
+ \<open>e2>0\<close>
+ 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 \<open>ball p r \<subseteq> s\<close>
+ by (auto intro!: holomorphic_intros)
+ then show "(anal has_contour_integral 0) (circlepath p e2)"
+ using e2_def \<open>r>0\<close>
+ 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 \<in> path_image (circlepath p e2)"
- then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto
- def wp\<equiv>"w-p"
- have "wp\<noteq>0" and "zp w \<noteq>0"
- unfolding wp_def using \<open>w\<noteq>p\<close> \<open>w\<in>ball p r\<close> 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\<equiv>"zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo"
- have po:"zo = Suc (zo - Suc 0) " using \<open>zo>0\<close> by auto
- have "(zp has_field_derivative (deriv zp w)) (at w)"
- using DERIV_deriv_iff_field_differentiable \<open>w\<in>ball p r\<close>
- 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 \<open>w\<noteq>p\<close> \<open>zo>0\<close> 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 \<in> path_image (circlepath p e2)"
+ then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto
+ define wp where "wp = w-p"
+ have "wp\<noteq>0" and "zp w \<noteq>0"
+ unfolding wp_def using \<open>w\<noteq>p\<close> \<open>w\<in>ball p r\<close> 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 \<open>zo>0\<close> by auto
+ have "(zp has_field_derivative (deriv zp w)) (at w)"
+ using DERIV_deriv_iff_field_differentiable \<open>w\<in>ball p r\<close>
+ 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 \<open>w\<noteq>p\<close> \<open>zo>0\<close> 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 \<in> path_image (circlepath p e2)"
@@ -3181,7 +3189,7 @@
qed
then obtain e3 where e3:"p\<in>zeros \<longrightarrow> e3>0 \<and> cball p e3 \<subseteq> ball p e1 \<and> cont_zero ff p e3"
by auto
- def e4\<equiv>"if p\<in>poles then e2 else if p\<in>zeros then e3 else e1"
+ define e4 where "e4 = (if p\<in>poles then e2 else if p\<in>zeros then e3 else e1)"
have "e4>0" using e2 e3 \<open>e1>0\<close> unfolding e4_def by auto
moreover have "avoid p e4" using e2 e3 \<open>e1>0\<close> e unfolding e4_def avoid_def by auto
moreover have "p\<in>poles \<longrightarrow> cont_pole ff p e4" and "p\<in>zeros \<longrightarrow> cont_zero ff p e4"
@@ -3191,19 +3199,19 @@
then obtain get_e where get_e:"\<forall>p\<in>s. get_e p>0 \<and> avoid p (get_e p)
\<and> (p\<in>poles \<longrightarrow> cont_pole ff p (get_e p)) \<and> (p\<in>zeros \<longrightarrow> cont_zero ff p (get_e p))"
by metis
- def cont\<equiv>"\<lambda>p. contour_integral (circlepath p (get_e p)) ff"
- def w\<equiv>"\<lambda>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 = (\<Sum>p\<in>pts. w p * cont p)"
unfolding cont_def w_def
- proof (rule Cauchy_theorem_singularities[OF \<open>open s\<close> connected finite _ \<open>valid_path g\<close> loop
- path_img homo])
- have "open (s - pts)" using open_Diff[OF _ finite_imp_closed[OF finite]] \<open>open s\<close> by auto
- then show "ff holomorphic_on s - pts" unfolding ff_def using f_holo \<open>poles \<subseteq> pts\<close> h_holo
- by (auto intro!: holomorphic_intros simp add:pts_def)
- next
- show "\<forall>p\<in>s. 0 < get_e p \<and> (\<forall>w\<in>cball p (get_e p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts))"
- using get_e using avoid_def by blast
- qed
+ proof (rule Cauchy_theorem_singularities[OF \<open>open s\<close> connected finite _ \<open>valid_path g\<close> loop
+ path_img homo])
+ have "open (s - pts)" using open_Diff[OF _ finite_imp_closed[OF finite]] \<open>open s\<close> by auto
+ then show "ff holomorphic_on s - pts" unfolding ff_def using f_holo \<open>poles \<subseteq> pts\<close> h_holo
+ by (auto intro!: holomorphic_intros simp add:pts_def)
+ next
+ show "\<forall>p\<in>s. 0 < get_e p \<and> (\<forall>w\<in>cball p (get_e p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts))"
+ using get_e using avoid_def by blast
+ qed
also have "... = (\<Sum>p\<in>zeros. w p * cont p) + (\<Sum>p\<in>poles. w p * cont p)"
using \<open>finite pts\<close> unfolding \<open>pts=zeros \<union> poles\<close>
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\<equiv>"\<lambda>p. winding_number \<gamma> p"
- def c\<equiv>"2 * complex_of_real pi * \<i>"
- def h\<equiv>"\<lambda>p. g p / f p + 1"
+ define w where "w p = winding_number \<gamma> p" for p
+ define c where "c = 2 * complex_of_real pi * \<i>"
+ define h where [abs_def]: "h p = g p / f p + 1" for p
obtain spikes
where "finite spikes" and spikes: "\<forall>x\<in>{0..1} - spikes. \<gamma> differentiable at x"
using \<open>valid_path \<gamma>\<close>
@@ -3366,8 +3374,8 @@
proof (rule vector_derivative_chain_at_general)
show "\<gamma> differentiable at x" using that \<open>valid_path \<gamma>\<close> spikes by auto
next
- def der\<equiv>"\<lambda>p. (deriv g p * f p - g p * deriv f p)/(f p * f p)"
- def t\<equiv>"\<gamma> 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 = \<gamma> x"
have "f t\<noteq>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\<in>s"
@@ -3412,29 +3420,29 @@
have "fg p\<noteq>0" and "f p\<noteq>0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def
by auto
have "h p\<noteq>0"
- proof (rule ccontr)
- assume "\<not> h p \<noteq> 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 "\<not> h p \<noteq> 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>open s\<close>] 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\<equiv>"\<lambda>p. (deriv g p * f p - g p * deriv f p)/(f p * f p)"
- have "p\<in>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 \<open>f p\<noteq>0\<close>
- apply (auto intro!: derivative_eq_intros)
- by (auto simp add: DERIV_deriv_iff_field_differentiable
- holomorphic_on_imp_differentiable_at[OF _ \<open>open s\<close>])
- 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\<in>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 \<open>f p\<noteq>0\<close>
+ apply (auto intro!: derivative_eq_intros)
+ by (auto simp add: DERIV_deriv_iff_field_differentiable
+ holomorphic_on_imp_differentiable_at[OF _ \<open>open s\<close>])
+ 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 \<open>h p\<noteq>0\<close> \<open>f p\<noteq>0\<close> \<open>fg p\<noteq>0\<close>)
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Mon Apr 25 16:09:26 2016 +0200
@@ -105,7 +105,7 @@
proof -
{
fix x :: "'n::euclidean_space"
- def y \<equiv> "(e / norm x) *\<^sub>R x"
+ define y where "y = (e / norm x) *\<^sub>R x"
then have "y \<in> 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: "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
"finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = (1::real)"
- def n \<equiv> "card s"
+ define n where "n = card s"
have "card s = 0 \<or> card s = 1 \<or> card s = 2 \<or> card s > 2" by auto
then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
proof (auto simp only: disjE)
@@ -838,7 +838,7 @@
assume "y = a + v" "v \<in> span {x - a |x. x \<in> s}"
then obtain t u where obt: "finite t" "t \<subseteq> {x - a |x. x \<in> s}" "a + (\<Sum>v\<in>t. u v *\<^sub>R v) = y"
unfolding span_explicit by auto
- def f \<equiv> "(\<lambda>x. x + a) ` t"
+ define f where "f = (\<lambda>x. x + a) ` t"
have f: "finite f" "f \<subseteq> s" "(\<Sum>v\<in>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 \<inter> {a} = {}" "f \<inter> - {a} = f"
@@ -989,7 +989,7 @@
{
fix x y
assume xy: "x \<in> S" "y \<in> 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 \<inter> s \<noteq> {}" "B \<inter> s \<noteq> {}"
then obtain a b where a: "a \<in> A" "a \<in> s" and b: "b \<in> B" "b \<in> s" by auto
- def f \<equiv> "\<lambda>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 \<equiv> "i - k1"
+ define j where "j = i - k1"
from i False have "j \<in> {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 \<subseteq> {x - a |x. x \<in> s}" "v\<in>S" "u v \<noteq> 0" "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0"
by auto
- def t \<equiv> "(\<lambda>x. x + a) ` S"
+ define t where "t = (\<lambda>x. x + a) ` S"
have inj: "inj_on (\<lambda>x. x + a) S"
unfolding inj_on_def by auto
@@ -2469,7 +2469,7 @@
"(\<lambda>x. -a+x) ` (S - {a}) \<subseteq> B \<and> B \<subseteq> (\<lambda>x. -a+x) ` V \<and> independent B \<and> (\<lambda>x. -a+x) ` V \<subseteq> span B"
using maximal_independent_subset_extend[of "(\<lambda>x. -a+x) ` (S-{a})" "(\<lambda>x. -a + x) ` V"] assms
by blast
- def T \<equiv> "(\<lambda>x. a+x) ` insert 0 B"
+ define T where "T = (\<lambda>x. a+x) ` insert 0 B"
then have "T = insert a ((\<lambda>x. a+x) ` B)"
by auto
then have "affine hull T = (\<lambda>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 \<in> B" by auto
- def Lb \<equiv> "span ((\<lambda>x. -a+x) ` (B-{a}))"
+ define Lb where "Lb = span ((\<lambda>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 \<in> B" by auto
- def Lb \<equiv> "span ((\<lambda>x. -a+x) ` (B-{a}))"
+ define Lb where "Lb = span ((\<lambda>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 \<in> S" using assms by auto
then have "a \<in> T" using assms by auto
- def LS \<equiv> "{y. \<exists>x \<in> S. (-a) + x = y}"
+ define LS where "LS = {y. \<exists>x \<in> S. (-a) + x = y}"
then have ls: "subspace LS" "affine_parallel S LS"
using assms parallel_subspace_explicit[of S a LS] \<open>a \<in> S\<close> by auto
then have h1: "int(dim LS) = aff_dim S"
using assms aff_dim_affine[of S LS] by auto
have "T \<noteq> {}" using assms by auto
- def LT \<equiv> "{y. \<exists>x \<in> T. (-a) + x = y}"
+ define LT where "LT = {y. \<exists>x \<in> T. (-a) + x = y}"
then have lt: "subspace LT \<and> affine_parallel T LT"
using assms parallel_subspace_explicit[of T a LT] \<open>a \<in> T\<close> by auto
then have "int(dim LT) = aff_dim T"
@@ -3062,7 +3062,7 @@
where a1 [simp]: "setsum a t = 1" and [simp]: "setsum (\<lambda>v. a v *\<^sub>R v) t = y"
and [simp]: "setsum b u = 1" "setsum (\<lambda>v. b v *\<^sub>R v) u = y"
by (auto simp: affine_hull_finite \<open>finite t\<close> \<open>finite u\<close>)
- def c \<equiv> "\<lambda>x. if x \<in> t then a x else if x \<in> u then -(b x) else 0"
+ define c where "c x = (if x \<in> t then a x else if x \<in> u then -(b x) else 0)" for x
have [simp]: "s \<inter> t = t" "s \<inter> - t \<inter> u = u" using assms by auto
have "setsum c s = 0"
by (simp add: c_def comm_monoid_add_class.setsum.If_cases \<open>finite s\<close> setsum_negf)
@@ -3172,8 +3172,8 @@
by blast
then obtain w v where wv: "setsum w s = 0" "v\<in>s" "w v \<noteq> 0" "(\<Sum>v\<in>s. w v *\<^sub>R v) = 0"
using affine_dependent_explicit_finite[OF obt(1)] by auto
- def i \<equiv> "(\<lambda>v. (u v) / (- w v)) ` {v\<in>s. w v < 0}"
- def t \<equiv> "Min i"
+ define i where "i = (\<lambda>v. (u v) / (- w v)) ` {v\<in>s. w v < 0}"
+ define t where "t = Min i"
have "\<exists>x\<in>s. w x < 0"
proof (rule ccontr, simp add: not_less)
assume as:"\<forall>x\<in>s. 0 \<le> w x"
@@ -3428,7 +3428,7 @@
{
fix x
assume x: "x \<in> affine hull S"
- def e \<equiv> "1::real"
+ define e :: real where "e = 1"
then have "e > 0" "ball x e \<inter> affine hull (affine hull S) \<subseteq> affine hull S"
using hull_hull[of _ S] by auto
then have "x \<in> rel_interior (affine hull S)"
@@ -3727,7 +3727,7 @@
qed
then obtain y where "y \<in> S" and y: "norm (y - x) * (1 - e) < e * d"
by auto
- def z \<equiv> "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 \<open>e > 0\<close>
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" "\<And>x. norm (f x) \<le> norm x * K"
using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto
- def e1 \<equiv> "1 / K"
+ define e1 where "e1 = 1 / K"
then have e1: "e1 > 0" "\<And>x. e1 * norm (f x) \<le> norm x"
using K pos_le_divide_eq[of e1] by auto
- def e \<equiv> "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 \<equiv> "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 "\<lambda>x e. e > 0 \<and> cball x e \<subseteq> s"] by auto
have "b ` t \<noteq> {}"
using obt by auto
- def i \<equiv> "b ` t"
+ define i where "i = b ` t"
show "\<exists>e > 0.
cball a e \<subseteq> {y. \<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>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 \<equiv> "SUP x:t. a \<bullet> x"
+ define k where "k = (SUP x:t. a \<bullet> 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 \<in> c. 0 < u x}" "finite {x \<in> c. 0 > u x}"
using assms(1) by auto
- def z \<equiv> "inverse (setsum u {x\<in>c. u x > 0}) *\<^sub>R setsum (\<lambda>x. u x *\<^sub>R x) {x\<in>c. u x > 0}"
+ define z where "z = inverse (setsum u {x\<in>c. u x > 0}) *\<^sub>R setsum (\<lambda>x. u x *\<^sub>R x) {x\<in>c. u x > 0}"
have "setsum u {x \<in> c. 0 < u x} \<noteq> 0"
proof (cases "u v \<ge> 0")
case False
@@ -5356,7 +5356,7 @@
using compact_imp_closed[OF assms(1)]
apply simp
done
- def pi \<equiv> "\<lambda>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 \<notin> frontier s"
unfolding frontier_straddle
apply (rule notI)
@@ -5373,7 +5373,7 @@
apply (intro continuous_intros)
apply simp
done
- def sphere \<equiv> "{x::'a. norm x = 1}"
+ define sphere :: "'a set" where "sphere = {x. norm x = 1}"
have pi: "\<And>x. x \<noteq> 0 \<Longrightarrow> pi x \<in> sphere" "\<And>x u. u>0 \<Longrightarrow> 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 \<in> s" "e > 0"
- def B \<equiv> "\<bar>b\<bar> + 1"
+ define B where "B = \<bar>b\<bar> + 1"
have B: "0 < B" "\<And>x. x\<in>s \<Longrightarrow> \<bar>f x\<bar> \<le> B"
unfolding B_def
defer
@@ -6312,7 +6312,7 @@
show "\<bar>f y - f x\<bar> < e"
proof (cases "y = x")
case False
- def t \<equiv> "k / norm (y - x)"
+ define t where "t = k / norm (y - x)"
have "2 < t" "0<t"
unfolding t_def using as False and \<open>k>0\<close>
by (auto simp add:field_simps)
@@ -6323,7 +6323,7 @@
using as
by (auto simp add: field_simps norm_minus_commute)
{
- def w \<equiv> "x + t *\<^sub>R (y - x)"
+ define w where "w = x + t *\<^sub>R (y - x)"
have "w \<in> s"
unfolding w_def
apply (rule k[unfolded subset_eq,rule_format])
@@ -6356,7 +6356,7 @@
}
moreover
{
- def w \<equiv> "x - t *\<^sub>R (y - x)"
+ define w where "w = x - t *\<^sub>R (y - x)"
have "w \<in> s"
unfolding w_def
apply (rule k[unfolded subset_eq,rule_format])
@@ -6408,7 +6408,7 @@
case True
fix y
assume y: "y \<in> cball x e"
- def z \<equiv> "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 \<in> cball x e"
@@ -6445,7 +6445,7 @@
assume "x \<in> s"
then obtain e where e: "cball x e \<subseteq> s" "e > 0"
using assms(1) unfolding open_contains_cball by auto
- def d \<equiv> "e / real DIM('a)"
+ define d where "d = e / real DIM('a)"
have "0 < d"
unfolding d_def using \<open>e > 0\<close> dimge1 by auto
let ?d = "(\<Sum>i\<in>Basis. d *\<^sub>R i)::'a"
@@ -6454,7 +6454,7 @@
and c1: "convex hull c \<subseteq> cball x e"
and c2: "cball x d \<subseteq> convex hull c"
proof
- def c \<equiv> "\<Sum>i\<in>Basis. (\<lambda>a. a *\<^sub>R i) ` {x\<bullet>i - d, x\<bullet>i + d}"
+ define c where "c = (\<Sum>i\<in>Basis. (\<lambda>a. a *\<^sub>R i) ` {x\<bullet>i - d, x\<bullet>i + d})"
show "finite c"
unfolding c_def by (simp add: finite_set_setsum)
have 1: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cbox (x \<bullet> i - d) (x \<bullet> i + d)}"
@@ -6490,7 +6490,7 @@
apply simp
done
qed
- def k \<equiv> "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 \<in> s" and y: "norm (y - x) * (1 - e) < e * d"
by auto
- def z \<equiv> "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 \<open>e > 0\<close>
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 \<equiv> "min 1 (e/norm (x - a))"
+ define e1 where "e1 = min 1 (e/norm (x - a))"
then have e1: "e1 > 0" "e1 \<le> 1" "e1 * norm (x - a) \<le> e"
using \<open>x \<noteq> a\<close> \<open>e > 0\<close> 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 \<le> 1" "z = y - e *\<^sub>R (y - x)"
proof -
- def e \<equiv> "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 \<inter> affine hull closure S \<le> 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 \<equiv> "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 \<in> cball z e"
using mem_cball y_def dist_norm[of z y] e by auto
have "x \<in> affine hull closure S"
@@ -8239,7 +8239,7 @@
assume x: "x \<in> affine hull S"
{
assume "x \<noteq> z"
- def m \<equiv> "1 + e1/norm(x-z)"
+ define m where "m = 1 + e1/norm(x-z)"
hence "m > 1" using e1 \<open>x \<noteq> z\<close> by auto
{
fix e
@@ -8274,7 +8274,7 @@
moreover
{
assume "x = z"
- def m \<equiv> "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 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
using assms by auto
- def y \<equiv> "(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 \<in> S" using e by auto
- def e1 \<equiv> "1/e"
+ define e1 where "e1 = 1/e"
then have "0 < e1 \<and> 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 \<and> z + e2 *\<^sub>R (z - x) \<in> S"
using r by auto
- def x1 \<equiv> "z + e1 *\<^sub>R (x - z)"
+ define x1 where [abs_def]: "x1 = z + e1 *\<^sub>R (x - z)"
then have x1: "x1 \<in> affine hull S"
using e1 hull_subset[of S] by auto
- def x2 \<equiv> "z + e2 *\<^sub>R (z - x)"
+ define x2 where [abs_def]: "x2 = z + e2 *\<^sub>R (z - x)"
then have x2: "x2 \<in> 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 \<in> S"
using **[rule_format, of "z-x"] by auto
- def e \<equiv> "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 \<in> S"
@@ -8427,7 +8427,7 @@
fix x
obtain e1 where e1: "e1 > 0" "z + e1 *\<^sub>R (z - x) \<in> S"
using r[rule_format, of "z-x"] by auto
- def e \<equiv> "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 \<in> S"
@@ -8508,11 +8508,11 @@
assume "y \<noteq> x"
{ fix e :: real
assume e: "e > 0"
- def e1 \<equiv> "min 1 (e/norm (y - x))"
+ define e1 where "e1 = min 1 (e/norm (y - x))"
then have e1: "e1 > 0" "e1 \<le> 1" "e1 * norm (y - x) \<le> e"
using \<open>y \<noteq> x\<close> \<open>e > 0\<close> le_divide_eq[of e1 e "norm (y - x)"]
by simp_all
- def z \<equiv> "y - e1 *\<^sub>R (y - x)"
+ define z where "z = y - e1 *\<^sub>R (y - x)"
{
fix S
assume "S \<in> I"
@@ -8624,7 +8624,7 @@
}
then obtain mS where
mS: "\<forall>S\<in>I. mS S > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> mS S \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)" by metis
- def e \<equiv> "Min (mS ` I)"
+ define e where "e = Min (mS ` I)"
then have "e \<in> mS ` I" using assms \<open>I \<noteq> {}\<close> by simp
then have "e > 1" using mS by auto
moreover have "\<forall>S\<in>I. e \<le> mS S"
@@ -9112,7 +9112,7 @@
have conv: "convex ({(1 :: real)} \<times> S)"
using convex_Times[of "{(1 :: real)}" S] assms convex_singleton[of "1 :: real"]
by auto
- def f \<equiv> "\<lambda>y. {z. (y, z) \<in> cone hull ({1 :: real} \<times> S)}"
+ define f where "f y = {z. (y, z) \<in> cone hull ({1 :: real} \<times> S)}" for y
then have *: "(c, x) \<in> rel_interior (cone hull ({(1 :: real)} \<times> S)) =
(c \<in> rel_interior {y. f y \<noteq> {}} \<and> x \<in> rel_interior (f c))"
apply (subst rel_interior_projection[of "cone hull ({(1 :: real)} \<times> S)" f c x])
@@ -9210,11 +9210,11 @@
{
fix x
assume "x \<in> S i"
- def c \<equiv> "\<lambda>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 \<open>finite I\<close> \<open>i \<in> I\<close> setsum.delta[of I i "\<lambda>j::'a. 1::real"]
by auto
- def s \<equiv> "\<lambda>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 "\<forall>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 (\<lambda>i. c i *\<^sub>R s i) I"
@@ -9243,7 +9243,7 @@
from xy obtain d t where
yc: "y = setsum (\<lambda>i. d i *\<^sub>R t i) I \<and> (\<forall>i\<in>I. d i \<ge> 0) \<and> setsum d I = 1 \<and> (\<forall>i\<in>I. t i \<in> S i)"
by auto
- def e \<equiv> "\<lambda>i. u * c i + v * d i"
+ define e where "e i = u * c i + v * d i" for i
have ge0: "\<forall>i\<in>I. e i \<ge> 0"
using e_def xc yc uv by simp
have "setsum (\<lambda>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 \<equiv> "\<lambda>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 \<in> I"
@@ -9318,8 +9319,8 @@
{u *\<^sub>R s + v *\<^sub>R t | u v s t. u \<ge> 0 \<and> v \<ge> 0 \<and> u + v = 1 \<and> s \<in> S \<and> t \<in> T}"
(is "?lhs = ?rhs")
proof
- def I \<equiv> "{1::nat, 2}"
- def s \<equiv> "\<lambda>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 "\<Union>(s ` I) = S \<union> T"
using s_def I_def by auto
then have "convex hull (\<Union>(s ` I)) = convex hull (S \<union> T)"
@@ -9414,7 +9415,7 @@
then obtain c xs where
x: "x = setsum (\<lambda>i. c i *\<^sub>R xs i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. xs i \<in> S i)"
using convex_hull_finite_union[of I S] assms by auto
- def s \<equiv> "\<lambda>i. c i *\<^sub>R xs i"
+ define s where "s i = c i *\<^sub>R xs i" for i
{
fix i
assume "i \<in> I"
@@ -9432,7 +9433,7 @@
assume "x \<in> ?rhs"
then obtain s where x: "x = setsum s I \<and> (\<forall>i\<in>I. s i \<in> S i)"
using set_setsum_alt[of I S] assms by auto
- def xs \<equiv> "\<lambda>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 (\<lambda>i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I"
using x assms by auto
moreover have "\<forall>i\<in>I. xs i \<in> S i"
@@ -9465,8 +9466,8 @@
and "T \<noteq> {}"
shows "convex hull (S \<union> T) = S + T"
proof -
- def I \<equiv> "{1::nat, 2}"
- def A \<equiv> "(\<lambda>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 "\<Union>(A ` I) = S \<union> T"
using A_def I_def by auto
then have "convex hull (\<Union>(A ` I)) = convex hull (S \<union> T)"
@@ -9500,11 +9501,11 @@
using convex_hull_empty rel_interior_empty by auto
next
case False
- def C0 \<equiv> "convex hull (\<Union>(S ` I))"
+ define C0 where "C0 = convex hull (\<Union>(S ` I))"
have "\<forall>i\<in>I. C0 \<ge> S i"
unfolding C0_def using hull_subset[of "\<Union>(S ` I)"] by auto
- def K0 \<equiv> "cone hull ({1 :: real} \<times> C0)"
- def K \<equiv> "\<lambda>i. cone hull ({1 :: real} \<times> S i)"
+ define K0 where "K0 = cone hull ({1 :: real} \<times> C0)"
+ define K where "K i = cone hull ({1 :: real} \<times> S i)" for i
have "\<forall>i\<in>I. K i \<noteq> {}"
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 (\<lambda>i. c i *\<^sub>R s i) I \<and>
(\<forall>i\<in>I. c i > 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> rel_interior (S i))"
by auto
- def k \<equiv> "\<lambda>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 \<in> rel_interior (K i)"
@@ -9653,7 +9654,7 @@
have "open (interior I)" by auto
from openE[OF this \<open>x \<in> interior I\<close>]
obtain e where e: "0 < e" "ball x e \<subseteq> interior I" .
- moreover def t \<equiv> "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 \<in> ball x e"
by (auto simp: dist_real_def field_simps split: split_min)
with \<open>x \<in> interior I\<close> e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
@@ -9661,7 +9662,7 @@
have "open (interior I)" by auto
from openE[OF this \<open>x \<in> interior I\<close>]
obtain e where "0 < e" "ball x e \<subseteq> interior I" .
- moreover def K \<equiv> "x - e / 2"
+ moreover define K where "K = x - e / 2"
with \<open>0 < e\<close> have "K \<in> ball x e" "K < x"
by (auto simp: dist_real_def)
ultimately have "K \<in> I" "K < x" "x \<in> I"
@@ -9688,7 +9689,7 @@
have "open (interior I)" by auto
from openE[OF this \<open>x \<in> interior I\<close>]
obtain e where e: "0 < e" "ball x e \<subseteq> interior I" .
- moreover def t \<equiv> "x + e / 2"
+ moreover define t where "t = x + e / 2"
ultimately have "x < t" "t \<in> ball x e"
by (auto simp: dist_real_def field_simps)
with \<open>x \<in> interior I\<close> e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
@@ -10767,7 +10768,7 @@
by blast
next
case False
- def k \<equiv> "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 \<open>e > 0\<close> 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 \<noteq> {}" "0 \<notin> S"
obtains a where "a \<in> span S" "a \<noteq> 0" "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> a \<bullet> x"
proof -
- def k \<equiv> "\<lambda>c::'a. {x. 0 \<le> c \<bullet> x}"
+ define k where [abs_def]: "k c = {x. 0 \<le> c \<bullet> x}" for c :: 'a
have *: "span S \<inter> frontier (cball 0 1) \<inter> \<Inter>f' \<noteq> {}"
if f': "finite f'" "f' \<subseteq> k ` S" for f'
proof -
@@ -11346,7 +11347,7 @@
proof -
obtain e where "0 < e" "y \<in> S" and e: "cball y e \<inter> affine hull S \<subseteq> S"
using \<open>y \<in> rel_interior S\<close> by (force simp: rel_interior_cball)
- def y' \<equiv> "y - (e / norm a) *\<^sub>R ((x + a) - x)"
+ define y' where "y' = y - (e / norm a) *\<^sub>R ((x + a) - x)"
have "y' \<in> cball y e"
unfolding y'_def using \<open>0 < e\<close> by force
moreover have "y' \<in> affine hull S"
@@ -11392,7 +11393,7 @@
"setsum v t = 1"
and eq: "(\<Sum>x\<in>t. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)" for u v
proof -
- def f \<equiv> "\<lambda>x. (if x \<in> s then u x else 0) - (if x \<in> t then v x else 0)"
+ define f where "f x = (if x \<in> s then u x else 0) - (if x \<in> t then v x else 0)" for x
have "setsum f (s \<union> t) = 0"
apply (simp add: f_def setsum_Un setsum_subtractf)
apply (simp add: setsum.inter_restrict [symmetric] Int_commute)
--- a/src/HOL/Multivariate_Analysis/Derivative.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy Mon Apr 25 16:09:26 2016 +0200
@@ -379,7 +379,7 @@
proof (rule, rule ccontr)
fix i :: 'a
assume i: "i \<in> Basis"
- def e \<equiv> "norm (f' i - f'' i)"
+ define e where "e = norm (f' i - f'' i)"
assume "f' i \<noteq> 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 \<equiv> "(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 \<equiv> "e / 2" with \<open>e > 0\<close> have "e2 > 0" by simp
+ define e2 where "e2 = e / 2"
+ with \<open>e > 0\<close> have "e2 > 0" by simp
let ?le = "\<lambda>x1. norm (f x1 - f a) \<le> \<phi> x1 - \<phi> a + e * (x1 - a) + e"
- def A \<equiv> "{x2. a \<le> x2 \<and> x2 \<le> b \<and> (\<forall>x1\<in>{a ..< x2}. ?le x1)}"
+ define A where "A = {x2. a \<le> x2 \<and> x2 \<le> b \<and> (\<forall>x1\<in>{a ..< x2}. ?le x1)}"
have A_subset: "A \<subseteq> {a .. b}" by (auto simp: A_def)
{
fix x2
@@ -817,7 +818,7 @@
have A_ivl: "\<And>x1 x2. x2 \<in> A \<Longrightarrow> x1 \<in> {a ..x2} \<Longrightarrow> x1 \<in> A"
by (simp add: A_def)
have [simp]: "bdd_above A" by (auto simp: A_def)
- def y \<equiv> "Sup A"
+ define y where "y = Sup A"
have "y \<le> b"
unfolding y_def
by (simp add: cSup_le_iff) (simp add: A_def)
@@ -878,7 +879,7 @@
by metis
from \<open>open S\<close> obtain d where d: "\<And>x. dist x y < d \<Longrightarrow> x \<in> S" "d > 0"
by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \<open>y \<in> S\<close>])
- def d' \<equiv> "min ((y + b)/2) (y + (d/2))"
+ define d' where "d' = min ((y + b)/2) (y + (d/2))"
have "d' \<in> A"
unfolding A_def
proof safe
@@ -942,7 +943,7 @@
by metis
from \<open>open S\<close> obtain d where d: "\<And>x. dist x y < d \<Longrightarrow> x \<in> S" "d > 0"
by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \<open>y \<in> S\<close>])
- def d' \<equiv> "min b (y + (d/2))"
+ define d' where "d' = min b (y + (d/2))"
have "d' \<in> A"
unfolding A_def
proof safe
@@ -1077,7 +1078,7 @@
assumes "x0 \<in> S"
shows "norm (f b - f a - f' x0 (b - a)) \<le> norm (b - a) * B"
proof -
- def g \<equiv> "\<lambda>x. f x - f' x0 x"
+ define g where [abs_def]: "g x = f x - f' x0 x" for x
have g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative (\<lambda>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"
"\<forall>z. norm (z - y) < d \<longrightarrow> norm (g z - g y - g' (z - y)) \<le> 1 / 2 * norm (g z - g y)"
using lem1 * by blast
- def B \<equiv> "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) \<le> 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 \<equiv> "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 \<in> ball a d" "y \<in> ball a d" "f x = f y"
- def ph \<equiv> "\<lambda>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' \<circ> (\<lambda>w. f' a w - (f w - f x))"
unfolding ph_def o_def
unfolding diff
@@ -2207,7 +2208,7 @@
shows "summable (\<lambda>n. f n x)" "((\<lambda>x. \<Sum>n. f n x) has_field_derivative (\<Sum>n. f' n x)) (at x)"
proof -
from \<open>x \<in> interior s\<close> have "x \<in> s" using interior_subset by blast
- def g' \<equiv> "\<lambda>x. \<Sum>i. f' i x"
+ define g' where [abs_def]: "g' x = (\<Sum>i. f' i x)" for x
from assms(3) have "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially"
by (simp add: uniformly_convergent_uniform_limit_iff suminf_eq_lim g'_def)
from has_field_derivative_series[OF assms(1,2) this assms(4,5)] obtain g where g:
@@ -2702,7 +2703,7 @@
shows "((\<lambda>(x, y). f x y) has_derivative (\<lambda>(tx, ty). fx x y tx + fy x y ty)) (at (x, y) within X \<times> Y)"
proof (safe intro!: has_derivativeI tendstoI, goal_cases)
case (2 e')
- def e\<equiv>"e' / 9"
+ define e where "e = e' / 9"
have "e > 0" using \<open>e' > 0\<close> by (simp add: e_def)
have "(x, y) \<in> X \<times> Y" using assms by auto
@@ -2720,7 +2721,7 @@
for x' y'
using \<open>0 < e\<close>
by (cases "(x', y') = (x, y)") auto
- def d \<equiv> "d' / sqrt 2"
+ define d where "d = d' / sqrt 2"
have "d > 0" using \<open>0 < d'\<close> by (simp add: d_def)
have d: "x' \<in> X \<Longrightarrow> y' \<in> Y \<Longrightarrow> dist x' x < d \<Longrightarrow> dist y' y < d \<Longrightarrow> dist (fy x' y') (fy x y) < e"
for x' y'
--- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Mon Apr 25 16:09:26 2016 +0200
@@ -119,7 +119,7 @@
assume "open {x}"
then obtain e where "0 < e" and e: "\<forall>y. dist y x < e \<longrightarrow> y = x"
unfolding open_dist by fast
- def y \<equiv> "x + scaleR (e/2) (SOME b. b \<in> Basis)"
+ define y where "y = x + scaleR (e/2) (SOME b. b \<in> Basis)"
have [simp]: "(SOME b. b \<in> Basis) \<in> Basis"
by (rule someI_ex) (auto simp: ex_in_conv)
from \<open>0 < e\<close> have "y \<noteq> x"
--- a/src/HOL/Multivariate_Analysis/Fashoda.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Fashoda.thy Mon Apr 25 16:09:26 2016 +0200
@@ -102,8 +102,10 @@
proof (rule ccontr)
assume "\<not> ?thesis"
note as = this[unfolded bex_simps,rule_format]
- def sqprojection \<equiv> "\<lambda>z::real^2. (inverse (infnorm z)) *\<^sub>R z"
- def negatex \<equiv> "\<lambda>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 \<Rightarrow> real^2"
+ where "negatex x = (vector [-(x$1), x$2])" for x
have lem1: "\<forall>z::real^2. infnorm (negatex z) = infnorm z"
unfolding negatex_def infnorm_2 vector_2 by auto
have lem2: "\<forall>z. z \<noteq> 0 \<longrightarrow> infnorm (sqprojection z) = 1"
@@ -349,7 +351,7 @@
obtains z where "z \<in> path_image f" and "z \<in> path_image g"
proof -
note assms=assms[unfolded path_def pathstart_def pathfinish_def path_image_def]
- def iscale \<equiv> "\<lambda>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} \<subseteq> {0..1}"
unfolding iscale_def by auto
have "\<exists>s\<in>{- 1..1}. \<exists>t\<in>{- 1..1}. (f \<circ> iscale) s = (g \<circ> iscale) t"
--- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Mon Apr 25 16:09:26 2016 +0200
@@ -328,12 +328,12 @@
fix x assume "x \<in> S"
then obtain e where "0 < e" and S: "\<forall>y. dist y x < e \<longrightarrow> y \<in> S"
using * by fast
- def r \<equiv> "\<lambda>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 \<open>0 < e\<close> have r: "\<forall>i. 0 < r i"
unfolding r_def by simp_all
from \<open>0 < e\<close> have e: "e = setL2 r UNIV"
unfolding r_def by (simp add: setL2_constant)
- def A \<equiv> "\<lambda>i. {y. dist (x $ i) y < r i}"
+ define A where "A i = {y. dist (x $ i) y < r i}" for i
have "\<forall>i. open (A i) \<and> x $ i \<in> A i"
unfolding A_def by (simp add: open_ball r)
moreover have "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> 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 \<equiv> "\<lambda>i. LEAST N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s"
- def M \<equiv> "Max (range N)"
+ define N where "N i = (LEAST N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s)" for i
+ define M where "M = Max (range N)"
have "\<And>i. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s"
using X \<open>0 < ?s\<close> by (rule metric_CauchyD)
hence "\<And>i. \<forall>m\<ge>N i. \<forall>n\<ge>N i. dist (X m $ i) (X n $ i) < ?s"
--- a/src/HOL/Multivariate_Analysis/Gamma.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Gamma.thy Mon Apr 25 16:09:26 2016 +0200
@@ -344,8 +344,8 @@
shows "uniformly_convergent_on (ball z d) (\<lambda>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'' \<equiv> "SUP t:ball z d. norm t + norm t^2"
- def e' \<equiv> "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 ((\<lambda>t. norm t + norm t^2) ` cball z d)"
by (intro compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros)
hence "bounded ((\<lambda>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 \<equiv> "max 2 (max (nat \<lceil>2 * (norm z + d)\<rceil>) M)"
+ define N where "N = max 2 (max (nat \<lceil>2 * (norm z + d)\<rceil>) M)"
{
from d have "\<lceil>2 * (cmod z + d)\<rceil> \<ge> \<lceil>0::real\<rceil>"
by (intro ceiling_mono mult_nonneg_nonneg add_nonneg_nonneg) simp_all
@@ -434,8 +434,8 @@
assumes z: "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"
shows "\<exists>d>0. uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n)"
proof -
- def d' \<equiv> "Re z"
- def d \<equiv> "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') \<in> \<int>\<^sub>\<le>\<^sub>0" if "d' \<le> 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 \<noteq> 0" and n: "n \<ge> 2"
shows "uniformly_convergent_on (ball z d) (\<lambda>k z. \<Sum>i<k. inverse ((z + of_nat i)^n))"
proof (rule weierstrass_m_test'_ev)
- def e \<equiv> "(1 + d / norm z)"
- def m \<equiv> "nat \<lceil>norm z * e\<rceil>"
+ define e where "e = (1 + d / norm z)"
+ define m where "m = nat \<lceil>norm z * e\<rceil>"
{
fix t assume t: "t \<in> ball z d"
have "norm t = norm (z + (t - z))" by simp
@@ -683,7 +683,7 @@
by blast+
let ?f' = "\<lambda>z k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))"
let ?f = "\<lambda>z k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))" and ?F' = "\<lambda>z. \<Sum>n. ?f' z n"
- def d \<equiv> "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 \<ge> d" by (auto simp add: complex_nonpos_Reals_iff d_def)
have ball: "Im t = 0 \<longrightarrow> 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 \<noteq> 0"
from z have z': "z \<noteq> 0" by auto
from no_nonpos_Int_in_ball'[OF z] guess d . note d = this
- def n' \<equiv> "Suc n"
+ define n' where "n' = Suc n"
from n have n': "n' \<ge> 2" by (simp add: n'_def)
have "((\<lambda>z. \<Sum>k. inverse ((z + of_nat k) ^ n')) has_field_derivative
(\<Sum>k. - of_nat n' * inverse ((z + of_nat k) ^ (n'+1)))) (at z)"
@@ -1763,10 +1763,10 @@
defines "a \<equiv> complex_of_real pi"
obtains h' where "continuous_on UNIV h'" "\<And>z. (h has_field_derivative (h' z)) (at z)"
proof -
- def f \<equiv> "\<lambda>n. a * of_real (cos_coeff (n+1) - sin_coeff (n+2))"
- def F \<equiv> "\<lambda>z. if z = 0 then 0 else (cos (a*z) - sin (a*z)/(a*z)) / z"
- def g \<equiv> "\<lambda>n. complex_of_real (sin_coeff (n+1))"
- def G \<equiv> "\<lambda>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 \<noteq> 0" unfolding a_def by simp
have "(\<lambda>n. f n * (a*z)^n) sums (F z) \<and> (\<lambda>n. g n * (a*z)^n) sums (G z)"
@@ -1803,13 +1803,13 @@
qed
note sums = conjunct1[OF this] conjunct2[OF this]
- def h2 \<equiv> "\<lambda>z. (\<Sum>n. f n * (a*z)^n) / (\<Sum>n. g n * (a*z)^n) +
- Digamma (1 + z) - Digamma (1 - z)"
- def POWSER \<equiv> "\<lambda>f z. (\<Sum>n. f n * (z^n :: complex))"
- def POWSER' \<equiv> "\<lambda>f z. (\<Sum>n. diffs f n * (z^n :: complex))"
-
- def h2' \<equiv> "\<lambda>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 = (\<Sum>n. f n * (a*z)^n) / (\<Sum>n. g n * (a*z)^n) + Digamma (1 + z) - Digamma (1 - z)" for z
+ define POWSER where [abs_def]: "POWSER f z = (\<Sum>n. f n * (z^n :: complex))" for f z
+ define POWSER' where [abs_def]: "POWSER' f z = (\<Sum>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 \<equiv> "\<i> * of_real (norm z + 1)"
+ define d where "d = \<i> * 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 (\<lambda>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'' \<equiv> "\<lambda>z. h2' (z - of_int \<lfloor>Re z\<rfloor>)"
+ define h2'' where "h2'' z = h2' (z - of_int \<lfloor>Re z\<rfloor>)" 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 \<equiv> "\<lfloor>Re z\<rfloor>"
- def A \<equiv> "{t. of_int r - 1 < Re t \<and> Re t < of_int r + 1}"
+ define r where "r = \<lfloor>Re z\<rfloor>"
+ define A where "A = {t. of_int r - 1 < Re t \<and> Re t < of_int r + 1}"
have "continuous_on A (\<lambda>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 = "\<lambda>z::complex. Gamma z * Gamma (1 - z) * sin (of_real pi * z)"
- def g \<equiv> "\<lambda>z::complex. if z \<in> \<int> then of_real pi else ?g z"
+ define g where [abs_def]: "g z = (if z \<in> \<int> then of_real pi else ?g z)" for z :: complex
let ?h = "\<lambda>z::complex. (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))"
- def h \<equiv> "\<lambda>z::complex. if z \<in> \<int> then 0 else ?h z"
+ define h where [abs_def]: "h z = (if z \<in> \<int> then 0 else ?h z)" for z :: complex
\<comment> \<open>@{term g} is periodic with period 1.\<close>
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 \<equiv> "\<lfloor>Re z / 2\<rfloor>"
+ define r where "r = \<lfloor>Re z / 2\<rfloor>"
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 \<equiv> "max 1 \<bar>Re z\<bar>"
- def B \<equiv> "{t. abs (Re t) \<le> m \<and> abs (Im t) \<le> abs (Im z)}"
+ define m where "m = max 1 \<bar>Re z\<bar>"
+ define B where "B = {t. abs (Re t) \<le> m \<and> abs (Im t) \<le> abs (Im z)}"
have "closed ({t. Re t \<ge> -m} \<inter> {t. Re t \<le> m} \<inter>
{t. Im t \<ge> -\<bar>Im z\<bar>} \<inter> {t. Im t \<le> \<bar>Im z\<bar>})"
(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 \<equiv> "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 ((\<lambda>z. norm (h' z)) ` B)"
@@ -2560,10 +2560,10 @@
theorem inverse_squares_sums: "(\<lambda>n. 1 / (n + 1)\<^sup>2) sums (pi\<^sup>2 / 6)"
proof -
- def P \<equiv> "\<lambda>x n. (\<Prod>k=1..n. 1 - x^2 / of_nat k^2 :: real)"
- def K \<equiv> "\<Sum>n. inverse (real_of_nat (Suc n))^2"
- def f \<equiv> "\<lambda>x. \<Sum>n. P x n / of_nat (Suc n)^2"
- def g \<equiv> "\<lambda>x. (1 - sin (pi * x) / (pi * x))"
+ define P where "P x n = (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)" for x :: real and n
+ define K where "K = (\<Sum>n. inverse (real_of_nat (Suc n))^2)"
+ define f where [abs_def]: "f x = (\<Sum>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: "(\<lambda>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 \<midarrow> 0 \<rightarrow> pi^2 / 6"
proof (rule Lim_transform_eventually)
- def f' \<equiv> "\<lambda>x. \<Sum>n. - sin_coeff (n+3) * pi ^ (n+2) * x^n"
+ define f' where [abs_def]: "f' x = (\<Sum>n. - sin_coeff (n+3) * pi ^ (n+2) * x^n)" for x
have "eventually (\<lambda>x. x \<noteq> (0::real)) (at 0)"
by (auto simp add: eventually_at intro!: exI[of _ 1])
thus "eventually (\<lambda>x. f' x = f x) (at 0)"
--- a/src/HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy Mon Apr 25 16:09:26 2016 +0200
@@ -73,7 +73,7 @@
assumes "norm z < 1"
shows "(\<lambda>n. (a gchoose n) * z^n) sums (1 + z) powr a"
proof -
- def K \<equiv> "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 = "\<lambda>n. a gchoose n" and ?f' = "diffs (\<lambda>n. a gchoose n)"
@@ -83,7 +83,7 @@
hence summable': "summable (\<lambda>n. ?f' n * z ^ n)" if "norm z < K" for z using that
by (intro termdiff_converges[of _ K]) simp_all
- def f \<equiv> "\<lambda>z. \<Sum>n. ?f n * z ^ n" and f' \<equiv> "\<lambda>z. \<Sum>n. ?f' n * z ^ n"
+ define f f' where [abs_def]: "f z = (\<Sum>n. ?f n * z ^ n)" "f' z = (\<Sum>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 = (\<Sum>n. ?f' n * z^n) + (\<Sum>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 "(\<Sum>n. ?f' n * z^n) = (\<Sum>n. of_nat (Suc n) * ?f (Suc n) * z^n)"
by (intro suminf_cong) (simp add: diffs_def)
also have "(\<Sum>n. ?f' n * z^Suc n) = (\<Sum>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 "\<dots> = a * f z" unfolding f_def
+ also have "\<dots> = 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 = (\<Sum>n. if n = 0 then 1 else 0)" unfolding f_def by (intro suminf_cong) simp
+ have "f 0 = (\<Sum>n. if n = 0 then 1 else 0)" unfolding f_f'_def by (intro suminf_cong) simp
also have "\<dots> = 1" using sums_single[of 0 "\<lambda>_. 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':
--- a/src/HOL/Multivariate_Analysis/Harmonic_Numbers.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Harmonic_Numbers.thy Mon Apr 25 16:09:26 2016 +0200
@@ -250,7 +250,7 @@
assumes "(x::real) > 0" "a > 0"
shows "ln (x + a) - ln x \<le> a * (inverse x + inverse (x + a))/2" (is "_ \<le> ?A")
proof -
- def f' \<equiv> "(inverse (x + a) - inverse x)/a"
+ define f' where "f' = (inverse (x + a) - inverse x)/a"
have f'_nonpos: "f' \<le> 0" using assms by (simp add: f'_def divide_simps)
let ?f = "\<lambda>t. (t - x) * f' + inverse x"
let ?F = "\<lambda>t. (t - x)^2 * f' / 2 + t * inverse x"
@@ -297,8 +297,8 @@
assumes "(x::real) > 0" "x < y"
shows "ln y - ln x \<ge> 2 * (y - x) / (x + y)" (is "_ \<ge> ?A")
proof -
- def m \<equiv> "(x+y)/2"
- def f' \<equiv> "-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 = "\<lambda>t. (t - m)^2 * f' / 2 + t / m"
from assms have "((\<lambda>t. (t - m) * f' + inverse m) has_integral (?F y - ?F x)) {x..y}"
@@ -337,9 +337,10 @@
and euler_mascheroni_upper:
"euler_mascheroni \<le> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 1))"
proof -
- def D \<equiv> "\<lambda>n. inverse (of_nat (n+1)) + ln (of_nat (n+1)) - ln (of_nat (n+2)) :: real"
+ define D :: "_ \<Rightarrow> real"
+ where "D n = inverse (of_nat (n+1)) + ln (of_nat (n+1)) - ln (of_nat (n+2))" for n
let ?g = "\<lambda>n. ln (of_nat (n+2)) - ln (of_nat (n+1)) - inverse (of_nat (n+1)) :: real"
- def inv \<equiv> "\<lambda>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: "(\<lambda>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 "\<dots> \<le> -(\<Sum>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 \<equiv> "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 "-(\<Sum>k. D (k + Suc n)) \<ge> -(\<Sum>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 \<equiv> "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 = "\<lambda>k. 2 * y ^ (2*k+1) / of_nat (2*k+1)"
note sums = ln_series_quadratic[OF x(1)]
- def c \<equiv> "inverse (2*y^(2*n+1))"
+ define c where "c = inverse (2*y^(2*n+1))"
let ?d = "c * (ln x - (\<Sum>k<n. ?f k))"
have "\<forall>k. y\<^sup>2^k / of_nat (2*(k+n)+1) \<le> 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 \<in> {approx..approx + 2*d}"
and ln_approx_abs: "abs (ln x - (approx + d)) \<le> d"
proof -
- def c \<equiv> "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 - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) \<in>
--- a/src/HOL/Multivariate_Analysis/Integration.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Integration.thy Mon Apr 25 16:09:26 2016 +0200
@@ -928,7 +928,7 @@
proof
let ?B = "\<lambda>f::'a\<Rightarrow>'a \<times> 'a.
cbox (\<Sum>i\<in>Basis. (fst (f i) \<bullet> i) *\<^sub>R i) (\<Sum>i\<in>Basis. (snd (f i) \<bullet> i) *\<^sub>R i)"
- def p \<equiv> "?B ` (Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)})"
+ define p where "p = ?B ` (Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)})"
show "cbox c d \<in> p"
unfolding p_def
@@ -2004,10 +2004,7 @@
apply (drule choice)
apply blast
done
- def AB \<equiv> "\<lambda>n. (f ^^ n) (a,b)"
- def A \<equiv> "\<lambda>n. fst(AB n)"
- def B \<equiv> "\<lambda>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" "\<And>n. \<not> P (cbox (A(Suc n)) (B(Suc n))) \<and>
(\<forall>i\<in>Basis. A(n)\<bullet>i \<le> A(Suc n)\<bullet>i \<and> A(Suc n)\<bullet>i \<le> B(Suc n)\<bullet>i \<and> B(Suc n)\<bullet>i \<le> B(n)\<bullet>i \<and>
2 * (B(Suc n)\<bullet>i - A(Suc n)\<bullet>i) \<le> B(n)\<bullet>i - A(n)\<bullet>i)" (is "\<And>n. ?P n")
@@ -3453,8 +3450,8 @@
and "f integrable_on (cbox a b \<inter> {x. x\<bullet>k \<ge> c})" (is ?t2)
proof -
guess y using assms(1) unfolding integrable_on_def .. note y=this
- def b' \<equiv> "\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i)*\<^sub>R i::'a"
- def a' \<equiv> "\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i)*\<^sub>R i::'a"
+ define b' where "b' = (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i)*\<^sub>R i)"
+ define a' where "a' = (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>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 \<equiv> "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 \<equiv> "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 \<equiv> "{l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
- def d2 \<equiv> "{l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
- def cb \<equiv> "(\<Sum>i\<in>Basis. (if i = k then c else b\<bullet>i) *\<^sub>R i)::'a"
- def ca \<equiv> "(\<Sum>i\<in>Basis. (if i = k then c else a\<bullet>i) *\<^sub>R i)::'a"
+ define d1 where "d1 = {l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
+ define d2 where "d2 = {l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
+ define cb where "cb = (\<Sum>i\<in>Basis. (if i = k then c else b\<bullet>i) *\<^sub>R i)"
+ define ca where "ca = (\<Sum>i\<in>Basis. (if i = k then c else a\<bullet>i) *\<^sub>R i)"
note division_points_psubset[OF \<open>d division_of cbox a b\<close> 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 \<inter> {x. x\<bullet>k \<le> c})"
@@ -4734,7 +4731,7 @@
done
next
case False
- def d \<equiv> "e / 3 / setprod (\<lambda>i. b\<bullet>i - a\<bullet>i) (Basis - {k})"
+ define d where "d = e / 3 / setprod (\<lambda>i. b\<bullet>i - a\<bullet>i) (Basis - {k})"
note False[unfolded content_eq_0 not_ex not_le, rule_format]
then have "\<And>x. x \<in> Basis \<Longrightarrow> b\<bullet>x > a\<bullet>x"
by (auto simp add:not_le)
@@ -5227,7 +5224,7 @@
next
fix x k
assume xk: "(x, k) \<in> p"
- def n \<equiv> "nat \<lfloor>norm (f x)\<rfloor>"
+ define n where "n = nat \<lfloor>norm (f x)\<rfloor>"
have *: "norm (f x) \<in> (\<lambda>(x, k). norm (f x)) ` p"
using xk by auto
have nfx: "real n \<le> norm (f x)" "norm (f x) \<le> real n + 1"
@@ -6085,7 +6082,7 @@
using assms
proof cases
assume p: "p \<noteq> 1"
- def p' \<equiv> "p - 2"
+ define p' where "p' = p - 2"
from assms p have p': "{..<p} = {..Suc p'}" "p = Suc (Suc p')"
by (auto simp: p'_def)
have *: "\<And>i. i \<le> p' \<Longrightarrow> Suc (Suc p' - i) = (Suc (Suc p') - i)"
@@ -6125,8 +6122,9 @@
case 1
interpret bounded_bilinear "scaleR::real\<Rightarrow>'a\<Rightarrow>'a"
by (rule bounded_bilinear_scaleR)
- def g \<equiv> "\<lambda>s. (b - s)^(p - 1)/fact (p - 1)"
- def Dg \<equiv> "\<lambda>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 \<open>p > 0\<close>
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\<bullet>i = d\<bullet>i"
using as unfolding k mem_box by (metis antisym)
- def y \<equiv> "\<Sum>j\<in>Basis. (if j = i then if c\<bullet>i \<le> (a\<bullet>i + b\<bullet>i) / 2 then c\<bullet>i +
- min e (b\<bullet>i - c\<bullet>i) / 2 else c\<bullet>i - min e (c\<bullet>i - a\<bullet>i) / 2 else x\<bullet>j) *\<^sub>R j"
+ define y where "y = (\<Sum>j\<in>Basis. (if j = i then if c\<bullet>i \<le> (a\<bullet>i + b\<bullet>i) / 2 then c\<bullet>i +
+ min e (b\<bullet>i - c\<bullet>i) / 2 else c\<bullet>i - min e (c\<bullet>i - a\<bullet>i) / 2 else x\<bullet>j) *\<^sub>R j)"
show "\<exists>x'\<in>\<Union>(s - {k}). x' \<noteq> x \<and> 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' \<equiv> "\<lambda>x. {y. g y \<in> d (g x)}"
+ define d' where "d' x = {y. g y \<in> d (g x)}" for x
have d': "\<And>x. d' x = {y. g y \<in> (d (g x))}"
unfolding d'_def ..
show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of h ` cbox a b \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>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 \<equiv> "\<lambda>x. ball x w \<inter> d1 x"
+ define d where [abs_def]: "d x = ball x w \<inter> 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 \<equiv> "\<lambda>x. if x \<le> t then d1 x \<inter> d2 x else d1 x"
+ define d3 where "d3 x = (if x \<le> t then d1 x \<inter> 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 \<subseteq> cbox a b"
shows "((\<lambda>x. if x \<in> box c d then f x else 0) has_integral i) (cbox a b)"
proof -
- def g \<equiv> "\<lambda>x. if x \<in>box c d then f x else 0"
+ define g where [abs_def]: "g x = (if x \<in>box c d then f x else 0)" for x
{
presume *: "cbox c d \<noteq> {} \<Longrightarrow> ?thesis"
show ?thesis
@@ -8264,8 +8262,8 @@
next
assume as: "\<forall>e>0. ?r e"
from this[rule_format,OF zero_less_one] guess C .. note C=conjunctD2[OF this,rule_format]
- def c \<equiv> "(\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)::'n"
- def d \<equiv> "(\<Sum>i\<in>Basis. max B C *\<^sub>R i)::'n"
+ define c :: 'n where "c = (\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)"
+ define d :: 'n where "d = (\<Sum>i\<in>Basis. max B C *\<^sub>R i)"
have c_d: "cbox a b \<subseteq> 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 \<equiv> "(\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)::'n"
- def d \<equiv> "(\<Sum>i\<in>Basis. max B C *\<^sub>R i)::'n"
+ define c :: 'n where "c = (\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)"
+ define d :: 'n where "d = (\<Sum>i\<in>Basis. max B C *\<^sub>R i)"
have c_d: "cbox a b \<subseteq> 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 \<equiv> "\<Sum>i\<in>Basis. min (a\<bullet>i) (- (max B1 B2)) *\<^sub>R i::'n"
- def d \<equiv> "\<Sum>i\<in>Basis. max (b\<bullet>i) (max B1 B2) *\<^sub>R i::'n"
+ define c :: 'n where "c = (\<Sum>i\<in>Basis. min (a\<bullet>i) (- (max B1 B2)) *\<^sub>R i)"
+ define d :: 'n where "d = (\<Sum>i\<in>Basis. max (b\<bullet>i) (max B1 B2) *\<^sub>R i)"
have *: "ball 0 B1 \<subseteq> cbox c d" "ball 0 B2 \<subseteq> 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 \<equiv> "q - snd ` p"
+ define r where "r = q - snd ` p"
have "snd ` p \<inter> r = {}"
unfolding r_def by auto
have r: "finite r"
@@ -9567,7 +9565,7 @@
fixes x :: "'a::ring_1"
shows "(1 - x) * setsum (\<lambda>i. x^i) {0 .. n} = (1 - x^(Suc n))"
proof -
- def y \<equiv> "1 - x"
+ define y where "y = 1 - x"
have "y * (\<Sum>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 \<equiv> "\<lambda>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' \<equiv> "{(x,k) | x k. \<exists>i l. x \<in> i \<and> i \<in> d \<and> (x,l) \<in> p \<and> k = i \<inter> l}"
+ define p' where "p' = {(x,k) | x k. \<exists>i l. x \<in> i \<and> i \<in> d \<and> (x,l) \<in> p \<and> k = i \<inter> 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' \<equiv> "{cbox u v \<inter> l |l. l \<in> snd ` p \<and> cbox u v \<inter> l \<noteq> {}}"
+ define d' where "d' = {cbox u v \<inter> l |l. l \<in> snd ` p \<and> cbox u v \<inter> l \<noteq> {}}"
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 \<in> X0" "open X0"
"\<forall>x\<in>X0 \<inter> U. \<forall>t \<in> C. dist (fx (x, t)) (fx (x0, t)) \<le> e"
proof -
- def psi \<equiv> "\<lambda>(x, t). dist (fx (x, t)) (fx (x0, t))"
- def W0 \<equiv> "{(x, t) \<in> U \<times> C. psi (x, t) < e}"
+ define psi where "psi = (\<lambda>(x, t). dist (fx (x, t)) (fx (x0, t)))"
+ define W0 where "W0 = {(x, t) \<in> U \<times> C. psi (x, t) < e}"
have W0_eq: "W0 = psi -` {..<e} \<inter> U \<times> C"
by (auto simp: vimage_def W0_def)
have "open {..<e}" by simp
@@ -11668,7 +11665,7 @@
proof (safe intro!: tendstoI)
fix e'::real and x
assume "e' > 0"
- def e \<equiv> "e' / (content (cbox a b) + 1)"
+ define e where "e = e' / (content (cbox a b) + 1)"
have "e > 0" using \<open>e' > 0\<close> by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos)
assume "x \<in> U"
from continuous_on_prod_compactE[OF cont_fx compact_cbox \<open>x \<in> U\<close> \<open>0 < e\<close>]
@@ -11744,7 +11741,7 @@
note [continuous_intros] = continuous_on_compose2[OF cont_f1]
fix e'::real
assume "e' > 0"
- def e \<equiv> "e' / (content (cbox a b) + 1)"
+ define e where "e = e' / (content (cbox a b) + 1)"
have "e > 0" using \<open>e' > 0\<close> by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos)
from continuous_on_prod_compactE[OF cont_fx compact_cbox \<open>x0 \<in> U\<close> \<open>e > 0\<close>]
obtain X0 where X0: "x0 \<in> X0" "open X0"
@@ -11916,7 +11913,7 @@
proof (rule tendstoI)
fix e::real
assume "e > 0"
- def e' \<equiv> "e / 2"
+ define e' where "e' = e / 2"
with \<open>e > 0\<close> have "e' > 0" by simp
then have "\<forall>\<^sub>F n in F. \<forall>x\<in>cbox a b. norm (f n x - g x) < e' / content (cbox a b)"
using u content_nonzero content_pos_le[of a b]
--- a/src/HOL/Multivariate_Analysis/Path_Connected.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy Mon Apr 25 16:09:26 2016 +0200
@@ -563,7 +563,7 @@
assumes "path(g1 +++ g2)" "path g2"
shows "pathfinish g1 = pathstart g2"
proof (rule ccontr)
- def e \<equiv> "dist (g1 1) (g2 0)"
+ define e where "e = dist (g1 1) (g2 0)"
assume Neg: "pathfinish g1 \<noteq> 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 \<subseteq> 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)) \<in> s" and "0 \<le> u" "u \<le> 1"
have CC: "1 \<le> 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)" \<comment>\<open>massive duplication with the proof above\<close>
+ define D where "D = B / norm(y - a)" \<comment>\<open>massive duplication with the proof above\<close>
{ fix u
assume u: "(1 - u) *\<^sub>R y + u *\<^sub>R (a + D *\<^sub>R (y - a)) \<in> s" and "0 \<le> u" "u \<le> 1"
have DD: "1 \<le> 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: "\<And>x. connected_component (- s) z x \<Longrightarrow> norm x < B"
by (metis mem_Collect_eq)
- def C \<equiv> "((B + 1 + norm z) / norm (z-a))"
+ define C where "C = (B + 1 + norm z) / norm (z-a)"
have "C > 0"
using \<open>0 < B\<close> zna by (simp add: C_def divide_simps add_strict_increasing)
have "\<bar>norm (z + C *\<^sub>R (z-a)) - norm (C *\<^sub>R (z-a))\<bar> \<le> norm z"
@@ -2751,7 +2751,7 @@
{ fix \<gamma>
assume "path \<gamma>" and pimg_sbs: "path_image \<gamma> - {pathfinish \<gamma>} \<subseteq> interior (- t)"
and pf: "pathfinish \<gamma> \<in> frontier t" and ps: "pathstart \<gamma> = a"
- def c \<equiv> "pathfinish \<gamma>"
+ define c where "c = pathfinish \<gamma>"
have "c \<in> -s" unfolding c_def using front pf by blast
moreover have "open (-s)" using s compact_imp_closed by blast
ultimately obtain \<epsilon>::real where "\<epsilon> > 0" and \<epsilon>: "cball c \<epsilon> \<subseteq> -s"
@@ -3260,8 +3260,10 @@
and geq: "\<forall>x. k1 (1, x) = g x" "\<forall>x. k2 (0, x) = g x"
and k12: "\<forall>x. k1 (0, x) = f x" "\<forall>x. k2 (1, x) = h x"
and P: "\<forall>t\<in>{0..1}. P (\<lambda>x. k1 (t, x))" "\<forall>t\<in>{0..1}. P (\<lambda>x. k2 (t, x))"
- def k \<equiv> "\<lambda>y. if fst y \<le> 1 / 2 then (k1 o (\<lambda>x. (2 *\<^sub>R fst x, snd x))) y
- else (k2 o (\<lambda>x. (2 *\<^sub>R fst x -1, snd x))) y"
+ define k where "k y =
+ (if fst y \<le> 1 / 2
+ then (k1 o (\<lambda>x. (2 *\<^sub>R fst x, snd x))) y
+ else (k2 o (\<lambda>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} \<times> X) k"
--- a/src/HOL/Multivariate_Analysis/Summation.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Summation.thy Mon Apr 25 16:09:26 2016 +0200
@@ -142,7 +142,7 @@
finally have "l \<ge> 0" by simp
with l obtain l' where l': "l = ereal l'" by (cases l) simp_all
- def c \<equiv> "(1 - l') / 2"
+ define c where "c = (1 - l') / 2"
from l and \<open>l \<ge> 0\<close> have c: "l + c > l" "l' + c \<ge> 0" "l' + c < 1" unfolding c_def
by (simp_all add: field_simps l')
have "\<forall>C>l. eventually (\<lambda>n. ereal (root n (norm (f n))) < C) sequentially"
@@ -176,7 +176,7 @@
also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
finally have l_nonneg: "l \<ge> 0" by simp
- def c \<equiv> "if l = \<infinity> then 2 else 1 + (real_of_ereal l - 1) / 2"
+ define c where "c = (if l = \<infinity> then 2 else 1 + (real_of_ereal l - 1) / 2)"
from l l_nonneg consider "l = \<infinity>" | "\<exists>l'. l = ereal l'" by (cases l) simp_all
hence c: "c > 1 \<and> 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" "\<And>n. norm (f n) \<le> K" using BseqE by blast
- def n \<equiv> "nat \<lceil>log c K\<rceil>"
+ define n where "n = nat \<lceil>log c K\<rceil>"
from unbounded have "\<exists>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: "\<And>n. f n \<ge> 0"
shows "summable f \<longleftrightarrow> summable (\<lambda>n. 2^n * f (2^n))"
proof -
- def f' \<equiv> "\<lambda>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 (\<lambda>n. f (Suc n))" by (intro decseq_SucI) simp
hence mono': "f n \<le> f m" if "m \<le> 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 \<equiv> "(if l = \<infinity> then 1 else real_of_ereal l / 2)"
+ define r where "r = (if l = \<infinity> then 1 else real_of_ereal l / 2)"
from l have "r > 0 \<and> 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 (\<lambda>n. p n * f n / f (Suc n) - p (Suc n) > r) sequentially"
@@ -420,7 +420,7 @@
have "Bseq (\<lambda>n. (\<Sum>k\<le>n + Suc m. f k))"
proof (rule BseqI')
fix k :: nat
- def n \<equiv> "k + Suc m"
+ define n where "n = k + Suc m"
have n: "n > m" by (simp add: n_def)
from r have "r * norm (\<Sum>k\<le>n. f k) = norm (\<Sum>k\<le>n. r * f k)"
@@ -588,7 +588,7 @@
assumes "ereal (norm z) < conv_radius f"
shows "summable (\<lambda>n. norm (f n * z ^ n))"
proof (rule root_test_convergence')
- def l \<equiv> "limsup (\<lambda>n. ereal (root n (norm (f n))))"
+ define l where "l = limsup (\<lambda>n. ereal (root n (norm (f n))))"
have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
finally have l_nonneg: "l \<ge> 0" .
@@ -626,7 +626,7 @@
assumes "ereal (norm z) > conv_radius f"
shows "\<not>summable (\<lambda>n. f n * z ^ n)"
proof (rule root_test_divergence)
- def l \<equiv> "limsup (\<lambda>n. ereal (root n (norm (f n))))"
+ define l where "l = limsup (\<lambda>n. ereal (root n (norm (f n))))"
have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
finally have l_nonneg: "l \<ge> 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 \<equiv> "if R = \<infinity> then conv_radius' + 1 else (real_of_ereal R + conv_radius') / 2"
+ define r where "r = (if R = \<infinity> then conv_radius' + 1 else (real_of_ereal R + conv_radius') / 2)"
from R conv_radius_nonneg[of f] have "0 < r \<and> ereal r < R \<and> 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 (\<lambda>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 \<equiv> "if conv_radius f = \<infinity> then R' + 1 else (R' + real_of_ereal (conv_radius f)) / 2"
+ define r where
+ "r = (if conv_radius f = \<infinity> then R' + 1 else (R' + real_of_ereal (conv_radius f)) / 2)"
from R conv_radius_nonneg[of f] have "r > R \<and> 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 \<noteq> 0"
with conv_radius_nonneg[of f] have pos: "conv_radius f > 0" by simp
- def r \<equiv> "if conv_radius f = \<infinity> then 1 else real_of_ereal (conv_radius f) / 2"
+ define r where "r = (if conv_radius f = \<infinity> then 1 else real_of_ereal (conv_radius f) / 2)"
from pos have r: "ereal r > 0 \<and> ereal r < conv_radius f"
by (cases "conv_radius f") (simp_all add: r_def)
hence "summable (\<lambda>n. f n * of_real r ^ n)" by (intro summable_in_conv_radius) simp
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Apr 25 16:09:26 2016 +0200
@@ -228,7 +228,8 @@
"\<And>a. a \<in> A' \<Longrightarrow> open a"
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> \<exists>a\<in>A'. a \<subseteq> S"
by (rule first_countable_basisE) blast
- def A \<equiv> "(\<lambda>N. \<Inter>((\<lambda>n. from_nat_into A' n) ` N)) ` (Collect finite::nat set set)"
+ define A where [abs_def]:
+ "A = (\<lambda>N. \<Inter>((\<lambda>n. from_nat_into A' n) ` N)) ` (Collect finite::nat set set)"
then show "\<exists>A. countable A \<and> (\<forall>a. a \<in> A \<longrightarrow> x \<in> a) \<and> (\<forall>a. a \<in> A \<longrightarrow> open a) \<and>
(\<forall>S. open S \<longrightarrow> x \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S)) \<and> (\<forall>a b. a \<in> A \<longrightarrow> b \<in> A \<longrightarrow> a \<inter> b \<in> A)"
proof (safe intro!: exI[where x=A])
@@ -385,7 +386,7 @@
instance second_countable_topology \<subseteq> first_countable_topology
proof
fix x :: 'a
- def B \<equiv> "SOME B::'a set set. countable B \<and> topological_basis B"
+ define B :: "'a set set" where "B = (SOME B. countable B \<and> 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 \<equiv> "{x. \<exists>a\<in>S. \<exists>d>0. (\<forall>y\<in>U. dist y a < d \<longrightarrow> y \<in> S) \<and> dist x a < d}"
+ define T where "T = {x. \<exists>a\<in>S. \<exists>d>0. (\<forall>y\<in>U. dist y a < d \<longrightarrow> y \<in> S) \<and> dist x a < d}"
have 1: "\<forall>x\<in>T. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> T"
unfolding T_def
apply clarsimp
@@ -1030,7 +1031,7 @@
assumes "e > 0"
shows "\<exists>a b. (\<forall>i\<in>Basis. a \<bullet> i \<in> \<rat> \<and> b \<bullet> i \<in> \<rat> ) \<and> x \<in> box a b \<and> box a b \<subseteq> ball x e"
proof -
- def e' \<equiv> "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 "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x \<bullet> i \<and> x \<bullet> i - y < e'" (is "\<forall>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 \<in> Basis" "s j \<in> S" by blast
- def y \<equiv> "rec_list
- (s j)
- (\<lambda>j _ Y. (\<Sum>i\<in>Basis. (if i = j then s i \<bullet> i else Y \<bullet> i) *\<^sub>R i))"
+ define y where
+ "y = rec_list (s j) (\<lambda>j _ Y. (\<Sum>i\<in>Basis. (if i = j then s i \<bullet> i else Y \<bullet> i) *\<^sub>R i))"
have "x = (\<Sum>i\<in>Basis. (if i \<in> set bs then s i \<bullet> i else s j \<bullet> i) *\<^sub>R i)"
using bs by (auto simp add: s(1)[symmetric] euclidean_representation)
also have [symmetric]: "y bs = \<dots>"
@@ -2534,7 +2534,7 @@
"\<And>i. x \<in> A i"
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially"
by blast
- def f \<equiv> "\<lambda>n. SOME y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y"
+ define f where "f n = (SOME y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y)" for n
{
fix n
from \<open>?lhs\<close> have "\<exists>y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y"
@@ -3137,8 +3137,8 @@
then obtain r where "0 < r" "\<forall>z. dist z y < r \<longrightarrow> z \<in> T"
unfolding open_dist by fast
(* choose point between x and y, within distance r of y. *)
- def k \<equiv> "min 1 (r / (2 * dist x y))"
- def z \<equiv> "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 @@
"\<And>i. l \<in> A i"
"\<And>S. open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially"
by blast
- def s \<equiv> "\<lambda>n i. SOME j. i < j \<and> f j \<in> A (Suc n)"
+ define s where "s n i = (SOME j. i < j \<and> f j \<in> A (Suc n))" for n i
{
fix n i
have "infinite (A (Suc n) \<inter> range f - f`{.. i})"
@@ -3691,7 +3691,7 @@
unfolding s_def by (auto intro: someI2_ex)
}
note s = this
- def r \<equiv> "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 \<noteq> {}"
by (auto simp: eventually_False)
- def Z \<equiv> "closure ` {A. eventually (\<lambda>x. x \<in> A) F}"
+ define Z where "Z = closure ` {A. eventually (\<lambda>x. x \<in> A) F}"
then have "\<forall>z\<in>Z. closed z"
by auto
moreover
@@ -4018,7 +4018,7 @@
next
fix A
assume A: "\<forall>a\<in>A. closed a" "\<forall>B\<subseteq>A. finite B \<longrightarrow> U \<inter> \<Inter>B \<noteq> {}" "U \<inter> \<Inter>A = {}"
- def F \<equiv> "INF a:insert U A. principal a"
+ define F where "F = (INF a:insert U A. principal a)"
have "F \<noteq> bot"
unfolding F_def
proof (rule INF_filter_not_bot)
@@ -4089,8 +4089,7 @@
fix A
assume A: "\<forall>a\<in>A. open a" "U \<subseteq> \<Union>A"
assume *: "\<forall>A. countable A \<longrightarrow> (\<forall>a\<in>A. open a) \<longrightarrow> U \<subseteq> \<Union>A \<longrightarrow> (\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T)"
-
- moreover def C \<equiv> "{b\<in>B. \<exists>a\<in>A. b \<inter> U \<subseteq> a}"
+ moreover define C where "C = {b\<in>B. \<exists>a\<in>A. b \<inter> U \<subseteq> a}"
ultimately have "countable C" "\<forall>a\<in>C. open a"
unfolding C_def using ccover by auto
moreover
@@ -4202,7 +4201,7 @@
by auto
then obtain X' where T: "\<And>T. T \<subseteq> A \<Longrightarrow> finite T \<Longrightarrow> X' T \<in> U - \<Union>T"
by metis
- def X \<equiv> "\<lambda>n. X' (from_nat_into A ` {.. n})"
+ define X where "X n = X' (from_nat_into A ` {.. n})" for n
have X: "\<And>n. X n \<in> U - (\<Union>i\<le>n. from_nat_into A i)"
using \<open>A \<noteq> {}\<close> unfolding X_def by (intro T) (auto intro: from_nat_into)
then have "range X \<subseteq> U"
@@ -4249,7 +4248,7 @@
"\<And>i. x \<in> A i"
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially"
by blast
- def s \<equiv> "\<lambda>n i. SOME j. i < j \<and> X j \<in> A (Suc n)"
+ define s where "s n i = (SOME j. i < j \<and> X j \<in> A (Suc n))" for n i
{
fix n i
have "\<exists>a. i < a \<and> X a \<in> A (Suc n)"
@@ -4270,7 +4269,7 @@
unfolding s_def by (auto intro: someI2_ex)
}
note s = this
- def r \<equiv> "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 \<subseteq> s"
shows "\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t)"
proof (rule ccontr)
- def C \<equiv> "(\<lambda>F. interior (F \<union> (- t))) ` {F. finite F \<and> F \<subseteq> t }"
+ define C where "C = (\<lambda>F. interior (F \<union> (- t))) ` {F. finite F \<and> F \<subseteq> t }"
note \<open>countably_compact s\<close>
moreover have "\<forall>t\<in>C. open t"
by (auto simp: C_def)
@@ -4442,7 +4441,7 @@
from seq_compact_imp_totally_bounded[OF \<open>seq_compact s\<close>]
obtain f where f: "\<forall>e>0. finite (f e) \<and> f e \<subseteq> s \<and> s \<subseteq> (\<Union>x\<in>f e. ball x e)"
unfolding choice_iff' ..
- def K \<equiv> "(\<lambda>(x, r). ball x r) ` ((\<Union>e \<in> \<rat> \<inter> {0 <..}. f e) \<times> \<rat>)"
+ define K where "K = (\<lambda>(x, r). ball x r) ` ((\<Union>e \<in> \<rat> \<inter> {0 <..}. f e) \<times> \<rat>)"
have "countably_compact s"
using \<open>seq_compact s\<close> 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:"((\<lambda>i. f (r1 (r2 i)) proj k) \<longlongrightarrow> l2) sequentially"
using bounded_imp_convergent_subsequence[of "\<lambda>i. f (r1 i) proj k"]
by (auto simp: o_def)
- def r \<equiv> "r1 \<circ> r2"
+ define r where "r = r1 \<circ> r2"
have r:"subseq r"
using r1 and r2 unfolding r_def o_def subseq_def by auto
moreover
- def l \<equiv> "unproj (\<lambda>i. if i = k then l2 else l1 proj i)::'a"
+ define l where "l = unproj (\<lambda>i. if i = k then l2 else l1 proj i)"
{
fix e::real
assume "e > 0"
@@ -4822,9 +4821,9 @@
fix f :: "nat \<Rightarrow> 'a"
assume f: "\<forall>n. f n \<in> s"
- def e \<equiv> "\<lambda>n. 1 / (2 * Suc n)"
+ define e where "e n = 1 / (2 * Suc n)" for n
then have [simp]: "\<And>n. 0 < e n" by auto
- def B \<equiv> "\<lambda>n U. SOME b. infinite {n. f n \<in> b} \<and> (\<exists>x. b \<subseteq> ball x (e n) \<inter> U)"
+ define B where "B n U = (SOME b. infinite {n. f n \<in> b} \<and> (\<exists>x. b \<subseteq> ball x (e n) \<inter> U))" for n U
{
fix n U
assume "infinite {n. f n \<in> U}"
@@ -4841,7 +4840,7 @@
}
note B = this
- def F \<equiv> "rec_nat (B 0 UNIV) B"
+ define F where "F = rec_nat (B 0 UNIV) B"
{
fix n
have "infinite {i. f i \<in> F n}"
@@ -4862,7 +4861,7 @@
by (simp add: set_eq_iff not_le conj_commute)
qed
- def t \<equiv> "rec_nat (sel 0 0) (\<lambda>n i. sel (Suc n) i)"
+ define t where "t = rec_nat (sel 0 0) (\<lambda>n i. sel (Suc n) i)"
have "subseq t"
unfolding subseq_Suc_iff by (simp add: t_def sel)
moreover have "\<forall>i. (f \<circ> t) i \<in> s"
@@ -5587,8 +5586,8 @@
using choice[of "\<lambda>d x. d>0 \<longrightarrow> fst x \<in> s \<and> snd x \<in> s \<and> dist (snd x) (fst x) < d \<and> \<not> dist (f (snd x)) (f (fst x)) < e"]
unfolding Bex_def
by (auto simp add: dist_commute)
- def x \<equiv> "\<lambda>n::nat. fst (fa (inverse (real n + 1)))"
- def y \<equiv> "\<lambda>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: "\<forall>n. x n \<in> s \<and> y n \<in> s"
and xy0: "\<forall>n. dist (x n) (y n) < inverse (real n + 1)"
and fxy:"\<forall>n. \<not> dist (f (x n)) (f (y n)) < e"
@@ -6552,7 +6551,8 @@
proof (cases, safe)
fix e :: real
assume "0 < e" "s \<noteq> {}"
- def [simp]: R \<equiv> "{(y, d). y \<in> s \<and> 0 < d \<and> ball y d \<inter> s \<subseteq> {x \<in> s. f x \<in> ball (f y) (e/2) } }"
+ define R where [simp]:
+ "R = {(y, d). y \<in> s \<and> 0 < d \<and> ball y d \<inter> s \<subseteq> {x \<in> s. f x \<in> ball (f y) (e/2)}}"
let ?b = "(\<lambda>(y, d). ball y (d/2))"
have "(\<forall>r\<in>R. open (?b r))" "s \<subseteq> (\<Union>r\<in>R. ?b r)"
proof safe
@@ -7451,13 +7451,13 @@
instance euclidean_space \<subseteq> second_countable_topology
proof
- def a \<equiv> "\<lambda>f :: 'a \<Rightarrow> (real \<times> real). \<Sum>i\<in>Basis. fst (f i) *\<^sub>R i"
+ define a where "a f = (\<Sum>i\<in>Basis. fst (f i) *\<^sub>R i)" for f :: "'a \<Rightarrow> real \<times> real"
then have a: "\<And>f. (\<Sum>i\<in>Basis. fst (f i) *\<^sub>R i) = a f"
by simp
- def b \<equiv> "\<lambda>f :: 'a \<Rightarrow> (real \<times> real). \<Sum>i\<in>Basis. snd (f i) *\<^sub>R i"
+ define b where "b f = (\<Sum>i\<in>Basis. snd (f i) *\<^sub>R i)" for f :: "'a \<Rightarrow> real \<times> real"
then have b: "\<And>f. (\<Sum>i\<in>Basis. snd (f i) *\<^sub>R i) = b f"
by simp
- def B \<equiv> "(\<lambda>f. box (a f) (b f)) ` (Basis \<rightarrow>\<^sub>E (\<rat> \<times> \<rat>))"
+ define B where "B = (\<lambda>f. box (a f) (b f)) ` (Basis \<rightarrow>\<^sub>E (\<rat> \<times> \<rat>))"
have "Ball B open" by (simp add: B_def open_box)
moreover have "(\<forall>A. open A \<longrightarrow> (\<exists>B'\<subseteq>B. \<Union>B' = A))"
@@ -7655,7 +7655,7 @@
{
fix x
assume as:"x \<in> cbox a b"
- def f \<equiv> "\<lambda>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 <e b \<longrightarrow> a <e f n \<longrightarrow> f n = x" and xc: "x \<noteq> ?c"
@@ -7712,7 +7712,7 @@
proof -
obtain b where "b>0" and b: "\<forall>x\<in>s. norm x \<le> b"
using assms[unfolded bounded_pos] by auto
- def a \<equiv> "(\<Sum>i\<in>Basis. (b + 1) *\<^sub>R i)::'a"
+ define a :: 'a where "a = (\<Sum>i\<in>Basis. (b + 1) *\<^sub>R i)"
{
fix x
assume "x \<in> s"
@@ -8062,7 +8062,7 @@
assumes "compact s" "continuous_on s f" "f ` s = t" "inj_on f s"
shows "\<exists>g. homeomorphism s t f g"
proof -
- def g \<equiv> "\<lambda>x. SOME y. y\<in>s \<and> f y = x"
+ define g where "g x = (SOME y. y\<in>s \<and> f y = x)" for x
have g: "\<forall>x\<in>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 \<in> s" by auto
- def z \<equiv> "\<lambda>n. (f ^^ n) z0"
+ define z where "z n = (f ^^ n) z0" for n
{
fix n :: nat
have "z n \<in> s" unfolding z_def
@@ -8498,7 +8498,7 @@
then show ?case using f by auto qed
} note z_in_s = this
- def d \<equiv> "dist (z 0) (z 1)"
+ define d where "d = dist (z 0) (z 1)"
have fzn:"\<And>n. f (z n) = z (Suc n)" unfolding z_def by auto
{
@@ -8608,7 +8608,7 @@
then obtain x where "x\<in>s" and x:"(z \<longlongrightarrow> x) sequentially"
using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto
- def e \<equiv> "dist (f x) x"
+ define e where "e = dist (f x) x"
have "e = 0"
proof (rule ccontr)
assume "e \<noteq> 0"
--- a/src/HOL/Multivariate_Analysis/Uniform_Limit.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Uniform_Limit.thy Mon Apr 25 16:09:26 2016 +0200
@@ -72,7 +72,7 @@
shows "(h \<longlongrightarrow> l) (at x within S)"
proof (rule tendstoI)
fix e :: real
- def e' \<equiv> "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 \<open>0 < e'\<close>]
--- a/src/HOL/Multivariate_Analysis/Weierstrass.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Weierstrass.thy Mon Apr 25 16:09:26 2016 +0200
@@ -218,7 +218,7 @@
have "t \<noteq> t0" using t t0 by auto
then obtain g where g: "g \<in> R" "g t \<noteq> g t0"
using separable t0 by (metis Diff_subset subset_eq t)
- def h \<equiv> "\<lambda>x. g x - g t0"
+ define h where [abs_def]: "h x = g x - g t0" for x
have "h \<in> R"
unfolding h_def by (fast intro: g const diff)
then have hsq: "(\<lambda>w. (h w)\<^sup>2) \<in> R"
@@ -232,7 +232,7 @@
also have "... \<le> normf (\<lambda>w. (h w)\<^sup>2)"
using t normf_upper [where x=t] continuous [OF hsq] by force
finally have nfp: "0 < normf (\<lambda>w. (h w)\<^sup>2)" .
- def p \<equiv> "\<lambda>x. (1 / normf (\<lambda>w. (h w)\<^sup>2)) * (h x)^2"
+ define p where [abs_def]: "p x = (1 / normf (\<lambda>w. (h w)\<^sup>2)) * (h x)^2" for x
have "p \<in> 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 \<equiv> "\<lambda>x. (1 / card subU) * (\<Sum>t \<in> subU. pf t x)"
+ define p where [abs_def]: "p x = (1 / card subU) * (\<Sum>t \<in> subU. pf t x)" for x
have pR: "p \<in> 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: "\<And>x. x \<in> s-U \<Longrightarrow> p x \<ge> delta0"
by (force simp: ball_def dist_norm dest: p01)
- def \<delta> \<equiv> "delta0/2"
+ define \<delta> where "\<delta> = delta0/2"
have "delta0 \<le> 1" using delta0 p01 [of t1] t1
by (force simp: ball_def dist_norm dest: p01)
with delta0 have \<delta>01: "0 < \<delta>" "\<delta> < 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 \<inter> s = {x\<in>s. p x < \<delta>/2}"
by blast
- def k \<equiv> "nat\<lfloor>1/\<delta>\<rfloor> + 1"
+ define k where "k = nat\<lfloor>1/\<delta>\<rfloor> + 1"
have "k>0" by (simp add: k_def)
have "k-1 \<le> 1/\<delta>"
using \<delta>01 by (simp add: k_def)
@@ -331,7 +331,7 @@
using \<delta>01 unfolding k_def by linarith
with \<delta>01 k2\<delta> have k\<delta>: "1 < k*\<delta>" "k*\<delta> < 2"
by (auto simp: divide_simps)
- def q \<equiv> "\<lambda>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 \<in> R" for n
by (simp add: q_def const diff power pR)
have q01: "\<And>n t. t \<in> s \<Longrightarrow> q n t \<in> {0..1}"
@@ -387,7 +387,8 @@
by (fastforce simp: field_simps)
finally have "q n t \<le> (1 / (real k * \<delta>)) ^ n " .
} note limitNonU = this
- def NN \<equiv> "\<lambda>e. 1 + nat \<lceil>max (ln e / ln (real k * \<delta> / 2)) (- ln e / ln (real k * \<delta>))\<rceil>"
+ define NN
+ where "NN e = 1 + nat \<lceil>max (ln e / ln (real k * \<delta> / 2)) (- ln e / ln (real k * \<delta>))\<rceil>" for e
have NN: "of_nat (NN e) > ln e / ln (real k * \<delta> / 2)" "of_nat (NN e) > - ln e / ln (real k * \<delta>)"
if "0<e" for e
unfolding NN_def by linarith+
@@ -464,7 +465,7 @@
"\<And>w. w \<in> A \<Longrightarrow> ff w \<in> R \<and> ff w ` s \<subseteq> {0..1} \<and>
(\<forall>x \<in> s \<inter> Vf w. ff w x < e / card subA) \<and> (\<forall>x \<in> s \<inter> B. ff w x > 1 - e / card subA)"
by metis
- def pff \<equiv> "\<lambda>x. (\<Prod>w \<in> subA. ff w x)"
+ define pff where [abs_def]: "pff x = (\<Prod>w \<in> subA. ff w x)" for x
have pffR: "pff \<in> 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 "\<exists>g \<in> R. \<forall>x\<in>s. \<bar>f x - g x\<bar> < 2*e"
proof -
- def n \<equiv> "1 + nat \<lceil>normf f / e\<rceil>"
- def A \<equiv> "\<lambda>j::nat. {x \<in> s. f x \<le> (j - 1/3)*e}"
- def B \<equiv> "\<lambda>j::nat. {x \<in> s. f x \<ge> (j + 1/3)*e}"
+ define n where "n = 1 + nat \<lceil>normf f / e\<rceil>"
+ define A where "A j = {x \<in> s. f x \<le> (j - 1/3)*e}" for j :: nat
+ define B where "B j = {x \<in> s. f x \<ge> (j + 1/3)*e}" for j :: nat
have ngt: "(n-1) * e \<ge> normf f" "n\<ge>1"
using e
apply (simp_all add: n_def field_simps of_nat_Suc)
@@ -591,7 +592,7 @@
and xfA: "\<And>x j. x \<in> A j \<Longrightarrow> xf j x < e/n"
and xfB: "\<And>x j. x \<in> B j \<Longrightarrow> xf j x > 1 - e/n"
by metis
- def g \<equiv> "\<lambda>x. e * (\<Sum>i\<le>n. xf i x)"
+ define g where [abs_def]: "g x = e * (\<Sum>i\<le>n. xf i x)" for x
have gR: "g \<in> R"
unfolding g_def by (fast intro: mult const setsum xfR)
have gge0: "\<And>x. x \<in> s \<Longrightarrow> g x \<ge> 0"
@@ -606,7 +607,7 @@
done
{ fix t
assume t: "t \<in> s"
- def j \<equiv> "LEAST j. t \<in> A j"
+ define j where "j = (LEAST j. t \<in> A j)"
have jn: "j \<le> n"
using t An by (simp add: Least_le j_def)
have Aj: "t \<in> A j"
--- a/src/HOL/Multivariate_Analysis/ex/Approximations.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/ex/Approximations.thy Mon Apr 25 16:09:26 2016 +0200
@@ -89,7 +89,7 @@
assumes "n > 0" "0 \<le> x" "x < 2"
shows "exp (x::real) - (\<Sum>k<n. x^k / fact k) \<in> {0..(2 * x^n / (2 - x)) / fact n}"
proof (unfold atLeastAtMost_iff, safe)
- def approx \<equiv> "(\<Sum>k<n. x^k / fact k)"
+ define approx where "approx = (\<Sum>k<n. x^k / fact k)"
have "(\<lambda>k. 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 \<equiv> "47388813395531028639296492901910937/82101866951584879688289000000000000 :: real"
- def u \<equiv> "142196984054132045946501548559032969 / 246305600854754639064867000000000000 :: real"
+ define l :: real
+ where "l = 47388813395531028639296492901910937/82101866951584879688289000000000000"
+ define u :: real
+ where "u = 142196984054132045946501548559032969 / 246305600854754639064867000000000000"
have impI: "P \<longrightarrow> 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 \<le> x" "x < 1" and n: "even n"
shows "arctan x - arctan_approx n x \<in> {0..x^(2*n+1) / (1-x^4)}"
proof -
- def c \<equiv> "\<lambda>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 "(\<lambda>k. (-1) ^ k * (1 / real (k * 2 + 1) * x^(k*2+1))) sums arctan x"
using arctan_series' by simp
also have "(\<lambda>k. (-1) ^ k * (1 / real (k * 2 + 1) * x^(k*2+1))) =
@@ -559,11 +561,13 @@
text \<open>We can now approximate pi to 22 decimals within a fraction of a second.\<close>
lemma pi_approx_75: "abs (pi - 3.1415926535897932384626 :: real) \<le> inverse (10^22)"
proof -
- def a \<equiv> "8295936325956147794769600190539918304 / 2626685325478320010006427764892578125 :: real"
- def b \<equiv> "8428294561696506782041394632 / 503593538783547230635598424135 :: real"
+ define a :: real
+ where "a = 8295936325956147794769600190539918304 / 2626685325478320010006427764892578125"
+ define b :: real
+ where "b = 8428294561696506782041394632 / 503593538783547230635598424135"
\<comment> \<open>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.}\<close>
- def eq \<equiv> "op = :: real \<Rightarrow> real \<Rightarrow> bool"
+ define eq :: "real \<Rightarrow> real \<Rightarrow> bool" where "eq = op ="
\<comment> \<open>Splitting the computation into several steps has the advantage that simplification can
be done in parallel\<close>
@@ -664,15 +668,18 @@
"abs (pi - 3.141592653589793238462643383279502884197169399375105821 :: real) \<le> inverse (10^54)"
(is "abs (pi - ?pi') \<le> _")
proof -
- def a \<equiv> "2829469759662002867886529831139137601191652261996513014734415222704732791803 /
- 1062141879292765061960538947347721564047051545995266466660439319087625011200 :: real"
- def b \<equiv> "13355545553549848714922837267299490903143206628621657811747118592 /
- 23792006023392488526789546722992491355941103837356113731091180925 :: real"
- def c \<equiv> "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 \<equiv> "op = :: real \<Rightarrow> real \<Rightarrow> bool"
+ define eq :: "real \<Rightarrow> real \<Rightarrow> bool" where "eq = op ="
have "abs (pi - pi_approx2 4) \<le> inverse (2^183)" by (rule pi_approx2') simp_all
also have "pi_approx2 4 = 48 * arctan_approx 24 (1 / 18) +
--- a/src/HOL/Nat.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Nat.thy Mon Apr 25 16:09:26 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 \<equiv> "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
--- a/src/HOL/Nonstandard_Analysis/HLim.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Nonstandard_Analysis/HLim.thy Mon Apr 25 16:09:26 2016 +0200
@@ -209,7 +209,7 @@
done
text\<open>NS continuity can be defined using NS Limit in
- similar fashion to standard def of continuity\<close>
+ similar fashion to standard definition of continuity\<close>
lemma isNSCont_NSLIM_iff: "(isNSCont f a) = (f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S (f a))"
by (blast intro: isNSCont_NSLIM NSLIM_isNSCont)
@@ -236,7 +236,7 @@
text\<open>Alternative definition of continuity\<close>
(* Prove equivalence between NS limits - *)
-(* seems easier than using standard def *)
+(* seems easier than using standard definition *)
lemma NSLIM_h_iff: "(f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L) = ((%h. f(a + h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S L)"
apply (simp add: NSLIM_def, auto)
apply (drule_tac x = "star_of a + x" in spec)
--- a/src/HOL/NthRoot.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/NthRoot.thy Mon Apr 25 16:09:26 2016 +0200
@@ -646,7 +646,7 @@
lemma LIMSEQ_root: "(\<lambda>n. root n n) \<longlonglongrightarrow> 1"
proof -
- def x \<equiv> "\<lambda>n. root n n - 1"
+ define x where "x n = root n n - 1" for n
have "x \<longlonglongrightarrow> sqrt 0"
proof (rule tendsto_sandwich[OF _ _ tendsto_const])
show "(\<lambda>x. sqrt (2 / x)) \<longlonglongrightarrow> sqrt 0"
@@ -684,7 +684,7 @@
shows "(\<lambda>n. root n c) \<longlonglongrightarrow> 1"
proof -
{ fix c :: real assume "1 \<le> c"
- def x \<equiv> "\<lambda>n. root n c - 1"
+ define x where "x n = root n c - 1" for n
have "x \<longlonglongrightarrow> 0"
proof (rule tendsto_sandwich[OF _ _ tendsto_const])
show "(\<lambda>n. c / n) \<longlonglongrightarrow> 0"
--- a/src/HOL/Number_Theory/Eratosthenes.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Number_Theory/Eratosthenes.thy Mon Apr 25 16:09:26 2016 +0200
@@ -96,7 +96,7 @@
with A B C show ?thesis by simp
next
case False
- def m \<equiv> "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 \<equiv> "Suc m" and w \<equiv> "Suc n"
+ { define v where "v = Suc m"
+ define w where "w = Suc n"
fix q
assume "m + n \<le> q"
then obtain r where q: "q = m + n + r" by (auto simp add: le_iff_add)
@@ -359,7 +360,7 @@
have "List.find (\<lambda>p. p \<ge> m) (primes_upto n) = Some (smallest_prime_beyond m)"
if assms: "m \<le> p" "prime p" "p \<le> n" for p
proof -
- def A \<equiv> "{p. p \<le> n \<and> prime p \<and> m \<le> p}"
+ define A where "A = {p. p \<le> n \<and> prime p \<and> m \<le> p}"
from assms have "smallest_prime_beyond m \<le> p"
by (auto intro: smallest_prime_beyond_smallest)
from this \<open>p \<le> n\<close> have *: "smallest_prime_beyond m \<le> n"
--- a/src/HOL/Number_Theory/Euclidean_Algorithm.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Number_Theory/Euclidean_Algorithm.thy Mon Apr 25 16:09:26 2016 +0200
@@ -159,8 +159,8 @@
next
case True
then obtain l\<^sub>0 where l\<^sub>0_props: "l\<^sub>0 \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l\<^sub>0)" by blast
- def n \<equiv> "LEAST n. \<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n"
- def l \<equiv> "SOME l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n"
+ define n where "n = (LEAST n. \<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n)"
+ define l where "l = (SOME l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n)"
have "\<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n"
apply (subst n_def)
apply (rule LeastI[of _ "euclidean_size l\<^sub>0"])
--- a/src/HOL/Number_Theory/Factorial_Ring.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Number_Theory/Factorial_Ring.thy Mon Apr 25 16:09:26 2016 +0200
@@ -258,7 +258,7 @@
by (rule dvd_mult_left)
with Suc \<open>is_prime p\<close> \<open>\<not> p dvd a\<close> have "p dvd b"
by (simp add: prime_dvd_mult_iff)
- moreover def c \<equiv> "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 \<noteq> 0"
by (auto intro!: is_prime_not_zeroI)
- def n \<equiv> "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 "\<not> p ^ Suc n dvd a"
proof -
- def N \<equiv> "{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 \<open>is_prime p\<close> have "inj (op ^ p)" .
then have "inj_on (op ^ p) U" for U
@@ -402,7 +402,7 @@
then show "\<not> p ^ Suc n dvd a"
by (simp add: n_M)
qed
- def b \<equiv> "a div p ^ n"
+ define b where "b = a div p ^ n"
with \<open>p ^ n dvd a\<close> have a: "a = p ^ n * b"
by simp
with \<open>\<not> p ^ Suc n dvd a\<close> have "\<not> p dvd b" and "b \<noteq> 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 \<equiv> "the (factorization b)"
+ define Q where "Q = the (factorization b)"
with \<open>b \<noteq> 0\<close> have [simp]: "factorization b = Some Q"
by simp
from \<open>a \<noteq> 0\<close> have "factorization a =
--- a/src/HOL/Power.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Power.thy Mon Apr 25 16:09:26 2016 +0200
@@ -90,7 +90,7 @@
case 0 then show ?case by (simp add: fun_eq_iff)
next
case (Suc n)
- def g \<equiv> "\<lambda>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
--- a/src/HOL/Probability/Binary_Product_Measure.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Binary_Product_Measure.thy Mon Apr 25 16:09:26 2016 +0200
@@ -995,7 +995,7 @@
finally show ?thesis by (simp add: top_unique)
next
case False
- def C' \<equiv> "fst ` C"
+ define C' where "C' = fst ` C"
have "\<infinity> = \<integral>\<^sup>+ x. ennreal (1 / Suc n) * indicator C' x \<partial>count_space UNIV"
using C'_def False by(simp add: nn_integral_cmult ennreal_mult_top)
also have "\<dots> = \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' x * indicator {SOME y. (x, y) \<in> C} y \<partial>count_space UNIV \<partial>count_space UNIV"
--- a/src/HOL/Probability/Bochner_Integration.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Bochner_Integration.thy Mon Apr 25 16:09:26 2016 +0200
@@ -24,7 +24,7 @@
obtain D :: "'b set" where "countable D" and D: "\<And>X. open X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> \<exists>d\<in>D. d \<in> X"
by (erule countable_dense_setE)
- def e \<equiv> "from_nat_into D"
+ define e where "e = from_nat_into D"
{ fix n x
obtain d where "d \<in> D" and d: "d \<in> 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 \<equiv> "\<lambda>m n. {x\<in>space M. dist (f x) (e n) < 1 / (Suc m) \<and> 1 / (Suc m) \<le> dist (f x) z}"
- def B \<equiv> "\<lambda>m. disjointed (A m)"
+ define A where [abs_def]: "A m n =
+ {x\<in>space M. dist (f x) (e n) < 1 / (Suc m) \<and> 1 / (Suc m) \<le> dist (f x) z}" for m n
+ define B where [abs_def]: "B m = disjointed (A m)" for m
- def m \<equiv> "\<lambda>N x. Max {m::nat. m \<le> N \<and> x \<in> (\<Union>n\<le>N. B m n)}"
- def F \<equiv> "\<lambda>N::nat. \<lambda>x. if (\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)) \<and> (\<exists>n\<le>N. x \<in> B (m N x) n)
- then e (LEAST n. x \<in> B (m N x) n) else z"
+ define m where [abs_def]: "m N x = Max {m. m \<le> N \<and> x \<in> (\<Union>n\<le>N. B m n)}" for N x
+ define F where [abs_def]: "F N x =
+ (if (\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)) \<and> (\<exists>n\<le>N. x \<in> B (m N x) n)
+ then e (LEAST n. x \<in> B (m N x) n) else z)" for N x
have B_imp_A[intro, simp]: "\<And>x m n. x \<in> B m n \<Longrightarrow> x \<in> A m n"
using disjointed_subset[of "A m" for m] unfolding B_def by auto
@@ -86,7 +88,7 @@
then have 1: "\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)" by auto
from m[OF this] obtain n where n: "m N x \<le> N" "n \<le> N" "x \<in> B (m N x) n" by auto
moreover
- def L \<equiv> "LEAST n. x \<in> B (m N x) n"
+ define L where "L = (LEAST n. x \<in> B (m N x) n)"
have "dist (f x) (e L) < 1 / Suc (m N x)"
proof -
have "x \<in> B (m N x) L"
@@ -175,7 +177,7 @@
sup: "\<And>x. (SUP i. U i x) = ennreal (u x)"
by blast
- def U' \<equiv> "\<lambda>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]: "\<And>i. simple_function M (U' i)"
using U by (auto intro!: simple_function_compose1[where g=enn2real])
@@ -268,7 +270,7 @@
assume non_empty: "\<exists>x\<in>space M. f x \<noteq> 0"
- def m \<equiv> "Min (f`space M - {0})"
+ define m where "m = Min (f`space M - {0})"
have "m \<in> 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]: "\<And>i. s i \<in> borel_measurable M"
using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
- def m \<equiv> "if space M = {} then 0 else Max ((\<lambda>x. norm (s i x))`space M)"
+ define m where "m = (if space M = {} then 0 else Max ((\<lambda>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' \<equiv> "\<lambda>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: "\<And>i x. x \<in> space M \<Longrightarrow> s' i x = s i x"
by simp
@@ -2600,7 +2602,8 @@
have "\<And>i. s i \<in> measurable (N \<Otimes>\<^sub>M M) (count_space UNIV)"
by (rule measurable_simple_function) fact
- def f' \<equiv> "\<lambda>i x. if integrable M (f x) then simple_bochner_integral M (\<lambda>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 (\<lambda>y. s i (x, y)) else 0)" for i x
{ fix i x assume "x \<in> space N"
then have "simple_bochner_integral M (\<lambda>y. s i (x, y)) =
--- a/src/HOL/Probability/Borel_Space.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Borel_Space.thy Mon Apr 25 16:09:26 2016 +0200
@@ -282,34 +282,35 @@
assumes "a \<le> b" "(c::real) \<le> d" "{c..d} \<subseteq> {g a..g b}"
obtains c' d' where "{a..b} \<inter> g -` {c..d} = {c'..d'}" "c' \<le> d'" "g c' = c" "g d' = d"
proof-
- let ?A = "{a..b} \<inter> g -` {c..d}"
- from IVT'[of g a c b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5)
- obtain c'' where c'': "c'' \<in> ?A" "g c'' = c" by auto
- from IVT'[of g a d b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5)
- obtain d'' where d'': "d'' \<in> ?A" "g d'' = d" by auto
- hence [simp]: "?A \<noteq> {}" by blast
+ let ?A = "{a..b} \<inter> g -` {c..d}"
+ from IVT'[of g a c b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5)
+ obtain c'' where c'': "c'' \<in> ?A" "g c'' = c" by auto
+ from IVT'[of g a d b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5)
+ obtain d'' where d'': "d'' \<in> ?A" "g d'' = d" by auto
+ hence [simp]: "?A \<noteq> {}" by blast
- def c' \<equiv> "Inf ?A" and d' \<equiv> "Sup ?A"
- have "?A \<subseteq> {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' \<in> ?A" "d' \<in> ?A" unfolding c'_def d'_def
- by ((intro closed_contains_Inf closed_contains_Sup, simp_all)[])+
- hence "{c'..d'} \<subseteq> ?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' \<le> d'" using c'd'_in_set(2) unfolding c'_def by (intro cInf_lower) auto
- moreover have "g c' \<le> c" "g d' \<ge> 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 \<subseteq> {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' \<in> ?A" "d' \<in> ?A" unfolding c'_def d'_def
+ by ((intro closed_contains_Inf closed_contains_Sup, simp_all)[])+
+ hence "{c'..d'} \<subseteq> ?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' \<le> d'" using c'd'_in_set(2) unfolding c'_def by (intro cInf_lower) auto
+ moreover have "g c' \<le> c" "g d' \<ge> 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 \<open>Generic Borel spaces\<close>
@@ -540,7 +541,7 @@
by (auto simp: topological_basis_def)
from B(2)[OF K] obtain m where m: "\<And>k. k \<in> K \<Longrightarrow> m k \<subseteq> B" "\<And>k. k \<in> K \<Longrightarrow> (\<Union>m k) = k"
by metis
- def U \<equiv> "(\<Union>k\<in>K. m k)"
+ define U where "U = (\<Union>k\<in>K. m k)"
with m have "countable U"
by (intro countable_subset[OF _ \<open>countable B\<close>]) auto
have "\<Union>U = (\<Union>A\<in>U. A)" by simp
@@ -1811,7 +1812,7 @@
assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M"
proof -
- def u' \<equiv> "\<lambda>x. lim (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)"
+ define u' where "u' x = lim (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)" for x
then have *: "\<And>x. lim (\<lambda>i. f i x) = (if Cauchy (\<lambda>i. f i x) then u' x else (THE x. False))"
by (auto simp: lim_def convergent_eq_cauchy[symmetric])
have "u' \<in> borel_measurable M"
--- a/src/HOL/Probability/Caratheodory.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Caratheodory.thy Mon Apr 25 16:09:26 2016 +0200
@@ -372,7 +372,7 @@
and Ble: "\<And>n. (\<Sum>i. f (B n i)) \<le> ?O (A n) + e * (1/2)^(Suc n)"
by (metis less_imp_le outer_measure_close[OF *])
- def C \<equiv> "case_prod B o prod_decode"
+ define C where "C = case_prod B o prod_decode"
from B have B_in_M: "\<And>i j. B i j \<in> M"
by (rule range_subsetD)
then have C: "range C \<subseteq> 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 \<equiv> "lambda_system \<Omega> (Pow \<Omega>) ?O"
+ define ls where "ls = lambda_system \<Omega> (Pow \<Omega>) ?O"
have mls: "measure_space \<Omega> 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 \<inter> F' j = {}"
by auto }
note F'_disj = this
- def F \<equiv> "\<lambda>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 "(\<Union>i. F i) = \<Union>C"
@@ -704,7 +704,7 @@
from generated_ringE[OF Un_A] guess C' . note C' = this
{ fix c assume "c \<in> C'"
- moreover def A \<equiv> "\<lambda>i. A' i \<inter> c"
+ moreover define A where [abs_def]: "A i = A' i \<inter> c" for i
ultimately have A: "range A \<subseteq> generated_ring" "disjoint_family A"
and Un_A: "(\<Union>i. A i) \<in> generated_ring"
using A' C'
@@ -722,7 +722,7 @@
have "\<exists>F'. bij_betw F' {..<card C} C"
by (rule finite_same_card_bij[OF _ \<open>finite C\<close>]) auto
then guess F .. note F = this
- def f \<equiv> "\<lambda>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 "\<forall>j. f j \<in> M"
--- a/src/HOL/Probability/Central_Limit_Theorem.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Central_Limit_Theorem.thy Mon Apr 25 16:09:26 2016 +0200
@@ -24,8 +24,8 @@
shows "weak_conv_m (\<lambda>n. distr M borel (\<lambda>x. S n x / sqrt (n * \<sigma>\<^sup>2))) std_normal_distribution"
proof -
let ?S' = "\<lambda>n x. S n x / sqrt (real n * \<sigma>\<^sup>2)"
- def \<phi> \<equiv> "\<lambda>n. char (distr M borel (?S' n))"
- def \<psi> \<equiv> "\<lambda>n t. char \<mu> (t / sqrt (\<sigma>\<^sup>2 * n))"
+ define \<phi> where "\<phi> n = char (distr M borel (?S' n))" for n
+ define \<psi> where "\<psi> n t = char \<mu> (t / sqrt (\<sigma>\<^sup>2 * n))" for n t
have X_rv [simp, measurable]: "\<And>n. random_variable borel (X n)"
using X_indep unfolding indep_vars_def2 by simp
@@ -48,7 +48,7 @@
hence n: "n \<ge> t^2 / 4" by (subst nat_ceiling_le_eq [symmetric])
let ?t = "t / sqrt (\<sigma>\<^sup>2 * n)"
- def \<psi>' \<equiv> "\<lambda>n i. char (distr M borel (\<lambda>x. X i x / sqrt (\<sigma>\<^sup>2 * n)))"
+ define \<psi>' where "\<psi>' n i = char (distr M borel (\<lambda>x. X i x / sqrt (\<sigma>\<^sup>2 * n)))" for n i
have *: "\<And>n i t. \<psi>' n i t = \<psi> n t"
unfolding \<psi>_def \<psi>'_def char_def
by (subst X_distrib [symmetric]) (auto simp: integral_distr)
--- a/src/HOL/Probability/Characteristic_Functions.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Characteristic_Functions.thy Mon Apr 25 16:09:26 2016 +0200
@@ -308,7 +308,7 @@
have integ_iexp: "integrable M (\<lambda>x. iexp (t * x))"
by (intro integrable_const_bound) auto
- def c \<equiv> "\<lambda>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: "\<And>k. k \<le> n \<Longrightarrow> integrable M (\<lambda>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 (\<lambda>x. iexp (t * x))"
by (intro integrable_const_bound) auto
- def c \<equiv> "\<lambda>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: "\<And>k. k \<le> n \<Longrightarrow> integrable M (\<lambda>x. c k x)"
unfolding c_def by (intro integrable_mult_right integrable_of_real integrable_moments)
--- a/src/HOL/Probability/Distributions.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Distributions.thy Mon Apr 25 16:09:26 2016 +0200
@@ -1216,7 +1216,7 @@
shows "(\<lambda>x. \<integral>\<^sup>+y. ennreal (normal_density 0 \<sigma> (x - y) * normal_density 0 \<tau> y) \<partial>lborel) =
normal_density 0 (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2))" (is "?LHS = ?RHS")
proof -
- def \<sigma>' \<equiv> "\<sigma>\<^sup>2" and \<tau>' \<equiv> "\<tau>\<^sup>2"
+ define \<sigma>' \<tau>' where "\<sigma>' = \<sigma>\<^sup>2" and "\<tau>' = \<tau>\<^sup>2"
then have [simp, arith]: "0 < \<sigma>'" "0 < \<tau>'"
by simp_all
let ?\<sigma> = "sqrt ((\<sigma>' * \<tau>') / (\<sigma>' + \<tau>'))"
--- a/src/HOL/Probability/Fin_Map.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Fin_Map.thy Mon Apr 25 16:09:26 2016 +0200
@@ -467,12 +467,12 @@
assume "Cauchy P"
then obtain Nd where Nd: "\<And>n. n \<ge> Nd \<Longrightarrow> dist (P n) (P Nd) < 1"
by (force simp: cauchy)
- def d \<equiv> "domain (P Nd)"
+ define d where "d = domain (P Nd)"
with Nd have dim: "\<And>n. n \<ge> Nd \<Longrightarrow> domain (P n) = d" using dist_le_1_imp_domain_eq by auto
have [simp]: "finite d" unfolding d_def by simp
- def p \<equiv> "\<lambda>i n. (P n) i"
- def q \<equiv> "\<lambda>i. lim (p i)"
- def Q \<equiv> "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: "\<And>i. i \<in> d \<Longrightarrow> q i = Q i" by (auto simp add: Q_def Abs_finmap_inverse)
{
fix i assume "i \<in> d"
@@ -505,7 +505,7 @@
from p[OF \<open>i \<in> d\<close>, THEN metric_LIMSEQ_D, OF \<open>0 < e\<close>]
show "\<exists>no. \<forall>n\<ge>no. dist (p i n) (q i) < e" .
qed then guess ni .. note ni = this
- def N \<equiv> "max Nd (Max (ni ` d))"
+ define N where "N = max Nd (Max (ni ` d))"
show "\<exists>N. \<forall>n\<ge>N. dist (P n) Q < e"
proof (safe intro!: exI[where x="N"])
fix n assume "N \<le> n"
@@ -602,7 +602,7 @@
thus "\<exists>y. x i \<in> y \<and> y \<subseteq> a i \<and> y \<in> basis_proj" by auto
qed
then guess B .. note B = this
- def B' \<equiv> "Pi' (domain x) (\<lambda>i. (B i)::'b set)"
+ define B' where "B' = Pi' (domain x) (\<lambda>i. (B i)::'b set)"
have "B' \<subseteq> Pi' (domain x) a" using B by (auto intro!: Pi'_mono simp: B'_def)
also note \<open>\<dots> \<subseteq> O'\<close>
finally show "\<exists>B'\<in>basis_finmap. x \<in> B' \<and> B' \<subseteq> O'" using B
@@ -1097,7 +1097,7 @@
proof (rule sigma_fprod_algebra_sigma_eq)
show "finite I" by simp
show "I \<noteq> {}" by fact
- def S'\<equiv>"from_nat_into S"
+ define S' where "S' = from_nat_into S"
show "(\<Union>j. S' j) = space borel"
using S
apply (auto simp add: from_nat_into countable_basis_proj S'_def basis_proj_def)
--- a/src/HOL/Probability/Finite_Product_Measure.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Finite_Product_Measure.thy Mon Apr 25 16:09:26 2016 +0200
@@ -440,7 +440,7 @@
obtain S where S: "\<And>i. i \<in> j \<Longrightarrow> S i \<subseteq> E i" "\<And>i. i \<in> j \<Longrightarrow> countable (S i)"
"\<And>i. i \<in> j \<Longrightarrow> \<Omega> i = \<Union>(S i)"
by (metis subset_eq \<Omega>_cover \<open>j \<subseteq> I\<close>)
- def A' \<equiv> "\<lambda>n. n(i := A)"
+ define A' where "A' n = n(i := A)" for n
then have A'_i: "\<And>n. A' n i = A"
by simp
{ fix n assume "n \<in> 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 \<equiv> "SOME i. i \<in> I"
+ define i where "i = (SOME i. i \<in> I)"
have i: "I \<noteq> {} \<Longrightarrow> i \<in> I"
unfolding i_def by (rule someI_ex) auto
- def A \<equiv> "\<lambda>n::nat. if I = {} then prod_emb I M {} (\<Pi>\<^sub>E i\<in>{}. {}) else prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))"
+ define A where "A n =
+ (if I = {} then prod_emb I M {} (\<Pi>\<^sub>E i\<in>{}. {}) else prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i)))"
+ for n :: nat
then show "range A \<subseteq> prod_algebra I M"
using prod_algebraI[of "{}" I "\<lambda>i. space (M i)" M] by (auto intro!: prod_algebraI i)
have "\<And>i. A i = space (PiM I M)"
@@ -954,7 +956,7 @@
proof (rule measure_eqI_PiM_finite[OF I refl P, symmetric])
show "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> (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 \<equiv> "\<lambda>n. \<Pi>\<^sub>E i\<in>I. C i n"
+ define A where "A n = (\<Pi>\<^sub>E i\<in>I. C i n)" for n
with C show "range A \<subseteq> prod_algebra I M" "\<And>i. emeasure (Pi\<^sub>M I M) (A i) \<noteq> \<infinity>" "(\<Union>i. A i) = space (PiM I M)"
by (auto intro!: prod_algebraI_finite simp: emeasure_PiM subset_eq ennreal_setprod_eq_top)
qed
--- a/src/HOL/Probability/Giry_Monad.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Giry_Monad.thy Mon Apr 25 16:09:26 2016 +0200
@@ -349,7 +349,8 @@
have [measurable]: "F i \<in> N \<rightarrow>\<^sub>M count_space UNIV" for i
using F(1) by (rule measurable_simple_function)
- def F' \<equiv> "\<lambda>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 "(\<lambda>M. F' M i) \<in> subprob_algebra N \<rightarrow>\<^sub>M borel" for i
proof (rule measurable_cong[THEN iffD2])
--- a/src/HOL/Probability/Helly_Selection.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Helly_Selection.thy Mon Apr 25 16:09:26 2016 +0200
@@ -43,7 +43,7 @@
thus "\<exists>s'. subseq s' \<and> (\<exists>l. (\<lambda>k. f (s (s' k)) (r n)) \<longlonglongrightarrow> l)"
by (auto simp: comp_def)
qed
- def d \<equiv> "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} \<noteq> {}" for x
by auto
- def F \<equiv> "\<lambda>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 \<mu>: "tight \<mu>" "subseq s"
shows "\<exists>r M. subseq r \<and> real_distribution M \<and> weak_conv_m (\<mu> \<circ> s \<circ> r) M"
proof -
- def f \<equiv> "\<lambda>k. cdf (\<mu> (s k))"
+ define f where "f k = cdf (\<mu> (s k))" for k
interpret \<mu>: real_distribution "\<mu> k" for k
using \<mu> unfolding tight_def by auto
@@ -275,7 +275,8 @@
subseq: "\<And>s \<nu>. subseq s \<Longrightarrow> real_distribution \<nu> \<Longrightarrow> weak_conv_m (\<mu> \<circ> s) \<nu> \<Longrightarrow> weak_conv_m (\<mu> \<circ> s) M"
shows "weak_conv_m \<mu> M"
proof (rule ccontr)
- def f \<equiv> "\<lambda>n. cdf (\<mu> n)" and F \<equiv> "cdf M"
+ define f where "f n = cdf (\<mu> n)" for n
+ define F where "F = cdf M"
assume "\<not> weak_conv_m \<mu> M"
then obtain x where x: "isCont F x" "\<not> (\<lambda>n. f n x) \<longlonglongrightarrow> F x"
--- a/src/HOL/Probability/Independent_Family.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Independent_Family.thy Mon Apr 25 16:09:26 2016 +0200
@@ -132,7 +132,7 @@
then have "indep_sets (?G J) K"
proof induct
case (insert j J)
- moreover def G \<equiv> "?G J"
+ moreover define G where "G = ?G J"
ultimately have G: "indep_sets G K" "\<And>i. i \<in> K \<Longrightarrow> G i \<subseteq> events" and "j \<in> K"
by (auto simp: indep_sets_def)
let ?D = "{E\<in>events. indep_sets (G(j := {E})) K }"
@@ -464,7 +464,7 @@
qed }
note L_inj = this
- def k \<equiv> "\<lambda>l. (SOME k. k \<in> K \<and> l \<in> L k)"
+ define k where "k l = (SOME k. k \<in> K \<and> l \<in> L k)" for l
{ fix x j l assume *: "j \<in> K" "l \<in> L j"
have "k l = j" unfolding k_def
proof (rule some_equality)
@@ -1265,7 +1265,7 @@
shows "(\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. X i \<omega>) \<partial>M) = (\<Prod>i\<in>I. \<integral>\<^sup>+\<omega>. X i \<omega> \<partial>M)"
proof cases
assume "I \<noteq> {}"
- def Y \<equiv> "\<lambda>i \<omega>. if i \<in> I then X i \<omega> else 0"
+ define Y where [abs_def]: "Y i \<omega> = (if i \<in> I then X i \<omega> else 0)" for i \<omega>
{ fix i have "i \<in> I \<Longrightarrow> random_variable borel (X i)"
using I(2) by (cases "i\<in>I") (auto simp: indep_vars_def) }
note rv_X = this
@@ -1302,7 +1302,7 @@
and indep_vars_integrable: "integrable M (\<lambda>\<omega>. (\<Prod>i\<in>I. X i \<omega>))" (is ?int)
proof (induct rule: case_split)
assume "I \<noteq> {}"
- def Y \<equiv> "\<lambda>i \<omega>. if i \<in> I then X i \<omega> else 0"
+ define Y where [abs_def]: "Y i \<omega> = (if i \<in> I then X i \<omega> else 0)" for i \<omega>
{ fix i have "i \<in> I \<Longrightarrow> random_variable borel (X i)"
using I(2) by (cases "i\<in>I") (auto simp: indep_vars_def) }
note rv_X = this[measurable]
--- a/src/HOL/Probability/Information.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Information.thy Mon Apr 25 16:09:26 2016 +0200
@@ -1671,9 +1671,14 @@
assumes X: "simple_function M X" and Y: "simple_function M Y"
shows "\<I>(X ; X | Y) = \<H>(X | Y)"
proof -
- def Py \<equiv> "\<lambda>x. if x \<in> Y`space M then measure M (Y -` {x} \<inter> space M) else 0"
- def Pxy \<equiv> "\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x))`space M then measure M ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M) else 0"
- def Pxxy \<equiv> "\<lambda>x. if x \<in> (\<lambda>x. (X x, X x, Y x))`space M then measure M ((\<lambda>x. (X x, X x, Y x)) -` {x} \<inter> space M) else 0"
+ define Py where "Py x = (if x \<in> Y`space M then measure M (Y -` {x} \<inter> space M) else 0)" for x
+ define Pxy where "Pxy x =
+ (if x \<in> (\<lambda>x. (X x, Y x))`space M then measure M ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M) else 0)"
+ for x
+ define Pxxy where "Pxxy x =
+ (if x \<in> (\<lambda>x. (X x, X x, Y x))`space M then measure M ((\<lambda>x. (X x, X x, Y x)) -` {x} \<inter> space M)
+ else 0)"
+ for x
let ?M = "X`space M \<times> X`space M \<times> Y`space M"
note XY = simple_function_Pair[OF X Y]
--- a/src/HOL/Probability/Interval_Integral.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Interval_Integral.thy Mon Apr 25 16:09:26 2016 +0200
@@ -81,7 +81,7 @@
(auto simp: incseq_def PInf)
next
case (real b')
- def d \<equiv> "b' - (if a = -\<infinity> then b' - 1 else real_of_ereal a)"
+ define d where "d = b' - (if a = -\<infinity> then b' - 1 else real_of_ereal a)"
with \<open>a < b\<close> have a': "0 < d"
by (cases a) (auto simp: real)
moreover
--- a/src/HOL/Probability/Lebesgue_Measure.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Lebesgue_Measure.thy Mon Apr 25 16:09:26 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' \<equiv> "{i. l i < r i}"
+ define S' where "S' = {i. l i < r i}"
obtain S :: "nat set" where
"S \<subseteq> S'" and finS: "finite S" and
Sprop: "{a'..b} \<subseteq> (\<Union>i \<in> S. {l i<..<r i + delta i})"
--- a/src/HOL/Probability/Measurable.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Measurable.thy Mon Apr 25 16:09:26 2016 +0200
@@ -559,7 +559,8 @@
unfolding measurable_def
proof safe
fix X
- def f \<equiv> "\<lambda>x. THE i. P i x" def undef \<equiv> "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 \<in> space M" "P i x" then have "f x = i"
unfolding f_def using unique by auto }
note f_eq = this
--- a/src/HOL/Probability/Measure_Space.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Measure_Space.thy Mon Apr 25 16:09:26 2016 +0200
@@ -567,7 +567,7 @@
by simp
next
assume "infinite I"
- def L \<equiv> "\<lambda>n. LEAST i. i \<in> I \<and> i \<ge> n"
+ define L where "L n = (LEAST i. i \<in> I \<and> i \<ge> n)" for n
have L: "L n \<in> I \<and> n \<le> L n" for n
unfolding L_def
proof (rule LeastI_ex)
--- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Mon Apr 25 16:09:26 2016 +0200
@@ -210,7 +210,8 @@
assumes u[measurable]: "u \<in> borel_measurable M"
shows "\<exists>f. incseq f \<and> (\<forall>i. (\<forall>x. f i x < top) \<and> simple_function M (f i)) \<and> u = (SUP i. f i)"
proof -
- def f \<equiv> "\<lambda>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 \<le> 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 \<inter> space M \<in> sets M"
shows "simple_function M (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function M ?IF")
proof -
- def F \<equiv> "\<lambda>x. f -` {x} \<inter> space M" and G \<equiv> "\<lambda>x. g -` {x} \<inter> space M"
+ define F where "F x = f -` {x} \<inter> space M" for x
+ define G where "G x = g -` {x} \<inter> space M" for x
show ?thesis unfolding simple_function_def
proof safe
have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto
@@ -1452,7 +1454,7 @@
assumes "f \<in> measurable M (count_space UNIV)"
shows "(\<integral>\<^sup>+x. of_nat (f x) \<partial>M) = (\<Sum>t. emeasure M {x\<in>space M. t < f x})"
proof -
- def F \<equiv> "\<lambda>i. {x\<in>space M. i < f x}"
+ define F where "F i = {x\<in>space M. i < f x}" for i
with assms have [measurable]: "\<And>i. F i \<in> sets M"
by auto
@@ -1637,7 +1639,7 @@
assumes f: "f \<in> measurable M (count_space UNIV)"
shows "(\<integral>\<^sup>+ x. ennreal_of_enat (f x) \<partial>M) = (\<Sum>t. emeasure M {x \<in> space M. t < f x})"
proof -
- def F \<equiv> "\<lambda>i::nat. {x\<in>space M. i < f x}"
+ define F where "F i = {x\<in>space M. i < f x}" for i :: nat
with assms have [measurable]: "\<And>i. F i \<in> sets M"
by auto
--- a/src/HOL/Probability/Probability_Mass_Function.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Probability_Mass_Function.thy Mon Apr 25 16:09:26 2016 +0200
@@ -1138,7 +1138,9 @@
from qr obtain qr where qr: "\<And>y z. (y, z) \<in> set_pmf qr \<Longrightarrow> S y z"
and q': "q = map_pmf fst qr" and r: "r = map_pmf snd qr" by cases auto
- def pr \<equiv> "bind_pmf pq (\<lambda>xy. bind_pmf (cond_pmf qr {yz. fst yz = snd xy}) (\<lambda>yz. return_pmf (fst xy, snd yz)))"
+ define pr where "pr =
+ bind_pmf pq (\<lambda>xy. bind_pmf (cond_pmf qr {yz. fst yz = snd xy})
+ (\<lambda>yz. return_pmf (fst xy, snd yz)))"
have pr_welldefined: "\<And>y. y \<in> q \<Longrightarrow> qr \<inter> {yz. fst yz = y} \<noteq> {}"
by (force simp: q')
--- a/src/HOL/Probability/Probability_Measure.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Probability_Measure.thy Mon Apr 25 16:09:26 2016 +0200
@@ -1147,7 +1147,7 @@
show "Int_stable (range atMost :: real set set)"
by (auto simp: Int_stable_def)
have vimage_eq: "\<And>a. (X -` {..a} \<inter> space M) = {x\<in>space M. X x \<le> a}" by auto
- def A \<equiv> "\<lambda>i::nat. {.. real i}"
+ define A where "A i = {.. real i}" for i :: nat
then show "range A \<subseteq> range atMost" "(\<Union>i. A i) = space lborel"
"\<And>i. emeasure (distr M lborel X) (A i) \<noteq> \<infinity>"
by (auto simp: real_arch_simple emeasure_distr vimage_eq M_eq)
--- a/src/HOL/Probability/Projective_Family.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Projective_Family.thy Mon Apr 25 16:09:26 2016 +0200
@@ -441,7 +441,8 @@
proof (rule PF.emeasure_lim[OF J subset_UNIV X])
fix J X' assume J[simp]: "\<And>i. finite (J i)" and X'[measurable]: "\<And>i. X' i \<in> sets (Pi\<^sub>M (J i) M)"
and dec: "decseq (\<lambda>i. PF.emb UNIV (J i) (X' i))"
- def X \<equiv> "\<lambda>n. (\<Inter>i\<in>{i. J i \<subseteq> {0..< n}}. PF.emb {0..<n} (J i) (X' i)) \<inter> space (PiM {0..<n} M)"
+ define X where "X n =
+ (\<Inter>i\<in>{i. J i \<subseteq> {0..< n}}. PF.emb {0..<n} (J i) (X' i)) \<inter> space (PiM {0..<n} M)" for n
have sets_X[measurable]: "X n \<in> sets (PiM {0..<n} M)" for n
by (cases "{i. J i \<subseteq> {0..< n}} = {}")
@@ -605,7 +606,8 @@
assume inf: "infinite (\<Union>i. J i)"
moreover have count: "countable (\<Union>i. J i)"
using 1(3) by (auto intro: countable_finite)
- def f \<equiv> "from_nat_into (\<Union>i. J i)" and t \<equiv> "to_nat_on (\<Union>i. J i)"
+ define f where "f = from_nat_into (\<Union>i. J i)"
+ define t where "t = to_nat_on (\<Union>i. J i)"
have ft[simp]: "x \<in> J i \<Longrightarrow> 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
--- a/src/HOL/Probability/Projective_Limit.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Projective_Limit.thy Mon Apr 25 16:09:26 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 \<le> 1"
by (cases "?a") (auto simp: top_unique)
- def Z \<equiv> "\<lambda>n. emb I (J n) (B n)"
+ define Z where "Z n = emb I (J n) (B n)" for n
have Z_mono: "n \<le> m \<Longrightarrow> Z m \<subseteq> Z n" for n m
unfolding Z_def using B[THEN antimonoD, of n m] .
have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"
@@ -133,13 +133,13 @@
unfolding Z_def by (auto intro!: generator.intros J)
have countable_UN_J: "countable (\<Union>n. J n)" by (simp add: countable_finite)
- def Utn \<equiv> "to_nat_on (\<Union>n. J n)"
+ define Utn where "Utn = to_nat_on (\<Union>n. J n)"
interpret function_to_finmap "J n" Utn "from_nat_into (\<Union>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 (\<Union>n. J n)"
unfolding Utn_def using countable_UN_J by (rule inj_on_to_nat_on)
hence inj_on_Utn_J: "\<And>n. inj_on Utn (J n)" by (rule subset_inj_on) auto
- def P' \<equiv> "\<lambda>n. mapmeasure n (P (J n)) (\<lambda>_. borel)"
+ define P' where "P' n = mapmeasure n (P (J n)) (\<lambda>_. 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 @@
"\<And>n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) \<le> ennreal (2 powr - real n) * ?a"
"\<And>n. compact (K' n)" "\<And>n. K' n \<subseteq> fm n ` B n"
unfolding choice_iff by blast
- def K \<equiv> "\<lambda>n. fm n -` K' n \<inter> space (Pi\<^sub>M (J n) (\<lambda>_. borel))"
+ define K where "K n = fm n -` K' n \<inter> space (Pi\<^sub>M (J n) (\<lambda>_. borel))" for n
have K_sets: "\<And>n. K n \<in> sets (Pi\<^sub>M (J n) (\<lambda>_. borel))"
unfolding K_def
using compact_imp_closed[OF \<open>compact (K' _)\<close>]
@@ -204,7 +204,7 @@
using \<open>x \<in> K n\<close> 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' \<equiv> "\<lambda>n. emb I (J n) (K n)"
+ define Z' where "Z' n = emb I (J n) (K n)" for n
have Z': "\<And>n. Z' n \<subseteq> Z n"
unfolding Z'_def Z_def
proof (rule prod_emb_mono, safe)
@@ -228,7 +228,7 @@
have "\<And>n. Z' n \<in> 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 \<equiv> "\<lambda>n. \<Inter>i\<in>{1..n}. Z' i"
+ define Y where "Y n = (\<Inter>i\<in>{1..n}. Z' i)" for n
hence "\<And>n k. Y (n + k) \<subseteq> Y n" by (induct_tac k) (auto simp: Y_def)
hence Y_mono: "\<And>n m. n \<le> m \<Longrightarrow> Y m \<subseteq> Y n" by (auto simp: le_iff_add)
have Y_Z': "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z' n" by (auto simp: Y_def)
--- a/src/HOL/Probability/Radon_Nikodym.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Radon_Nikodym.thy Mon Apr 25 16:09:26 2016 +0200
@@ -153,7 +153,7 @@
let ?A = "\<lambda>A. if (\<forall>B\<in>sets M. B \<subseteq> space M - A \<longrightarrow> -e < ?d B)
then {}
else (SOME B. B \<in> sets M \<and> B \<subseteq> space M - A \<and> ?d B \<le> -e)"
- def A \<equiv> "\<lambda>n. ((\<lambda>B. B \<union> ?A B) ^^ n) {}"
+ define A where "A n = ((\<lambda>B. B \<union> ?A B) ^^ n) {}" for n
have A_simps[simp]:
"A 0 = {}"
"\<And>n. A (Suc n) = (A n \<union> ?A (A n))" unfolding A_def by simp_all
@@ -296,7 +296,8 @@
shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> density M f = N"
proof -
interpret N: finite_measure N by fact
- def G \<equiv> "{g \<in> borel_measurable M. (\<forall>x. 0 \<le> g x) \<and> (\<forall>A\<in>sets M. (\<integral>\<^sup>+x. g x * indicator A x \<partial>M) \<le> N A)}"
+ define G where "G =
+ {g \<in> borel_measurable M. (\<forall>x. 0 \<le> g x) \<and> (\<forall>A\<in>sets M. (\<integral>\<^sup>+x. g x * indicator A x \<partial>M) \<le> N A)}"
{ fix f have "f \<in> G \<Longrightarrow> f \<in> borel_measurable M" by (auto simp: G_def) }
note this[measurable_dest]
have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto
@@ -366,7 +367,7 @@
from choice[OF this] obtain gs where "\<And>i. gs i \<in> G" "\<And>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 = "\<lambda>i x. Max ((\<lambda>n. gs n x) ` {..i})"
- def f \<equiv> "\<lambda>x. SUP i. ?g i x"
+ define f where [abs_def]: "f x = (SUP i. ?g i x)" for x
let ?F = "\<lambda>A x. f x * indicator A x"
have gs_not_empty: "\<And>i x. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto
{ fix i have "?g i \<in> G"
@@ -443,7 +444,7 @@
hence "(\<integral>\<^sup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<infinity>"
using M'.finite_emeasure_space by (auto simp: top_unique)
moreover
- def b \<equiv> "?M (space M) / emeasure M (space M) / 2"
+ define b where "b = ?M (space M) / emeasure M (space M) / 2"
ultimately have b: "b \<noteq> 0 \<and> 0 \<le> b \<and> b \<noteq> \<infinity>"
by (auto simp: ennreal_divide_eq_top_iff)
then have b: "b \<noteq> 0" "0 \<le> b" "0 < b" "b \<noteq> \<infinity>"
@@ -584,7 +585,7 @@
qed
let ?O_0 = "(\<Union>i. ?O i)"
have "?O_0 \<in> sets M" using Q' by auto
- def Q \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> Q' 0 | Suc n \<Rightarrow> ?O (Suc n) - ?O n"
+ define Q where "Q i = (case i of 0 \<Rightarrow> Q' 0 | Suc n \<Rightarrow> ?O (Suc n) - ?O n)" for i
{ fix i have "Q i \<in> 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 \<noteq> \<infinity>"
from sigma_finite guess Q . note Q = this
- def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. ennreal(of_nat (Suc n))}) \<inter> space M"
+ define A where "A i =
+ f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. ennreal(of_nat (Suc n))}) \<inter> space M" for i
{ fix i j have "A i \<inter> Q j \<in> sets M"
unfolding A_def using f Q
apply (rule_tac sets.Int)
--- a/src/HOL/Probability/Regularity.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Regularity.thy Mon Apr 25 16:09:26 2016 +0200
@@ -76,7 +76,8 @@
"\<exists>K \<in> {K. K \<subseteq> space M \<and> compact K}. emeasure M (space M) \<le> emeasure M K + ennreal e"
(is "?thesis e") if "0 < e" for e :: real
proof -
- def B \<equiv> "\<lambda>n. \<Union>i\<in>{0..k e (Suc n)}. cball (from_nat_into X i) (1 / Suc n)"
+ define B where [abs_def]:
+ "B n = (\<Union>i\<in>{0..k e (Suc n)}. cball (from_nat_into X i) (1 / Suc n))" for n
have "\<And>n. closed (B n)" by (auto simp: B_def)
hence [simp]: "\<And>n. B n \<in> sets M" by (simp add: sb)
from k[OF \<open>e > 0\<close> zero_less_Suc]
@@ -84,7 +85,7 @@
by (simp add: algebra_simps B_def finite_measure_compl)
hence B_compl_le: "\<And>n::nat. measure M (space M - B n) \<le> e * 2 powr - real (Suc n)"
by (simp add: finite_measure_compl)
- def K \<equiv> "\<Inter>n. B n"
+ define K where "K = (\<Inter>n. B n)"
from \<open>closed (B _)\<close> have "closed K" by (auto simp: K_def)
hence [simp]: "K \<in> sets M" by (simp add: sb)
have "measure M (space M) - measure M K = measure M (space M - K)"
--- a/src/HOL/Probability/Sigma_Algebra.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Sigma_Algebra.thy Mon Apr 25 16:09:26 2016 +0200
@@ -846,7 +846,7 @@
proof -
from a guess Ca .. note Ca = this
from b guess Cb .. note Cb = this
- def C \<equiv> "(\<lambda>(a,b). a \<inter> b)` (Ca\<times>Cb)"
+ define C where "C = (\<lambda>(a,b). a \<inter> b)` (Ca\<times>Cb)"
show ?thesis
proof
show "disjoint C"
@@ -1971,7 +1971,7 @@
with \<open>i\<in>I\<close> * show ?thesis
by simp
next
- def P \<equiv> "\<lambda>\<mu>'. (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>'"
+ define P where "P \<mu>' \<longleftrightarrow> (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>'" for \<mu>'
assume "\<not> (\<forall>i\<in>I. \<mu> i = 0)"
moreover
have "measure_space (space M) (sets M) \<mu>'"
--- a/src/HOL/Probability/Weak_Convergence.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/Weak_Convergence.thy Mon Apr 25 16:09:26 2016 +0200
@@ -212,11 +212,11 @@
using emeasure_lborel_countable[OF D_countable]
by (subst emeasure_restrict_space) auto
- def Y' \<equiv> "\<lambda>\<omega>. if \<omega> \<in> ?D then 0 else M.I \<omega>"
+ define Y' where "Y' \<omega> = (if \<omega> \<in> ?D then 0 else M.I \<omega>)" for \<omega>
have Y'_AE: "AE \<omega> in ?\<Omega>. Y' \<omega> = M.I \<omega>"
by (rule AE_I [OF _ D]) (auto simp: space_restrict_space sets_restrict_space_iff Y'_def)
- def Y_seq' \<equiv> "\<lambda>n \<omega>. if \<omega> \<in> ?D then 0 else \<mu>.I n \<omega>"
+ define Y_seq' where "Y_seq' n \<omega> = (if \<omega> \<in> ?D then 0 else \<mu>.I n \<omega>)" for n \<omega>
have Y_seq'_AE: "\<And>n. AE \<omega> in ?\<Omega>. Y_seq' n \<omega> = \<mu>.I n \<omega>"
by (rule AE_I [OF _ D]) (auto simp: space_restrict_space sets_restrict_space_iff Y_seq'_def)
--- a/src/HOL/Probability/ex/Dining_Cryptographers.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy Mon Apr 25 16:09:26 2016 +0200
@@ -64,7 +64,7 @@
have foldl_coin:
"\<not> ?XOR (\<lambda>c. coin dc c \<noteq> coin dc (c + 1)) n"
proof -
- def n' \<equiv> n \<comment> "Need to hide n, as it is hidden in coin"
+ define n' where "n' = n" \<comment> "Need to hide n, as it is hidden in coin"
have "?XOR (\<lambda>c. coin dc c \<noteq> coin dc (c + 1)) n'
= (coin dc 0 \<noteq> coin dc n')"
by (induct n') auto
@@ -81,7 +81,7 @@
next
assume "\<exists>k<n. payer dc = Some k"
then obtain k where "k < n" and "payer dc = Some k" by auto
- def l \<equiv> n \<comment> "Need to hide n, as it is hidden in coin, payer etc."
+ define l where "l = n" \<comment> "Need to hide n, as it is hidden in coin, payer etc."
have "?XOR (\<lambda>c. (payer dc = Some c) \<noteq> (coin dc c \<noteq> coin dc (c + 1))) l =
((k < l) \<noteq> ?XOR (\<lambda>c. (coin dc c \<noteq> coin dc (c + 1))) l)"
using \<open>payer dc = Some k\<close> by (induct l) auto
@@ -156,7 +156,7 @@
@{term i}.
\<close>
- def zs \<equiv> "map (\<lambda>p. if p \<in> {min i j<..max i j} then \<not> ys ! p else ys ! p) [0..<n]"
+ define zs where "zs = map (\<lambda>p. if p \<in> {min i j<..max i j} then \<not> ys ! p else ys ! p) [0..<n]"
hence [simp]: "length zs = n" by simp
hence [simp]: "0 < length zs" using n_gt_3 by simp
@@ -458,7 +458,7 @@
by (simp add: image_iff space_uniform_count_measure dc_crypto Bex_def) blast
then have xs: "inversion (Some i, xs) \<in> inversion`dc_crypto" and i: "Some i \<in> Some ` {0..<n}"
and x: "x = (inversion (Some i, xs), Some i)" by (simp_all add: payer_def dc_crypto)
- moreover def ys \<equiv> "inversion (Some i, xs)"
+ moreover define ys where "ys = inversion (Some i, xs)"
ultimately have ys: "ys \<in> inversion`dc_crypto"
and "Some i \<in> Some ` {0..<n}" "x = (ys, Some i)" by simp_all
then have "(\<lambda>x. (inversion x, payer x)) -` {x} \<inter> space (uniform_count_measure dc_crypto) =
--- a/src/HOL/Probability/ex/Measure_Not_CCC.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Probability/ex/Measure_Not_CCC.thy Mon Apr 25 16:09:26 2016 +0200
@@ -106,7 +106,7 @@
unfolding fun_eq_iff using Union.IH by metis
show ?case
proof (intro exI conjI)
- def G \<equiv> "\<lambda>j. (\<Union>i. if j \<in> J i then F i j else X i)"
+ define G where "G j = (\<Union>i. if j \<in> J i then F i j else X i)" for j
show "(\<Union>i. X i) \<in> sets M" "countable (\<Union>i. J i)" "G \<in> (\<Union>i. J i) \<rightarrow> sets M"
using XFJ by (auto simp: G_def Pi_iff)
show "UNION UNIV A = (UNIV - (\<Union>i. J i)) \<times> (\<Union>i. X i) \<union> (SIGMA j:\<Union>i. J i. \<Union>i. if j \<in> J i then F i j else X i)"
--- a/src/HOL/Real.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Real.thy Mon Apr 25 16:09:26 2016 +0200
@@ -775,7 +775,7 @@
obtain x where x: "x \<in> S" using assms(1) ..
obtain z where z: "\<forall>x\<in>S. x \<le> z" using assms(2) ..
- def P \<equiv> "\<lambda>x. \<forall>y\<in>S. y \<le> of_rat x"
+ define P where "P x \<longleftrightarrow> (\<forall>y\<in>S. y \<le> of_rat x)" for x
obtain a where a: "\<not> P a"
proof
have "of_int \<lfloor>x - 1\<rfloor> \<le> x - 1" by (rule of_int_floor_le)
@@ -797,11 +797,11 @@
qed
qed
- def avg \<equiv> "\<lambda>x y :: rat. x/2 + y/2"
- def bisect \<equiv> "\<lambda>(x, y). if P (avg x y) then (x, avg x y) else (avg x y, y)"
- def A \<equiv> "\<lambda>n. fst ((bisect ^^ n) (a, b))"
- def B \<equiv> "\<lambda>n. snd ((bisect ^^ n) (a, b))"
- def C \<equiv> "\<lambda>n. avg (A n) (B n)"
+ define avg where "avg x y = x/2 + y/2" for x y :: rat
+ define bisect where "bisect = (\<lambda>(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]: "\<And>n. A (Suc n) = (if P (C n) then A n else C n)"
@@ -1251,8 +1251,8 @@
from \<open>x<y\<close> 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 \<equiv> "\<lceil>y * real q\<rceil> - 1"
- def r \<equiv> "of_int p / real q"
+ define p where "p = \<lceil>y * real q\<rceil> - 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) \<le> r"
unfolding r_def p_def
--- a/src/HOL/Real_Vector_Spaces.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Real_Vector_Spaces.thy Mon Apr 25 16:09:26 2016 +0200
@@ -2070,7 +2070,7 @@
assumes X: "Cauchy X"
shows "convergent X"
proof -
- def S \<equiv> "{x::real. \<exists>N. \<forall>n\<ge>N. x < X n}"
+ define S :: "real set" where "S = {x. \<exists>N. \<forall>n\<ge>N. x < X n}"
then have mem_S: "\<And>N x. \<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S" by auto
{ fix N x assume N: "\<forall>n\<ge>N. X n < x"
--- a/src/HOL/Rings.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Rings.thy Mon Apr 25 16:09:26 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 \<equiv> "1 div a"
+ define b where "b = 1 div a"
then show "1 div a = b" by simp
from b_def \<open>is_unit a\<close> show "is_unit b" by simp
from \<open>is_unit a\<close> and \<open>is_unit b\<close> show "a \<noteq> 0" and "b \<noteq> 0" by auto
--- a/src/HOL/Series.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Series.thy Mon Apr 25 16:09:26 2016 +0200
@@ -84,7 +84,7 @@
shows "summable f = summable g"
proof -
from assms obtain N where N: "\<forall>n\<ge>N. f n = g n" by (auto simp: eventually_at_top_linorder)
- def C \<equiv> "(\<Sum>k<N. f k - g k)"
+ define C where "C = (\<Sum>k<N. f k - g k)"
from eventually_ge_at_top[of N]
have "eventually (\<lambda>n. setsum f {..<n} = C + setsum g {..<n}) sequentially"
proof eventually_elim
@@ -1048,7 +1048,7 @@
finally show "(\<lambda>n. \<Sum>k<n. f (g k)) \<longlonglongrightarrow> c" .
next
assume lim: "(\<lambda>n. \<Sum>k<n. f (g k)) \<longlonglongrightarrow> c"
- def g_inv \<equiv> "\<lambda>n. LEAST m. g m \<ge> n"
+ define g_inv where "g_inv n = (LEAST m. g m \<ge> n)" for n
from filterlim_subseq[OF subseq] have g_inv_ex: "\<exists>m. g m \<ge> n" for n
by (auto simp: filterlim_at_top eventually_at_top_linorder)
hence g_inv: "g (g_inv n) \<ge> n" for n unfolding g_inv_def by (rule LeastI_ex)
--- a/src/HOL/Topological_Spaces.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Topological_Spaces.thy Mon Apr 25 16:09:26 2016 +0200
@@ -1183,7 +1183,7 @@
from first_countable_basis[of x] obtain A :: "nat \<Rightarrow> 'a set" where
nhds: "\<And>i. open (A i)" "\<And>i. x \<in> A i"
and incl: "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> \<exists>i. A i \<subseteq> S" by auto
- def F \<equiv> "\<lambda>n. \<Inter>i\<le>n. A i"
+ define F where "F n = (\<Inter>i\<le>n. A i)" for n
show "\<exists>A. (\<forall>i. open (A i)) \<and> (\<forall>i. x \<in> A i) \<and>
(\<forall>S. open S \<longrightarrow> x \<in> S \<longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially)"
proof (safe intro!: exI[of _ F])
--- a/src/HOL/Transcendental.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Transcendental.thy Mon Apr 25 16:09:26 2016 +0200
@@ -893,7 +893,7 @@
from \<open>0 < r\<close> have "0 < ?r" by simp
let ?s = "\<lambda>n. SOME s. 0 < s \<and> (\<forall> x. x \<noteq> 0 \<and> \<bar> x \<bar> < s \<longrightarrow> \<bar> ?diff n x - f' x0 n \<bar> < ?r)"
- def S' \<equiv> "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 \<equiv> "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 \<le> x0 - a" and S_b: "S \<le> b - x0"
and "S \<le> S'" using x0_in_I and \<open>0 < S'\<close>
by auto
@@ -2248,7 +2248,7 @@
assumes "x > 0"
shows "DERIV (\<lambda>y. log b y) x :> 1 / (ln b * x)"
proof -
- def lb \<equiv> "1 / ln b"
+ define lb where "lb = 1 / ln b"
moreover have "DERIV (\<lambda>y. lb * ln y) x :> lb / x"
using \<open>x > 0\<close> by (auto intro!: derivative_eq_intros)
ultimately show ?thesis
@@ -4769,7 +4769,7 @@
proof (rule tendstoI)
fix e :: real
assume "0 < e"
- def y \<equiv> "pi/2 - min (pi/2) e"
+ define y where "y = pi/2 - min (pi/2) e"
then have y: "0 \<le> y" "y < pi/2" "pi/2 \<le> e + y"
using \<open>0 < e\<close> by auto
--- a/src/HOL/Wfrec.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Wfrec.thy Mon Apr 25 16:09:26 2016 +0200
@@ -37,7 +37,7 @@
lemma wfrec_unique: assumes "adm_wf R F" "wf R" shows "\<exists>!y. wfrec_rel R F x y"
using \<open>wf R\<close>
proof induct
- def f \<equiv> "\<lambda>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 "\<And>y z. (y, x) \<in> R \<Longrightarrow> wfrec_rel R F y z \<longleftrightarrow> z = f y"
unfolding f_def by (rule theI_unique)
--- a/src/HOL/Zorn.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/Zorn.thy Mon Apr 25 16:09:26 2016 +0200
@@ -609,7 +609,7 @@
proof -
\<comment> \<open>The initial segment relation on well-orders:\<close>
let ?WO = "{r::'a rel. Well_order r}"
- def I \<equiv> "init_seg_of \<inter> ?WO \<times> ?WO"
+ define I where "I = init_seg_of \<inter> ?WO \<times> ?WO"
have I_init: "I \<subseteq> init_seg_of" by (auto simp: I_def)
hence subch: "\<And>R. R \<in> Chains I \<Longrightarrow> chain\<^sub>\<subseteq> R"
unfolding init_seg_of_def chain_subset_def Chains_def by blast
@@ -732,7 +732,7 @@
shows "\<exists>f. \<forall>x. P f x (f x)"
proof (intro exI allI)
fix x
- def f \<equiv> "wfrec R (\<lambda>f x. SOME r. P f x r)"
+ define f where "f \<equiv> wfrec R (\<lambda>f x. SOME r. P f x r)"
from \<open>wf R\<close> show "P f x (f x)"
proof (induct x)
fix x assume "\<And>y. (y, x) \<in> R \<Longrightarrow> P f y (f y)"
--- a/src/HOL/ex/Ballot.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/ex/Ballot.thy Mon Apr 25 16:09:26 2016 +0200
@@ -77,13 +77,14 @@
card {V\<in>Pow {0..<a+b}. card V = a \<and> (\<forall>m\<in>{1..a+b}. card ({0..<m} \<inter> V) > m - card ({0..<m} \<inter> V))}"
(is "_ = card ?V")
proof -
- def P \<equiv> "\<lambda>j i. i < a + b \<and> j = Suc i"
+ define P where "P j i \<longleftrightarrow> i < a + b \<and> j = Suc i" for j i
have unique_P: "bi_unique P" and total_P: "\<And>m. m \<le> a + b \<Longrightarrow> rel_set P {1..m} {0..<m}"
by (auto simp add: bi_unique_def rel_set_def P_def Suc_le_eq gr0_conv_Suc)
have rel_fun_P: "\<And>R f g. (\<And>i. i < a+b \<Longrightarrow> R (f (Suc i)) (g i)) \<Longrightarrow> rel_fun P R f g"
by (simp add: rel_fun_def P_def)
- def R \<equiv> "\<lambda>f V. V \<subseteq> {0..<a+b} \<and> f \<in> extensional {1..a+b} \<and> (\<forall>i<a+b. i \<in> V \<longleftrightarrow> f (Suc i) = A)"
+ define R where "R f V \<longleftrightarrow>
+ V \<subseteq> {0..<a+b} \<and> f \<in> extensional {1..a+b} \<and> (\<forall>i<a+b. i \<in> V \<longleftrightarrow> f (Suc i) = A)" for f V
{ fix f g :: "nat \<Rightarrow> vote" assume "f \<in> extensional {1..a + b}" "g \<in> extensional {1..a + b}"
moreover assume "\<forall>i<a + b. (f (Suc i) = A) = (g (Suc i) = A)"
then have "\<forall>i<a + b. f (Suc i) = g (Suc i)"
--- a/src/HOL/ex/Erdoes_Szekeres.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/ex/Erdoes_Szekeres.thy Mon Apr 25 16:09:26 2016 +0200
@@ -80,7 +80,7 @@
(\<exists>S. S \<subseteq> {0..m * n} \<and> card S = n + 1 \<and> mono_on f (op \<ge>) S)"
proof (rule ccontr)
let ?max_subseq = "\<lambda>R k. Max (card ` {S. S \<subseteq> {0..k} \<and> mono_on f R S \<and> k \<in> S})"
- def phi == "\<lambda>k. (?max_subseq (op \<le>) k, ?max_subseq (op \<ge>) k)"
+ define phi where "phi k = (?max_subseq (op \<le>) k, ?max_subseq (op \<ge>) k)" for k
have one_member: "\<And>R k. reflp R \<Longrightarrow> {k} \<in> {S. S \<subseteq> {0..k} \<and> mono_on f R S \<and> k \<in> S}" by auto
--- a/src/HOL/ex/Gauge_Integration.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/ex/Gauge_Integration.thy Mon Apr 25 16:09:26 2016 +0200
@@ -379,9 +379,10 @@
and I2: "\<And> D. fine \<delta>2 (b,c) D \<Longrightarrow> \<bar> rsum D f - x2 \<bar> < (\<epsilon> / 2)"
using IntegralE [OF \<open>Integral (b, c) f x2\<close> \<open>0 < \<epsilon>/2\<close>] by auto
- def \<delta> \<equiv> "\<lambda> x. if x < b then min (\<delta>1 x) (b - x)
- else if x = b then min (\<delta>1 b) (\<delta>2 b)
- else min (\<delta>2 x) (x - b)"
+ define \<delta> where "\<delta> x =
+ (if x < b then min (\<delta>1 x) (b - x)
+ else if x = b then min (\<delta>1 b) (\<delta>2 b)
+ else min (\<delta>2 x) (x - b))" for x
have "gauge {a..c} \<delta>"
using \<delta>1_gauge \<delta>2_gauge unfolding \<delta>_def gauge_def by auto
@@ -410,8 +411,8 @@
let ?D1 = "take N D"
let ?D2 = "drop N D"
- def D1 \<equiv> "take N D @ [(d, t, b)]"
- def D2 \<equiv> "(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 \<open>N < length D\<close>]
have "fst (hd ?D2) = d" using \<open>D ! N = (d, t, e)\<close> by auto
--- a/src/HOL/ex/HarmonicSeries.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/HOL/ex/HarmonicSeries.thy Mon Apr 25 16:09:26 2016 +0200
@@ -279,7 +279,7 @@
assume sf: "summable ?f"
then obtain n::nat where ndef: "n = nat \<lceil>2 * ?s\<rceil>" by simp
then have ngt: "1 + real n/2 > ?s" by linarith
- def j \<equiv> "(2::nat)^n"
+ define j where "j = (2::nat)^n"
have "\<forall>m\<ge>j. 0 < ?f m" by simp
with sf have "(\<Sum>i<j. ?f i) < ?s" by (rule setsum_less_suminf)
then have "(\<Sum>i\<in>{Suc 0..<Suc j}. 1/(real i)) < ?s"
--- a/src/ZF/CardinalArith.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/ZF/CardinalArith.thy Mon Apr 25 16:09:26 2016 +0200
@@ -28,7 +28,7 @@
definition
jump_cardinal :: "i=>i" where
- \<comment>\<open>This def is more complex than Kunen's but it more easily proved to
+ \<comment>\<open>This definition is more complex than Kunen's but it more easily proved to
be a cardinal\<close>
"jump_cardinal(K) ==
\<Union>X\<in>Pow(K). {z. r \<in> Pow(K*K), well_ord(X,r) & z = ordertype(X,r)}"
--- a/src/ZF/Induct/Multiset.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/ZF/Induct/Multiset.thy Mon Apr 25 16:09:26 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 \<in> nat ==> 0 < k \<longrightarrow> n #+ k #- 1 = n #+ (k #- 1)"
by (erule nat_induct, auto)
--- a/src/ZF/QUniv.thy Sun Apr 24 21:31:14 2016 +0200
+++ b/src/ZF/QUniv.thy Mon Apr 25 16:09:26 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) \<subseteq> quniv(A)"
apply (rule arg_subset_eclose [THEN univ_mono, THEN subset_trans])
apply (rule univ_eclose_subset_quniv)