# HG changeset patch # User wenzelm # Date 1468438241 -7200 # Node ID a7c5074a025155b5c2e832038359bdb9341f8d2a # Parent ae33d1c2ab26c52c9c93858d309245f319a171f6# Parent 6e29fb72e65942174281178a4ed25684804d7c10 merged diff -r ae33d1c2ab26 -r a7c5074a0251 NEWS --- a/NEWS Wed Jul 13 21:00:03 2016 +0200 +++ b/NEWS Wed Jul 13 21:30:41 2016 +0200 @@ -79,13 +79,19 @@ * Refined folding mode "isabelle" based on Isar syntax: 'next' and 'qed' are treated as delimiters for fold structure. -* Improved support for indentation according to Isabelle outer syntax. -Action "indent-lines" (shortcut C+i) indents the current line according -to command keywords and some command substructure. Action +* Syntactic indentation according to Isabelle outer syntax. Action +"indent-lines" (shortcut C+i) indents the current line according to +command keywords and some command substructure. Action "isabelle.newline" (shortcut ENTER) indents the old and the new line according to command keywords only; see also option "jedit_indent_newline". +* Semantic indentation for unstructured proof scripts ('apply' etc.) via +number of subgoals. This requires information of ongoing document +processing and may thus lag behind, when the user is editing too +quickly; see also option "jedit_script_indent" and +"jedit_script_indent_limit". + * Action "isabelle.select-entity" (shortcut CS+ENTER) selects all occurences of the formal entity at the caret position. This facilitates systematic renaming. diff -r ae33d1c2ab26 -r a7c5074a0251 src/HOL/Library/AList.thy --- a/src/HOL/Library/AList.thy Wed Jul 13 21:00:03 2016 +0200 +++ b/src/HOL/Library/AList.thy Wed Jul 13 21:30:41 2016 +0200 @@ -73,8 +73,7 @@ @{term "update k' v' (update k v []) = [(k, v), (k', v')]"}.\ lemma update_swap: - "k \ k' \ - map_of (update k v (update k' v' al)) = map_of (update k' v' (update k v al))" + "k \ k' \ map_of (update k v (update k' v' al)) = map_of (update k' v' (update k v al))" by (simp add: update_conv' fun_eq_iff) lemma update_Some_unfold: @@ -85,8 +84,8 @@ lemma image_update [simp]: "x \ A \ map_of (update x y al) ` A = map_of al ` A" by (simp add: update_conv') -qualified definition updates - :: "'key list \ 'val list \ ('key \ 'val) list \ ('key \ 'val) list" +qualified definition updates :: + "'key list \ 'val list \ ('key \ 'val) list \ ('key \ 'val) list" where "updates ks vs = fold (case_prod update) (zip ks vs)" lemma updates_simps [simp]: @@ -216,8 +215,8 @@ subsection \\update_with_aux\ and \delete_aux\\ -qualified primrec update_with_aux - :: "'val \ 'key \ ('val \ 'val) \ ('key \ 'val) list \ ('key \ 'val) list" +qualified primrec update_with_aux :: + "'val \ 'key \ ('val \ 'val) \ ('key \ 'val) list \ ('key \ 'val) list" where "update_with_aux v k f [] = [(k, f v)]" | "update_with_aux v k f (p # ps) = @@ -257,7 +256,7 @@ lemma set_delete_aux: "distinct (map fst xs) \ set (delete_aux k xs) = set xs - {k} \ UNIV" apply (induct xs) - apply simp_all + apply simp_all apply clarsimp apply (fastforce intro: rev_image_eqI) done @@ -291,7 +290,7 @@ lemma map_of_delete_aux': "distinct (map fst xs) \ map_of (delete_aux k xs) = (map_of xs)(k := None)" apply (induct xs) - apply (fastforce simp add: map_of_eq_None_iff fun_upd_twist) + apply (fastforce simp add: map_of_eq_None_iff fun_upd_twist) apply (auto intro!: ext) apply (simp add: map_of_eq_None_iff) done @@ -318,9 +317,9 @@ proof show "map_of (restrict A al) k = ((map_of al)|` A) k" for k apply (induct al) - apply simp + apply simp apply (cases "k \ A") - apply auto + apply auto done qed diff -r ae33d1c2ab26 -r a7c5074a0251 src/HOL/Library/AList_Mapping.thy --- a/src/HOL/Library/AList_Mapping.thy Wed Jul 13 21:00:03 2016 +0200 +++ b/src/HOL/Library/AList_Mapping.thy Wed Jul 13 21:30:41 2016 +0200 @@ -51,8 +51,11 @@ proof - have *: "(a, b) \ set xs \ a \ fst ` set xs" for a b xs by (auto simp add: image_def intro!: bexI) - show ?thesis apply transfer - by (auto intro!: map_of_eqI) (auto dest!: map_of_eq_dom intro: *) + show ?thesis + apply transfer + apply (auto intro!: map_of_eqI) + apply (auto dest!: map_of_eq_dom intro: *) + done qed lemma map_values_Mapping [code]: @@ -72,8 +75,8 @@ apply (rule sym) subgoal for f xs ys x apply (cases "map_of xs x"; cases "map_of ys x"; simp) - apply (force simp: map_of_eq_None_iff combine_options_def option.the_def o_def image_iff - dest: map_of_SomeD split: option.splits)+ + apply (force simp: map_of_eq_None_iff combine_options_def option.the_def o_def image_iff + dest: map_of_SomeD split: option.splits)+ done done @@ -86,8 +89,8 @@ apply (rule sym) subgoal for f xs ys x apply (cases "map_of xs x"; cases "map_of ys x"; simp) - apply (force simp: map_of_eq_None_iff combine_options_def option.the_def o_def image_iff - dest: map_of_SomeD split: option.splits)+ + apply (force simp: map_of_eq_None_iff combine_options_def option.the_def o_def image_iff + dest: map_of_SomeD split: option.splits)+ done done @@ -106,7 +109,7 @@ apply transfer apply (rule ext) apply (subst map_of_filter_distinct) - apply (simp_all add: map_of_clearjunk split: option.split) + apply (simp_all add: map_of_clearjunk split: option.split) done lemma [code nbe]: "HOL.equal (x :: ('a, 'b) mapping) x \ True" diff -r ae33d1c2ab26 -r a7c5074a0251 src/HOL/Library/BigO.thy --- a/src/HOL/Library/BigO.thy Wed Jul 13 21:00:03 2016 +0200 +++ b/src/HOL/Library/BigO.thy Wed Jul 13 21:30:41 2016 +0200 @@ -5,33 +5,35 @@ section \Big O notation\ theory BigO - imports Complex_Main Function_Algebras Set_Algebras + imports + Complex_Main + Function_Algebras + Set_Algebras begin text \ -This library is designed to support asymptotic ``big O'' calculations, -i.e.~reasoning with expressions of the form $f = O(g)$ and $f = g + -O(h)$. An earlier version of this library is described in detail in -@{cite "Avigad-Donnelly"}. + This library is designed to support asymptotic ``big O'' calculations, + i.e.~reasoning with expressions of the form \f = O(g)\ and \f = g + O(h)\. + An earlier version of this library is described in detail in @{cite + "Avigad-Donnelly"}. + + The main changes in this version are as follows: -The main changes in this version are as follows: -\begin{itemize} -\item We have eliminated the \O\ operator on sets. (Most uses of this seem - to be inessential.) -\item We no longer use \+\ as output syntax for \+o\ -\item Lemmas involving \sumr\ have been replaced by more general lemmas - involving `\setsum\. -\item The library has been expanded, with e.g.~support for expressions of - the form \f < g + O(h)\. -\end{itemize} + \<^item> We have eliminated the \O\ operator on sets. (Most uses of this seem + to be inessential.) + \<^item> We no longer use \+\ as output syntax for \+o\ + \<^item> Lemmas involving \sumr\ have been replaced by more general lemmas + involving `\setsum\. + \<^item> The library has been expanded, with e.g.~support for expressions of + the form \f < g + O(h)\. -Note also since the Big O library includes rules that demonstrate set -inclusion, to use the automated reasoners effectively with the library -one should redeclare the theorem \subsetI\ as an intro rule, -rather than as an \intro!\ rule, for example, using -\<^theory_text>\declare subsetI [del, intro]\. + Note also since the Big O library includes rules that demonstrate set + inclusion, to use the automated reasoners effectively with the library one + should redeclare the theorem \subsetI\ as an intro rule, rather than as an + \intro!\ rule, for example, using \<^theory_text>\declare subsetI [del, intro]\. \ + subsection \Definitions\ definition bigo :: "('a \ 'b::linordered_idom) \ ('a \ 'b) set" ("(1O'(_'))") @@ -42,16 +44,16 @@ (\c. 0 < c \ (\x. \h x\ \ c * \f x\))" apply auto apply (case_tac "c = 0") - apply simp - apply (rule_tac x = "1" in exI) - apply simp + apply simp + apply (rule_tac x = "1" in exI) + apply simp apply (rule_tac x = "\c\" in exI) apply auto apply (subgoal_tac "c * \f x\ \ \c\ * \f x\") - apply (erule_tac x = x in allE) - apply force + apply (erule_tac x = x in allE) + apply force apply (rule mult_right_mono) - apply (rule abs_ge_self) + apply (rule abs_ge_self) apply (rule abs_ge_zero) done @@ -62,19 +64,19 @@ apply (auto simp add: bigo_alt_def) apply (rule_tac x = "ca * c" in exI) apply (rule conjI) - apply simp + apply simp apply (rule allI) apply (drule_tac x = "xa" in spec)+ apply (subgoal_tac "ca * \f xa\ \ ca * (c * \g xa\)") - apply (erule order_trans) - apply (simp add: ac_simps) + apply (erule order_trans) + apply (simp add: ac_simps) apply (rule mult_left_mono, assumption) apply (rule order_less_imp_le, assumption) done lemma bigo_refl [intro]: "f \ O(f)" - apply(auto simp add: bigo_def) - apply(rule_tac x = 1 in exI) + apply (auto simp add: bigo_def) + apply (rule_tac x = 1 in exI) apply simp done @@ -93,15 +95,15 @@ apply auto apply (simp add: ring_distribs func_plus) apply (rule order_trans) - apply (rule abs_triangle_ineq) + apply (rule abs_triangle_ineq) apply (rule add_mono) - apply force + apply force apply force done lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)" apply (rule equalityI) - apply (rule bigo_plus_self_subset) + apply (rule bigo_plus_self_subset) apply (rule set_zero_plus2) apply (rule bigo_zero) done @@ -112,73 +114,73 @@ apply (subst bigo_pos_const [symmetric])+ apply (rule_tac x = "\n. if \g n\ \ \f n\ then x n else 0" in exI) apply (rule conjI) - apply (rule_tac x = "c + c" in exI) - apply (clarsimp) - apply (subgoal_tac "c * \f xa + g xa\ \ (c + c) * \f xa\") - apply (erule_tac x = xa in allE) - apply (erule order_trans) - apply (simp) - apply (subgoal_tac "c * \f xa + g xa\ \ c * (\f xa\ + \g xa\)") - apply (erule order_trans) - apply (simp add: ring_distribs) - apply (rule mult_left_mono) - apply (simp add: abs_triangle_ineq) - apply (simp add: order_less_le) + apply (rule_tac x = "c + c" in exI) + apply (clarsimp) + apply (subgoal_tac "c * \f xa + g xa\ \ (c + c) * \f xa\") + apply (erule_tac x = xa in allE) + apply (erule order_trans) + apply (simp) + apply (subgoal_tac "c * \f xa + g xa\ \ c * (\f xa\ + \g xa\)") + apply (erule order_trans) + apply (simp add: ring_distribs) + apply (rule mult_left_mono) + apply (simp add: abs_triangle_ineq) + apply (simp add: order_less_le) apply (rule_tac x = "\n. if \f n\ < \g n\ then x n else 0" in exI) apply (rule conjI) - apply (rule_tac x = "c + c" in exI) - apply auto + apply (rule_tac x = "c + c" in exI) + apply auto apply (subgoal_tac "c * \f xa + g xa\ \ (c + c) * \g xa\") - apply (erule_tac x = xa in allE) - apply (erule order_trans) - apply simp + apply (erule_tac x = xa in allE) + apply (erule order_trans) + apply simp apply (subgoal_tac "c * \f xa + g xa\ \ c * (\f xa\ + \g xa\)") - apply (erule order_trans) - apply (simp add: ring_distribs) + apply (erule order_trans) + apply (simp add: ring_distribs) apply (rule mult_left_mono) - apply (rule abs_triangle_ineq) + apply (rule abs_triangle_ineq) apply (simp add: order_less_le) done lemma bigo_plus_subset2 [intro]: "A \ O(f) \ B \ O(f) \ A + B \ O(f)" apply (subgoal_tac "A + B \ O(f) + O(f)") - apply (erule order_trans) - apply simp + apply (erule order_trans) + apply simp apply (auto del: subsetI simp del: bigo_plus_idemp) done lemma bigo_plus_eq: "\x. 0 \ f x \ \x. 0 \ g x \ O(f + g) = O(f) + O(g)" apply (rule equalityI) - apply (rule bigo_plus_subset) + apply (rule bigo_plus_subset) apply (simp add: bigo_alt_def set_plus_def func_plus) apply clarify apply (rule_tac x = "max c ca" in exI) apply (rule conjI) - apply (subgoal_tac "c \ max c ca") - apply (erule order_less_le_trans) - apply assumption - apply (rule max.cobounded1) + apply (subgoal_tac "c \ max c ca") + apply (erule order_less_le_trans) + apply assumption + apply (rule max.cobounded1) apply clarify apply (drule_tac x = "xa" in spec)+ apply (subgoal_tac "0 \ f xa + g xa") - apply (simp add: ring_distribs) - apply (subgoal_tac "\a xa + b xa\ \ \a xa\ + \b xa\") - apply (subgoal_tac "\a xa\ + \b xa\ \ max c ca * f xa + max c ca * g xa") - apply force - apply (rule add_mono) - apply (subgoal_tac "c * f xa \ max c ca * f xa") - apply force - apply (rule mult_right_mono) - apply (rule max.cobounded1) - apply assumption - apply (subgoal_tac "ca * g xa \ max c ca * g xa") - apply force - apply (rule mult_right_mono) - apply (rule max.cobounded2) - apply assumption - apply (rule abs_triangle_ineq) + apply (simp add: ring_distribs) + apply (subgoal_tac "\a xa + b xa\ \ \a xa\ + \b xa\") + apply (subgoal_tac "\a xa\ + \b xa\ \ max c ca * f xa + max c ca * g xa") + apply force + apply (rule add_mono) + apply (subgoal_tac "c * f xa \ max c ca * f xa") + apply force + apply (rule mult_right_mono) + apply (rule max.cobounded1) + apply assumption + apply (subgoal_tac "ca * g xa \ max c ca * g xa") + apply force + apply (rule mult_right_mono) + apply (rule max.cobounded2) + apply assumption + apply (rule abs_triangle_ineq) apply (rule add_nonneg_nonneg) - apply assumption+ + apply assumption+ done lemma bigo_bounded_alt: "\x. 0 \ f x \ \x. f x \ c * g x \ f \ O(g)" @@ -197,7 +199,7 @@ lemma bigo_bounded2: "\x. lb x \ f x \ \x. f x \ lb x + g x \ f \ lb +o O(g)" apply (rule set_minus_imp_plus) apply (rule bigo_bounded) - apply (auto simp add: fun_Compl_def func_plus) + apply (auto simp add: fun_Compl_def func_plus) apply (drule_tac x = x in spec)+ apply force done @@ -218,8 +220,8 @@ lemma bigo_abs3: "O(f) = O(\x. \f x\)" apply (rule equalityI) - apply (rule bigo_elt_subset) - apply (rule bigo_abs2) + apply (rule bigo_elt_subset) + apply (rule bigo_abs2) apply (rule bigo_elt_subset) apply (rule bigo_abs) done @@ -229,13 +231,13 @@ apply (rule set_minus_imp_plus) apply (subst fun_diff_def) proof - - assume a: "f - g \ O(h)" + assume *: "f - g \ O(h)" have "(\x. \f x\ - \g x\) =o O(\x. \\f x\ - \g x\\)" by (rule bigo_abs2) also have "\ \ O(\x. \f x - g x\)" apply (rule bigo_elt_subset) apply (rule bigo_bounded) - apply force + apply force apply (rule allI) apply (rule abs_triangle_ineq3) done @@ -244,23 +246,23 @@ apply (subst fun_diff_def) apply (rule bigo_abs) done - also from a have "\ \ O(h)" + also from * have "\ \ O(h)" by (rule bigo_elt_subset) finally show "(\x. \f x\ - \g x\) \ O(h)". qed lemma bigo_abs5: "f =o O(g) \ (\x. \f x\) =o O(g)" - by (unfold bigo_def, auto) + by (auto simp: bigo_def) -lemma bigo_elt_subset2 [intro]: "f \ g +o O(h) \ O(f) \ O(g) + O(h)" +lemma bigo_elt_subset2 [intro]: + assumes *: "f \ g +o O(h)" + shows "O(f) \ O(g) + O(h)" proof - - assume "f \ g +o O(h)" - also have "\ \ O(g) + O(h)" + note * + also have "g +o O(h) \ O(g) + O(h)" by (auto del: subsetI) also have "\ = O(\x. \g x\) + O(\x. \h x\)" - apply (subst bigo_abs3 [symmetric])+ - apply (rule refl) - done + by (subst bigo_abs3 [symmetric])+ (rule refl) also have "\ = O((\x. \g x\) + (\x. \h x\))" by (rule bigo_plus_eq [symmetric]) auto finally have "f \ \" . @@ -280,11 +282,11 @@ apply (rule allI) apply (erule_tac x = x in allE)+ apply (subgoal_tac "c * ca * \f x * g x\ = (c * \f x\) * (ca * \g x\)") - apply (erule ssubst) - apply (subst abs_mult) - apply (rule mult_mono) - apply assumption+ - apply auto + apply (erule ssubst) + apply (subst abs_mult) + apply (rule mult_mono) + apply assumption+ + apply auto apply (simp add: ac_simps abs_mult) done @@ -294,14 +296,14 @@ apply auto apply (drule_tac x = x in spec) apply (subgoal_tac "\f x\ * \b x\ \ \f x\ * (c * \g x\)") - apply (force simp add: ac_simps) + apply (force simp add: ac_simps) apply (rule mult_left_mono, assumption) apply (rule abs_ge_zero) done lemma bigo_mult3: "f \ O(h) \ g \ O(j) \ f * g \ O(h * j)" apply (rule subsetD) - apply (rule bigo_mult) + apply (rule bigo_mult) apply (erule set_times_intro, assumption) done @@ -309,7 +311,7 @@ apply (drule set_plus_imp_minus) apply (rule set_minus_imp_plus) apply (drule bigo_mult3 [where g = g and j = g]) - apply (auto simp add: algebra_simps) + apply (auto simp add: algebra_simps) done lemma bigo_mult5: @@ -339,28 +341,25 @@ finally show "h \ f *o O(g)" . qed -lemma bigo_mult6: - fixes f :: "'a \ 'b::linordered_field" - shows "\x. f x \ 0 \ O(f * g) = f *o O(g)" +lemma bigo_mult6: "\x. f x \ 0 \ O(f * g) = f *o O(g)" + for f :: "'a \ 'b::linordered_field" apply (rule equalityI) - apply (erule bigo_mult5) + apply (erule bigo_mult5) apply (rule bigo_mult2) done -lemma bigo_mult7: - fixes f :: "'a \ 'b::linordered_field" - shows "\x. f x \ 0 \ O(f * g) \ O(f) * O(g)" +lemma bigo_mult7: "\x. f x \ 0 \ O(f * g) \ O(f) * O(g)" + for f :: "'a \ 'b::linordered_field" apply (subst bigo_mult6) - apply assumption + apply assumption apply (rule set_times_mono3) apply (rule bigo_refl) done -lemma bigo_mult8: - fixes f :: "'a \ 'b::linordered_field" - shows "\x. f x \ 0 \ O(f * g) = O(f) * O(g)" +lemma bigo_mult8: "\x. f x \ 0 \ O(f * g) = O(f) * O(g)" + for f :: "'a \ 'b::linordered_field" apply (rule equalityI) - apply (erule bigo_mult7) + apply (erule bigo_mult7) apply (rule bigo_mult) done @@ -377,65 +376,63 @@ lemma bigo_minus3: "O(- f) = O(f)" by (auto simp add: bigo_def fun_Compl_def) -lemma bigo_plus_absorb_lemma1: "f \ O(g) \ f +o O(g) \ O(g)" +lemma bigo_plus_absorb_lemma1: + assumes *: "f \ O(g)" + shows "f +o O(g) \ O(g)" proof - - assume a: "f \ O(g)" - show "f +o O(g) \ O(g)" + have "f \ O(f)" by auto + then have "f +o O(g) \ O(f) + O(g)" + by (auto del: subsetI) + also have "\ \ O(g) + O(g)" proof - - have "f \ O(f)" by auto - then have "f +o O(g) \ O(f) + O(g)" + from * have "O(f) \ O(g)" by (auto del: subsetI) - also have "\ \ O(g) + O(g)" - proof - - from a have "O(f) \ O(g)" by (auto del: subsetI) - then show ?thesis by (auto del: subsetI) - qed - also have "\ \ O(g)" by simp - finally show ?thesis . + then show ?thesis + by (auto del: subsetI) qed + also have "\ \ O(g)" by simp + finally show ?thesis . qed -lemma bigo_plus_absorb_lemma2: "f \ O(g) \ O(g) \ f +o O(g)" +lemma bigo_plus_absorb_lemma2: + assumes *: "f \ O(g)" + shows "O(g) \ f +o O(g)" proof - - assume a: "f \ O(g)" - show "O(g) \ f +o O(g)" - proof - - from a have "- f \ O(g)" - by auto - then have "- f +o O(g) \ O(g)" - by (elim bigo_plus_absorb_lemma1) - then have "f +o (- f +o O(g)) \ f +o O(g)" - by auto - also have "f +o (- f +o O(g)) = O(g)" - by (simp add: set_plus_rearranges) - finally show ?thesis . - qed + from * have "- f \ O(g)" + by auto + then have "- f +o O(g) \ O(g)" + by (elim bigo_plus_absorb_lemma1) + then have "f +o (- f +o O(g)) \ f +o O(g)" + by auto + also have "f +o (- f +o O(g)) = O(g)" + by (simp add: set_plus_rearranges) + finally show ?thesis . qed lemma bigo_plus_absorb [simp]: "f \ O(g) \ f +o O(g) = O(g)" apply (rule equalityI) - apply (erule bigo_plus_absorb_lemma1) + apply (erule bigo_plus_absorb_lemma1) apply (erule bigo_plus_absorb_lemma2) done lemma bigo_plus_absorb2 [intro]: "f \ O(g) \ A \ O(g) \ f +o A \ O(g)" apply (subgoal_tac "f +o A \ f +o O(g)") - apply force+ + apply force+ done lemma bigo_add_commute_imp: "f \ g +o O(h) \ g \ f +o O(h)" apply (subst set_minus_plus [symmetric]) apply (subgoal_tac "g - f = - (f - g)") - apply (erule ssubst) - apply (rule bigo_minus) - apply (subst set_minus_plus) - apply assumption + apply (erule ssubst) + apply (rule bigo_minus) + apply (subst set_minus_plus) + apply assumption apply (simp add: ac_simps) done lemma bigo_add_commute: "f \ g +o O(h) \ g \ f +o O(h)" apply (rule iffI) - apply (erule bigo_add_commute_imp)+ + apply (erule bigo_add_commute_imp)+ done lemma bigo_const1: "(\x. c) \ O(\x. 1)" @@ -446,27 +443,24 @@ apply (rule bigo_const1) done -lemma bigo_const3: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ (\x. 1) \ O(\x. c)" +lemma bigo_const3: "c \ 0 \ (\x. 1) \ O(\x. c)" + for c :: "'a::linordered_field" apply (simp add: bigo_def) apply (rule_tac x = "\inverse c\" in exI) apply (simp add: abs_mult [symmetric]) done -lemma bigo_const4: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ O(\x. 1) \ O(\x. c)" +lemma bigo_const4: "c \ 0 \ O(\x. 1) \ O(\x. c)" + for c :: "'a::linordered_field" apply (rule bigo_elt_subset) apply (rule bigo_const3) apply assumption done -lemma bigo_const [simp]: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ O(\x. c) = O(\x. 1)" +lemma bigo_const [simp]: "c \ 0 \ O(\x. c) = O(\x. 1)" + for c :: "'a::linordered_field" apply (rule equalityI) - apply (rule bigo_const2) + apply (rule bigo_const2) apply (rule bigo_const4) apply assumption done @@ -482,37 +476,33 @@ apply (rule bigo_const_mult1) done -lemma bigo_const_mult3: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ f \ O(\x. c * f x)" +lemma bigo_const_mult3: "c \ 0 \ f \ O(\x. c * f x)" + for c :: "'a::linordered_field" apply (simp add: bigo_def) apply (rule_tac x = "\inverse c\" in exI) apply (simp add: abs_mult mult.assoc [symmetric]) done -lemma bigo_const_mult4: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ O(f) \ O(\x. c * f x)" +lemma bigo_const_mult4: "c \ 0 \ O(f) \ O(\x. c * f x)" + for c :: "'a::linordered_field" apply (rule bigo_elt_subset) apply (rule bigo_const_mult3) apply assumption done -lemma bigo_const_mult [simp]: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ O(\x. c * f x) = O(f)" +lemma bigo_const_mult [simp]: "c \ 0 \ O(\x. c * f x) = O(f)" + for c :: "'a::linordered_field" apply (rule equalityI) - apply (rule bigo_const_mult2) + apply (rule bigo_const_mult2) apply (erule bigo_const_mult4) done -lemma bigo_const_mult5 [simp]: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ (\x. c) *o O(f) = O(f)" +lemma bigo_const_mult5 [simp]: "c \ 0 \ (\x. c) *o O(f) = O(f)" + for c :: "'a::linordered_field" apply (auto del: subsetI) - apply (rule order_trans) - apply (rule bigo_mult2) - apply (simp add: func_times) + apply (rule order_trans) + apply (rule bigo_mult2) + apply (simp add: func_times) apply (auto intro!: simp add: bigo_def elt_set_times_def func_times) apply (rule_tac x = "\y. inverse c * x y" in exI) apply (simp add: mult.assoc [symmetric] abs_mult) @@ -525,18 +515,19 @@ apply (rule_tac x = "ca * \c\" in exI) apply (rule allI) apply (subgoal_tac "ca * \c\ * \f x\ = \c\ * (ca * \f x\)") - apply (erule ssubst) - apply (subst abs_mult) - apply (rule mult_left_mono) - apply (erule spec) - apply simp - apply(simp add: ac_simps) + apply (erule ssubst) + apply (subst abs_mult) + apply (rule mult_left_mono) + apply (erule spec) + apply simp + apply (simp add: ac_simps) done -lemma bigo_const_mult7 [intro]: "f =o O(g) \ (\x. c * f x) =o O(g)" +lemma bigo_const_mult7 [intro]: + assumes *: "f =o O(g)" + shows "(\x. c * f x) =o O(g)" proof - - assume "f =o O(g)" - then have "(\x. c) * f =o (\x. c) *o O(g)" + from * have "(\x. c) * f =o (\x. c) *o O(g)" by auto also have "(\x. c) * f = (\x. c * f x)" by (simp add: func_times) @@ -546,10 +537,9 @@ qed lemma bigo_compose1: "f =o O(g) \ (\x. f (k x)) =o O(\x. g (k x))" - unfolding bigo_def by auto + by (auto simp: bigo_def) -lemma bigo_compose2: "f =o g +o O(h) \ - (\x. f (k x)) =o (\x. g (k x)) +o O(\x. h(k x))" +lemma bigo_compose2: "f =o g +o O(h) \ (\x. f (k x)) =o (\x. g (k x)) +o O(\x. h(k x))" apply (simp only: set_minus_plus [symmetric] fun_Compl_def func_plus) apply (drule bigo_compose1) apply (simp add: fun_diff_def) @@ -564,21 +554,21 @@ apply (auto simp add: bigo_def) apply (rule_tac x = "\c\" in exI) apply (subst abs_of_nonneg) back back - apply (rule setsum_nonneg) - apply force + apply (rule setsum_nonneg) + apply force apply (subst setsum_right_distrib) apply (rule allI) apply (rule order_trans) - apply (rule setsum_abs) + apply (rule setsum_abs) apply (rule setsum_mono) apply (rule order_trans) - apply (drule spec)+ - apply (drule bspec)+ - apply assumption+ - apply (drule bspec) - apply assumption+ + apply (drule spec)+ + apply (drule bspec)+ + apply assumption+ + apply (drule bspec) + apply assumption+ apply (rule mult_right_mono) - apply (rule abs_ge_self) + apply (rule abs_ge_self) apply force done @@ -586,7 +576,7 @@ \c. \x y. \f x y\ \ c * h x y \ (\x. \y \ A x. f x y) =o O(\x. \y \ A x. h x y)" apply (rule bigo_setsum_main) - apply force + apply force apply clarsimp apply (rule_tac x = c in exI) apply force @@ -600,8 +590,8 @@ lemma bigo_setsum3: "f =o O(h) \ (\x. \y \ A x. l x y * f (k x y)) =o O(\x. \y \ A x. \l x y * h (k x y)\)" apply (rule bigo_setsum1) - apply (rule allI)+ - apply (rule abs_ge_zero) + apply (rule allI)+ + apply (rule abs_ge_zero) apply (unfold bigo_def) apply auto apply (rule_tac x = c in exI) @@ -609,7 +599,7 @@ apply (subst abs_mult)+ apply (subst mult.left_commute) apply (rule mult_left_mono) - apply (erule spec) + apply (erule spec) apply (rule abs_ge_zero) done @@ -632,13 +622,13 @@ O(\x. \y \ A x. l x y * h (k x y))" apply (subgoal_tac "(\x. \y \ A x. l x y * h (k x y)) = (\x. \y \ A x. \l x y * h (k x y)\)") - apply (erule ssubst) - apply (erule bigo_setsum3) + apply (erule ssubst) + apply (erule bigo_setsum3) apply (rule ext) apply (rule setsum.cong) - apply (rule refl) + apply (rule refl) apply (subst abs_of_nonneg) - apply auto + apply auto done lemma bigo_setsum6: "f =o g +o O(h) \ \x y. 0 \ l x y \ @@ -651,9 +641,9 @@ apply (subst setsum_subtractf [symmetric]) apply (subst right_diff_distrib [symmetric]) apply (rule bigo_setsum5) - apply (subst fun_diff_def [symmetric]) - apply (drule set_plus_imp_minus) - apply auto + apply (subst fun_diff_def [symmetric]) + apply (drule set_plus_imp_minus) + apply auto done @@ -662,25 +652,24 @@ lemma bigo_useful_intro: "A \ O(f) \ B \ O(f) \ A + B \ O(f)" apply (subst bigo_plus_idemp [symmetric]) apply (rule set_plus_mono2) - apply assumption+ + apply assumption+ done lemma bigo_useful_add: "f =o O(h) \ g =o O(h) \ f + g =o O(h)" apply (subst bigo_plus_idemp [symmetric]) apply (rule set_plus_intro) - apply assumption+ + apply assumption+ done -lemma bigo_useful_const_mult: - fixes c :: "'a::linordered_field" - shows "c \ 0 \ (\x. c) * f =o O(h) \ f =o O(h)" +lemma bigo_useful_const_mult: "c \ 0 \ (\x. c) * f =o O(h) \ f =o O(h)" + for c :: "'a::linordered_field" apply (rule subsetD) - apply (subgoal_tac "(\x. 1 / c) *o O(h) \ O(h)") - apply assumption - apply (rule bigo_const_mult6) + apply (subgoal_tac "(\x. 1 / c) *o O(h) \ O(h)") + apply assumption + apply (rule bigo_const_mult6) apply (subgoal_tac "f = (\x. 1 / c) * ((\x. c) * f)") - apply (erule ssubst) - apply (erule set_times_intro2) + apply (erule ssubst) + apply (erule set_times_intro2) apply (simp add: func_times) done @@ -690,10 +679,10 @@ apply (rule_tac x = c in exI) apply auto apply (case_tac "x = 0") - apply simp + apply simp apply (subgoal_tac "x = Suc (x - 1)") - apply (erule ssubst) back - apply (erule spec) + apply (erule ssubst) back + apply (erule spec) apply simp done @@ -702,10 +691,10 @@ f 0 = g 0 \ f =o g +o O(h)" apply (rule set_minus_imp_plus) apply (rule bigo_fix) - apply (subst fun_diff_def) - apply (subst fun_diff_def [symmetric]) - apply (rule set_plus_imp_minus) - apply simp + apply (subst fun_diff_def) + apply (subst fun_diff_def [symmetric]) + apply (rule set_plus_imp_minus) + apply simp apply (simp add: fun_diff_def) done @@ -721,7 +710,7 @@ apply (rule_tac x = c in exI) apply (rule allI) apply (rule order_trans) - apply (erule spec)+ + apply (erule spec)+ done lemma bigo_lesseq2: "f =o O(h) \ \x. \g x\ \ f x \ g =o O(h)" @@ -729,7 +718,7 @@ apply (rule allI) apply (drule_tac x = x in spec) apply (rule order_trans) - apply assumption + apply assumption apply (rule abs_ge_self) done @@ -737,7 +726,7 @@ apply (erule bigo_lesseq2) apply (rule allI) apply (subst abs_of_nonneg) - apply (erule spec)+ + apply (erule spec)+ done lemma bigo_lesseq4: "f =o O(h) \ @@ -745,75 +734,72 @@ apply (erule bigo_lesseq1) apply (rule allI) apply (subst abs_of_nonneg) - apply (erule spec)+ + apply (erule spec)+ done lemma bigo_lesso1: "\x. f x \ g x \ f x. max (f x - g x) 0) = 0") - apply (erule ssubst) - apply (rule bigo_zero) + apply (erule ssubst) + apply (rule bigo_zero) apply (unfold func_zero) apply (rule ext) apply (simp split: split_max) done -lemma bigo_lesso2: "f =o g +o O(h) \ - \x. 0 \ k x \ \x. k x \ f x \ k \x. 0 \ k x \ \x. k x \ f x \ k k x - g x") - apply simp - apply (subst abs_of_nonneg) - apply (drule_tac x = x in spec) back - apply (simp add: algebra_simps) - apply (subst diff_conv_add_uminus)+ - apply (rule add_right_mono) - apply (erule spec) + apply simp + apply (subst abs_of_nonneg) + apply (drule_tac x = x in spec) back + apply (simp add: algebra_simps) + apply (subst diff_conv_add_uminus)+ + apply (rule add_right_mono) + apply (erule spec) apply (rule order_trans) - prefer 2 - apply (rule abs_ge_zero) + prefer 2 + apply (rule abs_ge_zero) apply (simp add: algebra_simps) done -lemma bigo_lesso3: "f =o g +o O(h) \ - \x. 0 \ k x \ \x. g x \ k x \ f \x. 0 \ k x \ \x. g x \ k x \ f f x - k x") - apply simp - apply (subst abs_of_nonneg) - apply (drule_tac x = x in spec) back - apply (simp add: algebra_simps) - apply (subst diff_conv_add_uminus)+ - apply (rule add_left_mono) - apply (rule le_imp_neg_le) - apply (erule spec) + apply simp + apply (subst abs_of_nonneg) + apply (drule_tac x = x in spec) back + apply (simp add: algebra_simps) + apply (subst diff_conv_add_uminus)+ + apply (rule add_left_mono) + apply (rule le_imp_neg_le) + apply (erule spec) apply (rule order_trans) - prefer 2 - apply (rule abs_ge_zero) + prefer 2 + apply (rule abs_ge_zero) apply (simp add: algebra_simps) done -lemma bigo_lesso4: - fixes k :: "'a \ 'b::linordered_field" - shows "f g =o h +o O(k) \ f g =o h +o O(k) \ f 'b::linordered_field" apply (unfold lesso_def) apply (drule set_plus_imp_minus) apply (drule bigo_abs5) back apply (simp add: fun_diff_def) apply (drule bigo_useful_add) - apply assumption + apply assumption apply (erule bigo_lesseq2) back apply (rule allI) apply (auto simp add: func_plus fun_diff_def algebra_simps split: split_max abs_split) @@ -826,7 +812,7 @@ apply (rule allI) apply (drule_tac x = x in spec) apply (subgoal_tac "\max (f x - g x) 0\ = max (f x - g x) 0") - apply (clarsimp simp add: algebra_simps) + apply (clarsimp simp add: algebra_simps) apply (rule abs_of_nonneg) apply (rule max.cobounded2) done @@ -834,39 +820,41 @@ lemma lesso_add: "f k (f + k) g \ 0 \ f \ (0::real)" +lemma bigo_LIMSEQ1: "f =o O(g) \ g \ 0 \ f \ 0" + for f g :: "nat \ real" apply (simp add: LIMSEQ_iff bigo_alt_def) apply clarify apply (drule_tac x = "r / c" in spec) apply (drule mp) - apply simp + apply simp apply clarify apply (rule_tac x = no in exI) apply (rule allI) apply (drule_tac x = n in spec)+ apply (rule impI) apply (drule mp) - apply assumption + apply assumption apply (rule order_le_less_trans) - apply assumption + apply assumption apply (rule order_less_le_trans) - apply (subgoal_tac "c * \g n\ < c * (r / c)") - apply assumption - apply (erule mult_strict_left_mono) - apply assumption + apply (subgoal_tac "c * \g n\ < c * (r / c)") + apply assumption + apply (erule mult_strict_left_mono) + apply assumption apply simp done -lemma bigo_LIMSEQ2: "f =o g +o O(h) \ h \ 0 \ f \ a \ g \ (a::real)" +lemma bigo_LIMSEQ2: "f =o g +o O(h) \ h \ 0 \ f \ a \ g \ a" + for f g h :: "nat \ real" apply (drule set_plus_imp_minus) apply (drule bigo_LIMSEQ1) - apply assumption + apply assumption apply (simp only: fun_diff_def) apply (erule Lim_transform2) apply assumption diff -r ae33d1c2ab26 -r a7c5074a0251 src/HOL/Library/Convex.thy --- a/src/HOL/Library/Convex.thy Wed Jul 13 21:00:03 2016 +0200 +++ b/src/HOL/Library/Convex.thy Wed Jul 13 21:30:41 2016 +0200 @@ -6,7 +6,7 @@ section \Convexity in real vector spaces\ theory Convex -imports Product_Vector + imports Product_Vector begin subsection \Convexity\ @@ -24,24 +24,27 @@ shows "u *\<^sub>R x + v *\<^sub>R y \ s" using assms unfolding convex_def by fast -lemma convex_alt: - "convex s \ (\x\s. \y\s. \u. 0 \ u \ u \ 1 \ ((1 - u) *\<^sub>R x + u *\<^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)" (is "_ \ ?alt") proof - assume alt[rule_format]: ?alt - { - fix x y and u v :: real - assume mem: "x \ s" "y \ s" - 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 - } - then show "convex s" - unfolding convex_def by auto -qed (auto simp: convex_def) + show "convex s" if alt: ?alt + proof - + { + fix x y and u v :: real + assume mem: "x \ s" "y \ s" + 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 [rule_format, OF mem] by auto + } + then show ?thesis + unfolding convex_def by auto + qed + show ?alt if "convex s" + using that by (auto simp: convex_def) +qed lemma convexD_alt: assumes "convex s" "a \ s" "b \ s" "0 \ u" "u \ 1" @@ -53,7 +56,7 @@ shows "((u/(u+v)) *\<^sub>R x + (v/(u+v)) *\<^sub>R y) \ S" apply (rule convexD) using assms - apply (simp_all add: zero_le_divide_iff add_divide_distrib [symmetric]) + apply (simp_all add: zero_le_divide_iff add_divide_distrib [symmetric]) done lemma convex_empty[intro,simp]: "convex {}" @@ -270,12 +273,12 @@ case False then show ?thesis using *[rule_format, of "{x, y}" "\ z. if z = x then 1 - \ else \"] ** - by auto + by auto next case True then show ?thesis using *[rule_format, of "{x, y}" "\ z. 1"] ** - by (auto simp: field_simps real_vector.scale_left_diff_distrib) + by (auto simp: field_simps real_vector.scale_left_diff_distrib) qed qed qed @@ -293,8 +296,8 @@ 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.If_cases if_distrib if_distrib_arg) + using sum[THEN spec[where x="\x. if x\t then u x else 0"]] as * + by (auto simp: assms setsum.If_cases if_distrib if_distrib_arg) qed (erule_tac x=s in allE, erule_tac x=u in allE, auto) @@ -306,39 +309,45 @@ lemma convex_onI [intro?]: assumes "\t x y. t > 0 \ t < 1 \ x \ A \ y \ A \ - f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" - shows "convex_on A f" + f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" + shows "convex_on A f" unfolding convex_on_def proof clarify - fix x y u v assume A: "x \ A" "y \ A" "(u::real) \ 0" "v \ 0" "u + v = 1" - from A(5) have [simp]: "v = 1 - u" by (simp add: algebra_simps) - from A(1-4) show "f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" using assms[of u y x] + fix x y + fix u v :: real + assume A: "x \ A" "y \ A" "u \ 0" "v \ 0" "u + v = 1" + from A(5) have [simp]: "v = 1 - u" + by (simp add: algebra_simps) + from A(1-4) show "f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" + using assms[of u y x] by (cases "u = 0 \ u = 1") (auto simp: algebra_simps) qed lemma convex_on_linorderI [intro?]: fixes A :: "('a::{linorder,real_vector}) set" assumes "\t x y. t > 0 \ t < 1 \ x \ A \ y \ A \ x < y \ - f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" - shows "convex_on A f" + f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" + shows "convex_on A f" proof - fix t x y assume A: "x \ A" "y \ A" "(t::real) > 0" "t < 1" - with assms[of t x y] assms[of "1 - t" y x] + fix x y + fix t :: real + assume A: "x \ A" "y \ A" "t > 0" "t < 1" + with assms [of t x y] assms [of "1 - t" y x] show "f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" by (cases x y rule: linorder_cases) (auto simp: algebra_simps) qed lemma convex_onD: assumes "convex_on A f" - shows "\t x y. t \ 0 \ t \ 1 \ x \ A \ y \ A \ - f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" - using assms unfolding convex_on_def by auto + shows "\t x y. t \ 0 \ t \ 1 \ x \ A \ y \ A \ + f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" + using assms by (auto simp: convex_on_def) lemma convex_onD_Icc: assumes "convex_on {x..y} f" "x \ (y :: _ :: {real_vector,preorder})" - shows "\t. t \ 0 \ t \ 1 \ - f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" - using assms(2) by (intro convex_onD[OF assms(1)]) simp_all + shows "\t. t \ 0 \ t \ 1 \ + f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" + using assms(2) by (intro convex_onD [OF assms(1)]) simp_all lemma convex_on_subset: "convex_on t f \ s \ t \ convex_on s f" unfolding convex_on_def by auto @@ -370,7 +379,8 @@ and "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)" + have *: "u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" + for u c fx v fy :: real by (simp add: field_simps) show ?thesis using assms(2) and mult_left_mono [OF _ assms(1)] unfolding convex_on_def and * by auto @@ -517,20 +527,24 @@ assume *: "y > 0" "x > 0" "\ \ 0" "\ \ 1" { assume "\ = 0" - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y = y" by simp - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using * 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 * by simp } moreover { assume "\ = 1" - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using * by simp + then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" + using * by simp } moreover { assume "\ \ 1" "\ \ 0" - then have "\ > 0" "(1 - \) > 0" using * by auto - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using * - by (auto simp: add_pos_pos) + then have "\ > 0" "(1 - \) > 0" + using * by auto + then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" + using * by (auto simp: add_pos_pos) } ultimately show "(1 - \) *\<^sub>R y + \ *\<^sub>R x > 0" by fastforce @@ -550,11 +564,14 @@ using assms proof (induct s arbitrary: a rule: finite_ne_induct) case (singleton i) - then have ai: "a i = 1" by auto - then show ?case by auto + then have ai: "a i = 1" + by auto + then show ?case + by auto next case (insert i s) - then have "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" @@ -593,8 +610,7 @@ unfolding setsum_divide_distrib by simp have "convex C" using insert by auto then have asum: "(\ j \ s. ?a j *\<^sub>R y j) \ C" - using insert convex_setsum[OF \finite s\ - \convex C\ a1 a_nonneg] by auto + using insert 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))" using a_nonneg a1 insert by blast have "f (\ j \ insert i s. a j *\<^sub>R y j) = f ((\ j \ s. a j *\<^sub>R y j) + a i *\<^sub>R y i)" @@ -611,10 +627,12 @@ using conv[of "y i" "(\ j \ s. ?a j *\<^sub>R y j)" "a i", OF yai(1) asum yai(2) ai1] by (auto simp: add.commute) also have "\ \ (1 - a i) * (\ j \ s. ?a j * f (y j)) + a i * f (y i)" - using add_right_mono[OF mult_left_mono[of _ _ "1 - a i", - OF asum_le less_imp_le[OF i0]], of "a i * f (y i)"] by simp + using add_right_mono [OF mult_left_mono [of _ _ "1 - a i", + OF asum_le less_imp_le[OF i0]], of "a i * f (y i)"] + by simp also have "\ = (\ j \ s. (1 - a i) * ?a j * f (y j)) + a i * f (y i)" - unfolding setsum_right_distrib[of "1 - a i" "\ j. ?a j * f (y j)"] using i0 by auto + unfolding setsum_right_distrib[of "1 - a i" "\ j. ?a j * f (y j)"] + using i0 by auto also have "\ = (\ j \ s. a j * f (y j)) + a i * f (y i)" using i0 by auto also have "\ = (\ j \ insert i s. a j * f (y j))" @@ -635,9 +653,9 @@ fix \ :: real assume *: "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" + have "0 \ u \ 0 \ v \ u + v = 1 \ f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" for u v by auto - from this[of "\" "1 - \", simplified] * + from this [of "\" "1 - \", simplified] * show "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" by auto next @@ -701,8 +719,8 @@ using * 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 *(2)] \\ \ 0\] - mult_left_mono[OF leq[OF xpos *(3)] \1 - \ \ 0\]] + using add_mono [OF mult_left_mono [OF leq [OF xpos *(2)] \\ \ 0\] + mult_left_mono [OF leq [OF xpos *(3)] \1 - \ \ 0\]] by auto then have "\ * f x + (1 - \) * f y - f ?x \ 0" by (auto simp: field_simps) @@ -728,14 +746,14 @@ have "?\ * x + (1 - ?\) * y = (y - z) * x / (y - x) + (1 - (y - z) / (y - x)) * y" by (auto simp: 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 by (simp only: add_divide_distrib) (auto simp: field_simps) also have "\ = z" using assms by (auto simp: field_simps) finally show ?thesis using comb by auto qed - show "z \ C" using z less assms - unfolding atLeastAtMost_iff le_less by auto + show "z \ C" + using z less assms by (auto simp: le_less) qed lemma f''_imp_f': @@ -744,20 +762,21 @@ 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" + and x: "x \ C" + and y: "y \ C" shows "f' x * (y - x) \ f y - f x" using assms proof - - { - fix x y :: real - assume *: "x \ C" "y \ C" "y > x" - then have ge: "y - x > 0" "y - x \ 0" + have less_imp: "f y - f x \ f' x * (y - x)" "f' y * (x - y) \ f x - f y" + if *: "x \ C" "y \ C" "y > x" for x y :: real + proof - + from * have ge: "y - x > 0" "y - x \ 0" by auto from * 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]]] + THEN f', THEN MVT2[OF \x < y\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] by auto then have "z1 \ C" using atMostAtLeast_subset_convex \convex C\ \x \ C\ \y \ C\ \x < y\ @@ -766,11 +785,11 @@ by (simp add: field_simps) obtain z2 where z2: "z2 > x" "z2 < z1" "f' z1 - f' x = (z1 - x) * f'' z2" using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \x \ C\ \z1 \ C\ \x < z1\], - THEN f'', THEN MVT2[OF \x < z1\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 + THEN f'', THEN MVT2[OF \x < z1\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 by auto obtain z3 where z3: "z3 > z1" "z3 < y" "f' y - f' z1 = (y - z1) * f'' z3" using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \z1 \ C\ \y \ C\ \z1 < y\], - THEN f'', THEN MVT2[OF \z1 < y\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 + THEN f'', THEN MVT2[OF \z1 < y\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 by auto have "f' y - (f x - f y) / (x - y) = f' y - f' z1" using * z1' by auto @@ -818,22 +837,18 @@ by (simp add: algebra_simps) 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" + then show "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" - then have"f y - f x \ f' x * (y - x)" - unfolding neq_iff using less_imp by auto - } - moreover - { - fix x y :: real - assume "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 + show ?thesis + proof (cases "x = y") + case True + with x y show ?thesis by auto + next + case False + with less_imp x y show ?thesis + by (auto simp: neq_iff) + qed qed lemma f''_ge0_imp_convex: @@ -855,10 +870,10 @@ using DERIV_log by auto then have f': "\z. z > 0 \ DERIV (\ z. - log b z) z :> - 1 / (ln b * z)" by (auto simp: DERIV_minus) - 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 \ + have "\z::real. z > 0 \ DERIV (\ z. (- 1 / ln b) * inverse z) z :> (- 1 / ln b) * (- (inverse z ^ Suc (Suc 0)))" by auto then have f''0: "\z::real. z > 0 \ @@ -866,9 +881,9 @@ unfolding inverse_eq_divide by (auto simp: mult.assoc) have f''_ge0: "\z::real. z > 0 \ 1 / (ln b * z * z) \ 0" using \b > 1\ by (auto intro!: less_imp_le) - from f''_ge0_imp_convex[OF pos_is_convex, - unfolded greaterThan_iff, OF f' f''0 f''_ge0] - show ?thesis by auto + from f''_ge0_imp_convex[OF pos_is_convex, unfolded greaterThan_iff, OF f' f''0 f''_ge0] + show ?thesis + by auto qed @@ -876,45 +891,59 @@ lemma convex_on_realI: assumes "connected A" - assumes "\x. x \ A \ (f has_real_derivative f' x) (at x)" - assumes "\x y. x \ A \ y \ A \ x \ y \ f' x \ f' y" - shows "convex_on A f" + and "\x. x \ A \ (f has_real_derivative f' x) (at x)" + and "\x y. x \ A \ y \ A \ x \ y \ f' x \ f' y" + shows "convex_on A f" proof (rule convex_on_linorderI) fix t x y :: real - assume t: "t > 0" "t < 1" and xy: "x \ A" "y \ A" "x < y" + assume t: "t > 0" "t < 1" + assume xy: "x \ A" "y \ A" "x < y" define z where "z = (1 - t) * x + t * y" - with \connected A\ and xy have ivl: "{x..y} \ A" using connected_contains_Icc by blast + with \connected A\ and xy have ivl: "{x..y} \ A" + using connected_contains_Icc by blast - from xy t have xz: "z > x" by (simp add: z_def algebra_simps) - have "y - z = (1 - t) * (y - x)" by (simp add: z_def algebra_simps) - also from xy t have "... > 0" by (intro mult_pos_pos) simp_all - finally have yz: "z < y" by simp + from xy t have xz: "z > x" + by (simp add: z_def algebra_simps) + have "y - z = (1 - t) * (y - x)" + by (simp add: z_def algebra_simps) + also from xy t have "\ > 0" + by (intro mult_pos_pos) simp_all + finally have yz: "z < y" + by simp from assms xz yz ivl t have "\\. \ > x \ \ < z \ f z - f x = (z - x) * f' \" by (intro MVT2) (auto intro!: assms(2)) - then obtain \ where \: "\ > x" "\ < z" "f' \ = (f z - f x) / (z - x)" by auto + then obtain \ where \: "\ > x" "\ < z" "f' \ = (f z - f x) / (z - x)" + by auto from assms xz yz ivl t have "\\. \ > z \ \ < y \ f y - f z = (y - z) * f' \" by (intro MVT2) (auto intro!: assms(2)) - then obtain \ where \: "\ > z" "\ < y" "f' \ = (f y - f z) / (y - z)" by auto + then obtain \ where \: "\ > z" "\ < y" "f' \ = (f y - f z) / (y - z)" + by auto from \(3) have "(f y - f z) / (y - z) = f' \" .. - also from \ \ ivl have "\ \ A" "\ \ A" by auto - with \ \ have "f' \ \ f' \" by (intro assms(3)) auto + also from \ \ ivl have "\ \ A" "\ \ A" + by auto + with \ \ have "f' \ \ f' \" + by (intro assms(3)) auto also from \(3) have "f' \ = (f z - f x) / (z - x)" . finally have "(f y - f z) * (z - x) \ (f z - f x) * (y - z)" using xz yz by (simp add: field_simps) - also have "z - x = t * (y - x)" by (simp add: z_def algebra_simps) - also have "y - z = (1 - t) * (y - x)" by (simp add: z_def algebra_simps) - finally have "(f y - f z) * t \ (f z - f x) * (1 - t)" using xy by simp - thus "(1 - t) * f x + t * f y \ f ((1 - t) *\<^sub>R x + t *\<^sub>R y)" + also have "z - x = t * (y - x)" + by (simp add: z_def algebra_simps) + also have "y - z = (1 - t) * (y - x)" + by (simp add: z_def algebra_simps) + finally have "(f y - f z) * t \ (f z - f x) * (1 - t)" + using xy by simp + then show "(1 - t) * f x + t * f y \ f ((1 - t) *\<^sub>R x + t *\<^sub>R y)" by (simp add: z_def algebra_simps) qed lemma convex_on_inverse: assumes "A \ {0<..}" - shows "convex_on A (inverse :: real \ real)" + shows "convex_on A (inverse :: real \ real)" proof (rule convex_on_subset[OF _ assms], intro convex_on_realI[of _ _ "\x. -inverse (x^2)"]) - fix u v :: real assume "u \ {0<..}" "v \ {0<..}" "u \ v" + fix u v :: real + assume "u \ {0<..}" "v \ {0<..}" "u \ v" with assms show "-inverse (u^2) \ -inverse (v^2)" by (intro le_imp_neg_le le_imp_inverse_le power_mono) (simp_all) qed (insert assms, auto intro!: derivative_eq_intros simp: divide_simps power2_eq_square) @@ -922,40 +951,47 @@ lemma convex_onD_Icc': assumes "convex_on {x..y} f" "c \ {x..y}" defines "d \ y - x" - shows "f c \ (f y - f x) / d * (c - x) + f x" -proof (cases y x rule: linorder_cases) - assume less: "x < y" - hence d: "d > 0" by (simp add: d_def) + shows "f c \ (f y - f x) / d * (c - x) + f x" +proof (cases x y rule: linorder_cases) + case less + then have d: "d > 0" + by (simp add: d_def) from assms(2) less have A: "0 \ (c - x) / d" "(c - x) / d \ 1" by (simp_all add: d_def divide_simps) - have "f c = f (x + (c - x) * 1)" by simp - also from less have "1 = ((y - x) / d)" by (simp add: d_def) + have "f c = f (x + (c - x) * 1)" + by simp + also from less have "1 = ((y - x) / d)" + by (simp add: d_def) also from d have "x + (c - x) * \ = (1 - (c - x) / d) *\<^sub>R x + ((c - x) / d) *\<^sub>R y" by (simp add: field_simps) - also have "f \ \ (1 - (c - x) / d) * f x + (c - x) / d * f y" using assms less - by (intro convex_onD_Icc) simp_all - also from d have "\ = (f y - f x) / d * (c - x) + f x" by (simp add: field_simps) + also have "f \ \ (1 - (c - x) / d) * f x + (c - x) / d * f y" + using assms less by (intro convex_onD_Icc) simp_all + also from d have "\ = (f y - f x) / d * (c - x) + f x" + by (simp add: field_simps) finally show ?thesis . qed (insert assms(2), simp_all) lemma convex_onD_Icc'': assumes "convex_on {x..y} f" "c \ {x..y}" defines "d \ y - x" - shows "f c \ (f x - f y) / d * (y - c) + f y" -proof (cases y x rule: linorder_cases) - assume less: "x < y" - hence d: "d > 0" by (simp add: d_def) + shows "f c \ (f x - f y) / d * (y - c) + f y" +proof (cases x y rule: linorder_cases) + case less + then have d: "d > 0" + by (simp add: d_def) from assms(2) less have A: "0 \ (y - c) / d" "(y - c) / d \ 1" by (simp_all add: d_def divide_simps) - have "f c = f (y - (y - c) * 1)" by simp - also from less have "1 = ((y - x) / d)" by (simp add: d_def) + have "f c = f (y - (y - c) * 1)" + by simp + also from less have "1 = ((y - x) / d)" + by (simp add: d_def) also from d have "y - (y - c) * \ = (1 - (1 - (y - c) / d)) *\<^sub>R x + (1 - (y - c) / d) *\<^sub>R y" by (simp add: field_simps) - also have "f \ \ (1 - (1 - (y - c) / d)) * f x + (1 - (y - c) / d) * f y" using assms less - by (intro convex_onD_Icc) (simp_all add: field_simps) - also from d have "\ = (f x - f y) / d * (y - c) + f y" by (simp add: field_simps) + also have "f \ \ (1 - (1 - (y - c) / d)) * f x + (1 - (y - c) / d) * f y" + using assms less by (intro convex_onD_Icc) (simp_all add: field_simps) + also from d have "\ = (f x - f y) / d * (y - c) + f y" + by (simp add: field_simps) finally show ?thesis . qed (insert assms(2), simp_all) - end diff -r ae33d1c2ab26 -r a7c5074a0251 src/HOL/Library/Mapping.thy --- a/src/HOL/Library/Mapping.thy Wed Jul 13 21:00:03 2016 +0200 +++ b/src/HOL/Library/Mapping.thy Wed Jul 13 21:30:41 2016 +0200 @@ -74,11 +74,11 @@ (\k. if k < length xs then Some (xs ! k) else None) (\k. if k < length ys then Some (ys ! k) else None)" apply induct - apply auto + apply auto unfolding rel_fun_def apply clarsimp apply (case_tac xa) - apply (auto dest: list_all2_lengthD list_all2_nthD) + apply (auto dest: list_all2_lengthD list_all2_nthD) done qed diff -r ae33d1c2ab26 -r a7c5074a0251 src/HOL/Library/Set_Algebras.thy --- a/src/HOL/Library/Set_Algebras.thy Wed Jul 13 21:00:03 2016 +0200 +++ b/src/HOL/Library/Set_Algebras.thy Wed Jul 13 21:30:41 2016 +0200 @@ -1,24 +1,26 @@ (* Title: HOL/Library/Set_Algebras.thy - Author: Jeremy Avigad and Kevin Donnelly; Florian Haftmann, TUM + Author: Jeremy Avigad + Author: Kevin Donnelly + Author: Florian Haftmann, TUM *) section \Algebraic operations on sets\ theory Set_Algebras -imports Main + imports Main begin text \ - This library lifts operations like addition and multiplication to - sets. It was designed to support asymptotic calculations. See the - comments at the top of theory \BigO\. + This library lifts operations like addition and multiplication to sets. It + was designed to support asymptotic calculations. See the comments at the top + of @{file "BigO.thy"}. \ instantiation set :: (plus) plus begin -definition plus_set :: "'a::plus set \ 'a set \ 'a set" where - set_plus_def: "A + B = {c. \a\A. \b\B. c = a + b}" +definition plus_set :: "'a::plus set \ 'a set \ 'a set" + where set_plus_def: "A + B = {c. \a\A. \b\B. c = a + b}" instance .. @@ -27,8 +29,8 @@ instantiation set :: (times) times begin -definition times_set :: "'a::times set \ 'a set \ 'a set" where - set_times_def: "A * B = {c. \a\A. \b\B. c = a * b}" +definition times_set :: "'a::times set \ 'a set \ 'a set" + where set_times_def: "A * B = {c. \a\A. \b\B. c = a * b}" instance .. @@ -37,8 +39,7 @@ instantiation set :: (zero) zero begin -definition - set_zero[simp]: "(0::'a::zero set) = {0}" +definition set_zero[simp]: "(0::'a::zero set) = {0}" instance .. @@ -47,21 +48,20 @@ instantiation set :: (one) one begin -definition - set_one[simp]: "(1::'a::one set) = {1}" +definition set_one[simp]: "(1::'a::one set) = {1}" instance .. end -definition elt_set_plus :: "'a::plus \ 'a set \ 'a set" (infixl "+o" 70) where - "a +o B = {c. \b\B. c = a + b}" +definition elt_set_plus :: "'a::plus \ 'a set \ 'a set" (infixl "+o" 70) + where "a +o B = {c. \b\B. c = a + b}" -definition elt_set_times :: "'a::times \ 'a set \ 'a set" (infixl "*o" 80) where - "a *o B = {c. \b\B. c = a * b}" +definition elt_set_times :: "'a::times \ 'a set \ 'a set" (infixl "*o" 80) + where "a *o B = {c. \b\B. c = a * b}" -abbreviation (input) elt_set_eq :: "'a \ 'a set \ bool" (infix "=o" 50) where - "x =o A \ x \ A" +abbreviation (input) elt_set_eq :: "'a \ 'a set \ bool" (infix "=o" 50) + where "x =o A \ x \ A" instance set :: (semigroup_add) semigroup_add by standard (force simp add: set_plus_def add.assoc) @@ -98,19 +98,21 @@ lemma set_plus_intro2 [intro]: "b \ C \ a + b \ a +o C" by (auto simp add: elt_set_plus_def) -lemma set_plus_rearrange: - "((a::'a::comm_monoid_add) +o C) + (b +o D) = (a + b) +o (C + D)" +lemma set_plus_rearrange: "(a +o C) + (b +o D) = (a + b) +o (C + D)" + for a b :: "'a::comm_monoid_add" apply (auto simp add: elt_set_plus_def set_plus_def ac_simps) apply (rule_tac x = "ba + bb" in exI) - apply (auto simp add: ac_simps) + apply (auto simp add: ac_simps) apply (rule_tac x = "aa + a" in exI) apply (auto simp add: ac_simps) done -lemma set_plus_rearrange2: "(a::'a::semigroup_add) +o (b +o C) = (a + b) +o C" +lemma set_plus_rearrange2: "a +o (b +o C) = (a + b) +o C" + for a b :: "'a::semigroup_add" by (auto simp add: elt_set_plus_def add.assoc) -lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) + C = a +o (B + C)" +lemma set_plus_rearrange3: "(a +o B) + C = a +o (B + C)" + for a :: "'a::semigroup_add" apply (auto simp add: elt_set_plus_def set_plus_def) apply (blast intro: ac_simps) apply (rule_tac x = "a + aa" in exI) @@ -121,7 +123,8 @@ apply (auto simp add: ac_simps) done -theorem set_plus_rearrange4: "C + ((a::'a::comm_monoid_add) +o D) = a +o (C + D)" +theorem set_plus_rearrange4: "C + (a +o D) = a +o (C + D)" + for a :: "'a::comm_monoid_add" apply (auto simp add: elt_set_plus_def set_plus_def ac_simps) apply (rule_tac x = "aa + ba" in exI) apply (auto simp add: ac_simps) @@ -133,13 +136,15 @@ lemma set_plus_mono [intro!]: "C \ D \ a +o C \ a +o D" by (auto simp add: elt_set_plus_def) -lemma set_plus_mono2 [intro]: "(C::'a::plus set) \ D \ E \ F \ C + E \ D + F" +lemma set_plus_mono2 [intro]: "C \ D \ E \ F \ C + E \ D + F" + for C D E F :: "'a::plus set" by (auto simp add: set_plus_def) lemma set_plus_mono3 [intro]: "a \ C \ a +o D \ C + D" by (auto simp add: elt_set_plus_def set_plus_def) -lemma set_plus_mono4 [intro]: "(a::'a::comm_monoid_add) \ C \ a +o D \ D + C" +lemma set_plus_mono4 [intro]: "a \ C \ a +o D \ D + C" + for a :: "'a::comm_monoid_add" by (auto simp add: elt_set_plus_def set_plus_def ac_simps) lemma set_plus_mono5: "a \ C \ B \ D \ a +o B \ C + D" @@ -166,33 +171,45 @@ apply auto done -lemma set_plus_mono4_b: "(a::'a::comm_monoid_add) : C \ x \ a +o D \ x \ D + C" +lemma set_plus_mono4_b: "a \ C \ x \ a +o D \ x \ D + C" + for a x :: "'a::comm_monoid_add" apply (frule set_plus_mono4) apply auto done -lemma set_zero_plus [simp]: "(0::'a::comm_monoid_add) +o C = C" +lemma set_zero_plus [simp]: "0 +o C = C" + for C :: "'a::comm_monoid_add set" by (auto simp add: elt_set_plus_def) -lemma set_zero_plus2: "(0::'a::comm_monoid_add) \ A \ B \ A + B" +lemma set_zero_plus2: "0 \ A \ B \ A + B" + for A B :: "'a::comm_monoid_add set" apply (auto simp add: set_plus_def) apply (rule_tac x = 0 in bexI) apply (rule_tac x = x in bexI) apply (auto simp add: ac_simps) done -lemma set_plus_imp_minus: "(a::'a::ab_group_add) : b +o C \ (a - b) \ C" +lemma set_plus_imp_minus: "a \ b +o C \ a - b \ C" + for a b :: "'a::ab_group_add" by (auto simp add: elt_set_plus_def ac_simps) -lemma set_minus_imp_plus: "(a::'a::ab_group_add) - b : C \ a \ b +o C" +lemma set_minus_imp_plus: "a - b \ C \ a \ b +o C" + for a b :: "'a::ab_group_add" apply (auto simp add: elt_set_plus_def ac_simps) apply (subgoal_tac "a = (a + - b) + b") - apply (rule bexI, assumption) - apply (auto simp add: ac_simps) + apply (rule bexI) + apply assumption + apply (auto simp add: ac_simps) done -lemma set_minus_plus: "(a::'a::ab_group_add) - b \ C \ a \ b +o C" - by (rule iffI, rule set_minus_imp_plus, assumption, rule set_plus_imp_minus) +lemma set_minus_plus: "a - b \ C \ a \ b +o C" + for a b :: "'a::ab_group_add" + apply (rule iffI) + apply (rule set_minus_imp_plus) + apply assumption + apply (rule set_plus_imp_minus) + apply assumption + done lemma set_times_intro [intro]: "a \ C \ b \ D \ a * b \ C * D" by (auto simp add: set_times_def) @@ -205,8 +222,8 @@ lemma set_times_intro2 [intro!]: "b \ C \ a * b \ a *o C" by (auto simp add: elt_set_times_def) -lemma set_times_rearrange: - "((a::'a::comm_monoid_mult) *o C) * (b *o D) = (a * b) *o (C * D)" +lemma set_times_rearrange: "(a *o C) * (b *o D) = (a * b) *o (C * D)" + for a b :: "'a::comm_monoid_mult" apply (auto simp add: elt_set_times_def set_times_def) apply (rule_tac x = "ba * bb" in exI) apply (auto simp add: ac_simps) @@ -214,12 +231,12 @@ apply (auto simp add: ac_simps) done -lemma set_times_rearrange2: - "(a::'a::semigroup_mult) *o (b *o C) = (a * b) *o C" +lemma set_times_rearrange2: "a *o (b *o C) = (a * b) *o C" + for a b :: "'a::semigroup_mult" by (auto simp add: elt_set_times_def mult.assoc) -lemma set_times_rearrange3: - "((a::'a::semigroup_mult) *o B) * C = a *o (B * C)" +lemma set_times_rearrange3: "(a *o B) * C = a *o (B * C)" + for a :: "'a::semigroup_mult" apply (auto simp add: elt_set_times_def set_times_def) apply (blast intro: ac_simps) apply (rule_tac x = "a * aa" in exI) @@ -230,8 +247,8 @@ apply (auto simp add: ac_simps) done -theorem set_times_rearrange4: - "C * ((a::'a::comm_monoid_mult) *o D) = a *o (C * D)" +theorem set_times_rearrange4: "C * (a *o D) = a *o (C * D)" + for a :: "'a::comm_monoid_mult" apply (auto simp add: elt_set_times_def set_times_def ac_simps) apply (rule_tac x = "aa * ba" in exI) apply (auto simp add: ac_simps) @@ -243,13 +260,15 @@ lemma set_times_mono [intro]: "C \ D \ a *o C \ a *o D" by (auto simp add: elt_set_times_def) -lemma set_times_mono2 [intro]: "(C::'a::times set) \ D \ E \ F \ C * E \ D * F" +lemma set_times_mono2 [intro]: "C \ D \ E \ F \ C * E \ D * F" + for C D E F :: "'a::times set" by (auto simp add: set_times_def) lemma set_times_mono3 [intro]: "a \ C \ a *o D \ C * D" by (auto simp add: elt_set_times_def set_times_def) -lemma set_times_mono4 [intro]: "(a::'a::comm_monoid_mult) : C \ a *o D \ D * C" +lemma set_times_mono4 [intro]: "a \ C \ a *o D \ D * C" + for a :: "'a::comm_monoid_mult" by (auto simp add: elt_set_times_def set_times_def ac_simps) lemma set_times_mono5: "a \ C \ B \ D \ a *o B \ C * D" @@ -276,30 +295,31 @@ apply auto done -lemma set_times_mono4_b: "(a::'a::comm_monoid_mult) \ C \ x \ a *o D \ x \ D * C" +lemma set_times_mono4_b: "a \ C \ x \ a *o D \ x \ D * C" + for a x :: "'a::comm_monoid_mult" apply (frule set_times_mono4) apply auto done -lemma set_one_times [simp]: "(1::'a::comm_monoid_mult) *o C = C" +lemma set_one_times [simp]: "1 *o C = C" + for C :: "'a::comm_monoid_mult set" by (auto simp add: elt_set_times_def) -lemma set_times_plus_distrib: - "(a::'a::semiring) *o (b +o C) = (a * b) +o (a *o C)" +lemma set_times_plus_distrib: "a *o (b +o C) = (a * b) +o (a *o C)" + for a b :: "'a::semiring" by (auto simp add: elt_set_plus_def elt_set_times_def ring_distribs) -lemma set_times_plus_distrib2: - "(a::'a::semiring) *o (B + C) = (a *o B) + (a *o C)" +lemma set_times_plus_distrib2: "a *o (B + C) = (a *o B) + (a *o C)" + for a :: "'a::semiring" apply (auto simp add: set_plus_def elt_set_times_def ring_distribs) apply blast apply (rule_tac x = "b + bb" in exI) apply (auto simp add: ring_distribs) done -lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) * D \ a *o D + C * D" - apply (auto simp add: - elt_set_plus_def elt_set_times_def set_times_def - set_plus_def ring_distribs) +lemma set_times_plus_distrib3: "(a +o C) * D \ a *o D + C * D" + for a :: "'a::semiring" + apply (auto simp: elt_set_plus_def elt_set_times_def set_times_def set_plus_def ring_distribs) apply auto done @@ -307,23 +327,25 @@ set_times_plus_distrib set_times_plus_distrib2 -lemma set_neg_intro: "(a::'a::ring_1) \ (- 1) *o C \ - a \ C" +lemma set_neg_intro: "a \ (- 1) *o C \ - a \ C" + for a :: "'a::ring_1" by (auto simp add: elt_set_times_def) -lemma set_neg_intro2: "(a::'a::ring_1) \ C \ - a \ (- 1) *o C" +lemma set_neg_intro2: "a \ C \ - a \ (- 1) *o C" + for a :: "'a::ring_1" by (auto simp add: elt_set_times_def) lemma set_plus_image: "S + T = (\(x, y). x + y) ` (S \ T)" - unfolding set_plus_def by (fastforce simp: image_iff) + by (fastforce simp: set_plus_def image_iff) lemma set_times_image: "S * T = (\(x, y). x * y) ` (S \ T)" - unfolding set_times_def by (fastforce simp: image_iff) + by (fastforce simp: set_times_def image_iff) lemma finite_set_plus: "finite s \ finite t \ finite (s + t)" - unfolding set_plus_image by simp + by (simp add: set_plus_image) lemma finite_set_times: "finite s \ finite t \ finite (s * t)" - unfolding set_times_image by simp + by (simp add: set_times_image) lemma set_setsum_alt: assumes fin: "finite I" diff -r ae33d1c2ab26 -r a7c5074a0251 src/HOL/Library/Stirling.thy --- a/src/HOL/Library/Stirling.thy Wed Jul 13 21:00:03 2016 +0200 +++ b/src/HOL/Library/Stirling.thy Wed Jul 13 21:30:41 2016 +0200 @@ -1,5 +1,9 @@ -(* Authors: Amine Chaieb & Florian Haftmann, TU Muenchen - with contributions by Lukas Bulwahn and Manuel Eberl*) +(* Title: HOL/Library/Stirling.thy + Author: Amine Chaieb + Author: Florian Haftmann + Author: Lukas Bulwahn + Author: Manuel Eberl +*) section \Stirling numbers of first and second kind\ @@ -10,102 +14,105 @@ subsection \Stirling numbers of the second kind\ fun Stirling :: "nat \ nat \ nat" -where - "Stirling 0 0 = 1" -| "Stirling 0 (Suc k) = 0" -| "Stirling (Suc n) 0 = 0" -| "Stirling (Suc n) (Suc k) = Suc k * Stirling n (Suc k) + Stirling n k" + where + "Stirling 0 0 = 1" + | "Stirling 0 (Suc k) = 0" + | "Stirling (Suc n) 0 = 0" + | "Stirling (Suc n) (Suc k) = Suc k * Stirling n (Suc k) + Stirling n k" -lemma Stirling_1 [simp]: - "Stirling (Suc n) (Suc 0) = 1" +lemma Stirling_1 [simp]: "Stirling (Suc n) (Suc 0) = 1" by (induct n) simp_all -lemma Stirling_less [simp]: - "n < k \ Stirling n k = 0" +lemma Stirling_less [simp]: "n < k \ Stirling n k = 0" by (induct n k rule: Stirling.induct) simp_all -lemma Stirling_same [simp]: - "Stirling n n = 1" +lemma Stirling_same [simp]: "Stirling n n = 1" by (induct n) simp_all -lemma Stirling_2_2: - "Stirling (Suc (Suc n)) (Suc (Suc 0)) = 2 ^ Suc n - 1" +lemma Stirling_2_2: "Stirling (Suc (Suc n)) (Suc (Suc 0)) = 2 ^ Suc n - 1" proof (induct n) - case 0 then show ?case by simp + case 0 + then show ?case by simp next case (Suc n) have "Stirling (Suc (Suc (Suc n))) (Suc (Suc 0)) = - 2 * Stirling (Suc (Suc n)) (Suc (Suc 0)) + Stirling (Suc (Suc n)) (Suc 0)" by simp + 2 * Stirling (Suc (Suc n)) (Suc (Suc 0)) + Stirling (Suc (Suc n)) (Suc 0)" + by simp also have "\ = 2 * (2 ^ Suc n - 1) + 1" by (simp only: Suc Stirling_1) also have "\ = 2 ^ Suc (Suc n) - 1" proof - - have "(2::nat) ^ Suc n - 1 > 0" by (induct n) simp_all - then have "2 * ((2::nat) ^ Suc n - 1) > 0" by simp - then have "2 \ 2 * ((2::nat) ^ Suc n)" by simp + have "(2::nat) ^ Suc n - 1 > 0" + by (induct n) simp_all + then have "2 * ((2::nat) ^ Suc n - 1) > 0" + by simp + then have "2 \ 2 * ((2::nat) ^ Suc n)" + by simp with add_diff_assoc2 [of 2 "2 * 2 ^ Suc n" 1] - have "2 * 2 ^ Suc n - 2 + (1::nat) = 2 * 2 ^ Suc n + 1 - 2" . - then show ?thesis by (simp add: nat_distrib) + have "2 * 2 ^ Suc n - 2 + (1::nat) = 2 * 2 ^ Suc n + 1 - 2" . + then show ?thesis + by (simp add: nat_distrib) qed finally show ?case by simp qed -lemma Stirling_2: - "Stirling (Suc n) (Suc (Suc 0)) = 2 ^ n - 1" +lemma Stirling_2: "Stirling (Suc n) (Suc (Suc 0)) = 2 ^ n - 1" using Stirling_2_2 by (cases n) simp_all + subsection \Stirling numbers of the first kind\ fun stirling :: "nat \ nat \ nat" -where - "stirling 0 0 = 1" -| "stirling 0 (Suc k) = 0" -| "stirling (Suc n) 0 = 0" -| "stirling (Suc n) (Suc k) = n * stirling n (Suc k) + stirling n k" + where + "stirling 0 0 = 1" + | "stirling 0 (Suc k) = 0" + | "stirling (Suc n) 0 = 0" + | "stirling (Suc n) (Suc k) = n * stirling n (Suc k) + stirling n k" lemma stirling_0 [simp]: "n > 0 \ stirling n 0 = 0" by (cases n) simp_all -lemma stirling_less [simp]: - "n < k \ stirling n k = 0" +lemma stirling_less [simp]: "n < k \ stirling n k = 0" by (induct n k rule: stirling.induct) simp_all -lemma stirling_same [simp]: - "stirling n n = 1" +lemma stirling_same [simp]: "stirling n n = 1" by (induct n) simp_all -lemma stirling_Suc_n_1: - "stirling (Suc n) (Suc 0) = fact n" +lemma stirling_Suc_n_1: "stirling (Suc n) (Suc 0) = fact n" by (induct n) auto -lemma stirling_Suc_n_n: - shows "stirling (Suc n) n = Suc n choose 2" -by (induct n) (auto simp add: numerals(2)) +lemma stirling_Suc_n_n: "stirling (Suc n) n = Suc n choose 2" + by (induct n) (auto simp add: numerals(2)) lemma stirling_Suc_n_2: assumes "n \ Suc 0" shows "stirling (Suc n) 2 = (\k=1..n. fact n div k)" -using assms + using assms proof (induct n) - case 0 from this show ?case by simp + case 0 + then show ?case by simp next case (Suc n) show ?case proof (cases n) - case 0 from this show ?thesis by (simp add: numerals(2)) + case 0 + then show ?thesis + by (simp add: numerals(2)) next case Suc - from this have geq1: "Suc 0 \ n" by simp + then have geq1: "Suc 0 \ n" + by simp have "stirling (Suc (Suc n)) 2 = Suc n * stirling (Suc n) 2 + stirling (Suc n) (Suc 0)" by (simp only: stirling.simps(4)[of "Suc n"] numerals(2)) - also have "... = Suc n * (\k=1..n. fact n div k) + fact n" + also have "\ = Suc n * (\k=1..n. fact n div k) + fact n" using Suc.hyps[OF geq1] by (simp only: stirling_Suc_n_1 of_nat_fact of_nat_add of_nat_mult) - also have "... = Suc n * (\k=1..n. fact n div k) + Suc n * fact n div Suc n" + also have "\ = Suc n * (\k=1..n. fact n div k) + Suc n * fact n div Suc n" by (metis nat.distinct(1) nonzero_mult_divide_cancel_left) - also have "... = (\k=1..n. fact (Suc n) div k) + fact (Suc n) div Suc n" + also have "\ = (\k=1..n. fact (Suc n) div k) + fact (Suc n) div Suc n" by (simp add: setsum_right_distrib div_mult_swap dvd_fact) - also have "... = (\k=1..Suc n. fact (Suc n) div k)" by simp + also have "\ = (\k=1..Suc n. fact (Suc n) div k)" + by simp finally show ?thesis . qed qed @@ -113,52 +120,60 @@ lemma of_nat_stirling_Suc_n_2: assumes "n \ Suc 0" shows "(of_nat (stirling (Suc n) 2)::'a::field_char_0) = fact n * (\k=1..n. (1 / of_nat k))" -using assms + using assms proof (induct n) - case 0 from this show ?case by simp + case 0 + then show ?case by simp next case (Suc n) show ?case proof (cases n) - case 0 from this show ?thesis by (auto simp add: numerals(2)) + case 0 + then show ?thesis + by (auto simp add: numerals(2)) next case Suc - from this have geq1: "Suc 0 \ n" by simp + then have geq1: "Suc 0 \ n" + by simp have "(of_nat (stirling (Suc (Suc n)) 2)::'a) = - of_nat (Suc n * stirling (Suc n) 2 + stirling (Suc n) (Suc 0))" + of_nat (Suc n * stirling (Suc n) 2 + stirling (Suc n) (Suc 0))" by (simp only: stirling.simps(4)[of "Suc n"] numerals(2)) - also have "... = of_nat (Suc n) * (fact n * (\k = 1..n. 1 / of_nat k)) + fact n" + also have "\ = of_nat (Suc n) * (fact n * (\k = 1..n. 1 / of_nat k)) + fact n" using Suc.hyps[OF geq1] by (simp only: stirling_Suc_n_1 of_nat_fact of_nat_add of_nat_mult) - also have "... = fact (Suc n) * (\k = 1..n. 1 / of_nat k) + fact (Suc n) * (1 / of_nat (Suc n))" + also have "\ = fact (Suc n) * (\k = 1..n. 1 / of_nat k) + fact (Suc n) * (1 / of_nat (Suc n))" using of_nat_neq_0 by auto - also have "... = fact (Suc n) * (\k = 1..Suc n. 1 / of_nat k)" + also have "\ = fact (Suc n) * (\k = 1..Suc n. 1 / of_nat k)" by (simp add: distrib_left) finally show ?thesis . qed qed -lemma setsum_stirling: - "(\k\n. stirling n k) = fact n" +lemma setsum_stirling: "(\k\n. stirling n k) = fact n" proof (induct n) case 0 - from this show ?case by simp + then show ?case by simp next case (Suc n) have "(\k\Suc n. stirling (Suc n) k) = stirling (Suc n) 0 + (\k\n. stirling (Suc n) (Suc k))" by (simp only: setsum_atMost_Suc_shift) - also have "\ = (\k\n. stirling (Suc n) (Suc k))" by simp - also have "\ = (\k\n. n * stirling n (Suc k) + stirling n k)" by simp + also have "\ = (\k\n. stirling (Suc n) (Suc k))" + by simp + also have "\ = (\k\n. n * stirling n (Suc k) + stirling n k)" + by simp also have "\ = n * (\k\n. stirling n (Suc k)) + (\k\n. stirling n k)" by (simp add: setsum.distrib setsum_right_distrib) also have "\ = n * fact n + fact n" proof - have "n * (\k\n. stirling n (Suc k)) = n * ((\k\Suc n. stirling n k) - stirling n 0)" by (metis add_diff_cancel_left' setsum_atMost_Suc_shift) - also have "\ = n * (\k\n. stirling n k)" by (cases n) simp+ - also have "\ = n * fact n" using Suc.hyps by simp + also have "\ = n * (\k\n. stirling n k)" + by (cases n) simp_all + also have "\ = n * fact n" + using Suc.hyps by simp finally have "n * (\k\n. stirling n (Suc k)) = n * fact n" . - moreover have "(\k\n. stirling n k) = fact n" using Suc.hyps . + moreover have "(\k\n. stirling n k) = fact n" + using Suc.hyps . ultimately show ?thesis by simp qed also have "\ = fact (Suc n)" by simp @@ -166,26 +181,29 @@ qed lemma stirling_pochhammer: - "(\k\n. of_nat (stirling n k) * x ^ k) = (pochhammer x n :: 'a :: comm_semiring_1)" -proof (induction n) + "(\k\n. of_nat (stirling n k) * x ^ k) = (pochhammer x n :: 'a::comm_semiring_1)" +proof (induct n) + case 0 + then show ?case by simp +next case (Suc n) have "of_nat (n * stirling n 0) = (0 :: 'a)" by (cases n) simp_all - hence "(\k\Suc n. of_nat (stirling (Suc n) k) * x ^ k) = - (of_nat (n * stirling n 0) * x ^ 0 + - (\i\n. of_nat (n * stirling n (Suc i)) * (x ^ Suc i))) + - (\i\n. of_nat (stirling n i) * (x ^ Suc i))" + then have "(\k\Suc n. of_nat (stirling (Suc n) k) * x ^ k) = + (of_nat (n * stirling n 0) * x ^ 0 + + (\i\n. of_nat (n * stirling n (Suc i)) * (x ^ Suc i))) + + (\i\n. of_nat (stirling n i) * (x ^ Suc i))" by (subst setsum_atMost_Suc_shift) (simp add: setsum.distrib ring_distribs) also have "\ = pochhammer x (Suc n)" by (subst setsum_atMost_Suc_shift [symmetric]) - (simp add: algebra_simps setsum.distrib setsum_right_distrib pochhammer_Suc Suc [symmetric]) + (simp add: algebra_simps setsum.distrib setsum_right_distrib pochhammer_Suc Suc [symmetric]) finally show ?case . -qed simp_all +qed text \A row of the Stirling number triangle\ -definition stirling_row :: "nat \ nat list" where - "stirling_row n = [stirling n k. k \ [0.. nat list" + where "stirling_row n = [stirling n k. k \ [0.. n \ stirling_row n ! k = stirling n k" by (simp add: stirling_row_def del: upt_Suc) @@ -200,82 +218,109 @@ lemma list_ext: assumes "length xs = length ys" assumes "\i. i < length xs \ xs ! i = ys ! i" - shows "xs = ys" -using assms + shows "xs = ys" + using assms proof (induction rule: list_induct2) + case Nil + then show ?case by simp +next case (Cons x xs y ys) - from Cons.prems[of 0] have "x = y" by simp - moreover from Cons.prems[of "Suc i" for i] have "xs = ys" by (intro Cons.IH) simp + from Cons.prems[of 0] have "x = y" + by simp + moreover from Cons.prems[of "Suc i" for i] have "xs = ys" + by (intro Cons.IH) simp ultimately show ?case by simp -qed simp_all +qed subsubsection \Efficient code\ text \ - Naively using the defining equations of the Stirling numbers of the first kind to - compute them leads to exponential run time due to repeated computations. - We can use memoisation to compute them row by row without repeating computations, at - the cost of computing a few unneeded values. + Naively using the defining equations of the Stirling numbers of the first + kind to compute them leads to exponential run time due to repeated + computations. We can use memoisation to compute them row by row without + repeating computations, at the cost of computing a few unneeded values. As a bonus, this is very efficient for applications where an entire row of - Stirling numbers is needed.. + Stirling numbers is needed. \ -definition zip_with_prev :: "('a \ 'a \ 'b) \ 'a \ 'a list \ 'b list" where - "zip_with_prev f x xs = map (\(x,y). f x y) (zip (x # xs) xs)" +definition zip_with_prev :: "('a \ 'a \ 'b) \ 'a \ 'a list \ 'b list" + where "zip_with_prev f x xs = map (\(x,y). f x y) (zip (x # xs) xs)" lemma zip_with_prev_altdef: "zip_with_prev f x xs = - (if xs = [] then [] else f x (hd xs) # [f (xs!i) (xs!(i+1)). i \ [0.. [0..i. f (xs ! i) (xs ! (i + 1))) [0..a b. a + n * b) y xs @ [1]" - by (induction xs arbitrary: y) (simp_all add: zip_with_prev_def) + by (induct xs arbitrary: y) (simp_all add: zip_with_prev_def) lemma stirling_row_code [code]: "stirling_row 0 = [1]" "stirling_row (Suc n) = stirling_row_aux n 0 (stirling_row n)" -proof - +proof goal_cases + case 1 + show ?case by (simp add: stirling_row_def) +next + case 2 have "stirling_row (Suc n) = - 0 # [stirling_row n ! i + stirling_row n ! (i+1) * n. i \ [0.. [0.. Suc n" by simp - then consider "i = 0" | j where "i > 0" "i \ n" | "i = Suc n" by linarith - thus ?case + from nth have "i \ Suc n" + by simp + then consider "i = 0 \ i = Suc n" | "i > 0" "i \ n" + by linarith + then show ?case proof cases - assume i: "i > 0" "i \ n" - from this show ?thesis by (cases i) (simp_all add: nth_append nth_stirling_row) - qed (simp_all add: nth_stirling_row nth_append) - qed simp + case 1 + then show ?thesis + by (auto simp: nth_stirling_row nth_append) + next + case 2 + then show ?thesis + by (cases i) (simp_all add: nth_append nth_stirling_row) + qed + next + case length + then show ?case by simp + qed also have "0 # [stirling_row n ! i + stirling_row n ! (i+1) * n. i \ [0..a b. a + n * b) 0 (stirling_row n) @ [1]" + zip_with_prev (\a b. a + n * b) 0 (stirling_row n) @ [1]" by (cases n) (auto simp add: zip_with_prev_altdef stirling_row_def hd_map simp del: upt_Suc) also have "\ = stirling_row_aux n 0 (stirling_row n)" by (simp add: stirling_row_aux_correct) - finally show "stirling_row (Suc n) = stirling_row_aux n 0 (stirling_row n)" . -qed (simp add: stirling_row_def) + finally show ?case . +qed lemma stirling_code [code]: - "stirling n k = (if k = 0 then if n = 0 then 1 else 0 - else if k > n then 0 else if k = n then 1 - else stirling_row n ! k)" + "stirling n k = + (if k = 0 then (if n = 0 then 1 else 0) + else if k > n then 0 + else if k = n then 1 + else stirling_row n ! k)" by (simp add: nth_stirling_row) end diff -r ae33d1c2ab26 -r a7c5074a0251 src/Pure/Isar/keyword.scala --- a/src/Pure/Isar/keyword.scala Wed Jul 13 21:00:03 2016 +0200 +++ b/src/Pure/Isar/keyword.scala Wed Jul 13 21:30:41 2016 +0200 @@ -66,6 +66,8 @@ val theory_body = Set(THY_LOAD, THY_DECL, THY_DECL_BLOCK, THY_GOAL) + val prf_script = Set(PRF_SCRIPT) + val proof = Set(QED, QED_SCRIPT, QED_BLOCK, QED_GLOBAL, PRF_GOAL, PRF_BLOCK, NEXT_BLOCK, PRF_OPEN, PRF_CLOSE, PRF_CHAIN, PRF_DECL, PRF_ASM, PRF_ASM_GOAL, PRF_SCRIPT, PRF_SCRIPT_GOAL, diff -r ae33d1c2ab26 -r a7c5074a0251 src/Pure/Isar/outer_syntax.scala --- a/src/Pure/Isar/outer_syntax.scala Wed Jul 13 21:00:03 2016 +0200 +++ b/src/Pure/Isar/outer_syntax.scala Wed Jul 13 21:30:41 2016 +0200 @@ -174,9 +174,9 @@ if (keywords.is_command(tok, Keyword.theory_goal)) (2, 1) else if (keywords.is_command(tok, Keyword.theory)) (1, 0) else if (keywords.is_command(tok, Keyword.proof_open)) (y + 2, y + 1) - else if (keywords.is_command(tok, Keyword.PRF_BLOCK == _)) (y + 2, y + 1) - else if (keywords.is_command(tok, Keyword.QED_BLOCK == _)) (y - 1, y - 2) - else if (keywords.is_command(tok, Keyword.PRF_CLOSE == _)) (y, y - 1) + else if (keywords.is_command(tok, Set(Keyword.PRF_BLOCK))) (y + 2, y + 1) + else if (keywords.is_command(tok, Set(Keyword.QED_BLOCK))) (y - 1, y - 2) + else if (keywords.is_command(tok, Set(Keyword.PRF_CLOSE))) (y, y - 1) else if (keywords.is_command(tok, Keyword.proof_close)) (y + 1, y - 1) else if (keywords.is_command(tok, Keyword.qed_global)) (1, 0) else (x, y) diff -r ae33d1c2ab26 -r a7c5074a0251 src/Pure/Isar/proof.ML --- a/src/Pure/Isar/proof.ML Wed Jul 13 21:00:03 2016 +0200 +++ b/src/Pure/Isar/proof.ML Wed Jul 13 21:30:41 2016 +0200 @@ -42,7 +42,6 @@ val raw_goal: state -> {context: context, facts: thm list, goal: thm} val goal: state -> {context: context, facts: thm list, goal: thm} val simple_goal: state -> {context: context, goal: thm} - val status_markup: state -> Markup.T val let_bind: (term list * term) list -> state -> state val let_bind_cmd: (string list * string) list -> state -> state val write: Syntax.mode -> (term * mixfix) list -> state -> state @@ -561,11 +560,6 @@ val (ctxt, (_, goal)) = get_goal (refine_insert using state); in {context = ctxt, goal = goal} end; -fun status_markup state = - (case try goal state of - SOME {goal, ...} => Markup.proof_state (Thm.nprems_of goal) - | NONE => Markup.empty); - fun method_error kind pos state = Seq.single (Proof_Display.method_error kind pos (raw_goal state)); diff -r ae33d1c2ab26 -r a7c5074a0251 src/Pure/Isar/token.scala --- a/src/Pure/Isar/token.scala Wed Jul 13 21:00:03 2016 +0200 +++ b/src/Pure/Isar/token.scala Wed Jul 13 21:30:41 2016 +0200 @@ -263,6 +263,7 @@ def is_begin: Boolean = is_keyword("begin") def is_end: Boolean = is_command("end") + def is_begin_or_command: Boolean = is_begin || is_command def content: String = if (kind == Token.Kind.STRING) Scan.Parsers.quoted_content("\"", source) diff -r ae33d1c2ab26 -r a7c5074a0251 src/Pure/PIDE/command.ML --- a/src/Pure/PIDE/command.ML Wed Jul 13 21:00:03 2016 +0200 +++ b/src/Pure/PIDE/command.ML Wed Jul 13 21:30:41 2016 +0200 @@ -209,17 +209,28 @@ fun status tr m = Toplevel.setmp_thread_position tr (fn () => Output.status (Markup.markup_only m)) (); -fun proof_status tr st = +fun command_indent tr st = (case try Toplevel.proof_of st of - SOME prf => status tr (Proof.status_markup prf) + SOME prf => + let val keywords = Thy_Header.get_keywords (Proof.theory_of prf) in + if Keyword.command_kind keywords (Toplevel.name_of tr) = SOME Keyword.prf_script then + (case try Proof.goal prf of + SOME {goal, ...} => + let val n = Thm.nprems_of goal + in if n > 1 then report tr (Markup.command_indent (n - 1)) else () end + | NONE => ()) + else () + end | NONE => ()); + fun eval_state keywords span tr ({state, ...}: eval_state) = let val _ = Thread_Attributes.expose_interrupt (); val st = reset_state keywords tr state; + val _ = command_indent tr st; val _ = status tr Markup.running; val (errs1, result) = run keywords true tr st; val errs2 = (case result of NONE => [] | SOME st' => check_cmts span tr st'); @@ -235,7 +246,6 @@ in {failed = true, command = tr, state = st} end | SOME st' => let - val _ = proof_status tr st'; val _ = status tr Markup.finished; in {failed = false, command = tr, state = st'} end) end; diff -r ae33d1c2ab26 -r a7c5074a0251 src/Pure/PIDE/markup.ML --- a/src/Pure/PIDE/markup.ML Wed Jul 13 21:00:03 2016 +0200 +++ b/src/Pure/PIDE/markup.ML Wed Jul 13 21:30:41 2016 +0200 @@ -155,8 +155,7 @@ val parse_command_timing_properties: Properties.T -> ({file: string, offset: int, name: string} * Time.time) option val timingN: string val timing: {elapsed: Time.time, cpu: Time.time, gc: Time.time} -> T - val subgoalsN: string - val proof_stateN: string val proof_state: int -> T + val command_indentN: string val command_indent: int -> T val goalN: string val goal: T val subgoalN: string val subgoal: string -> T val taskN: string @@ -576,10 +575,12 @@ | _ => NONE); -(* toplevel *) +(* indentation *) -val subgoalsN = "subgoals"; -val (proof_stateN, proof_state) = markup_int "proof_state" subgoalsN; +val (command_indentN, command_indent) = markup_int "command_indent" indentN; + + +(* goals *) val (goalN, goal) = markup_elem "goal"; val (subgoalN, subgoal) = markup_string "subgoal" nameN; diff -r ae33d1c2ab26 -r a7c5074a0251 src/Pure/PIDE/markup.scala --- a/src/Pure/PIDE/markup.scala Wed Jul 13 21:00:03 2016 +0200 +++ b/src/Pure/PIDE/markup.scala Wed Jul 13 21:30:41 2016 +0200 @@ -372,10 +372,17 @@ val COMMAND_TIMING = "command_timing" - /* toplevel */ + /* command indentation */ - val SUBGOALS = "subgoals" - val PROOF_STATE = "proof_state" + object Command_Indent + { + val name = "command_indent" + def unapply(markup: Markup): Option[Int] = + if (markup.name == name) Indent.unapply(markup.properties) else None + } + + + /* goals */ val GOAL = "goal" val SUBGOAL = "subgoal" diff -r ae33d1c2ab26 -r a7c5074a0251 src/Tools/jEdit/etc/options --- a/src/Tools/jEdit/etc/options Wed Jul 13 21:00:03 2016 +0200 +++ b/src/Tools/jEdit/etc/options Wed Jul 13 21:30:41 2016 +0200 @@ -45,6 +45,12 @@ public option jedit_indent_newline : bool = true -- "indentation of Isabelle keywords on ENTER (action isabelle.newline)" +public option jedit_indent_script : bool = true + -- "indent unstructured proof script ('apply' etc.) via number of subgoals" + +public option jedit_indent_script_limit : int = 20 + -- "maximum indentation of unstructured proof script ('apply' etc.)" + section "Completion" diff -r ae33d1c2ab26 -r a7c5074a0251 src/Tools/jEdit/src/isabelle.scala --- a/src/Tools/jEdit/src/isabelle.scala Wed Jul 13 21:00:03 2016 +0200 +++ b/src/Tools/jEdit/src/isabelle.scala Wed Jul 13 21:30:41 2016 +0200 @@ -265,6 +265,8 @@ { Isabelle.buffer_syntax(buffer) match { case Some(syntax) if buffer.isInstanceOf[Buffer] => + val keywords = syntax.keywords + val caret = text_area.getCaretPosition val line = text_area.getCaretLine val line_range = JEdit_Lib.line_range(buffer, line) @@ -282,9 +284,12 @@ val (tokens1, context1) = line_content(line_range.start, caret, context0) val (tokens2, _) = line_content(caret, line_range.stop, context1) - if (tokens1.nonEmpty && tokens1.head.is_command) buffer.indentLine(line, true) + if (tokens1.nonEmpty && + (tokens1.head.is_begin_or_command || keywords.is_quasi_command(tokens1.head))) + buffer.indentLine(line, true) - if (tokens2.isEmpty || tokens2.head.is_command) + if (tokens2.isEmpty || + tokens2.head.is_begin_or_command || keywords.is_quasi_command(tokens2.head)) JEdit_Lib.buffer_edit(buffer) { text_area.setSelectedText("\n") if (!buffer.indentLine(line + 1, true)) text_area.goToStartOfWhiteSpace(false) diff -r ae33d1c2ab26 -r a7c5074a0251 src/Tools/jEdit/src/rendering.scala --- a/src/Tools/jEdit/src/rendering.scala Wed Jul 13 21:00:03 2016 +0200 +++ b/src/Tools/jEdit/src/rendering.scala Wed Jul 13 21:30:41 2016 +0200 @@ -137,6 +137,9 @@ /* markup elements */ + private val indentation_elements = + Markup.Elements(Markup.Command_Indent.name) + private val semantic_completion_elements = Markup.Elements(Markup.COMPLETION, Markup.NO_COMPLETION) @@ -295,6 +298,16 @@ val markdown_item_color4 = color_value("markdown_item_color4") + /* indentation */ + + def indentation(range: Text.Range): Int = + snapshot.select(range, Rendering.indentation_elements, _ => + { + case Text.Info(_, XML.Elem(Markup.Command_Indent(i), _)) => Some(i) + case _ => None + }).headOption.map(_.info).getOrElse(0) + + /* completion */ def semantic_completion(completed_range: Option[Text.Range], range: Text.Range) diff -r ae33d1c2ab26 -r a7c5074a0251 src/Tools/jEdit/src/text_structure.scala --- a/src/Tools/jEdit/src/text_structure.scala Wed Jul 13 21:00:03 2016 +0200 +++ b/src/Tools/jEdit/src/text_structure.scala Wed Jul 13 21:30:41 2016 +0200 @@ -26,13 +26,13 @@ def iterator(line: Int, lim: Int = limit): Iterator[Text.Info[Token]] = { val it = Token_Markup.line_token_iterator(syntax, buffer, line, line + lim) - if (comments) it.filterNot(_.info.is_space) else it.filter(_.info.is_proper) + if (comments) it.filterNot(_.info.is_space) else it.filterNot(_.info.is_improper) } def reverse_iterator(line: Int, lim: Int = limit): Iterator[Text.Info[Token]] = { val it = Token_Markup.line_token_reverse_iterator(syntax, buffer, line, line - lim) - if (comments) it.filterNot(_.info.is_space) else it.filter(_.info.is_proper) + if (comments) it.filterNot(_.info.is_space) else it.filterNot(_.info.is_improper) } } @@ -52,31 +52,52 @@ val keywords = syntax.keywords val nav = new Navigator(syntax, buffer.asInstanceOf[Buffer], true) - def head_token(line: Int): Option[Token] = - nav.iterator(line, 1).toStream.headOption.map(_.info) - - def head_is_quasi_command(line: Int): Boolean = - head_token(line) match { - case None => false - case Some(tok) => keywords.is_quasi_command(tok) - } - - def prev_command: Option[Token] = - nav.reverse_iterator(prev_line, 1). - collectFirst({ case Text.Info(_, tok) if tok.is_command => tok }) - - def prev_span: Iterator[Token] = - nav.reverse_iterator(prev_line).map(_.info).takeWhile(tok => !tok.is_command) - - def prev_line_span: Iterator[Token] = - nav.reverse_iterator(prev_line, 1).map(_.info).takeWhile(tok => !tok.is_command) + val indent_size = buffer.getIndentSize def line_indent(line: Int): Int = if (line < 0 || line >= buffer.getLineCount) 0 else buffer.getCurrentIndentForLine(line, null) - val indent_size = buffer.getIndentSize + def line_head(line: Int): Option[Text.Info[Token]] = + nav.iterator(line, 1).toStream.headOption + + def head_is_quasi_command(line: Int): Boolean = + line_head(line) match { + case None => false + case Some(Text.Info(_, tok)) => keywords.is_quasi_command(tok) + } + + def prev_line_command: Option[Token] = + nav.reverse_iterator(prev_line, 1). + collectFirst({ case Text.Info(_, tok) if tok.is_begin_or_command => tok }) + + def prev_line_span: Iterator[Token] = + nav.reverse_iterator(prev_line, 1).map(_.info).takeWhile(tok => !tok.is_begin_or_command) + + def prev_span: Iterator[Token] = + nav.reverse_iterator(prev_line).map(_.info).takeWhile(tok => !tok.is_begin_or_command) + + + val script_indent: Text.Info[Token] => Int = + { + val opt_rendering: Option[Rendering] = + if (PIDE.options.value.bool("jedit_indent_script")) + GUI_Thread.now { + (for { + text_area <- JEdit_Lib.jedit_text_areas(buffer) + doc_view <- PIDE.document_view(text_area) + } yield doc_view.get_rendering).toStream.headOption + } + else None + val limit = PIDE.options.value.int("jedit_indent_script_limit") + (info: Text.Info[Token]) => + opt_rendering match { + case Some(rendering) if keywords.is_command(info.info, Keyword.prf_script) => + (rendering.indentation(info.range) min limit) max 0 + case _ => 0 + } + } def indent_indent(tok: Token): Int = if (keywords.is_command(tok, keyword_open)) indent_size @@ -84,9 +105,24 @@ else 0 def indent_offset(tok: Token): Int = - if (keywords.is_command(tok, Keyword.proof_enclose) || tok.is_begin) indent_size + if (keywords.is_command(tok, Keyword.proof_enclose)) indent_size else 0 + def indent_structure: Int = + nav.reverse_iterator(current_line - 1).scanLeft((0, false))( + { case ((ind, _), Text.Info(range, tok)) => + val ind1 = ind + indent_indent(tok) + if (tok.is_begin_or_command && !keywords.is_command(tok, Keyword.prf_script)) { + val line = buffer.getLineOfOffset(range.start) + line_head(line) match { + case Some(info) if info.info == tok => + (ind1 + indent_offset(tok) + line_indent(line), true) + case _ => (ind1, false) + } + } + else (ind1, false) + }).collectFirst({ case (i, true) => i }).getOrElse(0) + def indent_brackets: Int = (0 /: prev_line_span)( { case (i, tok) => @@ -98,35 +134,20 @@ if (prev_span.exists(keywords.is_quasi_command(_))) indent_size else 0 - def indent_structure: Int = - nav.reverse_iterator(current_line - 1).scanLeft((0, false))( - { case ((ind, _), Text.Info(range, tok)) => - val ind1 = ind + indent_indent(tok) - if (tok.is_command) { - val line = buffer.getLineOfOffset(range.start) - if (head_token(line) == Some(tok)) - (ind1 + indent_offset(tok) + line_indent(line), true) - else (ind1, false) - } - else (ind1, false) - }).collectFirst({ case (i, true) => i }).getOrElse(0) - - def nesting(it: Iterator[Token], open: Token => Boolean, close: Token => Boolean): Int = - (0 /: it)({ case (d, tok) => if (open(tok)) d + 1 else if (close(tok)) d - 1 else d }) - - def indent_begin: Int = - (nesting(nav.iterator(current_line - 1, 1).map(_.info), _.is_begin, _.is_end) max 0) * - indent_size - val indent = - head_token(current_line) match { + line_head(current_line) match { case None => indent_structure + indent_brackets + indent_extra - case Some(tok) => - if (keywords.is_before_command(tok) || - keywords.is_command(tok, Keyword.theory)) indent_begin - else if (tok.is_command) indent_structure + indent_begin - indent_offset(tok) + case Some(info @ Text.Info(range, tok)) => + if (tok.is_begin || + keywords.is_before_command(tok) || + keywords.is_command(tok, Keyword.theory)) 0 + else if (keywords.is_command(tok, Keyword.proof_enclose)) + indent_structure + script_indent(info) - indent_offset(tok) + else if (keywords.is_command(tok, Keyword.proof)) + (indent_structure + script_indent(info) - indent_offset(tok)) max indent_size + else if (tok.is_command) indent_structure - indent_offset(tok) else { - prev_command match { + prev_line_command match { case None => val extra = (keywords.is_quasi_command(tok), head_is_quasi_command(prev_line)) match { @@ -134,10 +155,10 @@ case (true, false) => - indent_extra case (false, true) => indent_extra } - line_indent(prev_line) - indent_offset(tok) + indent_brackets + extra + line_indent(prev_line) + indent_brackets + extra - indent_offset(tok) case Some(prev_tok) => - indent_structure - indent_offset(tok) - indent_offset(prev_tok) + - indent_brackets - indent_indent(prev_tok) + indent_size + indent_structure + indent_brackets + indent_size - indent_offset(tok) - + indent_offset(prev_tok) - indent_indent(prev_tok) } } }