# HG changeset patch # User wenzelm # Date 1348767329 -7200 # Node ID 9d34cfe1dbdf793084bf7ec7ef29e238f1fa40e6 # Parent 1a0f25c38629c61d69e587ea85aaf971f376671e# Parent 0009a6ebc83bbcc84eeb2003daf36228629ba968 merged diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/HOL/Library/Convex.thy --- a/src/HOL/Library/Convex.thy Thu Sep 27 18:58:15 2012 +0200 +++ b/src/HOL/Library/Convex.thy Thu Sep 27 19:35:29 2012 +0200 @@ -11,9 +11,8 @@ subsection {* Convexity. *} -definition - convex :: "'a::real_vector set \ bool" where - "convex s \ (\x\s. \y\s. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s)" +definition convex :: "'a::real_vector set \ bool" + where "convex s \ (\x\s. \y\s. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s)" lemma convex_alt: "convex s \ (\x\s. \y\s. \u. 0 \ u \ u \ 1 \ ((1 - u) *\<^sub>R x + u *\<^sub>R y) \ s)" @@ -21,10 +20,10 @@ proof assume alt[rule_format]: ?alt { fix x y and u v :: real assume mem: "x \ s" "y \ s" - assume "0 \ u" "0 \ v" "u + v = 1" - moreover hence "u = 1 - v" by auto + assume "0 \ u" "0 \ v" + moreover assume "u + v = 1" then have "u = 1 - v" by auto ultimately have "u *\<^sub>R x + v *\<^sub>R y \ s" using alt[OF mem] by auto } - thus "convex s" unfolding convex_def by auto + then show "convex s" unfolding convex_def by auto qed (auto simp: convex_def) lemma mem_convex: @@ -53,13 +52,13 @@ lemma convex_halfspace_ge: "convex {x. inner a x \ b}" proof - - have *:"{x. inner a x \ b} = {x. inner (-a) x \ -b}" by auto + have *: "{x. inner a x \ b} = {x. inner (-a) x \ -b}" by auto show ?thesis unfolding * using convex_halfspace_le[of "-a" "-b"] by auto qed lemma convex_hyperplane: "convex {x. inner a x = b}" -proof- - have *:"{x. inner a x = b} = {x. inner a x \ b} \ {x. inner a x \ b}" by auto +proof - + have *: "{x. inner a x = b} = {x. inner a x \ b} \ {x. inner a x \ b}" by auto show ?thesis using convex_halfspace_le convex_halfspace_ge by (auto intro!: convex_Int simp: *) qed @@ -74,78 +73,83 @@ lemma convex_real_interval: fixes a b :: "real" shows "convex {a..}" and "convex {..b}" - and "convex {a<..}" and "convex {.. inner 1 x}" by auto - thus 1: "convex {a..}" by (simp only: convex_halfspace_ge) + then show 1: "convex {a..}" by (simp only: convex_halfspace_ge) have "{..b} = {x. inner 1 x \ b}" by auto - thus 2: "convex {..b}" by (simp only: convex_halfspace_le) + then show 2: "convex {..b}" by (simp only: convex_halfspace_le) have "{a<..} = {x. a < inner 1 x}" by auto - thus 3: "convex {a<..}" by (simp only: convex_halfspace_gt) + then show 3: "convex {a<..}" by (simp only: convex_halfspace_gt) have "{.. {..b}" by auto - thus "convex {a..b}" by (simp only: convex_Int 1 2) + then show "convex {a..b}" by (simp only: convex_Int 1 2) have "{a<..b} = {a<..} \ {..b}" by auto - thus "convex {a<..b}" by (simp only: convex_Int 3 2) + then show "convex {a<..b}" by (simp only: convex_Int 3 2) have "{a.. {.. {.. i \ s. a i) = 1" - assumes "\ i. i \ s \ a i \ 0" and "\ i. i \ s \ y i \ C" + assumes "\i. i \ s \ a i \ 0" and "\i. i \ s \ y i \ C" shows "(\ j \ s. a j *\<^sub>R y j) \ C" -using assms -proof (induct s arbitrary:a rule:finite_induct) - case empty thus ?case by auto + using assms +proof (induct s arbitrary:a rule: finite_induct) + case empty + then show ?case by auto next case (insert i s) note asms = this { assume "a i = 1" - hence "(\ j \ s. a j) = 0" + then have "(\ j \ s. a j) = 0" using asms by auto - hence "\ j. j \ s \ a j = 0" + then have "\j. j \ s \ a j = 0" using setsum_nonneg_0[where 'b=real] asms by fastforce - hence ?case using asms by auto } + then have ?case using asms by auto } moreover { assume asm: "a i \ 1" from asms have yai: "y i \ C" "a i \ 0" by auto have fis: "finite (insert i s)" using asms by auto - hence ai1: "a i \ 1" using setsum_nonneg_leq_bound[of "insert i s" a 1] asms by simp - hence "a i < 1" using asm by auto - hence i0: "1 - a i > 0" by auto - let "?a j" = "a j / (1 - a i)" + then have ai1: "a i \ 1" using setsum_nonneg_leq_bound[of "insert i s" a 1] asms by simp + then have "a i < 1" using asm by auto + then have i0: "1 - a i > 0" by auto + let ?a = "\j. a j / (1 - a i)" { fix j assume "j \ s" - hence "?a j \ 0" + then have "?a j \ 0" using i0 asms divide_nonneg_pos - by fastforce } note a_nonneg = this + by fastforce + } note a_nonneg = this have "(\ j \ insert i s. a j) = 1" using asms by auto - hence "(\ j \ s. a j) = 1 - a i" using setsum.insert asms by fastforce - hence "(\ j \ s. a j) / (1 - a i) = 1" using i0 by auto - hence a1: "(\ j \ s. ?a j) = 1" unfolding setsum_divide_distrib by simp - from this asms - have "(\j\s. ?a j *\<^sub>R y j) \ C" using a_nonneg by fastforce - hence "a i *\<^sub>R y i + (1 - a i) *\<^sub>R (\ j \ s. ?a j *\<^sub>R y j) \ C" + then have "(\ j \ s. a j) = 1 - a i" using setsum.insert asms by fastforce + then have "(\ j \ s. a j) / (1 - a i) = 1" using i0 by auto + then have a1: "(\ j \ s. ?a j) = 1" unfolding setsum_divide_distrib by simp + with asms have "(\j\s. ?a j *\<^sub>R y j) \ C" using a_nonneg by fastforce + then have "a i *\<^sub>R y i + (1 - a i) *\<^sub>R (\ j \ s. ?a j *\<^sub>R y j) \ C" using asms[unfolded convex_def, rule_format] yai ai1 by auto - hence "a i *\<^sub>R y i + (\ j \ s. (1 - a i) *\<^sub>R (?a j *\<^sub>R y j)) \ C" + then have "a i *\<^sub>R y i + (\ j \ s. (1 - a i) *\<^sub>R (?a j *\<^sub>R y j)) \ C" using scaleR_right.setsum[of "(1 - a i)" "\ j. ?a j *\<^sub>R y j" s] by auto - hence "a i *\<^sub>R y i + (\ j \ s. a j *\<^sub>R y j) \ C" using i0 by auto - hence ?case using setsum.insert asms by auto } + then have "a i *\<^sub>R y i + (\ j \ s. a j *\<^sub>R y j) \ C" using i0 by auto + then have ?case using setsum.insert asms by auto + } ultimately show ?case by auto qed lemma convex: - shows "convex s \ (\(k::nat) u x. (\i. 1\i \ i\k \ 0 \ u i \ x i \s) \ (setsum u {1..k} = 1) - \ setsum (\i. u i *\<^sub>R x i) {1..k} \ s)" + "convex s \ (\(k::nat) u x. (\i. 1\i \ i\k \ 0 \ u i \ x i \s) \ (setsum u {1..k} = 1) + \ setsum (\i. u i *\<^sub>R x i) {1..k} \ s)" proof safe - fix k :: nat fix u :: "nat \ real" fix x + fix k :: nat + fix u :: "nat \ real" + fix x assume "convex s" "\i. 1 \ i \ i \ k \ 0 \ u i \ x i \ s" "setsum u {1..k} = 1" @@ -154,35 +158,39 @@ next assume asm: "\k u x. (\ i :: nat. 1 \ i \ i \ k \ 0 \ u i \ x i \ s) \ setsum u {1..k} = 1 \ (\i = 1..k. u i *\<^sub>R (x i :: 'a)) \ s" - { fix \ :: real fix x y :: 'a assume xy: "x \ s" "y \ s" assume mu: "\ \ 0" "\ \ 1" - let "?u i" = "if (i :: nat) = 1 then \ else 1 - \" - let "?x i" = "if (i :: nat) = 1 then x else y" + { fix \ :: real + fix x y :: 'a + assume xy: "x \ s" "y \ s" + assume mu: "\ \ 0" "\ \ 1" + let ?u = "\i. if (i :: nat) = 1 then \ else 1 - \" + let ?x = "\i. if (i :: nat) = 1 then x else y" have "{1 :: nat .. 2} \ - {x. x = 1} = {2}" by auto - hence card: "card ({1 :: nat .. 2} \ - {x. x = 1}) = 1" by simp - hence "setsum ?u {1 .. 2} = 1" + then have card: "card ({1 :: nat .. 2} \ - {x. x = 1}) = 1" by simp + then have "setsum ?u {1 .. 2} = 1" using setsum_cases[of "{(1 :: nat) .. 2}" "\ x. x = 1" "\ x. \" "\ x. 1 - \"] by auto - from this asm[rule_format, of "2" ?u ?x] - have s: "(\j \ {1..2}. ?u j *\<^sub>R ?x j) \ s" + with asm[rule_format, of "2" ?u ?x] have s: "(\j \ {1..2}. ?u j *\<^sub>R ?x j) \ s" using mu xy by auto have grarr: "(\j \ {Suc (Suc 0)..2}. ?u j *\<^sub>R ?x j) = (1 - \) *\<^sub>R y" using setsum_head_Suc[of "Suc (Suc 0)" 2 "\ j. (1 - \) *\<^sub>R y"] by auto from setsum_head_Suc[of "Suc 0" 2 "\ j. ?u j *\<^sub>R ?x j", simplified this] have "(\j \ {1..2}. ?u j *\<^sub>R ?x j) = \ *\<^sub>R x + (1 - \) *\<^sub>R y" by auto - hence "(1 - \) *\<^sub>R y + \ *\<^sub>R x \ s" using s by (auto simp:add_commute) } - thus "convex s" unfolding convex_alt by auto + then have "(1 - \) *\<^sub>R y + \ *\<^sub>R x \ s" using s by (auto simp:add_commute) + } + then show "convex s" unfolding convex_alt by auto qed lemma convex_explicit: fixes s :: "'a::real_vector set" shows "convex s \ - (\t u. finite t \ t \ s \ (\x\t. 0 \ u x) \ setsum u t = 1 \ setsum (\x. u x *\<^sub>R x) t \ s)" + (\t u. finite t \ t \ s \ (\x\t. 0 \ u x) \ setsum u t = 1 \ setsum (\x. u x *\<^sub>R x) t \ s)" proof safe - fix t fix u :: "'a \ real" + fix t + fix u :: "'a \ real" assume "convex s" "finite t" "t \ s" "\x\t. 0 \ u x" "setsum u t = 1" - thus "(\x\t. u x *\<^sub>R x) \ s" + then show "(\x\t. u x *\<^sub>R x) \ s" using convex_setsum[of t s u "\ x. x"] by auto next assume asm0: "\t. \ u. finite t \ t \ s \ (\x\t. 0 \ u x) @@ -190,39 +198,42 @@ show "convex s" unfolding convex_alt proof safe - fix x y fix \ :: real + fix x y + fix \ :: real assume asm: "x \ s" "y \ s" "0 \ \" "\ \ 1" { assume "x \ y" - hence "(1 - \) *\<^sub>R x + \ *\<^sub>R y \ s" + then have "(1 - \) *\<^sub>R x + \ *\<^sub>R y \ s" using asm0[rule_format, of "{x, y}" "\ z. if z = x then 1 - \ else \"] asm by auto } moreover { assume "x = y" - hence "(1 - \) *\<^sub>R x + \ *\<^sub>R y \ s" + then have "(1 - \) *\<^sub>R x + \ *\<^sub>R y \ s" using asm0[rule_format, of "{x, y}" "\ z. 1"] asm by (auto simp:field_simps real_vector.scale_left_diff_distrib) } ultimately show "(1 - \) *\<^sub>R x + \ *\<^sub>R y \ s" by blast qed qed -lemma convex_finite: assumes "finite s" +lemma convex_finite: + assumes "finite s" shows "convex s \ (\u. (\x\s. 0 \ u x) \ setsum u s = 1 \ setsum (\x. u x *\<^sub>R x) s \ s)" unfolding convex_explicit -proof (safe) - fix t u assume sum: "\u. (\x\s. 0 \ u x) \ setsum u s = 1 \ (\x\s. u x *\<^sub>R x) \ s" +proof safe + fix t u + assume sum: "\u. (\x\s. 0 \ u x) \ setsum u s = 1 \ (\x\s. u x *\<^sub>R x) \ s" and as: "finite t" "t \ s" "\x\t. 0 \ u x" "setsum u t = (1::real)" - have *:"s \ t = t" using as(2) by auto - have if_distrib_arg: "\P f g x. (if P then f else g) x = (if P then f x else g x)" by simp + have *: "s \ t = t" using as(2) by auto + have if_distrib_arg: "\P f g x. (if P then f else g) x = (if P then f x else g x)" + by simp show "(\x\t. u x *\<^sub>R x) \ s" using sum[THEN spec[where x="\x. if x\t then u x else 0"]] as * by (auto simp: assms setsum_cases if_distrib if_distrib_arg) qed (erule_tac x=s in allE, erule_tac x=u in allE, auto) -definition - convex_on :: "'a::real_vector set \ ('a \ real) \ bool" where - "convex_on s f \ - (\x\s. \y\s. \u\0. \v\0. u + v = 1 \ f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y)" +definition convex_on :: "'a::real_vector set \ ('a \ real) \ bool" + where "convex_on s f \ + (\x\s. \y\s. \u\0. \v\0. u + v = 1 \ f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y)" lemma convex_on_subset: "convex_on t f \ s \ t \ convex_on s f" unfolding convex_on_def by auto @@ -230,21 +241,29 @@ lemma convex_add[intro]: assumes "convex_on s f" "convex_on s g" shows "convex_on s (\x. f x + g x)" -proof- - { fix x y assume "x\s" "y\s" moreover - fix u v ::real assume "0 \ u" "0 \ v" "u + v = 1" - ultimately have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ (u * f x + v * f y) + (u * g x + v * g y)" - using assms unfolding convex_on_def by (auto simp add:add_mono) - hence "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ u * (f x + g x) + v * (f y + g y)" by (simp add: field_simps) } - thus ?thesis unfolding convex_on_def by auto +proof - + { fix x y + assume "x\s" "y\s" + moreover + fix u v :: real + assume "0 \ u" "0 \ v" "u + v = 1" + ultimately + have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ (u * f x + v * f y) + (u * g x + v * g y)" + using assms unfolding convex_on_def by (auto simp add: add_mono) + then have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ u * (f x + g x) + v * (f y + g y)" + by (simp add: field_simps) + } + then show ?thesis unfolding convex_on_def by auto qed lemma convex_cmul[intro]: assumes "0 \ (c::real)" "convex_on s f" shows "convex_on s (\x. c * f x)" proof- - have *:"\u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" by (simp add: field_simps) - show ?thesis using assms(2) and mult_left_mono [OF _ assms(1)] unfolding convex_on_def and * by auto + have *: "\u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" + by (simp add: field_simps) + show ?thesis using assms(2) and mult_left_mono [OF _ assms(1)] + unfolding convex_on_def and * by auto qed lemma convex_lower: @@ -254,7 +273,7 @@ let ?m = "max (f x) (f y)" have "u * f x + v * f y \ u * max (f x) (f y) + v * max (f x) (f y)" using assms(4,5) by (auto simp add: mult_left_mono add_mono) - also have "\ = max (f x) (f y)" using assms(6) unfolding distrib[THEN sym] by auto + also have "\ = max (f x) (f y)" using assms(6) unfolding distrib[symmetric] by auto finally show ?thesis using assms unfolding convex_on_def by fastforce qed @@ -262,24 +281,30 @@ lemma convex_distance[intro]: fixes s :: "'a::real_normed_vector set" shows "convex_on s (\x. dist a x)" -proof(auto simp add: convex_on_def dist_norm) - fix x y assume "x\s" "y\s" - fix u v ::real assume "0 \ u" "0 \ v" "u + v = 1" - have "a = u *\<^sub>R a + v *\<^sub>R a" unfolding scaleR_left_distrib[THEN sym] and `u+v=1` by simp - hence *:"a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))" +proof (auto simp add: convex_on_def dist_norm) + fix x y + assume "x\s" "y\s" + fix u v :: real + assume "0 \ u" "0 \ v" "u + v = 1" + have "a = u *\<^sub>R a + v *\<^sub>R a" + unfolding scaleR_left_distrib[symmetric] and `u+v=1` by simp + then have *: "a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))" by (auto simp add: algebra_simps) show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \ u * norm (a - x) + v * norm (a - y)" unfolding * using norm_triangle_ineq[of "u *\<^sub>R (a - x)" "v *\<^sub>R (a - y)"] using `0 \ u` `0 \ v` by auto qed + subsection {* Arithmetic operations on sets preserve convexity. *} + lemma convex_scaling: assumes "convex s" shows"convex ((\x. c *\<^sub>R x) ` s)" -using assms unfolding convex_def image_iff + using assms unfolding convex_def image_iff proof safe - fix x xa y xb :: "'a::real_vector" fix u v :: real + fix x xa y xb :: "'a::real_vector" + fix u v :: real assume asm: "\x\s. \y\s. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s" "xa \ s" "xb \ s" "0 \ u" "0 \ v" "u + v = 1" show "\x\s. u *\<^sub>R c *\<^sub>R xa + v *\<^sub>R c *\<^sub>R xb = c *\<^sub>R x" @@ -287,9 +312,10 @@ qed lemma convex_negations: "convex s \ convex ((\x. -x)` s)" -using assms unfolding convex_def image_iff + using assms unfolding convex_def image_iff proof safe - fix x xa y xb :: "'a::real_vector" fix u v :: real + fix x xa y xb :: "'a::real_vector" + fix u v :: real assume asm: "\x\s. \y\s. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s" "xa \ s" "xb \ s" "0 \ u" "0 \ v" "u + v = 1" show "\x\s. u *\<^sub>R - xa + v *\<^sub>R - xb = - x" @@ -299,10 +325,12 @@ lemma convex_sums: assumes "convex s" "convex t" shows "convex {x + y| x y. x \ s \ y \ t}" -using assms unfolding convex_def image_iff + using assms unfolding convex_def image_iff proof safe - fix xa xb ya yb assume xy:"xa\s" "xb\s" "ya\t" "yb\t" - fix u v ::real assume uv:"0 \ u" "0 \ v" "u + v = 1" + fix xa xb ya yb + assume xy:"xa\s" "xb\s" "ya\t" "yb\t" + fix u v :: real + assume uv: "0 \ u" "0 \ v" "u + v = 1" show "\x y. u *\<^sub>R (xa + ya) + v *\<^sub>R (xb + yb) = x + y \ x \ s \ y \ t" using exI[of _ "u *\<^sub>R xa + v *\<^sub>R xb"] exI[of _ "u *\<^sub>R ya + v *\<^sub>R yb"] assms[unfolded convex_def] uv xy by (auto simp add:scaleR_right_distrib) @@ -314,105 +342,120 @@ proof - have "{x - y| x y. x \ s \ y \ t} = {x + y |x y. x \ s \ y \ uminus ` t}" proof safe - fix x x' y assume "x' \ s" "y \ t" - thus "\x y'. x' - y = x + y' \ x \ s \ y' \ uminus ` t" + fix x x' y + assume "x' \ s" "y \ t" + then show "\x y'. x' - y = x + y' \ x \ s \ y' \ uminus ` t" using exI[of _ x'] exI[of _ "-y"] by auto next - fix x x' y y' assume "x' \ s" "y' \ t" - thus "\x y. x' + - y' = x - y \ x \ s \ y \ t" + fix x x' y y' + assume "x' \ s" "y' \ t" + then show "\x y. x' + - y' = x - y \ x \ s \ y \ t" using exI[of _ x'] exI[of _ y'] by auto qed - thus ?thesis using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto + then show ?thesis + using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto qed -lemma convex_translation: assumes "convex s" shows "convex ((\x. a + x) ` s)" -proof- have "{a + y |y. y \ s} = (\x. a + x) ` s" by auto - thus ?thesis using convex_sums[OF convex_singleton[of a] assms] by auto qed +lemma convex_translation: + assumes "convex s" + shows "convex ((\x. a + x) ` s)" +proof - + have "{a + y |y. y \ s} = (\x. a + x) ` s" by auto + then show ?thesis + using convex_sums[OF convex_singleton[of a] assms] by auto +qed -lemma convex_affinity: assumes "convex s" shows "convex ((\x. a + c *\<^sub>R x) ` s)" -proof- have "(\x. a + c *\<^sub>R x) ` s = op + a ` op *\<^sub>R c ` s" by auto - thus ?thesis using convex_translation[OF convex_scaling[OF assms], of a c] by auto qed +lemma convex_affinity: + assumes "convex s" + shows "convex ((\x. a + c *\<^sub>R x) ` s)" +proof - + have "(\x. a + c *\<^sub>R x) ` s = op + a ` op *\<^sub>R c ` s" by auto + then show ?thesis + using convex_translation[OF convex_scaling[OF assms], of a c] by auto +qed lemma convex_linear_image: assumes c:"convex s" and l:"bounded_linear f" shows "convex(f ` s)" -proof(auto simp add: convex_def) +proof (auto simp add: convex_def) interpret f: bounded_linear f by fact - fix x y assume xy:"x \ s" "y \ s" - fix u v ::real assume uv:"0 \ u" "0 \ v" "u + v = 1" + fix x y + assume xy: "x \ s" "y \ s" + fix u v :: real + assume uv: "0 \ u" "0 \ v" "u + v = 1" show "u *\<^sub>R f x + v *\<^sub>R f y \ f ` s" unfolding image_iff using bexI[of _ "u *\<^sub>R x + v *\<^sub>R y"] f.add f.scaleR c[unfolded convex_def] xy uv by auto qed -lemma pos_is_convex: - shows "convex {0 :: real <..}" -unfolding convex_alt +lemma pos_is_convex: "convex {0 :: real <..}" + unfolding convex_alt proof safe fix y x \ :: real assume asms: "y > 0" "x > 0" "\ \ 0" "\ \ 1" { assume "\ = 0" - hence "\ *\<^sub>R x + (1 - \) *\<^sub>R y = y" by simp - hence "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using asms by simp } + then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y = y" by simp + then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using asms by simp } moreover { assume "\ = 1" - hence "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using asms by simp } + then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using asms by simp } moreover { assume "\ \ 1" "\ \ 0" - hence "\ > 0" "(1 - \) > 0" using asms by auto - hence "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using asms + then have "\ > 0" "(1 - \) > 0" using asms by auto + then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using asms by (auto simp add: add_pos_pos mult_pos_pos) } ultimately show "(1 - \) *\<^sub>R y + \ *\<^sub>R x > 0" using assms by fastforce qed lemma convex_on_setsum: fixes a :: "'a \ real" - fixes y :: "'a \ 'b::real_vector" - fixes f :: "'b \ real" + and y :: "'a \ 'b::real_vector" + and f :: "'b \ real" assumes "finite s" "s \ {}" - assumes "convex_on C f" - assumes "convex C" - assumes "(\ i \ s. a i) = 1" - assumes "\ i. i \ s \ a i \ 0" - assumes "\ i. i \ s \ y i \ C" + and "convex_on C f" + and "convex C" + and "(\ i \ s. a i) = 1" + and "\i. i \ s \ a i \ 0" + and "\i. i \ s \ y i \ C" shows "f (\ i \ s. a i *\<^sub>R y i) \ (\ i \ s. a i * f (y i))" -using assms -proof (induct s arbitrary:a rule:finite_ne_induct) + using assms +proof (induct s arbitrary: a rule: finite_ne_induct) case (singleton i) - hence ai: "a i = 1" by auto - thus ?case by auto + then have ai: "a i = 1" by auto + then show ?case by auto next case (insert i s) note asms = this - hence "convex_on C f" by simp + then have "convex_on C f" by simp from this[unfolded convex_on_def, rule_format] - have conv: "\ x y \. \x \ C; y \ C; 0 \ \; \ \ 1\ - \ f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" + have conv: "\x y \. x \ C \ y \ C \ 0 \ \ \ \ \ 1 + \ f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" by simp { assume "a i = 1" - hence "(\ j \ s. a j) = 0" + then have "(\ j \ s. a j) = 0" using asms by auto - hence "\ j. j \ s \ a j = 0" + then have "\j. j \ s \ a j = 0" using setsum_nonneg_0[where 'b=real] asms by fastforce - hence ?case using asms by auto } + then have ?case using asms by auto } moreover { assume asm: "a i \ 1" from asms have yai: "y i \ C" "a i \ 0" by auto have fis: "finite (insert i s)" using asms by auto - hence ai1: "a i \ 1" using setsum_nonneg_leq_bound[of "insert i s" a] asms by simp - hence "a i < 1" using asm by auto - hence i0: "1 - a i > 0" by auto - let "?a j" = "a j / (1 - a i)" + then have ai1: "a i \ 1" using setsum_nonneg_leq_bound[of "insert i s" a] asms by simp + then have "a i < 1" using asm by auto + then have i0: "1 - a i > 0" by auto + let ?a = "\j. a j / (1 - a i)" { fix j assume "j \ s" - hence "?a j \ 0" + then have "?a j \ 0" using i0 asms divide_nonneg_pos - by fastforce } note a_nonneg = this + by fastforce } + note a_nonneg = this have "(\ j \ insert i s. a j) = 1" using asms by auto - hence "(\ j \ s. a j) = 1 - a i" using setsum.insert asms by fastforce - hence "(\ j \ s. a j) / (1 - a i) = 1" using i0 by auto - hence a1: "(\ j \ s. ?a j) = 1" unfolding setsum_divide_distrib by simp + then have "(\ j \ s. a j) = 1 - a i" using setsum.insert asms by fastforce + then have "(\ j \ s. a j) / (1 - a i) = 1" using i0 by auto + then have a1: "(\ j \ s. ?a j) = 1" unfolding setsum_divide_distrib by simp have "convex C" using asms by auto - hence asum: "(\ j \ s. ?a j *\<^sub>R y j) \ C" + then have asum: "(\ j \ s. ?a j *\<^sub>R y j) \ C" using asms convex_setsum[OF `finite s` `convex C` a1 a_nonneg] by auto have asum_le: "f (\ j \ s. ?a j *\<^sub>R y j) \ (\ j \ s. ?a j * f (y j))" @@ -423,7 +466,8 @@ also have "\ = f (((1 - a i) * inverse (1 - a i)) *\<^sub>R (\ j \ s. a j *\<^sub>R y j) + a i *\<^sub>R y i)" using i0 by auto also have "\ = f ((1 - a i) *\<^sub>R (\ j \ s. (a j * inverse (1 - a i)) *\<^sub>R y j) + a i *\<^sub>R y i)" - using scaleR_right.setsum[of "inverse (1 - a i)" "\ j. a j *\<^sub>R y j" s, symmetric] by (auto simp:algebra_simps) + using scaleR_right.setsum[of "inverse (1 - a i)" "\ j. a j *\<^sub>R y j" s, symmetric] + by (auto simp:algebra_simps) also have "\ = f ((1 - a i) *\<^sub>R (\ j \ s. ?a j *\<^sub>R y j) + a i *\<^sub>R y i)" by (auto simp: divide_inverse) also have "\ \ (1 - a i) *\<^sub>R f ((\ j \ s. ?a j *\<^sub>R y j)) + a i * f (y i)" @@ -448,27 +492,30 @@ (\ x \ C. \ y \ C. \ \ :: real. \ \ 0 \ \ \ 1 \ f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y)" proof safe - fix x y fix \ :: real + fix x y + fix \ :: real assume asms: "convex_on C f" "x \ C" "y \ C" "0 \ \" "\ \ 1" from this[unfolded convex_on_def, rule_format] - have "\ u v. \0 \ u; 0 \ v; u + v = 1\ \ f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" by auto + have "\u v. \0 \ u; 0 \ v; u + v = 1\ \ f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" by auto from this[of "\" "1 - \", simplified] asms - show "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) - \ \ * f x + (1 - \) * f y" by auto + show "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" by auto next assume asm: "\x\C. \y\C. \\. 0 \ \ \ \ \ 1 \ f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" - {fix x y fix u v :: real + { fix x y + fix u v :: real assume lasm: "x \ C" "y \ C" "u \ 0" "v \ 0" "u + v = 1" - hence[simp]: "1 - u = v" by auto + then have[simp]: "1 - u = v" by auto from asm[rule_format, of x y u] - have "f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" using lasm by auto } - thus "convex_on C f" unfolding convex_on_def by auto + have "f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" using lasm by auto + } + then show "convex_on C f" unfolding convex_on_def by auto qed lemma convex_on_diff: fixes f :: "real \ real" assumes f: "convex_on I f" and I: "x\I" "y\I" and t: "x < t" "t < y" - shows "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" "(f x - f y) / (x - y) \ (f t - f y) / (t - y)" + shows "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" + "(f x - f y) / (x - y) \ (f t - f y) / (t - y)" proof - def a \ "(t - y) / (x - y)" with t have "0 \ a" "0 \ 1 - a" by (auto simp: field_simps) @@ -488,46 +535,48 @@ lemma pos_convex_function: fixes f :: "real \ real" assumes "convex C" - assumes leq: "\ x y. \x \ C ; y \ C\ \ f' x * (y - x) \ f y - f x" + and leq: "\x y. \x \ C ; y \ C\ \ f' x * (y - x) \ f y - f x" shows "convex_on C f" -unfolding convex_on_alt[OF assms(1)] -using assms + unfolding convex_on_alt[OF assms(1)] + using assms proof safe fix x y \ :: real let ?x = "\ *\<^sub>R x + (1 - \) *\<^sub>R y" assume asm: "convex C" "x \ C" "y \ C" "\ \ 0" "\ \ 1" - hence "1 - \ \ 0" by auto - hence xpos: "?x \ C" using asm unfolding convex_alt by fastforce + then have "1 - \ \ 0" by auto + then have xpos: "?x \ C" using asm unfolding convex_alt by fastforce have geq: "\ * (f x - f ?x) + (1 - \) * (f y - f ?x) \ \ * f' ?x * (x - ?x) + (1 - \) * f' ?x * (y - ?x)" using add_mono[OF mult_left_mono[OF leq[OF xpos asm(2)] `\ \ 0`] mult_left_mono[OF leq[OF xpos asm(3)] `1 - \ \ 0`]] by auto - hence "\ * f x + (1 - \) * f y - f ?x \ 0" - by (auto simp add:field_simps) - thus "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" + then have "\ * f x + (1 - \) * f y - f ?x \ 0" + by (auto simp add: field_simps) + then show "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" using convex_on_alt by auto qed lemma atMostAtLeast_subset_convex: fixes C :: "real set" assumes "convex C" - assumes "x \ C" "y \ C" "x < y" + and "x \ C" "y \ C" "x < y" shows "{x .. y} \ C" proof safe fix z assume zasm: "z \ {x .. y}" { assume asm: "x < z" "z < y" - let "?\" = "(y - z) / (y - x)" - have "0 \ ?\" "?\ \ 1" using assms asm by (auto simp add:field_simps) - hence comb: "?\ * x + (1 - ?\) * y \ C" - using assms iffD1[OF convex_alt, rule_format, of C y x ?\] by (simp add:algebra_simps) + let ?\ = "(y - z) / (y - x)" + have "0 \ ?\" "?\ \ 1" using assms asm by (auto simp add: field_simps) + then have comb: "?\ * x + (1 - ?\) * y \ C" + using assms iffD1[OF convex_alt, rule_format, of C y x ?\] + by (simp add: algebra_simps) have "?\ * x + (1 - ?\) * y = (y - z) * x / (y - x) + (1 - (y - z) / (y - x)) * y" - by (auto simp add:field_simps) + by (auto simp add: field_simps) also have "\ = ((y - z) * x + (y - x - (y - z)) * y) / (y - x)" - using assms unfolding add_divide_distrib by (auto simp:field_simps) + using assms unfolding add_divide_distrib by (auto simp: field_simps) also have "\ = z" - using assms by (auto simp:field_simps) + using assms by (auto simp: field_simps) finally have "z \ C" - using comb by auto } note less = this + using comb by auto } + note less = this show "z \ C" using zasm less assms unfolding atLeastAtMost_iff le_less by auto qed @@ -535,21 +584,22 @@ lemma f''_imp_f': fixes f :: "real \ real" assumes "convex C" - assumes f': "\ x. x \ C \ DERIV f x :> (f' x)" - assumes f'': "\ x. x \ C \ DERIV f' x :> (f'' x)" - assumes pos: "\ x. x \ C \ f'' x \ 0" - assumes "x \ C" "y \ C" + and f': "\x. x \ C \ DERIV f x :> (f' x)" + and f'': "\x. x \ C \ DERIV f' x :> (f'' x)" + and pos: "\x. x \ C \ f'' x \ 0" + and "x \ C" "y \ C" shows "f' x * (y - x) \ f y - f x" -using assms + using assms proof - - { fix x y :: real assume asm: "x \ C" "y \ C" "y > x" - hence ge: "y - x > 0" "y - x \ 0" by auto + { fix x y :: real + assume asm: "x \ C" "y \ C" "y > x" + then have ge: "y - x > 0" "y - x \ 0" by auto from asm have le: "x - y < 0" "x - y \ 0" by auto then obtain z1 where z1: "z1 > x" "z1 < y" "f y - f x = (y - x) * f' z1" using subsetD[OF atMostAtLeast_subset_convex[OF `convex C` `x \ C` `y \ C` `x < y`], THEN f', THEN MVT2[OF `x < y`, rule_format, unfolded atLeastAtMost_iff[symmetric]]] by auto - hence "z1 \ C" using atMostAtLeast_subset_convex + then have "z1 \ C" using atMostAtLeast_subset_convex `convex C` `x \ C` `y \ C` `x < y` by fastforce from z1 have z1': "f x - f y = (x - y) * f' z1" by (simp add:field_simps) @@ -568,14 +618,14 @@ have A': "y - z1 \ 0" using z1 by auto have "z3 \ C" using z3 asm atMostAtLeast_subset_convex `convex C` `x \ C` `z1 \ C` `x < z1` by fastforce - hence B': "f'' z3 \ 0" using assms by auto + then have B': "f'' z3 \ 0" using assms by auto from A' B' have "(y - z1) * f'' z3 \ 0" using mult_nonneg_nonneg by auto from cool' this have "f' y - (f x - f y) / (x - y) \ 0" by auto from mult_right_mono_neg[OF this le(2)] have "f' y * (x - y) - (f x - f y) / (x - y) * (x - y) \ 0 * (x - y)" by (simp add: algebra_simps) - hence "f' y * (x - y) - (f x - f y) \ 0" using le by auto - hence res: "f' y * (x - y) \ f x - f y" by auto + then have "f' y * (x - y) - (f x - f y) \ 0" using le by auto + then have res: "f' y * (x - y) \ f x - f y" by auto have "(f y - f x) / (y - x) - f' x = f' z1 - f' x" using asm z1 by auto also have "\ = (z1 - x) * f'' z2" using z2 by auto @@ -583,30 +633,32 @@ have A: "z1 - x \ 0" using z1 by auto have "z2 \ C" using z2 z1 asm atMostAtLeast_subset_convex `convex C` `z1 \ C` `y \ C` `z1 < y` by fastforce - hence B: "f'' z2 \ 0" using assms by auto + then have B: "f'' z2 \ 0" using assms by auto from A B have "(z1 - x) * f'' z2 \ 0" using mult_nonneg_nonneg by auto from cool this have "(f y - f x) / (y - x) - f' x \ 0" by auto from mult_right_mono[OF this ge(2)] have "(f y - f x) / (y - x) * (y - x) - f' x * (y - x) \ 0 * (y - x)" by (simp add: algebra_simps) - hence "f y - f x - f' x * (y - x) \ 0" using ge by auto - hence "f y - f x \ f' x * (y - x)" "f' y * (x - y) \ f x - f y" + then have "f y - f x - f' x * (y - x) \ 0" using ge by auto + then have "f y - f x \ f' x * (y - x)" "f' y * (x - y) \ f x - f y" using res by auto } note less_imp = this - { fix x y :: real assume "x \ C" "y \ C" "x \ y" - hence"f y - f x \ f' x * (y - x)" + { fix x y :: real + assume "x \ C" "y \ C" "x \ y" + then have"f y - f x \ f' x * (y - x)" unfolding neq_iff using less_imp by auto } note neq_imp = this moreover - { fix x y :: real assume asm: "x \ C" "y \ C" "x = y" - hence "f y - f x \ f' x * (y - x)" by auto } + { fix x y :: real + assume asm: "x \ C" "y \ C" "x = y" + then have "f y - f x \ f' x * (y - x)" by auto } ultimately show ?thesis using assms by blast qed lemma f''_ge0_imp_convex: fixes f :: "real \ real" assumes conv: "convex C" - assumes f': "\ x. x \ C \ DERIV f x :> (f' x)" - assumes f'': "\ x. x \ C \ DERIV f' x :> (f'' x)" - assumes pos: "\ x. x \ C \ f'' x \ 0" + and f': "\x. x \ C \ DERIV f x :> (f' x)" + and f'': "\x. x \ C \ DERIV f' x :> (f'' x)" + and pos: "\x. x \ C \ f'' x \ 0" shows "convex_on C f" using f''_imp_f'[OF conv f' f'' pos] assms pos_convex_function by fastforce @@ -615,18 +667,19 @@ assumes "b > 1" shows "convex_on {0 <..} (\ x. - log b x)" proof - - have "\ z. z > 0 \ DERIV (log b) z :> 1 / (ln b * z)" using DERIV_log by auto - hence f': "\ z. z > 0 \ DERIV (\ z. - log b z) z :> - 1 / (ln b * z)" + have "\z. z > 0 \ DERIV (log b) z :> 1 / (ln b * z)" using DERIV_log by auto + then have f': "\z. z > 0 \ DERIV (\ z. - log b z) z :> - 1 / (ln b * z)" using DERIV_minus by auto - have "\ z :: real. z > 0 \ DERIV inverse z :> - (inverse z ^ Suc (Suc 0))" + have "\z :: real. z > 0 \ DERIV inverse z :> - (inverse z ^ Suc (Suc 0))" using less_imp_neq[THEN not_sym, THEN DERIV_inverse] by auto from this[THEN DERIV_cmult, of _ "- 1 / ln b"] - have "\ z :: real. z > 0 \ DERIV (\ z. (- 1 / ln b) * inverse z) z :> (- 1 / ln b) * (- (inverse z ^ Suc (Suc 0)))" + have "\z :: real. z > 0 \ + DERIV (\ z. (- 1 / ln b) * inverse z) z :> (- 1 / ln b) * (- (inverse z ^ Suc (Suc 0)))" by auto - hence f''0: "\ z :: real. z > 0 \ DERIV (\ z. - 1 / (ln b * z)) z :> 1 / (ln b * z * z)" + then have f''0: "\z :: real. z > 0 \ DERIV (\ z. - 1 / (ln b * z)) z :> 1 / (ln b * z * z)" unfolding inverse_eq_divide by (auto simp add: mult_assoc) - have f''_ge0: "\ z :: real. z > 0 \ 1 / (ln b * z * z) \ 0" - using `b > 1` by (auto intro!:less_imp_le simp add:divide_pos_pos[of 1] mult_pos_pos) + have f''_ge0: "\z :: real. z > 0 \ 1 / (ln b * z * z) \ 0" + using `b > 1` by (auto intro!:less_imp_le simp add: divide_pos_pos[of 1] mult_pos_pos) from f''_ge0_imp_convex[OF pos_is_convex, unfolded greaterThan_iff, OF f' f''0 f''_ge0] show ?thesis by auto diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Pure/General/file.scala --- a/src/Pure/General/file.scala Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Pure/General/file.scala Thu Sep 27 19:35:29 2012 +0200 @@ -105,7 +105,7 @@ /* copy */ def eq(file1: JFile, file2: JFile): Boolean = - file1.getCanonicalPath == file2.getCanonicalPath // FIXME prefer java.nio.file.Files.isSameFile of Java 1.7 + java.nio.file.Files.isSameFile(file1.toPath, file2.toPath) def copy(src: JFile, dst: JFile) { diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Pure/PIDE/command.scala --- a/src/Pure/PIDE/command.scala Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Pure/PIDE/command.scala Thu Sep 27 19:35:29 2012 +0200 @@ -23,6 +23,9 @@ val results: SortedMap[Long, XML.Tree] = SortedMap.empty, val markup: Markup_Tree = Markup_Tree.empty) { + def markup_to_XML: XML.Body = markup.to_XML(command.source) + + /* accumulate content */ private def add_status(st: Markup): State = copy(status = st :: status) diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Pure/PIDE/markup.scala --- a/src/Pure/PIDE/markup.scala Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Pure/PIDE/markup.scala Thu Sep 27 19:35:29 2012 +0200 @@ -22,7 +22,6 @@ /* elements */ val Empty = Markup("", Nil) - val Data = Markup("data", Nil) val Broken = Markup("broken", Nil) } diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Pure/PIDE/markup_tree.scala --- a/src/Pure/PIDE/markup_tree.scala Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Pure/PIDE/markup_tree.scala Thu Sep 27 19:35:29 2012 +0200 @@ -12,6 +12,7 @@ import javax.swing.tree.DefaultMutableTreeNode import scala.collection.immutable.SortedMap +import scala.collection.mutable import scala.annotation.tailrec @@ -65,6 +66,7 @@ /* XML representation */ + // FIXME decode markup body @tailrec private def strip_elems(markups: List[Markup], body: XML.Body): (List[Markup], XML.Body) = body match { case List(XML.Elem(markup1, body1)) => strip_elems(markup1 :: markups, body1) @@ -110,7 +112,6 @@ val start = Text.Range(range.start) val stop = Text.Range(range.stop) val bs = branches.range(start, stop) - // FIXME check after Scala 2.8.x branches.get(stop) match { case Some(end) if range overlaps end.range => bs + (end.range -> end) case _ => bs @@ -132,16 +133,10 @@ new Markup_Tree(Branches.empty, Entry(new_markup, this)) else { val body = overlapping(new_range) - if (body.forall(e => new_range.contains(e._1))) { - val rest = // branches -- body, modulo workarounds for Redblack in Scala 2.8.0 FIXME - if (body.size > 1) - (Branches.empty /: branches)((rest, entry) => - if (body.isDefinedAt(entry._1)) rest else rest + entry) - else branches - new Markup_Tree(rest, Entry(new_markup, new Markup_Tree(body))) - } - else { // FIXME split markup!? - System.err.println("Ignored overlapping markup information: " + new_markup + + if (body.forall(e => new_range.contains(e._1))) + new Markup_Tree(branches -- body.keys, Entry(new_markup, new Markup_Tree(body))) + else { + java.lang.System.err.println("Ignored overlapping markup information: " + new_markup + body.filter(e => !new_range.contains(e._1)).mkString("\n")) this } @@ -149,13 +144,45 @@ } } + def to_XML(root_range: Text.Range, text: CharSequence, filter: XML.Elem => Boolean): XML.Body = + { + def make_text(start: Text.Offset, stop: Text.Offset): XML.Body = + if (start == stop) Nil + else List(XML.Text(text.subSequence(start, stop).toString)) + + def make_elems(rev_markups: List[XML.Elem], body: XML.Body): XML.Body = + (body /: rev_markups) { + case (b, elem) => // FIXME encode markup body + if (filter(elem)) List(XML.Elem(elem.markup, b)) else b + } + + def make_body(elem_range: Text.Range, elem_markup: List[XML.Elem], entries: Branches.T) + : XML.Body = + { + val body = new mutable.ListBuffer[XML.Tree] + var last = elem_range.start + for ((range, entry) <- entries) { + val subrange = range.restrict(elem_range) + body ++= make_text(last, subrange.start) + body ++= make_body(subrange, entry.rev_markup, entry.subtree.overlapping(subrange)) + last = subrange.stop + } + body ++= make_text(last, elem_range.stop) + make_elems(elem_markup, body.toList) + } + make_body(root_range, Nil, overlapping(root_range)) + } + + def to_XML(text: CharSequence): XML.Body = + to_XML(Text.Range(0, text.length), text, (_: XML.Elem) => true) + def cumulate[A](root_range: Text.Range, root_info: A, result_elements: Option[Set[String]], result: PartialFunction[(A, Text.Markup), A]): Stream[Text.Info[A]] = { def results(x: A, entry: Entry): Option[A] = if (result_elements match { case Some(es) => es.exists(entry.elements) case None => true }) { val (y, changed) = - (entry.markup :\ (x, false))((info, res) => + ((x, false) /: entry.rev_markup)((res, info) => // FIXME proper order!? { val (y, changed) = res val arg = (y, Text.Info(entry.range, info)) diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Pure/PIDE/xml.scala --- a/src/Pure/PIDE/xml.scala Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Pure/PIDE/xml.scala Thu Sep 27 19:35:29 2012 +0200 @@ -7,7 +7,6 @@ package isabelle -import java.lang.System import java.util.WeakHashMap import java.lang.ref.WeakReference import javax.xml.parsers.DocumentBuilderFactory @@ -171,35 +170,6 @@ - /** document object model (W3C DOM) **/ - - def get_data(node: org.w3c.dom.Node): Option[XML.Tree] = - node.getUserData(Markup.Data.name) match { - case tree: XML.Tree => Some(tree) - case _ => None - } - - def document_node(doc: org.w3c.dom.Document, tree: Tree): org.w3c.dom.Node = - { - def DOM(tr: Tree): org.w3c.dom.Node = tr match { - case Elem(Markup.Data, List(data, t)) => - val node = DOM(t) - node.setUserData(Markup.Data.name, data, null) - node - case Elem(Markup(name, atts), ts) => - if (name == Markup.Data.name) - error("Malformed data element: " + tr.toString) - val node = doc.createElement(name) - for ((name, value) <- atts) node.setAttribute(name, value) - for (t <- ts) node.appendChild(DOM(t)) - node - case Text(txt) => doc.createTextNode(txt) - } - DOM(tree) - } - - - /** XML as data representation language **/ class XML_Atom(s: String) extends Exception(s) diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Pure/System/html5_panel.scala --- a/src/Pure/System/html5_panel.scala Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Pure/System/html5_panel.scala Thu Sep 27 19:35:29 2012 +0200 @@ -6,12 +6,11 @@ package isabelle -import com.sun.javafx.tk.{FontMetrics, Toolkit} import javafx.scene.Scene import javafx.scene.web.{WebView, WebEngine} import javafx.scene.input.KeyEvent -import javafx.scene.text.{Font, FontSmoothingType} +import javafx.scene.text.FontSmoothingType import javafx.scene.layout.{HBox, VBox, Priority} import javafx.geometry.{HPos, VPos, Insets} import javafx.event.EventHandler @@ -51,30 +50,8 @@ } -class HTML5_Panel(main_css: String, init_font_family: String, init_font_size: Int) - extends javafx.embed.swing.JFXPanel +class HTML5_Panel extends javafx.embed.swing.JFXPanel { - /* HTML/CSS template */ - - def template(font_family: String, font_size: Int): String = -""" - - - - - - - -""" - - - /* main Web view */ - private val future = JFX_Thread.future { val pane = new Web_View_Workaround @@ -93,68 +70,9 @@ }) setScene(new Scene(pane)) - - web_view.getEngine.loadContent(template(init_font_family, init_font_size)) pane } def web_view: WebView = future.join.web_view def web_engine: WebEngine = web_view.getEngine - - - /* internal state -- owned by JFX thread */ - - private var current_font_metrics: FontMetrics = null - private var current_font_family = "" - private var current_font_size: Int = 0 - private var current_margin: Int = 0 - private var current_body: XML.Body = Nil - - // FIXME move to pretty.scala (!?) - private def pretty_metric(metrics: FontMetrics): String => Double = - { - if (metrics == null) ((s: String) => s.length.toDouble) - else { - val unit = metrics.computeStringWidth(Pretty.space).toDouble - ((s: String) => if (s == "\n") 1.0 else metrics.computeStringWidth(s) / unit) - } - } - - def resize(font_family: String, font_size: Int): Unit = JFX_Thread.later { - val font = new Font(font_family, font_size) - val font_metrics = Toolkit.getToolkit().getFontLoader().getFontMetrics(font) - val margin = // FIXME Swing thread!? - (getWidth() / (font_metrics.computeStringWidth(Pretty.space) max 1.0f)).toInt max 20 - - if (current_font_metrics == null || - current_font_family != font_family || - current_font_size != font_size || - current_margin != margin) - { - current_font_metrics = font_metrics - current_font_family = font_family - current_font_size = font_size - current_margin = margin - refresh() - } - } - - def refresh(): Unit = JFX_Thread.later { render(current_body) } - - def render(body: XML.Body): Unit = JFX_Thread.later { - current_body = body - val html_body = - current_body.flatMap(div => - Pretty.formatted(List(div), current_margin, pretty_metric(current_font_metrics)) - .map(t => - XML.Elem(Markup(HTML.PRE, List((HTML.CLASS, Isabelle_Markup.MESSAGE))), - HTML.spans(t, false)))) // FIXME user data (!??!) - - // FIXME web_engine.loadContent(template(current_font_family, current_font_size)) - - val document = web_engine.getDocument - val html_root = document.getLastChild - html_root.removeChild(html_root.getLastChild) - html_root.appendChild(XML.document_node(document, XML.elem(HTML.BODY, html_body))) - } } diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Pure/Thy/html.scala --- a/src/Pure/Thy/html.scala Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Pure/Thy/html.scala Thu Sep 27 19:35:29 2012 +0200 @@ -29,6 +29,8 @@ } + /// FIXME unused stuff + // common elements and attributes val BODY = "body" @@ -55,14 +57,12 @@ def sup(txt: String): XML.Elem = XML.elem("sup", List(XML.Text(txt))) def bold(txt: String): XML.Elem = span("bold", List(XML.Text(txt))) - def spans(input: XML.Tree, original_data: Boolean = false): XML.Body = + def spans(input: XML.Tree): XML.Body = { def html_spans(tree: XML.Tree): XML.Body = tree match { case XML.Elem(m @ Markup(name, props), ts) => - val html_span = span(name, ts.flatMap(html_spans)) - if (original_data) List(XML.Elem(Markup.Data, List(tree, html_span))) - else List(html_span) + List(span(name, ts.flatMap(html_spans))) case XML.Text(txt) => val ts = new ListBuffer[XML.Tree] val t = new StringBuilder diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Tools/jEdit/etc/isabelle-jedit.css --- a/src/Tools/jEdit/etc/isabelle-jedit.css Thu Sep 27 18:58:15 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,17 +0,0 @@ -/* additional style file for Isabelle/jEdit output */ - -.message { margin-top: 0.3ex; background-color: #F0F0F0; } - -.writeln_message { } -.tracing_message { background-color: #F0F8FF; } -.warning_message { background-color: #EEE8AA; } -.error_message { background-color: #FFC1C1; } - -.intensify { background-color: #FFCC66; } - -.keyword { font-weight: bold; color: #009966; } -.operator { font-weight: bold; } -.command { font-weight: bold; color: #006699; } - -.sendback { background-color: #DCDCDC; } -.sendback:hover { background-color: #9DC75D; } diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Tools/jEdit/etc/settings --- a/src/Tools/jEdit/etc/settings Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Tools/jEdit/etc/settings Thu Sep 27 19:35:29 2012 +0200 @@ -10,8 +10,6 @@ -Dcom.apple.mrj.application.apple.menu.about.name=Isabelle/jEdit -Dscala.repl.no-threads=true" -JEDIT_STYLE_SHEETS="$ISABELLE_HOME/etc/isabelle.css:$JEDIT_HOME/etc/isabelle-jedit.css:$ISABELLE_HOME_USER/etc/isabelle.css:$ISABELLE_HOME_USER/etc/isabelle-jedit.css" - ISABELLE_JEDIT_OPTIONS="" ISABELLE_TOOLS="$ISABELLE_TOOLS:$JEDIT_HOME/lib/Tools" diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Tools/jEdit/lib/Tools/jedit --- a/src/Tools/jEdit/lib/Tools/jedit Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Tools/jEdit/lib/Tools/jedit Thu Sep 27 19:35:29 2012 +0200 @@ -24,7 +24,6 @@ "src/jedit_thy_load.scala" "src/jedit_options.scala" "src/output_dockable.scala" - "src/output1_dockable.scala" "src/plugin.scala" "src/pretty_text_area.scala" "src/protocol_dockable.scala" diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Tools/jEdit/src/Isabelle.props --- a/src/Tools/jEdit/src/Isabelle.props Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Tools/jEdit/src/Isabelle.props Thu Sep 27 19:35:29 2012 +0200 @@ -40,11 +40,10 @@ #menu actions plugin.isabelle.jedit.Plugin.menu.label=Isabelle -plugin.isabelle.jedit.Plugin.menu=isabelle.session-panel isabelle.output-panel isabelle.graphview-panel isabelle.output1-panel isabelle.raw-output-panel isabelle.protocol-panel isabelle.readme-panel isabelle.syslog-panel +plugin.isabelle.jedit.Plugin.menu=isabelle.session-panel isabelle.output-panel isabelle.graphview-panel isabelle.raw-output-panel isabelle.protocol-panel isabelle.readme-panel isabelle.syslog-panel isabelle.session-panel.label=Prover Session panel isabelle.output-panel.label=Output panel isabelle.graphview-panel.label=Graphview panel -isabelle.output1-panel.label=Output1 panel isabelle.raw-output-panel.label=Raw Output panel isabelle.protocol-panel.label=Protocol panel isabelle.readme-panel.label=README panel @@ -54,7 +53,6 @@ isabelle-session.title=Prover Session isabelle-output.title=Output isabelle-graphview.title=Graphview -isabelle-output1.title=Output1 isabelle-raw-output.title=Raw Output isabelle-protocol.title=Protocol isabelle-readme.title=README diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Tools/jEdit/src/actions.xml --- a/src/Tools/jEdit/src/actions.xml Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Tools/jEdit/src/actions.xml Thu Sep 27 19:35:29 2012 +0200 @@ -22,11 +22,6 @@ wm.addDockableWindow("isabelle-output"); - - - wm.addDockableWindow("isabelle-output1"); - - wm.addDockableWindow("isabelle-graphview"); diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Tools/jEdit/src/dockables.xml --- a/src/Tools/jEdit/src/dockables.xml Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Tools/jEdit/src/dockables.xml Thu Sep 27 19:35:29 2012 +0200 @@ -17,9 +17,6 @@ new isabelle.jedit.Graphview_Dockable(view, position); - - new isabelle.jedit.Output1_Dockable(view, position); - new isabelle.jedit.Raw_Output_Dockable(view, position); diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Tools/jEdit/src/html_panel.scala --- a/src/Tools/jEdit/src/html_panel.scala Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Tools/jEdit/src/html_panel.scala Thu Sep 27 19:35:29 2012 +0200 @@ -9,195 +9,28 @@ import isabelle._ -import java.lang.System import java.io.StringReader -import java.awt.{Font, BorderLayout, Dimension, GraphicsEnvironment, Toolkit, FontMetrics} -import java.awt.event.MouseEvent import java.util.logging.{Logger, Level} -import org.w3c.dom.html2.HTMLElement - import org.lobobrowser.html.parser.{DocumentBuilderImpl, InputSourceImpl} import org.lobobrowser.html.gui.HtmlPanel -import org.lobobrowser.html.domimpl.{HTMLDocumentImpl, HTMLStyleElementImpl, NodeImpl} import org.lobobrowser.html.test.{SimpleHtmlRendererContext, SimpleUserAgentContext} -import scala.actors.Actor._ - - -object HTML_Panel -{ - sealed abstract class Event { val element: HTMLElement; val mouse: MouseEvent } - case class Context_Menu(val element: HTMLElement, mouse: MouseEvent) extends Event - case class Mouse_Click(val element: HTMLElement, mouse: MouseEvent) extends Event - case class Double_Click(val element: HTMLElement, mouse: MouseEvent) extends Event - case class Mouse_Over(val element: HTMLElement, mouse: MouseEvent) extends Event - case class Mouse_Out(val element: HTMLElement, mouse: MouseEvent) extends Event -} - -class HTML_Panel(initial_font_family: String, initial_font_size: Int) extends HtmlPanel +class HTML_Panel extends HtmlPanel { - /** Lobo setup **/ - - /* global logging */ + Swing_Thread.require() Logger.getLogger("org.lobobrowser").setLevel(Level.WARNING) - - /* pixel size -- cf. org.lobobrowser.html.style.HtmlValues.getFontSize */ - - val screen_resolution = - if (GraphicsEnvironment.isHeadless()) 72 - else Toolkit.getDefaultToolkit().getScreenResolution() - - def lobo_px(raw_px: Int): Int = raw_px * 96 / screen_resolution - def raw_px(lobo_px: Int): Int = (lobo_px * screen_resolution + 95) / 96 - - - /* contexts and event handling */ - - protected val handler: PartialFunction[HTML_Panel.Event, Unit] = Map.empty - private val ucontext = new SimpleUserAgentContext private val rcontext = new SimpleHtmlRendererContext(this, ucontext) - { - private def handle(event: HTML_Panel.Event): Boolean = - if (handler.isDefinedAt(event)) { handler(event); false } - else true - - override def onContextMenu(elem: HTMLElement, event: MouseEvent): Boolean = - handle(HTML_Panel.Context_Menu(elem, event)) - override def onMouseClick(elem: HTMLElement, event: MouseEvent): Boolean = - handle(HTML_Panel.Mouse_Click(elem, event)) - override def onDoubleClick(elem: HTMLElement, event: MouseEvent): Boolean = - handle(HTML_Panel.Double_Click(elem, event)) - override def onMouseOver(elem: HTMLElement, event: MouseEvent) - { handle(HTML_Panel.Mouse_Over(elem, event)) } - override def onMouseOut(elem: HTMLElement, event: MouseEvent) - { handle(HTML_Panel.Mouse_Out(elem, event)) } - } - private val builder = new DocumentBuilderImpl(ucontext, rcontext) - - /* document template with style sheets */ - - private val template_head = - """ - - - - - - - -""" - - private def template(font_family: String, font_size: Int): String = - template_head + - "body { font-family: " + font_family + "; font-size: " + raw_px(font_size) + "px; }" + - template_tail - - - /** main actor **/ - - /* internal messages */ - - private case class Resize(font_family: String, font_size: Int) - private case class Render_Document(url: String, text: String) - private case class Render(body: XML.Body) - private case class Render_Sync(body: XML.Body) - private case object Refresh - - private val main_actor = actor { - - /* internal state */ - - var current_font_metrics: FontMetrics = null - var current_font_family = "" - var current_font_size: Int = 0 - var current_margin: Int = 0 - var current_body: XML.Body = Nil - - def resize(font_family: String, font_size: Int) - { - val font = new Font(font_family, Font.PLAIN, lobo_px(raw_px(font_size))) - val (font_metrics, margin) = - Swing_Thread.now { - val metrics = getFontMetrics(font) - (metrics, (getWidth() / (metrics.charWidth(Pretty.spc) max 1) - 4) max 20) - } - if (current_font_metrics == null || - current_font_family != font_family || - current_font_size != font_size || - current_margin != margin) - { - current_font_metrics = font_metrics - current_font_family = font_family - current_font_size = font_size - current_margin = margin - refresh() - } - } - - def refresh() { render(current_body) } - - def render_document(url: String, text: String) - { - val doc = builder.parse(new InputSourceImpl(new StringReader(text), url)) - Swing_Thread.later { setDocument(doc, rcontext) } - } - - def render(body: XML.Body) - { - current_body = body - val html_body = - current_body.flatMap(div => - Pretty.formatted(List(div), current_margin, Pretty.font_metric(current_font_metrics)) - .map(t => - XML.Elem(Markup(HTML.PRE, List((HTML.CLASS, Isabelle_Markup.MESSAGE))), - HTML.spans(t, true)))) - val doc = - builder.parse( - new InputSourceImpl( - new StringReader(template(current_font_family, current_font_size)), "http://localhost")) - doc.removeChild(doc.getLastChild()) - doc.appendChild(XML.document_node(doc, XML.elem(HTML.BODY, html_body))) - Swing_Thread.later { setDocument(doc, rcontext) } - } - - - /* main loop */ - - resize(initial_font_family, initial_font_size) - - loop { - react { - case Resize(font_family, font_size) => resize(font_family, font_size) - case Refresh => refresh() - case Render_Document(url, text) => render_document(url, text) - case Render(body) => render(body) - case Render_Sync(body) => render(body); reply(()) - case bad => System.err.println("main_actor: ignoring bad message " + bad) - } - } + def render_document(url: String, html_text: String) + { + val doc = builder.parse(new InputSourceImpl(new StringReader(html_text), url)) + Swing_Thread.later { setDocument(doc, rcontext) } } - - - /* external methods */ - - def resize(font_family: String, font_size: Int) { main_actor ! Resize(font_family, font_size) } - def refresh() { main_actor ! Refresh } - def render_document(url: String, text: String) { main_actor ! Render_Document(url, text) } - def render(body: XML.Body) { main_actor ! Render(body) } - def render_sync(body: XML.Body) { main_actor !? Render_Sync(body) } } diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Tools/jEdit/src/output1_dockable.scala --- a/src/Tools/jEdit/src/output1_dockable.scala Thu Sep 27 18:58:15 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,188 +0,0 @@ -/* Title: Tools/jEdit/src/output1_dockable.scala - Author: Makarius - -Dockable window with result message output. -*/ - -package isabelle.jedit - - -import isabelle._ - -import scala.actors.Actor._ - -import scala.swing.{FlowPanel, Button, CheckBox} -import scala.swing.event.ButtonClicked - -import java.lang.System -import java.awt.BorderLayout -import java.awt.event.{ComponentEvent, ComponentAdapter} - -import org.gjt.sp.jedit.View - - -class Output1_Dockable(view: View, position: String) extends Dockable(view, position) -{ - Swing_Thread.require() - - - /* component state -- owned by Swing thread */ - - private var zoom_factor = 100 - private var show_tracing = false - private var do_update = true - private var current_state = Command.empty.init_state - private var current_body: XML.Body = Nil - - - /* HTML panel */ - - private val html_panel = - new HTML_Panel(Isabelle.font_family(), scala.math.round(Isabelle.font_size())) - { - override val handler: PartialFunction[HTML_Panel.Event, Unit] = - { - case HTML_Panel.Mouse_Click(elem, event) - if Protocol.Sendback.unapply(elem.getUserData(Markup.Data.name)).isDefined => - val sendback = Protocol.Sendback.unapply(elem.getUserData(Markup.Data.name)).get - Document_View(view.getTextArea) match { - case Some(doc_view) => - doc_view.rich_text_area.robust_body() { - val cmd = current_state.command - val model = doc_view.model - val buffer = model.buffer - val snapshot = model.snapshot() - snapshot.node.command_start(cmd) match { - case Some(start) if !snapshot.is_outdated => - val text = Pretty.string_of(sendback) - try { - buffer.beginCompoundEdit() - buffer.remove(start, cmd.proper_range.length) - buffer.insert(start, text) - } - finally { buffer.endCompoundEdit() } - case _ => - } - } - case None => - } - } - } - - set_content(html_panel) - - - private def handle_resize() - { - Swing_Thread.require() - - html_panel.resize(Isabelle.font_family(), - scala.math.round(Isabelle.font_size() * zoom_factor / 100)) - } - - private def handle_update(follow: Boolean, restriction: Option[Set[Command]]) - { - Swing_Thread.require() - - val new_state = - if (follow) { - Document_View(view.getTextArea) match { - case Some(doc_view) => - val snapshot = doc_view.model.snapshot() - snapshot.node.command_at(doc_view.text_area.getCaretPosition).map(_._1) match { - case Some(cmd) => snapshot.state.command_state(snapshot.version, cmd) - case None => Command.empty.init_state - } - case None => Command.empty.init_state - } - } - else current_state - - val new_body = - if (!restriction.isDefined || restriction.get.contains(new_state.command)) - new_state.results.iterator.map(_._2) - .filter(msg => !Protocol.is_tracing(msg) || show_tracing).toList // FIXME not scalable - else current_body - - if (new_body != current_body) html_panel.render(new_body) - - current_state = new_state - current_body = new_body - } - - - /* main actor */ - - private val main_actor = actor { - loop { - react { - case Session.Global_Options => - Swing_Thread.later { handle_resize() } - case changed: Session.Commands_Changed => - Swing_Thread.later { handle_update(do_update, Some(changed.commands)) } - case Session.Caret_Focus => - Swing_Thread.later { handle_update(do_update, None) } - case bad => System.err.println("Output_Dockable: ignoring bad message " + bad) - } - } - } - - override def init() - { - Swing_Thread.require() - - Isabelle.session.global_options += main_actor - Isabelle.session.commands_changed += main_actor - Isabelle.session.caret_focus += main_actor - handle_update(true, None) - } - - override def exit() - { - Swing_Thread.require() - - Isabelle.session.global_options -= main_actor - Isabelle.session.commands_changed -= main_actor - Isabelle.session.caret_focus -= main_actor - delay_resize.revoke() - } - - - /* resize */ - - private val delay_resize = - Swing_Thread.delay_first( - Time.seconds(Isabelle.options.real("editor_update_delay"))) { handle_resize() } - - addComponentListener(new ComponentAdapter { - override def componentResized(e: ComponentEvent) { delay_resize.invoke() } - }) - - - /* controls */ - - private val zoom = new Library.Zoom_Box(factor => { zoom_factor = factor; handle_resize() }) - zoom.tooltip = "Zoom factor for basic font size" - - private val tracing = new CheckBox("Tracing") { - reactions += { - case ButtonClicked(_) => show_tracing = this.selected; handle_update(do_update, None) } - } - tracing.selected = show_tracing - tracing.tooltip = "Indicate output of tracing messages" - - private val auto_update = new CheckBox("Auto update") { - reactions += { - case ButtonClicked(_) => do_update = this.selected; handle_update(do_update, None) } - } - auto_update.selected = do_update - auto_update.tooltip = "Indicate automatic update following cursor movement" - - private val update = new Button("Update") { - reactions += { case ButtonClicked(_) => handle_update(true, None) } - } - update.tooltip = "Update display according to the command at cursor position" - - private val controls = new FlowPanel(FlowPanel.Alignment.Right)(zoom, tracing, auto_update, update) - add(controls.peer, BorderLayout.NORTH) -} diff -r 1a0f25c38629 -r 9d34cfe1dbdf src/Tools/jEdit/src/readme_dockable.scala --- a/src/Tools/jEdit/src/readme_dockable.scala Thu Sep 27 18:58:15 2012 +0200 +++ b/src/Tools/jEdit/src/readme_dockable.scala Thu Sep 27 19:35:29 2012 +0200 @@ -14,10 +14,11 @@ class README_Dockable(view: View, position: String) extends Dockable(view, position) { - private val readme = new HTML_Panel("SansSerif", 14) + Swing_Thread.require() + + private val readme = new HTML_Panel private val readme_path = Path.explode("$JEDIT_HOME/README.html") - readme.render_document( - Isabelle_System.platform_file_url(readme_path), File.try_read(List(readme_path))) + readme.render_document(Isabelle_System.platform_file_url(readme_path), File.read(readme_path)) set_content(readme) }