# HG changeset patch # User haftmann # Date 1369489469 -7200 # Node ID eff000cab70f21ddca804bd95c90f1460b94f39b # Parent 88a69da5d3faaeacf66a94b6465b4f1106c1506e weaker precendence of syntax for big intersection and union on sets diff -r 88a69da5d3fa -r eff000cab70f NEWS --- 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 diff -r 88a69da5d3fa -r eff000cab70f src/HOL/BNF/BNF_Comp.thy --- 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: "\x. |fset x| \o fbd" and gset_bd: "\x. |gset x| \o gbd" - shows "|\fset ` gset x| \o gbd *c fbd" + shows "|\(fset ` gset x)| \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: "\f ` insert a B = f a \ \f ` B" +lemma Union_image_insert: "\(f ` insert a B) = f a \ \(f ` B)" by simp -lemma Union_image_empty: "A \ \f ` {} = A" +lemma Union_image_empty: "A \ \(f ` {}) = A" by simp lemma image_o_collect: "collect ((\f. image g o f) ` F) = image g o collect F" @@ -54,10 +54,10 @@ lemma conj_subset_def: "A \ {x. P x \ Q x} = (A \ {x. P x} \ A \ {x. Q x})" by blast -lemma UN_image_subset: "\f ` g x \ X = (g x \ {x. f x \ X})" +lemma UN_image_subset: "\(f ` g x) \ X = (g x \ {x. f x \ X})" by blast -lemma comp_set_bd_Union_o_collect: "|\\(\f. f x) ` X| \o hbd \ |(Union \ collect X) x| \o hbd" +lemma comp_set_bd_Union_o_collect: "|\\((\f. f x) ` X)| \o hbd \ |(Union \ collect X) x| \o hbd" by (unfold o_apply collect_def SUP_def) lemma wpull_cong: diff -r 88a69da5d3fa -r eff000cab70f src/HOL/Complete_Lattices.thy --- 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 \ \S" notation (xsymbols) - Inter ("\_" [90] 90) + Inter ("\_" [900] 900) lemma Inter_eq: "\A = {x. \B \ A. x \ B}" @@ -934,7 +934,7 @@ "Union S \ \S" notation (xsymbols) - Union ("\_" [90] 90) + Union ("\_" [900] 900) lemma Union_eq: "\A = {x. \B \ A. x \ B}" diff -r 88a69da5d3fa -r eff000cab70f src/HOL/Hoare_Parallel/OG_Tran.thy --- 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 \ \s. {t. (Some c, s) -*\ (None, t)}" definition ann_SEM :: "'a ann_com \ 'a set \ 'a set" where - "ann_SEM c S \ \ann_sem c ` S" + "ann_SEM c S \ \(ann_sem c ` S)" definition sem :: "'a com \ 'a \ 'a set" where "sem c \ \s. {t. \Ts. (c, s) -P*\ (Parallel Ts, t) \ All_None Ts}" definition SEM :: "'a com \ 'a set \ 'a set" where - "SEM c S \ \sem c ` S " + "SEM c S \ \(sem c ` S)" abbreviation Omega :: "'a com" ("\" 63) where "\ \ While UNIV UNIV (Basic id)" diff -r 88a69da5d3fa -r eff000cab70f src/HOL/List.thy --- 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 "\i. i \Collect (?k_list k) \ card (A - set i) = card A - k" by (simp add: card_Diff_subset distinct_card) moreover have "{xs. ?k_list (Suc k) xs} = - (\(xs, n). n#xs) ` \(\xs. {xs} \ (A - set xs)) ` {xs. ?k_list k xs}" + (\(xs, n). n#xs) ` \((\xs. {xs} \ (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 diff -r 88a69da5d3fa -r eff000cab70f src/HOL/Multivariate_Analysis/Integration.thy --- 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 \ p"]) proof - have *: "\s f t. s \ {} \ (\i\s. f i \ i = t) \ t = \ (f ` s) \ (\s)" by auto - have *: "{a..b} = \ (\i. \(q i - {i})) ` p \ \p" + have *: "{a..b} = \ ((\i. \(q i - {i})) ` p) \ \p" apply (rule *[OF False]) proof fix i @@ -1032,7 +1032,7 @@ fix k assume k: "k\p" have *: "\u t s. u \ s \ s \ t = {} \ u \ t = {}" by auto - show "interior (\(\i. \(q i - {i})) ` p) \ interior k = {}" + show "interior (\ ((\i. \(q i - {i})) ` p)) \ interior k = {}" apply (rule *[of _ "interior (\(q k - {k}))"]) defer apply (subst Int_commute) @@ -1044,7 +1044,7 @@ show "\t\q k - {k}. interior k \ interior t = {}" using qk(5) using q(2)[OF k] by auto have *: "\x s. x \ s \ \s \ x" by auto - show "interior (\(\i. \(q i - {i})) ` p) \ interior (\(q k - {k}))" + show "interior (\ ((\i. \(q i - {i})) ` p)) \ interior (\(q k - {k}))" apply (rule interior_mono *)+ using k by auto qed @@ -1131,7 +1131,7 @@ lemma elementary_union_interval: assumes "p division_of \p" obtains q where "q division_of ({a..b::'a::ordered_euclidean_space} \ \p)" proof- note assm=division_ofD[OF assms] - have lem1:"\f s. \\ (f ` s) = \(\x.\(f x)) ` s" by auto + have lem1:"\f s. \\ (f ` s) = \((\x.\(f x)) ` s)" by auto have lem2:"\f s. f \ {} \ \{s \ t |t. t \ f} = s \ \f" by auto { presume "p={} \ thesis" "{a..b} = {} \ thesis" "{a..b} \ {} \ interior {a..b} = {} \ thesis" "p\{} \ interior {a..b}\{} \ {a..b} \ {} \ 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 "\snd ` s = i" "finite (snd ` s)" using assm by auto + show "\(snd ` s) = i" "finite (snd ` s)" using assm by auto fix k assume k:"k \ snd ` s" then obtain xk where xk:"(xk, k) \ s" by auto thus "k \ i" "k \ {}" "\a b. k = {a..b}" using assm apply- by fastforce+ fix k' assume k':"k' \ snd ` s" "k \ k'" from this(1) obtain xk' where xk':"(xk', k') \ 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 \(snd ` s)" proof(rule division_ofI) note assm=tagged_partial_division_ofD[OF assms] - show "finite (snd ` s)" "\snd ` s = \snd ` s" using assm by auto + show "finite (snd ` s)" "\(snd ` s) = \(snd ` s)" using assm by auto fix k assume k:"k \ snd ` s" then obtain xk where xk:"(xk, k) \ s" by auto - thus "k\{}" "\a b. k = {a..b}" "k \ \snd ` s" using assm by auto + thus "k\{}" "\a b. k = {a..b}" "k \ \(snd ` s)" using assm by auto fix k' assume k':"k' \ snd ` s" "k \ k'" from this(1) obtain xk' where xk':"(xk', k') \ s" by auto thus "interior k \ interior k' = {}" apply-apply(rule assm(5)) apply(rule xk xk')+ using k' by auto qed @@ -1355,13 +1355,13 @@ shows "\(pfn ` iset) tagged_division_of (\iset)" proof(rule tagged_division_ofI) note assm = tagged_division_ofD[OF assms(2)[rule_format]] - show "finite (\pfn ` iset)" apply(rule finite_Union) using assms by auto - have "\{k. \x. (x, k) \ \pfn ` iset} = \(\i. \{k. \x. (x, k) \ pfn i}) ` iset" by blast + show "finite (\(pfn ` iset))" apply(rule finite_Union) using assms by auto + have "\{k. \x. (x, k) \ \(pfn ` iset)} = \((\i. \{k. \x. (x, k) \ pfn i}) ` iset)" by blast also have "\ = \iset" using assm(6) by auto - finally show "\{k. \x. (x, k) \ \pfn ` iset} = \iset" . - fix x k assume xk:"(x,k)\\pfn ` iset" then obtain i where i:"i \ iset" "(x, k) \ pfn i" by auto + finally show "\{k. \x. (x, k) \ \(pfn ` iset)} = \iset" . + fix x k assume xk:"(x,k)\\(pfn ` iset)" then obtain i where i:"i \ iset" "(x, k) \ pfn i" by auto show "x\k" "\a b. k = {a..b}" "k \ \iset" using assm(2-4)[OF i] using i(1) by auto - fix x' k' assume xk':"(x',k')\\pfn ` iset" "(x, k) \ (x', k')" then obtain i' where i':"i' \ iset" "(x', k') \ pfn i'" by auto + fix x' k' assume xk':"(x',k')\\(pfn ` iset)" "(x, k) \ (x', k')" then obtain i' where i':"i' \ iset" "(x', k') \ pfn i'" by auto have *:"\a b. i\i' \ a\ i \ b\ i' \ interior a \ interior b = {}" using i(1) i'(1) using assms(3)[rule_format] interior_mono by blast show "interior k \ interior k' = {}" apply(cases "i=i'") @@ -4975,7 +4975,7 @@ shows "norm(setsum (\(x,k). content k *\<^sub>R f x - integral k f) p) \ e" (is "?x \ e") proof- { presume "\k. 0 ?x \ 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 "\snd ` p \ {a..b}" using p'(3) by fastforce + have "\(snd ` p) \ {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 \ "q - snd ` p" have "snd ` p \ 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 \ \qq ` r" have "norm ((\(x, k)\?p. content k *\<^sub>R f x) - integral {a..b} f) < e" + let ?p = "p \ \(qq ` r)" have "norm ((\(x, k)\?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 \ \qq ` r tagged_division_of \snd ` p \ \r" + have "p \ \(qq ` r) tagged_division_of \(snd ` p) \ \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 "\snd ` p \ \r = {a..b}" "{qq i |i. i \ r} = qq ` r" + moreover have "\(snd ` p) \ \r = {a..b}" "{qq i |i. i \ 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 ((\(x, k)\p. content k *\<^sub>R f x) + (\(x, k)\\qq ` r. content k *\<^sub>R f x) - + hence "norm ((\(x, k)\p. content k *\<^sub>R f x) + (\(x, k)\\(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)\p" "(x,l)\qq k" "k\r" diff -r 88a69da5d3fa -r eff000cab70f src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- 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 \ A" thus "x \ a" "open a" using A'(4)[OF open_UNIV] by (auto simp: A_def intro: A' from_nat_into) next - let ?int = "\N. \from_nat_into A' ` N" + let ?int = "\N. \(from_nat_into A' ` N)" fix a b assume "a \ A" "b \ A" then obtain N M where "a = ?int N" "b = ?int M" "finite (N \ M)" by (auto simp: A_def) thus "a \ b \ A" by (auto simp: A_def intro!: image_eqI[where x="N \ M"]) @@ -2617,7 +2617,7 @@ proof (safe intro!: compact_eq_heine_borel[THEN iffD2]) fix A assume "compact U" and A: "\a\A. closed a" "U \ \A = {}" and fi: "\B \ A. finite B \ U \ \B \ {}" - from A have "(\a\uminus`A. open a) \ U \ \uminus`A" + from A have "(\a\uminus`A. open a) \ U \ \(uminus`A)" by auto with `compact U` obtain B where "B \ A" "finite (uminus`B)" "U \ \(uminus`B)" unfolding compact_eq_heine_borel by (metis subset_image_iff) @@ -2627,7 +2627,7 @@ fix A assume ?R and cover: "\a\A. open a" "U \ \A" from cover have "U \ \(uminus`A) = {}" "\a\uminus`A. closed a" by auto - with `?R` obtain B where "B \ A" "finite (uminus`B)" "U \ \uminus`B = {}" + with `?R` obtain B where "B \ A" "finite (uminus`B)" "U \ \(uminus`B) = {}" by (metis subset_image_iff) then show "\T\A. finite T \ U \ \T" by (auto intro!: exI[of _ B] inj_setminus dest: finite_imageD) @@ -3041,7 +3041,7 @@ assumes "seq_compact s" shows "\e>0. \k. finite k \ k \ s \ s \ (\((\x. ball x e) ` k))" proof(rule, rule, rule ccontr) - fix e::real assume "e>0" and assm:"\ (\k. finite k \ k \ s \ s \ \(\x. ball x e) ` k)" + fix e::real assume "e>0" and assm:"\ (\k. finite k \ k \ s \ s \ \((\x. ball x e) ` k))" def x \ "helper_1 s e" { fix n have "x n \ s \ (\m dist (x m) (x n) < e)" @@ -4586,7 +4586,7 @@ moreover with c have "(\y\D. a y) \ t \ (\y\D. c y)" by (fastforce simp: subset_eq) ultimately show "\a. open a \ x \ a \ (\d\C. finite d \ a \ t \ \d)" - using c by (intro exI[of _ "c`D"] exI[of _ "\a`D"] conjI) (auto intro!: open_INT) + using c by (intro exI[of _ "c`D"] exI[of _ "\(a`D)"] conjI) (auto intro!: open_INT) qed then obtain a d where a: "\x\s. open (a x)" "s \ (\x\s. a x)" and d: "\x. x \ s \ d x \ C \ finite (d x) \ a x \ t \ \d x" diff -r 88a69da5d3fa -r eff000cab70f src/HOL/Probability/Caratheodory.thy --- 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: "\i. i \ I \ A i \ M" "disjoint_family_on A I" "finite I" "UNION I A \ M" shows "f (UNION I A) = (\i\I. f (A i))" proof - - have "A`I \ M" "disjoint (A`I)" "finite (A`I)" "\A`I \ M" + have "A`I \ M" "disjoint (A`I)" "finite (A`I)" "\(A`I) \ M" using A unfolding SUP_def by (auto simp: disjoint_family_on_disjoint_image) - with `volume M f` have "f (\A`I) = (\a\A`I. f a)" + with `volume M f` have "f (\(A`I)) = (\a\A`I. f a)" unfolding volume_def by blast also have "\ = (\i\I. f (A i))" proof (subst setsum_reindex_nonzero) diff -r 88a69da5d3fa -r eff000cab70f src/HOL/Probability/Radon_Nikodym.thy --- 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)) \ ?a" unfolding SUP_def proof (safe intro!: Sup_mono, unfold bex_simps) fix i - have *: "(\Q' ` {..i}) = ?O i" by auto + have *: "(\(Q' ` {..i})) = ?O i" by auto then show "\x. (x \ sets M \ N x \ \) \ - emeasure M (\Q' ` {..i}) \ emeasure M x" + emeasure M (\(Q' ` {..i})) \ emeasure M x" using O_in_G[of i] by (auto intro!: exI[of _ "?O i"]) qed qed diff -r 88a69da5d3fa -r eff000cab70f src/HOL/Probability/Regularity.thy --- 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 \ \(\x. ball x e') ` ?k" unfolding K_def B_def using n by force - ultimately show "\k. finite k \ K \ \(\x. ball x e') ` k" by blast + moreover have "K \ \((\x. ball x e') ` ?k)" unfolding K_def B_def using n by force + ultimately show "\k. finite k \ K \ \((\x. ball x e') ` k)" by blast qed ultimately show "?thesis e " by (auto simp: sU)