weaker precendence of syntax for big intersection and union on sets
authorhaftmann
Sat, 25 May 2013 15:44:29 +0200
changeset 52141 eff000cab70f
parent 52140 88a69da5d3fa
child 52148 893b15200ec1
weaker precendence of syntax for big intersection and union on sets
NEWS
src/HOL/BNF/BNF_Comp.thy
src/HOL/Complete_Lattices.thy
src/HOL/Hoare_Parallel/OG_Tran.thy
src/HOL/List.thy
src/HOL/Multivariate_Analysis/Integration.thy
src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
src/HOL/Probability/Caratheodory.thy
src/HOL/Probability/Radon_Nikodym.thy
src/HOL/Probability/Regularity.thy
--- a/NEWS	Sat May 25 15:44:08 2013 +0200
+++ b/NEWS	Sat May 25 15:44:29 2013 +0200
@@ -57,6 +57,9 @@
 
 *** HOL ***
 
+* Weaker precendence of syntax for big intersection and union on sets,
+in accordance with corresponding lattice operations.  INCOMPATIBILITY.
+
 * Nested case expressions are now translated in a separate check
   phase rather than during parsing. The data for case combinators
   is separated from the datatype package. The declaration attribute
--- a/src/HOL/BNF/BNF_Comp.thy	Sat May 25 15:44:08 2013 +0200
+++ b/src/HOL/BNF/BNF_Comp.thy	Sat May 25 15:44:29 2013 +0200
@@ -24,7 +24,7 @@
   assumes fbd_Card_order: "Card_order fbd" and
     fset_bd: "\<And>x. |fset x| \<le>o fbd" and
     gset_bd: "\<And>x. |gset x| \<le>o gbd"
-  shows "|\<Union>fset ` gset x| \<le>o gbd *c fbd"
+  shows "|\<Union>(fset ` gset x)| \<le>o gbd *c fbd"
 apply (subst sym[OF SUP_def])
 apply (rule ordLeq_transitive)
 apply (rule card_of_UNION_Sigma)
@@ -42,10 +42,10 @@
 apply (rule Card_order_cprod)
 done
 
-lemma Union_image_insert: "\<Union>f ` insert a B = f a \<union> \<Union>f ` B"
+lemma Union_image_insert: "\<Union>(f ` insert a B) = f a \<union> \<Union>(f ` B)"
 by simp
 
-lemma Union_image_empty: "A \<union> \<Union>f ` {} = A"
+lemma Union_image_empty: "A \<union> \<Union>(f ` {}) = A"
 by simp
 
 lemma image_o_collect: "collect ((\<lambda>f. image g o f) ` F) = image g o collect F"
@@ -54,10 +54,10 @@
 lemma conj_subset_def: "A \<subseteq> {x. P x \<and> Q x} = (A \<subseteq> {x. P x} \<and> A \<subseteq> {x. Q x})"
 by blast
 
-lemma UN_image_subset: "\<Union>f ` g x \<subseteq> X = (g x \<subseteq> {x. f x \<subseteq> X})"
+lemma UN_image_subset: "\<Union>(f ` g x) \<subseteq> X = (g x \<subseteq> {x. f x \<subseteq> X})"
 by blast
 
-lemma comp_set_bd_Union_o_collect: "|\<Union>\<Union>(\<lambda>f. f x) ` X| \<le>o hbd \<Longrightarrow> |(Union \<circ> collect X) x| \<le>o hbd"
+lemma comp_set_bd_Union_o_collect: "|\<Union>\<Union>((\<lambda>f. f x) ` X)| \<le>o hbd \<Longrightarrow> |(Union \<circ> collect X) x| \<le>o hbd"
 by (unfold o_apply collect_def SUP_def)
 
 lemma wpull_cong:
--- a/src/HOL/Complete_Lattices.thy	Sat May 25 15:44:08 2013 +0200
+++ b/src/HOL/Complete_Lattices.thy	Sat May 25 15:44:29 2013 +0200
@@ -751,7 +751,7 @@
   "Inter S \<equiv> \<Sqinter>S"
   
 notation (xsymbols)
-  Inter  ("\<Inter>_" [90] 90)
+  Inter  ("\<Inter>_" [900] 900)
 
 lemma Inter_eq:
   "\<Inter>A = {x. \<forall>B \<in> A. x \<in> B}"
@@ -934,7 +934,7 @@
   "Union S \<equiv> \<Squnion>S"
 
 notation (xsymbols)
-  Union  ("\<Union>_" [90] 90)
+  Union  ("\<Union>_" [900] 900)
 
 lemma Union_eq:
   "\<Union>A = {x. \<exists>B \<in> A. x \<in> B}"
--- a/src/HOL/Hoare_Parallel/OG_Tran.thy	Sat May 25 15:44:08 2013 +0200
+++ b/src/HOL/Hoare_Parallel/OG_Tran.thy	Sat May 25 15:44:29 2013 +0200
@@ -90,13 +90,13 @@
   "ann_sem c \<equiv> \<lambda>s. {t. (Some c, s) -*\<rightarrow> (None, t)}"
 
 definition ann_SEM :: "'a ann_com \<Rightarrow> 'a set \<Rightarrow> 'a set" where
-  "ann_SEM c S \<equiv> \<Union>ann_sem c ` S"  
+  "ann_SEM c S \<equiv> \<Union>(ann_sem c ` S)"
 
 definition sem :: "'a com \<Rightarrow> 'a \<Rightarrow> 'a set" where
   "sem c \<equiv> \<lambda>s. {t. \<exists>Ts. (c, s) -P*\<rightarrow> (Parallel Ts, t) \<and> All_None Ts}"
 
 definition SEM :: "'a com \<Rightarrow> 'a set \<Rightarrow> 'a set" where
-  "SEM c S \<equiv> \<Union>sem c ` S "
+  "SEM c S \<equiv> \<Union>(sem c ` S)"
 
 abbreviation Omega :: "'a com"    ("\<Omega>" 63)
   where "\<Omega> \<equiv> While UNIV UNIV (Basic id)"
--- a/src/HOL/List.thy	Sat May 25 15:44:08 2013 +0200
+++ b/src/HOL/List.thy	Sat May 25 15:44:29 2013 +0200
@@ -4312,7 +4312,7 @@
   moreover have "\<And>i. i \<in>Collect (?k_list k) \<Longrightarrow> card (A - set i) = card A - k"
     by (simp add: card_Diff_subset distinct_card)
   moreover have "{xs. ?k_list (Suc k) xs} =
-      (\<lambda>(xs, n). n#xs) ` \<Union>(\<lambda>xs. {xs} \<times> (A - set xs)) ` {xs. ?k_list k xs}"
+      (\<lambda>(xs, n). n#xs) ` \<Union>((\<lambda>xs. {xs} \<times> (A - set xs)) ` {xs. ?k_list k xs})"
     by (auto simp: length_Suc_conv)
   moreover
   have "Suc (card A - Suc k) = card A - k" using Suc.prems by simp
--- a/src/HOL/Multivariate_Analysis/Integration.thy	Sat May 25 15:44:08 2013 +0200
+++ b/src/HOL/Multivariate_Analysis/Integration.thy	Sat May 25 15:44:29 2013 +0200
@@ -1015,7 +1015,7 @@
     apply (rule that[of "d \<union> p"])
   proof -
     have *: "\<And>s f t. s \<noteq> {} \<Longrightarrow> (\<forall>i\<in>s. f i \<union> i = t) \<Longrightarrow> t = \<Inter> (f ` s) \<union> (\<Union>s)" by auto
-    have *: "{a..b} = \<Inter> (\<lambda>i. \<Union>(q i - {i})) ` p \<union> \<Union>p"
+    have *: "{a..b} = \<Inter> ((\<lambda>i. \<Union>(q i - {i})) ` p) \<union> \<Union>p"
       apply (rule *[OF False])
     proof
       fix i
@@ -1032,7 +1032,7 @@
       fix k
       assume k: "k\<in>p"
       have *: "\<And>u t s. u \<subseteq> s \<Longrightarrow> s \<inter> t = {} \<Longrightarrow> u \<inter> t = {}" by auto
-      show "interior (\<Inter>(\<lambda>i. \<Union>(q i - {i})) ` p) \<inter> interior k = {}"
+      show "interior (\<Inter> ((\<lambda>i. \<Union>(q i - {i})) ` p)) \<inter> interior k = {}"
         apply (rule *[of _ "interior (\<Union>(q k - {k}))"])
         defer
         apply (subst Int_commute)
@@ -1044,7 +1044,7 @@
         show "\<forall>t\<in>q k - {k}. interior k \<inter> interior t = {}"
           using qk(5) using q(2)[OF k] by auto
         have *: "\<And>x s. x \<in> s \<Longrightarrow> \<Inter>s \<subseteq> x" by auto
-        show "interior (\<Inter>(\<lambda>i. \<Union>(q i - {i})) ` p) \<subseteq> interior (\<Union>(q k - {k}))"
+        show "interior (\<Inter> ((\<lambda>i. \<Union>(q i - {i})) ` p)) \<subseteq> interior (\<Union>(q k - {k}))"
           apply (rule interior_mono *)+
           using k by auto
       qed
@@ -1131,7 +1131,7 @@
 lemma elementary_union_interval: assumes "p division_of \<Union>p"
   obtains q where "q division_of ({a..b::'a::ordered_euclidean_space} \<union> \<Union>p)" proof-
   note assm=division_ofD[OF assms]
-  have lem1:"\<And>f s. \<Union>\<Union> (f ` s) = \<Union>(\<lambda>x.\<Union>(f x)) ` s" by auto
+  have lem1:"\<And>f s. \<Union>\<Union> (f ` s) = \<Union>((\<lambda>x.\<Union>(f x)) ` s)" by auto
   have lem2:"\<And>f s. f \<noteq> {} \<Longrightarrow> \<Union>{s \<union> t |t. t \<in> f} = s \<union> \<Union>f" by auto
 { presume "p={} \<Longrightarrow> thesis" "{a..b} = {} \<Longrightarrow> thesis" "{a..b} \<noteq> {} \<Longrightarrow> interior {a..b} = {} \<Longrightarrow> thesis"
     "p\<noteq>{} \<Longrightarrow> interior {a..b}\<noteq>{} \<Longrightarrow> {a..b} \<noteq> {} \<Longrightarrow> thesis"
@@ -1282,7 +1282,7 @@
 
 lemma division_of_tagged_division: assumes"s tagged_division_of i"  shows "(snd ` s) division_of i"
 proof(rule division_ofI) note assm=tagged_division_ofD[OF assms]
-  show "\<Union>snd ` s = i" "finite (snd ` s)" using assm by auto
+  show "\<Union>(snd ` s) = i" "finite (snd ` s)" using assm by auto
   fix k assume k:"k \<in> snd ` s" then obtain xk where xk:"(xk, k) \<in> s" by auto
   thus  "k \<subseteq> i" "k \<noteq> {}" "\<exists>a b. k = {a..b}" using assm apply- by fastforce+
   fix k' assume k':"k' \<in> snd ` s" "k \<noteq> k'" from this(1) obtain xk' where xk':"(xk', k') \<in> s" by auto
@@ -1292,9 +1292,9 @@
 lemma partial_division_of_tagged_division: assumes "s tagged_partial_division_of i"
   shows "(snd ` s) division_of \<Union>(snd ` s)"
 proof(rule division_ofI) note assm=tagged_partial_division_ofD[OF assms]
-  show "finite (snd ` s)" "\<Union>snd ` s = \<Union>snd ` s" using assm by auto
+  show "finite (snd ` s)" "\<Union>(snd ` s) = \<Union>(snd ` s)" using assm by auto
   fix k assume k:"k \<in> snd ` s" then obtain xk where xk:"(xk, k) \<in> s" by auto
-  thus "k\<noteq>{}" "\<exists>a b. k = {a..b}" "k \<subseteq> \<Union>snd ` s" using assm by auto
+  thus "k\<noteq>{}" "\<exists>a b. k = {a..b}" "k \<subseteq> \<Union>(snd ` s)" using assm by auto
   fix k' assume k':"k' \<in> snd ` s" "k \<noteq> k'" from this(1) obtain xk' where xk':"(xk', k') \<in> s" by auto
   thus "interior k \<inter> interior k' = {}" apply-apply(rule assm(5)) apply(rule xk xk')+ using k' by auto
 qed
@@ -1355,13 +1355,13 @@
   shows "\<Union>(pfn ` iset) tagged_division_of (\<Union>iset)"
 proof(rule tagged_division_ofI)
   note assm = tagged_division_ofD[OF assms(2)[rule_format]]
-  show "finite (\<Union>pfn ` iset)" apply(rule finite_Union) using assms by auto
-  have "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>pfn ` iset} = \<Union>(\<lambda>i. \<Union>{k. \<exists>x. (x, k) \<in> pfn i}) ` iset" by blast 
+  show "finite (\<Union>(pfn ` iset))" apply(rule finite_Union) using assms by auto
+  have "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>((\<lambda>i. \<Union>{k. \<exists>x. (x, k) \<in> pfn i}) ` iset)" by blast 
   also have "\<dots> = \<Union>iset" using assm(6) by auto
-  finally show "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>pfn ` iset} = \<Union>iset" . 
-  fix x k assume xk:"(x,k)\<in>\<Union>pfn ` iset" then obtain i where i:"i \<in> iset" "(x, k) \<in> pfn i" by auto
+  finally show "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>iset" . 
+  fix x k assume xk:"(x,k)\<in>\<Union>(pfn ` iset)" then obtain i where i:"i \<in> iset" "(x, k) \<in> pfn i" by auto
   show "x\<in>k" "\<exists>a b. k = {a..b}" "k \<subseteq> \<Union>iset" using assm(2-4)[OF i] using i(1) by auto
-  fix x' k' assume xk':"(x',k')\<in>\<Union>pfn ` iset" "(x, k) \<noteq> (x', k')" then obtain i' where i':"i' \<in> iset" "(x', k') \<in> pfn i'" by auto
+  fix x' k' assume xk':"(x',k')\<in>\<Union>(pfn ` iset)" "(x, k) \<noteq> (x', k')" then obtain i' where i':"i' \<in> iset" "(x', k') \<in> pfn i'" by auto
   have *:"\<And>a b. i\<noteq>i' \<Longrightarrow> a\<subseteq> i \<Longrightarrow> b\<subseteq> i' \<Longrightarrow> interior a \<inter> interior b = {}" using i(1) i'(1)
     using assms(3)[rule_format] interior_mono by blast
   show "interior k \<inter> interior k' = {}" apply(cases "i=i'")
@@ -4975,7 +4975,7 @@
   shows "norm(setsum (\<lambda>(x,k). content k *\<^sub>R f x - integral k f) p) \<le> e" (is "?x \<le> e")
 proof-  { presume "\<And>k. 0<k \<Longrightarrow> ?x \<le> e + k" thus ?thesis by (blast intro: field_le_epsilon) }
   fix k::real assume k:"k>0" note p' = tagged_partial_division_ofD[OF p(1)]
-  have "\<Union>snd ` p \<subseteq> {a..b}" using p'(3) by fastforce
+  have "\<Union>(snd ` p) \<subseteq> {a..b}" 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" have "snd ` p \<inter> r = {}" unfolding r_def by auto
@@ -4994,11 +4994,11 @@
     thus ?case apply(rule_tac x=qq in exI) using dd(2)[of qq] unfolding fine_inter uv by auto qed
   from bchoice[OF this] guess qq .. note qq=this[rule_format]
 
-  let ?p = "p \<union> \<Union>qq ` r" have "norm ((\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) - integral {a..b} f) < e"
+  let ?p = "p \<union> \<Union>(qq ` r)" have "norm ((\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) - integral {a..b} f) < e"
     apply(rule assms(4)[rule_format])
   proof show "d fine ?p" apply(rule fine_union,rule p) apply(rule fine_unions) using qq by auto 
     note * = tagged_partial_division_of_union_self[OF p(1)]
-    have "p \<union> \<Union>qq ` r tagged_division_of \<Union>snd ` p \<union> \<Union>r"
+    have "p \<union> \<Union>(qq ` r) tagged_division_of \<Union>(snd ` p) \<union> \<Union>r"
     proof(rule tagged_division_union[OF * tagged_division_unions])
       show "finite r" by fact case goal2 thus ?case using qq by auto
     next case goal3 thus ?case apply(rule,rule,rule) apply(rule q'(5)) unfolding r_def by auto
@@ -5006,11 +5006,11 @@
         apply(rule,rule q') defer apply(rule,subst Int_commute) 
         apply(rule inter_interior_unions_intervals) apply(rule finite_imageI,rule p',rule) defer
         apply(rule,rule q') using q(1) p' unfolding r_def by auto qed
-    moreover have "\<Union>snd ` p \<union> \<Union>r = {a..b}" "{qq i |i. i \<in> r} = qq ` r"
+    moreover have "\<Union>(snd ` p) \<union> \<Union>r = {a..b}" "{qq i |i. i \<in> r} = qq ` r"
       unfolding Union_Un_distrib[THEN sym] r_def using q by auto
     ultimately show "?p tagged_division_of {a..b}" by fastforce qed
 
-  hence "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>\<Union>qq ` r. content k *\<^sub>R f x) -
+  hence "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>\<Union>(qq ` r). content k *\<^sub>R f x) -
     integral {a..b} f) < e" apply(subst setsum_Un_zero[THEN sym]) apply(rule p') prefer 3 
     apply assumption apply rule apply(rule finite_imageI,rule r) apply safe apply(drule qq)
   proof- fix x l k assume as:"(x,l)\<in>p" "(x,l)\<in>qq k" "k\<in>r"
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Sat May 25 15:44:08 2013 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Sat May 25 15:44:29 2013 +0200
@@ -220,7 +220,7 @@
     fix a assume "a \<in> A"
     thus "x \<in> a" "open a" using A'(4)[OF open_UNIV] by (auto simp: A_def intro: A' from_nat_into)
   next
-    let ?int = "\<lambda>N. \<Inter>from_nat_into A' ` N"
+    let ?int = "\<lambda>N. \<Inter>(from_nat_into A' ` N)"
     fix a b assume "a \<in> A" "b \<in> A"
     then obtain N M where "a = ?int N" "b = ?int M" "finite (N \<union> M)" by (auto simp: A_def)
     thus "a \<inter> b \<in> A" by (auto simp: A_def intro!: image_eqI[where x="N \<union> M"])
@@ -2617,7 +2617,7 @@
 proof (safe intro!: compact_eq_heine_borel[THEN iffD2])
   fix A assume "compact U" and A: "\<forall>a\<in>A. closed a" "U \<inter> \<Inter>A = {}"
     and fi: "\<forall>B \<subseteq> A. finite B \<longrightarrow> U \<inter> \<Inter>B \<noteq> {}"
-  from A have "(\<forall>a\<in>uminus`A. open a) \<and> U \<subseteq> \<Union>uminus`A"
+  from A have "(\<forall>a\<in>uminus`A. open a) \<and> U \<subseteq> \<Union>(uminus`A)"
     by auto
   with `compact U` obtain B where "B \<subseteq> A" "finite (uminus`B)" "U \<subseteq> \<Union>(uminus`B)"
     unfolding compact_eq_heine_borel by (metis subset_image_iff)
@@ -2627,7 +2627,7 @@
   fix A assume ?R and cover: "\<forall>a\<in>A. open a" "U \<subseteq> \<Union>A"
   from cover have "U \<inter> \<Inter>(uminus`A) = {}" "\<forall>a\<in>uminus`A. closed a"
     by auto
-  with `?R` obtain B where "B \<subseteq> A" "finite (uminus`B)" "U \<inter> \<Inter>uminus`B = {}"
+  with `?R` obtain B where "B \<subseteq> A" "finite (uminus`B)" "U \<inter> \<Inter>(uminus`B) = {}"
     by (metis subset_image_iff)
   then show "\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T"
     by  (auto intro!: exI[of _ B] inj_setminus dest: finite_imageD)
@@ -3041,7 +3041,7 @@
   assumes "seq_compact s"
   shows "\<forall>e>0. \<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k))"
 proof(rule, rule, rule ccontr)
-  fix e::real assume "e>0" and assm:"\<not> (\<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k)"
+  fix e::real assume "e>0" and assm:"\<not> (\<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> \<Union>((\<lambda>x. ball x e) ` k))"
   def x \<equiv> "helper_1 s e"
   { fix n
     have "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)"
@@ -4586,7 +4586,7 @@
     moreover with c have "(\<Inter>y\<in>D. a y) \<times> t \<subseteq> (\<Union>y\<in>D. c y)"
       by (fastforce simp: subset_eq)
     ultimately show "\<exists>a. open a \<and> x \<in> a \<and> (\<exists>d\<subseteq>C. finite d \<and> a \<times> t \<subseteq> \<Union>d)"
-      using c by (intro exI[of _ "c`D"] exI[of _ "\<Inter>a`D"] conjI) (auto intro!: open_INT)
+      using c by (intro exI[of _ "c`D"] exI[of _ "\<Inter>(a`D)"] conjI) (auto intro!: open_INT)
   qed
   then obtain a d where a: "\<forall>x\<in>s. open (a x)" "s \<subseteq> (\<Union>x\<in>s. a x)"
     and d: "\<And>x. x \<in> s \<Longrightarrow> d x \<subseteq> C \<and> finite (d x) \<and> a x \<times> t \<subseteq> \<Union>d x"
--- a/src/HOL/Probability/Caratheodory.thy	Sat May 25 15:44:08 2013 +0200
+++ b/src/HOL/Probability/Caratheodory.thy	Sat May 25 15:44:29 2013 +0200
@@ -707,9 +707,9 @@
   assumes A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M" "disjoint_family_on A I" "finite I" "UNION I A \<in> M"
   shows "f (UNION I A) = (\<Sum>i\<in>I. f (A i))"
 proof -
-  have "A`I \<subseteq> M" "disjoint (A`I)" "finite (A`I)" "\<Union>A`I \<in> M"
+  have "A`I \<subseteq> M" "disjoint (A`I)" "finite (A`I)" "\<Union>(A`I) \<in> M"
     using A unfolding SUP_def by (auto simp: disjoint_family_on_disjoint_image)
-  with `volume M f` have "f (\<Union>A`I) = (\<Sum>a\<in>A`I. f a)"
+  with `volume M f` have "f (\<Union>(A`I)) = (\<Sum>a\<in>A`I. f a)"
     unfolding volume_def by blast
   also have "\<dots> = (\<Sum>i\<in>I. f (A i))"
   proof (subst setsum_reindex_nonzero)
--- a/src/HOL/Probability/Radon_Nikodym.thy	Sat May 25 15:44:08 2013 +0200
+++ b/src/HOL/Probability/Radon_Nikodym.thy	Sat May 25 15:44:29 2013 +0200
@@ -571,9 +571,9 @@
     show "(SUP i. emeasure M (?O i)) \<le> ?a" unfolding SUP_def
     proof (safe intro!: Sup_mono, unfold bex_simps)
       fix i
-      have *: "(\<Union>Q' ` {..i}) = ?O i" by auto
+      have *: "(\<Union>(Q' ` {..i})) = ?O i" by auto
       then show "\<exists>x. (x \<in> sets M \<and> N x \<noteq> \<infinity>) \<and>
-        emeasure M (\<Union>Q' ` {..i}) \<le> emeasure M x"
+        emeasure M (\<Union>(Q' ` {..i})) \<le> emeasure M x"
         using O_in_G[of i] by (auto intro!: exI[of _ "?O i"])
     qed
   qed
--- a/src/HOL/Probability/Regularity.thy	Sat May 25 15:44:08 2013 +0200
+++ b/src/HOL/Probability/Regularity.thy	Sat May 25 15:44:29 2013 +0200
@@ -209,8 +209,8 @@
       from nat_approx_posE[OF this] guess n . note n = this
       let ?k = "from_nat_into X ` {0..k e (Suc n)}"
       have "finite ?k" by simp
-      moreover have "K \<subseteq> \<Union>(\<lambda>x. ball x e') ` ?k" unfolding K_def B_def using n by force
-      ultimately show "\<exists>k. finite k \<and> K \<subseteq> \<Union>(\<lambda>x. ball x e') ` k" by blast
+      moreover have "K \<subseteq> \<Union>((\<lambda>x. ball x e') ` ?k)" unfolding K_def B_def using n by force
+      ultimately show "\<exists>k. finite k \<and> K \<subseteq> \<Union>((\<lambda>x. ball x e') ` k)" by blast
     qed
     ultimately
     show "?thesis e " by (auto simp: sU)