nipkow@8924: (* Title: HOL/SetInterval.thy nipkow@8924: ID: $Id$ ballarin@13735: Author: Tobias Nipkow and Clemens Ballarin paulson@14485: Additions by Jeremy Avigad in March 2004 paulson@8957: Copyright 2000 TU Muenchen nipkow@8924: ballarin@13735: lessThan, greaterThan, atLeast, atMost and two-sided intervals nipkow@8924: *) nipkow@8924: wenzelm@14577: header {* Set intervals *} wenzelm@14577: nipkow@15131: theory SetInterval haftmann@25919: imports Int nipkow@15131: begin nipkow@8924: nipkow@24691: context ord nipkow@24691: begin nipkow@24691: definition haftmann@25062: lessThan :: "'a => 'a set" ("(1{..<_})") where haftmann@25062: "{.. 'a set" ("(1{.._})") where haftmann@25062: "{..u} == {x. x \ u}" nipkow@24691: nipkow@24691: definition haftmann@25062: greaterThan :: "'a => 'a set" ("(1{_<..})") where haftmann@25062: "{l<..} == {x. l 'a set" ("(1{_..})") where haftmann@25062: "{l..} == {x. l\x}" nipkow@24691: nipkow@24691: definition haftmann@25062: greaterThanLessThan :: "'a => 'a => 'a set" ("(1{_<..<_})") where haftmann@25062: "{l<.. 'a => 'a set" ("(1{_..<_})") where haftmann@25062: "{l.. 'a => 'a set" ("(1{_<.._})") where haftmann@25062: "{l<..u} == {l<..} Int {..u}" nipkow@24691: nipkow@24691: definition haftmann@25062: atLeastAtMost :: "'a => 'a => 'a set" ("(1{_.._})") where haftmann@25062: "{l..u} == {l..} Int {..u}" nipkow@24691: nipkow@24691: end nipkow@8924: ballarin@13735: nipkow@15048: text{* A note of warning when using @{term"{.. nat => 'b set => 'b set" ("(3UN _<=_./ _)" 10) kleing@14418: "@UNION_less" :: "nat => nat => 'b set => 'b set" ("(3UN _<_./ _)" 10) kleing@14418: "@INTER_le" :: "nat => nat => 'b set => 'b set" ("(3INT _<=_./ _)" 10) kleing@14418: "@INTER_less" :: "nat => nat => 'b set => 'b set" ("(3INT _<_./ _)" 10) kleing@14418: kleing@14418: syntax (input) kleing@14418: "@UNION_le" :: "nat => nat => 'b set => 'b set" ("(3\ _\_./ _)" 10) kleing@14418: "@UNION_less" :: "nat => nat => 'b set => 'b set" ("(3\ _<_./ _)" 10) kleing@14418: "@INTER_le" :: "nat => nat => 'b set => 'b set" ("(3\ _\_./ _)" 10) kleing@14418: "@INTER_less" :: "nat => nat => 'b set => 'b set" ("(3\ _<_./ _)" 10) kleing@14418: kleing@14418: syntax (xsymbols) wenzelm@14846: "@UNION_le" :: "nat \ nat => 'b set => 'b set" ("(3\(00\<^bsub>_ \ _\<^esub>)/ _)" 10) wenzelm@14846: "@UNION_less" :: "nat \ nat => 'b set => 'b set" ("(3\(00\<^bsub>_ < _\<^esub>)/ _)" 10) wenzelm@14846: "@INTER_le" :: "nat \ nat => 'b set => 'b set" ("(3\(00\<^bsub>_ \ _\<^esub>)/ _)" 10) wenzelm@14846: "@INTER_less" :: "nat \ nat => 'b set => 'b set" ("(3\(00\<^bsub>_ < _\<^esub>)/ _)" 10) kleing@14418: kleing@14418: translations kleing@14418: "UN i<=n. A" == "UN i:{..n}. A" nipkow@15045: "UN i atLeast y) = (y \ (x::'a::order))" paulson@15418: by (blast intro: order_trans) paulson@13850: paulson@13850: lemma atLeast_eq_iff [iff]: paulson@15418: "(atLeast x = atLeast y) = (x = (y::'a::linorder))" paulson@13850: by (blast intro: order_antisym order_trans) paulson@13850: paulson@13850: lemma greaterThan_subset_iff [iff]: paulson@15418: "(greaterThan x \ greaterThan y) = (y \ (x::'a::linorder))" paulson@15418: apply (auto simp add: greaterThan_def) paulson@15418: apply (subst linorder_not_less [symmetric], blast) paulson@13850: done paulson@13850: paulson@13850: lemma greaterThan_eq_iff [iff]: paulson@15418: "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))" paulson@15418: apply (rule iffI) paulson@15418: apply (erule equalityE) paulson@15418: apply (simp_all add: greaterThan_subset_iff) paulson@13850: done paulson@13850: paulson@15418: lemma atMost_subset_iff [iff]: "(atMost x \ atMost y) = (x \ (y::'a::order))" paulson@13850: by (blast intro: order_trans) paulson@13850: paulson@15418: lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::linorder))" paulson@13850: by (blast intro: order_antisym order_trans) paulson@13850: paulson@13850: lemma lessThan_subset_iff [iff]: paulson@15418: "(lessThan x \ lessThan y) = (x \ (y::'a::linorder))" paulson@15418: apply (auto simp add: lessThan_def) paulson@15418: apply (subst linorder_not_less [symmetric], blast) paulson@13850: done paulson@13850: paulson@13850: lemma lessThan_eq_iff [iff]: paulson@15418: "(lessThan x = lessThan y) = (x = (y::'a::linorder))" paulson@15418: apply (rule iffI) paulson@15418: apply (erule equalityE) paulson@15418: apply (simp_all add: lessThan_subset_iff) ballarin@13735: done ballarin@13735: ballarin@13735: paulson@13850: subsection {*Two-sided intervals*} ballarin@13735: nipkow@24691: context ord nipkow@24691: begin nipkow@24691: paulson@24286: lemma greaterThanLessThan_iff [simp,noatp]: haftmann@25062: "(i : {l<.. {m..n} = {}"; nipkow@24691: by (auto simp add: atLeastAtMost_def atMost_def atLeast_def) nipkow@24691: haftmann@25062: lemma atLeastLessThan_empty[simp]: "n \ m ==> {m.. k ==> {k<..l} = {}" nipkow@17719: by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def) nipkow@17719: haftmann@25062: lemma greaterThanLessThan_empty[simp]:"l \ k ==> {k<..l} = {}" nipkow@17719: by(auto simp:greaterThanLessThan_def greaterThan_def lessThan_def) nipkow@17719: haftmann@25062: lemma atLeastAtMost_singleton [simp]: "{a..a} = {a}" nipkow@24691: by (auto simp add: atLeastAtMost_def atMost_def atLeast_def) nipkow@24691: nipkow@24691: end paulson@14485: paulson@14485: subsection {* Intervals of natural numbers *} paulson@14485: paulson@15047: subsubsection {* The Constant @{term lessThan} *} paulson@15047: paulson@14485: lemma lessThan_0 [simp]: "lessThan (0::nat) = {}" paulson@14485: by (simp add: lessThan_def) paulson@14485: paulson@14485: lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)" paulson@14485: by (simp add: lessThan_def less_Suc_eq, blast) paulson@14485: paulson@14485: lemma lessThan_Suc_atMost: "lessThan (Suc k) = atMost k" paulson@14485: by (simp add: lessThan_def atMost_def less_Suc_eq_le) paulson@14485: paulson@14485: lemma UN_lessThan_UNIV: "(UN m::nat. lessThan m) = UNIV" paulson@14485: by blast paulson@14485: paulson@15047: subsubsection {* The Constant @{term greaterThan} *} paulson@15047: paulson@14485: lemma greaterThan_0 [simp]: "greaterThan 0 = range Suc" paulson@14485: apply (simp add: greaterThan_def) paulson@14485: apply (blast dest: gr0_conv_Suc [THEN iffD1]) paulson@14485: done paulson@14485: paulson@14485: lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}" paulson@14485: apply (simp add: greaterThan_def) paulson@14485: apply (auto elim: linorder_neqE) paulson@14485: done paulson@14485: paulson@14485: lemma INT_greaterThan_UNIV: "(INT m::nat. greaterThan m) = {}" paulson@14485: by blast paulson@14485: paulson@15047: subsubsection {* The Constant @{term atLeast} *} paulson@15047: paulson@14485: lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV" paulson@14485: by (unfold atLeast_def UNIV_def, simp) paulson@14485: paulson@14485: lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}" paulson@14485: apply (simp add: atLeast_def) paulson@14485: apply (simp add: Suc_le_eq) paulson@14485: apply (simp add: order_le_less, blast) paulson@14485: done paulson@14485: paulson@14485: lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k" paulson@14485: by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le) paulson@14485: paulson@14485: lemma UN_atLeast_UNIV: "(UN m::nat. atLeast m) = UNIV" paulson@14485: by blast paulson@14485: paulson@15047: subsubsection {* The Constant @{term atMost} *} paulson@15047: paulson@14485: lemma atMost_0 [simp]: "atMost (0::nat) = {0}" paulson@14485: by (simp add: atMost_def) paulson@14485: paulson@14485: lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)" paulson@14485: apply (simp add: atMost_def) paulson@14485: apply (simp add: less_Suc_eq order_le_less, blast) paulson@14485: done paulson@14485: paulson@14485: lemma UN_atMost_UNIV: "(UN m::nat. atMost m) = UNIV" paulson@14485: by blast paulson@14485: paulson@15047: subsubsection {* The Constant @{term atLeastLessThan} *} paulson@15047: nipkow@28068: text{*The orientation of the following 2 rules is tricky. The lhs is nipkow@24449: defined in terms of the rhs. Hence the chosen orientation makes sense nipkow@24449: in this theory --- the reverse orientation complicates proofs (eg nipkow@24449: nontermination). But outside, when the definition of the lhs is rarely nipkow@24449: used, the opposite orientation seems preferable because it reduces a nipkow@24449: specific concept to a more general one. *} nipkow@28068: paulson@15047: lemma atLeast0LessThan: "{0::nat.. n then insert n {m.. Suc n \ {m..Suc n} = insert (Suc n) {m..n}" nipkow@15554: by (auto simp add: atLeastAtMost_def) nipkow@15554: nipkow@16733: subsubsection {* Image *} nipkow@16733: nipkow@16733: lemma image_add_atLeastAtMost: nipkow@16733: "(%n::nat. n+k) ` {i..j} = {i+k..j+k}" (is "?A = ?B") nipkow@16733: proof nipkow@16733: show "?A \ ?B" by auto nipkow@16733: next nipkow@16733: show "?B \ ?A" nipkow@16733: proof nipkow@16733: fix n assume a: "n : ?B" webertj@20217: hence "n - k : {i..j}" by auto nipkow@16733: moreover have "n = (n - k) + k" using a by auto nipkow@16733: ultimately show "n : ?A" by blast nipkow@16733: qed nipkow@16733: qed nipkow@16733: nipkow@16733: lemma image_add_atLeastLessThan: nipkow@16733: "(%n::nat. n+k) ` {i.. ?B" by auto nipkow@16733: next nipkow@16733: show "?B \ ?A" nipkow@16733: proof nipkow@16733: fix n assume a: "n : ?B" webertj@20217: hence "n - k : {i.. finite N" nipkow@28068: apply (rule finite_subset) nipkow@28068: apply (rule_tac [2] finite_lessThan, auto) nipkow@28068: done nipkow@28068: nipkow@28068: lemma finite_less_ub: nipkow@28068: "!!f::nat=>nat. (!!n. n \ f n) ==> finite {n. f n \ u}" nipkow@28068: by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans) paulson@14485: nipkow@24853: text{* Any subset of an interval of natural numbers the size of the nipkow@24853: subset is exactly that interval. *} nipkow@24853: nipkow@24853: lemma subset_card_intvl_is_intvl: nipkow@24853: "A <= {k.. A = {k.. \h. bij_betw h {0.. \h. bij_betw h M {0.. u ==> nipkow@15045: {(0::int).. u") paulson@14485: apply (subst image_atLeastZeroLessThan_int, assumption) paulson@14485: apply (rule finite_imageI) paulson@14485: apply auto paulson@14485: done paulson@14485: nipkow@15045: lemma finite_atLeastLessThan_int [iff]: "finite {l.. u") paulson@14485: apply (subst image_atLeastZeroLessThan_int, assumption) paulson@14485: apply (subst card_image) paulson@14485: apply (auto simp add: inj_on_def) paulson@14485: done paulson@14485: nipkow@15045: lemma card_atLeastLessThan_int [simp]: "card {l.. k < (i::nat)}" bulwahn@27656: proof - bulwahn@27656: have "{k. P k \ k < i} \ {.. M" bulwahn@27656: shows "card {k \ M. k < Suc i} \ 0" bulwahn@27656: proof - bulwahn@27656: from zero_in_M have "{k \ M. k < Suc i} \ {}" by auto bulwahn@27656: with finite_M_bounded_by_nat show ?thesis by (auto simp add: card_eq_0_iff) bulwahn@27656: qed bulwahn@27656: bulwahn@27656: lemma card_less_Suc2: "0 \ M \ card {k. Suc k \ M \ k < i} = card {k \ M. k < Suc i}" bulwahn@27656: apply (rule card_bij_eq [of "Suc" _ _ "\x. x - 1"]) bulwahn@27656: apply simp bulwahn@27656: apply fastsimp bulwahn@27656: apply auto bulwahn@27656: apply (rule inj_on_diff_nat) bulwahn@27656: apply auto bulwahn@27656: apply (case_tac x) bulwahn@27656: apply auto bulwahn@27656: apply (case_tac xa) bulwahn@27656: apply auto bulwahn@27656: apply (case_tac xa) bulwahn@27656: apply auto bulwahn@27656: apply (auto simp add: finite_M_bounded_by_nat) bulwahn@27656: done bulwahn@27656: bulwahn@27656: lemma card_less_Suc: bulwahn@27656: assumes zero_in_M: "0 \ M" bulwahn@27656: shows "Suc (card {k. Suc k \ M \ k < i}) = card {k \ M. k < Suc i}" bulwahn@27656: proof - bulwahn@27656: from assms have a: "0 \ {k \ M. k < Suc i}" by simp bulwahn@27656: hence c: "{k \ M. k < Suc i} = insert 0 ({k \ M. k < Suc i} - {0})" bulwahn@27656: by (auto simp only: insert_Diff) bulwahn@27656: have b: "{k \ M. k < Suc i} - {0} = {k \ M - {0}. k < Suc i}" by auto bulwahn@27656: from finite_M_bounded_by_nat[of "\x. x \ M" "Suc i"] have "Suc (card {k. Suc k \ M \ k < i}) = card (insert 0 ({k \ M. k < Suc i} - {0}))" bulwahn@27656: apply (subst card_insert) bulwahn@27656: apply simp_all bulwahn@27656: apply (subst b) bulwahn@27656: apply (subst card_less_Suc2[symmetric]) bulwahn@27656: apply simp_all bulwahn@27656: done bulwahn@27656: with c show ?thesis by simp bulwahn@27656: qed bulwahn@27656: paulson@14485: paulson@13850: subsection {*Lemmas useful with the summation operator setsum*} paulson@13850: ballarin@16102: text {* For examples, see Algebra/poly/UnivPoly2.thy *} ballarin@13735: wenzelm@14577: subsubsection {* Disjoint Unions *} ballarin@13735: wenzelm@14577: text {* Singletons and open intervals *} ballarin@13735: ballarin@13735: lemma ivl_disj_un_singleton: nipkow@15045: "{l::'a::linorder} Un {l<..} = {l..}" nipkow@15045: "{.. {l} Un {l<.. {l<.. {l} Un {l<..u} = {l..u}" nipkow@15045: "(l::'a::linorder) <= u ==> {l.. {..l} Un {l<.. {.. {..l} Un {l<..u} = {..u}" nipkow@15045: "(l::'a::linorder) <= u ==> {.. {l<..u} Un {u<..} = {l<..}" nipkow@15045: "(l::'a::linorder) < u ==> {l<.. {l..u} Un {u<..} = {l..}" nipkow@15045: "(l::'a::linorder) <= u ==> {l.. {l<.. {l<..m} Un {m<.. {l.. {l..m} Un {m<.. {l<.. {l<..m} Un {m<..u} = {l<..u}" nipkow@15045: "[| (l::'a::linorder) <= m; m <= u |] ==> {l.. {l..m} Un {m<..u} = {l..u}" ballarin@14398: by auto ballarin@13735: ballarin@13735: lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two ballarin@13735: wenzelm@14577: subsubsection {* Disjoint Intersections *} ballarin@13735: wenzelm@14577: text {* Singletons and open intervals *} ballarin@13735: ballarin@13735: lemma ivl_disj_int_singleton: nipkow@15045: "{l::'a::order} Int {l<..} = {}" nipkow@15045: "{.. n \ {i.. {m.. i | m \ i & j \ (n::'a::linorder))" nipkow@15542: apply(auto simp:linorder_not_le) nipkow@15542: apply(rule ccontr) nipkow@15542: apply(insert linorder_le_less_linear[of i n]) nipkow@15542: apply(clarsimp simp:linorder_not_le) nipkow@15542: apply(fastsimp) nipkow@15542: done nipkow@15542: nipkow@15041: nipkow@15042: subsection {* Summation indexed over intervals *} nipkow@15042: nipkow@15042: syntax nipkow@15042: "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10) nipkow@15048: "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10) nipkow@16052: "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<_./ _)" [0,0,10] 10) nipkow@16052: "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<=_./ _)" [0,0,10] 10) nipkow@15042: syntax (xsymbols) nipkow@15042: "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) nipkow@15048: "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) nipkow@16052: "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) nipkow@16052: "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) nipkow@15042: syntax (HTML output) nipkow@15042: "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) nipkow@15048: "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) nipkow@16052: "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) nipkow@16052: "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) nipkow@15056: syntax (latex_sum output) nipkow@15052: "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" nipkow@15052: ("(3\<^raw:$\sum_{>_ = _\<^raw:}^{>_\<^raw:}$> _)" [0,0,0,10] 10) nipkow@15052: "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" nipkow@15052: ("(3\<^raw:$\sum_{>_ = _\<^raw:}^{<>_\<^raw:}$> _)" [0,0,0,10] 10) nipkow@16052: "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" nipkow@16052: ("(3\<^raw:$\sum_{>_ < _\<^raw:}$> _)" [0,0,10] 10) nipkow@15052: "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" nipkow@16052: ("(3\<^raw:$\sum_{>_ \ _\<^raw:}$> _)" [0,0,10] 10) nipkow@15041: nipkow@15048: translations nipkow@15048: "\x=a..b. t" == "setsum (%x. t) {a..b}" nipkow@15048: "\x=a..i\n. t" == "setsum (\i. t) {..n}" nipkow@15048: "\ii. t) {..x\{a..b}. e"} & @{term"\x=a..b. e"} & @{term[mode=latex_sum]"\x=a..b. e"}\\ nipkow@15056: @{term[source]"\x\{a..x=a..x=a..x\{..b}. e"} & @{term"\x\b. e"} & @{term[mode=latex_sum]"\x\b. e"}\\ nipkow@15056: @{term[source]"\x\{..xxx::nat=0..xa = c; b = d; !!x. \ c \ x; x < d \ \ f x = g x \ \ nipkow@15542: setsum f {a..i \ Suc n. f i) = (\i \ n. f i) + f(Suc n)" nipkow@16052: by (simp add:atMost_Suc add_ac) nipkow@16052: nipkow@16041: lemma setsum_lessThan_Suc[simp]: "(\i < Suc n. f i) = (\i < n. f i) + f n" nipkow@16041: by (simp add:lessThan_Suc add_ac) nipkow@15041: nipkow@15911: lemma setsum_cl_ivl_Suc[simp]: nipkow@15561: "setsum f {m..Suc n} = (if Suc n < m then 0 else setsum f {m..n} + f(Suc n))" nipkow@15561: by (auto simp:add_ac atLeastAtMostSuc_conv) nipkow@15561: nipkow@15911: lemma setsum_op_ivl_Suc[simp]: nipkow@15561: "setsum f {m.. nipkow@15561: (\i=n..m+1. f i) = (\i=n..m. f i) + f(m + 1)" nipkow@15561: by (auto simp:add_ac atLeastAtMostSuc_conv) nipkow@16041: *) nipkow@28068: nipkow@28068: lemma setsum_head: nipkow@28068: fixes n :: nat nipkow@28068: assumes mn: "m <= n" nipkow@28068: shows "(\x\{m..n}. P x) = P m + (\x\{m<..n}. P x)" (is "?lhs = ?rhs") nipkow@28068: proof - nipkow@28068: from mn nipkow@28068: have "{m..n} = {m} \ {m<..n}" nipkow@28068: by (auto intro: ivl_disj_un_singleton) nipkow@28068: hence "?lhs = (\x\{m} \ {m<..n}. P x)" nipkow@28068: by (simp add: atLeast0LessThan) nipkow@28068: also have "\ = ?rhs" by simp nipkow@28068: finally show ?thesis . nipkow@28068: qed nipkow@28068: nipkow@28068: lemma setsum_head_Suc: nipkow@28068: "m \ n \ setsum f {m..n} = f m + setsum f {Suc m..n}" nipkow@28068: by (simp add: setsum_head atLeastSucAtMost_greaterThanAtMost) nipkow@28068: nipkow@28068: lemma setsum_head_upt_Suc: nipkow@28068: "m < n \ setsum f {m.. m \ n; n \ p \ \ nipkow@15539: setsum f {m.. 'a::ab_group_add" nipkow@15539: shows "\ m \ n; n \ p \ \ nipkow@15539: setsum f {m.. setsum f {Suc 0..k} = setsum f {0..k}" nipkow@28068: by(simp add:setsum_head_Suc) kleing@19106: nipkow@28068: lemma setsum_shift_lb_Suc0_0_upt: nipkow@28068: "f(0::nat) = 0 \ setsum f {Suc 0.. (\i=0..i\{1..n}. of_nat i) = kleing@19469: of_nat n*((of_nat n)+1)" kleing@19469: proof (induct n) kleing@19469: case 0 kleing@19469: show ?case by simp kleing@19469: next kleing@19469: case (Suc n) nipkow@23477: then show ?case by (simp add: ring_simps) kleing@19469: qed kleing@19469: kleing@19469: theorem arith_series_general: huffman@23277: "((1::'a::comm_semiring_1) + 1) * (\i\{.. 1" kleing@19469: let ?I = "\i. of_nat i" and ?n = "of_nat n" kleing@19469: have kleing@19469: "(\i\{..i\{..i\{.. = ?n*a + (\i\{.. = (?n*a + d*(\i\{1.. = (1+1)*?n*a + d*(1+1)*(\i\{1..i\{1..n - 1}. ?I i) = ((1+1)*?n*a + d*?I (n - 1)*?I n)" kleing@19469: by (simp only: mult_ac gauss_sum [of "n - 1"]) huffman@23431: (simp add: mult_ac trans [OF add_commute of_nat_Suc [symmetric]]) kleing@19469: finally show ?thesis by (simp add: mult_ac add_ac right_distrib) kleing@19469: next kleing@19469: assume "\(n > 1)" kleing@19469: hence "n = 1 \ n = 0" by auto kleing@19469: thus ?thesis by (auto simp: mult_ac right_distrib) kleing@19469: qed kleing@19469: kleing@19469: lemma arith_series_nat: kleing@19469: "Suc (Suc 0) * (\i\{..i\{..i\{..i\{..nat" kleing@19022: shows kleing@19022: "\x. Q x \ P x \ kleing@19022: (\xxxxx