# HG changeset patch # User wenzelm # Date 1375883038 -7200 # Node ID 3695ce0f9f96001f8b7d57739885b7ac3fc23a77 # Parent 34c47bc771f28ee1d7dff4b0d01e3cd1988ba137# Parent 73e32ed924b3eeec85966e00deeb8aaad0593d92 merged diff -r 34c47bc771f2 -r 3695ce0f9f96 CONTRIBUTORS --- a/CONTRIBUTORS Wed Aug 07 15:40:59 2013 +0200 +++ b/CONTRIBUTORS Wed Aug 07 15:43:58 2013 +0200 @@ -6,6 +6,10 @@ Contributions to this Isabelle version -------------------------------------- +* Summer 2013: Christian Sternagel, JAIST + Improved support for adhoc overloading of constants, including + documentation and examples. + * May 2013: Florian Haftmann, TUM Ephemeral interpretation in local theories. diff -r 34c47bc771f2 -r 3695ce0f9f96 NEWS --- a/NEWS Wed Aug 07 15:40:59 2013 +0200 +++ b/NEWS Wed Aug 07 15:43:58 2013 +0200 @@ -106,6 +106,9 @@ *** HOL *** +* Improved support for adhoc overloading of constants (see also +isar-ref manual and ~~/src/HOL/ex/Adhoc_Overloading_Examples.thy). + * Attibute 'code': 'code' now declares concrete and abstract code equations uniformly. Use explicit 'code equation' and 'code abstract' to distinguish both when desired. diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Doc/IsarRef/HOL_Specific.thy --- a/src/Doc/IsarRef/HOL_Specific.thy Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Doc/IsarRef/HOL_Specific.thy Wed Aug 07 15:43:58 2013 +0200 @@ -1,5 +1,5 @@ theory HOL_Specific -imports Base Main "~~/src/HOL/Library/Old_Recdef" +imports Base Main "~~/src/HOL/Library/Old_Recdef" "~~/src/Tools/Adhoc_Overloading" begin chapter {* Higher-Order Logic *} @@ -588,7 +588,7 @@ There are no fixed syntactic restrictions on the body of the function, but the induced functional must be provably monotonic - wrt.\ the underlying order. The monotonicitity proof is performed + wrt.\ the underlying order. The monotonicity proof is performed internally, and the definition is rejected when it fails. The proof can be influenced by declaring hints using the @{attribute (HOL) partial_function_mono} attribute. @@ -622,7 +622,7 @@ @{const "partial_function_definitions"} appropriately. \item @{attribute (HOL) partial_function_mono} declares rules for - use in the internal monononicity proofs of partial function + use in the internal monotonicity proofs of partial function definitions. \end{description} @@ -1288,7 +1288,7 @@ "morphisms"} specification provides alternative names. @{command (HOL) "quotient_type"} requires the user to prove that the relation is an equivalence relation (predicate @{text equivp}), unless the - user specifies explicitely @{text partial} in which case the + user specifies explicitly @{text partial} in which case the obligation is @{text part_equivp}. A quotient defined with @{text partial} is weaker in the sense that less things can be proved automatically. @@ -1318,7 +1318,7 @@ and @{method (HOL) "descending"} continues in the same way as @{method (HOL) "lifting"} does. @{method (HOL) "descending"} tries to solve the arising regularization, injection and cleaning - subgoals with the analoguous method @{method (HOL) + subgoals with the analogous method @{method (HOL) "descending_setup"} which leaves the four unsolved subgoals. \item @{method (HOL) "partiality_descending"} finds the regularized @@ -1416,6 +1416,46 @@ problem, one should use @{command (HOL) "ax_specification"}. *} +section {* Adhoc overloading of constants *} + +text {* + \begin{tabular}{rcll} + @{command_def "adhoc_overloading"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "no_adhoc_overloading"} & : & @{text "local_theory \ local_theory"} \\ + @{attribute_def "show_variants"} & : & @{text "attribute"} & default @{text false} \\ + \end{tabular} + + \medskip + + Adhoc overloading allows to overload a constant depending on + its type. Typically this involves the introduction of an + uninterpreted constant (used for input and output) and the addition + of some variants (used internally). For examples see + @{file "~~/src/HOL/ex/Adhoc_Overloading_Examples.thy"} and + @{file "~~/src/HOL/Library/Monad_Syntax.thy"}. + + @{rail " + (@@{command adhoc_overloading} | @@{command no_adhoc_overloading}) + (@{syntax nameref} (@{syntax term} + ) + @'and') + "} + + \begin{description} + + \item @{command "adhoc_overloading"}~@{text "c v\<^isub>1 ... v\<^isub>n"} + associates variants with an existing constant. + + \item @{command "no_adhoc_overloading"} is similar to + @{command "adhoc_overloading"}, but removes the specified variants + from the present context. + + \item @{attribute "show_variants"} controls printing of variants + of overloaded constants. If enabled, the internally used variants + are printed instead of their respective overloaded constants. This + is occasionally useful to check whether the system agrees with a + user's expectations about derived variants. + + \end{description} +*} chapter {* Proof tools *} @@ -1553,7 +1593,7 @@ equations in the code generator. The option @{text "no_code"} of the command @{command (HOL) "setup_lifting"} can turn off that behavior and causes that code certificate theorems generated by - @{command (HOL) "lift_definition"} are not registred as abstract + @{command (HOL) "lift_definition"} are not registered as abstract code equations. \item @{command (HOL) "lift_definition"} @{text "f :: \ is t"} @@ -1607,7 +1647,7 @@ files. \item @{attribute (HOL) reflexivity_rule} registers a theorem that shows - that a relator respects reflexivity and left-totality. For exampless + that a relator respects reflexivity and left-totality. For examples see @{file "~~/src/HOL/Library/Quotient_List.thy"} or other Quotient_*.thy The property is used in generation of abstraction function equations. @@ -2009,7 +2049,7 @@ counterexamples using a series of assignments for its free variables; by default the first subgoal is tested, an other can be selected explicitly using an optional goal index. Assignments can - be chosen exhausting the search space upto a given size, or using a + be chosen exhausting the search space up to a given size, or using a fixed number of random assignments in the search space, or exploring the search space symbolically using narrowing. By default, quickcheck uses exhaustive testing. A number of configuration @@ -2373,12 +2413,12 @@ internally. Constants may be specified by giving them literally, referring to - all executable contants within a certain theory by giving @{text + all executable constants within a certain theory by giving @{text "name.*"}, or referring to \emph{all} executable constants currently available by giving @{text "*"}. By default, for each involved theory one corresponding name space - module is generated. Alternativly, a module name may be specified + module is generated. Alternatively, a module name may be specified after the @{keyword "module_name"} keyword; then \emph{all} code is placed in this module. diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Doc/IsarRef/document/root.tex --- a/src/Doc/IsarRef/document/root.tex Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Doc/IsarRef/document/root.tex Wed Aug 07 15:43:58 2013 +0200 @@ -40,7 +40,8 @@ Lars Noschinski, \\ David von Oheimb, Larry Paulson, - Sebastian Skalberg + Sebastian Skalberg, \\ + Christian Sternagel } \makeindex diff -r 34c47bc771f2 -r 3695ce0f9f96 src/HOL/Library/Formal_Power_Series.thy --- a/src/HOL/Library/Formal_Power_Series.thy Wed Aug 07 15:40:59 2013 +0200 +++ b/src/HOL/Library/Formal_Power_Series.thy Wed Aug 07 15:43:58 2013 +0200 @@ -101,9 +101,9 @@ lemma fps_mult_nth: "(f * g) $ n = (\i=0..n. f$i * g$(n - i))" unfolding fps_times_def by simp -declare atLeastAtMost_iff[presburger] -declare Bex_def[presburger] -declare Ball_def[presburger] +declare atLeastAtMost_iff [presburger] +declare Bex_def [presburger] +declare Ball_def [presburger] lemma mult_delta_left: fixes x y :: "'a::mult_zero" @@ -117,6 +117,7 @@ lemma cond_value_iff: "f (if b then x else y) = (if b then f x else f y)" by auto + lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)" by auto @@ -125,13 +126,15 @@ instance fps :: (semigroup_add) semigroup_add proof - fix a b c :: "'a fps" show "a + b + c = a + (b + c)" + fix a b c :: "'a fps" + show "a + b + c = a + (b + c)" by (simp add: fps_ext add_assoc) qed instance fps :: (ab_semigroup_add) ab_semigroup_add proof - fix a b :: "'a fps" show "a + b = b + a" + fix a b :: "'a fps" + show "a + b = b + a" by (simp add: fps_ext add_commute) qed @@ -140,11 +143,12 @@ shows "(\j=0..k. \i=0..j. f i (j - i) (n - j)) = (\j=0..k. \i=0..k - j. f j i (n - j - i))" proof (induct k) - case 0 show ?case by simp + case 0 + show ?case by simp next - case (Suc k) thus ?case - by (simp add: Suc_diff_le setsum_addf add_assoc - cong: strong_setsum_cong) + case (Suc k) + then show ?case + by (simp add: Suc_diff_le setsum_addf add_assoc cong: strong_setsum_cong) qed instance fps :: (semiring_0) semigroup_mult @@ -156,9 +160,8 @@ have "(\j=0..n. \i=0..j. a$i * b$(j - i) * c$(n - j)) = (\j=0..n. \i=0..n - j. a$j * b$i * c$(n - j - i))" by (rule fps_mult_assoc_lemma) - thus "((a * b) * c) $ n = (a * (b * c)) $ n" - by (simp add: fps_mult_nth setsum_right_distrib - setsum_left_distrib mult_assoc) + then show "((a * b) * c) $ n = (a * (b * c)) $ n" + by (simp add: fps_mult_nth setsum_right_distrib setsum_left_distrib mult_assoc) qed qed @@ -171,9 +174,10 @@ show "{0..n} = (\i. n - i) ` {0..n}" by (auto, rule_tac x="n - x" in image_eqI, simp_all) next - fix i assume "i \ {0..n}" - hence "n - (n - i) = i" by simp - thus "f (n - i) i = f (n - i) (n - (n - i))" by simp + fix i + assume "i \ {0..n}" + then have "n - (n - i) = i" by simp + then show "f (n - i) i = f (n - i) (n - (n - i))" by simp qed instance fps :: (comm_semiring_0) ab_semigroup_mult @@ -184,73 +188,61 @@ fix n :: nat have "(\i=0..n. a$i * b$(n - i)) = (\i=0..n. a$(n - i) * b$i)" by (rule fps_mult_commute_lemma) - thus "(a * b) $ n = (b * a) $ n" + then show "(a * b) $ n = (b * a) $ n" by (simp add: fps_mult_nth mult_commute) qed qed instance fps :: (monoid_add) monoid_add proof - fix a :: "'a fps" show "0 + a = a " - by (simp add: fps_ext) -next - fix a :: "'a fps" show "a + 0 = a " - by (simp add: fps_ext) + fix a :: "'a fps" + show "0 + a = a" by (simp add: fps_ext) + show "a + 0 = a" by (simp add: fps_ext) qed instance fps :: (comm_monoid_add) comm_monoid_add proof - fix a :: "'a fps" show "0 + a = a " - by (simp add: fps_ext) + fix a :: "'a fps" + show "0 + a = a" by (simp add: fps_ext) qed instance fps :: (semiring_1) monoid_mult proof - fix a :: "'a fps" show "1 * a = a" + fix a :: "'a fps" + show "1 * a = a" by (simp add: fps_ext fps_mult_nth mult_delta_left setsum_delta) -next - fix a :: "'a fps" show "a * 1 = a" + show "a * 1 = a" by (simp add: fps_ext fps_mult_nth mult_delta_right setsum_delta') qed instance fps :: (cancel_semigroup_add) cancel_semigroup_add proof fix a b c :: "'a fps" - assume "a + b = a + c" then show "b = c" - by (simp add: expand_fps_eq) -next - fix a b c :: "'a fps" - assume "b + a = c + a" then show "b = c" - by (simp add: expand_fps_eq) + { assume "a + b = a + c" then show "b = c" by (simp add: expand_fps_eq) } + { assume "b + a = c + a" then show "b = c" by (simp add: expand_fps_eq) } qed instance fps :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add proof fix a b c :: "'a fps" - assume "a + b = a + c" then show "b = c" - by (simp add: expand_fps_eq) + assume "a + b = a + c" + then show "b = c" by (simp add: expand_fps_eq) qed instance fps :: (cancel_comm_monoid_add) cancel_comm_monoid_add .. instance fps :: (group_add) group_add proof - fix a :: "'a fps" show "- a + a = 0" - by (simp add: fps_ext) -next - fix a b :: "'a fps" show "a - b = a + - b" - by (simp add: fps_ext diff_minus) + fix a b :: "'a fps" + show "- a + a = 0" by (simp add: fps_ext) + show "a - b = a + - b" by (simp add: fps_ext diff_minus) qed instance fps :: (ab_group_add) ab_group_add proof - fix a :: "'a fps" - show "- a + a = 0" - by (simp add: fps_ext) -next fix a b :: "'a fps" - show "a - b = a + - b" - by (simp add: fps_ext) + show "- a + a = 0" by (simp add: fps_ext) + show "a - b = a + - b" by (simp add: fps_ext) qed instance fps :: (zero_neq_one) zero_neq_one @@ -261,19 +253,15 @@ fix a b c :: "'a fps" show "(a + b) * c = a * c + b * c" by (simp add: expand_fps_eq fps_mult_nth distrib_right setsum_addf) -next - fix a b c :: "'a fps" show "a * (b + c) = a * b + a * c" by (simp add: expand_fps_eq fps_mult_nth distrib_left setsum_addf) qed instance fps :: (semiring_0) semiring_0 proof - fix a:: "'a fps" show "0 * a = 0" - by (simp add: fps_ext fps_mult_nth) -next - fix a:: "'a fps" show "a * 0 = 0" - by (simp add: fps_ext fps_mult_nth) + fix a:: "'a fps" + show "0 * a = 0" by (simp add: fps_ext fps_mult_nth) + show "a * 0 = 0" by (simp add: fps_ext fps_mult_nth) qed instance fps :: (semiring_0_cancel) semiring_0_cancel .. @@ -284,7 +272,7 @@ by (simp add: expand_fps_eq) lemma fps_nonzero_nth_minimal: - "f \ 0 \ (\n. f $ n \ 0 \ (\m 0 \ (\n. f $ n \ 0 \ (\m < n. f $ m = 0))" proof let ?n = "LEAST n. f $ n \ 0" assume "f \ 0" @@ -304,18 +292,18 @@ lemma fps_eq_iff: "f = g \ (\n. f $ n = g $n)" by (rule expand_fps_eq) -lemma fps_setsum_nth: "(setsum f S) $ n = setsum (\k. (f k) $ n) S" +lemma fps_setsum_nth: "setsum f S $ n = setsum (\k. (f k) $ n) S" proof (cases "finite S") - assume "\ finite S" then show ?thesis by simp + case True + then show ?thesis by (induct set: finite) auto next - assume "finite S" - then show ?thesis by (induct set: finite) auto + case False + then show ?thesis by simp qed subsection{* Injection of the basic ring elements and multiplication by scalars *} -definition - "fps_const c = Abs_fps (\n. if n = 0 then c else 0)" +definition "fps_const c = Abs_fps (\n. if n = 0 then c else 0)" lemma fps_nth_fps_const [simp]: "fps_const c $ n = (if n = 0 then c else 0)" unfolding fps_const_def by simp @@ -331,8 +319,10 @@ lemma fps_const_add [simp]: "fps_const (c::'a\monoid_add) + fps_const d = fps_const (c + d)" by (simp add: fps_ext) + lemma fps_const_sub [simp]: "fps_const (c::'a\group_add) - fps_const d = fps_const (c - d)" by (simp add: fps_ext) + lemma fps_const_mult[simp]: "fps_const (c::'a\ring) * fps_const d = fps_const (c * d)" by (simp add: fps_eq_iff fps_mult_nth setsum_0') @@ -429,18 +419,20 @@ case 0 thus ?case by (simp add: X_def fps_eq_iff) next case (Suc k) - {fix m + { + fix m have "(X^Suc k) $ m = (if m = 0 then (0::'a) else (X^k) $ (m - 1))" - by (simp add: power_Suc del: One_nat_def) - then have "(X^Suc k) $ m = (if m = Suc k then (1::'a) else 0)" - using Suc.hyps by (auto cong del: if_weak_cong)} + by (simp del: One_nat_def) + then have "(X^Suc k) $ m = (if m = Suc k then (1::'a) else 0)" + using Suc.hyps by (auto cong del: if_weak_cong) + } then show ?case by (simp add: fps_eq_iff) qed lemma X_power_mult_nth: "(X^k * (f :: ('a::comm_ring_1) fps)) $n = (if n < k then 0 else f $ (n - k))" apply (induct k arbitrary: n) - apply (simp) + apply simp unfolding power_Suc mult_assoc apply (case_tac n) apply auto @@ -452,7 +444,7 @@ - + subsection{* Formal Power series form a metric space *} definition (in dist) ball_def: "ball x r = {y. dist y x < r}" @@ -460,8 +452,9 @@ instantiation fps :: (comm_ring_1) dist begin -definition dist_fps_def: - "dist (a::'a fps) b = (if (\n. a$n \ b$n) then inverse (2 ^ The (leastP (\n. a$n \ b$n))) else 0)" +definition + dist_fps_def: "dist (a::'a fps) b = + (if (\n. a$n \ b$n) then inverse (2 ^ The (leastP (\n. a$n \ b$n))) else 0)" lemma dist_fps_ge0: "dist (a::'a fps) b \ 0" by (simp add: dist_fps_def) @@ -479,18 +472,22 @@ lemma fps_nonzero_least_unique: assumes a0: "a \ 0" shows "\! n. leastP (\n. a$n \ 0) n" -proof- - from fps_nonzero_nth_minimal[of a] a0 - obtain n where n: "a$n \ 0" "\m < n. a$m = 0" by blast - from n have ln: "leastP (\n. a$n \ 0) n" - by (auto simp add: leastP_def setge_def not_le[symmetric]) +proof - + from fps_nonzero_nth_minimal [of a] a0 + obtain n where "a$n \ 0" "\m < n. a$m = 0" by blast + then have ln: "leastP (\n. a$n \ 0) n" + by (auto simp add: leastP_def setge_def not_le [symmetric]) moreover - {fix m assume "leastP (\n. a$n \ 0) m" + { + fix m + assume "leastP (\n. a $ n \ 0) m" then have "m = n" using ln apply (auto simp add: leastP_def setge_def) apply (erule allE[where x=n]) apply (erule allE[where x=m]) - by simp} + apply simp + done + } ultimately show ?thesis by blast qed @@ -507,21 +504,21 @@ instance proof - fix S :: "'a fps set" + fix S :: "'a fps set" show "open S = (\x\S. \e>0. \y. dist y x < e \ y \ S)" by (auto simp add: open_fps_def ball_def subset_eq) next { fix a b :: "'a fps" { - assume ab: "a = b" - then have "\ (\n. a$n \ b$n)" by simp + assume "a = b" + then have "\ (\n. a $ n \ b $ n)" by simp then have "dist a b = 0" by (simp add: dist_fps_def) } moreover { assume d: "dist a b = 0" - then have "\n. a$n = b$n" + then have "\n. a$n = b$n" by - (rule ccontr, simp add: dist_fps_def) then have "a = b" by (simp add: fps_eq_iff) } @@ -531,23 +528,25 @@ from th have th'[simp]: "\a::'a fps. dist a a = 0" by simp fix a b c :: "'a fps" { - assume ab: "a = b" then have d0: "dist a b = 0" unfolding th . - then have "dist a b \ dist a c + dist b c" - using dist_fps_ge0[of a c] dist_fps_ge0[of b c] by simp + assume "a = b" + then have "dist a b = 0" unfolding th . + then have "dist a b \ dist a c + dist b c" + using dist_fps_ge0 [of a c] dist_fps_ge0 [of b c] by simp } moreover { - assume c: "c = a \ c = b" + assume "c = a \ c = b" then have "dist a b \ dist a c + dist b c" - by (cases "c=a") (simp_all add: th dist_fps_sym) + by (cases "c = a") (simp_all add: th dist_fps_sym) } moreover - {assume ab: "a \ b" and ac: "a \ c" and bc: "b \ c" + { + assume ab: "a \ b" and ac: "a \ c" and bc: "b \ c" let ?P = "\a b n. a$n \ b$n" - from fps_eq_least_unique[OF ab] fps_eq_least_unique[OF ac] + from fps_eq_least_unique[OF ab] fps_eq_least_unique[OF ac] fps_eq_least_unique[OF bc] - obtain nab nac nbc where nab: "leastP (?P a b) nab" - and nac: "leastP (?P a c) nac" + obtain nab nac nbc where nab: "leastP (?P a b) nab" + and nac: "leastP (?P a c) nac" and nbc: "leastP (?P b c) nbc" by blast from nab have nab': "\m. m < nab \ a$m = b$m" "a$nab \ b$nab" by (auto simp add: leastP_def setge_def) @@ -558,41 +557,46 @@ have th0: "\(a::'a fps) b. a \ b \ (\n. a$n \ b$n)" by (simp add: fps_eq_iff) - from ab ac bc nab nac nbc - have dab: "dist a b = inverse (2 ^ nab)" - and dac: "dist a c = inverse (2 ^ nac)" + from ab ac bc nab nac nbc + have dab: "dist a b = inverse (2 ^ nab)" + and dac: "dist a c = inverse (2 ^ nac)" and dbc: "dist b c = inverse (2 ^ nbc)" unfolding th0 apply (simp_all add: dist_fps_def) apply (erule the1_equality[OF fps_eq_least_unique[OF ab]]) apply (erule the1_equality[OF fps_eq_least_unique[OF ac]]) - by (erule the1_equality[OF fps_eq_least_unique[OF bc]]) + apply (erule the1_equality[OF fps_eq_least_unique[OF bc]]) + done from ab ac bc have nz: "dist a b \ 0" "dist a c \ 0" "dist b c \ 0" unfolding th by simp_all from nz have pos: "dist a b > 0" "dist a c > 0" "dist b c > 0" - using dist_fps_ge0[of a b] dist_fps_ge0[of a c] dist_fps_ge0[of b c] + using dist_fps_ge0[of a b] dist_fps_ge0[of a c] dist_fps_ge0[of b c] by auto have th1: "\n. (2::real)^n >0" by auto - {assume h: "dist a b > dist a c + dist b c" + { + assume h: "dist a b > dist a c + dist b c" then have gt: "dist a b > dist a c" "dist a b > dist b c" using pos by auto from gt have gtn: "nab < nbc" "nab < nac" unfolding dab dbc dac by (auto simp add: th1) from nac'(1)[OF gtn(2)] nbc'(1)[OF gtn(1)] - have "a$nab = b$nab" by simp - with nab'(2) have False by simp} + have "a $ nab = b $ nab" by simp + with nab'(2) have False by simp + } then have "dist a b \ dist a c + dist b c" - by (auto simp add: not_le[symmetric]) } + by (auto simp add: not_le[symmetric]) + } ultimately show "dist a b \ dist a c + dist b c" by blast qed - + end text{* The infinite sums and justification of the notation in textbooks*} -lemma reals_power_lt_ex: assumes xp: "x > 0" and y1: "(y::real) > 1" +lemma reals_power_lt_ex: + assumes xp: "x > 0" and y1: "(y::real) > 1" shows "\k>0. (1/y)^k < x" -proof- +proof - have yp: "y > 0" using y1 by simp from reals_Archimedean2[of "max 0 (- log y x) + 1"] obtain k::nat where k: "real k > max 0 (- log y x) + 1" by blast @@ -605,62 +609,72 @@ then have "exp (real k * ln y + ln x) > exp 0" by (simp add: mult_ac) then have "y ^ k * x > 1" - unfolding exp_zero exp_add exp_real_of_nat_mult - exp_ln[OF xp] exp_ln[OF yp] by simp - then have "x > (1/y)^k" using yp + unfolding exp_zero exp_add exp_real_of_nat_mult exp_ln [OF xp] exp_ln [OF yp] + by simp + then have "x > (1 / y)^k" using yp by (simp add: field_simps nonzero_power_divide) then show ?thesis using kp by blast qed + lemma X_nth[simp]: "X$n = (if n = 1 then 1 else 0)" by (simp add: X_def) + lemma X_power_nth[simp]: "(X^k) $n = (if n = k then 1 else (0::'a::comm_ring_1))" by (simp add: X_power_iff) - + lemma fps_sum_rep_nth: "(setsum (%i. fps_const(a$i)*X^i) {0..m})$n = (if n \ m then a$n else (0::'a::comm_ring_1))" - apply (auto simp add: fps_eq_iff fps_setsum_nth X_power_nth cond_application_beta cond_value_iff - cong del: if_weak_cong) + apply (auto simp add: fps_setsum_nth cond_value_iff cong del: if_weak_cong) apply (simp add: setsum_delta') done - -lemma fps_notation: + +lemma fps_notation: "(%n. setsum (%i. fps_const(a$i) * X^i) {0..n}) ----> a" (is "?s ----> a") -proof- - {fix r:: real - assume rp: "r > 0" - have th0: "(2::real) > 1" by simp - from reals_power_lt_ex[OF rp th0] - obtain n0 where n0: "(1/2)^n0 < r" "n0 > 0" by blast - {fix n::nat - assume nn0: "n \ n0" - then have thnn0: "(1/2)^n <= (1/2 :: real)^n0" +proof - + { + fix r:: real + assume rp: "r > 0" + have th0: "(2::real) > 1" by simp + from reals_power_lt_ex[OF rp th0] + obtain n0 where n0: "(1/2)^n0 < r" "n0 > 0" by blast + { + fix n::nat + assume nn0: "n \ n0" + then have thnn0: "(1/2)^n <= (1/2 :: real)^n0" + by (auto intro: power_decreasing) + { + assume "?s n = a" + then have "dist (?s n) a < r" + unfolding dist_eq_0_iff[of "?s n" a, symmetric] + using rp by (simp del: dist_eq_0_iff) + } + moreover + { + assume neq: "?s n \ a" + from fps_eq_least_unique[OF neq] + obtain k where k: "leastP (\i. ?s n $ i \ a$i) k" by blast + have th0: "\(a::'a fps) b. a \ b \ (\n. a$n \ b$n)" + by (simp add: fps_eq_iff) + from neq have dth: "dist (?s n) a = (1/2)^k" + unfolding th0 dist_fps_def + unfolding the1_equality[OF fps_eq_least_unique[OF neq], OF k] + by (auto simp add: inverse_eq_divide power_divide) + + from k have kn: "k > n" + by (simp add: leastP_def setge_def fps_sum_rep_nth split:split_if_asm) + then have "dist (?s n) a < (1/2)^n" unfolding dth + by (auto intro: power_strict_decreasing) + also have "\ <= (1/2)^n0" using nn0 by (auto intro: power_decreasing) - {assume "?s n = a" then have "dist (?s n) a < r" - unfolding dist_eq_0_iff[of "?s n" a, symmetric] - using rp by (simp del: dist_eq_0_iff)} - moreover - {assume neq: "?s n \ a" - from fps_eq_least_unique[OF neq] - obtain k where k: "leastP (\i. ?s n $ i \ a$i) k" by blast - have th0: "\(a::'a fps) b. a \ b \ (\n. a$n \ b$n)" - by (simp add: fps_eq_iff) - from neq have dth: "dist (?s n) a = (1/2)^k" - unfolding th0 dist_fps_def - unfolding the1_equality[OF fps_eq_least_unique[OF neq], OF k] - by (auto simp add: inverse_eq_divide power_divide) - - from k have kn: "k > n" - by (simp add: leastP_def setge_def fps_sum_rep_nth split:split_if_asm) - then have "dist (?s n) a < (1/2)^n" unfolding dth - by (auto intro: power_strict_decreasing) - also have "\ <= (1/2)^n0" using nn0 - by (auto intro: power_decreasing) - also have "\ < r" using n0 by simp - finally have "dist (?s n) a < r" .} - ultimately have "dist (?s n) a < r" by blast} - then have "\n0. \ n \ n0. dist (?s n) a < r " by blast} - then show ?thesis unfolding LIMSEQ_def by blast - qed + also have "\ < r" using n0 by simp + finally have "dist (?s n) a < r" . + } + ultimately have "dist (?s n) a < r" by blast + } + then have "\n0. \ n \ n0. dist (?s n) a < r " by blast + } + then show ?thesis unfolding LIMSEQ_def by blast +qed subsection{* Inverses of formal power series *} @@ -669,52 +683,58 @@ instantiation fps :: ("{comm_monoid_add, inverse, times, uminus}") inverse begin -fun natfun_inverse:: "'a fps \ nat \ 'a" where +fun natfun_inverse:: "'a fps \ nat \ 'a" +where "natfun_inverse f 0 = inverse (f$0)" | "natfun_inverse f n = - inverse (f$0) * setsum (\i. f$i * natfun_inverse f (n - i)) {1..n}" -definition fps_inverse_def: - "inverse f = (if f $ 0 = 0 then 0 else Abs_fps (natfun_inverse f))" - -definition fps_divide_def: "divide = (\(f::'a fps) g. f * inverse g)" +definition + fps_inverse_def: "inverse f = (if f $ 0 = 0 then 0 else Abs_fps (natfun_inverse f))" + +definition + fps_divide_def: "divide = (\(f::'a fps) g. f * inverse g)" instance .. end -lemma fps_inverse_zero[simp]: +lemma fps_inverse_zero [simp]: "inverse (0 :: 'a::{comm_monoid_add,inverse, times, uminus} fps) = 0" by (simp add: fps_ext fps_inverse_def) -lemma fps_inverse_one[simp]: "inverse (1 :: 'a::{division_ring,zero_neq_one} fps) = 1" +lemma fps_inverse_one [simp]: "inverse (1 :: 'a::{division_ring,zero_neq_one} fps) = 1" apply (auto simp add: expand_fps_eq fps_inverse_def) - by (case_tac n, auto) - -lemma inverse_mult_eq_1 [intro]: assumes f0: "f$0 \ (0::'a::field)" + apply (case_tac n) + apply auto + done + +lemma inverse_mult_eq_1 [intro]: + assumes f0: "f$0 \ (0::'a::field)" shows "inverse f * f = 1" -proof- +proof - have c: "inverse f * f = f * inverse f" by (simp add: mult_commute) from f0 have ifn: "\n. inverse f $ n = natfun_inverse f n" by (simp add: fps_inverse_def) from f0 have th0: "(inverse f * f) $ 0 = 1" by (simp add: fps_mult_nth fps_inverse_def) - {fix n::nat assume np: "n >0 " + { + fix n :: nat + assume np: "n > 0" from np have eq: "{0..n} = {0} \ {1 .. n}" by auto have d: "{0} \ {1 .. n} = {}" by auto - from f0 np have th0: "- (inverse f$n) = + from f0 np have th0: "- (inverse f $ n) = (setsum (\i. f$i * natfun_inverse f (n - i)) {1..n}) / (f$0)" - by (cases n, simp, simp add: divide_inverse fps_inverse_def) + by (cases n) (simp_all add: divide_inverse fps_inverse_def) from th0[symmetric, unfolded nonzero_divide_eq_eq[OF f0]] - have th1: "setsum (\i. f$i * natfun_inverse f (n - i)) {1..n} = - - (f$0) * (inverse f)$n" + have th1: "setsum (\i. f$i * natfun_inverse f (n - i)) {1..n} = - (f$0) * (inverse f)$n" by (simp add: field_simps) have "(f * inverse f) $ n = (\i = 0..n. f $i * natfun_inverse f (n - i))" unfolding fps_mult_nth ifn .. - also have "\ = f$0 * natfun_inverse f n - + (\i = 1..n. f$i * natfun_inverse f (n-i))" + also have "\ = f$0 * natfun_inverse f n + (\i = 1..n. f$i * natfun_inverse f (n-i))" by (simp add: eq) also have "\ = 0" unfolding th1 ifn by simp - finally have "(inverse f * f)$n = 0" unfolding c . } + finally have "(inverse f * f)$n = 0" unfolding c . + } with th0 show ?thesis by (simp add: fps_eq_iff) qed @@ -722,28 +742,34 @@ by (simp add: fps_inverse_def nonzero_imp_inverse_nonzero) lemma fps_inverse_eq_0_iff[simp]: "inverse f = (0:: ('a::field) fps) \ f $0 = 0" -proof- - {assume "f$0 = 0" hence "inverse f = 0" by (simp add: fps_inverse_def)} +proof - + { + assume "f$0 = 0" + then have "inverse f = 0" by (simp add: fps_inverse_def) + } moreover - {assume h: "inverse f = 0" and c: "f $0 \ 0" - from inverse_mult_eq_1[OF c] h have False by simp} + { + assume h: "inverse f = 0" and c: "f $0 \ 0" + from inverse_mult_eq_1[OF c] h have False by simp + } ultimately show ?thesis by blast qed lemma fps_inverse_idempotent[intro]: assumes f0: "f$0 \ (0::'a::field)" shows "inverse (inverse f) = f" -proof- +proof - from f0 have if0: "inverse f $ 0 \ 0" by simp from inverse_mult_eq_1[OF f0] inverse_mult_eq_1[OF if0] - have th0: "inverse f * f = inverse f * inverse (inverse f)" by (simp add: mult_ac) + have "inverse f * f = inverse f * inverse (inverse f)" + by (simp add: mult_ac) then show ?thesis using f0 unfolding mult_cancel_left by simp qed lemma fps_inverse_unique: assumes f0: "f$0 \ (0::'a::field)" and fg: "f*g = 1" shows "inverse f = g" -proof- +proof - from inverse_mult_eq_1[OF f0] fg have th0: "inverse f * f = g * f" by (simp add: mult_ac) then show ?thesis using f0 unfolding mult_cancel_right @@ -755,8 +781,9 @@ apply (rule fps_inverse_unique) apply simp apply (simp add: fps_eq_iff fps_mult_nth) -proof(clarsimp) - fix n::nat assume n: "n > 0" +proof clarsimp + fix n :: nat + assume n: "n > 0" let ?f = "\i. if n = i then (1\'a) else if n - i = 1 then - 1 else 0" let ?g = "\i. if i = n then 1 else if i=n - 1 then - 1 else 0" let ?h = "\i. if i=n - 1 then - 1 else 0" @@ -771,7 +798,8 @@ unfolding th1 apply (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] del: One_nat_def) unfolding th2 - by(simp add: setsum_delta) + apply (simp add: setsum_delta) + done qed subsection{* Formal Derivatives, and the MacLaurin theorem around 0*} @@ -789,9 +817,9 @@ lemma fps_deriv_mult[simp]: fixes f :: "('a :: comm_ring_1) fps" shows "fps_deriv (f * g) = f * fps_deriv g + fps_deriv f * g" -proof- +proof - let ?D = "fps_deriv" - {fix n::nat + { fix n::nat let ?Zn = "{0 ..n}" let ?Zn1 = "{0 .. n + 1}" let ?f = "\i. i + 1" @@ -801,25 +829,33 @@ of_nat (i+1)* f $ (i+1) * g $ (n - i)" let ?h = "\i. of_nat i * g $ i * f $ ((n+1) - i) + of_nat i* f $ i * g $ ((n + 1) - i)" - {fix k assume k: "k \ {0..n}" - have "?h (k + 1) = ?g k" using k by auto} + { + fix k + assume k: "k \ {0..n}" + have "?h (k + 1) = ?g k" using k by auto + } note th0 = this have eq': "{0..n +1}- {1 .. n+1} = {0}" by auto - have s0: "setsum (\i. of_nat i * f $ i * g $ (n + 1 - i)) ?Zn1 = setsum (\i. of_nat (n + 1 - i) * f $ (n + 1 - i) * g $ i) ?Zn1" + have s0: "setsum (\i. of_nat i * f $ i * g $ (n + 1 - i)) ?Zn1 = + setsum (\i. of_nat (n + 1 - i) * f $ (n + 1 - i) * g $ i) ?Zn1" apply (rule setsum_reindex_cong[where f="\i. n + 1 - i"]) apply (simp add: inj_on_def Ball_def) apply presburger apply (rule set_eqI) apply (presburger add: image_iff) - by simp - have s1: "setsum (\i. f $ i * g $ (n + 1 - i)) ?Zn1 = setsum (\i. f $ (n + 1 - i) * g $ i) ?Zn1" + apply simp + done + have s1: "setsum (\i. f $ i * g $ (n + 1 - i)) ?Zn1 = + setsum (\i. f $ (n + 1 - i) * g $ i) ?Zn1" apply (rule setsum_reindex_cong[where f="\i. n + 1 - i"]) apply (simp add: inj_on_def Ball_def) apply presburger apply (rule set_eqI) apply (presburger add: image_iff) - by simp - have "(f * ?D g + ?D f * g)$n = (?D g * f + ?D f * g)$n" by (simp only: mult_commute) + apply simp + done + have "(f * ?D g + ?D f * g)$n = (?D g * f + ?D f * g)$n" + by (simp only: mult_commute) also have "\ = (\i = 0..n. ?g i)" by (simp add: fps_mult_nth setsum_addf[symmetric]) also have "\ = setsum ?h {1..n+1}" @@ -829,14 +865,17 @@ apply simp apply (simp add: subset_eq) unfolding eq' - by simp + apply simp + done also have "\ = (fps_deriv (f * g)) $ n" apply (simp only: fps_deriv_nth fps_mult_nth setsum_addf) unfolding s0 s1 unfolding setsum_addf[symmetric] setsum_right_distrib apply (rule setsum_cong2) - by (auto simp add: of_nat_diff field_simps) - finally have "(f * ?D g + ?D f * g) $ n = ?D (f*g) $ n" .} + apply (auto simp add: of_nat_diff field_simps) + done + finally have "(f * ?D g + ?D f * g) $ n = ?D (f*g) $ n" . + } then show ?thesis unfolding fps_eq_iff by auto qed @@ -845,6 +884,7 @@ lemma fps_deriv_neg[simp]: "fps_deriv (- (f:: ('a:: comm_ring_1) fps)) = - (fps_deriv f)" by (simp add: fps_eq_iff fps_deriv_def) + lemma fps_deriv_add[simp]: "fps_deriv ((f:: ('a::comm_ring_1) fps) + g) = fps_deriv f + fps_deriv g" using fps_deriv_linear[of 1 f 1 g] by simp @@ -871,11 +911,14 @@ lemma fps_deriv_setsum: "fps_deriv (setsum f S) = setsum (\i. fps_deriv (f i :: ('a::comm_ring_1) fps)) S" proof- - { assume "\ finite S" hence ?thesis by simp } + { + assume "\ finite S" + then have ?thesis by simp + } moreover { assume fS: "finite S" - have ?thesis by (induct rule: finite_induct[OF fS]) simp_all + have ?thesis by (induct rule: finite_induct [OF fS]) simp_all } ultimately show ?thesis by blast qed @@ -883,23 +926,29 @@ lemma fps_deriv_eq_0_iff[simp]: "fps_deriv f = 0 \ (f = fps_const (f$0 :: 'a::{idom,semiring_char_0}))" proof- - {assume "f= fps_const (f$0)" hence "fps_deriv f = fps_deriv (fps_const (f$0))" by simp - hence "fps_deriv f = 0" by simp } + { + assume "f = fps_const (f$0)" + then have "fps_deriv f = fps_deriv (fps_const (f$0))" by simp + then have "fps_deriv f = 0" by simp + } moreover - {assume z: "fps_deriv f = 0" - hence "\n. (fps_deriv f)$n = 0" by simp - hence "\n. f$(n+1) = 0" by (simp del: of_nat_Suc of_nat_add One_nat_def) - hence "f = fps_const (f$0)" + { + assume z: "fps_deriv f = 0" + then have "\n. (fps_deriv f)$n = 0" by simp + then have "\n. f$(n+1) = 0" by (simp del: of_nat_Suc of_nat_add One_nat_def) + then have "f = fps_const (f$0)" apply (clarsimp simp add: fps_eq_iff fps_const_def) apply (erule_tac x="n - 1" in allE) - by simp} + apply simp + done + } ultimately show ?thesis by blast qed lemma fps_deriv_eq_iff: fixes f:: "('a::{idom,semiring_char_0}) fps" shows "fps_deriv f = fps_deriv g \ (f = fps_const(f$0 - g$0) + g)" -proof- +proof - have "fps_deriv f = fps_deriv g \ fps_deriv (f - g) = 0" by simp also have "\ \ f - g = fps_const ((f-g)$0)" unfolding fps_deriv_eq_0_iff .. finally show ?thesis by (simp add: field_simps) @@ -907,7 +956,8 @@ lemma fps_deriv_eq_iff_ex: "(fps_deriv f = fps_deriv g) \ (\(c::'a::{idom,semiring_char_0}). f = fps_const c + g)" - apply auto unfolding fps_deriv_eq_iff + apply auto + unfolding fps_deriv_eq_iff apply blast done @@ -957,12 +1007,15 @@ lemma fps_nth_deriv_setsum: "fps_nth_deriv n (setsum f S) = setsum (\i. fps_nth_deriv n (f i :: ('a::comm_ring_1) fps)) S" -proof- - { assume "\ finite S" hence ?thesis by simp } +proof - + { + assume "\ finite S" + then have ?thesis by simp + } moreover { assume fS: "finite S" - have ?thesis by (induct rule: finite_induct[OF fS]) simp_all + have ?thesis by (induct rule: finite_induct[OF fS]) simp_all } ultimately show ?thesis by blast qed @@ -977,13 +1030,15 @@ by (induct n) (auto simp add: expand_fps_eq fps_mult_nth) lemma fps_power_first_eq: "(a:: 'a::comm_ring_1 fps)$0 =1 \ a^n $ 1 = of_nat n * a$1" -proof(induct n) - case 0 thus ?case by simp +proof (induct n) + case 0 + then show ?case by simp next case (Suc n) note h = Suc.hyps[OF `a$0 = 1`] show ?case unfolding power_Suc fps_mult_nth - using h `a$0 = 1` fps_power_zeroth_eq_one[OF `a$0=1`] by (simp add: field_simps) + using h `a$0 = 1` fps_power_zeroth_eq_one[OF `a$0=1`] + by (simp add: field_simps) qed lemma startsby_one_power:"a $ 0 = (1::'a::comm_ring_1) \ a^n $ 0 = 1" @@ -993,44 +1048,56 @@ by (induct n) (auto simp add: fps_mult_nth) lemma startsby_power:"a $0 = (v::'a::{comm_ring_1}) \ a^n $0 = v^n" - by (induct n) (auto simp add: fps_mult_nth power_Suc) - -lemma startsby_zero_power_iff[simp]: - "a^n $0 = (0::'a::{idom}) \ (n \ 0 \ a$0 = 0)" -apply (rule iffI) -apply (induct n) -apply (auto simp add: fps_mult_nth) -apply (rule startsby_zero_power, simp_all) -done + by (induct n) (auto simp add: fps_mult_nth) + +lemma startsby_zero_power_iff[simp]: "a^n $0 = (0::'a::{idom}) \ (n \ 0 \ a$0 = 0)" + apply (rule iffI) + apply (induct n) + apply (auto simp add: fps_mult_nth) + apply (rule startsby_zero_power, simp_all) + done lemma startsby_zero_power_prefix: assumes a0: "a $0 = (0::'a::idom)" shows "\n < k. a ^ k $ n = 0" using a0 proof(induct k rule: nat_less_induct) - fix k assume H: "\m (\n'a)" + fix k + assume H: "\m (\n'a)" let ?ths = "\m 0" - have "a ^k $ m = (a^l * a) $m" by (simp add: k power_Suc mult_commute) + { + assume m0: "m \ 0" + have "a ^k $ m = (a^l * a) $m" by (simp add: k mult_commute) also have "\ = (\i = 0..m. a ^ l $ i * a $ (m - i))" by (simp add: fps_mult_nth) - also have "\ = 0" apply (rule setsum_0') + also have "\ = 0" + apply (rule setsum_0') apply auto apply (case_tac "x = m") - using a0 - apply simp + using a0 apply simp apply (rule H[rule_format]) - using a0 k mk by auto - finally have "a^k $ m = 0" .} - ultimately have "a^k $ m = 0" by blast} - hence ?ths by blast} - ultimately show ?ths by (cases k, auto) + using a0 k mk apply auto + done + finally have "a^k $ m = 0" . + } + ultimately have "a^k $ m = 0" by blast + } + then have ?ths by blast + } + ultimately show ?ths by (cases k) auto qed lemma startsby_zero_setsum_depends: @@ -1039,16 +1106,20 @@ apply (rule setsum_mono_zero_right) using kn apply auto apply (rule startsby_zero_power_prefix[rule_format, OF a0]) - by arith - -lemma startsby_zero_power_nth_same: assumes a0: "a$0 = (0::'a::{idom})" + apply arith + done + +lemma startsby_zero_power_nth_same: + assumes a0: "a$0 = (0::'a::{idom})" shows "a^n $ n = (a$1) ^ n" -proof(induct n) - case 0 thus ?case by (simp add: power_0) +proof (induct n) + case 0 + then show ?case by (simp add: power_0) next case (Suc n) - have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)" by (simp add: field_simps power_Suc) - also have "\ = setsum (\i. a^n$i * a $ (Suc n - i)) {0.. Suc n}" by (simp add: fps_mult_nth) + have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)" by (simp add: field_simps) + also have "\ = setsum (\i. a^n$i * a $ (Suc n - i)) {0.. Suc n}" + by (simp add: fps_mult_nth) also have "\ = setsum (\i. a^n$i * a $ (Suc n - i)) {n .. Suc n}" apply (rule setsum_mono_zero_right) apply simp @@ -1058,23 +1129,28 @@ apply arith done also have "\ = a^n $ n * a$1" using a0 by simp - finally show ?case using Suc.hyps by (simp add: power_Suc) + finally show ?case using Suc.hyps by simp qed lemma fps_inverse_power: fixes a :: "('a::{field}) fps" shows "inverse (a^n) = inverse a ^ n" -proof- - {assume a0: "a$0 = 0" - hence eq: "inverse a = 0" by (simp add: fps_inverse_def) - {assume "n = 0" hence ?thesis by simp} +proof - + { + assume a0: "a$0 = 0" + then have eq: "inverse a = 0" by (simp add: fps_inverse_def) + { assume "n = 0" hence ?thesis by simp } moreover - {assume n: "n > 0" + { + assume n: "n > 0" from startsby_zero_power[OF a0 n] eq a0 n have ?thesis - by (simp add: fps_inverse_def)} - ultimately have ?thesis by blast} + by (simp add: fps_inverse_def) + } + ultimately have ?thesis by blast + } moreover - {assume a0: "a$0 \ 0" + { + assume a0: "a$0 \ 0" have ?thesis apply (rule fps_inverse_unique) apply (simp add: a0) @@ -1082,16 +1158,18 @@ apply (rule ssubst[where t = "a * inverse a" and s= 1]) apply simp_all apply (subst mult_commute) - by (rule inverse_mult_eq_1[OF a0])} + apply (rule inverse_mult_eq_1[OF a0]) + done + } ultimately show ?thesis by blast qed lemma fps_deriv_power: "fps_deriv (a ^ n) = fps_const (of_nat n :: 'a:: comm_ring_1) * fps_deriv a * a ^ (n - 1)" apply (induct n) - apply (auto simp add: power_Suc field_simps fps_const_add[symmetric] simp del: fps_const_add) + apply (auto simp add: field_simps fps_const_add[symmetric] simp del: fps_const_add) apply (case_tac n) - apply (auto simp add: power_Suc field_simps) + apply (auto simp add: field_simps) done lemma fps_inverse_deriv: @@ -1379,7 +1457,7 @@ apply (rule bexI[where x = "?m"]) apply (rule exI[where x = "?xs"]) apply (rule exI[where x = "?ys"]) - using ls l + using ls l apply (auto simp add: natpermute_def l_take_drop simp del: append_take_drop_id) apply simp done @@ -1719,7 +1797,7 @@ ultimately show "?r ^ Suc k $ n = a $n" by (cases n, auto) qed } then have ?thesis using r0 by (simp add: fps_eq_iff)} -moreover +moreover { assume h: "(fps_radical r (Suc k) a) ^ (Suc k) = a" hence "((fps_radical r (Suc k) a) ^ (Suc k))$0 = a$0" by simp then have "(r (Suc k) (a$0)) ^ Suc k = a$0" @@ -1998,15 +2076,15 @@ {assume ?rhs then have "?r (a/b) $ 0 = (?r a / ?r b)$0" by simp - then have ?lhs using k a0 b0 rb0' + then have ?lhs using k a0 b0 rb0' by (simp add: fps_divide_def fps_mult_nth fps_inverse_def divide_inverse) } moreover {assume h: ?lhs - from a0 b0 have ab0[simp]: "(a/b)$0 = a$0 / b$0" + from a0 b0 have ab0[simp]: "(a/b)$0 = a$0 / b$0" by (simp add: fps_divide_def fps_mult_nth divide_inverse fps_inverse_def) have th0: "r k ((a/b)$0) ^ k = (a/b)$0" by (simp add: h nonzero_power_divide[OF rb0'] ra0 rb0 del: k) - from a0 b0 ra0' rb0' kp h + from a0 b0 ra0' rb0' kp h have th1: "r k ((a / b) $ 0) = (fps_radical r k a / fps_radical r k b) $ 0" by (simp add: fps_divide_def fps_mult_nth fps_inverse_def divide_inverse del: k) from a0 b0 ra0' rb0' kp have ab0': "(a / b) $ 0 \ 0" @@ -2360,7 +2438,7 @@ then have "(?ia oo b) * (a oo b) * ?iab = 1 * ?iab" by simp then have "(?ia oo b) * (?iab * (a oo b)) = ?iab" by simp -then show ?thesis +then show ?thesis unfolding inverse_mult_eq_1[OF ab0] by simp qed @@ -2374,11 +2452,11 @@ shows "(Abs_fps (\n. 1)) oo a = 1/(1 - a)" (is "?one oo a = _") proof- have o0: "?one $ 0 \ 0" by simp - have th0: "(1 - X) $ 0 \ (0::'a)" by simp + have th0: "(1 - X) $ 0 \ (0::'a)" by simp from fps_inverse_gp[where ?'a = 'a] have "inverse ?one = 1 - X" by (simp add: fps_eq_iff) hence "inverse (inverse ?one) = inverse (1 - X)" by simp - hence th: "?one = 1/(1 - X)" unfolding fps_inverse_idempotent[OF o0] + hence th: "?one = 1/(1 - X)" unfolding fps_inverse_idempotent[OF o0] by (simp add: fps_divide_def) show ?thesis unfolding th unfolding fps_divide_compose[OF a0 th0] @@ -2402,8 +2480,8 @@ by (simp add: ab0 fps_compose_def) have th0: "(?r a oo b) ^ (Suc k) = a oo b" unfolding fps_compose_power[OF b0] - unfolding iffD1[OF power_radical[of a r k], OF a0 ra0] .. - from iffD1[OF radical_unique[where r=r and k=k and b= ?ab and a = "?r a oo b", OF rab0(2) th00 rab0(1)], OF th0] show ?thesis . + unfolding iffD1[OF power_radical[of a r k], OF a0 ra0] .. + from iffD1[OF radical_unique[where r=r and k=k and b= ?ab and a = "?r a oo b", OF rab0(2) th00 rab0(1)], OF th0] show ?thesis . qed lemma fps_const_mult_apply_left: @@ -2485,7 +2563,7 @@ show "?dia = inverse ?d" by simp qed -lemma fps_inv_idempotent: +lemma fps_inv_idempotent: assumes a0: "a$0 = 0" and a1: "a$1 \ 0" shows "fps_inv (fps_inv a) = a" proof- @@ -2495,7 +2573,7 @@ have X0: "X$0 = 0" by simp from fps_inv[OF ra0 ra1] have "?r (?r a) oo ?r a = X" . then have "?r (?r a) oo ?r a oo a = X oo a" by simp - then have "?r (?r a) oo (?r a oo a) = a" + then have "?r (?r a) oo (?r a oo a) = a" unfolding X_fps_compose_startby0[OF a0] unfolding fps_compose_assoc[OF a0 ra0, symmetric] . then show ?thesis unfolding fps_inv[OF a0 a1] by simp @@ -2509,7 +2587,7 @@ let ?r = "fps_ginv" from c0 have rca0: "?r c a $0 = 0" by (simp add: fps_ginv_def) from a1 c1 have rca1: "?r c a $ 1 \ 0" by (simp add: fps_ginv_def field_simps) - from fps_ginv[OF rca0 rca1] + from fps_ginv[OF rca0 rca1] have "?r b (?r c a) oo ?r c a = b" . then have "?r b (?r c a) oo ?r c a oo a = b oo a" by simp then have "?r b (?r c a) oo (?r c a oo a) = b oo a" @@ -2537,7 +2615,7 @@ from fps_ginv[OF a0 a1, of b] have "?d (?ia oo a) = fps_deriv b" by simp then have "(?d ?ia oo a) * ?d a = ?d b" unfolding fps_compose_deriv[OF a0] . then have "(?d ?ia oo a) * ?d a * inverse (?d a) = ?d b * inverse (?d a)" by simp - then have "(?d ?ia oo a) * (inverse (?d a) * ?d a) = ?d b / ?d a" + then have "(?d ?ia oo a) * (inverse (?d a) * ?d a) = ?d b / ?d a" by (simp add: fps_divide_def) then have "(?d ?ia oo a) oo ?iXa = (?d b / ?d a) oo ?iXa " unfolding inverse_mult_eq_1[OF da0] by simp @@ -2649,7 +2727,7 @@ by (induct n) (auto simp add: field_simps E_add_mult power_Suc) lemma radical_E: - assumes r: "r (Suc k) 1 = 1" + assumes r: "r (Suc k) 1 = 1" shows "fps_radical r (Suc k) (E (c::'a::{field_char_0})) = E (c / of_nat (Suc k))" proof- let ?ck = "(c / of_nat (Suc k))" @@ -2660,24 +2738,24 @@ have th: "r (Suc k) (E c $0) ^ Suc k = E c $ 0" "r (Suc k) (E c $ 0) = E ?ck $ 0" "E c $ 0 \ 0" using r by simp_all from th0 radical_unique[where r=r and k=k, OF th] - show ?thesis by auto + show ?thesis by auto qed -lemma Ec_E1_eq: +lemma Ec_E1_eq: "E (1::'a::{field_char_0}) oo (fps_const c * X) = E c" apply (auto simp add: fps_eq_iff E_def fps_compose_def power_mult_distrib) by (simp add: cond_value_iff cond_application_beta setsum_delta' cong del: if_weak_cong) text{* The generalized binomial theorem as a consequence of @{thm E_add_mult} *} -lemma gbinomial_theorem: +lemma gbinomial_theorem: "((a::'a::{field_char_0, field_inverse_zero})+b) ^ n = (\k=0..n. of_nat (n choose k) * a^k * b^(n-k))" proof- - from E_add_mult[of a b] + from E_add_mult[of a b] have "(E (a + b)) $ n = (E a * E b)$n" by simp then have "(a + b) ^ n = (\i\nat = 0\nat..n. a ^ i * b ^ (n - i) * (of_nat (fact n) / of_nat (fact i * fact (n - i))))" by (simp add: field_simps fps_mult_nth of_nat_mult[symmetric] setsum_right_distrib) - then show ?thesis + then show ?thesis apply simp apply (rule setsum_cong2) apply simp @@ -2693,11 +2771,11 @@ subsubsection{* Logarithmic series *} -lemma Abs_fps_if_0: +lemma Abs_fps_if_0: "Abs_fps(%n. if n=0 then (v::'a::ring_1) else f n) = fps_const v + X * Abs_fps (%n. f (Suc n))" by (auto simp add: fps_eq_iff) -definition L:: "'a::{field_char_0} \ 'a fps" where +definition L:: "'a::{field_char_0} \ 'a fps" where "L c \ fps_const (1/c) * Abs_fps (\n. if n = 0 then 0 else (- 1) ^ (n - 1) / of_nat n)" lemma fps_deriv_L: "fps_deriv (L c) = fps_const (1/c) * inverse (1 + X)" @@ -2722,7 +2800,7 @@ finally have eq: "fps_deriv (E a - 1) oo fps_inv (E a - 1) = fps_const a * (X + 1)" . from fps_inv_deriv[OF b0 b1, unfolded eq] have "fps_deriv (fps_inv ?b) = fps_const (inverse a) / (X + 1)" - using a + using a by (simp add: fps_const_inverse eq fps_divide_def fps_inverse_mult) hence "fps_deriv ?l = fps_deriv ?r" by (simp add: fps_deriv_L add_commute fps_divide_def divide_inverse) @@ -2730,7 +2808,7 @@ by (simp add: L_nth fps_inv_def) qed -lemma L_mult_add: +lemma L_mult_add: assumes c0: "c\0" and d0: "d\0" shows "L c + L d = fps_const (c+d) * L (c*d)" (is "?r = ?l") @@ -2768,12 +2846,12 @@ by (simp add: field_simps) finally have eq: "?l = ?r \ ?lhs" by simp moreover - {assume h: "?l = ?r" + {assume h: "?l = ?r" {fix n from h have lrn: "?l $ n = ?r$n" by simp - - from lrn - have "a$ Suc n = ((c - of_nat n) / of_nat (Suc n)) * a $n" + + from lrn + have "a$ Suc n = ((c - of_nat n) / of_nat (Suc n)) * a $n" apply (simp add: field_simps del: of_nat_Suc) by (cases n, simp_all add: field_simps del: of_nat_Suc) } @@ -2796,7 +2874,7 @@ moreover {assume h: ?rhs have th00:"\x y. x * (a$0 * y) = a$0 * (x*y)" by (simp add: mult_commute) - have "?l = ?r" + have "?l = ?r" apply (subst h) apply (subst (2) h) apply (clarsimp simp add: fps_eq_iff field_simps) @@ -2856,10 +2934,10 @@ lemma binomial_Vandermonde: "setsum (\k. (a choose k) * (b choose (n - k))) {0..n} = (a + b) choose n" using gbinomial_Vandermonde[of "(of_nat a)" "of_nat b" n] - + apply (simp only: binomial_gbinomial[symmetric] of_nat_mult[symmetric] of_nat_setsum[symmetric] of_nat_add[symmetric]) by simp - + lemma binomial_Vandermonde_same: "setsum (\k. (n choose k)^2) {0..n} = (2*n) choose n" using binomial_Vandermonde[of n n n,symmetric] unfolding mult_2 apply (simp add: power2_eq_square) @@ -2879,22 +2957,22 @@ {assume c:"pochhammer (b - of_nat n + 1) n = 0" then obtain j where j: "j < n" "b - of_nat n + 1 = - of_nat j" unfolding pochhammer_eq_0_iff by blast - from j have "b = of_nat n - of_nat j - of_nat 1" + from j have "b = of_nat n - of_nat j - of_nat 1" by (simp add: algebra_simps) - then have "b = of_nat (n - j - 1)" + then have "b = of_nat (n - j - 1)" using j kn by (simp add: of_nat_diff) with b have False using j by auto} - then have nz: "pochhammer (1 + b - of_nat n) n \ 0" + then have nz: "pochhammer (1 + b - of_nat n) n \ 0" by (auto simp add: algebra_simps) - - from nz kn [simplified] have nz': "pochhammer (1 + b - of_nat n) k \ 0" + + from nz kn [simplified] have nz': "pochhammer (1 + b - of_nat n) k \ 0" by (rule pochhammer_neq_0_mono) - {assume k0: "k = 0 \ n =0" - then have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" + {assume k0: "k = 0 \ n =0" + then have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" using kn by (cases "k=0", simp_all add: gbinomial_pochhammer)} moreover - {assume n0: "n \ 0" and k0: "k \ 0" + {assume n0: "n \ 0" and k0: "k \ 0" then obtain m where m: "n = Suc m" by (cases n, auto) from k0 obtain h where h: "k = Suc h" by (cases k, auto) {assume kn: "k = n" @@ -2905,27 +2983,27 @@ by (simp add: field_simps power_add[symmetric])} moreover {assume nk: "k \ n" - have m1nk: "?m1 n = setprod (%i. - 1) {0..m}" + have m1nk: "?m1 n = setprod (%i. - 1) {0..m}" "?m1 k = setprod (%i. - 1) {0..h}" by (simp_all add: setprod_constant m h) from kn nk have kn': "k < n" by simp have bnz0: "pochhammer (b - of_nat n + 1) k \ 0" - using bn0 kn + using bn0 kn unfolding pochhammer_eq_0_iff apply auto apply (erule_tac x= "n - ka - 1" in allE) by (auto simp add: algebra_simps of_nat_diff) - have eq1: "setprod (%k. (1::'a) + of_nat m - of_nat k) {0 .. h} = setprod of_nat {Suc (m - h) .. Suc m}" + have eq1: "setprod (%k. (1::'a) + of_nat m - of_nat k) {0 .. h} = setprod of_nat {Suc (m - h) .. Suc m}" apply (rule strong_setprod_reindex_cong[where f="%k. Suc m - k "]) using kn' h m apply (auto simp add: inj_on_def image_def) apply (rule_tac x="Suc m - x" in bexI) apply (simp_all add: of_nat_diff) done - + have th1: "(?m1 k * ?p (of_nat n) k) / ?f n = 1 / of_nat(fact (n - k))" - unfolding m1nk - + unfolding m1nk + unfolding m h pochhammer_Suc_setprod apply (simp add: field_simps del: fact_Suc id_def minus_one) unfolding fact_altdef_nat id_def @@ -2939,21 +3017,21 @@ apply auto done have th20: "?m1 n * ?p b n = setprod (%i. b - of_nat i) {0..m}" - unfolding m1nk + unfolding m1nk unfolding m h pochhammer_Suc_setprod unfolding setprod_timesf[symmetric] apply (rule setprod_cong) apply auto done have th21:"pochhammer (b - of_nat n + 1) k = setprod (%i. b - of_nat i) {n - k .. n - 1}" - unfolding h m + unfolding h m unfolding pochhammer_Suc_setprod apply (rule strong_setprod_reindex_cong[where f="%k. n - 1 - k"]) using kn apply (auto simp add: inj_on_def m h image_def) apply (rule_tac x= "m - x" in bexI) by (auto simp add: of_nat_diff) - + have "?m1 n * ?p b n = pochhammer (b - of_nat n + 1) k * setprod (%i. b - of_nat i) {0.. n - k - 1}" unfolding th20 th21 unfolding h m @@ -2963,24 +3041,24 @@ apply (rule setprod_cong) apply auto done - then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k = setprod (%i. b - of_nat i) {0.. n - k - 1}" + then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k = setprod (%i. b - of_nat i) {0.. n - k - 1}" using nz' by (simp add: field_simps) have "(?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k) = ((?m1 k * ?p (of_nat n) k) / ?f n) * ((?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k)" using bnz0 by (simp add: field_simps) - also have "\ = b gchoose (n - k)" + also have "\ = b gchoose (n - k)" unfolding th1 th2 using kn' by (simp add: gbinomial_def) finally have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" by simp} ultimately have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" by (cases "k =n", auto)} ultimately have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" "pochhammer (1 + b - of_nat n) k \ 0 " - using nz' + using nz' apply (cases "n=0", auto) by (cases "k", auto)} note th00 = this have "?r = ((a + b) gchoose n) * (of_nat (fact n)/ (?m1 n * pochhammer (- b) n))" - unfolding gbinomial_pochhammer + unfolding gbinomial_pochhammer using bn0 by (auto simp add: field_simps) also have "\ = ?l" unfolding gbinomial_Vandermonde[symmetric] @@ -2991,9 +3069,9 @@ apply (drule th00(2)) by (simp add: field_simps power_add[symmetric]) finally show ?thesis by simp -qed - - +qed + + lemma Vandermonde_pochhammer: fixes a :: "'a::field_char_0" assumes c: "ALL i : {0..< n}. c \ - of_nat i" @@ -3211,17 +3289,17 @@ proof- {fix n::nat {assume en: "even n" - from en obtain m where m: "n = 2*m" + from en obtain m where m: "n = 2*m" unfolding even_mult_two_ex by blast - - have "?l $n = ?r$n" + + have "?l $n = ?r$n" by (simp add: m fps_sin_def fps_cos_def power_mult_distrib power_mult power_minus)} moreover {assume on: "odd n" - from on obtain m where m: "n = 2*m + 1" - unfolding odd_nat_equiv_def2 by (auto simp add: mult_2) - have "?l $n = ?r$n" + from on obtain m where m: "n = 2*m + 1" + unfolding odd_nat_equiv_def2 by (auto simp add: mult_2) + have "?l $n = ?r$n" by (simp add: m fps_sin_def fps_cos_def power_mult_distrib power_mult power_minus)} ultimately have "?l $n = ?r$n" by blast} @@ -3240,7 +3318,7 @@ lemma fps_cos_Eii: "fps_cos c = (E (ii * c) + E (- ii * c)) / fps_const 2" proof- - have th: "fps_cos c + fps_cos c = fps_cos c * fps_const 2" + have th: "fps_cos c + fps_cos c = fps_cos c * fps_const 2" by (simp add: numeral_fps_const) show ?thesis unfolding Eii_sin_cos minus_mult_commute @@ -3251,7 +3329,7 @@ lemma fps_sin_Eii: "fps_sin c = (E (ii * c) - E (- ii * c)) / fps_const (2*ii)" proof- - have th: "fps_const \ * fps_sin c + fps_const \ * fps_sin c = fps_sin c * fps_const (2 * ii)" + have th: "fps_const \ * fps_sin c + fps_const \ * fps_sin c = fps_sin c * fps_const (2 * ii)" by (simp add: fps_eq_iff numeral_fps_const) show ?thesis unfolding Eii_sin_cos minus_mult_commute @@ -3287,7 +3365,7 @@ foldr (\b r. r * pochhammer b n) bs (of_nat (fact n))" by (simp add: foldl_mult_start foldr_mult_foldl) -lemma F_E[simp]: "F [] [] c = E c" +lemma F_E[simp]: "F [] [] c = E c" by (simp add: fps_eq_iff) lemma F_1_0[simp]: "F [1] [] c = 1/(1 - fps_const c * X)" @@ -3340,7 +3418,7 @@ lemma XDp_fps_integral[simp]:"XDp 0 (fps_integral a c) = X * a" by (simp add: fps_eq_iff fps_integral_def) -lemma F_minus_nat: +lemma F_minus_nat: "F [- of_nat n] [- of_nat (n + m)] (c::'a::{field_char_0, field_inverse_zero}) $ k = (if k <= n then pochhammer (- of_nat n) k * c ^ k / (pochhammer (- of_nat (n + m)) k * of_nat (fact k)) else 0)" "F [- of_nat m] [- of_nat (m + n)] (c::'a::{field_char_0, field_inverse_zero}) $ k = (if k <= m then pochhammer (- of_nat m) k * c ^ k / @@ -3352,19 +3430,19 @@ apply (subst setsum_insert[symmetric]) by (auto simp add: not_less setsum_head_Suc) -lemma pochhammer_rec_if: +lemma pochhammer_rec_if: "pochhammer a n = (if n = 0 then 1 else a * pochhammer (a + 1) (n - 1))" by (cases n, simp_all add: pochhammer_rec) -lemma XDp_foldr_nth[simp]: "foldr (%c r. XDp c o r) cs (%c. XDp c a) c0 $ n = +lemma XDp_foldr_nth[simp]: "foldr (%c r. XDp c o r) cs (%c. XDp c a) c0 $ n = foldr (%c r. (c + of_nat n) * r) cs (c0 + of_nat n) * a$n" by (induct cs arbitrary: c0) (auto simp add: algebra_simps) lemma genric_XDp_foldr_nth: - assumes + assumes f: "ALL n c a. f c a $ n = (of_nat n + k c) * a$n" - shows "foldr (%c r. f c o r) cs (%c. g c a) c0 $ n = + shows "foldr (%c r. f c o r) cs (%c. g c a) c0 $ n = foldr (%c r. (k c + of_nat n) * r) cs (g c0 a $ n)" by (induct cs arbitrary: c0) (auto simp add: algebra_simps f) diff -r 34c47bc771f2 -r 3695ce0f9f96 src/HOL/ex/Adhoc_Overloading_Examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Adhoc_Overloading_Examples.thy Wed Aug 07 15:43:58 2013 +0200 @@ -0,0 +1,250 @@ +(* Title: HOL/ex/Adhoc_Overloading_Examples.thy + Author: Christian Sternagel +*) + +header {* Ad Hoc Overloading *} + +theory Adhoc_Overloading_Examples +imports + Main + "~~/src/Tools/Adhoc_Overloading" + "~~/src/HOL/Library/Infinite_Set" +begin + +text {*Adhoc overloading allows to overload a constant depending on +its type. Typically this involves to introduce an uninterpreted +constant (used for input and output) and then add some variants (used +internally).*} + +subsection {* Plain Ad Hoc Overloading *} + +text {*Consider the type of first-order terms.*} +datatype ('a, 'b) "term" = + Var 'b | + Fun 'a "('a, 'b) term list" + +text {*The set of variables of a term might be computed as follows.*} +fun term_vars :: "('a, 'b) term \ 'b set" where + "term_vars (Var x) = {x}" | + "term_vars (Fun f ts) = \set (map term_vars ts)" + +text {*However, also for \emph{rules} (i.e., pairs of terms) and term +rewrite systems (i.e., sets of rules), the set of variables makes +sense. Thus we introduce an unspecified constant @{text vars}.*} + +consts vars :: "'a \ 'b set" + +text {*Which is then overloaded with variants for terms, rules, and TRSs.*} +adhoc_overloading + vars term_vars + +value "vars (Fun ''f'' [Var 0, Var 1])" + +fun rule_vars :: "('a, 'b) term \ ('a, 'b) term \ 'b set" where + "rule_vars (l, r) = vars l \ vars r" + +adhoc_overloading + vars rule_vars + +value "vars (Var 1, Var 0)" + +definition trs_vars :: "(('a, 'b) term \ ('a, 'b) term) set \ 'b set" where + "trs_vars R = \(rule_vars ` R)" + +adhoc_overloading + vars trs_vars + +value "vars {(Var 1, Var 0)}" + +text {*Sometimes it is necessary to add explicit type constraints +before a variant can be determined.*} +(*value "vars R" (*has multiple instances*)*) +value "vars (R :: (('a, 'b) term \ ('a, 'b) term) set)" + +text {*It is also possible to remove variants.*} +no_adhoc_overloading + vars term_vars rule_vars + +(*value "vars (Var 1)" (*does not have an instance*)*) + +text {*As stated earlier, the overloaded constant is only used for +input and output. Internally, always a variant is used, as can be +observed by the configuration option @{text show_variants}.*} + +adhoc_overloading + vars term_vars + +declare [[show_variants]] + +term "vars (Var 1)" (*which yields: "term_vars (Var 1)"*) + + +subsection {* Adhoc Overloading inside Locales *} + +text {*As example we use permutations that are parametrized over an +atom type @{typ "'a"}.*} + +definition perms :: "('a \ 'a) set" where + "perms = {f. bij f \ finite {x. f x \ x}}" + +typedef 'a perm = "perms :: ('a \ 'a) set" + by (default) (auto simp: perms_def) + +text {*First we need some auxiliary lemmas.*} +lemma permsI [Pure.intro]: + assumes "bij f" and "MOST x. f x = x" + shows "f \ perms" + using assms by (auto simp: perms_def) (metis MOST_iff_finiteNeg) + +lemma perms_imp_bij: + "f \ perms \ bij f" + by (simp add: perms_def) + +lemma perms_imp_MOST_eq: + "f \ perms \ MOST x. f x = x" + by (simp add: perms_def) (metis MOST_iff_finiteNeg) + +lemma id_perms [simp]: + "id \ perms" + "(\x. x) \ perms" + by (auto simp: perms_def bij_def) + +lemma perms_comp [simp]: + assumes f: "f \ perms" and g: "g \ perms" + shows "(f \ g) \ perms" + apply (intro permsI bij_comp) + apply (rule perms_imp_bij [OF g]) + apply (rule perms_imp_bij [OF f]) + apply (rule MOST_rev_mp [OF perms_imp_MOST_eq [OF g]]) + apply (rule MOST_rev_mp [OF perms_imp_MOST_eq [OF f]]) + by simp + +lemma perms_inv: + assumes f: "f \ perms" + shows "inv f \ perms" + apply (rule permsI) + apply (rule bij_imp_bij_inv) + apply (rule perms_imp_bij [OF f]) + apply (rule MOST_mono [OF perms_imp_MOST_eq [OF f]]) + apply (erule subst, rule inv_f_f) + by (rule bij_is_inj [OF perms_imp_bij [OF f]]) + +lemma bij_Rep_perm: "bij (Rep_perm p)" + using Rep_perm [of p] unfolding perms_def by simp + +instantiation perm :: (type) group_add +begin + +definition "0 = Abs_perm id" +definition "- p = Abs_perm (inv (Rep_perm p))" +definition "p + q = Abs_perm (Rep_perm p \ Rep_perm q)" +definition "(p1::'a perm) - p2 = p1 + - p2" + +lemma Rep_perm_0: "Rep_perm 0 = id" + unfolding zero_perm_def by (simp add: Abs_perm_inverse) + +lemma Rep_perm_add: + "Rep_perm (p1 + p2) = Rep_perm p1 \ Rep_perm p2" + unfolding plus_perm_def by (simp add: Abs_perm_inverse Rep_perm) + +lemma Rep_perm_uminus: + "Rep_perm (- p) = inv (Rep_perm p)" + unfolding uminus_perm_def by (simp add: Abs_perm_inverse perms_inv Rep_perm) + +instance + apply default + unfolding Rep_perm_inject [symmetric] + unfolding minus_perm_def + unfolding Rep_perm_add + unfolding Rep_perm_uminus + unfolding Rep_perm_0 + by (simp_all add: o_assoc inv_o_cancel [OF bij_is_inj [OF bij_Rep_perm]]) + +end + +lemmas Rep_perm_simps = + Rep_perm_0 + Rep_perm_add + Rep_perm_uminus + + +section {* Permutation Types *} + +text {*We want to be able to apply permutations to arbitrary types. To +this end we introduce a constant @{text PERMUTE} together with +convenient infix syntax.*} + +consts PERMUTE :: "'a perm \ 'b \ 'b" (infixr "\" 75) + +text {*Then we add a locale for types @{typ 'b} that support +appliciation of permutations.*} +locale permute = + fixes permute :: "'a perm \ 'b \ 'b" + assumes permute_zero [simp]: "permute 0 x = x" + and permute_plus [simp]: "permute (p + q) x = permute p (permute q x)" +begin + +adhoc_overloading + PERMUTE permute + +end + +text {*Permuting atoms.*} +definition permute_atom :: "'a perm \ 'a \ 'a" where + "permute_atom p a = (Rep_perm p) a" + +adhoc_overloading + PERMUTE permute_atom + +interpretation atom_permute: permute permute_atom + by (default) (simp add: permute_atom_def Rep_perm_simps)+ + +text {*Permuting permutations.*} +definition permute_perm :: "'a perm \ 'a perm \ 'a perm" where + "permute_perm p q = p + q - p" + +adhoc_overloading + PERMUTE permute_perm + +interpretation perm_permute: permute permute_perm + by (default) (simp add: diff_minus minus_add add_assoc permute_perm_def)+ + +text {*Permuting functions.*} +locale fun_permute = + dom: permute perm1 + ran: permute perm2 + for perm1 :: "'a perm \ 'b \ 'b" + and perm2 :: "'a perm \ 'c \ 'c" +begin + +adhoc_overloading + PERMUTE perm1 perm2 + +definition permute_fun :: "'a perm \ ('b \ 'c) \ ('b \ 'c)" where + "permute_fun p f = (\x. p \ (f (-p \ x)))" + +adhoc_overloading + PERMUTE permute_fun + +end + +sublocale fun_permute \ permute permute_fun + by (unfold_locales, auto simp: permute_fun_def) + (metis dom.permute_plus minus_add) + +lemma "(Abs_perm id :: nat perm) \ Suc 0 = Suc 0" + unfolding permute_atom_def + by (metis Rep_perm_0 id_apply zero_perm_def) + +interpretation atom_fun_permute: fun_permute permute_atom permute_atom + by (unfold_locales) + +adhoc_overloading + PERMUTE atom_fun_permute.permute_fun + +lemma "(Abs_perm id :: 'a perm) \ id = id" + unfolding atom_fun_permute.permute_fun_def + unfolding permute_atom_def + by (metis Rep_perm_0 id_def inj_imp_inv_eq inj_on_id uminus_perm_def zero_perm_def) + +end + diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Pure/General/symbol.scala --- a/src/Pure/General/symbol.scala Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Pure/General/symbol.scala Wed Aug 07 15:43:58 2013 +0200 @@ -193,7 +193,7 @@ if (min <= c && c <= max) { matcher.region(i, len).lookingAt val x = matcher.group - result.append(table.get(x) getOrElse x) + result.append(table.getOrElse(x, x)) i = matcher.end } else { result.append(c); i += 1 } @@ -255,7 +255,7 @@ } val groups: List[(String, List[Symbol])] = - symbols.map({ case (sym, props) => (sym, props.get("group") getOrElse "unsorted") }) + symbols.map({ case (sym, props) => (sym, props.getOrElse("group", "unsorted")) }) .groupBy(_._2).toList.map({ case (group, list) => (group, list.map(_._1)) }) .sortBy(_._1) diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Pure/PIDE/document.scala --- a/src/Pure/PIDE/document.scala Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Pure/PIDE/document.scala Wed Aug 07 15:43:58 2013 +0200 @@ -15,48 +15,6 @@ { /** document structure **/ - /* overlays -- print functions with arguments */ - - type Overlay = (Command, (String, List[String])) - - object Overlays - { - val empty = new Overlays(Map.empty) - } - - final class Overlays private(rep: Map[Command, List[(String, List[String])]]) - { - def commands: Set[Command] = rep.keySet - def is_empty: Boolean = rep.isEmpty - - def insert(cmd: Command, fn: (String, List[String])): Overlays = - { - val fns = rep.get(cmd) getOrElse Nil - if (fns.contains(fn)) this - else new Overlays(rep + (cmd -> (fn :: fns))) - } - - def remove(cmd: Command, fn: (String, List[String])): Overlays = - rep.get(cmd) match { - case Some(fns) => - if (fns.contains(fn)) { - fns.filterNot(_ == fn) match { - case Nil => new Overlays(rep - cmd) - case fns1 => new Overlays(rep + (cmd -> fns1)) - } - } - else this - case None => this - } - - def dest: List[Overlay] = - (for { - (cmd, fns) <- rep.iterator - fn <- fns - } yield (cmd, fn)).toList - } - - /* individual nodes */ type Edit[A, B] = (Node.Name, Node.Edit[A, B]) @@ -65,6 +23,11 @@ object Node { + val empty: Node = new Node() + + + /* header and name */ + sealed case class Header( imports: List[Name], keywords: Thy_Header.Keywords, @@ -84,6 +47,7 @@ def compare(name1: Name, name2: Name): Int = name1.node compare name2.node } } + sealed case class Name(node: String, dir: String, theory: String) { override def hashCode: Int = node.hashCode @@ -95,6 +59,51 @@ override def toString: String = theory } + + /* overlays -- print functions with arguments */ + + type Overlay = (Command, (String, List[String])) + + object Overlays + { + val empty = new Overlays(Map.empty) + } + + final class Overlays private(rep: Map[Command, List[(String, List[String])]]) + { + def commands: Set[Command] = rep.keySet + def is_empty: Boolean = rep.isEmpty + + def insert(cmd: Command, over: (String, List[String])): Overlays = + { + val overs = rep.getOrElse(cmd, Nil) + if (overs.contains(over)) this + else new Overlays(rep + (cmd -> (over :: overs))) + } + + def remove(cmd: Command, over: (String, List[String])): Overlays = + rep.get(cmd) match { + case Some(overs) => + if (overs.contains(over)) { + overs.filterNot(_ == over) match { + case Nil => new Overlays(rep - cmd) + case overs1 => new Overlays(rep + (cmd -> overs1)) + } + } + else this + case None => this + } + + def dest: List[Overlay] = + (for { + (cmd, overs) <- rep.iterator + over <- overs + } yield (cmd, over)).toList + } + + + /* edits */ + sealed abstract class Edit[A, B] { def foreach(f: A => Unit) @@ -112,6 +121,9 @@ type Perspective_Text = Perspective[Text.Edit, Text.Perspective] type Perspective_Command = Perspective[Command.Edit, Command.Perspective] + + /* commands */ + def command_starts(commands: Iterator[Command], offset: Text.Offset = 0) : Iterator[(Command, Text.Offset)] = { @@ -124,14 +136,12 @@ } private val block_size = 1024 - - val empty: Node = new Node() } final class Node private( val header: Node.Header = Node.bad_header("Bad theory header"), val perspective: Node.Perspective_Command = - Node.Perspective(false, Command.Perspective.empty, Overlays.empty), + Node.Perspective(false, Command.Perspective.empty, Node.Overlays.empty), val commands: Linear_Set[Command] = Linear_Set.empty) { def clear: Node = new Node(header = header) @@ -310,19 +320,19 @@ val node: Node val is_outdated: Boolean def convert(i: Text.Offset): Text.Offset + def revert(i: Text.Offset): Text.Offset def convert(range: Text.Range): Text.Range - def revert(i: Text.Offset): Text.Offset def revert(range: Text.Range): Text.Range def eq_content(other: Snapshot): Boolean def cumulate_markup[A]( range: Text.Range, info: A, elements: Option[Set[String]], - result: Command.State => PartialFunction[(A, Text.Markup), A]): Stream[Text.Info[A]] + result: Command.State => (A, Text.Markup) => Option[A]): Stream[Text.Info[A]] def select_markup[A]( range: Text.Range, elements: Option[Set[String]], - result: Command.State => PartialFunction[Text.Markup, A]): Stream[Text.Info[A]] + result: Command.State => Text.Markup => Option[A]): Stream[Text.Info[A]] } type Assign_Update = @@ -554,33 +564,30 @@ }) def cumulate_markup[A](range: Text.Range, info: A, elements: Option[Set[String]], - result: Command.State => PartialFunction[(A, Text.Markup), A]): Stream[Text.Info[A]] = + result: Command.State => (A, Text.Markup) => Option[A]): Stream[Text.Info[A]] = { val former_range = revert(range) for { (command, command_start) <- node.command_range(former_range).toStream st = state.command_state(version, command) res = result(st) - Text.Info(r0, a) <- st.markup. - cumulate[A]((former_range - command_start).restrict(command.range), info, elements, - { - case (a, Text.Info(r0, b)) - if res.isDefinedAt((a, Text.Info(convert(r0 + command_start), b))) => - res((a, Text.Info(convert(r0 + command_start), b))) - }) + Text.Info(r0, a) <- st.markup.cumulate[A]( + (former_range - command_start).restrict(command.range), info, elements, + { case (a, Text.Info(r0, b)) => res(a, Text.Info(convert(r0 + command_start), b)) }) } yield Text.Info(convert(r0 + command_start), a) } def select_markup[A](range: Text.Range, elements: Option[Set[String]], - result: Command.State => PartialFunction[Text.Markup, A]): Stream[Text.Info[A]] = + result: Command.State => Text.Markup => Option[A]): Stream[Text.Info[A]] = { - def result1(st: Command.State) = + def result1(st: Command.State): (Option[A], Text.Markup) => Option[Option[A]] = { val res = result(st) - new PartialFunction[(Option[A], Text.Markup), Option[A]] { - def isDefinedAt(arg: (Option[A], Text.Markup)): Boolean = res.isDefinedAt(arg._2) - def apply(arg: (Option[A], Text.Markup)): Option[A] = Some(res(arg._2)) - } + (_: Option[A], x: Text.Markup) => + res(x) match { + case None => None + case some => Some(some) + } } for (Text.Info(r, Some(x)) <- cumulate_markup(range, None, elements, result1)) yield Text.Info(r, x) diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Pure/PIDE/markup_tree.scala --- a/src/Pure/PIDE/markup_tree.scala Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Pure/PIDE/markup_tree.scala Wed Aug 07 15:43:58 2013 +0200 @@ -223,7 +223,7 @@ 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]] = + result: (A, Text.Markup) => Option[A]): Stream[Text.Info[A]] = { val notable: Elements => Boolean = result_elements match { @@ -233,15 +233,12 @@ def results(x: A, entry: Entry): Option[A] = { - val (y, changed) = - // FIXME proper cumulation order (including status markup) (!?) - ((x, false) /: entry.rev_markup)((res, info) => - { - val (y, changed) = res - val arg = (y, Text.Info(entry.range, info)) - if (result.isDefinedAt(arg)) (result(arg), true) - else res - }) + var y = x + var changed = false + for { + info <- entry.rev_markup // FIXME proper cumulation order (including status markup) (!?) + y1 <- result(y, Text.Info(entry.range, info)) + } { y = y1; changed = true } if (changed) Some(y) else None } diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Pure/PIDE/xml.scala --- a/src/Pure/PIDE/xml.scala Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Pure/PIDE/xml.scala Wed Aug 07 15:43:58 2013 +0200 @@ -22,6 +22,9 @@ sealed abstract class Tree { override def toString = string_of_tree(this) } case class Elem(markup: Markup, body: List[Tree]) extends Tree + { + def name: String = markup.name + } case class Text(content: String) extends Tree def elem(name: String, body: List[Tree]) = Elem(Markup(name, Nil), body) diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Pure/Thy/thy_syntax.scala --- a/src/Pure/Thy/thy_syntax.scala Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Pure/Thy/thy_syntax.scala Wed Aug 07 15:43:58 2013 +0200 @@ -95,7 +95,7 @@ def command_perspective( node: Document.Node, perspective: Text.Perspective, - overlays: Document.Overlays): (Command.Perspective, Command.Perspective) = + overlays: Document.Node.Overlays): (Command.Perspective, Command.Perspective) = { if (perspective.is_empty && overlays.is_empty) (Command.Perspective.empty, Command.Perspective.empty) diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Pure/Tools/ml_statistics.scala --- a/src/Pure/Tools/ml_statistics.scala Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Pure/Tools/ml_statistics.scala Wed Aug 07 15:43:58 2013 +0200 @@ -85,7 +85,7 @@ // rising edges -- relative speed val speeds = for ((key, value) <- props; a <- Library.try_unprefix("time", key)) yield { - val (x0, y0, s0) = last_edge.get(a) getOrElse (0.0, 0.0, 0.0) + val (x0, y0, s0) = last_edge.getOrElse(a, (0.0, 0.0, 0.0)) val x1 = time val y1 = java.lang.Double.parseDouble(value) diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Tools/Adhoc_Overloading.thy --- a/src/Tools/Adhoc_Overloading.thy Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Tools/Adhoc_Overloading.thy Wed Aug 07 15:43:58 2013 +0200 @@ -2,7 +2,7 @@ Author: Christian Sternagel, University of Innsbruck *) -header {* Ad-hoc overloading of constants based on their types *} +header {* Adhoc overloading of constants based on their types *} theory Adhoc_Overloading imports Pure diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Tools/adhoc_overloading.ML --- a/src/Tools/adhoc_overloading.ML Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Tools/adhoc_overloading.ML Wed Aug 07 15:43:58 2013 +0200 @@ -1,7 +1,7 @@ (* Author: Alexander Krauss, TU Muenchen Author: Christian Sternagel, University of Innsbruck -Ad-hoc overloading of constants based on their types. +Adhoc overloading of constants based on their types. *) signature ADHOC_OVERLOADING = @@ -41,7 +41,7 @@ "in term " :: quote (Syntax.string_of_term ctxt' t) :: (if null instances then "no instances" else "multiple instances:") :: - map (Syntax.string_of_term ctxt') instances) + map (Markup.markup Markup.item o Syntax.string_of_term ctxt') instances) |> error end; @@ -227,12 +227,12 @@ val _ = Outer_Syntax.local_theory @{command_spec "adhoc_overloading"} - "add ad-hoc overloading for constants / fixed variables" + "add adhoc overloading for constants / fixed variables" (Parse.and_list1 (Parse.const -- Scan.repeat Parse.term) >> adhoc_overloading_cmd true); val _ = Outer_Syntax.local_theory @{command_spec "no_adhoc_overloading"} - "add ad-hoc overloading for constants / fixed variables" + "add adhoc overloading for constants / fixed variables" (Parse.and_list1 (Parse.const -- Scan.repeat Parse.term) >> adhoc_overloading_cmd false); end; diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Tools/jEdit/src/document_model.scala --- a/src/Tools/jEdit/src/document_model.scala Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Tools/jEdit/src/document_model.scala Wed Aug 07 15:43:58 2013 +0200 @@ -79,7 +79,7 @@ /* overlays */ - private var overlays = Document.Overlays.empty + private var overlays = Document.Node.Overlays.empty def insert_overlay(command: Command, name: String, args: List[String]): Unit = Swing_Thread.require { overlays = overlays.insert(command, (name, args)) } @@ -104,7 +104,7 @@ } val empty_perspective: Document.Node.Perspective_Text = - Document.Node.Perspective(false, Text.Perspective.empty, Document.Overlays.empty) + Document.Node.Perspective(false, Text.Perspective.empty, Document.Node.Overlays.empty) def node_perspective(): Document.Node.Perspective_Text = { diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Tools/jEdit/src/find_dockable.scala --- a/src/Tools/jEdit/src/find_dockable.scala Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Tools/jEdit/src/find_dockable.scala Wed Aug 07 15:43:58 2013 +0200 @@ -15,10 +15,10 @@ import scala.swing.event.ButtonClicked import java.awt.BorderLayout -import java.awt.event.{ComponentEvent, ComponentAdapter} +import java.awt.event.{ComponentEvent, ComponentAdapter, KeyEvent} import org.gjt.sp.jedit.View -import org.gjt.sp.jedit.gui.HistoryTextArea +import org.gjt.sp.jedit.gui.HistoryTextField class Find_Dockable(view: View, position: String) extends Dockable(view, position) @@ -91,17 +91,23 @@ /* controls */ - private val query = new HistoryTextArea("isabelle-find-theorems") { + private def clicked { find_theorems.apply_query(List(query.getText)) } + + private val query = new HistoryTextField("isabelle-find-theorems") { + override def processKeyEvent(evt: KeyEvent) + { + if (evt.getID == KeyEvent.KEY_PRESSED && evt.getKeyCode == KeyEvent.VK_ENTER) clicked + super.processKeyEvent(evt) + } { val max = getPreferredSize; max.width = Integer.MAX_VALUE; setMaximumSize(max) } - setColumns(25) - setRows(1) + setColumns(40) } private val query_wrapped = Component.wrap(query) private val apply_query = new Button("Apply") { tooltip = "Find theorems meeting specified criteria" - reactions += { case ButtonClicked(_) => find_theorems.apply_query(List(query.getText)) } + reactions += { case ButtonClicked(_) => clicked } } private val locate_query = new Button("Locate") { diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Tools/jEdit/src/query_operation.scala --- a/src/Tools/jEdit/src/query_operation.scala Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Tools/jEdit/src/query_operation.scala Wed Aug 07 15:43:58 2013 +0200 @@ -141,7 +141,7 @@ /* status */ def get_status(name: String, status: Status.Value): Option[Status.Value] = - results.collectFirst({ case List(XML.Elem(m, _)) if m.name == name => status }) + results.collectFirst({ case List(elem: XML.Elem) if elem.name == name => status }) val new_status = get_status(Markup.FINISHED, Status.FINISHED) orElse diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Tools/jEdit/src/rendering.scala --- a/src/Tools/jEdit/src/rendering.scala Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Tools/jEdit/src/rendering.scala Wed Aug 07 15:43:58 2013 +0200 @@ -154,11 +154,12 @@ snapshot.cumulate_markup[(Protocol.Status, Int)]( range, (Protocol.Status.init, 0), Some(overview_include), _ => { - case ((status, pri), Text.Info(_, XML.Elem(markup, _))) - if overview_include(markup.name) => - if (markup.name == Markup.WARNING || markup.name == Markup.ERROR) - (status, pri max Rendering.message_pri(markup.name)) - else (Protocol.command_status(status, markup), pri) + case ((status, pri), Text.Info(_, elem)) => + if (elem.name == Markup.WARNING || elem.name == Markup.ERROR) + Some((status, pri max Rendering.message_pri(elem.name))) + else if (overview_include(elem.name)) + Some((Protocol.command_status(status, elem.markup), pri)) + else None }) if (results.isEmpty) None else { @@ -188,8 +189,10 @@ { snapshot.select_markup(range, Some(highlight_include), _ => { - case Text.Info(info_range, XML.Elem(Markup(name, _), _)) if highlight_include(name) => - Text.Info(snapshot.convert(info_range), highlight_color) + case Text.Info(info_range, elem) => + if (highlight_include(elem.name)) + Some(Text.Info(snapshot.convert(info_range), highlight_color)) + else None }) match { case Text.Info(_, info) #:: _ => Some(info) case _ => None } } @@ -203,7 +206,8 @@ case (links, Text.Info(info_range, XML.Elem(Markup.Path(name), _))) if Path.is_ok(name) => val jedit_file = PIDE.thy_load.append(snapshot.node_name.dir, Path.explode(name)) - Text.Info(snapshot.convert(info_range), Hyperlink(jedit_file, 0, 0)) :: links + val link = Text.Info(snapshot.convert(info_range), Hyperlink(jedit_file, 0, 0)) + Some(link :: links) case (links, Text.Info(info_range, XML.Elem(Markup(Markup.ENTITY, props), _))) if !props.exists( @@ -216,8 +220,10 @@ Isabelle_System.source_file(Path.explode(name)) match { case Some(path) => val jedit_file = Isabelle_System.platform_path(path) - Text.Info(snapshot.convert(info_range), Hyperlink(jedit_file, line, 0)) :: links - case None => links + val link = + Text.Info(snapshot.convert(info_range), Hyperlink(jedit_file, line, 0)) + Some(link :: links) + case None => None } case Position.Def_Id_Offset(id, offset) => @@ -227,13 +233,17 @@ node.commands.iterator.takeWhile(_ != command).map(_.source) ++ Iterator.single(command.source(Text.Range(0, command.decode(offset)))) val (line, column) = ((1, 1) /: sources)(Symbol.advance_line_column) - Text.Info(snapshot.convert(info_range), - Hyperlink(command.node_name.node, line, column)) :: links - case None => links + val link = + Text.Info(snapshot.convert(info_range), + Hyperlink(command.node_name.node, line, column)) + Some(link :: links) + case None => None } - case _ => links + case _ => None } + + case _ => None }) match { case Text.Info(_, info :: _) #:: _ => Some(info) case _ => None } } @@ -246,10 +256,13 @@ { case Text.Info(info_range, elem @ Protocol.Dialog(_, serial, _)) if !command_state.results.defined(serial) => - Text.Info(snapshot.convert(info_range), elem) - case Text.Info(info_range, elem @ XML.Elem(Markup(name, _), _)) - if name == Markup.BROWSER || name == Markup.GRAPHVIEW || name == Markup.SENDBACK => - Text.Info(snapshot.convert(info_range), elem) + Some(Text.Info(snapshot.convert(info_range), elem)) + case Text.Info(info_range, elem) => + if (elem.name == Markup.BROWSER || + elem.name == Markup.GRAPHVIEW || + elem.name == Markup.SENDBACK) + Some(Text.Info(snapshot.convert(info_range), elem)) + else None }) match { case Text.Info(_, info) #:: _ => Some(info) case _ => None } @@ -257,7 +270,7 @@ { val results = snapshot.select_markup[Command.Results](range, None, command_state => - { case _ => command_state.results }).map(_.info) + { case _ => Some(command_state.results) }).map(_.info) (Command.Results.empty /: results)(_ ++ _) } @@ -272,12 +285,14 @@ if name == Markup.WRITELN || name == Markup.WARNING || name == Markup.ERROR => val entry: Command.Results.Entry = (serial -> XML.Elem(Markup(Markup.message(name), props), body)) - Text.Info(snapshot.convert(info_range), entry) :: msgs + Some(Text.Info(snapshot.convert(info_range), entry) :: msgs) case (msgs, Text.Info(info_range, msg @ XML.Elem(Markup(Markup.BAD, _), body))) if !body.isEmpty => val entry: Command.Results.Entry = (Document_ID.make(), msg) - Text.Info(snapshot.convert(info_range), entry) :: msgs + Some(Text.Info(snapshot.convert(info_range), entry) :: msgs) + + case _ => None }).toList.flatMap(_.info) if (results.isEmpty) None else { @@ -328,7 +343,7 @@ range, Text.Info(range, (Timing.zero, Nil)), Some(tooltip_elements), _ => { case (Text.Info(r, (t1, info)), Text.Info(_, XML.Elem(Markup.Timing(t2), _))) => - Text.Info(r, (t1 + t2, info)) + Some(Text.Info(r, (t1 + t2, info))) case (prev, Text.Info(r, XML.Elem(Markup.Entity(kind, name), _))) => val kind1 = space_explode('_', kind).mkString(" ") val txt1 = kind1 + " " + quote(name) @@ -337,19 +352,20 @@ if (kind == Markup.COMMAND && t.elapsed.seconds >= timing_threshold) "\n" + t.message else "" - add(prev, r, (true, XML.Text(txt1 + txt2))) + Some(add(prev, r, (true, XML.Text(txt1 + txt2)))) case (prev, Text.Info(r, XML.Elem(Markup.Path(name), _))) if Path.is_ok(name) => val jedit_file = PIDE.thy_load.append(snapshot.node_name.dir, Path.explode(name)) - add(prev, r, (true, XML.Text("file " + quote(jedit_file)))) + Some(add(prev, r, (true, XML.Text("file " + quote(jedit_file))))) case (prev, Text.Info(r, XML.Elem(Markup(name, _), body))) if name == Markup.SORTING || name == Markup.TYPING => - add(prev, r, (true, pretty_typing("::", body))) + Some(add(prev, r, (true, pretty_typing("::", body)))) case (prev, Text.Info(r, XML.Elem(Markup(Markup.ML_TYPING, _), body))) => - add(prev, r, (false, pretty_typing("ML:", body))) - case (prev, Text.Info(r, XML.Elem(Markup(name, _), _))) - if tooltips.isDefinedAt(name) => - add(prev, r, (true, XML.Text(tooltips(name)))) + Some(add(prev, r, (false, pretty_typing("ML:", body)))) + case (prev, Text.Info(r, XML.Elem(Markup(name, _), _))) => + if (tooltips.isDefinedAt(name)) + Some(add(prev, r, (true, XML.Text(tooltips(name))))) + else None }).toList.map(_.info) results.map(_.info).flatMap(_._2) match { @@ -383,15 +399,17 @@ { case (pri, Text.Info(_, XML.Elem(Markup(Markup.WRITELN, _), List(XML.Elem(Markup(Markup.INFORMATION, _), _))))) => - pri max Rendering.information_pri + Some(pri max Rendering.information_pri) case (pri, Text.Info(_, XML.Elem(Markup(Markup.WARNING, _), body))) => body match { case List(XML.Elem(Markup(Markup.LEGACY, _), _)) => - pri max Rendering.legacy_pri - case _ => pri max Rendering.warning_pri + Some(pri max Rendering.legacy_pri) + case _ => + Some(pri max Rendering.warning_pri) } case (pri, Text.Info(_, XML.Elem(Markup(Markup.ERROR, _), _))) => - pri max Rendering.error_pri + Some(pri max Rendering.error_pri) + case _ => None }) val pri = (0 /: results) { case (p1, Text.Info(_, p2)) => p1 max p2 } gutter_icons.get(pri) @@ -404,19 +422,19 @@ Rendering.warning_pri -> warning_color, Rendering.error_pri -> error_color) - private val squiggly_elements = Set(Markup.WRITELN, Markup.WARNING, Markup.ERROR) + private val squiggly_include = Set(Markup.WRITELN, Markup.WARNING, Markup.ERROR) def squiggly_underline(range: Text.Range): Stream[Text.Info[Color]] = { val results = - snapshot.cumulate_markup[Int](range, 0, Some(squiggly_elements), _ => + snapshot.cumulate_markup[Int](range, 0, Some(squiggly_include), _ => { - case (pri, Text.Info(_, msg @ XML.Elem(Markup(name, _), _))) - if name == Markup.WRITELN || - name == Markup.WARNING || - name == Markup.ERROR => - if (Protocol.is_information(msg)) pri max Rendering.information_pri - else pri max Rendering.message_pri(name) + case (pri, Text.Info(_, elem)) => + if (Protocol.is_information(elem)) + Some(pri max Rendering.information_pri) + else if (squiggly_include.contains(elem.name)) + Some(pri max Rendering.message_pri(elem.name)) + else None }) for { Text.Info(r, pri) <- results @@ -441,20 +459,22 @@ val results = snapshot.cumulate_markup[Int](range, 0, Some(line_background_elements), _ => { - case (pri, Text.Info(_, XML.Elem(Markup(Markup.INFORMATION, _), _))) => - pri max Rendering.information_pri - case (pri, Text.Info(_, XML.Elem(Markup(name, _), _))) - if name == Markup.WRITELN_MESSAGE || - name == Markup.TRACING_MESSAGE || - name == Markup.WARNING_MESSAGE || - name == Markup.ERROR_MESSAGE => pri max Rendering.message_pri(name) + case (pri, Text.Info(_, elem)) => + val name = elem.name + if (name == Markup.INFORMATION) + Some(pri max Rendering.information_pri) + else if (name == Markup.WRITELN_MESSAGE || name == Markup.TRACING_MESSAGE || + elem.name == Markup.WARNING_MESSAGE || name == Markup.ERROR_MESSAGE) + Some(pri max Rendering.message_pri(name)) + else None }) val pri = (0 /: results) { case (p1, Text.Info(_, p2)) => p1 max p2 } val is_separator = pri > 0 && snapshot.cumulate_markup[Boolean](range, false, Some(Set(Markup.SEPARATOR)), _ => { - case (_, Text.Info(_, XML.Elem(Markup(Markup.SEPARATOR, _), _))) => true + case (_, Text.Info(_, elem)) => + if (elem.name == Markup.SEPARATOR) Some(true) else None }).exists(_.info) message_colors.get(pri).map((_, is_separator)) @@ -483,20 +503,21 @@ { case (((Some(status), color), Text.Info(_, XML.Elem(markup, _)))) if (Protocol.command_status_markup(markup.name)) => - (Some(Protocol.command_status(status, markup)), color) + Some((Some(Protocol.command_status(status, markup)), color)) case (_, Text.Info(_, XML.Elem(Markup(Markup.BAD, _), _))) => - (None, Some(bad_color)) + Some((None, Some(bad_color))) case (_, Text.Info(_, XML.Elem(Markup(Markup.INTENSIFY, _), _))) => - (None, Some(intensify_color)) + Some((None, Some(intensify_color))) case (acc, Text.Info(_, Protocol.Dialog(_, serial, result))) => command_state.results.get(serial) match { case Some(Protocol.Dialog_Result(res)) if res == result => - (None, Some(active_result_color)) + Some((None, Some(active_result_color))) case _ => - (None, Some(active_color)) + Some((None, Some(active_color))) } - case (_, Text.Info(_, XML.Elem(markup, _))) if active_include(markup.name) => - (None, Some(active_color)) + case (_, Text.Info(_, elem)) => + if (active_include(elem.name)) Some((None, Some(active_color))) + else None }) color <- (result match { @@ -513,27 +534,29 @@ def background2(range: Text.Range): Stream[Text.Info[Color]] = snapshot.select_markup(range, Some(Set(Markup.TOKEN_RANGE)), _ => { - case Text.Info(_, XML.Elem(Markup(Markup.TOKEN_RANGE, _), _)) => light_color + case Text.Info(_, elem) => + if (elem.name == Markup.TOKEN_RANGE) Some(light_color) else None }) def bullet(range: Text.Range): Stream[Text.Info[Color]] = snapshot.select_markup(range, Some(Set(Markup.BULLET)), _ => { - case Text.Info(_, XML.Elem(Markup(Markup.BULLET, _), _)) => bullet_color + case Text.Info(_, elem) => + if (elem.name == Markup.BULLET) Some(bullet_color) else None }) - private val foreground_elements = + private val foreground_include = Set(Markup.STRING, Markup.ALTSTRING, Markup.VERBATIM, Markup.ANTIQ) def foreground(range: Text.Range): Stream[Text.Info[Color]] = - snapshot.select_markup(range, Some(foreground_elements), _ => + snapshot.select_markup(range, Some(foreground_include), _ => { - case Text.Info(_, XML.Elem(Markup(Markup.STRING, _), _)) => quoted_color - case Text.Info(_, XML.Elem(Markup(Markup.ALTSTRING, _), _)) => quoted_color - case Text.Info(_, XML.Elem(Markup(Markup.VERBATIM, _), _)) => quoted_color - case Text.Info(_, XML.Elem(Markup(Markup.ANTIQ, _), _)) => antiquoted_color + case Text.Info(_, elem) => + if (elem.name == Markup.ANTIQ) Some(antiquoted_color) + else if (foreground_include.contains(elem.name)) Some(quoted_color) + else None }) @@ -570,8 +593,7 @@ else snapshot.cumulate_markup(range, color, Some(text_color_elements), _ => { - case (_, Text.Info(_, XML.Elem(Markup(m, _), _))) - if text_colors.isDefinedAt(m) => text_colors(m) + case (_, Text.Info(_, elem)) => text_colors.get(elem.name) }) } @@ -583,7 +605,7 @@ def fold_depth(range: Text.Range): Stream[Text.Info[Int]] = snapshot.cumulate_markup[Int](range, 0, Some(fold_depth_include), _ => { - case (depth, Text.Info(_, XML.Elem(Markup(name, _), _))) - if fold_depth_include(name) => depth + 1 + case (depth, Text.Info(_, elem)) => + if (fold_depth_include(elem.name)) Some(depth + 1) else None }) } diff -r 34c47bc771f2 -r 3695ce0f9f96 src/Tools/jEdit/src/timing_dockable.scala --- a/src/Tools/jEdit/src/timing_dockable.scala Wed Aug 07 15:40:59 2013 +0200 +++ b/src/Tools/jEdit/src/timing_dockable.scala Wed Aug 07 15:43:58 2013 +0200 @@ -167,7 +167,7 @@ case None => Document.Node.Name.empty case Some(doc_view) => doc_view.model.name } - val timing = nodes_timing.get(name) getOrElse Protocol.empty_node_timing + val timing = nodes_timing.getOrElse(name, Protocol.empty_node_timing) val theories = (for ((node_name, node_timing) <- nodes_timing.toList if !node_timing.commands.isEmpty)