(* Title: HOL/SetInterval.thy ID: $Id: SetInterval.thy,v 1.61 2008/02/21 16:33:59 nipkow Exp $ Author: Tobias Nipkow and Clemens Ballarin Additions by Jeremy Avigad in March 2004 Copyright 2000 TU Muenchen lessThan, greaterThan, atLeast, atMost and two-sided intervals *) header {* Set intervals *} theory SetInterval imports Int begin context ord begin definition lessThan :: "'a => 'a set" ("(1{..<_})") where "{..<u} == {x. x < u}" definition atMost :: "'a => 'a set" ("(1{.._})") where "{..u} == {x. x ≤ u}" definition greaterThan :: "'a => 'a set" ("(1{_<..})") where "{l<..} == {x. l<x}" definition atLeast :: "'a => 'a set" ("(1{_..})") where "{l..} == {x. l≤x}" definition greaterThanLessThan :: "'a => 'a => 'a set" ("(1{_<..<_})") where "{l<..<u} == {l<..} Int {..<u}" definition atLeastLessThan :: "'a => 'a => 'a set" ("(1{_..<_})") where "{l..<u} == {l..} Int {..<u}" definition greaterThanAtMost :: "'a => 'a => 'a set" ("(1{_<.._})") where "{l<..u} == {l<..} Int {..u}" definition atLeastAtMost :: "'a => 'a => 'a set" ("(1{_.._})") where "{l..u} == {l..} Int {..u}" end (* constdefs lessThan :: "('a::ord) => 'a set" ("(1{..<_})") "{..<u} == {x. x<u}" atMost :: "('a::ord) => 'a set" ("(1{.._})") "{..u} == {x. x<=u}" greaterThan :: "('a::ord) => 'a set" ("(1{_<..})") "{l<..} == {x. l<x}" atLeast :: "('a::ord) => 'a set" ("(1{_..})") "{l..} == {x. l<=x}" greaterThanLessThan :: "['a::ord, 'a] => 'a set" ("(1{_<..<_})") "{l<..<u} == {l<..} Int {..<u}" atLeastLessThan :: "['a::ord, 'a] => 'a set" ("(1{_..<_})") "{l..<u} == {l..} Int {..<u}" greaterThanAtMost :: "['a::ord, 'a] => 'a set" ("(1{_<.._})") "{l<..u} == {l<..} Int {..u}" atLeastAtMost :: "['a::ord, 'a] => 'a set" ("(1{_.._})") "{l..u} == {l..} Int {..u}" *) text{* A note of warning when using @{term"{..<n}"} on type @{typ nat}: it is equivalent to @{term"{0::nat..<n}"} but some lemmas involving @{term"{m..<n}"} may not exist in @{term"{..<n}"}-form as well. *} syntax "@UNION_le" :: "nat => nat => 'b set => 'b set" ("(3UN _<=_./ _)" 10) "@UNION_less" :: "nat => nat => 'b set => 'b set" ("(3UN _<_./ _)" 10) "@INTER_le" :: "nat => nat => 'b set => 'b set" ("(3INT _<=_./ _)" 10) "@INTER_less" :: "nat => nat => 'b set => 'b set" ("(3INT _<_./ _)" 10) syntax (input) "@UNION_le" :: "nat => nat => 'b set => 'b set" ("(3\<Union> _≤_./ _)" 10) "@UNION_less" :: "nat => nat => 'b set => 'b set" ("(3\<Union> _<_./ _)" 10) "@INTER_le" :: "nat => nat => 'b set => 'b set" ("(3\<Inter> _≤_./ _)" 10) "@INTER_less" :: "nat => nat => 'b set => 'b set" ("(3\<Inter> _<_./ _)" 10) syntax (xsymbols) "@UNION_le" :: "nat => nat => 'b set => 'b set" ("(3\<Union>(00_ ≤ _)/ _)" 10) "@UNION_less" :: "nat => nat => 'b set => 'b set" ("(3\<Union>(00_ < _)/ _)" 10) "@INTER_le" :: "nat => nat => 'b set => 'b set" ("(3\<Inter>(00_ ≤ _)/ _)" 10) "@INTER_less" :: "nat => nat => 'b set => 'b set" ("(3\<Inter>(00_ < _)/ _)" 10) translations "UN i<=n. A" == "UN i:{..n}. A" "UN i<n. A" == "UN i:{..<n}. A" "INT i<=n. A" == "INT i:{..n}. A" "INT i<n. A" == "INT i:{..<n}. A" subsection {* Various equivalences *} lemma (in ord) lessThan_iff [iff]: "(i: lessThan k) = (i<k)" by (simp add: lessThan_def) lemma Compl_lessThan [simp]: "!!k:: 'a::linorder. -lessThan k = atLeast k" apply (auto simp add: lessThan_def atLeast_def) done lemma single_Diff_lessThan [simp]: "!!k:: 'a::order. {k} - lessThan k = {k}" by auto lemma (in ord) greaterThan_iff [iff]: "(i: greaterThan k) = (k<i)" by (simp add: greaterThan_def) lemma Compl_greaterThan [simp]: "!!k:: 'a::linorder. -greaterThan k = atMost k" by (auto simp add: greaterThan_def atMost_def) lemma Compl_atMost [simp]: "!!k:: 'a::linorder. -atMost k = greaterThan k" apply (subst Compl_greaterThan [symmetric]) apply (rule double_complement) done lemma (in ord) atLeast_iff [iff]: "(i: atLeast k) = (k<=i)" by (simp add: atLeast_def) lemma Compl_atLeast [simp]: "!!k:: 'a::linorder. -atLeast k = lessThan k" by (auto simp add: lessThan_def atLeast_def) lemma (in ord) atMost_iff [iff]: "(i: atMost k) = (i<=k)" by (simp add: atMost_def) lemma atMost_Int_atLeast: "!!n:: 'a::order. atMost n Int atLeast n = {n}" by (blast intro: order_antisym) subsection {* Logical Equivalences for Set Inclusion and Equality *} lemma atLeast_subset_iff [iff]: "(atLeast x ⊆ atLeast y) = (y ≤ (x::'a::order))" by (blast intro: order_trans) lemma atLeast_eq_iff [iff]: "(atLeast x = atLeast y) = (x = (y::'a::linorder))" by (blast intro: order_antisym order_trans) lemma greaterThan_subset_iff [iff]: "(greaterThan x ⊆ greaterThan y) = (y ≤ (x::'a::linorder))" apply (auto simp add: greaterThan_def) apply (subst linorder_not_less [symmetric], blast) done lemma greaterThan_eq_iff [iff]: "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))" apply (rule iffI) apply (erule equalityE) apply (simp_all add: greaterThan_subset_iff) done lemma atMost_subset_iff [iff]: "(atMost x ⊆ atMost y) = (x ≤ (y::'a::order))" by (blast intro: order_trans) lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::linorder))" by (blast intro: order_antisym order_trans) lemma lessThan_subset_iff [iff]: "(lessThan x ⊆ lessThan y) = (x ≤ (y::'a::linorder))" apply (auto simp add: lessThan_def) apply (subst linorder_not_less [symmetric], blast) done lemma lessThan_eq_iff [iff]: "(lessThan x = lessThan y) = (x = (y::'a::linorder))" apply (rule iffI) apply (erule equalityE) apply (simp_all add: lessThan_subset_iff) done subsection {*Two-sided intervals*} context ord begin lemma greaterThanLessThan_iff [simp,noatp]: "(i : {l<..<u}) = (l < i & i < u)" by (simp add: greaterThanLessThan_def) lemma atLeastLessThan_iff [simp,noatp]: "(i : {l..<u}) = (l <= i & i < u)" by (simp add: atLeastLessThan_def) lemma greaterThanAtMost_iff [simp,noatp]: "(i : {l<..u}) = (l < i & i <= u)" by (simp add: greaterThanAtMost_def) lemma atLeastAtMost_iff [simp,noatp]: "(i : {l..u}) = (l <= i & i <= u)" by (simp add: atLeastAtMost_def) text {* The above four lemmas could be declared as iffs. If we do so, a call to blast in Hyperreal/Star.ML, lemma @{text STAR_Int} seems to take forever (more than one hour). *} end subsubsection{* Emptyness and singletons *} context order begin lemma atLeastAtMost_empty [simp]: "n < m ==> {m..n} = {}"; by (auto simp add: atLeastAtMost_def atMost_def atLeast_def) lemma atLeastLessThan_empty[simp]: "n ≤ m ==> {m..<n} = {}" by (auto simp add: atLeastLessThan_def) lemma greaterThanAtMost_empty[simp]:"l ≤ k ==> {k<..l} = {}" by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def) lemma greaterThanLessThan_empty[simp]:"l ≤ k ==> {k<..l} = {}" by(auto simp:greaterThanLessThan_def greaterThan_def lessThan_def) lemma atLeastAtMost_singleton [simp]: "{a..a} = {a}" by (auto simp add: atLeastAtMost_def atMost_def atLeast_def) end subsection {* Intervals of natural numbers *} subsubsection {* The Constant @{term lessThan} *} lemma lessThan_0 [simp]: "lessThan (0::nat) = {}" by (simp add: lessThan_def) lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)" by (simp add: lessThan_def less_Suc_eq, blast) lemma lessThan_Suc_atMost: "lessThan (Suc k) = atMost k" by (simp add: lessThan_def atMost_def less_Suc_eq_le) lemma UN_lessThan_UNIV: "(UN m::nat. lessThan m) = UNIV" by blast subsubsection {* The Constant @{term greaterThan} *} lemma greaterThan_0 [simp]: "greaterThan 0 = range Suc" apply (simp add: greaterThan_def) apply (blast dest: gr0_conv_Suc [THEN iffD1]) done lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}" apply (simp add: greaterThan_def) apply (auto elim: linorder_neqE) done lemma INT_greaterThan_UNIV: "(INT m::nat. greaterThan m) = {}" by blast subsubsection {* The Constant @{term atLeast} *} lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV" by (unfold atLeast_def UNIV_def, simp) lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}" apply (simp add: atLeast_def) apply (simp add: Suc_le_eq) apply (simp add: order_le_less, blast) done lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k" by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le) lemma UN_atLeast_UNIV: "(UN m::nat. atLeast m) = UNIV" by blast subsubsection {* The Constant @{term atMost} *} lemma atMost_0 [simp]: "atMost (0::nat) = {0}" by (simp add: atMost_def) lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)" apply (simp add: atMost_def) apply (simp add: less_Suc_eq order_le_less, blast) done lemma UN_atMost_UNIV: "(UN m::nat. atMost m) = UNIV" by blast subsubsection {* The Constant @{term atLeastLessThan} *} text{*The orientation of the following rule is tricky. The lhs is defined in terms of the rhs. Hence the chosen orientation makes sense in this theory --- the reverse orientation complicates proofs (eg nontermination). But outside, when the definition of the lhs is rarely used, the opposite orientation seems preferable because it reduces a specific concept to a more general one. *} lemma atLeast0LessThan: "{0::nat..<n} = {..<n}" by(simp add:lessThan_def atLeastLessThan_def) declare atLeast0LessThan[symmetric, code unfold] lemma atLeastLessThan0: "{m..<0::nat} = {}" by (simp add: atLeastLessThan_def) subsubsection {* Intervals of nats with @{term Suc} *} text{*Not a simprule because the RHS is too messy.*} lemma atLeastLessThanSuc: "{m..<Suc n} = (if m ≤ n then insert n {m..<n} else {})" by (auto simp add: atLeastLessThan_def) lemma atLeastLessThan_singleton [simp]: "{m..<Suc m} = {m}" by (auto simp add: atLeastLessThan_def) (* lemma atLeast_sum_LessThan [simp]: "{m + k..<k::nat} = {}" by (induct k, simp_all add: atLeastLessThanSuc) lemma atLeastSucLessThan [simp]: "{Suc n..<n} = {}" by (auto simp add: atLeastLessThan_def) *) lemma atLeastLessThanSuc_atLeastAtMost: "{l..<Suc u} = {l..u}" by (simp add: lessThan_Suc_atMost atLeastAtMost_def atLeastLessThan_def) lemma atLeastSucAtMost_greaterThanAtMost: "{Suc l..u} = {l<..u}" by (simp add: atLeast_Suc_greaterThan atLeastAtMost_def greaterThanAtMost_def) lemma atLeastSucLessThan_greaterThanLessThan: "{Suc l..<u} = {l<..<u}" by (simp add: atLeast_Suc_greaterThan atLeastLessThan_def greaterThanLessThan_def) lemma atLeastAtMostSuc_conv: "m ≤ Suc n ==> {m..Suc n} = insert (Suc n) {m..n}" by (auto simp add: atLeastAtMost_def) subsubsection {* Image *} lemma image_add_atLeastAtMost: "(%n::nat. n+k) ` {i..j} = {i+k..j+k}" (is "?A = ?B") proof show "?A ⊆ ?B" by auto next show "?B ⊆ ?A" proof fix n assume a: "n : ?B" hence "n - k : {i..j}" by auto moreover have "n = (n - k) + k" using a by auto ultimately show "n : ?A" by blast qed qed lemma image_add_atLeastLessThan: "(%n::nat. n+k) ` {i..<j} = {i+k..<j+k}" (is "?A = ?B") proof show "?A ⊆ ?B" by auto next show "?B ⊆ ?A" proof fix n assume a: "n : ?B" hence "n - k : {i..<j}" by auto moreover have "n = (n - k) + k" using a by auto ultimately show "n : ?A" by blast qed qed corollary image_Suc_atLeastAtMost[simp]: "Suc ` {i..j} = {Suc i..Suc j}" using image_add_atLeastAtMost[where k=1] by simp corollary image_Suc_atLeastLessThan[simp]: "Suc ` {i..<j} = {Suc i..<Suc j}" using image_add_atLeastLessThan[where k=1] by simp lemma image_add_int_atLeastLessThan: "(%x. x + (l::int)) ` {0..<u-l} = {l..<u}" apply (auto simp add: image_def) apply (rule_tac x = "x - l" in bexI) apply auto done subsubsection {* Finiteness *} lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..<k}" by (induct k) (simp_all add: lessThan_Suc) lemma finite_atMost [iff]: fixes k :: nat shows "finite {..k}" by (induct k) (simp_all add: atMost_Suc) lemma finite_greaterThanLessThan [iff]: fixes l :: nat shows "finite {l<..<u}" by (simp add: greaterThanLessThan_def) lemma finite_atLeastLessThan [iff]: fixes l :: nat shows "finite {l..<u}" by (simp add: atLeastLessThan_def) lemma finite_greaterThanAtMost [iff]: fixes l :: nat shows "finite {l<..u}" by (simp add: greaterThanAtMost_def) lemma finite_atLeastAtMost [iff]: fixes l :: nat shows "finite {l..u}" by (simp add: atLeastAtMost_def) lemma bounded_nat_set_is_finite: "(ALL i:N. i < (n::nat)) ==> finite N" -- {* A bounded set of natural numbers is finite. *} apply (rule finite_subset) apply (rule_tac [2] finite_lessThan, auto) done text{* Any subset of an interval of natural numbers the size of the subset is exactly that interval. *} lemma subset_card_intvl_is_intvl: "A <= {k..<k+card A} ==> A = {k..<k+card A}" (is "PROP ?P") proof cases assume "finite A" thus "PROP ?P" proof(induct A rule:finite_linorder_induct) case empty thus ?case by auto next case (insert A b) moreover hence "b ~: A" by auto moreover have "A <= {k..<k+card A}" and "b = k+card A" using `b ~: A` insert by fastsimp+ ultimately show ?case by auto qed next assume "~finite A" thus "PROP ?P" by simp qed subsubsection {* Cardinality *} lemma card_lessThan [simp]: "card {..<u} = u" by (induct u, simp_all add: lessThan_Suc) lemma card_atMost [simp]: "card {..u} = Suc u" by (simp add: lessThan_Suc_atMost [THEN sym]) lemma card_atLeastLessThan [simp]: "card {l..<u} = u - l" apply (subgoal_tac "card {l..<u} = card {..<u-l}") apply (erule ssubst, rule card_lessThan) apply (subgoal_tac "(%x. x + l) ` {..<u-l} = {l..<u}") apply (erule subst) apply (rule card_image) apply (simp add: inj_on_def) apply (auto simp add: image_def atLeastLessThan_def lessThan_def) apply (rule_tac x = "x - l" in exI) apply arith done lemma card_atLeastAtMost [simp]: "card {l..u} = Suc u - l" by (subst atLeastLessThanSuc_atLeastAtMost [THEN sym], simp) lemma card_greaterThanAtMost [simp]: "card {l<..u} = u - l" by (subst atLeastSucAtMost_greaterThanAtMost [THEN sym], simp) lemma card_greaterThanLessThan [simp]: "card {l<..<u} = u - Suc l" by (subst atLeastSucLessThan_greaterThanLessThan [THEN sym], simp) lemma ex_bij_betw_nat_finite: "finite M ==> ∃h. bij_betw h {0..<card M} M" apply(drule finite_imp_nat_seg_image_inj_on) apply(auto simp:atLeast0LessThan[symmetric] lessThan_def[symmetric] card_image bij_betw_def) done lemma ex_bij_betw_finite_nat: "finite M ==> ∃h. bij_betw h M {0..<card M}" by (blast dest: ex_bij_betw_nat_finite bij_betw_inv) subsection {* Intervals of integers *} lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l..<u+1} = {l..(u::int)}" by (auto simp add: atLeastAtMost_def atLeastLessThan_def) lemma atLeastPlusOneAtMost_greaterThanAtMost_int: "{l+1..u} = {l<..(u::int)}" by (auto simp add: atLeastAtMost_def greaterThanAtMost_def) lemma atLeastPlusOneLessThan_greaterThanLessThan_int: "{l+1..<u} = {l<..<u::int}" by (auto simp add: atLeastLessThan_def greaterThanLessThan_def) subsubsection {* Finiteness *} lemma image_atLeastZeroLessThan_int: "0 ≤ u ==> {(0::int)..<u} = int ` {..<nat u}" apply (unfold image_def lessThan_def) apply auto apply (rule_tac x = "nat x" in exI) apply (auto simp add: zless_nat_conj zless_nat_eq_int_zless [THEN sym]) done lemma finite_atLeastZeroLessThan_int: "finite {(0::int)..<u}" apply (case_tac "0 ≤ u") apply (subst image_atLeastZeroLessThan_int, assumption) apply (rule finite_imageI) apply auto done lemma finite_atLeastLessThan_int [iff]: "finite {l..<u::int}" apply (subgoal_tac "(%x. x + l) ` {0..<u-l} = {l..<u}") apply (erule subst) apply (rule finite_imageI) apply (rule finite_atLeastZeroLessThan_int) apply (rule image_add_int_atLeastLessThan) done lemma finite_atLeastAtMost_int [iff]: "finite {l..(u::int)}" by (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym], simp) lemma finite_greaterThanAtMost_int [iff]: "finite {l<..(u::int)}" by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp) lemma finite_greaterThanLessThan_int [iff]: "finite {l<..<u::int}" by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp) subsubsection {* Cardinality *} lemma card_atLeastZeroLessThan_int: "card {(0::int)..<u} = nat u" apply (case_tac "0 ≤ u") apply (subst image_atLeastZeroLessThan_int, assumption) apply (subst card_image) apply (auto simp add: inj_on_def) done lemma card_atLeastLessThan_int [simp]: "card {l..<u} = nat (u - l)" apply (subgoal_tac "card {l..<u} = card {0..<u-l}") apply (erule ssubst, rule card_atLeastZeroLessThan_int) apply (subgoal_tac "(%x. x + l) ` {0..<u-l} = {l..<u}") apply (erule subst) apply (rule card_image) apply (simp add: inj_on_def) apply (rule image_add_int_atLeastLessThan) done lemma card_atLeastAtMost_int [simp]: "card {l..u} = nat (u - l + 1)" apply (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym]) apply (auto simp add: compare_rls) done lemma card_greaterThanAtMost_int [simp]: "card {l<..u} = nat (u - l)" by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp) lemma card_greaterThanLessThan_int [simp]: "card {l<..<u} = nat (u - (l + 1))" by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp) subsection {*Lemmas useful with the summation operator setsum*} text {* For examples, see Algebra/poly/UnivPoly2.thy *} subsubsection {* Disjoint Unions *} text {* Singletons and open intervals *} lemma ivl_disj_un_singleton: "{l::'a::linorder} Un {l<..} = {l..}" "{..<u} Un {u::'a::linorder} = {..u}" "(l::'a::linorder) < u ==> {l} Un {l<..<u} = {l..<u}" "(l::'a::linorder) < u ==> {l<..<u} Un {u} = {l<..u}" "(l::'a::linorder) <= u ==> {l} Un {l<..u} = {l..u}" "(l::'a::linorder) <= u ==> {l..<u} Un {u} = {l..u}" by auto text {* One- and two-sided intervals *} lemma ivl_disj_un_one: "(l::'a::linorder) < u ==> {..l} Un {l<..<u} = {..<u}" "(l::'a::linorder) <= u ==> {..<l} Un {l..<u} = {..<u}" "(l::'a::linorder) <= u ==> {..l} Un {l<..u} = {..u}" "(l::'a::linorder) <= u ==> {..<l} Un {l..u} = {..u}" "(l::'a::linorder) <= u ==> {l<..u} Un {u<..} = {l<..}" "(l::'a::linorder) < u ==> {l<..<u} Un {u..} = {l<..}" "(l::'a::linorder) <= u ==> {l..u} Un {u<..} = {l..}" "(l::'a::linorder) <= u ==> {l..<u} Un {u..} = {l..}" by auto text {* Two- and two-sided intervals *} lemma ivl_disj_un_two: "[| (l::'a::linorder) < m; m <= u |] ==> {l<..<m} Un {m..<u} = {l<..<u}" "[| (l::'a::linorder) <= m; m < u |] ==> {l<..m} Un {m<..<u} = {l<..<u}" "[| (l::'a::linorder) <= m; m <= u |] ==> {l..<m} Un {m..<u} = {l..<u}" "[| (l::'a::linorder) <= m; m < u |] ==> {l..m} Un {m<..<u} = {l..<u}" "[| (l::'a::linorder) < m; m <= u |] ==> {l<..<m} Un {m..u} = {l<..u}" "[| (l::'a::linorder) <= m; m <= u |] ==> {l<..m} Un {m<..u} = {l<..u}" "[| (l::'a::linorder) <= m; m <= u |] ==> {l..<m} Un {m..u} = {l..u}" "[| (l::'a::linorder) <= m; m <= u |] ==> {l..m} Un {m<..u} = {l..u}" by auto lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two subsubsection {* Disjoint Intersections *} text {* Singletons and open intervals *} lemma ivl_disj_int_singleton: "{l::'a::order} Int {l<..} = {}" "{..<u} Int {u} = {}" "{l} Int {l<..<u} = {}" "{l<..<u} Int {u} = {}" "{l} Int {l<..u} = {}" "{l..<u} Int {u} = {}" by simp+ text {* One- and two-sided intervals *} lemma ivl_disj_int_one: "{..l::'a::order} Int {l<..<u} = {}" "{..<l} Int {l..<u} = {}" "{..l} Int {l<..u} = {}" "{..<l} Int {l..u} = {}" "{l<..u} Int {u<..} = {}" "{l<..<u} Int {u..} = {}" "{l..u} Int {u<..} = {}" "{l..<u} Int {u..} = {}" by auto text {* Two- and two-sided intervals *} lemma ivl_disj_int_two: "{l::'a::order<..<m} Int {m..<u} = {}" "{l<..m} Int {m<..<u} = {}" "{l..<m} Int {m..<u} = {}" "{l..m} Int {m<..<u} = {}" "{l<..<m} Int {m..u} = {}" "{l<..m} Int {m<..u} = {}" "{l..<m} Int {m..u} = {}" "{l..m} Int {m<..u} = {}" by auto lemmas ivl_disj_int = ivl_disj_int_singleton ivl_disj_int_one ivl_disj_int_two subsubsection {* Some Differences *} lemma ivl_diff[simp]: "i ≤ n ==> {i..<m} - {i..<n} = {n..<(m::'a::linorder)}" by(auto) subsubsection {* Some Subset Conditions *} lemma ivl_subset [simp,noatp]: "({i..<j} ⊆ {m..<n}) = (j ≤ i | m ≤ i & j ≤ (n::'a::linorder))" apply(auto simp:linorder_not_le) apply(rule ccontr) apply(insert linorder_le_less_linear[of i n]) apply(clarsimp simp:linorder_not_le) apply(fastsimp) done subsection {* Summation indexed over intervals *} syntax "_from_to_setsum" :: "idt => 'a => 'a => 'b => 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_setsum" :: "idt => 'a => 'a => 'b => 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10) "_upt_setsum" :: "idt => 'a => 'b => 'b" ("(SUM _<_./ _)" [0,0,10] 10) "_upto_setsum" :: "idt => 'a => 'b => 'b" ("(SUM _<=_./ _)" [0,0,10] 10) syntax (xsymbols) "_from_to_setsum" :: "idt => 'a => 'a => 'b => 'b" ("(3∑_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_setsum" :: "idt => 'a => 'a => 'b => 'b" ("(3∑_ = _..<_./ _)" [0,0,0,10] 10) "_upt_setsum" :: "idt => 'a => 'b => 'b" ("(3∑_<_./ _)" [0,0,10] 10) "_upto_setsum" :: "idt => 'a => 'b => 'b" ("(3∑_≤_./ _)" [0,0,10] 10) syntax (HTML output) "_from_to_setsum" :: "idt => 'a => 'a => 'b => 'b" ("(3∑_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_setsum" :: "idt => 'a => 'a => 'b => 'b" ("(3∑_ = _..<_./ _)" [0,0,0,10] 10) "_upt_setsum" :: "idt => 'a => 'b => 'b" ("(3∑_<_./ _)" [0,0,10] 10) "_upto_setsum" :: "idt => 'a => 'b => 'b" ("(3∑_≤_./ _)" [0,0,10] 10) syntax (latex_sum output) "_from_to_setsum" :: "idt => 'a => 'a => 'b => 'b" ("(3$\sum_{_ = _}^{_}$ _)" [0,0,0,10] 10) "_from_upto_setsum" :: "idt => 'a => 'a => 'b => 'b" ("(3$\sum_{_ = _}^{<_}$ _)" [0,0,0,10] 10) "_upt_setsum" :: "idt => 'a => 'b => 'b" ("(3$\sum_{_ < _}$ _)" [0,0,10] 10) "_upto_setsum" :: "idt => 'a => 'b => 'b" ("(3$\sum_{_ ≤ _}$ _)" [0,0,10] 10) translations "∑x=a..b. t" == "setsum (%x. t) {a..b}" "∑x=a..<b. t" == "setsum (%x. t) {a..<b}" "∑i≤n. t" == "setsum (λi. t) {..n}" "∑i<n. t" == "setsum (λi. t) {..<n}" text{* The above introduces some pretty alternative syntaxes for summation over intervals: \begin{center} \begin{tabular}{lll} Old & New & \LaTeX\\ @{term[source]"∑x∈{a..b}. e"} & @{term"∑x=a..b. e"} & @{term[mode=latex_sum]"∑x=a..b. e"}\\ @{term[source]"∑x∈{a..<b}. e"} & @{term"∑x=a..<b. e"} & @{term[mode=latex_sum]"∑x=a..<b. e"}\\ @{term[source]"∑x∈{..b}. e"} & @{term"∑x≤b. e"} & @{term[mode=latex_sum]"∑x≤b. e"}\\ @{term[source]"∑x∈{..<b}. e"} & @{term"∑x<b. e"} & @{term[mode=latex_sum]"∑x<b. e"} \end{tabular} \end{center} The left column shows the term before introduction of the new syntax, the middle column shows the new (default) syntax, and the right column shows a special syntax. The latter is only meaningful for latex output and has to be activated explicitly by setting the print mode to @{text latex_sum} (e.g.\ via @{text "mode = latex_sum"} in antiquotations). It is not the default \LaTeX\ output because it only works well with italic-style formulae, not tt-style. Note that for uniformity on @{typ nat} it is better to use @{term"∑x::nat=0..<n. e"} rather than @{text"∑x<n. e"}: @{text setsum} may not provide all lemmas available for @{term"{m..<n}"} also in the special form for @{term"{..<n}"}. *} text{* This congruence rule should be used for sums over intervals as the standard theorem @{text[source]setsum_cong} does not work well with the simplifier who adds the unsimplified premise @{term"x:B"} to the context. *} lemma setsum_ivl_cong: "[|a = c; b = d; !!x. [| c ≤ x; x < d |] ==> f x = g x |] ==> setsum f {a..<b} = setsum g {c..<d}" by(rule setsum_cong, simp_all) (* FIXME why are the following simp rules but the corresponding eqns on intervals are not? *) lemma setsum_atMost_Suc[simp]: "(∑i ≤ Suc n. f i) = (∑i ≤ n. f i) + f(Suc n)" by (simp add:atMost_Suc add_ac) lemma setsum_lessThan_Suc[simp]: "(∑i < Suc n. f i) = (∑i < n. f i) + f n" by (simp add:lessThan_Suc add_ac) lemma setsum_cl_ivl_Suc[simp]: "setsum f {m..Suc n} = (if Suc n < m then 0 else setsum f {m..n} + f(Suc n))" by (auto simp:add_ac atLeastAtMostSuc_conv) lemma setsum_op_ivl_Suc[simp]: "setsum f {m..<Suc n} = (if n < m then 0 else setsum f {m..<n} + f(n))" by (auto simp:add_ac atLeastLessThanSuc) (* lemma setsum_cl_ivl_add_one_nat: "(n::nat) <= m + 1 ==> (∑i=n..m+1. f i) = (∑i=n..m. f i) + f(m + 1)" by (auto simp:add_ac atLeastAtMostSuc_conv) *) lemma setsum_add_nat_ivl: "[| m ≤ n; n ≤ p |] ==> setsum f {m..<n} + setsum f {n..<p} = setsum f {m..<p::nat}" by (simp add:setsum_Un_disjoint[symmetric] ivl_disj_int ivl_disj_un) lemma setsum_diff_nat_ivl: fixes f :: "nat => 'a::ab_group_add" shows "[| m ≤ n; n ≤ p |] ==> setsum f {m..<p} - setsum f {m..<n} = setsum f {n..<p}" using setsum_add_nat_ivl [of m n p f,symmetric] apply (simp add: add_ac) done subsection{* Shifting bounds *} lemma setsum_shift_bounds_nat_ivl: "setsum f {m+k..<n+k} = setsum (%i. f(i + k)){m..<n::nat}" by (induct "n", auto simp:atLeastLessThanSuc) lemma setsum_shift_bounds_cl_nat_ivl: "setsum f {m+k..n+k} = setsum (%i. f(i + k)){m..n::nat}" apply (insert setsum_reindex[OF inj_on_add_nat, where h=f and B = "{m..n}"]) apply (simp add:image_add_atLeastAtMost o_def) done corollary setsum_shift_bounds_cl_Suc_ivl: "setsum f {Suc m..Suc n} = setsum (%i. f(Suc i)){m..n}" by (simp add:setsum_shift_bounds_cl_nat_ivl[where k=1,simplified]) corollary setsum_shift_bounds_Suc_ivl: "setsum f {Suc m..<Suc n} = setsum (%i. f(Suc i)){m..<n}" by (simp add:setsum_shift_bounds_nat_ivl[where k=1,simplified]) lemma setsum_head: fixes n :: nat assumes mn: "m <= n" shows "(∑x∈{m..n}. P x) = P m + (∑x∈{m<..n}. P x)" (is "?lhs = ?rhs") proof - from mn have "{m..n} = {m} ∪ {m<..n}" by (auto intro: ivl_disj_un_singleton) hence "?lhs = (∑x∈{m} ∪ {m<..n}. P x)" by (simp add: atLeast0LessThan) also have "… = ?rhs" by simp finally show ?thesis . qed lemma setsum_head_upt: fixes m::nat assumes m: "0 < m" shows "(∑x<m. P x) = P 0 + (∑x∈{1..<m}. P x)" proof - have "(∑x<m. P x) = (∑x∈{0..<m}. P x)" by (simp add: atLeast0LessThan) also from m have "… = (∑x∈{0..m - 1}. P x)" by (cases m) (auto simp add: atLeastLessThanSuc_atLeastAtMost) also have "… = P 0 + (∑x∈{0<..m - 1}. P x)" by (simp add: setsum_head) also from m have "{0<..m - 1} = {1..<m}" by (cases m) (auto simp add: atLeastLessThanSuc_atLeastAtMost) finally show ?thesis . qed subsection {* The formula for geometric sums *} lemma geometric_sum: "x ~= 1 ==> (∑i=0..<n. x ^ i) = (x ^ n - 1) / (x - 1::'a::{field, recpower})" by (induct "n") (simp_all add:field_simps power_Suc) subsection {* The formula for arithmetic sums *} lemma gauss_sum: "((1::'a::comm_semiring_1) + 1)*(∑i∈{1..n}. of_nat i) = of_nat n*((of_nat n)+1)" proof (induct n) case 0 show ?case by simp next case (Suc n) then show ?case by (simp add: ring_simps) qed theorem arith_series_general: "((1::'a::comm_semiring_1) + 1) * (∑i∈{..<n}. a + of_nat i * d) = of_nat n * (a + (a + of_nat(n - 1)*d))" proof cases assume ngt1: "n > 1" let ?I = "λi. of_nat i" and ?n = "of_nat n" have "(∑i∈{..<n}. a+?I i*d) = ((∑i∈{..<n}. a) + (∑i∈{..<n}. ?I i*d))" by (rule setsum_addf) also from ngt1 have "… = ?n*a + (∑i∈{..<n}. ?I i*d)" by simp also from ngt1 have "… = (?n*a + d*(∑i∈{1..<n}. ?I i))" by (simp add: setsum_right_distrib setsum_head_upt mult_ac) also have "(1+1)*… = (1+1)*?n*a + d*(1+1)*(∑i∈{1..<n}. ?I i)" by (simp add: left_distrib right_distrib) also from ngt1 have "{1..<n} = {1..n - 1}" by (cases n) (auto simp: atLeastLessThanSuc_atLeastAtMost) also from ngt1 have "(1+1)*?n*a + d*(1+1)*(∑i∈{1..n - 1}. ?I i) = ((1+1)*?n*a + d*?I (n - 1)*?I n)" by (simp only: mult_ac gauss_sum [of "n - 1"]) (simp add: mult_ac trans [OF add_commute of_nat_Suc [symmetric]]) finally show ?thesis by (simp add: mult_ac add_ac right_distrib) next assume "¬(n > 1)" hence "n = 1 ∨ n = 0" by auto thus ?thesis by (auto simp: mult_ac right_distrib) qed lemma arith_series_nat: "Suc (Suc 0) * (∑i∈{..<n}. a+i*d) = n * (a + (a+(n - 1)*d))" proof - have "((1::nat) + 1) * (∑i∈{..<n::nat}. a + of_nat(i)*d) = of_nat(n) * (a + (a + of_nat(n - 1)*d))" by (rule arith_series_general) thus ?thesis by (auto simp add: of_nat_id) qed lemma arith_series_int: "(2::int) * (∑i∈{..<n}. a + of_nat i * d) = of_nat n * (a + (a + of_nat(n - 1)*d))" proof - have "((1::int) + 1) * (∑i∈{..<n}. a + of_nat i * d) = of_nat(n) * (a + (a + of_nat(n - 1)*d))" by (rule arith_series_general) thus ?thesis by simp qed lemma sum_diff_distrib: fixes P::"nat=>nat" shows "∀x. Q x ≤ P x ==> (∑x<n. P x) - (∑x<n. Q x) = (∑x<n. P x - Q x)" proof (induct n) case 0 show ?case by simp next case (Suc n) let ?lhs = "(∑x<n. P x) - (∑x<n. Q x)" let ?rhs = "∑x<n. P x - Q x" from Suc have "?lhs = ?rhs" by simp moreover from Suc have "?lhs + P n - Q n = ?rhs + (P n - Q n)" by simp moreover from Suc have "(∑x<n. P x) + P n - ((∑x<n. Q x) + Q n) = ?rhs + (P n - Q n)" by (subst diff_diff_left[symmetric], subst diff_add_assoc2) (auto simp: diff_add_assoc2 intro: setsum_mono) ultimately show ?case by simp qed end
lemma lessThan_iff:
(i ∈ {..<k}) = (i < k)
lemma Compl_lessThan:
- {..<k} = {k..}
lemma single_Diff_lessThan:
{k} - {..<k} = {k}
lemma greaterThan_iff:
(i ∈ {k<..}) = (k < i)
lemma Compl_greaterThan:
- {k<..} = {..k}
lemma Compl_atMost:
- {..k} = {k<..}
lemma atLeast_iff:
(i ∈ {k..}) = (k ≤ i)
lemma Compl_atLeast:
- {k..} = {..<k}
lemma atMost_iff:
(i ∈ {..k}) = (i ≤ k)
lemma atMost_Int_atLeast:
{..n} ∩ {n..} = {n}
lemma atLeast_subset_iff:
({x..} ⊆ {y..}) = (y ≤ x)
lemma atLeast_eq_iff:
({x..} = {y..}) = (x = y)
lemma greaterThan_subset_iff:
({x<..} ⊆ {y<..}) = (y ≤ x)
lemma greaterThan_eq_iff:
({x<..} = {y<..}) = (x = y)
lemma atMost_subset_iff:
({..x} ⊆ {..y}) = (x ≤ y)
lemma atMost_eq_iff:
({..x} = {..y}) = (x = y)
lemma lessThan_subset_iff:
({..<x} ⊆ {..<y}) = (x ≤ y)
lemma lessThan_eq_iff:
({..<x} = {..<y}) = (x = y)
lemma greaterThanLessThan_iff:
(i ∈ {l<..<u}) = (l < i ∧ i < u)
lemma atLeastLessThan_iff:
(i ∈ {l..<u}) = (l ≤ i ∧ i < u)
lemma greaterThanAtMost_iff:
(i ∈ {l<..u}) = (l < i ∧ i ≤ u)
lemma atLeastAtMost_iff:
(i ∈ {l..u}) = (l ≤ i ∧ i ≤ u)
lemma atLeastAtMost_empty:
n < m ==> {m..n} = {}
lemma atLeastLessThan_empty:
n ≤ m ==> {m..<n} = {}
lemma greaterThanAtMost_empty:
l ≤ k ==> {k<..l} = {}
lemma greaterThanLessThan_empty:
l ≤ k ==> {k<..l} = {}
lemma atLeastAtMost_singleton:
{a..a} = {a}
lemma lessThan_0:
{..<0} = {}
lemma lessThan_Suc:
{..<Suc k} = insert k {..<k}
lemma lessThan_Suc_atMost:
{..<Suc k} = {..k}
lemma UN_lessThan_UNIV:
(UN m. {..<m}) = UNIV
lemma greaterThan_0:
{0<..} = range Suc
lemma greaterThan_Suc:
{Suc k<..} = {k<..} - {Suc k}
lemma INT_greaterThan_UNIV:
(INT m. {m<..}) = {}
lemma atLeast_0:
{0..} = UNIV
lemma atLeast_Suc:
{Suc k..} = {k..} - {k}
lemma atLeast_Suc_greaterThan:
{Suc k..} = {k<..}
lemma UN_atLeast_UNIV:
(UN m. {m..}) = UNIV
lemma atMost_0:
{..0} = {0}
lemma atMost_Suc:
{..Suc k} = insert (Suc k) {..k}
lemma UN_atMost_UNIV:
(UN m. {..m}) = UNIV
lemma atLeast0LessThan:
{0..<n} = {..<n}
lemma atLeastLessThan0:
{m..<0} = {}
lemma atLeastLessThanSuc:
{m..<Suc n} = (if m ≤ n then insert n {m..<n} else {})
lemma atLeastLessThan_singleton:
{m..<Suc m} = {m}
lemma atLeastLessThanSuc_atLeastAtMost:
{l..<Suc u} = {l..u}
lemma atLeastSucAtMost_greaterThanAtMost:
{Suc l..u} = {l<..u}
lemma atLeastSucLessThan_greaterThanLessThan:
{Suc l..<u} = {l<..<u}
lemma atLeastAtMostSuc_conv:
m ≤ Suc n ==> {m..Suc n} = insert (Suc n) {m..n}
lemma image_add_atLeastAtMost:
(λn. n + k) ` {i..j} = {i + k..j + k}
lemma image_add_atLeastLessThan:
(λn. n + k) ` {i..<j} = {i + k..<j + k}
corollary image_Suc_atLeastAtMost:
Suc ` {i..j} = {Suc i..Suc j}
corollary image_Suc_atLeastLessThan:
Suc ` {i..<j} = {Suc i..<Suc j}
lemma image_add_int_atLeastLessThan:
(λx. x + l) ` {0..<u - l} = {l..<u}
lemma finite_lessThan:
finite {..<k}
lemma finite_atMost:
finite {..k}
lemma finite_greaterThanLessThan:
finite {l<..<u}
lemma finite_atLeastLessThan:
finite {l..<u}
lemma finite_greaterThanAtMost:
finite {l<..u}
lemma finite_atLeastAtMost:
finite {l..u}
lemma bounded_nat_set_is_finite:
∀i∈N. i < n ==> finite N
lemma subset_card_intvl_is_intvl:
A ⊆ {k..<k + card A} ==> A = {k..<k + card A}
lemma card_lessThan:
card {..<u} = u
lemma card_atMost:
card {..u} = Suc u
lemma card_atLeastLessThan:
card {l..<u} = u - l
lemma card_atLeastAtMost:
card {l..u} = Suc u - l
lemma card_greaterThanAtMost:
card {l<..u} = u - l
lemma card_greaterThanLessThan:
card {l<..<u} = u - Suc l
lemma ex_bij_betw_nat_finite:
finite M ==> ∃h. bij_betw h {0..<card M} M
lemma ex_bij_betw_finite_nat:
finite M ==> ∃h. bij_betw h M {0..<card M}
lemma atLeastLessThanPlusOne_atLeastAtMost_int:
{l..<u + 1} = {l..u}
lemma atLeastPlusOneAtMost_greaterThanAtMost_int:
{l + 1..u} = {l<..u}
lemma atLeastPlusOneLessThan_greaterThanLessThan_int:
{l + 1..<u} = {l<..<u}
lemma image_atLeastZeroLessThan_int:
0 ≤ u ==> {0..<u} = int ` {..<nat u}
lemma finite_atLeastZeroLessThan_int:
finite {0..<u}
lemma finite_atLeastLessThan_int:
finite {l..<u}
lemma finite_atLeastAtMost_int:
finite {l..u}
lemma finite_greaterThanAtMost_int:
finite {l<..u}
lemma finite_greaterThanLessThan_int:
finite {l<..<u}
lemma card_atLeastZeroLessThan_int:
card {0..<u} = nat u
lemma card_atLeastLessThan_int:
card {l..<u} = nat (u - l)
lemma card_atLeastAtMost_int:
card {l..u} = nat (u - l + 1)
lemma card_greaterThanAtMost_int:
card {l<..u} = nat (u - l)
lemma card_greaterThanLessThan_int:
card {l<..<u} = nat (u - (l + 1))
lemma ivl_disj_un_singleton:
{l} ∪ {l<..} = {l..}
{..<u} ∪ {u} = {..u}
l < u ==> {l} ∪ {l<..<u} = {l..<u}
l < u ==> {l<..<u} ∪ {u} = {l<..u}
l ≤ u ==> {l} ∪ {l<..u} = {l..u}
l ≤ u ==> {l..<u} ∪ {u} = {l..u}
lemma ivl_disj_un_one:
l < u ==> {..l} ∪ {l<..<u} = {..<u}
l ≤ u ==> {..<l} ∪ {l..<u} = {..<u}
l ≤ u ==> {..l} ∪ {l<..u} = {..u}
l ≤ u ==> {..<l} ∪ {l..u} = {..u}
l ≤ u ==> {l<..u} ∪ {u<..} = {l<..}
l < u ==> {l<..<u} ∪ {u..} = {l<..}
l ≤ u ==> {l..u} ∪ {u<..} = {l..}
l ≤ u ==> {l..<u} ∪ {u..} = {l..}
lemma ivl_disj_un_two:
[| l < m; m ≤ u |] ==> {l<..<m} ∪ {m..<u} = {l<..<u}
[| l ≤ m; m < u |] ==> {l<..m} ∪ {m<..<u} = {l<..<u}
[| l ≤ m; m ≤ u |] ==> {l..<m} ∪ {m..<u} = {l..<u}
[| l ≤ m; m < u |] ==> {l..m} ∪ {m<..<u} = {l..<u}
[| l < m; m ≤ u |] ==> {l<..<m} ∪ {m..u} = {l<..u}
[| l ≤ m; m ≤ u |] ==> {l<..m} ∪ {m<..u} = {l<..u}
[| l ≤ m; m ≤ u |] ==> {l..<m} ∪ {m..u} = {l..u}
[| l ≤ m; m ≤ u |] ==> {l..m} ∪ {m<..u} = {l..u}
lemma ivl_disj_un:
{l} ∪ {l<..} = {l..}
{..<u} ∪ {u} = {..u}
l < u ==> {l} ∪ {l<..<u} = {l..<u}
l < u ==> {l<..<u} ∪ {u} = {l<..u}
l ≤ u ==> {l} ∪ {l<..u} = {l..u}
l ≤ u ==> {l..<u} ∪ {u} = {l..u}
l < u ==> {..l} ∪ {l<..<u} = {..<u}
l ≤ u ==> {..<l} ∪ {l..<u} = {..<u}
l ≤ u ==> {..l} ∪ {l<..u} = {..u}
l ≤ u ==> {..<l} ∪ {l..u} = {..u}
l ≤ u ==> {l<..u} ∪ {u<..} = {l<..}
l < u ==> {l<..<u} ∪ {u..} = {l<..}
l ≤ u ==> {l..u} ∪ {u<..} = {l..}
l ≤ u ==> {l..<u} ∪ {u..} = {l..}
[| l < m; m ≤ u |] ==> {l<..<m} ∪ {m..<u} = {l<..<u}
[| l ≤ m; m < u |] ==> {l<..m} ∪ {m<..<u} = {l<..<u}
[| l ≤ m; m ≤ u |] ==> {l..<m} ∪ {m..<u} = {l..<u}
[| l ≤ m; m < u |] ==> {l..m} ∪ {m<..<u} = {l..<u}
[| l < m; m ≤ u |] ==> {l<..<m} ∪ {m..u} = {l<..u}
[| l ≤ m; m ≤ u |] ==> {l<..m} ∪ {m<..u} = {l<..u}
[| l ≤ m; m ≤ u |] ==> {l..<m} ∪ {m..u} = {l..u}
[| l ≤ m; m ≤ u |] ==> {l..m} ∪ {m<..u} = {l..u}
lemma ivl_disj_int_singleton:
{l} ∩ {l<..} = {}
{..<u} ∩ {u} = {}
{l} ∩ {l<..<u} = {}
{l<..<u} ∩ {u} = {}
{l} ∩ {l<..u} = {}
{l..<u} ∩ {u} = {}
lemma ivl_disj_int_one:
{..l} ∩ {l<..<u} = {}
{..<l} ∩ {l..<u} = {}
{..l} ∩ {l<..u} = {}
{..<l} ∩ {l..u} = {}
{l<..u} ∩ {u<..} = {}
{l<..<u} ∩ {u..} = {}
{l..u} ∩ {u<..} = {}
{l..<u} ∩ {u..} = {}
lemma ivl_disj_int_two:
{l<..<m} ∩ {m..<u} = {}
{l<..m} ∩ {m<..<u} = {}
{l..<m} ∩ {m..<u} = {}
{l..m} ∩ {m<..<u} = {}
{l<..<m} ∩ {m..u} = {}
{l<..m} ∩ {m<..u} = {}
{l..<m} ∩ {m..u} = {}
{l..m} ∩ {m<..u} = {}
lemma ivl_disj_int:
{l} ∩ {l<..} = {}
{..<u} ∩ {u} = {}
{l} ∩ {l<..<u} = {}
{l<..<u} ∩ {u} = {}
{l} ∩ {l<..u} = {}
{l..<u} ∩ {u} = {}
{..l} ∩ {l<..<u} = {}
{..<l} ∩ {l..<u} = {}
{..l} ∩ {l<..u} = {}
{..<l} ∩ {l..u} = {}
{l<..u} ∩ {u<..} = {}
{l<..<u} ∩ {u..} = {}
{l..u} ∩ {u<..} = {}
{l..<u} ∩ {u..} = {}
{l<..<m} ∩ {m..<u} = {}
{l<..m} ∩ {m<..<u} = {}
{l..<m} ∩ {m..<u} = {}
{l..m} ∩ {m<..<u} = {}
{l<..<m} ∩ {m..u} = {}
{l<..m} ∩ {m<..u} = {}
{l..<m} ∩ {m..u} = {}
{l..m} ∩ {m<..u} = {}
lemma ivl_diff:
i ≤ n ==> {i..<m} - {i..<n} = {n..<m}
lemma ivl_subset:
({i..<j} ⊆ {m..<n}) = (j ≤ i ∨ m ≤ i ∧ j ≤ n)
lemma setsum_ivl_cong:
[| a = c; b = d; !!x. [| c ≤ x; x < d |] ==> f x = g x |]
==> setsum f {a..<b} = setsum g {c..<d}
lemma setsum_atMost_Suc:
setsum f {..Suc n} = setsum f {..n} + f (Suc n)
lemma setsum_lessThan_Suc:
setsum f {..<Suc n} = setsum f {..<n} + f n
lemma setsum_cl_ivl_Suc:
setsum f {m..Suc n} = (if Suc n < m then 0::'a else setsum f {m..n} + f (Suc n))
lemma setsum_op_ivl_Suc:
setsum f {m..<Suc n} = (if n < m then 0::'a else setsum f {m..<n} + f n)
lemma setsum_add_nat_ivl:
[| m ≤ n; n ≤ p |] ==> setsum f {m..<n} + setsum f {n..<p} = setsum f {m..<p}
lemma setsum_diff_nat_ivl:
[| m ≤ n; n ≤ p |] ==> setsum f {m..<p} - setsum f {m..<n} = setsum f {n..<p}
lemma setsum_shift_bounds_nat_ivl:
setsum f {m + k..<n + k} = (∑i = m..<n. f (i + k))
lemma setsum_shift_bounds_cl_nat_ivl:
setsum f {m + k..n + k} = (∑i = m..n. f (i + k))
corollary setsum_shift_bounds_cl_Suc_ivl:
setsum f {Suc m..Suc n} = (∑i = m..n. f (Suc i))
corollary setsum_shift_bounds_Suc_ivl:
setsum f {Suc m..<Suc n} = (∑i = m..<n. f (Suc i))
lemma setsum_head:
m ≤ n ==> setsum P {m..n} = P m + setsum P {m<..n}
lemma setsum_head_upt:
0 < m ==> setsum P {..<m} = P 0 + setsum P {1..<m}
lemma geometric_sum:
x ≠ (1::'a) ==> setsum (op ^ x) {0..<n} = (x ^ n - (1::'a)) / (x - (1::'a))
lemma gauss_sum:
((1::'a) + (1::'a)) * setsum of_nat {1..n} = of_nat n * (of_nat n + (1::'a))
theorem arith_series_general:
((1::'a) + (1::'a)) * (∑i<n. a + of_nat i * d) =
of_nat n * (a + (a + of_nat (n - 1) * d))
lemma arith_series_nat:
Suc (Suc 0) * (∑i<n. a + i * d) = n * (a + (a + (n - 1) * d))
lemma arith_series_int:
2 * (∑i<n. a + int i * d) = int n * (a + (a + int (n - 1) * d))
lemma sum_diff_distrib:
∀x. Q x ≤ P x ==> setsum P {..<n} - setsum Q {..<n} = (∑x<n. P x - Q x)